Page MenuHomeIsabelle/Phabricator

No OneTemporary

This file is larger than 256 KB, so syntax highlighting was skipped.
diff --git a/thys/AWN/OClosed_Transfer.thy b/thys/AWN/OClosed_Transfer.thy
--- a/thys/AWN/OClosed_Transfer.thy
+++ b/thys/AWN/OClosed_Transfer.thy
@@ -1,1784 +1,1784 @@
(* Title: OClosed_Transfer.thy
License: BSD 2-Clause. See LICENSE.
Author: Timothy Bourke
*)
section "Transfer open results onto closed models"
theory OClosed_Transfer
imports Closed OClosed_Lifting
begin
locale openproc =
fixes np :: "ip \<Rightarrow> ('s, ('m::msg) seq_action) automaton"
and onp :: "ip \<Rightarrow> ((ip \<Rightarrow> 'g) \<times> 'l, 'm seq_action) automaton"
and sr :: "'s \<Rightarrow> ('g \<times> 'l)"
assumes init: "{ (\<sigma>, \<zeta>) |\<sigma> \<zeta> s. s \<in> init (np i)
\<and> (\<sigma> i, \<zeta>) = sr s
\<and> (\<forall>j. j\<noteq>i \<longrightarrow> \<sigma> j \<in> (fst o sr) ` init (np j)) } \<subseteq> init (onp i)"
and init_notempty: "\<forall>j. init (np j) \<noteq> {}"
and trans: "\<And>s a s' \<sigma> \<sigma>'. \<lbrakk> \<sigma> i = fst (sr s);
\<sigma>' i = fst (sr s');
(s, a, s') \<in> trans (np i) \<rbrakk>
\<Longrightarrow> ((\<sigma>, snd (sr s)), a, (\<sigma>', snd (sr s'))) \<in> trans (onp i)"
begin
lemma init_pnet_p_NodeS:
assumes "NodeS i s R \<in> init (pnet np p)"
shows "p = \<langle>i; R\<rangle>"
using assms by (cases p) (auto simp add: node_comps)
lemma init_pnet_p_SubnetS:
assumes "SubnetS s1 s2 \<in> init (pnet np p)"
obtains p1 p2 where "p = (p1 \<parallel> p2)"
and "s1 \<in> init (pnet np p1)"
and "s2 \<in> init (pnet np p2)"
using assms by (cases p) (auto simp add: node_comps)
lemma init_pnet_fst_sr_netgmap:
assumes "s \<in> init (pnet np p)"
and "i \<in> net_ips s"
and "wf_net_tree p"
shows "the (fst (netgmap sr s) i) \<in> (fst \<circ> sr) ` init (np i)"
using assms proof (induction s arbitrary: p)
fix ii s R\<^sub>i p
assume "NodeS ii s R\<^sub>i \<in> init (pnet np p)"
and "i \<in> net_ips (NodeS ii s R\<^sub>i)"
and "wf_net_tree p"
note this(1)
moreover then have "p = \<langle>ii; R\<^sub>i\<rangle>"
by (rule init_pnet_p_NodeS)
ultimately have "s \<in> init (np ii)"
by (clarsimp simp: node_comps)
with \<open>i \<in> net_ips (NodeS ii s R\<^sub>i)\<close>
show "the (fst (netgmap sr (NodeS ii s R\<^sub>i)) i) \<in> (fst \<circ> sr) ` init (np i)"
by clarsimp
next
fix s1 s2 p
assume IH1: "\<And>p. s1 \<in> init (pnet np p)
\<Longrightarrow> i \<in> net_ips s1
\<Longrightarrow> wf_net_tree p
\<Longrightarrow> the (fst (netgmap sr s1) i) \<in> (fst \<circ> sr) ` init (np i)"
and IH2: "\<And>p. s2 \<in> init (pnet np p)
\<Longrightarrow> i \<in> net_ips s2
\<Longrightarrow> wf_net_tree p
\<Longrightarrow> the (fst (netgmap sr s2) i) \<in> (fst \<circ> sr) ` init (np i)"
and "SubnetS s1 s2 \<in> init (pnet np p)"
and "i \<in> net_ips (SubnetS s1 s2)"
and "wf_net_tree p"
from this(3) obtain p1 p2 where "p = (p1 \<parallel> p2)"
and "s1 \<in> init (pnet np p1)"
and "s2 \<in> init (pnet np p2)"
by (rule init_pnet_p_SubnetS)
from this(1) and \<open>wf_net_tree p\<close> have "wf_net_tree p1"
and "wf_net_tree p2"
and "net_tree_ips p1 \<inter> net_tree_ips p2 = {}"
by auto
from \<open>i \<in> net_ips (SubnetS s1 s2)\<close> have "i \<in> net_ips s1 \<or> i \<in> net_ips s2"
by simp
thus "the (fst (netgmap sr (SubnetS s1 s2)) i) \<in> (fst \<circ> sr) ` init (np i)"
proof
assume "i \<in> net_ips s1"
hence "i \<notin> net_ips s2"
proof -
from \<open>s1 \<in> init (pnet np p1)\<close> and \<open>i \<in> net_ips s1\<close> have "i\<in>net_tree_ips p1" ..
with \<open>net_tree_ips p1 \<inter> net_tree_ips p2 = {}\<close> have "i\<notin>net_tree_ips p2" by auto
with \<open>s2 \<in> init (pnet np p2)\<close> show ?thesis ..
qed
moreover from \<open>s1 \<in> init (pnet np p1)\<close> \<open>i \<in> net_ips s1\<close> and \<open>wf_net_tree p1\<close>
have "the (fst (netgmap sr s1) i) \<in> (fst \<circ> sr) ` init (np i)"
by (rule IH1)
ultimately show ?thesis by simp
next
assume "i \<in> net_ips s2"
moreover with \<open>s2 \<in> init (pnet np p2)\<close> have "the (fst (netgmap sr s2) i) \<in> (fst \<circ> sr) ` init (np i)"
using \<open>wf_net_tree p2\<close> by (rule IH2)
moreover from \<open>s2 \<in> init (pnet np p2)\<close> and \<open>i \<in> net_ips s2\<close> have "i\<in>net_tree_ips p2" ..
ultimately show ?thesis by simp
qed
qed
lemma init_lifted:
assumes "wf_net_tree p"
shows "{ (\<sigma>, snd (netgmap sr s)) |\<sigma> s. s \<in> init (pnet np p)
\<and> (\<forall>i. if i\<in>net_tree_ips p then \<sigma> i = the (fst (netgmap sr s) i)
else \<sigma> i \<in> (fst o sr) ` init (np i)) } \<subseteq> init (opnet onp p)"
using assms proof (induction p)
fix i R
assume "wf_net_tree \<langle>i; R\<rangle>"
show "{(\<sigma>, snd (netgmap sr s)) |\<sigma> s. s \<in> init (pnet np \<langle>i; R\<rangle>)
\<and> (\<forall>j. if j \<in> net_tree_ips \<langle>i; R\<rangle> then \<sigma> j = the (fst (netgmap sr s) j)
else \<sigma> j \<in> (fst \<circ> sr) ` init (np j))} \<subseteq> init (opnet onp \<langle>i; R\<rangle>)"
by (clarsimp simp add: node_comps onode_comps)
(rule subsetD [OF init], auto)
next
fix p1 p2
assume IH1: "wf_net_tree p1
\<Longrightarrow> {(\<sigma>, snd (netgmap sr s)) |\<sigma> s. s \<in> init (pnet np p1)
\<and> (\<forall>i. if i \<in> net_tree_ips p1 then \<sigma> i = the (fst (netgmap sr s) i)
else \<sigma> i \<in> (fst \<circ> sr) ` init (np i))} \<subseteq> init (opnet onp p1)"
(is "_ \<Longrightarrow> ?S1 \<subseteq> _")
and IH2: "wf_net_tree p2
\<Longrightarrow> {(\<sigma>, snd (netgmap sr s)) |\<sigma> s. s \<in> init (pnet np p2)
\<and> (\<forall>i. if i \<in> net_tree_ips p2 then \<sigma> i = the (fst (netgmap sr s) i)
else \<sigma> i \<in> (fst \<circ> sr) ` init (np i))} \<subseteq> init (opnet onp p2)"
(is "_ \<Longrightarrow> ?S2 \<subseteq> _")
and "wf_net_tree (p1 \<parallel> p2)"
from this(3) have "wf_net_tree p1"
and "wf_net_tree p2"
and "net_tree_ips p1 \<inter> net_tree_ips p2 = {}" by auto
show "{(\<sigma>, snd (netgmap sr s)) |\<sigma> s. s \<in> init (pnet np (p1 \<parallel> p2))
\<and> (\<forall>i. if i \<in> net_tree_ips (p1 \<parallel> p2) then \<sigma> i = the (fst (netgmap sr s) i)
else \<sigma> i \<in> (fst \<circ> sr) ` init (np i))} \<subseteq> init (opnet onp (p1 \<parallel> p2))"
proof (rule, clarsimp simp only: split_paired_all pnet.simps automaton.simps)
fix \<sigma> s1 s2
assume \<sigma>_desc: "\<forall>i. if i \<in> net_tree_ips (p1 \<parallel> p2)
then \<sigma> i = the (fst (netgmap sr (SubnetS s1 s2)) i)
else \<sigma> i \<in> (fst \<circ> sr) ` init (np i)"
and "s1 \<in> init (pnet np p1)"
and "s2 \<in> init (pnet np p2)"
from this(2-3) have "net_ips s1 = net_tree_ips p1"
and "net_ips s2 = net_tree_ips p2" by auto
have "(\<sigma>, snd (netgmap sr s1)) \<in> ?S1"
proof -
{ fix i
assume "i \<in> net_tree_ips p1"
with \<open>net_tree_ips p1 \<inter> net_tree_ips p2 = {}\<close> have "i \<notin> net_tree_ips p2" by auto
with \<open>s2 \<in> init (pnet np p2)\<close> have "i \<notin> net_ips s2" ..
hence "the ((fst (netgmap sr s1) ++ fst (netgmap sr s2)) i) = the (fst (netgmap sr s1) i)"
by simp
}
moreover
{ fix i
assume "i \<notin> net_tree_ips p1"
have "\<sigma> i \<in> (fst \<circ> sr) ` init (np i)"
proof (cases "i \<in> net_tree_ips p2")
assume "i \<notin> net_tree_ips p2"
with \<open>i \<notin> net_tree_ips p1\<close> and \<sigma>_desc show ?thesis
by (auto dest: spec [of _ i])
next
assume "i \<in> net_tree_ips p2"
with \<open>s2 \<in> init (pnet np p2)\<close> have "i \<in> net_ips s2" ..
with \<open>s2 \<in> init (pnet np p2)\<close> have "the (fst (netgmap sr s2) i) \<in> (fst \<circ> sr) ` init (np i)"
using \<open>wf_net_tree p2\<close> by (rule init_pnet_fst_sr_netgmap)
with \<open>i\<in>net_tree_ips p2\<close> and \<open>i\<in>net_ips s2\<close> show ?thesis
using \<sigma>_desc by simp
qed
}
ultimately show ?thesis
using \<open>s1 \<in> init (pnet np p1)\<close> and \<sigma>_desc by auto
qed
hence "(\<sigma>, snd (netgmap sr s1)) \<in> init (opnet onp p1)"
by (rule subsetD [OF IH1 [OF \<open>wf_net_tree p1\<close>]])
have "(\<sigma>, snd (netgmap sr s2)) \<in> ?S2"
proof -
{ fix i
assume "i \<in> net_tree_ips p2"
with \<open>s2 \<in> init (pnet np p2)\<close> have "i \<in> net_ips s2" ..
hence "the ((fst (netgmap sr s1) ++ fst (netgmap sr s2)) i) = the (fst (netgmap sr s2) i)"
by simp
}
moreover
{ fix i
assume "i \<notin> net_tree_ips p2"
have "\<sigma> i \<in> (fst \<circ> sr) ` init (np i)"
proof (cases "i \<in> net_tree_ips p1")
assume "i \<notin> net_tree_ips p1"
with \<open>i \<notin> net_tree_ips p2\<close> and \<sigma>_desc show ?thesis
by (auto dest: spec [of _ i])
next
assume "i \<in> net_tree_ips p1"
with \<open>s1 \<in> init (pnet np p1)\<close> have "i \<in> net_ips s1" ..
with \<open>s1 \<in> init (pnet np p1)\<close> have "the (fst (netgmap sr s1) i) \<in> (fst \<circ> sr) ` init (np i)"
using \<open>wf_net_tree p1\<close> by (rule init_pnet_fst_sr_netgmap)
moreover from \<open>s2 \<in> init (pnet np p2)\<close> and \<open>i \<notin> net_tree_ips p2\<close> have "i\<notin>net_ips s2" ..
ultimately show ?thesis
using \<open>i\<in>net_tree_ips p1\<close> \<open>i\<in>net_ips s1\<close> and \<open>i\<notin>net_tree_ips p2\<close> \<sigma>_desc by simp
qed
}
ultimately show ?thesis
using \<open>s2 \<in> init (pnet np p2)\<close> and \<sigma>_desc by auto
qed
hence "(\<sigma>, snd (netgmap sr s2)) \<in> init (opnet onp p2)"
by (rule subsetD [OF IH2 [OF \<open>wf_net_tree p2\<close>]])
with \<open>(\<sigma>, snd (netgmap sr s1)) \<in> init (opnet onp p1)\<close>
show "(\<sigma>, snd (netgmap sr (SubnetS s1 s2))) \<in> init (opnet onp (p1 \<parallel> p2))"
using \<open>net_tree_ips p1 \<inter> net_tree_ips p2 = {}\<close>
\<open>net_ips s1 = net_tree_ips p1\<close>
\<open>net_ips s2 = net_tree_ips p2\<close> by simp
qed
qed
lemma init_pnet_opnet [elim]:
assumes "wf_net_tree p"
and "s \<in> init (pnet np p)"
shows "netgmap sr s \<in> netmask (net_tree_ips p) ` init (opnet onp p)"
proof -
from \<open>wf_net_tree p\<close>
have "{ (\<sigma>, snd (netgmap sr s)) |\<sigma> s. s \<in> init (pnet np p)
\<and> (\<forall>i. if i\<in>net_tree_ips p then \<sigma> i = the (fst (netgmap sr s) i)
else \<sigma> i \<in> (fst o sr) ` init (np i)) } \<subseteq> init (opnet onp p)"
(is "?S \<subseteq> _")
by (rule init_lifted)
hence "netmask (net_tree_ips p) ` ?S \<subseteq> netmask (net_tree_ips p) ` init (opnet onp p)"
by (rule image_mono)
moreover have "netgmap sr s \<in> netmask (net_tree_ips p) ` ?S"
proof -
{ fix i
from init_notempty have "\<exists>s. s \<in> (fst \<circ> sr) ` init (np i)" by auto
hence "(SOME x. x \<in> (fst \<circ> sr) ` init (np i)) \<in> (fst \<circ> sr) ` init (np i)" ..
}
with \<open>s \<in> init (pnet np p)\<close> and init_notempty
have "(\<lambda>i. if i \<in> net_tree_ips p
then the (fst (netgmap sr s) i)
else SOME x. x \<in> (fst \<circ> sr) ` init (np i), snd (netgmap sr s)) \<in> ?S"
(is "?s \<in> ?S") by auto
moreover have "netgmap sr s = netmask (net_tree_ips p) ?s"
proof (intro prod_eqI ext)
fix i
show "fst (netgmap sr s) i = fst (netmask (net_tree_ips p) ?s) i"
proof (cases "i \<in> net_tree_ips p")
assume "i \<in> net_tree_ips p"
with \<open>s\<in>init (pnet np p)\<close> have "i\<in>net_ips s" ..
hence "Some (the (fst (netgmap sr s) i)) = fst (netgmap sr s) i"
by (rule some_the_fst_netgmap)
with \<open>i\<in>net_tree_ips p\<close> show ?thesis
by simp
next
assume "i \<notin> net_tree_ips p"
moreover with \<open>s\<in>init (pnet np p)\<close> have "i\<notin>net_ips s" ..
ultimately show ?thesis
by simp
qed
qed simp
ultimately show ?thesis
by (rule rev_image_eqI)
qed
ultimately show ?thesis
by (rule rev_subsetD [rotated])
qed
lemma transfer_connect:
assumes "(s, connect(i, i'), s') \<in> trans (pnet np n)"
and "s \<in> reachable (pnet np n) TT"
and "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)"
and "wf_net_tree n"
obtains \<sigma>' \<zeta>' where "((\<sigma>, \<zeta>), connect(i, i'), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)"
and "\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
proof atomize_elim
from assms have "((\<sigma>, snd (netgmap sr s)), connect(i, i'), (\<sigma>, snd (netgmap sr s'))) \<in> trans (opnet onp n)
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>, snd (netgmap sr s'))"
proof (induction n arbitrary: s s' \<zeta>)
fix ii R\<^sub>i ns ns' \<zeta>
assume "(ns, connect(i, i'), ns') \<in> trans (pnet np \<langle>ii; R\<^sub>i\<rangle>)"
and "netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)"
from this(1) have "(ns, connect(i, i'), ns') \<in> node_sos (trans (np ii))"
by (simp add: node_comps)
moreover then obtain ni s s' R R' where "ns = NodeS ni s R"
and "ns' = NodeS ni s' R'" ..
ultimately have "(NodeS ni s R, connect(i, i'), NodeS ni s' R') \<in> node_sos (trans (np ii))"
by simp
moreover then have "s' = s" by auto
ultimately have "((\<sigma>, NodeS ni (snd (sr s)) R), connect(i, i'), (\<sigma>, NodeS ni (snd (sr s)) R'))
\<in> onode_sos (trans (onp ii))"
by - (rule node_connectTE', auto intro!: onode_sos.intros [simplified])
with \<open>ns = NodeS ni s R\<close> \<open>ns' = NodeS ni s' R'\<close> \<open>s' = s\<close>
and \<open>netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)\<close>
show "((\<sigma>, snd (netgmap sr ns)), connect(i, i'), (\<sigma>, snd (netgmap sr ns'))) \<in> trans (opnet onp \<langle>ii; R\<^sub>i\<rangle>)
\<and> netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, snd (netgmap sr ns'))"
by (simp add: onode_comps)
next
fix n1 n2 s s' \<zeta>
assume IH1: "\<And>s s' \<zeta>. (s, connect(i, i'), s') \<in> trans (pnet np n1)
\<Longrightarrow> s \<in> reachable (pnet np n1) TT
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n1) (\<sigma>, \<zeta>)
\<Longrightarrow> wf_net_tree n1
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), connect(i, i'), (\<sigma>, snd (netgmap sr s'))) \<in> trans (opnet onp n1)
\<and> netgmap sr s' = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s'))"
and IH2: "\<And>s s' \<zeta>. (s, connect(i, i'), s') \<in> trans (pnet np n2)
\<Longrightarrow> s \<in> reachable (pnet np n2) TT
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n2) (\<sigma>, \<zeta>)
\<Longrightarrow> wf_net_tree n2
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), connect(i, i'), (\<sigma>, snd (netgmap sr s'))) \<in> trans (opnet onp n2)
\<and> netgmap sr s' = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s'))"
and tr: "(s, connect(i, i'), s') \<in> trans (pnet np (n1 \<parallel> n2))"
and sr: "s \<in> reachable (pnet np (n1 \<parallel> n2)) TT"
and nm: "netgmap sr s = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)"
and "wf_net_tree (n1 \<parallel> n2)"
from this(3) have "(s, connect(i, i'), s') \<in> pnet_sos (trans (pnet np n1))
(trans (pnet np n2))"
by simp
then obtain s1 s1' s2 s2' where "s = SubnetS s1 s2"
and "s' = SubnetS s1' s2'"
and "(s1, connect(i, i'), s1') \<in> trans (pnet np n1)"
and "(s2, connect(i, i'), s2') \<in> trans (pnet np n2)"
by (rule partial_connectTE) auto
from this(1) and nm have "netgmap sr (SubnetS s1 s2) = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)"
by simp
from \<open>wf_net_tree (n1 \<parallel> n2)\<close> have "wf_net_tree n1" and "wf_net_tree n2"
and "net_tree_ips n1 \<inter> net_tree_ips n2 = {}" by auto
from sr \<open>s = SubnetS s1 s2\<close> have "s1 \<in> reachable (pnet np n1) TT" by (metis subnet_reachable(1))
hence "net_ips s1 = net_tree_ips n1" by (rule pnet_net_ips_net_tree_ips)
from sr \<open>s = SubnetS s1 s2\<close> have "s2 \<in> reachable (pnet np n2) TT" by (metis subnet_reachable(2))
hence "net_ips s2 = net_tree_ips n2" by (rule pnet_net_ips_net_tree_ips)
from nm \<open>s = SubnetS s1 s2\<close>
have "netgmap sr (SubnetS s1 s2) = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)" by simp
hence "netgmap sr s1 = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1))"
using \<open>net_tree_ips n1 \<inter> net_tree_ips n2 = {}\<close> \<open>net_ips s1 = net_tree_ips n1\<close>
and \<open>net_ips s2 = net_tree_ips n2\<close> by (rule netgmap_subnet_split1)
with \<open>(s1, connect(i, i'), s1') \<in> trans (pnet np n1)\<close>
and \<open>s1 \<in> reachable (pnet np n1) TT\<close>
have "((\<sigma>, snd (netgmap sr s1)), connect(i, i'), (\<sigma>, snd (netgmap sr s1'))) \<in> trans (opnet onp n1)"
and "netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1'))"
using \<open>wf_net_tree n1\<close> unfolding atomize_conj by (rule IH1)
from \<open>netgmap sr (SubnetS s1 s2) = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)\<close>
\<open>net_ips s1 = net_tree_ips n1\<close> and \<open>net_ips s2 = net_tree_ips n2\<close>
have "netgmap sr s2 = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2))"
by (rule netgmap_subnet_split2)
with \<open>(s2, connect(i, i'), s2') \<in> trans (pnet np n2)\<close>
and \<open>s2 \<in> reachable (pnet np n2) TT\<close>
have "((\<sigma>, snd (netgmap sr s2)), connect(i, i'), (\<sigma>, snd (netgmap sr s2'))) \<in> trans (opnet onp n2)"
and "netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2'))"
using \<open>wf_net_tree n2\<close> unfolding atomize_conj by (rule IH2)
have "((\<sigma>, snd (netgmap sr s)), connect(i, i'), (\<sigma>, snd (netgmap sr s')))
\<in> trans (opnet onp (n1 \<parallel> n2))"
proof -
from \<open>((\<sigma>, snd (netgmap sr s1)), connect(i, i'), (\<sigma>, snd (netgmap sr s1'))) \<in> trans (opnet onp n1)\<close>
and \<open>((\<sigma>, snd (netgmap sr s2)), connect(i, i'), (\<sigma>, snd (netgmap sr s2'))) \<in> trans (opnet onp n2)\<close>
have "((\<sigma>, SubnetS (snd (netgmap sr s1)) (snd (netgmap sr s2))), connect(i, i'),
(\<sigma>, SubnetS (snd (netgmap sr s1')) (snd (netgmap sr s2'))))
\<in> opnet_sos (trans (opnet onp n1)) (trans (opnet onp n2))"
by (rule opnet_connect)
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> show ?thesis by simp
qed
moreover from \<open>netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1'))\<close>
\<open>netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2'))\<close>
\<open>s' = SubnetS s1' s2'\<close>
have "netgmap sr s' = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, snd (netgmap sr s'))" ..
ultimately show "((\<sigma>, snd (netgmap sr s)), connect(i, i'), (\<sigma>, snd (netgmap sr s')))
\<in> trans (opnet onp (n1 \<parallel> n2))
\<and> netgmap sr s' = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, snd (netgmap sr s'))" ..
qed
moreover from \<open>netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)\<close> have "\<zeta> = snd (netgmap sr s)" by simp
ultimately show " \<exists>\<sigma>' \<zeta>'. ((\<sigma>, \<zeta>), connect(i, i'), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)
\<and> (\<forall>j. j \<notin> net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j)
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')" by auto
qed
lemma transfer_disconnect:
assumes "(s, disconnect(i, i'), s') \<in> trans (pnet np n)"
and "s \<in> reachable (pnet np n) TT"
and "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)"
and "wf_net_tree n"
obtains \<sigma>' \<zeta>' where "((\<sigma>, \<zeta>), disconnect(i, i'), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)"
and "\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
proof atomize_elim
from assms have "((\<sigma>, snd (netgmap sr s)), disconnect(i, i'), (\<sigma>, snd (netgmap sr s'))) \<in> trans (opnet onp n)
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>, snd (netgmap sr s'))"
proof (induction n arbitrary: s s' \<zeta>)
fix ii R\<^sub>i ns ns' \<zeta>
assume "(ns, disconnect(i, i'), ns') \<in> trans (pnet np \<langle>ii; R\<^sub>i\<rangle>)"
and "netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)"
from this(1) have "(ns, disconnect(i, i'), ns') \<in> node_sos (trans (np ii))"
by (simp add: node_comps)
moreover then obtain ni s s' R R' where "ns = NodeS ni s R"
and "ns' = NodeS ni s' R'" ..
ultimately have "(NodeS ni s R, disconnect(i, i'), NodeS ni s' R') \<in> node_sos (trans (np ii))"
by simp
moreover then have "s' = s" by auto
ultimately have "((\<sigma>, NodeS ni (snd (sr s)) R), disconnect(i, i'), (\<sigma>, NodeS ni (snd (sr s)) R'))
\<in> onode_sos (trans (onp ii))"
by - (rule node_disconnectTE', auto intro!: onode_sos.intros [simplified])
with \<open>ns = NodeS ni s R\<close> \<open>ns' = NodeS ni s' R'\<close> \<open>s' = s\<close>
and \<open>netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)\<close>
show "((\<sigma>, snd (netgmap sr ns)), disconnect(i, i'), (\<sigma>, snd (netgmap sr ns'))) \<in> trans (opnet onp \<langle>ii; R\<^sub>i\<rangle>)
\<and> netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, snd (netgmap sr ns'))"
by (simp add: onode_comps)
next
fix n1 n2 s s' \<zeta>
assume IH1: "\<And>s s' \<zeta>. (s, disconnect(i, i'), s') \<in> trans (pnet np n1)
\<Longrightarrow> s \<in> reachable (pnet np n1) TT
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n1) (\<sigma>, \<zeta>)
\<Longrightarrow> wf_net_tree n1
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), disconnect(i, i'), (\<sigma>, snd (netgmap sr s'))) \<in> trans (opnet onp n1)
\<and> netgmap sr s' = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s'))"
and IH2: "\<And>s s' \<zeta>. (s, disconnect(i, i'), s') \<in> trans (pnet np n2)
\<Longrightarrow> s \<in> reachable (pnet np n2) TT
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n2) (\<sigma>, \<zeta>)
\<Longrightarrow> wf_net_tree n2
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), disconnect(i, i'), (\<sigma>, snd (netgmap sr s'))) \<in> trans (opnet onp n2)
\<and> netgmap sr s' = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s'))"
and tr: "(s, disconnect(i, i'), s') \<in> trans (pnet np (n1 \<parallel> n2))"
and sr: "s \<in> reachable (pnet np (n1 \<parallel> n2)) TT"
and nm: "netgmap sr s = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)"
and "wf_net_tree (n1 \<parallel> n2)"
from this(3) have "(s, disconnect(i, i'), s') \<in> pnet_sos (trans (pnet np n1))
(trans (pnet np n2))"
by simp
then obtain s1 s1' s2 s2' where "s = SubnetS s1 s2"
and "s' = SubnetS s1' s2'"
and "(s1, disconnect(i, i'), s1') \<in> trans (pnet np n1)"
and "(s2, disconnect(i, i'), s2') \<in> trans (pnet np n2)"
by (rule partial_disconnectTE) auto
from this(1) and nm have "netgmap sr (SubnetS s1 s2) = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)"
by simp
from \<open>wf_net_tree (n1 \<parallel> n2)\<close> have "wf_net_tree n1" and "wf_net_tree n2"
and "net_tree_ips n1 \<inter> net_tree_ips n2 = {}" by auto
from sr \<open>s = SubnetS s1 s2\<close> have "s1 \<in> reachable (pnet np n1) TT" by (metis subnet_reachable(1))
hence "net_ips s1 = net_tree_ips n1" by (rule pnet_net_ips_net_tree_ips)
from sr \<open>s = SubnetS s1 s2\<close> have "s2 \<in> reachable (pnet np n2) TT" by (metis subnet_reachable(2))
hence "net_ips s2 = net_tree_ips n2" by (rule pnet_net_ips_net_tree_ips)
from nm \<open>s = SubnetS s1 s2\<close>
have "netgmap sr (SubnetS s1 s2) = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)" by simp
hence "netgmap sr s1 = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1))"
using \<open>net_tree_ips n1 \<inter> net_tree_ips n2 = {}\<close> \<open>net_ips s1 = net_tree_ips n1\<close>
and \<open>net_ips s2 = net_tree_ips n2\<close> by (rule netgmap_subnet_split1)
with \<open>(s1, disconnect(i, i'), s1') \<in> trans (pnet np n1)\<close>
and \<open>s1 \<in> reachable (pnet np n1) TT\<close>
have "((\<sigma>, snd (netgmap sr s1)), disconnect(i, i'), (\<sigma>, snd (netgmap sr s1'))) \<in> trans (opnet onp n1)"
and "netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1'))"
using \<open>wf_net_tree n1\<close> unfolding atomize_conj by (rule IH1)
from \<open>netgmap sr (SubnetS s1 s2) = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)\<close>
\<open>net_ips s1 = net_tree_ips n1\<close> and \<open>net_ips s2 = net_tree_ips n2\<close>
have "netgmap sr s2 = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2))"
by (rule netgmap_subnet_split2)
with \<open>(s2, disconnect(i, i'), s2') \<in> trans (pnet np n2)\<close>
and \<open>s2 \<in> reachable (pnet np n2) TT\<close>
have "((\<sigma>, snd (netgmap sr s2)), disconnect(i, i'), (\<sigma>, snd (netgmap sr s2'))) \<in> trans (opnet onp n2)"
and "netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2'))"
using \<open>wf_net_tree n2\<close> unfolding atomize_conj by (rule IH2)
have "((\<sigma>, snd (netgmap sr s)), disconnect(i, i'), (\<sigma>, snd (netgmap sr s')))
\<in> trans (opnet onp (n1 \<parallel> n2))"
proof -
from \<open>((\<sigma>, snd (netgmap sr s1)), disconnect(i, i'), (\<sigma>, snd (netgmap sr s1'))) \<in> trans (opnet onp n1)\<close>
and \<open>((\<sigma>, snd (netgmap sr s2)), disconnect(i, i'), (\<sigma>, snd (netgmap sr s2'))) \<in> trans (opnet onp n2)\<close>
have "((\<sigma>, SubnetS (snd (netgmap sr s1)) (snd (netgmap sr s2))), disconnect(i, i'),
(\<sigma>, SubnetS (snd (netgmap sr s1')) (snd (netgmap sr s2'))))
\<in> opnet_sos (trans (opnet onp n1)) (trans (opnet onp n2))"
by (rule opnet_disconnect)
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> show ?thesis by simp
qed
moreover from \<open>netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1'))\<close>
\<open>netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2'))\<close>
\<open>s' = SubnetS s1' s2'\<close>
have "netgmap sr s' = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, snd (netgmap sr s'))" ..
ultimately show "((\<sigma>, snd (netgmap sr s)), disconnect(i, i'), (\<sigma>, snd (netgmap sr s')))
\<in> trans (opnet onp (n1 \<parallel> n2))
\<and> netgmap sr s' = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, snd (netgmap sr s'))" ..
qed
moreover from \<open>netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)\<close> have "\<zeta> = snd (netgmap sr s)" by simp
ultimately show "\<exists>\<sigma>' \<zeta>'. ((\<sigma>, \<zeta>), disconnect(i, i'), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)
\<and> (\<forall>j. j \<notin> net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j)
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')" by auto
qed
lemma transfer_tau:
assumes "(s, \<tau>, s') \<in> trans (pnet np n)"
and "s \<in> reachable (pnet np n) TT"
and "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)"
and "wf_net_tree n"
obtains \<sigma>' \<zeta>' where "((\<sigma>, \<zeta>), \<tau>, (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)"
and "\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
proof atomize_elim
from assms(4,2,1) obtain i where "i\<in>net_ips s"
and "\<forall>j. j\<noteq>i \<longrightarrow> netmap s' j = netmap s j"
and "net_ip_action np \<tau> i n s s'"
by (metis pnet_tau_single_node)
from this(2) have "\<forall>j. j\<noteq>i \<longrightarrow> fst (netgmap sr s') j = fst (netgmap sr s) j"
by (clarsimp intro!: netmap_is_fst_netgmap')
from \<open>(s, \<tau>, s') \<in> trans (pnet np n)\<close> have "net_ips s' = net_ips s"
by (rule pnet_maintains_dom [THEN sym])
define \<sigma>' where "\<sigma>' j = (if j = i then the (fst (netgmap sr s') i) else \<sigma> j)" for j
from \<open>\<forall>j. j\<noteq>i \<longrightarrow> fst (netgmap sr s') j = fst (netgmap sr s) j\<close>
and \<open>netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)\<close>
have "\<forall>j. j\<noteq>i \<longrightarrow> \<sigma>' j = \<sigma> j"
unfolding \<sigma>'_def by clarsimp
from assms(2) have "net_ips s = net_tree_ips n"
by (rule pnet_net_ips_net_tree_ips)
from \<open>netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)\<close>
have "\<zeta> = snd (netgmap sr s)" by simp
from \<open>\<forall>j. j\<noteq>i \<longrightarrow> fst (netgmap sr s') j = fst (netgmap sr s) j\<close> \<open>i \<in> net_ips s\<close>
\<open>net_ips s = net_tree_ips n\<close> \<open>net_ips s' = net_ips s\<close>
\<open>netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)\<close>
have "fst (netgmap sr s') = fst (netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s')))"
unfolding \<sigma>'_def [abs_def] by - (rule ext, clarsimp)
hence "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))"
by (rule prod_eqI, simp)
with assms(1, 3)
have "((\<sigma>, snd (netgmap sr s)), \<tau>, (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n)"
using assms(2,4) \<open>i\<in>net_ips s\<close> and \<open>net_ip_action np \<tau> i n s s'\<close>
proof (induction n arbitrary: s s' \<zeta>)
fix ii R\<^sub>i ns ns' \<zeta>
assume "(ns, \<tau>, ns') \<in> trans (pnet np \<langle>ii; R\<^sub>i\<rangle>)"
and nsr: "ns \<in> reachable (pnet np \<langle>ii; R\<^sub>i\<rangle>) TT"
and "netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)"
and "netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>', snd (netgmap sr ns'))"
and "i\<in>net_ips ns"
from this(1) have "(ns, \<tau>, ns') \<in> node_sos (trans (np ii))"
by (simp add: node_comps)
moreover with nsr obtain s s' R R' where "ns = NodeS ii s R"
and "ns' = NodeS ii s' R'"
by (metis net_node_reachable_is_node node_tauTE')
moreover from \<open>i \<in> net_ips ns\<close> and \<open>ns = NodeS ii s R\<close> have "ii = i" by simp
ultimately have ntr: "(NodeS i s R, \<tau>, NodeS i s' R') \<in> node_sos (trans (np i))"
by simp
hence "R' = R" by (metis net_state.inject(1) node_tauTE')
from ntr obtain a where "(s, a, s') \<in> trans (np i)"
and "(\<exists>d. a = \<not>unicast d \<and> d\<notin>R) \<or> (a = \<tau>)"
by (rule node_tauTE') auto
from \<open>netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)\<close> \<open>ns = NodeS ii s R\<close> and \<open>ii = i\<close>
have "\<sigma> i = fst (sr s)" by simp (metis map_upd_Some_unfold)
moreover from \<open>netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>', snd (netgmap sr ns'))\<close>
\<open>ns' = NodeS ii s' R'\<close> and \<open>ii = i\<close>
have "\<sigma>' i = fst (sr s')"
unfolding \<sigma>'_def [abs_def] by clarsimp (hypsubst_thin,
metis (full_types, lifting) fun_upd_same option.sel)
ultimately have "((\<sigma>, snd (sr s)), a, (\<sigma>', snd (sr s'))) \<in> trans (onp i)"
using \<open>(s, a, s') \<in> trans (np i)\<close> by (rule trans)
from \<open>(\<exists>d. a = \<not>unicast d \<and> d\<notin>R) \<or> (a = \<tau>)\<close> \<open>\<forall>j. j\<noteq>i \<longrightarrow> \<sigma>' j = \<sigma> j\<close> \<open>R'=R\<close>
and \<open>((\<sigma>, snd (sr s)), a, (\<sigma>', snd (sr s'))) \<in> trans (onp i)\<close>
have "((\<sigma>, NodeS i (snd (sr s)) R), \<tau>, (\<sigma>', NodeS i (snd (sr s')) R')) \<in> onode_sos (trans (onp i))"
by (metis onode_sos.onode_notucast onode_sos.onode_tau)
with \<open>ns = NodeS ii s R\<close> \<open>ns' = NodeS ii s' R'\<close> \<open>ii = i\<close>
show "((\<sigma>, snd (netgmap sr ns)), \<tau>, (\<sigma>', snd (netgmap sr ns'))) \<in> trans (opnet onp \<langle>ii; R\<^sub>i\<rangle>)"
by (simp add: onode_comps)
next
fix n1 n2 s s' \<zeta>
assume IH1: "\<And>s s' \<zeta>. (s, \<tau>, s') \<in> trans (pnet np n1)
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n1) (\<sigma>, \<zeta>)
\<Longrightarrow> netgmap sr s' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s'))
\<Longrightarrow> s \<in> reachable (pnet np n1) TT
\<Longrightarrow> wf_net_tree n1
\<Longrightarrow> i\<in>net_ips s
\<Longrightarrow> net_ip_action np \<tau> i n1 s s'
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), \<tau>, (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n1)"
and IH2: "\<And>s s' \<zeta>. (s, \<tau>, s') \<in> trans (pnet np n2)
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n2) (\<sigma>, \<zeta>)
\<Longrightarrow> netgmap sr s' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s'))
\<Longrightarrow> s \<in> reachable (pnet np n2) TT
\<Longrightarrow> wf_net_tree n2
\<Longrightarrow> i\<in>net_ips s
\<Longrightarrow> net_ip_action np \<tau> i n2 s s'
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), \<tau>, (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n2)"
and tr: "(s, \<tau>, s') \<in> trans (pnet np (n1 \<parallel> n2))"
and sr: "s \<in> reachable (pnet np (n1 \<parallel> n2)) TT"
and nm: "netgmap sr s = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)"
and nm': "netgmap sr s' = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>', snd (netgmap sr s'))"
and "wf_net_tree (n1 \<parallel> n2)"
and "i\<in>net_ips s"
and "net_ip_action np \<tau> i (n1 \<parallel> n2) s s'"
from tr have "(s, \<tau>, s') \<in> pnet_sos (trans (pnet np n1)) (trans (pnet np n2))" by simp
then obtain s1 s1' s2 s2' where "s = SubnetS s1 s2"
and "s' = SubnetS s1' s2'"
by (rule partial_tauTE) auto
from this(1) and nm have "netgmap sr (SubnetS s1 s2) = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)"
by simp
from \<open>s' = SubnetS s1' s2'\<close> and nm'
have "netgmap sr (SubnetS s1' s2') = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>', snd (netgmap sr s'))"
by simp
from \<open>wf_net_tree (n1 \<parallel> n2)\<close> have "wf_net_tree n1"
and "wf_net_tree n2"
and "net_tree_ips n1 \<inter> net_tree_ips n2 = {}" by auto
from sr [simplified \<open>s = SubnetS s1 s2\<close>] have "s1 \<in> reachable (pnet np n1) TT"
by (rule subnet_reachable(1))
hence "net_ips s1 = net_tree_ips n1" by (rule pnet_net_ips_net_tree_ips)
from sr [simplified \<open>s = SubnetS s1 s2\<close>] have "s2 \<in> reachable (pnet np n2) TT"
by (rule subnet_reachable(2))
hence "net_ips s2 = net_tree_ips n2" by (rule pnet_net_ips_net_tree_ips)
from nm [simplified \<open>s = SubnetS s1 s2\<close>]
\<open>net_tree_ips n1 \<inter> net_tree_ips n2 = {}\<close>
\<open>net_ips s1 = net_tree_ips n1\<close>
\<open>net_ips s2 = net_tree_ips n2\<close>
have "netgmap sr s1 = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1))"
by (rule netgmap_subnet_split1)
from nm [simplified \<open>s = SubnetS s1 s2\<close>]
\<open>net_ips s1 = net_tree_ips n1\<close>
\<open>net_ips s2 = net_tree_ips n2\<close>
have "netgmap sr s2 = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2))"
by (rule netgmap_subnet_split2)
from \<open>i\<in>net_ips s\<close> and \<open>s = SubnetS s1 s2\<close> have "i\<in>net_ips s1 \<or> i\<in>net_ips s2" by auto
thus "((\<sigma>, snd (netgmap sr s)), \<tau>, (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp (n1 \<parallel> n2))"
proof
assume "i\<in>net_ips s1"
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> \<open>net_ip_action np \<tau> i (n1 \<parallel> n2) s s'\<close>
have "(s1, \<tau>, s1') \<in> trans (pnet np n1)"
and "net_ip_action np \<tau> i n1 s1 s1'"
and "s2' = s2" by simp_all
from \<open>net_ips s1 = net_tree_ips n1\<close> and \<open>(s1, \<tau>, s1') \<in> trans (pnet np n1)\<close>
have "net_ips s1' = net_tree_ips n1" by (metis pnet_maintains_dom)
from nm' [simplified \<open>s' = SubnetS s1' s2'\<close> \<open>s2' = s2\<close>]
\<open>net_tree_ips n1 \<inter> net_tree_ips n2 = {}\<close>
\<open>net_ips s1' = net_tree_ips n1\<close>
\<open>net_ips s2 = net_tree_ips n2\<close>
have "netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s1'))"
by (rule netgmap_subnet_split1)
from \<open>(s1, \<tau>, s1') \<in> trans (pnet np n1)\<close>
\<open>netgmap sr s1 = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1))\<close>
\<open>netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s1'))\<close>
\<open>s1 \<in> reachable (pnet np n1) TT\<close>
\<open>wf_net_tree n1\<close>
\<open>i\<in>net_ips s1\<close>
\<open>net_ip_action np \<tau> i n1 s1 s1'\<close>
have "((\<sigma>, snd (netgmap sr s1)), \<tau>, (\<sigma>', snd (netgmap sr s1'))) \<in> trans (opnet onp n1)"
by (rule IH1)
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> \<open>s2' = s2\<close> show ?thesis
by (simp del: step_node_tau) (erule opnet_tau1)
next
assume "i\<in>net_ips s2"
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> \<open>net_ip_action np \<tau> i (n1 \<parallel> n2) s s'\<close>
have "(s2, \<tau>, s2') \<in> trans (pnet np n2)"
and "net_ip_action np \<tau> i n2 s2 s2'"
and "s1' = s1" by simp_all
from \<open>net_ips s2 = net_tree_ips n2\<close> and \<open>(s2, \<tau>, s2') \<in> trans (pnet np n2)\<close>
have "net_ips s2' = net_tree_ips n2" by (metis pnet_maintains_dom)
from nm' [simplified \<open>s' = SubnetS s1' s2'\<close> \<open>s1' = s1\<close>]
\<open>net_ips s1 = net_tree_ips n1\<close>
\<open>net_ips s2' = net_tree_ips n2\<close>
have "netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s2'))"
by (rule netgmap_subnet_split2)
from \<open>(s2, \<tau>, s2') \<in> trans (pnet np n2)\<close>
\<open>netgmap sr s2 = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2))\<close>
\<open>netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s2'))\<close>
\<open>s2 \<in> reachable (pnet np n2) TT\<close>
\<open>wf_net_tree n2\<close>
\<open>i\<in>net_ips s2\<close>
\<open>net_ip_action np \<tau> i n2 s2 s2'\<close>
have "((\<sigma>, snd (netgmap sr s2)), \<tau>, (\<sigma>', snd (netgmap sr s2'))) \<in> trans (opnet onp n2)"
by (rule IH2)
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> \<open>s1' = s1\<close> show ?thesis
by (simp del: step_node_tau) (erule opnet_tau2)
qed
qed
with \<open>\<zeta> = snd (netgmap sr s)\<close> have "((\<sigma>, \<zeta>), \<tau>, (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n)"
by simp
moreover from \<open>\<forall>j. j\<noteq>i \<longrightarrow> \<sigma>' j = \<sigma> j\<close> \<open>i \<in> net_ips s\<close> \<open>\<zeta> = snd (netgmap sr s)\<close>
have "\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j" by (metis net_ips_netgmap)
ultimately have "((\<sigma>, \<zeta>), \<tau>, (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n)
\<and> (\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j)
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))"
using \<open>netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))\<close> by simp
thus "\<exists>\<sigma>' \<zeta>'. ((\<sigma>, \<zeta>), \<tau>, (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)
\<and> (\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j)
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')" by auto
qed
lemma transfer_deliver:
assumes "(s, i:deliver(d), s') \<in> trans (pnet np n)"
and "s \<in> reachable (pnet np n) TT"
and "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)"
and "wf_net_tree n"
obtains \<sigma>' \<zeta>' where "((\<sigma>, \<zeta>), i:deliver(d), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)"
and "\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
proof atomize_elim
from assms(4,2,1) obtain "i\<in>net_ips s"
and "\<forall>j. j\<noteq>i \<longrightarrow> netmap s' j = netmap s j"
and "net_ip_action np (i:deliver(d)) i n s s'"
by (metis delivered_to_net_ips pnet_deliver_single_node)
from this(2) have "\<forall>j. j\<noteq>i \<longrightarrow> fst (netgmap sr s') j = fst (netgmap sr s) j"
by (clarsimp intro!: netmap_is_fst_netgmap')
from \<open>(s, i:deliver(d), s') \<in> trans (pnet np n)\<close> have "net_ips s' = net_ips s"
by (rule pnet_maintains_dom [THEN sym])
define \<sigma>' where "\<sigma>' j = (if j = i then the (fst (netgmap sr s') i) else \<sigma> j)" for j
from \<open>\<forall>j. j\<noteq>i \<longrightarrow> fst (netgmap sr s') j = fst (netgmap sr s) j\<close>
and \<open>netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)\<close>
have "\<forall>j. j\<noteq>i \<longrightarrow> \<sigma>' j = \<sigma> j"
unfolding \<sigma>'_def by clarsimp
from assms(2) have "net_ips s = net_tree_ips n"
by (rule pnet_net_ips_net_tree_ips)
from \<open>netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)\<close>
have "\<zeta> = snd (netgmap sr s)" by simp
from \<open>\<forall>j. j\<noteq>i \<longrightarrow> fst (netgmap sr s') j = fst (netgmap sr s) j\<close> \<open>i \<in> net_ips s\<close>
\<open>net_ips s = net_tree_ips n\<close> \<open>net_ips s' = net_ips s\<close>
\<open>netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)\<close>
have "fst (netgmap sr s') = fst (netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s')))"
unfolding \<sigma>'_def [abs_def] by - (rule ext, clarsimp)
hence "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))"
by (rule prod_eqI, simp)
with assms(1, 3)
have "((\<sigma>, snd (netgmap sr s)), i:deliver(d), (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n)"
using assms(2,4) \<open>i\<in>net_ips s\<close> and \<open>net_ip_action np (i:deliver(d)) i n s s'\<close>
proof (induction n arbitrary: s s' \<zeta>)
fix ii R\<^sub>i ns ns' \<zeta>
assume "(ns, i:deliver(d), ns') \<in> trans (pnet np \<langle>ii; R\<^sub>i\<rangle>)"
and nsr: "ns \<in> reachable (pnet np \<langle>ii; R\<^sub>i\<rangle>) TT"
and "netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)"
and "netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>', snd (netgmap sr ns'))"
and "i\<in>net_ips ns"
from this(1) have "(ns, i:deliver(d), ns') \<in> node_sos (trans (np ii))"
by (simp add: node_comps)
moreover with nsr obtain s s' R R' where "ns = NodeS ii s R"
and "ns' = NodeS ii s' R'"
by (metis net_node_reachable_is_node node_sos_dest)
moreover from \<open>i \<in> net_ips ns\<close> and \<open>ns = NodeS ii s R\<close> have "ii = i" by simp
ultimately have ntr: "(NodeS i s R, i:deliver(d), NodeS i s' R') \<in> node_sos (trans (np i))"
by simp
hence "R' = R" by (metis net_state.inject(1) node_deliverTE')
from ntr have "(s, deliver d, s') \<in> trans (np i)"
by (rule node_deliverTE') simp
from \<open>netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)\<close> \<open>ns = NodeS ii s R\<close> and \<open>ii = i\<close>
have "\<sigma> i = fst (sr s)" by simp (metis map_upd_Some_unfold)
moreover from \<open>netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>', snd (netgmap sr ns'))\<close>
\<open>ns' = NodeS ii s' R'\<close> and \<open>ii = i\<close>
have "\<sigma>' i = fst (sr s')"
unfolding \<sigma>'_def [abs_def] by clarsimp (hypsubst_thin,
metis (lifting, full_types) fun_upd_same option.sel)
ultimately have "((\<sigma>, snd (sr s)), deliver d, (\<sigma>', snd (sr s'))) \<in> trans (onp i)"
using \<open>(s, deliver d, s') \<in> trans (np i)\<close> by (rule trans)
with \<open>\<forall>j. j\<noteq>i \<longrightarrow> \<sigma>' j = \<sigma> j\<close> \<open>R'=R\<close>
have "((\<sigma>, NodeS i (snd (sr s)) R), i:deliver(d), (\<sigma>', NodeS i (snd (sr s')) R'))
\<in> onode_sos (trans (onp i))"
by (metis onode_sos.onode_deliver)
with \<open>ns = NodeS ii s R\<close> \<open>ns' = NodeS ii s' R'\<close> \<open>ii = i\<close>
show "((\<sigma>, snd (netgmap sr ns)), i:deliver(d), (\<sigma>', snd (netgmap sr ns'))) \<in> trans (opnet onp \<langle>ii; R\<^sub>i\<rangle>)"
by (simp add: onode_comps)
next
fix n1 n2 s s' \<zeta>
assume IH1: "\<And>s s' \<zeta>. (s, i:deliver(d), s') \<in> trans (pnet np n1)
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n1) (\<sigma>, \<zeta>)
\<Longrightarrow> netgmap sr s' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s'))
\<Longrightarrow> s \<in> reachable (pnet np n1) TT
\<Longrightarrow> wf_net_tree n1
\<Longrightarrow> i\<in>net_ips s
\<Longrightarrow> net_ip_action np (i:deliver(d)) i n1 s s'
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), i:deliver(d), (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n1)"
and IH2: "\<And>s s' \<zeta>. (s, i:deliver(d), s') \<in> trans (pnet np n2)
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n2) (\<sigma>, \<zeta>)
\<Longrightarrow> netgmap sr s' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s'))
\<Longrightarrow> s \<in> reachable (pnet np n2) TT
\<Longrightarrow> wf_net_tree n2
\<Longrightarrow> i\<in>net_ips s
\<Longrightarrow> net_ip_action np (i:deliver(d)) i n2 s s'
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), i:deliver(d), (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n2)"
and tr: "(s, i:deliver(d), s') \<in> trans (pnet np (n1 \<parallel> n2))"
and sr: "s \<in> reachable (pnet np (n1 \<parallel> n2)) TT"
and nm: "netgmap sr s = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)"
and nm': "netgmap sr s' = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>', snd (netgmap sr s'))"
and "wf_net_tree (n1 \<parallel> n2)"
and "i\<in>net_ips s"
and "net_ip_action np (i:deliver(d)) i (n1 \<parallel> n2) s s'"
from tr have "(s, i:deliver(d), s') \<in> pnet_sos (trans (pnet np n1)) (trans (pnet np n2))" by simp
then obtain s1 s1' s2 s2' where "s = SubnetS s1 s2"
and "s' = SubnetS s1' s2'"
by (rule partial_deliverTE) auto
from this(1) and nm have "netgmap sr (SubnetS s1 s2) = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)"
by simp
from \<open>s' = SubnetS s1' s2'\<close> and nm'
have "netgmap sr (SubnetS s1' s2') = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>', snd (netgmap sr s'))"
by simp
from \<open>wf_net_tree (n1 \<parallel> n2)\<close> have "wf_net_tree n1"
and "wf_net_tree n2"
and "net_tree_ips n1 \<inter> net_tree_ips n2 = {}" by auto
from sr [simplified \<open>s = SubnetS s1 s2\<close>] have "s1 \<in> reachable (pnet np n1) TT"
by (rule subnet_reachable(1))
hence "net_ips s1 = net_tree_ips n1" by (rule pnet_net_ips_net_tree_ips)
from sr [simplified \<open>s = SubnetS s1 s2\<close>] have "s2 \<in> reachable (pnet np n2) TT"
by (rule subnet_reachable(2))
hence "net_ips s2 = net_tree_ips n2" by (rule pnet_net_ips_net_tree_ips)
from nm [simplified \<open>s = SubnetS s1 s2\<close>]
\<open>net_tree_ips n1 \<inter> net_tree_ips n2 = {}\<close>
\<open>net_ips s1 = net_tree_ips n1\<close>
\<open>net_ips s2 = net_tree_ips n2\<close>
have "netgmap sr s1 = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1))"
by (rule netgmap_subnet_split1)
from nm [simplified \<open>s = SubnetS s1 s2\<close>]
\<open>net_ips s1 = net_tree_ips n1\<close>
\<open>net_ips s2 = net_tree_ips n2\<close>
have "netgmap sr s2 = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2))"
by (rule netgmap_subnet_split2)
from \<open>i\<in>net_ips s\<close> and \<open>s = SubnetS s1 s2\<close> have "i\<in>net_ips s1 \<or> i\<in>net_ips s2" by auto
thus "((\<sigma>, snd (netgmap sr s)), i:deliver(d), (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp (n1 \<parallel> n2))"
proof
assume "i\<in>net_ips s1"
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> \<open>net_ip_action np (i:deliver(d)) i (n1 \<parallel> n2) s s'\<close>
have "(s1, i:deliver(d), s1') \<in> trans (pnet np n1)"
and "net_ip_action np (i:deliver(d)) i n1 s1 s1'"
and "s2' = s2" by simp_all
from \<open>net_ips s1 = net_tree_ips n1\<close> and \<open>(s1, i:deliver(d), s1') \<in> trans (pnet np n1)\<close>
have "net_ips s1' = net_tree_ips n1" by (metis pnet_maintains_dom)
from nm' [simplified \<open>s' = SubnetS s1' s2'\<close> \<open>s2' = s2\<close>]
\<open>net_tree_ips n1 \<inter> net_tree_ips n2 = {}\<close>
\<open>net_ips s1' = net_tree_ips n1\<close>
\<open>net_ips s2 = net_tree_ips n2\<close>
have "netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s1'))"
by (rule netgmap_subnet_split1)
from \<open>(s1, i:deliver(d), s1') \<in> trans (pnet np n1)\<close>
\<open>netgmap sr s1 = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1))\<close>
\<open>netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s1'))\<close>
\<open>s1 \<in> reachable (pnet np n1) TT\<close>
\<open>wf_net_tree n1\<close>
\<open>i\<in>net_ips s1\<close>
\<open>net_ip_action np (i:deliver(d)) i n1 s1 s1'\<close>
have "((\<sigma>, snd (netgmap sr s1)), i:deliver(d), (\<sigma>', snd (netgmap sr s1'))) \<in> trans (opnet onp n1)"
by (rule IH1)
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> \<open>s2' = s2\<close> show ?thesis
by simp (erule opnet_deliver1)
next
assume "i\<in>net_ips s2"
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> \<open>net_ip_action np (i:deliver(d)) i (n1 \<parallel> n2) s s'\<close>
have "(s2, i:deliver(d), s2') \<in> trans (pnet np n2)"
and "net_ip_action np (i:deliver(d)) i n2 s2 s2'"
and "s1' = s1" by simp_all
from \<open>net_ips s2 = net_tree_ips n2\<close> and \<open>(s2, i:deliver(d), s2') \<in> trans (pnet np n2)\<close>
have "net_ips s2' = net_tree_ips n2" by (metis pnet_maintains_dom)
from nm' [simplified \<open>s' = SubnetS s1' s2'\<close> \<open>s1' = s1\<close>]
\<open>net_ips s1 = net_tree_ips n1\<close>
\<open>net_ips s2' = net_tree_ips n2\<close>
have "netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s2'))"
by (rule netgmap_subnet_split2)
from \<open>(s2, i:deliver(d), s2') \<in> trans (pnet np n2)\<close>
\<open>netgmap sr s2 = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2))\<close>
\<open>netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s2'))\<close>
\<open>s2 \<in> reachable (pnet np n2) TT\<close>
\<open>wf_net_tree n2\<close>
\<open>i\<in>net_ips s2\<close>
\<open>net_ip_action np (i:deliver(d)) i n2 s2 s2'\<close>
have "((\<sigma>, snd (netgmap sr s2)), i:deliver(d), (\<sigma>', snd (netgmap sr s2'))) \<in> trans (opnet onp n2)"
by (rule IH2)
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> \<open>s1' = s1\<close> show ?thesis
by simp (erule opnet_deliver2)
qed
qed
with \<open>\<zeta> = snd (netgmap sr s)\<close> have "((\<sigma>, \<zeta>), i:deliver(d), (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n)"
by simp
moreover from \<open>\<forall>j. j\<noteq>i \<longrightarrow> \<sigma>' j = \<sigma> j\<close> \<open>i \<in> net_ips s\<close> \<open>\<zeta> = snd (netgmap sr s)\<close>
have "\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j" by (metis net_ips_netgmap)
ultimately have "((\<sigma>, \<zeta>), i:deliver(d), (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n)
\<and> (\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j)
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))"
using \<open>netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))\<close> by simp
thus "\<exists>\<sigma>' \<zeta>'. ((\<sigma>, \<zeta>), i:deliver(d), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)
\<and> (\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j)
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')" by auto
qed
lemma transfer_arrive':
assumes "(s, H\<not>K:arrive(m), s') \<in> trans (pnet np n)"
and "s \<in> reachable (pnet np n) TT"
and "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
and "wf_net_tree n"
shows "((\<sigma>, \<zeta>), H\<not>K:arrive(m), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)"
proof -
from assms(2) have "net_ips s = net_tree_ips n" ..
with assms(1) have "net_ips s' = net_tree_ips n"
by (metis pnet_maintains_dom)
from \<open>netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)\<close>
have "\<zeta> = snd (netgmap sr s)" by simp
from \<open>netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')\<close>
have "\<zeta>' = snd (netgmap sr s')"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))"
by simp_all
from assms(1-3) \<open>netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))\<close> assms(5)
have "((\<sigma>, snd (netgmap sr s)), H\<not>K:arrive(m), (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n)"
proof (induction n arbitrary: s s' \<zeta> H K)
fix ii R\<^sub>i ns ns' \<zeta> H K
assume "(ns, H\<not>K:arrive(m), ns') \<in> trans (pnet np \<langle>ii; R\<^sub>i\<rangle>)"
and nsr: "ns \<in> reachable (pnet np \<langle>ii; R\<^sub>i\<rangle>) TT"
and "netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)"
and "netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>', snd (netgmap sr ns'))"
from this(1) have "(ns, H\<not>K:arrive(m), ns') \<in> node_sos (trans (np ii))"
by (simp add: node_comps)
moreover with nsr obtain s s' R where "ns = NodeS ii s R"
and "ns' = NodeS ii s' R"
by (metis net_node_reachable_is_node node_arriveTE')
ultimately have "(NodeS ii s R, H\<not>K:arrive(m), NodeS ii s' R) \<in> node_sos (trans (np ii))"
by simp
from this(1) have "((\<sigma>, NodeS ii (snd (sr s)) R), H\<not>K:arrive(m), (\<sigma>', NodeS ii (snd (sr s')) R))
\<in> onode_sos (trans (onp ii))"
proof (rule node_arriveTE)
assume "(s, receive m, s') \<in> trans (np ii)"
and "H = {ii}"
and "K = {}"
from \<open>netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)\<close> and \<open>ns = NodeS ii s R\<close>
have "\<sigma> ii = fst (sr s)"
by simp (metis map_upd_Some_unfold)
moreover from \<open>netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>', snd (netgmap sr ns'))\<close>
and \<open>ns' = NodeS ii s' R\<close>
have "\<sigma>' ii = fst (sr s')" by simp (metis map_upd_Some_unfold)
ultimately have "((\<sigma>, snd (sr s)), receive m, (\<sigma>', snd (sr s'))) \<in> trans (onp ii)"
using \<open>(s, receive m, s') \<in> trans (np ii)\<close> by (rule trans)
hence "((\<sigma>, NodeS ii (snd (sr s)) R), {ii}\<not>{}:arrive(m), (\<sigma>', NodeS ii (snd (sr s')) R))
\<in> onode_sos (trans (onp ii))"
by (rule onode_receive)
with \<open>H={ii}\<close> and \<open>K={}\<close>
show "((\<sigma>, NodeS ii (snd (sr s)) R), H\<not>K:arrive(m), (\<sigma>', NodeS ii (snd (sr s')) R))
\<in> onode_sos (trans (onp ii))"
by simp
next
assume "H = {}"
- and "s = s'"
+ and "s' = s"
and "K = {ii}"
- from \<open>s = s'\<close> \<open>netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>', snd (netgmap sr ns'))\<close>
+ from \<open>s' = s\<close> \<open>netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>', snd (netgmap sr ns'))\<close>
\<open>netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)\<close>
\<open>ns = NodeS ii s R\<close> and \<open>ns' = NodeS ii s' R\<close>
have "\<sigma>' ii = \<sigma> ii" by simp (metis option.sel)
hence "((\<sigma>, NodeS ii (snd (sr s)) R), {}\<not>{ii}:arrive(m), (\<sigma>', NodeS ii (snd (sr s)) R))
\<in> onode_sos (trans (onp ii))"
by (rule onode_arrive)
- with \<open>H={}\<close> \<open>K={ii}\<close> and \<open>s = s'\<close>
- show "((\<sigma>, NodeS ii (snd (sr s)) R), H\<not>K:arrive(m), (\<sigma>', NodeS ii (snd (sr s')) R))
- \<in> onode_sos (trans (onp ii))"
- by simp
+ with \<open>H={}\<close> \<open>K={ii}\<close> and \<open>s' = s\<close>
+ show "((\<sigma>, NodeS ii (snd (sr s)) R), H\<not>K:arrive(m), (\<sigma>', NodeS ii (snd (sr s')) R))
+ \<in> onode_sos (trans (onp ii))"
+ by simp
qed
with \<open>ns = NodeS ii s R\<close> \<open>ns' = NodeS ii s' R\<close>
show "((\<sigma>, snd (netgmap sr ns)), H\<not>K:arrive(m), (\<sigma>', snd (netgmap sr ns')))
\<in> trans (opnet onp \<langle>ii; R\<^sub>i\<rangle>)"
by (simp add: onode_comps)
next
fix n1 n2 s s' \<zeta> H K
assume IH1: "\<And>s s' \<zeta> H K. (s, H\<not>K:arrive(m), s') \<in> trans (pnet np n1)
\<Longrightarrow> s \<in> reachable (pnet np n1) TT
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n1) (\<sigma>, \<zeta>)
\<Longrightarrow> netgmap sr s' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s'))
\<Longrightarrow> wf_net_tree n1
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), H\<not>K:arrive(m), \<sigma>', snd (netgmap sr s'))
\<in> trans (opnet onp n1)"
and IH2: "\<And>s s' \<zeta> H K. (s, H\<not>K:arrive(m), s') \<in> trans (pnet np n2)
\<Longrightarrow> s \<in> reachable (pnet np n2) TT
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n2) (\<sigma>, \<zeta>)
\<Longrightarrow> netgmap sr s' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s'))
\<Longrightarrow> wf_net_tree n2
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), H\<not>K:arrive(m), \<sigma>', snd (netgmap sr s'))
\<in> trans (opnet onp n2)"
and "(s, H\<not>K:arrive(m), s') \<in> trans (pnet np (n1 \<parallel> n2))"
and sr: "s \<in> reachable (pnet np (n1 \<parallel> n2)) TT"
and nm: "netgmap sr s = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)"
and nm': "netgmap sr s' = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>', snd (netgmap sr s'))"
and "wf_net_tree (n1 \<parallel> n2)"
from this(3) have "(s, H\<not>K:arrive(m), s') \<in> pnet_sos (trans (pnet np n1))
(trans (pnet np n2))"
by simp
thus "((\<sigma>, snd (netgmap sr s)), H\<not>K:arrive(m), (\<sigma>', snd (netgmap sr s')))
\<in> trans (opnet onp (n1 \<parallel> n2))"
proof (rule partial_arriveTE)
fix s1 s1' s2 s2' H1 H2 K1 K2
assume "s = SubnetS s1 s2"
and "s' = SubnetS s1' s2'"
and tr1: "(s1, H1\<not>K1:arrive(m), s1') \<in> trans (pnet np n1)"
and tr2: "(s2, H2\<not>K2:arrive(m), s2') \<in> trans (pnet np n2)"
and "H = H1 \<union> H2"
and "K = K1 \<union> K2"
from \<open>wf_net_tree (n1 \<parallel> n2)\<close> have "wf_net_tree n1"
and "wf_net_tree n2"
and "net_tree_ips n1 \<inter> net_tree_ips n2 = {}" by auto
from sr [simplified \<open>s = SubnetS s1 s2\<close>] have "s1 \<in> reachable (pnet np n1) TT"
by (rule subnet_reachable(1))
hence "net_ips s1 = net_tree_ips n1" by (rule pnet_net_ips_net_tree_ips)
with tr1 have "net_ips s1' = net_tree_ips n1" by (metis pnet_maintains_dom)
from sr [simplified \<open>s = SubnetS s1 s2\<close>] have "s2 \<in> reachable (pnet np n2) TT"
by (rule subnet_reachable(2))
hence "net_ips s2 = net_tree_ips n2" by (rule pnet_net_ips_net_tree_ips)
with tr2 have "net_ips s2' = net_tree_ips n2" by (metis pnet_maintains_dom)
from \<open>(s1, H1\<not>K1:arrive(m), s1') \<in> trans (pnet np n1)\<close>
\<open>s1 \<in> reachable (pnet np n1) TT\<close>
have "((\<sigma>, snd (netgmap sr s1)), H1\<not>K1:arrive(m), (\<sigma>', snd (netgmap sr s1')))
\<in> trans (opnet onp n1)"
proof (rule IH1 [OF _ _ _ _ \<open>wf_net_tree n1\<close>])
from nm [simplified \<open>s = SubnetS s1 s2\<close>]
\<open>net_tree_ips n1 \<inter> net_tree_ips n2 = {}\<close>
\<open>net_ips s1 = net_tree_ips n1\<close>
\<open>net_ips s2 = net_tree_ips n2\<close>
show "netgmap sr s1 = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1))"
by (rule netgmap_subnet_split1)
next
from nm' [simplified \<open>s' = SubnetS s1' s2'\<close>]
\<open>net_tree_ips n1 \<inter> net_tree_ips n2 = {}\<close>
\<open>net_ips s1' = net_tree_ips n1\<close>
\<open>net_ips s2' = net_tree_ips n2\<close>
show "netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s1'))"
by (rule netgmap_subnet_split1)
qed
moreover from \<open>(s2, H2\<not>K2:arrive(m), s2') \<in> trans (pnet np n2)\<close>
\<open>s2 \<in> reachable (pnet np n2) TT\<close>
have "((\<sigma>, snd (netgmap sr s2)), H2\<not>K2:arrive(m), (\<sigma>', snd (netgmap sr s2')))
\<in> trans (opnet onp n2)"
proof (rule IH2 [OF _ _ _ _ \<open>wf_net_tree n2\<close>])
from nm [simplified \<open>s = SubnetS s1 s2\<close>]
\<open>net_ips s1 = net_tree_ips n1\<close>
\<open>net_ips s2 = net_tree_ips n2\<close>
show "netgmap sr s2 = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2))"
by (rule netgmap_subnet_split2)
next
from nm' [simplified \<open>s' = SubnetS s1' s2'\<close>]
\<open>net_ips s1' = net_tree_ips n1\<close>
\<open>net_ips s2' = net_tree_ips n2\<close>
show "netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s2'))"
by (rule netgmap_subnet_split2)
qed
ultimately show "((\<sigma>, snd (netgmap sr s)), H\<not>K:arrive(m), (\<sigma>', snd (netgmap sr s')))
\<in> trans (opnet onp (n1 \<parallel> n2))"
using \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> \<open>H = H1 \<union> H2\<close> \<open>K = K1 \<union> K2\<close>
by simp (rule opnet_sos.opnet_arrive)
qed
qed
with \<open>\<zeta> = snd (netgmap sr s)\<close> and \<open>\<zeta>' = snd (netgmap sr s')\<close>
show "((\<sigma>, \<zeta>), H\<not>K:arrive(m), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)"
by simp
qed
lemma transfer_arrive:
assumes "(s, H\<not>K:arrive(m), s') \<in> trans (pnet np n)"
and "s \<in> reachable (pnet np n) TT"
and "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)"
and "wf_net_tree n"
obtains \<sigma>' \<zeta>' where "((\<sigma>, \<zeta>), H\<not>K:arrive(m), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)"
and "\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
proof atomize_elim
define \<sigma>' where "\<sigma>' i = (if i\<in>net_tree_ips n then the (fst (netgmap sr s') i) else \<sigma> i)" for i
from assms(2) have "net_ips s = net_tree_ips n"
by (rule pnet_net_ips_net_tree_ips)
with assms(1) have "net_ips s' = net_tree_ips n"
by (metis pnet_maintains_dom)
have "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))"
proof (rule prod_eqI)
from \<open>net_ips s' = net_tree_ips n\<close>
show "fst (netgmap sr s') = fst (netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s')))"
unfolding \<sigma>'_def [abs_def] by - (rule ext, clarsimp)
qed simp
moreover with assms(1-3)
have "((\<sigma>, \<zeta>), H\<not>K:arrive(m), (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n)"
using \<open>wf_net_tree n\<close> by (rule transfer_arrive')
moreover have "\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j"
proof -
have "\<forall>j. j\<notin>net_tree_ips n \<longrightarrow> \<sigma>' j = \<sigma> j" unfolding \<sigma>'_def by simp
with assms(3) and \<open>net_ips s = net_tree_ips n\<close>
show ?thesis
by clarsimp (metis (mono_tags) net_ips_netgmap snd_conv)
qed
ultimately show "\<exists>\<sigma>' \<zeta>'. ((\<sigma>, \<zeta>), H\<not>K:arrive(m), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)
\<and> (\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j)
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')" by auto
qed
lemma transfer_cast:
assumes "(s, mR:*cast(m), s') \<in> trans (pnet np n)"
and "s \<in> reachable (pnet np n) TT"
and "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)"
and "wf_net_tree n"
obtains \<sigma>' \<zeta>' where "((\<sigma>, \<zeta>), mR:*cast(m), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)"
and "\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
proof atomize_elim
define \<sigma>' where "\<sigma>' i = (if i\<in>net_tree_ips n then the (fst (netgmap sr s') i) else \<sigma> i)" for i
from assms(2) have "net_ips s = net_tree_ips n" ..
with assms(1) have "net_ips s' = net_tree_ips n"
by (metis pnet_maintains_dom)
have "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))"
proof (rule prod_eqI)
from \<open>net_ips s' = net_tree_ips n\<close>
show "fst (netgmap sr s') = fst (netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s')))"
unfolding \<sigma>'_def [abs_def] by - (rule ext, clarsimp simp add: some_the_fst_netgmap)
qed simp
from \<open>net_ips s' = net_tree_ips n\<close> and \<open>net_ips s = net_tree_ips n\<close>
have "\<forall>j. j\<notin>net_ips (snd (netgmap sr s)) \<longrightarrow> \<sigma>' j = \<sigma> j"
unfolding \<sigma>'_def by simp
from \<open>netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)\<close>
have "\<zeta> = snd (netgmap sr s)" by simp
from assms(1-3) \<open>netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))\<close> assms(4)
have "((\<sigma>, snd (netgmap sr s)), mR:*cast(m), (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n)"
proof (induction n arbitrary: s s' \<zeta> mR)
fix ii R\<^sub>i ns ns' \<zeta> mR
assume "(ns, mR:*cast(m), ns') \<in> trans (pnet np \<langle>ii; R\<^sub>i\<rangle>)"
and nsr: "ns \<in> reachable (pnet np \<langle>ii; R\<^sub>i\<rangle>) TT"
and "netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)"
and "netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>', snd (netgmap sr ns'))"
from this(1) have "(ns, mR:*cast(m), ns') \<in> node_sos (trans (np ii))"
by (simp add: node_comps)
moreover with nsr obtain s s' R where "ns = NodeS ii s R"
and "ns' = NodeS ii s' R"
by (metis net_node_reachable_is_node node_castTE')
ultimately have "(NodeS ii s R, mR:*cast(m), NodeS ii s' R) \<in> node_sos (trans (np ii))"
by simp
from \<open>netgmap sr ns = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>, \<zeta>)\<close> and \<open>ns = NodeS ii s R\<close>
have "\<sigma> ii = fst (sr s)"
by simp (metis map_upd_Some_unfold)
from \<open>netgmap sr ns' = netmask (net_tree_ips \<langle>ii; R\<^sub>i\<rangle>) (\<sigma>', snd (netgmap sr ns'))\<close>
and \<open>ns' = NodeS ii s' R\<close>
have "\<sigma>' ii = fst (sr s')" by simp (metis map_upd_Some_unfold)
from \<open>(NodeS ii s R, mR:*cast(m), NodeS ii s' R) \<in> node_sos (trans (np ii))\<close>
have "((\<sigma>, NodeS ii (snd (sr s)) R), mR:*cast(m), (\<sigma>', NodeS ii (snd (sr s')) R))
\<in> onode_sos (trans (onp ii))"
proof (rule node_castTE)
assume "(s, broadcast m, s') \<in> trans (np ii)"
- and "R = mR"
+ and "mR = R"
from \<open>\<sigma> ii = fst (sr s)\<close> \<open>\<sigma>' ii = fst (sr s')\<close> and this(1)
have "((\<sigma>, snd (sr s)), broadcast m, (\<sigma>', snd (sr s'))) \<in> trans (onp ii)"
by (rule trans)
hence "((\<sigma>, NodeS ii (snd (sr s)) R), R:*cast(m), (\<sigma>', NodeS ii (snd (sr s')) R))
\<in> onode_sos (trans (onp ii))"
by (rule onode_bcast)
- with \<open>R=mR\<close> show "((\<sigma>, NodeS ii (snd (sr s)) R), mR:*cast(m), (\<sigma>', NodeS ii (snd (sr s')) R))
+ with \<open>mR = R\<close> show "((\<sigma>, NodeS ii (snd (sr s)) R), mR:*cast(m), (\<sigma>', NodeS ii (snd (sr s')) R))
\<in> onode_sos (trans (onp ii))"
by simp
next
fix D
assume "(s, groupcast D m, s') \<in> trans (np ii)"
and "mR = R \<inter> D"
from \<open>\<sigma> ii = fst (sr s)\<close> \<open>\<sigma>' ii = fst (sr s')\<close> and this(1)
have "((\<sigma>, snd (sr s)), groupcast D m, (\<sigma>', snd (sr s'))) \<in> trans (onp ii)"
by (rule trans)
hence "((\<sigma>, NodeS ii (snd (sr s)) R), (R \<inter> D):*cast(m), (\<sigma>', NodeS ii (snd (sr s')) R))
\<in> onode_sos (trans (onp ii))"
by (rule onode_gcast)
with \<open>mR = R \<inter> D\<close> show "((\<sigma>, NodeS ii (snd (sr s)) R), mR:*cast(m), (\<sigma>', NodeS ii (snd (sr s')) R))
\<in> onode_sos (trans (onp ii))"
by simp
next
fix d
assume "(s, unicast d m, s') \<in> trans (np ii)"
and "d \<in> R"
and "mR = {d}"
from \<open>\<sigma> ii = fst (sr s)\<close> \<open>\<sigma>' ii = fst (sr s')\<close> and this(1)
have "((\<sigma>, snd (sr s)), unicast d m, (\<sigma>', snd (sr s'))) \<in> trans (onp ii)"
by (rule trans)
hence "((\<sigma>, NodeS ii (snd (sr s)) R), {d}:*cast(m), (\<sigma>', NodeS ii (snd (sr s')) R))
\<in> onode_sos (trans (onp ii))"
using \<open>d\<in>R\<close> by (rule onode_ucast)
with \<open>mR={d}\<close> show "((\<sigma>, NodeS ii (snd (sr s)) R), mR:*cast(m), (\<sigma>', NodeS ii (snd (sr s')) R))
\<in> onode_sos (trans (onp ii))"
by simp
qed
with \<open>ns = NodeS ii s R\<close> \<open>ns' = NodeS ii s' R\<close>
show "((\<sigma>, snd (netgmap sr ns)), mR:*cast(m), (\<sigma>', snd (netgmap sr ns')))
\<in> trans (opnet onp \<langle>ii; R\<^sub>i\<rangle>)"
by (simp add: onode_comps)
next
fix n1 n2 s s' \<zeta> mR
assume IH1: "\<And>s s' \<zeta> mR. (s, mR:*cast(m), s') \<in> trans (pnet np n1)
\<Longrightarrow> s \<in> reachable (pnet np n1) TT
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n1) (\<sigma>, \<zeta>)
\<Longrightarrow> netgmap sr s' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s'))
\<Longrightarrow> wf_net_tree n1
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), mR:*cast(m), \<sigma>', snd (netgmap sr s'))
\<in> trans (opnet onp n1)"
and IH2: "\<And>s s' \<zeta> mR. (s, mR:*cast(m), s') \<in> trans (pnet np n2)
\<Longrightarrow> s \<in> reachable (pnet np n2) TT
\<Longrightarrow> netgmap sr s = netmask (net_tree_ips n2) (\<sigma>, \<zeta>)
\<Longrightarrow> netgmap sr s' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s'))
\<Longrightarrow> wf_net_tree n2
\<Longrightarrow> ((\<sigma>, snd (netgmap sr s)), mR:*cast(m), \<sigma>', snd (netgmap sr s'))
\<in> trans (opnet onp n2)"
and "(s, mR:*cast(m), s') \<in> trans (pnet np (n1 \<parallel> n2))"
and sr: "s \<in> reachable (pnet np (n1 \<parallel> n2)) TT"
and nm: "netgmap sr s = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>, \<zeta>)"
and nm': "netgmap sr s' = netmask (net_tree_ips (n1 \<parallel> n2)) (\<sigma>', snd (netgmap sr s'))"
and "wf_net_tree (n1 \<parallel> n2)"
from this(3) have "(s, mR:*cast(m), s') \<in> pnet_sos (trans (pnet np n1)) (trans (pnet np n2))"
by simp
then obtain s1 s1' s2 s2' H K
where "s = SubnetS s1 s2"
and "s' = SubnetS s1' s2'"
and "H \<subseteq> mR"
and "K \<inter> mR = {}"
and trtr: "((s1, mR:*cast(m), s1') \<in> trans (pnet np n1)
\<and> (s2, H\<not>K:arrive(m), s2') \<in> trans (pnet np n2))
\<or> ((s1, H\<not>K:arrive(m), s1') \<in> trans (pnet np n1)
\<and> (s2, mR:*cast(m), s2') \<in> trans (pnet np n2))"
by (rule partial_castTE) metis+
from \<open>wf_net_tree (n1 \<parallel> n2)\<close> have "wf_net_tree n1"
and "wf_net_tree n2"
and "net_tree_ips n1 \<inter> net_tree_ips n2 = {}" by auto
from sr [simplified \<open>s = SubnetS s1 s2\<close>] have "s1 \<in> reachable (pnet np n1) TT"
by (rule subnet_reachable(1))
hence "net_ips s1 = net_tree_ips n1" by (rule pnet_net_ips_net_tree_ips)
with trtr have "net_ips s1' = net_tree_ips n1" by (metis pnet_maintains_dom)
from sr [simplified \<open>s = SubnetS s1 s2\<close>] have "s2 \<in> reachable (pnet np n2) TT"
by (rule subnet_reachable(2))
hence "net_ips s2 = net_tree_ips n2" by (rule pnet_net_ips_net_tree_ips)
with trtr have "net_ips s2' = net_tree_ips n2" by (metis pnet_maintains_dom)
from nm [simplified \<open>s = SubnetS s1 s2\<close>]
\<open>net_tree_ips n1 \<inter> net_tree_ips n2 = {}\<close>
\<open>net_ips s1 = net_tree_ips n1\<close>
\<open>net_ips s2 = net_tree_ips n2\<close>
have "netgmap sr s1 = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1))"
by (rule netgmap_subnet_split1)
from nm' [simplified \<open>s' = SubnetS s1' s2'\<close>]
\<open>net_tree_ips n1 \<inter> net_tree_ips n2 = {}\<close>
\<open>net_ips s1' = net_tree_ips n1\<close>
\<open>net_ips s2' = net_tree_ips n2\<close>
have "netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s1'))"
by (rule netgmap_subnet_split1)
from nm [simplified \<open>s = SubnetS s1 s2\<close>]
\<open>net_ips s1 = net_tree_ips n1\<close>
\<open>net_ips s2 = net_tree_ips n2\<close>
have "netgmap sr s2 = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2))"
by (rule netgmap_subnet_split2)
from nm' [simplified \<open>s' = SubnetS s1' s2'\<close>]
\<open>net_ips s1' = net_tree_ips n1\<close>
\<open>net_ips s2' = net_tree_ips n2\<close>
have "netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s2'))"
by (rule netgmap_subnet_split2)
from trtr show "((\<sigma>, snd (netgmap sr s)), mR:*cast(m), (\<sigma>', snd (netgmap sr s')))
\<in> trans (opnet onp (n1 \<parallel> n2))"
proof (elim disjE conjE)
assume "(s1, mR:*cast(m), s1') \<in> trans (pnet np n1)"
and "(s2, H\<not>K:arrive(m), s2') \<in> trans (pnet np n2)"
from \<open>(s1, mR:*cast(m), s1') \<in> trans (pnet np n1)\<close>
\<open>s1 \<in> reachable (pnet np n1) TT\<close>
\<open>netgmap sr s1 = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1))\<close>
\<open>netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s1'))\<close>
\<open>wf_net_tree n1\<close>
have "((\<sigma>, snd (netgmap sr s1)), mR:*cast(m), (\<sigma>', snd (netgmap sr s1'))) \<in> trans (opnet onp n1)"
by (rule IH1)
moreover from \<open>(s2, H\<not>K:arrive(m), s2') \<in> trans (pnet np n2)\<close>
\<open>s2 \<in> reachable (pnet np n2) TT\<close>
\<open>netgmap sr s2 = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2))\<close>
\<open>netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s2'))\<close>
\<open>wf_net_tree n2\<close>
have "((\<sigma>, snd (netgmap sr s2)), H\<not>K:arrive(m), (\<sigma>', snd (netgmap sr s2'))) \<in> trans (opnet onp n2)"
by (rule transfer_arrive')
ultimately have "((\<sigma>, SubnetS (snd (netgmap sr s1)) (snd (netgmap sr s2))), mR:*cast(m),
(\<sigma>', SubnetS (snd (netgmap sr s1')) (snd (netgmap sr s2'))))
\<in> opnet_sos (trans (opnet onp n1)) (trans (opnet onp n2))"
using \<open>H \<subseteq> mR\<close> and \<open>K \<inter> mR = {}\<close> by (rule opnet_sos.intros(1))
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> show ?thesis by simp
next
assume "(s1, H\<not>K:arrive(m), s1') \<in> trans (pnet np n1)"
and "(s2, mR:*cast(m), s2') \<in> trans (pnet np n2)"
from \<open>(s1, H\<not>K:arrive(m), s1') \<in> trans (pnet np n1)\<close>
\<open>s1 \<in> reachable (pnet np n1) TT\<close>
\<open>netgmap sr s1 = netmask (net_tree_ips n1) (\<sigma>, snd (netgmap sr s1))\<close>
\<open>netgmap sr s1' = netmask (net_tree_ips n1) (\<sigma>', snd (netgmap sr s1'))\<close>
\<open>wf_net_tree n1\<close>
have "((\<sigma>, snd (netgmap sr s1)), H\<not>K:arrive(m), (\<sigma>', snd (netgmap sr s1'))) \<in> trans (opnet onp n1)"
by (rule transfer_arrive')
moreover from \<open>(s2, mR:*cast(m), s2') \<in> trans (pnet np n2)\<close>
\<open>s2 \<in> reachable (pnet np n2) TT\<close>
\<open>netgmap sr s2 = netmask (net_tree_ips n2) (\<sigma>, snd (netgmap sr s2))\<close>
\<open>netgmap sr s2' = netmask (net_tree_ips n2) (\<sigma>', snd (netgmap sr s2'))\<close>
\<open>wf_net_tree n2\<close>
have "((\<sigma>, snd (netgmap sr s2)), mR:*cast(m), (\<sigma>', snd (netgmap sr s2'))) \<in> trans (opnet onp n2)"
by (rule IH2)
ultimately have "((\<sigma>, SubnetS (snd (netgmap sr s1)) (snd (netgmap sr s2))), mR:*cast(m),
(\<sigma>', SubnetS (snd (netgmap sr s1')) (snd (netgmap sr s2'))))
\<in> opnet_sos (trans (opnet onp n1)) (trans (opnet onp n2))"
using \<open>H \<subseteq> mR\<close> and \<open>K \<inter> mR = {}\<close> by (rule opnet_sos.intros(2))
with \<open>s = SubnetS s1 s2\<close> \<open>s' = SubnetS s1' s2'\<close> show ?thesis by simp
qed
qed
with \<open>\<zeta> = snd (netgmap sr s)\<close> have "((\<sigma>, \<zeta>), mR:*cast(m), (\<sigma>', snd (netgmap sr s'))) \<in> trans (opnet onp n)"
by simp
moreover from \<open>\<forall>j. j\<notin>net_ips (snd (netgmap sr s)) \<longrightarrow> \<sigma>' j = \<sigma> j\<close> \<open>\<zeta> = snd (netgmap sr s)\<close>
have "\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j" by simp
moreover note \<open>netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', snd (netgmap sr s'))\<close>
ultimately show "\<exists>\<sigma>' \<zeta>'. ((\<sigma>, \<zeta>), mR:*cast(m), (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)
\<and> (\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j)
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
by auto
qed
lemma transfer_pnet_action:
assumes "s \<in> reachable (pnet np n) TT"
and "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)"
and "wf_net_tree n"
and "(s, a, s') \<in> trans (pnet np n)"
obtains \<sigma>' \<zeta>' where "((\<sigma>, \<zeta>), a, (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)"
and "\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
proof atomize_elim
show "\<exists>\<sigma>' \<zeta>'. ((\<sigma>, \<zeta>), a, (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)
\<and> (\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j)
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
proof (cases a)
case node_cast
with assms(4) show ?thesis
by (auto elim!: transfer_cast [OF _ assms(1-3)])
next
case node_deliver
with assms(4) show ?thesis
by (auto elim!: transfer_deliver [OF _ assms(1-3)])
next
case node_arrive
with assms(4) show ?thesis
by (auto elim!: transfer_arrive [OF _ assms(1-3)])
next
case node_connect
with assms(4) show ?thesis
by (auto elim!: transfer_connect [OF _ assms(1-3)])
next
case node_disconnect
with assms(4) show ?thesis
by (auto elim!: transfer_disconnect [OF _ assms(1-3)])
next
case node_newpkt
with assms(4) have False by (metis pnet_never_newpkt)
thus ?thesis ..
next
case node_tau
with assms(4) show ?thesis
by (auto elim!: transfer_tau [OF _ assms(1-3), simplified])
qed
qed
lemma transfer_action_pnet_closed:
assumes "(s, a, s') \<in> trans (closed (pnet np n))"
obtains a' where "(s, a', s') \<in> trans (pnet np n)"
and "\<And>\<sigma> \<zeta> \<sigma>' \<zeta>'. \<lbrakk> ((\<sigma>, \<zeta>), a', (\<sigma>', \<zeta>')) \<in> trans (opnet onp n);
(\<forall>j. j\<notin>net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j) \<rbrakk>
\<Longrightarrow> ((\<sigma>, \<zeta>), a, (\<sigma>', \<zeta>')) \<in> trans (oclosed (opnet onp n))"
proof (atomize_elim)
from assms have "(s, a, s') \<in> cnet_sos (trans (pnet np n))" by simp
thus "\<exists>a'. (s, a', s') \<in> trans (pnet np n)
\<and> (\<forall>\<sigma> \<zeta> \<sigma>' \<zeta>'. ((\<sigma>, \<zeta>), a', (\<sigma>', \<zeta>')) \<in> trans (opnet onp n)
\<longrightarrow> (\<forall>j. j \<notin> net_ips \<zeta> \<longrightarrow> \<sigma>' j = \<sigma> j)
\<longrightarrow> ((\<sigma>, \<zeta>), a, (\<sigma>', \<zeta>')) \<in> trans (oclosed (opnet onp n)))"
proof cases
case (cnet_cast R m) thus ?thesis
by (auto intro!: exI [where x="R:*cast(m)"] dest!: ocnet_cast)
qed (auto intro!: ocnet_sos.intros [simplified])
qed
lemma transfer_action:
assumes "s \<in> reachable (closed (pnet np n)) TT"
and "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)"
and "wf_net_tree n"
and "(s, a, s') \<in> trans (closed (pnet np n))"
obtains \<sigma>' \<zeta>' where "((\<sigma>, \<zeta>), a, (\<sigma>', \<zeta>')) \<in> trans (oclosed (opnet onp n))"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
proof atomize_elim
from assms(1) have "s \<in> reachable (pnet np n) TT" ..
from assms(4)
show "\<exists>\<sigma>' \<zeta>'. ((\<sigma>, \<zeta>), a, (\<sigma>', \<zeta>')) \<in> trans (oclosed (opnet onp n))
\<and> netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
by (cases a)
((elim transfer_action_pnet_closed
transfer_pnet_action [OF \<open>s \<in> reachable (pnet np n) TT\<close> assms(2-3)])?,
(auto intro!: exI)[1])+
qed
lemma pnet_reachable_transfer':
assumes "wf_net_tree n"
and "s \<in> reachable (closed (pnet np n)) TT"
shows "netgmap sr s \<in> netmask (net_tree_ips n) ` oreachable (oclosed (opnet onp n)) (\<lambda>_ _ _. True) U"
(is " _ \<in> ?f ` ?oreachable n")
using assms(2) proof induction
fix s
assume "s \<in> init (closed (pnet np n))"
hence "s \<in> init (pnet np n)" by simp
with \<open>wf_net_tree n\<close> have "netgmap sr s \<in> netmask (net_tree_ips n) ` init (opnet onp n)"
by (rule init_pnet_opnet)
hence "netgmap sr s \<in> netmask (net_tree_ips n) ` init (oclosed (opnet onp n))"
by simp
moreover have "netmask (net_tree_ips n) ` init (oclosed (opnet onp n))
\<subseteq> netmask (net_tree_ips n) ` ?oreachable n"
by (intro image_mono subsetI) (rule oreachable_init)
ultimately show "netgmap sr s \<in> netmask (net_tree_ips n) ` ?oreachable n"
by (rule rev_subsetD)
next
fix s a s'
assume "s \<in> reachable (closed (pnet np n)) TT"
and "netgmap sr s \<in> netmask (net_tree_ips n) ` ?oreachable n"
and "(s, a, s') \<in> trans (closed (pnet np n))"
from this(2) obtain \<sigma> \<zeta> where "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)"
and "(\<sigma>, \<zeta>) \<in> ?oreachable n"
by clarsimp
from \<open>s \<in> reachable (closed (pnet np n)) TT\<close> this(1) \<open>wf_net_tree n\<close>
and \<open>(s, a, s') \<in> trans (closed (pnet np n))\<close>
obtain \<sigma>' \<zeta>' where "((\<sigma>, \<zeta>), a, (\<sigma>', \<zeta>')) \<in> trans (oclosed (opnet onp n))"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
by (rule transfer_action)
from \<open>(\<sigma>, \<zeta>) \<in> ?oreachable n\<close> and this(1) have "(\<sigma>', \<zeta>') \<in> ?oreachable n"
by (rule oreachable_local) simp
with \<open>netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')\<close>
show "netgmap sr s' \<in> netmask (net_tree_ips n) ` ?oreachable n" by (rule image_eqI)
qed
definition
someinit :: "nat \<Rightarrow> 'g"
where
"someinit i \<equiv> SOME x. x \<in> (fst o sr) ` init (np i)"
definition
initmissing :: "((nat \<Rightarrow> 'g option) \<times> 'a) \<Rightarrow> (nat \<Rightarrow> 'g) \<times> 'a"
where
"initmissing \<sigma> = (\<lambda>i. case (fst \<sigma>) i of None \<Rightarrow> someinit i | Some s \<Rightarrow> s, snd \<sigma>)"
lemma initmissing_def':
"initmissing = apfst (default someinit)"
by (auto simp add: initmissing_def default_def)
lemma netmask_initmissing_netgmap:
"netmask (net_ips s) (initmissing (netgmap sr s)) = netgmap sr s"
proof (intro prod_eqI ext)
fix i
show "fst (netmask (net_ips s) (initmissing (netgmap sr s))) i = fst (netgmap sr s) i"
unfolding initmissing_def by (clarsimp split: option.split)
qed (simp add: initmissing_def)
lemma snd_initmissing [simp]:
"snd (initmissing x)= snd x"
unfolding initmissing_def by simp
lemma initmnissing_snd_netgmap [simp]:
assumes "initmissing (netgmap sr s) = (\<sigma>, \<zeta>)"
shows "snd (netgmap sr s) = \<zeta>"
using assms unfolding initmissing_def by simp
lemma in_net_ips_fst_init_missing [simp]:
assumes "i \<in> net_ips s"
shows "fst (initmissing (netgmap sr s)) i = the (fst (netgmap sr s) i)"
using assms unfolding initmissing_def by (clarsimp split: option.split)
lemma not_in_net_ips_fst_init_missing [simp]:
assumes "i \<notin> net_ips s"
shows "fst (initmissing (netgmap sr s)) i = someinit i"
using assms unfolding initmissing_def by (clarsimp split: option.split)
lemma initmissing_oreachable_netmask [elim]:
assumes "initmissing (netgmap sr s) \<in> oreachable (oclosed (opnet onp n)) (\<lambda>_ _ _. True) U"
(is "_ \<in> ?oreachable n")
and "net_ips s = net_tree_ips n"
shows "netgmap sr s \<in> netmask (net_tree_ips n) ` ?oreachable n"
proof -
obtain \<sigma> \<zeta> where "initmissing (netgmap sr s) = (\<sigma>, \<zeta>)" by (metis surj_pair)
with assms(1) have "(\<sigma>, \<zeta>) \<in> ?oreachable n" by simp
have "netgmap sr s = netmask (net_ips s) (\<sigma>, \<zeta>)"
proof (intro prod_eqI ext)
fix i
show "fst (netgmap sr s) i = fst (netmask (net_ips s) (\<sigma>, \<zeta>)) i"
proof (cases "i\<in>net_ips s")
assume "i\<in>net_ips s"
hence "fst (initmissing (netgmap sr s)) i = the (fst (netgmap sr s) i)"
by (rule in_net_ips_fst_init_missing)
moreover from \<open>i\<in>net_ips s\<close> have "Some (the (fst (netgmap sr s) i)) = fst (netgmap sr s) i"
by (rule some_the_fst_netgmap)
ultimately show ?thesis
using \<open>initmissing (netgmap sr s) = (\<sigma>, \<zeta>)\<close> by simp
qed simp
next
from \<open>initmissing (netgmap sr s) = (\<sigma>, \<zeta>)\<close>
show "snd (netgmap sr s) = snd (netmask (net_ips s) (\<sigma>, \<zeta>))"
by simp
qed
with assms(2) have "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)" by simp
moreover from \<open>(\<sigma>, \<zeta>) \<in> ?oreachable n\<close>
have "netmask (net_ips s) (\<sigma>, \<zeta>) \<in> netmask (net_ips s) ` ?oreachable n"
by (rule imageI)
ultimately show ?thesis
by (simp only: assms(2))
qed
lemma pnet_reachable_transfer:
assumes "wf_net_tree n"
and "s \<in> reachable (closed (pnet np n)) TT"
shows "initmissing (netgmap sr s) \<in> oreachable (oclosed (opnet onp n)) (\<lambda>_ _ _. True) U"
(is " _ \<in> ?oreachable n")
using assms(2) proof induction
fix s
assume "s \<in> init (closed (pnet np n))"
hence "s \<in> init (pnet np n)" by simp
from \<open>wf_net_tree n\<close> have "initmissing (netgmap sr s) \<in> init (opnet onp n)"
proof (rule init_lifted [THEN subsetD], intro CollectI exI conjI allI)
show "initmissing (netgmap sr s) = (fst (initmissing (netgmap sr s)), snd (netgmap sr s))"
by (metis snd_initmissing surjective_pairing)
next
from \<open>s \<in> init (pnet np n)\<close> show "s \<in> init (pnet np n)" ..
next
fix i
show "if i \<in> net_tree_ips n
then (fst (initmissing (netgmap sr s))) i = the (fst (netgmap sr s) i)
else (fst (initmissing (netgmap sr s))) i \<in> (fst \<circ> sr) ` init (np i)"
proof (cases "i \<in> net_tree_ips n", simp_all only: if_True if_False)
assume "i \<in> net_tree_ips n"
with \<open>s \<in> init (pnet np n)\<close> have "i \<in> net_ips s" ..
thus "fst (initmissing (netgmap sr s)) i = the (fst (netgmap sr s) i)" by simp
next
assume "i \<notin> net_tree_ips n"
with \<open>s \<in> init (pnet np n)\<close> have "i \<notin> net_ips s" ..
hence "fst (initmissing (netgmap sr s)) i = someinit i" by simp
moreover have "someinit i \<in> (fst \<circ> sr) ` init (np i)"
unfolding someinit_def proof (rule someI_ex)
from init_notempty show "\<exists>x. x \<in> (fst o sr) ` init (np i)" by auto
qed
ultimately show "fst (initmissing (netgmap sr s)) i \<in> (fst \<circ> sr) ` init (np i)"
by simp
qed
qed
hence "initmissing (netgmap sr s) \<in> init (oclosed (opnet onp n))" by simp
thus "initmissing (netgmap sr s) \<in> ?oreachable n" ..
next
fix s a s'
assume "s \<in> reachable (closed (pnet np n)) TT"
and "(s, a, s') \<in> trans (closed (pnet np n))"
and "initmissing (netgmap sr s) \<in> ?oreachable n"
from this(1) have "s \<in> reachable (pnet np n) TT" ..
hence "net_ips s = net_tree_ips n" by (rule pnet_net_ips_net_tree_ips)
with \<open>initmissing (netgmap sr s) \<in> ?oreachable n\<close>
have "netgmap sr s \<in> netmask (net_tree_ips n) ` ?oreachable n"
by (rule initmissing_oreachable_netmask)
obtain \<sigma> \<zeta> where "(\<sigma>, \<zeta>) = initmissing (netgmap sr s)" by (metis surj_pair)
with \<open>initmissing (netgmap sr s) \<in> ?oreachable n\<close>
have "(\<sigma>, \<zeta>) \<in> ?oreachable n" by simp
from \<open>(\<sigma>, \<zeta>) = initmissing (netgmap sr s)\<close> and \<open>net_ips s = net_tree_ips n\<close> [symmetric]
have "netgmap sr s = netmask (net_tree_ips n) (\<sigma>, \<zeta>)"
by (clarsimp simp add: netmask_initmissing_netgmap)
with \<open>s \<in> reachable (closed (pnet np n)) TT\<close>
obtain \<sigma>' \<zeta>' where "((\<sigma>, \<zeta>), a, (\<sigma>', \<zeta>')) \<in> trans (oclosed (opnet onp n))"
and "netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')"
using \<open>wf_net_tree n\<close> and \<open>(s, a, s') \<in> trans (closed (pnet np n))\<close>
by (rule transfer_action)
from \<open>(\<sigma>, \<zeta>) \<in> ?oreachable n\<close> have "net_ips \<zeta> = net_tree_ips n"
by (rule opnet_net_ips_net_tree_ips [OF oclosed_oreachable_oreachable])
with \<open>((\<sigma>, \<zeta>), a, (\<sigma>', \<zeta>')) \<in> trans (oclosed (opnet onp n))\<close>
have "\<forall>j. j\<notin>net_tree_ips n \<longrightarrow> \<sigma>' j = \<sigma> j"
by (clarsimp elim!: ocomplete_no_change)
have "initmissing (netgmap sr s') = (\<sigma>', \<zeta>')"
proof (intro prod_eqI ext)
fix i
from \<open>netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')\<close>
\<open>\<forall>j. j\<notin>net_tree_ips n \<longrightarrow> \<sigma>' j = \<sigma> j\<close>
\<open>(\<sigma>, \<zeta>) = initmissing (netgmap sr s)\<close>
\<open>net_ips s = net_tree_ips n\<close>
show "fst (initmissing (netgmap sr s')) i = fst (\<sigma>', \<zeta>') i"
unfolding initmissing_def by simp
next
from \<open>netgmap sr s' = netmask (net_tree_ips n) (\<sigma>', \<zeta>')\<close>
show "snd (initmissing (netgmap sr s')) = snd (\<sigma>', \<zeta>')" by simp
qed
moreover from \<open>(\<sigma>, \<zeta>) \<in> ?oreachable n\<close> \<open>((\<sigma>, \<zeta>), a, (\<sigma>', \<zeta>')) \<in> trans (oclosed (opnet onp n))\<close>
have "(\<sigma>', \<zeta>') \<in> ?oreachable n"
by (rule oreachable_local) (rule TrueI)
ultimately show "initmissing (netgmap sr s') \<in> ?oreachable n"
by simp
qed
definition
netglobal :: "((nat \<Rightarrow> 'g) \<Rightarrow> bool) \<Rightarrow> 's net_state \<Rightarrow> bool"
where
"netglobal P \<equiv> (\<lambda>s. P (fst (initmissing (netgmap sr s))))"
lemma netglobalsimp [simp]:
"netglobal P s = P (fst (initmissing (netgmap sr s)))"
unfolding netglobal_def by simp
lemma netglobalE [elim]:
assumes "netglobal P s"
and "\<And>\<sigma>. \<lbrakk> P \<sigma>; fst (initmissing (netgmap sr s)) = \<sigma> \<rbrakk> \<Longrightarrow> Q \<sigma>"
shows "netglobal Q s"
using assms by simp
lemma netglobal_weakenE [elim]:
assumes "p \<TTurnstile> netglobal P"
and "\<And>\<sigma>. P \<sigma> \<Longrightarrow> Q \<sigma>"
shows "p \<TTurnstile> netglobal Q"
using assms(1) proof (rule invariant_weakenE)
fix s
assume "netglobal P s"
thus "netglobal Q s"
by (rule netglobalE) (erule assms(2))
qed
lemma close_opnet:
assumes "wf_net_tree n"
and "oclosed (opnet onp n) \<Turnstile> (\<lambda>_ _ _. True, U \<rightarrow>) global P"
shows "closed (pnet np n) \<TTurnstile> netglobal P"
unfolding invariant_def proof
fix s
assume "s \<in> reachable (closed (pnet np n)) TT"
with assms(1)
have "initmissing (netgmap sr s) \<in> oreachable (oclosed (opnet onp n)) (\<lambda>_ _ _. True) U"
by (rule pnet_reachable_transfer)
with assms(2) have "global P (initmissing (netgmap sr s))" ..
thus "netglobal P s" by simp
qed
end
locale openproc_parq =
op?: openproc np onp sr for np :: "ip \<Rightarrow> ('s, ('m::msg) seq_action) automaton" and onp sr
+ fixes qp :: "('t, 'm seq_action) automaton"
assumes init_qp_notempty: "init qp \<noteq> {}"
sublocale openproc_parq \<subseteq> openproc "\<lambda>i. np i \<langle>\<langle> qp"
"\<lambda>i. onp i \<langle>\<langle>\<^bsub>i\<^esub> qp"
"\<lambda>(p, q). (fst (sr p), (snd (sr p), q))"
proof unfold_locales
fix i
show "{ (\<sigma>, \<zeta>) |\<sigma> \<zeta> s. s \<in> init (np i \<langle>\<langle> qp)
\<and> (\<sigma> i, \<zeta>) = ((\<lambda>(p, q). (fst (sr p), (snd (sr p), q))) s)
\<and> (\<forall>j. j\<noteq>i \<longrightarrow> \<sigma> j \<in> (fst o (\<lambda>(p, q). (fst (sr p), (snd (sr p), q))))
` init (np j \<langle>\<langle> qp)) } \<subseteq> init (onp i \<langle>\<langle>\<^bsub>i\<^esub> qp)"
(is "?S \<subseteq> _")
proof
fix s
assume "s \<in> ?S"
then obtain \<sigma> p lq
where "s = (\<sigma>, (snd (sr p), lq))"
and "lq \<in> init qp"
and "p \<in> init (np i)"
and "\<sigma> i = fst (sr p)"
and "\<forall>j. j \<noteq> i \<longrightarrow> \<sigma> j \<in> (fst \<circ> (\<lambda>(p, q). (fst (sr p), snd (sr p), q)))
` (init (np j) \<times> init qp)"
by auto
from this(5) have "\<forall>j. j \<noteq> i \<longrightarrow> \<sigma> j \<in> (fst \<circ> sr) ` init (np j)"
by auto
with \<open>p \<in> init (np i)\<close> and \<open>\<sigma> i = fst (sr p)\<close> have "(\<sigma>, snd (sr p)) \<in> init (onp i)"
by - (rule init [THEN subsetD], auto)
with \<open>lq\<in> init qp\<close> have "((\<sigma>, snd (sr p)), lq) \<in> init (onp i) \<times> init qp"
by simp
hence "(\<sigma>, (snd (sr p), lq)) \<in> extg ` (init (onp i) \<times> init qp)"
by (rule rev_image_eqI) simp
with \<open>s = (\<sigma>, (snd (sr p), lq))\<close> show "s \<in> init (onp i \<langle>\<langle>\<^bsub>i\<^esub> qp)"
by simp
qed
next
fix i s a s' \<sigma> \<sigma>'
assume "\<sigma> i = fst ((\<lambda>(p, q). (fst (sr p), (snd (sr p), q))) s)"
and "\<sigma>' i = fst ((\<lambda>(p, q). (fst (sr p), (snd (sr p), q))) s')"
and "(s, a, s') \<in> trans (np i \<langle>\<langle> qp)"
then obtain p q p' q' where "s = (p, q)"
and "s' = (p', q')"
and "\<sigma> i = fst (sr p)"
and "\<sigma>' i = fst (sr p')"
by (clarsimp split: prod.split_asm)
from this(1-2) and \<open>(s, a, s') \<in> trans (np i \<langle>\<langle> qp)\<close>
have "((p, q), a, (p', q')) \<in> parp_sos (trans (np i)) (trans qp)" by simp
hence "((\<sigma>, (snd (sr p), q)), a, (\<sigma>', (snd (sr p'), q'))) \<in> trans (onp i \<langle>\<langle>\<^bsub>i\<^esub> qp)"
proof cases
assume "q' = q"
and "(p, a, p') \<in> trans (np i)"
and "\<And>m. a \<noteq> receive m"
from \<open>\<sigma> i = fst (sr p)\<close> and \<open>\<sigma>' i = fst (sr p')\<close> this(2)
have "((\<sigma>, snd (sr p)), a, (\<sigma>', snd (sr p'))) \<in> trans (onp i)" by (rule trans)
with \<open>q' = q\<close> and \<open>\<And>m. a \<noteq> receive m\<close>
show "((\<sigma>, snd (sr p), q), a, (\<sigma>', (snd (sr p'), q'))) \<in> trans (onp i \<langle>\<langle>\<^bsub>i\<^esub> qp)"
by (auto elim!: oparleft)
next
assume "p' = p"
and "(q, a, q') \<in> trans qp"
and "\<And>m. a \<noteq> send m"
with \<open>\<sigma> i = fst (sr p)\<close> and \<open>\<sigma>' i = fst (sr p')\<close>
show "((\<sigma>, snd (sr p), q), a, (\<sigma>', (snd (sr p'), q'))) \<in> trans (onp i \<langle>\<langle>\<^bsub>i\<^esub> qp)"
by (auto elim!: oparright)
next
fix m
assume "a = \<tau>"
and "(p, receive m, p') \<in> trans (np i)"
and "(q, send m, q') \<in> trans qp"
from \<open>\<sigma> i = fst (sr p)\<close> and \<open>\<sigma>' i = fst (sr p')\<close> this(2)
have "((\<sigma>, snd (sr p)), receive m, (\<sigma>', snd (sr p'))) \<in> trans (onp i)"
by (rule trans)
with \<open>(q, send m, q') \<in> trans qp\<close> and \<open>a = \<tau>\<close>
show "((\<sigma>, snd (sr p), q), a, (\<sigma>', (snd (sr p'), q'))) \<in> trans (onp i \<langle>\<langle>\<^bsub>i\<^esub> qp)"
by (simp del: step_seq_tau) (rule oparboth)
qed
with \<open>s = (p, q)\<close> \<open>s' = (p', q')\<close>
show "((\<sigma>, snd ((\<lambda>(p, q). (fst (sr p), (snd (sr p), q))) s)), a,
(\<sigma>', snd ((\<lambda>(p, q). (fst (sr p), (snd (sr p), q))) s'))) \<in> trans (onp i \<langle>\<langle>\<^bsub>i\<^esub> qp)"
by simp
next
show "\<forall>j. init (np j \<langle>\<langle> qp) \<noteq> {}"
by (clarsimp simp add: init_notempty init_qp_notempty)
qed
end
diff --git a/thys/Aggregation_Algebras/Minimum_Spanning_Trees.thy b/thys/Aggregation_Algebras/Minimum_Spanning_Trees.thy
--- a/thys/Aggregation_Algebras/Minimum_Spanning_Trees.thy
+++ b/thys/Aggregation_Algebras/Minimum_Spanning_Trees.thy
@@ -1,1070 +1,1070 @@
(* Title: Minimum Spanning Tree Algorithms
Author: Walter Guttmann
Maintainer: Walter Guttmann <walter.guttmann at canterbury.ac.nz>
*)
section \<open>Minimum Spanning Tree Algorithms\<close>
text \<open>
In this theory we prove the total-correctness of Kruskal's and Prim's minimum spanning tree algorithms.
Specifications and algorithms work in Stone-Kleene relation algebras extended by operations for aggregation and minimisation.
The algorithms are implemented in a simple imperative language and their proof uses Hoare Logic.
The correctness proofs are discussed in \cite{Guttmann2016c,Guttmann2018a,Guttmann2018b}.
\<close>
theory Minimum_Spanning_Trees
imports Hoare_Logic Aggregation_Algebras
begin
no_notation
trancl ("(_\<^sup>+)" [1000] 999)
context m_kleene_algebra
begin
subsection \<open>Kruskal's Minimum Spanning Tree Algorithm\<close>
text \<open>
The total-correctness proof of Kruskal's minimum spanning tree algorithm uses the following steps \cite{Guttmann2018b}.
We first establish that the algorithm terminates and constructs a spanning tree.
This is a constructive proof of the existence of a spanning tree; any spanning tree algorithm could be used for this.
We then conclude that a minimum spanning tree exists.
This is necessary to establish the invariant for the actual correctness proof, which shows that Kruskal's algorithm produces a minimum spanning tree.
\<close>
definition "spanning_forest f g \<equiv> forest f \<and> f \<le> --g \<and> components g \<le> forest_components f \<and> regular f"
definition "minimum_spanning_forest f g \<equiv> spanning_forest f g \<and> (\<forall>u . spanning_forest u g \<longrightarrow> sum (f \<sqinter> g) \<le> sum (u \<sqinter> g))"
definition "kruskal_spanning_invariant f g h \<equiv> symmetric g \<and> h = h\<^sup>T \<and> g \<sqinter> --h = h \<and> spanning_forest f (-h \<sqinter> g)"
definition "kruskal_invariant f g h \<equiv> kruskal_spanning_invariant f g h \<and> (\<exists>w . minimum_spanning_forest w g \<and> f \<le> w \<squnion> w\<^sup>T)"
text \<open>
We first show two verification conditions which are used in both correctness proofs.
\<close>
lemma kruskal_vc_1:
assumes "symmetric g"
shows "kruskal_spanning_invariant bot g g"
proof (unfold kruskal_spanning_invariant_def, intro conjI)
show "symmetric g"
using assms by simp
next
show "g = g\<^sup>T"
using assms by simp
next
show "g \<sqinter> --g = g"
using inf.sup_monoid.add_commute selection_closed_id by simp
next
show "spanning_forest bot (-g \<sqinter> g)"
using star.circ_transitive_equal spanning_forest_def by simp
qed
lemma kruskal_vc_2:
assumes "kruskal_spanning_invariant f g h"
and "h \<noteq> bot"
and "card { x . regular x \<and> x \<le> --h } = n"
shows "(minarc h \<le> -forest_components f \<longrightarrow> kruskal_spanning_invariant ((f \<sqinter> -(top * minarc h * f\<^sup>T\<^sup>\<star>)) \<squnion> (f \<sqinter> top * minarc h * f\<^sup>T\<^sup>\<star>)\<^sup>T \<squnion> minarc h) g (h \<sqinter> -minarc h \<sqinter> -minarc h\<^sup>T)
\<and> card { x . regular x \<and> x \<le> --h \<and> x \<le> -minarc h \<and> x \<le> -minarc h\<^sup>T } < n) \<and>
(\<not> minarc h \<le> -forest_components f \<longrightarrow> kruskal_spanning_invariant f g (h \<sqinter> -minarc h \<sqinter> -minarc h\<^sup>T)
\<and> card { x . regular x \<and> x \<le> --h \<and> x \<le> -minarc h \<and> x \<le> -minarc h\<^sup>T } < n)"
proof -
let ?e = "minarc h"
let ?f = "(f \<sqinter> -(top * ?e * f\<^sup>T\<^sup>\<star>)) \<squnion> (f \<sqinter> top * ?e * f\<^sup>T\<^sup>\<star>)\<^sup>T \<squnion> ?e"
let ?h = "h \<sqinter> -?e \<sqinter> -?e\<^sup>T"
let ?F = "forest_components f"
let ?n1 = "card { x . regular x \<and> x \<le> --h }"
let ?n2 = "card { x . regular x \<and> x \<le> --h \<and> x \<le> -?e \<and> x \<le> -?e\<^sup>T }"
have 1: "regular f \<and> regular ?e"
by (metis assms(1) kruskal_spanning_invariant_def spanning_forest_def minarc_regular)
hence 2: "regular ?f \<and> regular ?F \<and> regular (?e\<^sup>T)"
using regular_closed_star regular_conv_closed regular_mult_closed by simp
have 3: "\<not> ?e \<le> -?e"
using assms(2) inf.orderE minarc_bot_iff by fastforce
have "?n2 < ?n1"
apply (rule psubset_card_mono)
using finite_regular apply simp
using 1 3 kruskal_spanning_invariant_def minarc_below by auto
hence 4: "?n2 < n"
using assms(3) by simp
show "(?e \<le> -?F \<longrightarrow> kruskal_spanning_invariant ?f g ?h \<and> ?n2 < n) \<and> (\<not> ?e \<le> -?F \<longrightarrow> kruskal_spanning_invariant f g ?h \<and> ?n2 < n)"
proof (rule conjI)
have 5: "injective ?f"
apply (rule kruskal_injective_inv)
using assms(1) kruskal_spanning_invariant_def spanning_forest_def apply simp
apply (simp add: covector_mult_closed)
apply (simp add: comp_associative comp_isotone star.right_plus_below_circ)
apply (meson mult_left_isotone order_lesseq_imp star_outer_increasing top.extremum)
using assms(1,2) kruskal_spanning_invariant_def kruskal_injective_inv_2 minarc_arc spanning_forest_def apply simp
using assms(2) arc_injective minarc_arc apply blast
using assms(1,2) kruskal_spanning_invariant_def kruskal_injective_inv_3 minarc_arc spanning_forest_def by simp
show "?e \<le> -?F \<longrightarrow> kruskal_spanning_invariant ?f g ?h \<and> ?n2 < n"
proof
assume 6: "?e \<le> -?F"
have 7: "equivalence ?F"
using assms(1) kruskal_spanning_invariant_def forest_components_equivalence spanning_forest_def by simp
have "?e\<^sup>T * top * ?e\<^sup>T = ?e\<^sup>T"
using assms(2) by (simp add: arc_top_arc minarc_arc)
hence "?e\<^sup>T * top * ?e\<^sup>T \<le> -?F"
using 6 7 conv_complement conv_isotone by fastforce
hence 8: "?e * ?F * ?e = bot"
using le_bot triple_schroeder_p by simp
show "kruskal_spanning_invariant ?f g ?h \<and> ?n2 < n"
proof (unfold kruskal_spanning_invariant_def, intro conjI)
show "symmetric g"
using assms(1) kruskal_spanning_invariant_def by simp
next
show "?h = ?h\<^sup>T"
using assms(1) by (simp add: conv_complement conv_dist_inf inf_commute inf_left_commute kruskal_spanning_invariant_def)
next
show "g \<sqinter> --?h = ?h"
using 1 2 by (metis (hide_lams) assms(1) kruskal_spanning_invariant_def inf_assoc pp_dist_inf)
next
show "spanning_forest ?f (-?h \<sqinter> g)"
proof (unfold spanning_forest_def, intro conjI)
show "injective ?f"
using 5 by simp
next
show "acyclic ?f"
apply (rule kruskal_acyclic_inv)
using assms(1) kruskal_spanning_invariant_def spanning_forest_def apply simp
apply (simp add: covector_mult_closed)
using 8 assms(1) kruskal_spanning_invariant_def spanning_forest_def kruskal_acyclic_inv_1 apply simp
using 8 apply (metis comp_associative mult_left_sub_dist_sup_left star.circ_loop_fixpoint sup_commute le_bot)
using 6 by (simp add: p_antitone_iff)
next
show "?f \<le> --(-?h \<sqinter> g)"
apply (rule kruskal_subgraph_inv)
using assms(1) kruskal_spanning_invariant_def spanning_forest_def apply simp
using assms(1) apply (metis kruskal_spanning_invariant_def minarc_below order.trans pp_isotone_inf)
using assms(1) kruskal_spanning_invariant_def apply simp
using assms(1) kruskal_spanning_invariant_def by simp
next
show "components (-?h \<sqinter> g) \<le> forest_components ?f"
apply (rule kruskal_spanning_inv)
using 5 apply simp
using 1 regular_closed_star regular_conv_closed regular_mult_closed apply simp
using 1 apply simp
using assms(1) kruskal_spanning_invariant_def spanning_forest_def by simp
next
show "regular ?f"
using 2 by simp
qed
next
show "?n2 < n"
using 4 by simp
qed
qed
next
show "\<not> ?e \<le> -?F \<longrightarrow> kruskal_spanning_invariant f g ?h \<and> ?n2 < n"
proof
assume "\<not> ?e \<le> -?F"
hence 9: "?e \<le> ?F"
using 2 assms(2) arc_in_partition minarc_arc by fastforce
show "kruskal_spanning_invariant f g ?h \<and> ?n2 < n"
proof (unfold kruskal_spanning_invariant_def, intro conjI)
show "symmetric g"
using assms(1) kruskal_spanning_invariant_def by simp
next
show "?h = ?h\<^sup>T"
using assms(1) by (simp add: conv_complement conv_dist_inf inf_commute inf_left_commute kruskal_spanning_invariant_def)
next
show "g \<sqinter> --?h = ?h"
using 1 2 by (metis (hide_lams) assms(1) kruskal_spanning_invariant_def inf_assoc pp_dist_inf)
next
show "spanning_forest f (-?h \<sqinter> g)"
proof (unfold spanning_forest_def, intro conjI)
show "injective f"
using assms(1) kruskal_spanning_invariant_def spanning_forest_def by simp
next
show "acyclic f"
using assms(1) kruskal_spanning_invariant_def spanning_forest_def by simp
next
have "f \<le> --(-h \<sqinter> g)"
using assms(1) kruskal_spanning_invariant_def spanning_forest_def by simp
also have "... \<le> --(-?h \<sqinter> g)"
using comp_inf.mult_right_isotone inf.sup_monoid.add_commute inf_left_commute p_antitone_inf pp_isotone by presburger
finally show "f \<le> --(-?h \<sqinter> g)"
by simp
next
show "components (-?h \<sqinter> g) \<le> ?F"
apply (rule kruskal_spanning_inv_1)
using 9 apply simp
using 1 apply simp
using assms(1) kruskal_spanning_invariant_def spanning_forest_def apply simp
using assms(1) kruskal_spanning_invariant_def forest_components_equivalence spanning_forest_def by simp
next
show "regular f"
using 1 by simp
qed
next
show "?n2 < n"
using 4 by simp
qed
qed
qed
qed
text \<open>
The following result shows that Kruskal's algorithm terminates and constructs a spanning tree.
We cannot yet show that this is a minimum spanning tree.
\<close>
theorem kruskal_spanning:
"VARS e f h
[ symmetric g ]
f := bot;
h := g;
WHILE h \<noteq> bot
INV { kruskal_spanning_invariant f g h }
VAR { card { x . regular x \<and> x \<le> --h } }
DO e := minarc h;
IF e \<le> -forest_components f THEN
f := (f \<sqinter> -(top * e * f\<^sup>T\<^sup>\<star>)) \<squnion> (f \<sqinter> top * e * f\<^sup>T\<^sup>\<star>)\<^sup>T \<squnion> e
ELSE
SKIP
FI;
h := h \<sqinter> -e \<sqinter> -e\<^sup>T
OD
[ spanning_forest f g ]"
apply vcg_tc_simp
using kruskal_vc_1 apply simp
using kruskal_vc_2 apply blast
using kruskal_spanning_invariant_def by auto
text \<open>
Because we have shown total correctness, we conclude that a spanning tree exists.
\<close>
lemma kruskal_exists_spanning:
"symmetric g \<Longrightarrow> \<exists>f . spanning_forest f g"
using tc_extract_function kruskal_spanning by blast
text \<open>
This implies that a minimum spanning tree exists, which is used in the subsequent correctness proof.
\<close>
lemma kruskal_exists_minimal_spanning:
assumes "symmetric g"
shows "\<exists>f . minimum_spanning_forest f g"
proof -
let ?s = "{ f . spanning_forest f g }"
have "\<exists>m\<in>?s . \<forall>z\<in>?s . sum (m \<sqinter> g) \<le> sum (z \<sqinter> g)"
apply (rule finite_set_minimal)
using finite_regular spanning_forest_def apply simp
using assms kruskal_exists_spanning apply simp
using sum_linear by simp
thus ?thesis
using minimum_spanning_forest_def by simp
qed
text \<open>
Kruskal's minimum spanning tree algorithm terminates and is correct.
This is the same algorithm that is used in the previous correctness proof, with the same precondition and variant, but with a different invariant and postcondition.
\<close>
theorem kruskal:
"VARS e f h
[ symmetric g ]
f := bot;
h := g;
WHILE h \<noteq> bot
INV { kruskal_invariant f g h }
VAR { card { x . regular x \<and> x \<le> --h } }
DO e := minarc h;
IF e \<le> -forest_components f THEN
f := (f \<sqinter> -(top * e * f\<^sup>T\<^sup>\<star>)) \<squnion> (f \<sqinter> top * e * f\<^sup>T\<^sup>\<star>)\<^sup>T \<squnion> e
ELSE
SKIP
FI;
h := h \<sqinter> -e \<sqinter> -e\<^sup>T
OD
[ minimum_spanning_forest f g ]"
-proof vcg_tc_simp
+using [[simproc del: defined_all]] proof vcg_tc_simp
assume "symmetric g"
thus "kruskal_invariant bot g g"
using kruskal_vc_1 kruskal_exists_minimal_spanning kruskal_invariant_def by simp
next
fix n f h
let ?e = "minarc h"
let ?f = "(f \<sqinter> -(top * ?e * f\<^sup>T\<^sup>\<star>)) \<squnion> (f \<sqinter> top * ?e * f\<^sup>T\<^sup>\<star>)\<^sup>T \<squnion> ?e"
let ?h = "h \<sqinter> -?e \<sqinter> -?e\<^sup>T"
let ?F = "forest_components f"
let ?n1 = "card { x . regular x \<and> x \<le> --h }"
let ?n2 = "card { x . regular x \<and> x \<le> --h \<and> x \<le> -?e \<and> x \<le> -?e\<^sup>T }"
assume 1: "kruskal_invariant f g h \<and> h \<noteq> bot \<and> ?n1 = n"
from 1 obtain w where 2: "minimum_spanning_forest w g \<and> f \<le> w \<squnion> w\<^sup>T"
using kruskal_invariant_def by auto
hence 3: "regular f \<and> regular w \<and> regular ?e"
using 1 by (metis kruskal_invariant_def kruskal_spanning_invariant_def minimum_spanning_forest_def spanning_forest_def minarc_regular)
show "(?e \<le> -?F \<longrightarrow> kruskal_invariant ?f g ?h \<and> ?n2 < n) \<and> (\<not> ?e \<le> -?F \<longrightarrow> kruskal_invariant f g ?h \<and> ?n2 < n)"
proof (rule conjI)
show "?e \<le> -?F \<longrightarrow> kruskal_invariant ?f g ?h \<and> ?n2 < n"
proof
assume 4: "?e \<le> -?F"
have 5: "equivalence ?F"
using 1 kruskal_invariant_def kruskal_spanning_invariant_def forest_components_equivalence spanning_forest_def by simp
have "?e\<^sup>T * top * ?e\<^sup>T = ?e\<^sup>T"
using 1 by (simp add: arc_top_arc minarc_arc)
hence "?e\<^sup>T * top * ?e\<^sup>T \<le> -?F"
using 4 5 conv_complement conv_isotone by fastforce
hence 6: "?e * ?F * ?e = bot"
using le_bot triple_schroeder_p by simp
show "kruskal_invariant ?f g ?h \<and> ?n2 < n"
proof (unfold kruskal_invariant_def, intro conjI)
show "kruskal_spanning_invariant ?f g ?h"
using 1 4 kruskal_vc_2 kruskal_invariant_def by simp
next
show "\<exists>w . minimum_spanning_forest w g \<and> ?f \<le> w \<squnion> w\<^sup>T"
proof
let ?p = "w \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>"
let ?v = "(w \<sqinter> -(top * ?e * w\<^sup>T\<^sup>\<star>)) \<squnion> ?p\<^sup>T"
have 7: "regular ?p"
using 3 regular_closed_star regular_conv_closed regular_mult_closed by simp
have 8: "injective ?v"
apply (rule kruskal_exchange_injective_inv_1)
using 2 minimum_spanning_forest_def spanning_forest_def apply simp
apply (simp add: covector_mult_closed)
apply (simp add: comp_associative comp_isotone star.right_plus_below_circ)
using 1 2 kruskal_injective_inv_3 minarc_arc minimum_spanning_forest_def spanning_forest_def by simp
have 9: "components g \<le> forest_components ?v"
apply (rule kruskal_exchange_spanning_inv_1)
using 8 apply simp
using 7 apply simp
using 2 minimum_spanning_forest_def spanning_forest_def by simp
have 10: "spanning_forest ?v g"
proof (unfold spanning_forest_def, intro conjI)
show "injective ?v"
using 8 by simp
next
show "acyclic ?v"
apply (rule kruskal_exchange_acyclic_inv_1)
using 2 minimum_spanning_forest_def spanning_forest_def apply simp
by (simp add: covector_mult_closed)
next
show "?v \<le> --g"
apply (rule sup_least)
using 2 inf.coboundedI1 minimum_spanning_forest_def spanning_forest_def apply simp
using 1 2 by (metis kruskal_invariant_def kruskal_spanning_invariant_def conv_complement conv_dist_inf order.trans inf.absorb2 inf.cobounded1 minimum_spanning_forest_def spanning_forest_def)
next
show "components g \<le> forest_components ?v"
using 9 by simp
next
show "regular ?v"
using 3 regular_closed_star regular_conv_closed regular_mult_closed by simp
qed
have 11: "sum (?v \<sqinter> g) = sum (w \<sqinter> g)"
proof -
have "sum (?v \<sqinter> g) = sum (w \<sqinter> -(top * ?e * w\<^sup>T\<^sup>\<star>) \<sqinter> g) + sum (?p\<^sup>T \<sqinter> g)"
using 2 by (metis conv_complement conv_top epm_8 inf_import_p inf_top_right regular_closed_top vector_top_closed minimum_spanning_forest_def spanning_forest_def sum_disjoint)
also have "... = sum (w \<sqinter> -(top * ?e * w\<^sup>T\<^sup>\<star>) \<sqinter> g) + sum (?p \<sqinter> g)"
using 1 kruskal_invariant_def kruskal_spanning_invariant_def sum_symmetric by simp
also have "... = sum (((w \<sqinter> -(top * ?e * w\<^sup>T\<^sup>\<star>)) \<squnion> ?p) \<sqinter> g)"
using inf_commute inf_left_commute sum_disjoint by simp
also have "... = sum (w \<sqinter> g)"
using 3 7 maddux_3_11_pp by simp
finally show ?thesis
by simp
qed
have 12: "?v \<squnion> ?v\<^sup>T = w \<squnion> w\<^sup>T"
proof -
have "?v \<squnion> ?v\<^sup>T = (w \<sqinter> -?p) \<squnion> ?p\<^sup>T \<squnion> (w\<^sup>T \<sqinter> -?p\<^sup>T) \<squnion> ?p"
using conv_complement conv_dist_inf conv_dist_sup inf_import_p sup_assoc by simp
also have "... = w \<squnion> w\<^sup>T"
using 3 7 conv_complement conv_dist_inf inf_import_p maddux_3_11_pp sup_monoid.add_assoc sup_monoid.add_commute by simp
finally show ?thesis
by simp
qed
have 13: "?v * ?e\<^sup>T = bot"
apply (rule kruskal_reroot_edge)
using 1 apply (simp add: minarc_arc)
using 2 minimum_spanning_forest_def spanning_forest_def by simp
have "?v \<sqinter> ?e \<le> ?v \<sqinter> top * ?e"
using inf.sup_right_isotone top_left_mult_increasing by simp
also have "... \<le> ?v * (top * ?e)\<^sup>T"
using covector_restrict_comp_conv covector_mult_closed vector_top_closed by simp
finally have 14: "?v \<sqinter> ?e = bot"
using 13 by (metis conv_dist_comp mult_assoc le_bot mult_left_zero)
let ?d = "?v \<sqinter> top * ?e\<^sup>T * ?v\<^sup>T\<^sup>\<star> \<sqinter> ?F * ?e\<^sup>T * top \<sqinter> top * ?e * -?F"
let ?w = "(?v \<sqinter> -?d) \<squnion> ?e"
have 15: "regular ?d"
using 3 regular_closed_star regular_conv_closed regular_mult_closed by simp
have 16: "?F \<le> -?d"
apply (rule kruskal_edge_between_components_1)
using 5 apply simp
using 1 conv_dist_comp minarc_arc mult_assoc by simp
have 17: "f \<squnion> f\<^sup>T \<le> (?v \<sqinter> -?d \<sqinter> -?d\<^sup>T) \<squnion> (?v\<^sup>T \<sqinter> -?d \<sqinter> -?d\<^sup>T)"
apply (rule kruskal_edge_between_components_2)
using 16 apply simp
using 1 kruskal_invariant_def kruskal_spanning_invariant_def spanning_forest_def apply simp
using 2 12 by (metis conv_dist_sup conv_involutive conv_isotone le_supI sup_commute)
show "minimum_spanning_forest ?w g \<and> ?f \<le> ?w \<squnion> ?w\<^sup>T"
proof (intro conjI)
have 18: "?e\<^sup>T \<le> ?v\<^sup>\<star>"
apply (rule kruskal_edge_arc_1[where g=g and h=h])
using minarc_below apply simp
using 1 apply (metis kruskal_invariant_def kruskal_spanning_invariant_def inf_le1)
using 1 kruskal_invariant_def kruskal_spanning_invariant_def apply simp
using 9 apply simp
using 13 by simp
have 19: "arc ?d"
apply (rule kruskal_edge_arc)
using 5 apply simp
using 10 spanning_forest_def apply blast
using 1 apply (simp add: minarc_arc)
using 3 apply (metis conv_complement pp_dist_star regular_mult_closed)
using 2 8 12 apply (simp add: kruskal_forest_components_inf)
using 10 spanning_forest_def apply simp
using 13 apply simp
using 6 apply simp
using 18 by simp
show "minimum_spanning_forest ?w g"
proof (unfold minimum_spanning_forest_def, intro conjI)
have "(?v \<sqinter> -?d) * ?e\<^sup>T \<le> ?v * ?e\<^sup>T"
using inf_le1 mult_left_isotone by simp
hence "(?v \<sqinter> -?d) * ?e\<^sup>T = bot"
using 13 le_bot by simp
hence 20: "?e * (?v \<sqinter> -?d)\<^sup>T = bot"
using conv_dist_comp conv_involutive conv_bot by force
have 21: "injective ?w"
apply (rule injective_sup)
using 8 apply (simp add: injective_inf_closed)
using 20 apply simp
using 1 arc_injective minarc_arc by blast
show "spanning_forest ?w g"
proof (unfold spanning_forest_def, intro conjI)
show "injective ?w"
using 21 by simp
next
show "acyclic ?w"
apply (rule kruskal_exchange_acyclic_inv_2)
using 10 spanning_forest_def apply blast
using 8 apply simp
using inf.coboundedI1 apply simp
using 19 apply simp
using 1 apply (simp add: minarc_arc)
using inf.cobounded2 inf.coboundedI1 apply simp
using 13 by simp
next
have "?w \<le> ?v \<squnion> ?e"
using inf_le1 sup_left_isotone by simp
also have "... \<le> --g \<squnion> ?e"
using 10 sup_left_isotone spanning_forest_def by blast
also have "... \<le> --g \<squnion> --h"
by (simp add: le_supI2 minarc_below)
also have "... = --g"
using 1 by (metis kruskal_invariant_def kruskal_spanning_invariant_def pp_isotone_inf sup.orderE)
finally show "?w \<le> --g"
by simp
next
have 22: "?d \<le> (?v \<sqinter> -?d)\<^sup>T\<^sup>\<star> * ?e\<^sup>T * top"
apply (rule kruskal_exchange_spanning_inv_2)
using 8 apply simp
using 13 apply (metis semiring.mult_not_zero star_absorb star_simulation_right_equal)
using 17 apply simp
by (simp add: inf.coboundedI1)
have "components g \<le> forest_components ?v"
using 10 spanning_forest_def by auto
also have "... \<le> forest_components ?w"
apply (rule kruskal_exchange_forest_components_inv)
using 21 apply simp
using 15 apply simp
using 1 apply (simp add: arc_top_arc minarc_arc)
apply (simp add: inf.coboundedI1)
using 13 apply simp
using 8 apply simp
apply (simp add: le_infI1)
using 22 by simp
finally show "components g \<le> forest_components ?w"
by simp
next
show "regular ?w"
using 3 7 regular_conv_closed by simp
qed
next
have 23: "?e \<sqinter> g \<noteq> bot"
using 1 by (metis kruskal_invariant_def kruskal_spanning_invariant_def comp_inf.semiring.mult_zero_right inf.sup_monoid.add_assoc inf.sup_monoid.add_commute minarc_bot_iff minarc_meet_bot)
have "g \<sqinter> -h \<le> (g \<sqinter> -h)\<^sup>\<star>"
using star.circ_increasing by simp
also have "... \<le> (--(g \<sqinter> -h))\<^sup>\<star>"
using pp_increasing star_isotone by blast
also have "... \<le> ?F"
using 1 kruskal_invariant_def kruskal_spanning_invariant_def inf.sup_monoid.add_commute spanning_forest_def by simp
finally have 24: "g \<sqinter> -h \<le> ?F"
by simp
have "?d \<le> --g"
using 10 inf.coboundedI1 spanning_forest_def by blast
hence "?d \<le> --g \<sqinter> -?F"
using 16 inf.boundedI p_antitone_iff by simp
also have "... = --(g \<sqinter> -?F)"
by simp
also have "... \<le> --h"
using 24 p_shunting_swap pp_isotone by fastforce
finally have 25: "?d \<le> --h"
by simp
have "?d = bot \<longrightarrow> top = bot"
using 19 by (metis mult_left_zero mult_right_zero)
hence "?d \<noteq> bot"
using 1 le_bot by auto
hence 26: "?d \<sqinter> h \<noteq> bot"
using 25 by (metis inf.absorb_iff2 inf_commute pseudo_complement)
have "sum (?e \<sqinter> g) = sum (?e \<sqinter> --h \<sqinter> g)"
by (simp add: inf.absorb1 minarc_below)
also have "... = sum (?e \<sqinter> h)"
using 1 by (metis kruskal_invariant_def kruskal_spanning_invariant_def inf.left_commute inf.sup_monoid.add_commute)
also have "... \<le> sum (?d \<sqinter> h)"
using 19 26 minarc_min by simp
also have "... = sum (?d \<sqinter> (--h \<sqinter> g))"
using 1 kruskal_invariant_def kruskal_spanning_invariant_def inf_commute by simp
also have "... = sum (?d \<sqinter> g)"
using 25 by (simp add: inf.absorb2 inf_assoc inf_commute)
finally have 27: "sum (?e \<sqinter> g) \<le> sum (?d \<sqinter> g)"
by simp
have "?v \<sqinter> ?e \<sqinter> -?d = bot"
using 14 by simp
hence "sum (?w \<sqinter> g) = sum (?v \<sqinter> -?d \<sqinter> g) + sum (?e \<sqinter> g)"
using sum_disjoint inf_commute inf_assoc by simp
also have "... \<le> sum (?v \<sqinter> -?d \<sqinter> g) + sum (?d \<sqinter> g)"
using 23 27 sum_plus_right_isotone by simp
also have "... = sum (((?v \<sqinter> -?d) \<squnion> ?d) \<sqinter> g)"
using sum_disjoint inf_le2 pseudo_complement by simp
also have "... = sum ((?v \<squnion> ?d) \<sqinter> (-?d \<squnion> ?d) \<sqinter> g)"
by (simp add: sup_inf_distrib2)
also have "... = sum ((?v \<squnion> ?d) \<sqinter> g)"
using 15 by (metis inf_top_right stone)
also have "... = sum (?v \<sqinter> g)"
by (simp add: inf.sup_monoid.add_assoc)
finally have "sum (?w \<sqinter> g) \<le> sum (?v \<sqinter> g)"
by simp
thus "\<forall>u . spanning_forest u g \<longrightarrow> sum (?w \<sqinter> g) \<le> sum (u \<sqinter> g)"
using 2 11 minimum_spanning_forest_def by auto
qed
next
have "?f \<le> f \<squnion> f\<^sup>T \<squnion> ?e"
using conv_dist_inf inf_le1 sup_left_isotone sup_mono by presburger
also have "... \<le> (?v \<sqinter> -?d \<sqinter> -?d\<^sup>T) \<squnion> (?v\<^sup>T \<sqinter> -?d \<sqinter> -?d\<^sup>T) \<squnion> ?e"
using 17 sup_left_isotone by simp
also have "... \<le> (?v \<sqinter> -?d) \<squnion> (?v\<^sup>T \<sqinter> -?d \<sqinter> -?d\<^sup>T) \<squnion> ?e"
using inf.cobounded1 sup_inf_distrib2 by presburger
also have "... = ?w \<squnion> (?v\<^sup>T \<sqinter> -?d \<sqinter> -?d\<^sup>T)"
by (simp add: sup_assoc sup_commute)
also have "... \<le> ?w \<squnion> (?v\<^sup>T \<sqinter> -?d\<^sup>T)"
using inf.sup_right_isotone inf_assoc sup_right_isotone by simp
also have "... \<le> ?w \<squnion> ?w\<^sup>T"
using conv_complement conv_dist_inf conv_dist_sup sup_right_isotone by simp
finally show "?f \<le> ?w \<squnion> ?w\<^sup>T"
by simp
qed
qed
next
show "?n2 < n"
using 1 kruskal_vc_2 kruskal_invariant_def by auto
qed
qed
next
show "\<not> ?e \<le> -?F \<longrightarrow> kruskal_invariant f g ?h \<and> ?n2 < n"
using 1 kruskal_vc_2 kruskal_invariant_def by auto
qed
next
fix f g h
assume 28: "kruskal_invariant f g h \<and> h = bot"
hence 29: "spanning_forest f g"
using kruskal_invariant_def kruskal_spanning_invariant_def by auto
from 28 obtain w where 30: "minimum_spanning_forest w g \<and> f \<le> w \<squnion> w\<^sup>T"
using kruskal_invariant_def by auto
hence "w = w \<sqinter> --g"
by (simp add: inf.absorb1 minimum_spanning_forest_def spanning_forest_def)
also have "... \<le> w \<sqinter> components g"
by (metis inf.sup_right_isotone star.circ_increasing)
also have "... \<le> w \<sqinter> f\<^sup>T\<^sup>\<star> * f\<^sup>\<star>"
using 29 spanning_forest_def inf.sup_right_isotone by simp
also have "... \<le> f \<squnion> f\<^sup>T"
apply (rule cancel_separate_6[where z=w and y="w\<^sup>T"])
using 30 minimum_spanning_forest_def spanning_forest_def apply simp
using 30 apply (metis conv_dist_inf conv_dist_sup conv_involutive inf.cobounded2 inf.orderE)
using 30 apply (simp add: sup_commute)
using 30 minimum_spanning_forest_def spanning_forest_def apply simp
using 30 by (metis acyclic_star_below_complement comp_inf.mult_right_isotone inf_p le_bot minimum_spanning_forest_def spanning_forest_def)
finally have 31: "w \<le> f \<squnion> f\<^sup>T"
by simp
have "sum (f \<sqinter> g) = sum ((w \<squnion> w\<^sup>T) \<sqinter> (f \<sqinter> g))"
using 30 by (metis inf_absorb2 inf.assoc)
also have "... = sum (w \<sqinter> (f \<sqinter> g)) + sum (w\<^sup>T \<sqinter> (f \<sqinter> g))"
using 30 inf.commute acyclic_asymmetric sum_disjoint minimum_spanning_forest_def spanning_forest_def by simp
also have "... = sum (w \<sqinter> (f \<sqinter> g)) + sum (w \<sqinter> (f\<^sup>T \<sqinter> g\<^sup>T))"
by (metis conv_dist_inf conv_involutive sum_conv)
also have "... = sum (f \<sqinter> (w \<sqinter> g)) + sum (f\<^sup>T \<sqinter> (w \<sqinter> g))"
using 28 inf.commute inf.assoc kruskal_invariant_def kruskal_spanning_invariant_def by simp
also have "... = sum ((f \<squnion> f\<^sup>T) \<sqinter> (w \<sqinter> g))"
using 29 acyclic_asymmetric inf.sup_monoid.add_commute sum_disjoint spanning_forest_def by simp
also have "... = sum (w \<sqinter> g)"
using 31 by (metis inf_absorb2 inf.assoc)
finally show "minimum_spanning_forest f g"
using 29 30 minimum_spanning_forest_def by simp
qed
subsection \<open>Prim's Minimum Spanning Tree Algorithm\<close>
text \<open>
The total-correctness proof of Prim's minimum spanning tree algorithm has the same overall structure as the proof of Kruskal's algorithm.
The partial-correctness proof is discussed in \cite{Guttmann2016c,Guttmann2018a}.
\<close>
abbreviation "component g r \<equiv> r\<^sup>T * (--g)\<^sup>\<star>"
definition "spanning_tree t g r \<equiv> forest t \<and> t \<le> (component g r)\<^sup>T * (component g r) \<sqinter> --g \<and> component g r \<le> r\<^sup>T * t\<^sup>\<star> \<and> regular t"
definition "minimum_spanning_tree t g r \<equiv> spanning_tree t g r \<and> (\<forall>u . spanning_tree u g r \<longrightarrow> sum (t \<sqinter> g) \<le> sum (u \<sqinter> g))"
definition "prim_precondition g r \<equiv> g = g\<^sup>T \<and> injective r \<and> vector r \<and> regular r"
definition "prim_spanning_invariant t v g r \<equiv> prim_precondition g r \<and> v\<^sup>T = r\<^sup>T * t\<^sup>\<star> \<and> spanning_tree t (v * v\<^sup>T \<sqinter> g) r"
definition "prim_invariant t v g r \<equiv> prim_spanning_invariant t v g r \<and> (\<exists>w . minimum_spanning_tree w g r \<and> t \<le> w)"
lemma span_tree_split:
assumes "vector r"
shows "t \<le> (component g r)\<^sup>T * (component g r) \<sqinter> --g \<longleftrightarrow> (t \<le> (component g r)\<^sup>T \<and> t \<le> component g r \<and> t \<le> --g)"
proof -
have "(component g r)\<^sup>T * (component g r) = (component g r)\<^sup>T \<sqinter> component g r"
by (metis assms conv_involutive covector_mult_closed vector_conv_covector vector_covector)
thus ?thesis
by simp
qed
lemma span_tree_component:
assumes "spanning_tree t g r"
shows "component g r = component t r"
using assms by (simp add: antisym mult_right_isotone star_isotone spanning_tree_def)
text \<open>
We first show three verification conditions which are used in both correctness proofs.
\<close>
lemma prim_vc_1:
assumes "prim_precondition g r"
shows "prim_spanning_invariant bot r g r"
proof (unfold prim_spanning_invariant_def, intro conjI)
show "prim_precondition g r"
using assms by simp
next
show "r\<^sup>T = r\<^sup>T * bot\<^sup>\<star>"
by (simp add: star_absorb)
next
let ?ss = "r * r\<^sup>T \<sqinter> g"
show "spanning_tree bot ?ss r"
proof (unfold spanning_tree_def, intro conjI)
show "injective bot"
by simp
next
show "acyclic bot"
by simp
next
show "bot \<le> (component ?ss r)\<^sup>T * (component ?ss r) \<sqinter> --?ss"
by simp
next
have "component ?ss r \<le> component (r * r\<^sup>T) r"
by (simp add: mult_right_isotone star_isotone)
also have "... \<le> r\<^sup>T * 1\<^sup>\<star>"
using assms by (metis inf.eq_iff p_antitone regular_one_closed star_sub_one prim_precondition_def)
also have "... = r\<^sup>T * bot\<^sup>\<star>"
by (simp add: star.circ_zero star_one)
finally show "component ?ss r \<le> r\<^sup>T * bot\<^sup>\<star>"
.
next
show "regular bot"
by simp
qed
qed
lemma prim_vc_2:
assumes "prim_spanning_invariant t v g r"
and "v * -v\<^sup>T \<sqinter> g \<noteq> bot"
and "card { x . regular x \<and> x \<le> component g r \<and> x \<le> -v\<^sup>T } = n"
shows "prim_spanning_invariant (t \<squnion> minarc (v * -v\<^sup>T \<sqinter> g)) (v \<squnion> minarc (v * -v\<^sup>T \<sqinter> g)\<^sup>T * top) g r \<and> card { x . regular x \<and> x \<le> component g r \<and> x \<le> -(v \<squnion> minarc (v * -v\<^sup>T \<sqinter> g)\<^sup>T * top)\<^sup>T } < n"
proof -
let ?vcv = "v * -v\<^sup>T \<sqinter> g"
let ?e = "minarc ?vcv"
let ?t = "t \<squnion> ?e"
let ?v = "v \<squnion> ?e\<^sup>T * top"
let ?c = "component g r"
let ?g = "--g"
let ?n1 = "card { x . regular x \<and> x \<le> ?c \<and> x \<le> -v\<^sup>T }"
let ?n2 = "card { x . regular x \<and> x \<le> ?c \<and> x \<le> -?v\<^sup>T }"
have 1: "regular v \<and> regular (v * v\<^sup>T) \<and> regular (?v * ?v\<^sup>T) \<and> regular (top * ?e)"
using assms(1) by (metis prim_spanning_invariant_def spanning_tree_def prim_precondition_def regular_conv_closed regular_closed_star regular_mult_closed conv_involutive regular_closed_top regular_closed_sup minarc_regular)
hence 2: "t \<le> v * v\<^sup>T \<sqinter> ?g"
using assms(1) by (metis prim_spanning_invariant_def spanning_tree_def inf_pp_commute inf.boundedE)
hence 3: "t \<le> v * v\<^sup>T"
by simp
have 4: "t \<le> ?g"
using 2 by simp
have 5: "?e \<le> v * -v\<^sup>T \<sqinter> ?g"
using 1 by (metis minarc_below pp_dist_inf regular_mult_closed regular_closed_p)
hence 6: "?e \<le> v * -v\<^sup>T"
by simp
have 7: "vector v"
using assms(1) prim_spanning_invariant_def prim_precondition_def by (simp add: covector_mult_closed vector_conv_covector)
hence 8: "?e \<le> v"
using 6 by (metis conv_complement inf.boundedE vector_complement_closed vector_covector)
have 9: "?e * t = bot"
using 3 6 7 et(1) by blast
have 10: "?e * t\<^sup>T = bot"
using 3 6 7 et(2) by simp
have 11: "arc ?e"
using assms(2) minarc_arc by simp
have "r\<^sup>T \<le> r\<^sup>T * t\<^sup>\<star>"
by (metis mult_right_isotone order_refl semiring.mult_not_zero star.circ_separate_mult_1 star_absorb)
hence 12: "r\<^sup>T \<le> v\<^sup>T"
using assms(1) by (simp add: prim_spanning_invariant_def)
have 13: "vector r \<and> injective r \<and> v\<^sup>T = r\<^sup>T * t\<^sup>\<star>"
using assms(1) prim_spanning_invariant_def prim_precondition_def minimum_spanning_tree_def spanning_tree_def reachable_restrict by simp
have "g = g\<^sup>T"
using assms(1) prim_invariant_def prim_spanning_invariant_def prim_precondition_def by simp
hence 14: "?g\<^sup>T = ?g"
using conv_complement by simp
show "prim_spanning_invariant ?t ?v g r \<and> ?n2 < n"
proof (rule conjI)
show "prim_spanning_invariant ?t ?v g r"
proof (unfold prim_spanning_invariant_def, intro conjI)
show "prim_precondition g r"
using assms(1) prim_spanning_invariant_def by simp
next
show "?v\<^sup>T = r\<^sup>T * ?t\<^sup>\<star>"
using assms(1) 6 7 9 by (simp add: reachable_inv prim_spanning_invariant_def prim_precondition_def spanning_tree_def)
next
let ?G = "?v * ?v\<^sup>T \<sqinter> g"
show "spanning_tree ?t ?G r"
proof (unfold spanning_tree_def, intro conjI)
show "injective ?t"
using assms(1) 10 11 by (simp add: injective_inv prim_spanning_invariant_def spanning_tree_def)
next
show "acyclic ?t"
using assms(1) 3 6 7 acyclic_inv prim_spanning_invariant_def spanning_tree_def by simp
next
show "?t \<le> (component ?G r)\<^sup>T * (component ?G r) \<sqinter> --?G"
using 1 2 5 7 13 prim_subgraph_inv inf_pp_commute mst_subgraph_inv_2 by auto
next
show "component (?v * ?v\<^sup>T \<sqinter> g) r \<le> r\<^sup>T * ?t\<^sup>\<star>"
proof -
have 15: "r\<^sup>T * (v * v\<^sup>T \<sqinter> ?g)\<^sup>\<star> \<le> r\<^sup>T * t\<^sup>\<star>"
using assms(1) 1 by (metis prim_spanning_invariant_def spanning_tree_def inf_pp_commute)
have "component (?v * ?v\<^sup>T \<sqinter> g) r = r\<^sup>T * (?v * ?v\<^sup>T \<sqinter> ?g)\<^sup>\<star>"
using 1 by simp
also have "... \<le> r\<^sup>T * ?t\<^sup>\<star>"
using 2 6 7 11 12 13 14 15 by (metis span_inv)
finally show ?thesis
.
qed
next
show "regular ?t"
using assms(1) by (metis prim_spanning_invariant_def spanning_tree_def regular_closed_sup minarc_regular)
qed
qed
next
have 16: "top * ?e \<le> ?c"
proof -
have "top * ?e = top * ?e\<^sup>T * ?e"
using 11 by (metis arc_top_edge mult_assoc)
also have "... \<le> v\<^sup>T * ?e"
using 7 8 by (metis conv_dist_comp conv_isotone mult_left_isotone symmetric_top_closed)
also have "... \<le> v\<^sup>T * ?g"
using 5 mult_right_isotone by auto
also have "... = r\<^sup>T * t\<^sup>\<star> * ?g"
using 13 by simp
also have "... \<le> r\<^sup>T * ?g\<^sup>\<star> * ?g"
using 4 by (simp add: mult_left_isotone mult_right_isotone star_isotone)
also have "... \<le> ?c"
by (simp add: comp_associative mult_right_isotone star.right_plus_below_circ)
finally show ?thesis
by simp
qed
have 17: "top * ?e \<le> -v\<^sup>T"
using 6 7 by (simp add: schroeder_4_p vTeT)
have 18: "\<not> top * ?e \<le> -(top * ?e)"
by (metis assms(2) inf.orderE minarc_bot_iff conv_complement_sub_inf inf_p inf_top.left_neutral p_bot symmetric_top_closed vector_top_closed)
have 19: "-?v\<^sup>T = -v\<^sup>T \<sqinter> -(top * ?e)"
by (simp add: conv_dist_comp conv_dist_sup)
hence 20: "\<not> top * ?e \<le> -?v\<^sup>T"
using 18 by simp
have "?n2 < ?n1"
apply (rule psubset_card_mono)
using finite_regular apply simp
using 1 16 17 19 20 by auto
thus "?n2 < n"
using assms(3) by simp
qed
qed
lemma prim_vc_3:
assumes "prim_spanning_invariant t v g r"
and "v * -v\<^sup>T \<sqinter> g = bot"
shows "spanning_tree t g r"
proof -
let ?g = "--g"
have 1: "regular v \<and> regular (v * v\<^sup>T)"
using assms(1) by (metis prim_spanning_invariant_def spanning_tree_def prim_precondition_def regular_conv_closed regular_closed_star regular_mult_closed conv_involutive)
have 2: "v * -v\<^sup>T \<sqinter> ?g = bot"
using assms(2) pp_inf_bot_iff pp_pp_inf_bot_iff by simp
have 3: "v\<^sup>T = r\<^sup>T * t\<^sup>\<star> \<and> vector v"
using assms(1) by (simp add: covector_mult_closed prim_invariant_def prim_spanning_invariant_def vector_conv_covector prim_precondition_def)
have 4: "t \<le> v * v\<^sup>T \<sqinter> ?g"
using assms(1) 1 by (metis prim_spanning_invariant_def inf_pp_commute spanning_tree_def inf.boundedE)
have "r\<^sup>T * (v * v\<^sup>T \<sqinter> ?g)\<^sup>\<star> \<le> r\<^sup>T * t\<^sup>\<star>"
using assms(1) 1 by (metis prim_spanning_invariant_def inf_pp_commute spanning_tree_def)
hence 5: "component g r = v\<^sup>T"
using 1 2 3 4 by (metis span_post)
have "regular (v * v\<^sup>T)"
using assms(1) by (metis prim_spanning_invariant_def spanning_tree_def prim_precondition_def regular_conv_closed regular_closed_star regular_mult_closed conv_involutive)
hence 6: "t \<le> v * v\<^sup>T \<sqinter> ?g"
by (metis assms(1) prim_spanning_invariant_def spanning_tree_def inf_pp_commute inf.boundedE)
show "spanning_tree t g r"
apply (unfold spanning_tree_def, intro conjI)
using assms(1) prim_spanning_invariant_def spanning_tree_def apply simp
using assms(1) prim_spanning_invariant_def spanning_tree_def apply simp
using 5 6 apply simp
using assms(1) 5 prim_spanning_invariant_def apply simp
using assms(1) prim_spanning_invariant_def spanning_tree_def by simp
qed
text \<open>
The following result shows that Prim's algorithm terminates and constructs a spanning tree.
We cannot yet show that this is a minimum spanning tree.
\<close>
theorem prim_spanning:
"VARS t v e
[ prim_precondition g r ]
t := bot;
v := r;
WHILE v * -v\<^sup>T \<sqinter> g \<noteq> bot
INV { prim_spanning_invariant t v g r }
VAR { card { x . regular x \<and> x \<le> component g r \<sqinter> -v\<^sup>T } }
DO e := minarc (v * -v\<^sup>T \<sqinter> g);
t := t \<squnion> e;
v := v \<squnion> e\<^sup>T * top
OD
[ spanning_tree t g r ]"
apply vcg_tc_simp
apply (simp add: prim_vc_1)
using prim_vc_2 apply blast
using prim_vc_3 by auto
text \<open>
Because we have shown total correctness, we conclude that a spanning tree exists.
\<close>
lemma prim_exists_spanning:
"prim_precondition g r \<Longrightarrow> \<exists>t . spanning_tree t g r"
using tc_extract_function prim_spanning by blast
text \<open>
This implies that a minimum spanning tree exists, which is used in the subsequent correctness proof.
\<close>
lemma prim_exists_minimal_spanning:
assumes "prim_precondition g r"
shows "\<exists>t . minimum_spanning_tree t g r"
proof -
let ?s = "{ t . spanning_tree t g r }"
have "\<exists>m\<in>?s . \<forall>z\<in>?s . sum (m \<sqinter> g) \<le> sum (z \<sqinter> g)"
apply (rule finite_set_minimal)
using finite_regular spanning_tree_def apply simp
using assms prim_exists_spanning apply simp
using sum_linear by simp
thus ?thesis
using minimum_spanning_tree_def by simp
qed
text \<open>
Prim's minimum spanning tree algorithm terminates and is correct.
This is the same algorithm that is used in the previous correctness proof, with the same precondition and variant, but with a different invariant and postcondition.
\<close>
theorem prim:
"VARS t v e
[ prim_precondition g r \<and> (\<exists>w . minimum_spanning_tree w g r) ]
t := bot;
v := r;
WHILE v * -v\<^sup>T \<sqinter> g \<noteq> bot
INV { prim_invariant t v g r }
VAR { card { x . regular x \<and> x \<le> component g r \<sqinter> -v\<^sup>T } }
DO e := minarc (v * -v\<^sup>T \<sqinter> g);
t := t \<squnion> e;
v := v \<squnion> e\<^sup>T * top
OD
[ minimum_spanning_tree t g r ]"
-proof vcg_tc_simp
+using [[simproc del: defined_all]] proof vcg_tc_simp
assume "prim_precondition g r \<and> (\<exists>w . minimum_spanning_tree w g r)"
thus "prim_invariant bot r g r"
using prim_invariant_def prim_vc_1 by simp
next
fix t v n
let ?vcv = "v * -v\<^sup>T \<sqinter> g"
let ?vv = "v * v\<^sup>T \<sqinter> g"
let ?e = "minarc ?vcv"
let ?t = "t \<squnion> ?e"
let ?v = "v \<squnion> ?e\<^sup>T * top"
let ?c = "component g r"
let ?g = "--g"
let ?n1 = "card { x . regular x \<and> x \<le> ?c \<and> x \<le> -v\<^sup>T }"
let ?n2 = "card { x . regular x \<and> x \<le> ?c \<and> x \<le> -?v\<^sup>T }"
assume 1: "prim_invariant t v g r \<and> ?vcv \<noteq> bot \<and> ?n1 = n"
hence 2: "regular v \<and> regular (v * v\<^sup>T)"
by (metis (no_types, hide_lams) prim_invariant_def prim_spanning_invariant_def spanning_tree_def prim_precondition_def regular_conv_closed regular_closed_star regular_mult_closed conv_involutive)
have 3: "t \<le> v * v\<^sup>T \<sqinter> ?g"
using 1 2 by (metis (no_types, hide_lams) prim_invariant_def prim_spanning_invariant_def spanning_tree_def inf_pp_commute inf.boundedE)
hence 4: "t \<le> v * v\<^sup>T"
by simp
have 5: "t \<le> ?g"
using 3 by simp
have 6: "?e \<le> v * -v\<^sup>T \<sqinter> ?g"
using 2 by (metis minarc_below pp_dist_inf regular_mult_closed regular_closed_p)
hence 7: "?e \<le> v * -v\<^sup>T"
by simp
have 8: "vector v"
using 1 prim_invariant_def prim_spanning_invariant_def prim_precondition_def by (simp add: covector_mult_closed vector_conv_covector)
have 9: "arc ?e"
using 1 minarc_arc by simp
from 1 obtain w where 10: "minimum_spanning_tree w g r \<and> t \<le> w"
by (metis prim_invariant_def)
hence 11: "vector r \<and> injective r \<and> v\<^sup>T = r\<^sup>T * t\<^sup>\<star> \<and> forest w \<and> t \<le> w \<and> w \<le> ?c\<^sup>T * ?c \<sqinter> ?g \<and> r\<^sup>T * (?c\<^sup>T * ?c \<sqinter> ?g)\<^sup>\<star> \<le> r\<^sup>T * w\<^sup>\<star>"
using 1 2 prim_invariant_def prim_spanning_invariant_def prim_precondition_def minimum_spanning_tree_def spanning_tree_def reachable_restrict by simp
hence 12: "w * v \<le> v"
using predecessors_reachable reachable_restrict by auto
have 13: "g = g\<^sup>T"
using 1 prim_invariant_def prim_spanning_invariant_def prim_precondition_def by simp
hence 14: "?g\<^sup>T = ?g"
using conv_complement by simp
show "prim_invariant ?t ?v g r \<and> ?n2 < n"
proof (unfold prim_invariant_def, intro conjI)
show "prim_spanning_invariant ?t ?v g r"
using 1 prim_invariant_def prim_vc_2 by blast
next
show "\<exists>w . minimum_spanning_tree w g r \<and> ?t \<le> w"
proof
let ?f = "w \<sqinter> v * -v\<^sup>T \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>"
let ?p = "w \<sqinter> -v * -v\<^sup>T \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>"
let ?fp = "w \<sqinter> -v\<^sup>T \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>"
let ?w = "(w \<sqinter> -?fp) \<squnion> ?p\<^sup>T \<squnion> ?e"
have 15: "regular ?f \<and> regular ?fp \<and> regular ?w"
using 2 10 by (metis regular_conv_closed regular_closed_star regular_mult_closed regular_closed_top regular_closed_inf regular_closed_sup minarc_regular minimum_spanning_tree_def spanning_tree_def)
show "minimum_spanning_tree ?w g r \<and> ?t \<le> ?w"
proof (intro conjI)
show "minimum_spanning_tree ?w g r"
proof (unfold minimum_spanning_tree_def, intro conjI)
show "spanning_tree ?w g r"
proof (unfold spanning_tree_def, intro conjI)
show "injective ?w"
using 7 8 9 11 exchange_injective by blast
next
show "acyclic ?w"
using 7 8 11 12 exchange_acyclic by blast
next
show "?w \<le> ?c\<^sup>T * ?c \<sqinter> --g"
proof -
have 16: "w \<sqinter> -?fp \<le> ?c\<^sup>T * ?c \<sqinter> --g"
using 10 by (simp add: le_infI1 minimum_spanning_tree_def spanning_tree_def)
have "?p\<^sup>T \<le> w\<^sup>T"
by (simp add: conv_isotone inf.sup_monoid.add_assoc)
also have "... \<le> (?c\<^sup>T * ?c \<sqinter> --g)\<^sup>T"
using 11 conv_order by simp
also have "... = ?c\<^sup>T * ?c \<sqinter> --g"
using 2 14 conv_dist_comp conv_dist_inf by simp
finally have 17: "?p\<^sup>T \<le> ?c\<^sup>T * ?c \<sqinter> --g"
.
have "?e \<le> ?c\<^sup>T * ?c \<sqinter> ?g"
using 5 6 11 mst_subgraph_inv by auto
thus ?thesis
using 16 17 by simp
qed
next
show "?c \<le> r\<^sup>T * ?w\<^sup>\<star>"
proof -
have "?c \<le> r\<^sup>T * w\<^sup>\<star>"
using 10 minimum_spanning_tree_def spanning_tree_def by simp
also have "... \<le> r\<^sup>T * ?w\<^sup>\<star>"
using 4 7 8 10 11 12 15 by (metis mst_reachable_inv)
finally show ?thesis
.
qed
next
show "regular ?w"
using 15 by simp
qed
next
have 18: "?f \<squnion> ?p = ?fp"
using 2 8 epm_1 by fastforce
have "arc (w \<sqinter> --v * -v\<^sup>T \<sqinter> top * ?e * w\<^sup>T\<^sup>\<star>)"
using 5 6 8 9 11 12 reachable_restrict arc_edge by auto
hence 19: "arc ?f"
using 2 by simp
hence "?f = bot \<longrightarrow> top = bot"
by (metis mult_left_zero mult_right_zero)
hence "?f \<noteq> bot"
using 1 le_bot by auto
hence "?f \<sqinter> v * -v\<^sup>T \<sqinter> ?g \<noteq> bot"
using 2 11 by (simp add: inf.absorb1 le_infI1)
hence "g \<sqinter> (?f \<sqinter> v * -v\<^sup>T) \<noteq> bot"
using inf_commute pp_inf_bot_iff by simp
hence 20: "?f \<sqinter> ?vcv \<noteq> bot"
by (simp add: inf_assoc inf_commute)
hence 21: "?f \<sqinter> g = ?f \<sqinter> ?vcv"
using 2 by (simp add: inf_assoc inf_commute inf_left_commute)
have 22: "?e \<sqinter> g = minarc ?vcv \<sqinter> ?vcv"
using 7 by (simp add: inf.absorb2 inf.assoc inf.commute)
hence 23: "sum (?e \<sqinter> g) \<le> sum (?f \<sqinter> g)"
using 15 19 20 21 by (simp add: minarc_min)
have "?e \<noteq> bot"
using 20 comp_inf.semiring.mult_not_zero semiring.mult_not_zero by blast
hence 24: "?e \<sqinter> g \<noteq> bot"
using 22 minarc_meet_bot by auto
have "sum (?w \<sqinter> g) = sum (w \<sqinter> -?fp \<sqinter> g) + sum (?p\<^sup>T \<sqinter> g) + sum (?e \<sqinter> g)"
using 7 8 10 by (metis sum_disjoint_3 epm_8 epm_9 epm_10 minimum_spanning_tree_def spanning_tree_def)
also have "... = sum (((w \<sqinter> -?fp) \<squnion> ?p\<^sup>T) \<sqinter> g) + sum (?e \<sqinter> g)"
using 11 by (metis epm_8 sum_disjoint)
also have "... \<le> sum (((w \<sqinter> -?fp) \<squnion> ?p\<^sup>T) \<sqinter> g) + sum (?f \<sqinter> g)"
using 23 24 by (simp add: sum_plus_right_isotone)
also have "... = sum (w \<sqinter> -?fp \<sqinter> g) + sum (?p\<^sup>T \<sqinter> g) + sum (?f \<sqinter> g)"
using 11 by (metis epm_8 sum_disjoint)
also have "... = sum (w \<sqinter> -?fp \<sqinter> g) + sum (?p \<sqinter> g) + sum (?f \<sqinter> g)"
using 13 sum_symmetric by auto
also have "... = sum (((w \<sqinter> -?fp) \<squnion> ?p \<squnion> ?f) \<sqinter> g)"
using 2 8 by (metis sum_disjoint_3 epm_11 epm_12 epm_13)
also have "... = sum (w \<sqinter> g)"
using 2 8 15 18 epm_2 by force
finally have "sum (?w \<sqinter> g) \<le> sum (w \<sqinter> g)"
.
thus "\<forall>u . spanning_tree u g r \<longrightarrow> sum (?w \<sqinter> g) \<le> sum (u \<sqinter> g)"
using 10 order_lesseq_imp minimum_spanning_tree_def by auto
qed
next
show "?t \<le> ?w"
using 4 8 10 mst_extends_new_tree by simp
qed
qed
next
show "?n2 < n"
using 1 prim_invariant_def prim_vc_2 by auto
qed
next
fix t v
let ?g = "--g"
assume 25: "prim_invariant t v g r \<and> v * -v\<^sup>T \<sqinter> g = bot"
hence 26: "regular v"
by (metis prim_invariant_def prim_spanning_invariant_def spanning_tree_def prim_precondition_def regular_conv_closed regular_closed_star regular_mult_closed conv_involutive)
from 25 obtain w where 27: "minimum_spanning_tree w g r \<and> t \<le> w"
by (metis prim_invariant_def)
have "spanning_tree t g r"
using 25 prim_invariant_def prim_vc_3 by blast
hence "component g r = v\<^sup>T"
by (metis 25 prim_invariant_def span_tree_component prim_spanning_invariant_def spanning_tree_def)
hence 28: "w \<le> v * v\<^sup>T"
using 26 27 by (simp add: minimum_spanning_tree_def spanning_tree_def inf_pp_commute)
have "vector r \<and> injective r \<and> forest w"
using 25 27 by (simp add: prim_invariant_def prim_spanning_invariant_def prim_precondition_def minimum_spanning_tree_def spanning_tree_def)
hence "w = t"
using 25 27 28 prim_invariant_def prim_spanning_invariant_def mst_post by blast
thus "minimum_spanning_tree t g r"
using 27 by simp
qed
end
end
diff --git a/thys/Buildings/Algebra.thy b/thys/Buildings/Algebra.thy
--- a/thys/Buildings/Algebra.thy
+++ b/thys/Buildings/Algebra.thy
@@ -1,3095 +1,3095 @@
section \<open>Algebra\<close>
text \<open>
In this section, we develop the necessary algebra for developing the theory of Coxeter systems,
including groups, quotient groups, free groups, group presentations, and words in a group over a
set of generators.
\<close>
theory Algebra
imports Prelim
begin
subsection \<open>Miscellaneous algebra facts\<close>
lemma times2_conv_add: "(j::nat) + j = 2*j"
by (induct j) auto
lemma (in comm_semiring_1) odd_n0: "odd m \<Longrightarrow> m\<noteq>0"
using dvd_0_right by fast
lemma (in semigroup_add) add_assoc4: "a + b + c + d = a + (b + c + d)"
using add.assoc by simp
lemmas (in monoid_add) sum_list_map_cong =
arg_cong[OF map_cong, OF refl, of _ _ _ sum_list]
context group_add
begin
lemma map_uminus_order2:
"\<forall>s\<in>set ss. s+s=0 \<Longrightarrow> map (uminus) ss = ss"
by (induct ss) (auto simp add: minus_unique)
lemma uminus_sum_list: "- sum_list as = sum_list (map uminus (rev as))"
by (induct as) (auto simp add: minus_add)
lemma uminus_sum_list_order2:
"\<forall>s\<in>set ss. s+s=0 \<Longrightarrow> - sum_list ss = sum_list (rev ss)"
using uminus_sum_list map_uminus_order2 by simp
end (* context group_add *)
subsection \<open>The type of permutations of a type\<close>
text \<open>
Here we construct a type consisting of all bijective functions on a type. This is the
prototypical example of a group, where the group operation is composition, and every group can
be embedded into such a type. It is for this purpose that we construct this type, so that we may
confer upon suitable subsets of types that are not of class @{class group_add} the properties of
that class, via a suitable injective correspondence to this permutation type.
\<close>
typedef 'a permutation = "{f::'a\<Rightarrow>'a. bij f}"
morphisms permutation Abs_permutation
by fast
setup_lifting type_definition_permutation
abbreviation permutation_apply :: "'a permutation \<Rightarrow> 'a \<Rightarrow> 'a " (infixr "\<rightarrow>" 90)
where "p \<rightarrow> a \<equiv> permutation p a"
abbreviation permutation_image :: "'a permutation \<Rightarrow> 'a set \<Rightarrow> 'a set"
(infixr "`\<rightarrow>" 90)
where "p `\<rightarrow> A \<equiv> permutation p ` A"
lemma permutation_eq_image: "a `\<rightarrow> A = a `\<rightarrow> B \<Longrightarrow> A=B"
using permutation[of a] inj_eq_image[OF bij_is_inj] by auto
instantiation permutation :: (type) zero
begin
lift_definition zero_permutation :: "'a permutation" is "id::'a\<Rightarrow>'a" by simp
instance ..
end
instantiation permutation :: (type) plus
begin
lift_definition plus_permutation :: "'a permutation \<Rightarrow> 'a permutation \<Rightarrow> 'a permutation"
is "comp"
using bij_comp
by fast
instance ..
end
lemma plus_permutation_abs_eq:
"bij f \<Longrightarrow> bij g \<Longrightarrow>
Abs_permutation f + Abs_permutation g = Abs_permutation (f\<circ>g)"
by (simp add: plus_permutation.abs_eq eq_onp_same_args)
instance permutation :: (type) semigroup_add
proof
fix a b c :: "'a permutation" show "a + b + c = a + (b + c)"
using comp_assoc[of "permutation a" "permutation b" "permutation c"]
by transfer simp
qed
instance permutation :: (type) monoid_add
proof
fix a :: "'a permutation"
show "0 + a = a" by transfer simp
show "a + 0 = a" by transfer simp
qed
instantiation permutation :: (type) uminus
begin
lift_definition uminus_permutation :: "'a permutation \<Rightarrow> 'a permutation"
is "\<lambda>f. the_inv f"
using bij_betw_the_inv_into
by fast
instance ..
end
instantiation permutation :: (type) minus
begin
lift_definition minus_permutation :: "'a permutation \<Rightarrow> 'a permutation \<Rightarrow> 'a permutation"
is "\<lambda>f g. f \<circ> (the_inv g)"
using bij_betw_the_inv_into bij_comp
by fast
instance ..
end
lemma minus_permutation_abs_eq:
"bij f \<Longrightarrow> bij g \<Longrightarrow>
Abs_permutation f - Abs_permutation g = Abs_permutation (f \<circ> the_inv g)"
by (simp add: minus_permutation.abs_eq eq_onp_same_args)
instance permutation :: (type) group_add
proof
fix a b :: "'a permutation"
show "- a + a = 0" using the_inv_leftinv[of "permutation a"] by transfer simp
show "a + - b = a - b" by transfer simp
qed
subsection \<open>Natural action of @{typ nat} on types of class @{class monoid_add}\<close>
subsubsection \<open>Translation from class @{class power}.\<close>
text \<open>
Here we translate the @{class power} class to apply to types of class @{class monoid_add}.
\<close>
context monoid_add
begin
sublocale nataction: power 0 plus .
sublocale add_mult_translate: monoid_mult 0 plus
by unfold_locales (auto simp add: add.assoc)
abbreviation nataction :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infix "+^" 80)
where "a+^n \<equiv> nataction.power a n"
lemmas nataction_2 = add_mult_translate.power2_eq_square
lemmas nataction_Suc2 = add_mult_translate.power_Suc2
lemma alternating_sum_list_conv_nataction:
"sum_list (alternating_list (2*n) s t) = (s+t)+^n"
by (induct n) (auto simp add: nataction_Suc2[THEN sym])
lemma nataction_add_flip: "(a+b)+^(Suc n) = a + (b+a)+^n + b"
using nataction_Suc2 add.assoc by (induct n arbitrary: a b) auto
end (* context monoid_add *)
lemma (in group_add) nataction_add_eq0_flip:
assumes "(a+b)+^n = 0"
shows "(b+a)+^n = 0"
proof (cases n)
case (Suc k) with assms show ?thesis
using nataction_add_flip add.assoc[of "-a" "a+b" "(a+b)+^k"] by simp
qed simp
subsubsection \<open>Additive order of an element\<close>
context monoid_add
begin
definition add_order :: "'a \<Rightarrow> nat"
where "add_order a \<equiv> if (\<exists>n>0. a+^n = 0) then
(LEAST n. n>0 \<and> a+^n = 0) else 0"
lemma add_order: "a+^(add_order a) = 0"
using LeastI_ex[of "\<lambda>n. n>0 \<and> a+^n = 0"] add_order_def by simp
lemma add_order_least: "n>0 \<Longrightarrow> a+^n = 0 \<Longrightarrow> add_order a \<le> n"
using Least_le[of "\<lambda>n. n>0 \<and> a+^n = 0"] add_order_def by simp
lemma add_order_equality:
"\<lbrakk> n>0; a+^n = 0; (\<And>m. m>0 \<Longrightarrow> a+^m = 0 \<Longrightarrow> n\<le>m) \<rbrakk> \<Longrightarrow>
add_order a = n"
using Least_equality[of "\<lambda>n. n>0 \<and> a+^n = 0"] add_order_def by auto
lemma add_order0: "add_order 0 = 1"
using add_order_equality by simp
lemma add_order_gt0: "(add_order a > 0) = (\<exists>n>0. a+^n = 0)"
using LeastI_ex[of "\<lambda>n. n>0 \<and> a+^n = 0"] add_order_def by simp
lemma add_order_eq0: "add_order a = 0 \<Longrightarrow> n>0 \<Longrightarrow> a+^n \<noteq> 0"
using add_order_gt0 by force
lemma less_add_order_eq_0:
assumes "a+^k = 0" "k < add_order a"
shows "k = 0"
proof (cases "k=0")
case False
moreover with assms(1) have "\<exists>n>0. a+^n = 0" by fast
ultimately show ?thesis
using assms add_order_def not_less_Least[of k "\<lambda>n. n>0 \<and> a+^n = 0"]
by auto
qed simp
lemma less_add_order_eq_0_contra: "k>0 \<Longrightarrow> k < add_order a \<Longrightarrow> a+^k \<noteq> 0"
using less_add_order_eq_0 by fast
lemma add_order_relator: "add_order (a+^(add_order a)) = 1"
using add_order by (auto intro: add_order_equality)
abbreviation pair_relator_list :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list"
where "pair_relator_list s t \<equiv> alternating_list (2*add_order (s+t)) s t"
abbreviation pair_relator_halflist :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list"
where "pair_relator_halflist s t \<equiv> alternating_list (add_order (s+t)) s t"
abbreviation pair_relator_halflist2 :: "'a \<Rightarrow> 'a \<Rightarrow> 'a list"
where "pair_relator_halflist2 s t \<equiv>
(if even (add_order (s+t)) then pair_relator_halflist s t else
pair_relator_halflist t s)"
lemma sum_list_pair_relator_list: "sum_list (pair_relator_list s t) = 0"
by (auto simp add: add_order alternating_sum_list_conv_nataction)
end (* context monoid_add *)
context group_add
begin
lemma add_order_add_eq1: "add_order (s+t) = 1 \<Longrightarrow> t = -s"
using add_order[of "s+t"] by (simp add: minus_unique)
lemma add_order_add_sym: "add_order (t+s) = add_order (s+t)"
proof (cases "add_order (t+s) = 0" "add_order (s+t) = 0" rule: two_cases)
case one thus ?thesis
using add_order nataction_add_eq0_flip[of s t] add_order_eq0 by auto
next
case other thus ?thesis
using add_order nataction_add_eq0_flip[of t s] add_order_eq0 by auto
next
case neither thus ?thesis
using add_order[of "s+t"] add_order[of "t+s"]
nataction_add_eq0_flip[of s t] nataction_add_eq0_flip[of t s]
add_order_least[of "add_order (s+t)"] add_order_least[of "add_order (t+s)"]
by fastforce
qed simp
lemma pair_relator_halflist_append:
"pair_relator_halflist s t @ pair_relator_halflist2 s t = pair_relator_list s t"
using alternating_list_split[of "add_order (s+t)" "add_order (s+t)" s t]
by (auto simp add: times2_conv_add add_order_add_sym)
lemma rev_pair_relator_list: "rev (pair_relator_list s t) = pair_relator_list t s"
by (simp add:rev_alternating_list add_order_add_sym)
lemma pair_relator_halflist2_conv_rev_pair_relator_halflist:
"pair_relator_halflist2 s t = rev (pair_relator_halflist t s)"
by (auto simp add: add_order_add_sym rev_alternating_list)
end (* context group_add *)
subsection \<open>Partial sums of a list\<close>
text \<open>
Here we construct a list that collects the results of adding the elements of a given list
together one-by-one.
\<close>
context monoid_add
begin
primrec sums :: "'a list \<Rightarrow> 'a list"
where
"sums [] = [0]"
| "sums (x#xs) = 0 # map ((+) x) (sums xs)"
lemma length_sums: "length (sums xs) = Suc (length xs)"
by (induct xs) auto
lemma sums_snoc: "sums (xs@[x]) = sums xs @ [sum_list (xs@[x])]"
by (induct xs) (auto simp add: add.assoc)
lemma sums_append2:
"sums (xs@ys) = butlast (sums xs) @ map ((+) (sum_list xs)) (sums ys)"
proof (induct ys rule: rev_induct)
case Nil show ?case by (cases xs rule: rev_cases) (auto simp add: sums_snoc)
next
case (snoc y ys) thus ?case using sums_snoc[of "xs@ys"] by (simp add: sums_snoc)
qed
lemma sums_Cons_conv_append_tl:
"sums (x#xs) = 0 # x # map ((+) x) (tl (sums xs))"
by (cases xs) auto
lemma pullback_sums_map_middle2:
"map F (sums xs) = ds@[d,e]@es \<Longrightarrow>
\<exists>as a bs. xs = as@[a]@bs \<and> map F (sums as) = ds@[d] \<and>
d = F (sum_list as) \<and> e = F (sum_list (as@[a]))"
proof (induct xs es rule: list_induct2_snoc)
case (Nil2 xs)
show ?case
proof (cases xs rule: rev_cases)
case Nil with Nil2 show ?thesis by simp
next
case (snoc ys y) have ys: "xs = ys@[y]" by fact
with Nil2(1) have y: "map F (sums ys) = ds@[d]" "e = F (sum_list (ys@[y]))"
by (auto simp add: sums_snoc)
show ?thesis
proof (cases ys rule: rev_cases)
case Nil
with ys y have
"xs = []@[y]@[]" "map F (sums []) = ds@[d]"
"d = F (sum_list [])" "e = F (sum_list ([]@[y]))"
by auto
thus ?thesis by fast
next
case (snoc zs z)
with y(1) have z: "map F (sums zs) = ds" "d = F (sum_list (zs@[z]))"
by (auto simp add: sums_snoc)
from z(1) ys y snoc have
"xs = (zs@[z])@[y]@[]" "map F (sums (zs@[z])) = ds@[d]"
"e = F (sum_list ((zs@[z])@[y]))"
by auto
with z(2) show ?thesis by fast
qed
qed
next
case snoc thus ?case by (fastforce simp add: sums_snoc)
qed simp
lemma pullback_sums_map_middle3:
"map F (sums xs) = ds@[d,e,f]@fs \<Longrightarrow>
\<exists>as a b bs. xs = as@[a,b]@bs \<and> d = F (sum_list as) \<and>
e = F (sum_list (as@[a])) \<and> f = F (sum_list (as@[a,b]))"
proof (induct xs fs rule: list_induct2_snoc)
case (Nil2 xs)
show ?case
proof (cases xs rule: rev_cases)
case Nil with Nil2 show ?thesis by simp
next
case (snoc ys y)
with Nil2 have y: "map F (sums ys) = ds@[d,e]" "f = F (sum_list (ys@[y]))"
by (auto simp add: sums_snoc)
from y(1) obtain as a bs where asabs:
"ys = as@[a]@bs" "map F (sums as) = ds@[d]"
"d = F (sum_list as)" "e = F (sum_list (as@[a]))"
using pullback_sums_map_middle2[of F ys ds]
by fastforce
have "bs = []"
proof-
from y(1) asabs(1,2) have "Suc (length bs) = Suc 0"
by (auto simp add: sums_append2 map_butlast length_sums[THEN sym])
thus ?thesis by fast
qed
with snoc asabs(1) y(2) have "xs = as@[a,y]@[]" "f = F (sum_list (as@[a,y]))"
by auto
with asabs(3,4) show ?thesis by fast
qed
next
case snoc thus ?case by (fastforce simp add: sums_snoc)
qed simp
lemma pullback_sums_map_double_middle2:
assumes "map F (sums xs) = ds@[d,e]@es@[f,g]@gs"
shows "\<exists>as a bs b cs. xs = as@[a]@bs@[b]@cs \<and> d = F (sum_list as) \<and>
e = F (sum_list (as@[a])) \<and> f = F (sum_list (as@[a]@bs)) \<and>
g = F (sum_list (as@[a]@bs@[b]))"
proof-
from assms obtain As b cs where Asbcs:
"xs = As@[b]@cs" "map F (sums As) = ds@[d,e]@es@[f]"
"f = F (sum_list As)" "g = F (sum_list (As@[b]))"
using pullback_sums_map_middle2[of F xs "ds@[d,e]@es"]
by fastforce
from Asbcs show ?thesis
using pullback_sums_map_middle2[of F As ds d e "es@[f]"] by fastforce
qed
end (* context monoid_add *)
subsection \<open>Sums of alternating lists\<close>
lemma (in group_add) uminus_sum_list_alternating_order2:
"s+s=0 \<Longrightarrow> t+t=0 \<Longrightarrow> - sum_list (alternating_list n s t) =
sum_list (if even n then alternating_list n t s else alternating_list n s t)"
using uminus_sum_list_order2 set_alternating_list[of n] rev_alternating_list[of n s]
by fastforce
context monoid_add
begin
lemma alternating_order2_cancel_1left:
"s+s=0 \<Longrightarrow>
sum_list (s # (alternating_list (Suc n) s t)) = sum_list (alternating_list n t s)"
using add.assoc[of s s] alternating_list_Suc_Cons[of n s] by simp
lemma alternating_order2_cancel_2left:
"s+s=0 \<Longrightarrow> t+t=0 \<Longrightarrow>
sum_list (t # s # (alternating_list (Suc (Suc n)) s t)) =
sum_list (alternating_list n s t)"
using alternating_order2_cancel_1left[of s "Suc n"]
alternating_order2_cancel_1left[of t n]
by simp
lemma alternating_order2_even_cancel_right:
assumes st : "s+s=0" "t+t=0"
and even_n: "even n"
shows "m \<le> n \<Longrightarrow> sum_list (alternating_list n s t @ alternating_list m t s) =
sum_list (alternating_list (n-m) s t)"
proof (induct n arbitrary: m rule: nat_even_induct, rule even_n)
case (SucSuc k) with st show ?case
using alternating_order2_cancel_2left[of t s]
by (cases m rule: nat_cases_2Suc) auto
qed simp
end (* context monoid_add *)
subsection \<open>Conjugation in @{class group_add}\<close>
subsubsection \<open>Abbreviations and basic facts\<close>
context group_add
begin
abbreviation lconjby :: "'a\<Rightarrow>'a\<Rightarrow>'a"
where "lconjby x y \<equiv> x+y-x"
abbreviation rconjby :: "'a\<Rightarrow>'a\<Rightarrow>'a"
where "rconjby x y \<equiv> -x+y+x"
lemma lconjby_add: "lconjby (x+y) z = lconjby x (lconjby y z)"
by (auto simp add: algebra_simps)
lemma rconjby_add: "rconjby (x+y) z = rconjby y (rconjby x z)"
by (simp add: minus_add add.assoc[THEN sym])
lemma add_rconjby: "rconjby x y + rconjby x z = rconjby x (y+z)"
by (simp add: add.assoc)
lemma lconjby_uminus: "lconjby x (-y) = - lconjby x y"
using minus_unique[of "lconjby x y", THEN sym] by (simp add: algebra_simps)
lemma rconjby_uminus: "rconjby x (-y) = - rconjby x y"
using minus_unique[of "rconjby x y"] add_assoc4[of "rconjby x y" "-x" "-y" x] by simp
lemma lconjby_rconjby: "lconjby x (rconjby x y) = y"
by (simp add: algebra_simps)
lemma rconjby_lconjby: "rconjby x (lconjby x y) = y"
by (simp add: algebra_simps)
lemma lconjby_inj: "inj (lconjby x)"
using rconjby_lconjby by (fast intro: inj_on_inverseI)
lemma rconjby_inj: "inj (rconjby x)"
using lconjby_rconjby by (fast intro: inj_on_inverseI)
lemma lconjby_surj: "surj (lconjby x)"
using lconjby_rconjby surjI[of "lconjby x"] by fast
lemma lconjby_bij: "bij (lconjby x)"
unfolding bij_def using lconjby_inj lconjby_surj by fast
lemma the_inv_lconjby: "the_inv (lconjby x) = (rconjby x)"
using bij_betw_f_the_inv_into_f[OF lconjby_bij, of _ x] lconjby_rconjby
by (force intro: inj_onD[OF lconjby_inj, of x])
lemma lconjby_eq_conv_rconjby_eq: "w = lconjby x y \<Longrightarrow> y = rconjby x w"
using the_inv_lconjby the_inv_into_f_f[OF lconjby_inj] by force
lemma rconjby_order2: "s+s = 0 \<Longrightarrow> rconjby x s + rconjby x s = 0"
by (simp add: add_rconjby)
lemma rconjby_order2_eq_lconjby:
assumes "s+s=0"
shows "rconjby s = lconjby s"
proof-
have "rconjby s = lconjby (-s)" by simp
with assms show ?thesis using minus_unique by simp
qed
lemma lconjby_alternating_list_order2:
assumes "s+s=0" "t+t=0"
shows "lconjby (sum_list (alternating_list k s t)) (if even k then s else t) =
sum_list (alternating_list (Suc (2*k)) s t)"
proof (induct k rule: nat_induct_step2)
case (SucSuc m)
have "lconjby (sum_list (alternating_list (Suc (Suc m)) s t))
(if even (Suc (Suc m)) then s else t) = s + t +
lconjby (sum_list (alternating_list m s t)) (if even m then s else t) - t - s"
using alternating_list_SucSuc_ConsCons[of m s t]
by (simp add: algebra_simps)
also from assms SucSuc
have "\<dots> = sum_list (alternating_list (Suc (2*Suc (Suc m))) s t)"
using alternating_list_SucSuc_ConsCons[of "Suc (2*m)" s t]
sum_list.append[of "alternating_list (Suc (2*Suc m)) s t" "[t]"]
by (simp add: algebra_simps)
finally show ?case by fast
qed (auto simp add: assms(1) algebra_simps)
end (* context group_add *)
subsubsection \<open>The conjugation sequence\<close>
text \<open>
Given a list in @{class group_add}, we create a new list by conjugating each term by all the
previous terms. This sequence arises in Coxeter systems.
\<close>
context group_add
begin
primrec lconjseq :: "'a list \<Rightarrow> 'a list"
where
"lconjseq [] = []"
| "lconjseq (x#xs) = x # (map (lconjby x) (lconjseq xs))"
lemma length_lconjseq: "length (lconjseq xs) = length xs"
by (induct xs) auto
lemma lconjseq_snoc: "lconjseq (xs@[x]) = lconjseq xs @ [lconjby (sum_list xs) x]"
by (induct xs) (auto simp add: lconjby_add)
lemma lconjseq_append:
"lconjseq (xs@ys) = lconjseq xs @ (map (lconjby (sum_list xs)) (lconjseq ys))"
proof (induct ys rule: rev_induct)
case (snoc y ys) thus ?case
using lconjseq_snoc[of "xs@ys"] lconjseq_snoc[of ys] by (simp add: lconjby_add)
qed simp
lemma lconjseq_alternating_order2_repeats':
fixes s t :: 'a
defines altst: "altst \<equiv> \<lambda>n. alternating_list n s t"
and altts: "altts \<equiv> \<lambda>n. alternating_list n t s"
assumes st : "s+s=0" "t+t=0" "(s+t)+^k = 0"
shows "map (lconjby (sum_list (altst k)))
(lconjseq (if even k then altst m else altts m)) = lconjseq (altst m)"
proof (induct m)
case (Suc j)
with altst altts
have "map (lconjby (sum_list (altst k)))
(lconjseq (if even k then altst (Suc j) else altts (Suc j))) =
lconjseq (altst j) @
[lconjby (sum_list (altst k @ (if even k then altst j else altts j)))
(if even k then (if even j then s else t) else (if even j then t else s))]"
by (auto simp add: lconjseq_snoc lconjby_add)
also from altst altts st(1,2)
have "\<dots> = lconjseq (altst j) @ [sum_list (altst (Suc (2*(k+j))))]"
using lconjby_alternating_list_order2[of s t "k+j"]
by (cases "even k")
(auto simp add: alternating_list_append[of k])
finally show ?case using altst st
by (auto simp add:
alternating_list_append(1)[THEN sym]
alternating_sum_list_conv_nataction
lconjby_alternating_list_order2 lconjseq_snoc
)
qed (simp add: altst altts)
lemma lconjseq_alternating_order2_repeats:
fixes s t :: 'a and k :: nat
defines altst: "altst \<equiv> \<lambda>n. alternating_list n s t"
and altts: "altts \<equiv> \<lambda>n. alternating_list n t s"
assumes st: "s+s=0" "t+t=0" "(s+t)+^k = 0"
shows "lconjseq (altst (2*k)) = lconjseq (altst k) @ lconjseq (altst k)"
proof-
from altst altts
have "lconjseq (altst (2*k)) = lconjseq (altst k) @
map (lconjby (sum_list (altst k)))
(lconjseq (if even k then altst k else altts k))"
using alternating_list_append[THEN sym, of k k s t]
by (auto simp add: times2_conv_add lconjseq_append)
with altst altts st show ?thesis
using lconjseq_alternating_order2_repeats'[of s t k k] by auto
qed
lemma even_count_lconjseq_alternating_order2:
fixes s t :: 'a
assumes "s+s=0" "t+t=0" "(s+t)+^k = 0"
shows "even (count_list (lconjseq (alternating_list (2*k) s t)) x)"
proof-
define xs where xs: "xs \<equiv> lconjseq (alternating_list (2*k) s t)"
with assms obtain as where "xs = as@as"
using lconjseq_alternating_order2_repeats by fast
hence "count_list xs x = 2 * (count_list as x)"
by (simp add: count_list_append times2_conv_add)
with xs show ?thesis by simp
qed
lemma order2_hd_in_lconjseq_deletion:
shows "s+s=0 \<Longrightarrow> s \<in> set (lconjseq ss)
\<Longrightarrow> \<exists>as b bs. ss = as@[b]@bs \<and> sum_list (s#ss) = sum_list (as@bs)"
proof (induct ss arbitrary: s rule: rev_induct)
case (snoc t ts) show ?case
proof (cases "s \<in> set (lconjseq ts)")
case True
with snoc(1,2) obtain as b bs
where asbbs: "ts = as @[b]@bs" "sum_list (s#ts) = sum_list (as@bs)"
by fastforce
from asbbs(2) have "sum_list (s#ts@[t]) = sum_list (as@(bs@[t]))"
using sum_list.append[of "s#ts" "[t]"] sum_list.append[of "as@bs" "[t]"] by simp
with asbbs(1) show ?thesis by fastforce
next
case False
with snoc(3) have s: "s = lconjby (sum_list ts) t" by (simp add: lconjseq_snoc)
with snoc(2) have "t+t=0"
using lconjby_eq_conv_rconjby_eq[of s "sum_list ts" t]
rconjby_order2[of s "sum_list ts"]
by simp
moreover from s have "sum_list (s#ts@[t]) = sum_list ts + t + t"
using add.assoc[of "sum_list ts + t - sum_list ts" "sum_list ts"]
by (simp add: algebra_simps)
ultimately have "sum_list (s#ts@[t]) = sum_list (ts@[])"
by (simp add: algebra_simps)
thus ?thesis by fast
qed
qed simp
end (* context group_add *)
subsubsection \<open>The action on signed @{class group_add} elements\<close>
text \<open>
Here we construct an action of a group on itself by conjugation, where group elements are
endowed with an auxiliary sign by pairing with a boolean element. In multiple applications of
this action, the auxiliary sign helps keep track of how many times the elements conjugating and
being conjugated are the same. This action arises in exploring reduced expressions of group
elements as words in a set of generators of order two (in particular, in a Coxeter group).
\<close>
type_synonym 'a signed = "'a\<times>bool"
definition signed_funaction :: "('a\<Rightarrow>'a\<Rightarrow>'a) \<Rightarrow> 'a \<Rightarrow> 'a signed \<Rightarrow> 'a signed"
where "signed_funaction f s x \<equiv> map_prod (f s) (\<lambda>b. b \<noteq> (fst x = s)) x"
\<comment> \<open>so the sign of @{term x} is flipped precisely when its first component is equal to
@{term s}\<close>
context group_add
begin
abbreviation "signed_lconjaction \<equiv> signed_funaction lconjby"
abbreviation "signed_rconjaction \<equiv> signed_funaction rconjby"
lemmas signed_lconjactionD = signed_funaction_def[of lconjby]
lemmas signed_rconjactionD = signed_funaction_def[of rconjby]
abbreviation signed_lconjpermutation :: "'a \<Rightarrow> 'a signed permutation"
where "signed_lconjpermutation s \<equiv> Abs_permutation (signed_lconjaction s)"
abbreviation signed_list_lconjaction :: "'a list \<Rightarrow> 'a signed \<Rightarrow> 'a signed"
where "signed_list_lconjaction ss \<equiv> foldr signed_lconjaction ss"
lemma signed_lconjaction_fst: "fst (signed_lconjaction s x) = lconjby s (fst x)"
using signed_lconjactionD by simp
lemma signed_lconjaction_rconjaction:
"signed_lconjaction s (signed_rconjaction s x) = x"
proof-
obtain a::'a and b::bool where "x = (a,b)" by fastforce
thus ?thesis
using signed_lconjactionD signed_rconjactionD injD[OF rconjby_inj, of s a]
lconjby_rconjby[of s a]
by auto
qed
lemma signed_rconjaction_by_order2_eq_lconjaction:
"s+s=0 \<Longrightarrow> signed_rconjaction s = signed_lconjaction s"
using signed_funaction_def[of lconjby s] signed_funaction_def[of rconjby s]
rconjby_order2_eq_lconjby[of s]
by auto
lemma inj_signed_lconjaction: "inj (signed_lconjaction s)"
proof (rule injI)
fix x y assume 1: "signed_lconjaction s x = signed_lconjaction s y"
moreover obtain a1 a2 :: 'a and b1 b2 :: bool
where xy: "x = (a1,b1)" "y = (a2,b2)"
by fastforce
ultimately show "x=y"
using injD[OF lconjby_inj, of s a1 a2] signed_lconjactionD
by (cases "a1=s" "a2=s" rule: two_cases) auto
qed
lemma surj_signed_lconjaction: "surj (signed_lconjaction s)"
using signed_lconjaction_rconjaction[THEN sym] by fast
lemma bij_signed_lconjaction: "bij (signed_lconjaction s)"
using inj_signed_lconjaction surj_signed_lconjaction by (fast intro: bijI)
lemma the_inv_signed_lconjaction:
"the_inv (signed_lconjaction s) = signed_rconjaction s"
proof
fix x
show "the_inv (signed_lconjaction s) x = signed_rconjaction s x"
proof (rule the_inv_into_f_eq, rule inj_signed_lconjaction)
show "signed_lconjaction s (signed_rconjaction s x) = x"
using signed_lconjaction_rconjaction by fast
qed (simp add: surj_signed_lconjaction)
qed
lemma the_inv_signed_lconjaction_by_order2:
"s+s=0 \<Longrightarrow> the_inv (signed_lconjaction s) = signed_lconjaction s"
using the_inv_signed_lconjaction signed_rconjaction_by_order2_eq_lconjaction
by simp
lemma signed_list_lconjaction_fst:
"fst (signed_list_lconjaction ss x) = lconjby (sum_list ss) (fst x)"
using signed_lconjaction_fst lconjby_add by (induct ss) auto
lemma signed_list_lconjaction_snd:
shows "\<forall>s\<in>set ss. s+s=0 \<Longrightarrow> snd (signed_list_lconjaction ss x)
= (if even (count_list (lconjseq (rev ss)) (fst x)) then snd x else \<not>snd x)"
proof (induct ss)
case (Cons s ss) hence prevcase:
"snd (signed_list_lconjaction ss x) =
(if even (count_list (lconjseq (rev ss)) (fst x)) then snd x else \<not> snd x)"
by simp
have 1: "snd (signed_list_lconjaction (s # ss) x) =
snd (signed_lconjaction s (signed_list_lconjaction ss x))"
by simp
show ?case
proof (cases "fst (signed_list_lconjaction ss x) = s")
case True
with 1 prevcase
have "snd (signed_list_lconjaction (s # ss) x) =
(if even (count_list (lconjseq (rev ss)) (fst x)) then \<not> snd x else snd x)"
by (simp add: signed_lconjactionD)
with True Cons(2) show ?thesis
by (simp add:
signed_list_lconjaction_fst lconjby_eq_conv_rconjby_eq
uminus_sum_list_order2[THEN sym] lconjseq_snoc count_list_snoc
)
next
case False
hence "rconjby (sum_list ss) (lconjby (sum_list ss) (fst x)) \<noteq>
rconjby (sum_list ss) s"
by (simp add: signed_list_lconjaction_fst)
with Cons(2)
have "count_list (lconjseq (rev (s#ss))) (fst x) =
count_list (lconjseq (rev ss)) (fst x)"
by (simp add:
rconjby_lconjby uminus_sum_list_order2[THEN sym]
lconjseq_snoc count_list_snoc
)
moreover from False 1 prevcase
have "snd (signed_list_lconjaction (s # ss) x) =
(if even (count_list (lconjseq (rev ss)) (fst x)) then snd x else \<not> snd x)"
by (simp add: signed_lconjactionD)
ultimately show ?thesis by simp
qed
qed simp
end (* context group_add *)
subsection \<open>Cosets\<close>
subsubsection \<open>Basic facts\<close>
lemma set_zero_plus' [simp]: "(0::'a::monoid_add) +o C = C"
\<comment> \<open>lemma @{text "Set_Algebras.set_zero_plus"} is restricted to types of class
@{class comm_monoid_add}; here is a version in @{class monoid_add}.\<close>
by (auto simp add: elt_set_plus_def)
lemma lcoset_0: "(w::'a::monoid_add) +o 0 = {w}"
using elt_set_plus_def[of w] by simp
lemma lcoset_refl: "(0::'a::monoid_add) \<in> A \<Longrightarrow> a \<in> a +o A"
using elt_set_plus_def by force
lemma lcoset_eq_reps_subset:
"(a::'a::group_add) +o A \<subseteq> a +o B \<Longrightarrow> A \<subseteq> B"
using elt_set_plus_def[of a] by auto
lemma lcoset_eq_reps: "(a::'a::group_add) +o A = a +o B \<Longrightarrow> A = B"
using lcoset_eq_reps_subset[of a A B] lcoset_eq_reps_subset[of a B A] by auto
lemma lcoset_inj_on: "inj ((+o) (a::'a::group_add))"
using lcoset_eq_reps inj_onI[of UNIV "(+o) a"] by auto
lemma lcoset_conv_set: "(a::'g::group_add) \<in> b +o A \<Longrightarrow> -b + a \<in> A"
by (auto simp add: elt_set_plus_def)
subsubsection \<open>The supset order on cosets\<close>
lemma supset_lbound_lcoset_shift:
"supset_lbound_of X Y B \<Longrightarrow>
ordering.lbound_of (\<supseteq>) (a +o X) (a +o Y) (a +o B)"
using ordering.lbound_of_def[OF supset_poset, of X Y B]
by (fast intro: ordering.lbound_ofI supset_poset)
lemma supset_glbound_in_of_lcoset_shift:
fixes P :: "'a::group_add set set"
assumes "supset_glbound_in_of P X Y B"
shows "supset_glbound_in_of ((+o) a ` P) (a +o X) (a +o Y) (a +o B)"
using ordering.glbound_in_ofD_in[OF supset_poset, OF assms]
ordering.glbound_in_ofD_lbound[OF supset_poset, OF assms]
supset_lbound_lcoset_shift[of X Y B a]
supset_lbound_lcoset_shift[of "a +o X" "a +o Y" _ "-a"]
ordering.glbound_in_ofD_glbound[OF supset_poset, OF assms]
ordering.glbound_in_ofI[
OF supset_poset, of "a +o B" "(+o) a ` P" "a +o X" "a +o Y"
]
by (fastforce simp add: set_plus_rearrange2)
subsubsection \<open>The afforded partition\<close>
definition lcoset_rel :: "'a::{uminus,plus} set \<Rightarrow> ('a\<times>'a) set"
where "lcoset_rel A \<equiv> {(x,y). -x + y \<in> A}"
lemma lcoset_relI: "-x+y \<in> A \<Longrightarrow> (x,y) \<in> lcoset_rel A"
using lcoset_rel_def by fast
subsection \<open>Groups\<close>
text \<open>We consider groups as closed sets in a type of class @{class group_add}.\<close>
subsubsection \<open>Locale definition and basic facts\<close>
locale Group =
fixes G :: "'g::group_add set"
assumes nonempty : "G \<noteq> {}"
and diff_closed: "\<And>g h. g \<in> G \<Longrightarrow> h \<in> G \<Longrightarrow> g - h \<in> G"
begin
abbreviation Subgroup :: "'g set \<Rightarrow> bool"
where "Subgroup H \<equiv> Group H \<and> H \<subseteq> G"
lemma SubgroupD1: "Subgroup H \<Longrightarrow> Group H" by fast
lemma zero_closed : "0 \<in> G"
proof-
from nonempty obtain g where "g \<in> G" by fast
hence "g - g \<in> G" using diff_closed by fast
thus ?thesis by simp
qed
lemma uminus_closed: "g\<in>G \<Longrightarrow> -g\<in>G"
using zero_closed diff_closed[of 0 g] by simp
lemma add_closed: "g\<in>G \<Longrightarrow> h\<in>G \<Longrightarrow> g+h \<in> G"
using uminus_closed[of h] diff_closed[of g "-h"] by simp
lemma uminus_add_closed: "g \<in> G \<Longrightarrow> h \<in> G \<Longrightarrow> -g + h \<in> G"
using uminus_closed add_closed by fast
lemma lconjby_closed: "g\<in>G \<Longrightarrow> x\<in>G \<Longrightarrow> lconjby g x \<in> G"
using add_closed diff_closed by fast
lemma lconjby_set_closed: "g\<in>G \<Longrightarrow> A\<subseteq>G \<Longrightarrow> lconjby g ` A \<subseteq> G"
using lconjby_closed by fast
lemma set_lconjby_subset_closed:
"H\<subseteq>G \<Longrightarrow> A\<subseteq>G \<Longrightarrow> (\<Union>h\<in>H. lconjby h ` A) \<subseteq> G"
using lconjby_set_closed[of _ A] by fast
lemma sum_list_map_closed: "set (map f as) \<subseteq> G \<Longrightarrow> (\<Sum>a\<leftarrow>as. f a) \<in> G"
using zero_closed add_closed by (induct as) auto
lemma sum_list_closed: "set as \<subseteq> G \<Longrightarrow> sum_list as \<in> G"
using sum_list_map_closed by force
end (* context Group *)
subsubsection \<open>Sets with a suitable binary operation\<close>
text \<open>
We have chosen to only consider groups in types of class @{class group_add} so that we can take
advantage of all the algebra lemmas already proven in @{theory HOL.Groups}, as well as
constructs like @{const sum_list}. The following locale builds a bridge between this restricted
view of groups and the usual notion of a binary operation on a set satisfying the group axioms,
by constructing an injective map into type @{type permutation} (which is of class
@{class group_add} with respect to the composition operation) that respects the group operation.
This bridge will be necessary to define quotient groups, in particular.
\<close>
locale BinOpSetGroup =
fixes G :: "'a set"
and binop :: "'a\<Rightarrow>'a\<Rightarrow>'a"
and e :: "'a"
assumes closed : "g\<in>G \<Longrightarrow> h\<in>G \<Longrightarrow> binop g h \<in> G"
and assoc :
"\<lbrakk> g\<in>G; h\<in>G; k\<in>G \<rbrakk> \<Longrightarrow> binop (binop g h) k = binop g (binop h k)"
and identity: "e\<in>G" "g\<in>G \<Longrightarrow> binop g e = g" "g\<in>G \<Longrightarrow> binop e g = g"
and inverses: "g\<in>G \<Longrightarrow> \<exists>h\<in>G. binop g h = e \<and> binop h g = e"
begin
lemma unique_identity1: "g\<in>G \<Longrightarrow> \<forall>x\<in>G. binop g x = x \<Longrightarrow> g = e"
using identity(1,2) by auto
lemma unique_inverse:
assumes "g\<in>G"
shows "\<exists>!h. h\<in>G \<and> binop g h = e \<and> binop h g = e"
proof (rule ex_ex1I)
from assms show "\<exists>h. h \<in> G \<and> binop g h = e \<and> binop h g = e"
using inverses by fast
next
fix h k
assume "h\<in>G \<and> binop g h = e \<and> binop h g = e" "k\<in>G \<and>
binop g k = e \<and> binop k g = e"
hence h: "h\<in>G" "binop g h = e" "binop h g = e"
and k: "k\<in>G" "binop g k = e" "binop k g = e"
by auto
from assms h(1,3) k(1,2) show "h=k" using identity(2,3) assoc by force
qed
abbreviation "G_perm g \<equiv> restrict1 (binop g) G"
definition Abs_G_perm :: "'a \<Rightarrow> 'a permutation"
where "Abs_G_perm g \<equiv> Abs_permutation (G_perm g)"
abbreviation "\<pp> \<equiv> Abs_G_perm" \<comment> \<open>the injection into type @{type permutation}\<close>
abbreviation "\<ii>\<pp> \<equiv> the_inv_into G \<pp>" \<comment> \<open>the reverse correspondence\<close>
abbreviation "pG \<equiv> \<pp>`G" \<comment> \<open>the resulting @{const Group} of type @{type permutation}\<close>
lemma G_perm_comp:
"g\<in>G \<Longrightarrow> h\<in>G \<Longrightarrow> G_perm g \<circ> G_perm h = G_perm (binop g h)"
using closed by (auto simp add: assoc)
definition the_inverse :: "'a \<Rightarrow> 'a"
where "the_inverse g \<equiv> (THE h. h\<in>G \<and> binop g h = e \<and> binop h g = e)"
abbreviation "\<ii> \<equiv> the_inverse"
lemma the_inverseD:
assumes "g\<in>G"
shows "\<ii> g \<in> G" "binop g (\<ii> g) = e" "binop (\<ii> g) g = e"
using assms theI'[OF unique_inverse]
unfolding the_inverse_def
by auto
lemma binop_G_comp_binop_\<ii>G: "g\<in>G \<Longrightarrow> x\<in>G \<Longrightarrow> binop g (binop (\<ii> g) x) = x"
using the_inverseD(1) assoc[of g "\<ii> g" x] by (simp add: identity(3) the_inverseD(2))
lemma bij_betw_binop_G:
assumes "g\<in>G"
shows "bij_betw (binop g) G G"
unfolding bij_betw_def
proof
show "inj_on (binop g) G"
proof (rule inj_onI)
fix h k assume hk: "h\<in>G" "k\<in>G" "binop g h = binop g k"
with assms have "binop (binop (\<ii> g) g) h = binop (binop (\<ii> g) g) k"
using the_inverseD(1) by (simp add: assoc)
with assms hk(1,2) show "h=k" using the_inverseD(3) identity by simp
qed
show "binop g ` G = G"
proof
from assms show "binop g ` G \<subseteq> G" using closed by fast
from assms show "binop g ` G \<supseteq> G"
using binop_G_comp_binop_\<ii>G[THEN sym] the_inverseD(1) closed by fast
qed
qed
lemma the_inv_into_G_binop_G:
assumes "g\<in>G" "x\<in>G"
shows "the_inv_into G (binop g) x = binop (\<ii> g) x"
proof (rule the_inv_into_f_eq)
from assms(1) show "inj_on (binop g) G"
using bij_betw_imp_inj_on[OF bij_betw_binop_G] by fast
from assms show "binop g (binop (\<ii> g) x) = x"
using binop_G_comp_binop_\<ii>G by fast
from assms show "binop (\<ii> g) x \<in> G" using closed the_inverseD(1) by fast
qed
lemma restrict1_the_inv_into_G_binop_G:
"g\<in>G \<Longrightarrow> restrict1 (the_inv_into G (binop g)) G = G_perm (\<ii> g)"
using the_inv_into_G_binop_G by auto
lemma bij_G_perm: "g\<in>G \<Longrightarrow> bij (G_perm g)"
using set_permutation_bij_restrict1 bij_betw_binop_G by fast
lemma G_perm_apply: "g\<in>G \<Longrightarrow> x\<in>G \<Longrightarrow> \<pp> g \<rightarrow> x = binop g x"
using Abs_G_perm_def Abs_permutation_inverse bij_G_perm by fastforce
lemma G_perm_apply_identity: "g\<in>G \<Longrightarrow> \<pp> g \<rightarrow> e = g"
using G_perm_apply identity(1,2) by simp
lemma the_inv_G_perm:
"g\<in>G \<Longrightarrow> the_inv (G_perm g) = G_perm (\<ii> g)"
using set_permutation_the_inv_restrict1 bij_betw_binop_G
restrict1_the_inv_into_G_binop_G
by fastforce
lemma Abs_G_perm_diff:
"g\<in>G \<Longrightarrow> h\<in>G \<Longrightarrow> \<pp> g - \<pp> h = \<pp> (binop g (\<ii> h))"
using Abs_G_perm_def minus_permutation_abs_eq[OF bij_G_perm bij_G_perm]
the_inv_G_perm G_perm_comp the_inverseD(1)
by simp
lemma Group: "Group pG"
using identity(1) Abs_G_perm_diff the_inverseD(1) closed by unfold_locales auto
lemma inj_on_\<pp>_G: "inj_on \<pp> G"
proof (rule inj_onI)
fix x y assume xy: "x\<in>G" "y\<in>G" "\<pp> x = \<pp> y"
hence "Abs_permutation (G_perm (binop x (\<ii> y))) = Abs_permutation id"
using Abs_G_perm_diff Abs_G_perm_def
by (fastforce simp add: zero_permutation.abs_eq)
moreover from xy(1,2) have 1: "binop x (\<ii> y) \<in> G"
using bij_id closed the_inverseD(1) by fast
ultimately have 2: "G_perm (binop x (\<ii> y)) = id"
using Abs_permutation_inject[of "G_perm (binop x (\<ii> y))"] bij_G_perm bij_id
by simp
have "\<forall>z\<in>G. binop (binop x (\<ii> y)) z = z"
proof
fix z assume "z\<in>G"
thus "binop (binop x (\<ii> y)) z = z" using fun_cong[OF 2, of z] by simp
qed
with xy(1,2) have "binop x (binop (\<ii> y) y) = y"
using unique_identity1[OF 1] the_inverseD(1) by (simp add: assoc)
with xy(1,2) show "x = y" using the_inverseD(3) identity(2) by simp
qed
lemma homs:
"\<And>g h. g\<in>G \<Longrightarrow> h\<in>G \<Longrightarrow> \<pp> (binop g h) = \<pp> g + \<pp> h"
"\<And>x y. x\<in>pG \<Longrightarrow> y\<in>pG \<Longrightarrow> binop (\<ii>\<pp> x) (\<ii>\<pp> y) = \<ii>\<pp> (x+y)"
proof-
show 1: "\<And>g h. g\<in>G \<Longrightarrow> h\<in>G \<Longrightarrow> \<pp> (binop g h) = \<pp> g + \<pp> h"
using Abs_G_perm_def G_perm_comp
plus_permutation_abs_eq[OF bij_G_perm bij_G_perm]
by simp
show "\<And>x y. x\<in>pG \<Longrightarrow> y\<in>pG \<Longrightarrow> binop (\<ii>\<pp> x) (\<ii>\<pp> y) = \<ii>\<pp> (x+y)"
proof-
fix x y assume "x\<in>pG" "y\<in>pG"
moreover hence "\<ii>\<pp> (\<pp> (binop (\<ii>\<pp> x) (\<ii>\<pp> y))) = \<ii>\<pp> (x + y)"
using 1 the_inv_into_into[OF inj_on_\<pp>_G] f_the_inv_into_f[OF inj_on_\<pp>_G]
by simp
ultimately show "binop (\<ii>\<pp> x) (\<ii>\<pp> y) = \<ii>\<pp> (x+y)"
using the_inv_into_into[OF inj_on_\<pp>_G] closed the_inv_into_f_f[OF inj_on_\<pp>_G]
by simp
qed
qed
lemmas inv_correspondence_into =
the_inv_into_into[OF inj_on_\<pp>_G, of _ G, simplified]
lemma inv_correspondence_conv_apply: "x \<in> pG \<Longrightarrow> \<ii>\<pp> x = x\<rightarrow>e"
using G_perm_apply_identity inj_on_\<pp>_G by (auto intro: the_inv_into_f_eq)
end (* context BinOpSetGroup *)
subsubsection \<open>Cosets of a @{const Group}\<close>
context Group
begin
lemma lcoset_refl: "a \<in> a +o G"
using lcoset_refl zero_closed by fast
lemma lcoset_el_reduce:
assumes "a \<in> G"
shows "a +o G = G"
proof (rule seteqI)
fix x assume "x \<in> a +o G"
from this obtain g where "g\<in>G" "x = a+g" using elt_set_plus_def[of a] by auto
with assms show "x\<in>G" by (simp add: add_closed)
next
fix x assume "x\<in>G"
with assms have "-a + x \<in> G" by (simp add: uminus_add_closed)
thus "x \<in> a +o G" using elt_set_plus_def by force
qed
lemma lcoset_el_reduce0: "0 \<in> a +o G \<Longrightarrow> a +o G = G"
using elt_set_plus_def[of a G] minus_unique uminus_closed[of "-a"]
lcoset_el_reduce
by fastforce
lemma lcoset_subgroup_imp_eq_reps:
"Group H \<Longrightarrow> w +o H \<subseteq> w' +o G \<Longrightarrow> w' +o G = w +o G"
using Group.lcoset_refl[of H w] lcoset_conv_set[of w] lcoset_el_reduce
set_plus_rearrange2[of w' "-w'+w" G]
by force
lemma lcoset_closed: "a\<in>G \<Longrightarrow> A\<subseteq>G \<Longrightarrow> a +o A \<subseteq> G"
using elt_set_plus_def[of a] add_closed by auto
lemma lcoset_rel_sym: "sym (lcoset_rel G)"
proof (rule symI)
fix a b show "(a,b) \<in> lcoset_rel G \<Longrightarrow> (b,a) \<in> lcoset_rel G"
using uminus_closed minus_add[of "-a" b] lcoset_rel_def[of G] by fastforce
qed
lemma lcoset_rel_trans: "trans (lcoset_rel G)"
proof (rule transI)
fix x y z assume xy: "(x,y) \<in> lcoset_rel G" and yz: "(y,z) \<in> lcoset_rel G"
from this obtain g g' where "g\<in>G" "-x+y = g" "g'\<in>G" "-y+z = g'"
using lcoset_rel_def[of G] by fast
thus "(x, z) \<in> lcoset_rel G"
using add.assoc[of g "-y" z] add_closed lcoset_rel_def[of G] by auto
qed
abbreviation LCoset_rel :: "'g set \<Rightarrow> ('g\<times>'g) set"
where "LCoset_rel H \<equiv> lcoset_rel H \<inter> (G\<times>G)"
lemma refl_on_LCoset_rel: "0\<in>H \<Longrightarrow> refl_on G (LCoset_rel H)"
using lcoset_rel_def by (fastforce intro: refl_onI)
lemmas subgroup_refl_on_LCoset_rel =
refl_on_LCoset_rel[OF Group.zero_closed, OF SubgroupD1]
lemmas LCoset_rel_quotientI = quotientI[of _ G "LCoset_rel _"]
lemmas LCoset_rel_quotientE = quotientE[of _ G "LCoset_rel _"]
lemma lcoset_subgroup_rel_equiv:
"Subgroup H \<Longrightarrow> equiv G (LCoset_rel H)"
using Group.lcoset_rel_sym sym_sym sym_Int Group.lcoset_rel_trans trans_sym
trans_Int subgroup_refl_on_LCoset_rel
by (blast intro: equivI)
lemma trivial_LCoset: "H\<subseteq>G \<Longrightarrow> H = LCoset_rel H `` {0}"
using zero_closed unfolding lcoset_rel_def by auto
end (* context Group *)
subsubsection \<open>The @{const Group} generated by a set\<close>
inductive_set genby :: "'a::group_add set \<Rightarrow> 'a set" ("\<langle>_\<rangle>")
for S :: "'a set"
where
genby_0_closed : "0\<in>\<langle>S\<rangle>" \<comment> \<open>just in case @{term S} is empty\<close>
| genby_genset_closed: "s\<in>S \<Longrightarrow> s\<in>\<langle>S\<rangle>"
| genby_diff_closed : "w\<in>\<langle>S\<rangle> \<Longrightarrow> w'\<in>\<langle>S\<rangle> \<Longrightarrow> w - w' \<in> \<langle>S\<rangle>"
lemma genby_Group: "Group \<langle>S\<rangle>"
using genby_0_closed genby_diff_closed by unfold_locales fast
lemmas genby_uminus_closed = Group.uminus_closed [OF genby_Group]
lemmas genby_add_closed = Group.add_closed [OF genby_Group]
lemmas genby_uminus_add_closed = Group.uminus_add_closed [OF genby_Group]
lemmas genby_lcoset_refl = Group.lcoset_refl [OF genby_Group]
lemmas genby_lcoset_el_reduce = Group.lcoset_el_reduce [OF genby_Group]
lemmas genby_lcoset_el_reduce0 = Group.lcoset_el_reduce0 [OF genby_Group]
lemmas genby_lcoset_closed = Group.lcoset_closed [OF genby_Group]
lemmas genby_lcoset_subgroup_imp_eq_reps =
Group.lcoset_subgroup_imp_eq_reps[OF genby_Group, OF genby_Group]
lemma genby_genset_subset: "S \<subseteq> \<langle>S\<rangle>"
using genby_genset_closed by fast
lemma genby_uminus_genset_subset: "uminus ` S \<subseteq> \<langle>S\<rangle>"
using genby_genset_subset genby_uminus_closed by auto
lemma genby_in_sum_list_lists:
fixes S
defines S_sum_lists: "S_sum_lists \<equiv> (\<Union>ss\<in>lists (S \<union> uminus ` S). {sum_list ss})"
shows "w \<in> \<langle>S\<rangle> \<Longrightarrow> w \<in> S_sum_lists"
proof (erule genby.induct)
have "0 = sum_list []" by simp
with S_sum_lists show "0 \<in> S_sum_lists" by blast
next
fix s assume "s\<in>S"
hence "[s] \<in> lists (S \<union> uminus ` S)" by simp
moreover have "s = sum_list [s]" by simp
ultimately show "s \<in> S_sum_lists" using S_sum_lists by blast
next
fix w w' assume ww': "w \<in> S_sum_lists" "w' \<in> S_sum_lists"
with S_sum_lists obtain ss ts
where ss: "ss \<in> lists (S \<union> uminus ` S)" "w = sum_list ss"
and ts: "ts \<in> lists (S \<union> uminus ` S)" "w' = sum_list ts"
by fastforce
from ss(2) ts(2) have "w-w' = sum_list (ss @ map uminus (rev ts))"
by (simp add: diff_conv_add_uminus uminus_sum_list)
moreover from ss(1) ts(1)
have "ss @ map uminus (rev ts) \<in> lists (S \<union> uminus ` S)"
by fastforce
ultimately show "w - w' \<in> S_sum_lists" using S_sum_lists by fast
qed
lemma sum_list_lists_in_genby: "ss \<in> lists (S \<union> uminus ` S) \<Longrightarrow> sum_list ss \<in> \<langle>S\<rangle>"
proof (induct ss)
case Nil show ?case using genby_0_closed by simp
next
case (Cons s ss) thus ?case
using genby_genset_subset[of S] genby_uminus_genset_subset
genby_add_closed[of s S "sum_list ss"]
by auto
qed
lemma sum_list_lists_in_genby_sym:
"uminus ` S \<subseteq> S \<Longrightarrow> ss \<in> lists S \<Longrightarrow> sum_list ss \<in> \<langle>S\<rangle>"
using sum_list_lists_in_genby by fast
lemma genby_eq_sum_lists: "\<langle>S\<rangle> = (\<Union>ss\<in>lists (S \<union> uminus ` S). {sum_list ss})"
using genby_in_sum_list_lists sum_list_lists_in_genby by fast
lemma genby_mono: "T \<subseteq> S \<Longrightarrow> \<langle>T\<rangle> \<subseteq> \<langle>S\<rangle>"
using genby_eq_sum_lists[of T] genby_eq_sum_lists[of S] by force
lemma (in Group) genby_closed:
assumes "S \<subseteq> G"
shows "\<langle>S\<rangle> \<subseteq> G"
proof
fix x show "x \<in> \<langle>S\<rangle> \<Longrightarrow> x \<in> G"
proof (erule genby.induct, rule zero_closed)
from assms show "\<And>s. s\<in>S \<Longrightarrow> s\<in>G" by fast
show "\<And>w w'. w\<in>G \<Longrightarrow> w'\<in>G \<Longrightarrow> w-w' \<in> G" using diff_closed by fast
qed
qed
lemma (in Group) genby_subgroup: "S \<subseteq> G \<Longrightarrow> Subgroup \<langle>S\<rangle>"
using genby_closed genby_Group by simp
lemma genby_sym_eq_sum_lists:
"uminus ` S \<subseteq> S \<Longrightarrow> \<langle>S\<rangle> = (\<Union>ss\<in>lists S. {sum_list ss})"
using lists_mono genby_eq_sum_lists[of S] by force
lemma genby_empty': "w \<in> \<langle>{}\<rangle> \<Longrightarrow> w = 0"
proof (erule genby.induct) qed auto
lemma genby_order2':
assumes "s+s=0"
shows "w \<in> \<langle>{s}\<rangle> \<Longrightarrow> w = 0 \<or> w = s"
proof (erule genby.induct)
fix w w' assume "w = 0 \<or> w = s" "w' = 0 \<or> w' = s"
with assms show "w - w' = 0 \<or> w - w' = s"
by (cases "w'=0") (auto simp add: minus_unique)
qed auto
lemma genby_order2: "s+s=0 \<Longrightarrow> \<langle>{s}\<rangle> = {0,s}"
using genby_order2'[of s] genby_0_closed genby_genset_closed by auto
lemma genby_empty: "\<langle>{}\<rangle> = 0"
using genby_empty' genby_0_closed by auto
lemma genby_lcoset_order2: "s+s=0 \<Longrightarrow> w +o \<langle>{s}\<rangle> = {w,w+s}"
using elt_set_plus_def[of w] by (auto simp add: genby_order2)
lemma genby_lcoset_empty: "(w::'a::group_add) +o \<langle>{}\<rangle> = {w}"
proof-
have "\<langle>{}::'a set\<rangle> = (0::'a set)" using genby_empty by fast
thus ?thesis using lcoset_0 by simp
qed
lemma (in Group) genby_set_lconjby_set_lconjby_closed:
fixes A :: "'g set"
defines "S \<equiv> (\<Union>g\<in>G. lconjby g ` A)"
assumes "g\<in>G"
shows "x \<in> \<langle>S\<rangle> \<Longrightarrow> lconjby g x \<in> \<langle>S\<rangle>"
proof (erule genby.induct)
show "lconjby g 0 \<in> \<langle>S\<rangle>" using genby_0_closed by simp
from assms show "\<And>s. s \<in> S \<Longrightarrow> lconjby g s \<in> \<langle>S\<rangle>"
using add_closed genby_genset_closed[of _ S] by (force simp add: lconjby_add)
next
fix w w'
assume ww': "lconjby g w \<in> \<langle>S\<rangle>" "lconjby g w' \<in> \<langle>S\<rangle>"
have "lconjby g (w - w') = lconjby g w + lconjby g (-w')"
by (simp add: algebra_simps)
with ww' show "lconjby g (w - w') \<in> \<langle>S\<rangle>"
using lconjby_uminus[of g] diff_conv_add_uminus[of _ "lconjby g w'"]
genby_diff_closed
by fastforce
qed
lemma (in Group) genby_set_lconjby_set_rconjby_closed:
fixes A :: "'g set"
defines "S \<equiv> (\<Union>g\<in>G. lconjby g ` A)"
assumes "g\<in>G" "x \<in> \<langle>S\<rangle>"
shows "rconjby g x \<in> \<langle>S\<rangle>"
using assms uminus_closed genby_set_lconjby_set_lconjby_closed
by fastforce
subsubsection \<open>Homomorphisms and isomorphisms\<close>
locale GroupHom = Group G
for G :: "'g::group_add set"
+ fixes T :: "'g \<Rightarrow> 'h::group_add"
assumes hom : "g \<in> G \<Longrightarrow> g' \<in> G \<Longrightarrow> T (g + g') = T g + T g'"
and supp: "supp T \<subseteq> G"
begin
lemma im_zero: "T 0 = 0"
using zero_closed hom[of 0 0] add_diff_cancel[of "T 0" "T 0"] by simp
lemma im_uminus: "T (- g) = - T g"
using im_zero hom[of g "- g"] uminus_closed[of g] minus_unique[of "T g"]
uminus_closed[of "-g"] supp suppI_contra[of g T]
suppI_contra[of "-g" T]
by fastforce
lemma im_uminus_add: "g \<in> G \<Longrightarrow> g' \<in> G \<Longrightarrow> T (-g + g') = - T g + T g'"
by (simp add: uminus_closed hom im_uminus)
lemma im_diff: "g \<in> G \<Longrightarrow> g' \<in> G \<Longrightarrow> T (g - g') = T g - T g'"
using hom uminus_closed hom[of g "-g'"] im_uminus by simp
lemma im_lconjby: "x \<in> G \<Longrightarrow> g \<in> G \<Longrightarrow> T (lconjby x g) = lconjby (T x) (T g)"
using add_closed by (simp add: im_diff hom)
lemma im_sum_list_map:
"set (map f as) \<subseteq> G \<Longrightarrow> T (\<Sum>a\<leftarrow>as. f a) = (\<Sum>a\<leftarrow>as. T (f a))"
using hom im_zero sum_list_closed by (induct as) auto
lemma comp:
assumes "GroupHom H S" "T`G \<subseteq> H"
shows "GroupHom G (S \<circ> T)"
proof
fix g g' assume "g \<in> G" "g' \<in> G"
with hom assms(2) show "(S \<circ> T) (g + g') = (S \<circ> T) g + (S \<circ> T) g'"
using GroupHom.hom[OF assms(1)] by fastforce
next
from supp have "\<And>g. g \<notin> G \<Longrightarrow> (S \<circ> T) g = 0"
using suppI_contra GroupHom.im_zero[OF assms(1)] by fastforce
thus "supp (S \<circ> T) \<subseteq> G" using suppD_contra by fast
qed
end (* context GroupHom *)
definition ker :: "('a\<Rightarrow>'b::zero) \<Rightarrow> 'a set"
where "ker f = {a. f a = 0}"
lemma ker_subset_ker_restrict0: "ker f \<subseteq> ker (restrict0 f A)"
unfolding ker_def by auto
context GroupHom
begin
abbreviation "Ker \<equiv> ker T \<inter> G"
lemma uminus_add_in_Ker_eq_eq_im:
"g\<in>G \<Longrightarrow> h\<in>G \<Longrightarrow> (-g + h \<in> Ker) = (T g = T h)"
using neg_equal_iff_equal
by (simp add: uminus_add_closed ker_def im_uminus_add eq_neg_iff_add_eq_0)
end (* context GroupHom *)
locale UGroupHom = GroupHom UNIV T
for T :: "'g::group_add \<Rightarrow> 'h::group_add"
begin
lemmas im_zero = im_zero
lemmas im_uminus = im_uminus
lemma hom: "T (g+g') = T g + T g'"
using hom by simp
lemma im_diff: "T (g - g') = T g - T g'"
using im_diff by simp
lemma im_lconjby: "T (lconjby x g) = lconjby (T x) (T g)"
using im_lconjby by simp
lemma restrict0:
assumes "Group G"
shows "GroupHom G (restrict0 T G)"
proof (intro_locales, rule assms, unfold_locales)
from hom
show "\<And>g g'. g \<in> G \<Longrightarrow> g' \<in> G \<Longrightarrow>
restrict0 T G (g + g') = restrict0 T G g + restrict0 T G g'"
using Group.add_closed[OF assms]
by auto
show "supp (restrict0 T G) \<subseteq> G" using supp_restrict0[of G T] by fast
qed
end (* context UGroupHom *)
lemma UGroupHomI:
assumes "\<And>g g'. T (g + g') = T g + T g'"
shows "UGroupHom T"
using assms
by unfold_locales auto
locale GroupIso = GroupHom G T
for G :: "'g::group_add set"
and T :: "'g \<Rightarrow> 'h::group_add"
+ assumes inj_on: "inj_on T G"
lemma (in GroupHom) isoI:
assumes "\<And>k. k\<in>G \<Longrightarrow> T k = 0 \<Longrightarrow> k=0"
shows "GroupIso G T"
proof (unfold_locales, rule inj_onI)
fix x y from assms show "\<lbrakk> x\<in>G; y\<in>G; T x = T y \<rbrakk> \<Longrightarrow> x = y"
using im_diff diff_closed by force
qed
text \<open>
In a @{const BinOpSetGroup}, any map from the set into a type of class @{class group_add} that respects the
binary operation induces a @{const GroupHom}.
\<close>
abbreviation (in BinOpSetGroup) "lift_hom T \<equiv> restrict0 (T \<circ> \<ii>\<pp>) pG"
lemma (in BinOpSetGroup) lift_hom:
fixes T :: "'a \<Rightarrow> 'b::group_add"
assumes "\<forall>g\<in>G. \<forall>h\<in>G. T (binop g h) = T g + T h"
shows "GroupHom pG (lift_hom T)"
proof (intro_locales, rule Group, unfold_locales)
from assms
show "\<And>x y. x\<in>pG \<Longrightarrow> y\<in>pG \<Longrightarrow>
lift_hom T (x+y) = lift_hom T x + lift_hom T y"
using Group.add_closed[OF Group] inv_correspondence_into
by (simp add: homs(2)[THEN sym])
qed (rule supp_restrict0)
subsubsection \<open>Normal subgroups\<close>
definition rcoset_rel :: "'a::{minus,plus} set \<Rightarrow> ('a\<times>'a) set"
where "rcoset_rel A \<equiv> {(x,y). x-y \<in> A}"
context Group
begin
lemma rcoset_rel_conv_lcoset_rel:
"rcoset_rel G = map_prod uminus uminus ` (lcoset_rel G)"
proof (rule set_eqI)
fix x :: "'g\<times>'g"
obtain a b where ab: "x=(a,b)" by fastforce
hence "(x \<in> rcoset_rel G) = (a-b \<in> G)" using rcoset_rel_def by auto
also have "\<dots> = ( (-b,-a) \<in> lcoset_rel G )"
using uminus_closed lcoset_rel_def by fastforce
finally
show "(x \<in> rcoset_rel G) = (x \<in> map_prod uminus uminus ` (lcoset_rel G))"
using ab symD[OF lcoset_rel_sym] map_prod_def
by force
qed
lemma rcoset_rel_sym: "sym (rcoset_rel G)"
using rcoset_rel_conv_lcoset_rel map_prod_sym lcoset_rel_sym by simp
abbreviation RCoset_rel :: "'g set \<Rightarrow> ('g\<times>'g) set"
where "RCoset_rel H \<equiv> rcoset_rel H \<inter> (G\<times>G)"
definition normal :: "'g set \<Rightarrow> bool"
where "normal H \<equiv> (\<forall>g\<in>G. LCoset_rel H `` {g} = RCoset_rel H `` {g})"
lemma normalI:
assumes "Group H" "\<forall>g\<in>G. \<forall>h\<in>H. \<exists>h'\<in>H. g+h = h'+g"
"\<forall>g\<in>G. \<forall>h\<in>H. \<exists>h'\<in>H. h+g = g+h'"
shows "normal H"
unfolding normal_def
proof
fix g assume g: "g\<in>G"
show "LCoset_rel H `` {g} = RCoset_rel H `` {g}"
proof (rule seteqI)
fix x assume "x \<in> LCoset_rel H `` {g}"
with g have x: "x\<in>G" "-g+x \<in> H" unfolding lcoset_rel_def by auto
from g x(2) assms(2) obtain h where h: "h\<in>H" "g-x = -h"
by (fastforce simp add: algebra_simps)
with assms(1) g x(1) show "x \<in> RCoset_rel H `` {g}"
using Group.uminus_closed unfolding rcoset_rel_def by simp
next
fix x assume "x \<in> RCoset_rel H `` {g}"
with g have x: "x\<in>G" "g-x \<in> H" unfolding rcoset_rel_def by auto
with assms(3) obtain h where h: "h\<in>H" "-g+x = -h"
by (fastforce simp add: algebra_simps minus_add)
with assms(1) g x(1) show "x \<in> LCoset_rel H `` {g}"
using Group.uminus_closed unfolding lcoset_rel_def by simp
qed
qed
lemma normal_lconjby_closed:
"\<lbrakk> Subgroup H; normal H; g\<in>G; h\<in>H \<rbrakk> \<Longrightarrow> lconjby g h \<in> H"
using lcoset_relI[of g "g+h" H] add_closed[of g h] normal_def[of H]
symD[OF Group.rcoset_rel_sym, of H g "g+h"] rcoset_rel_def[of H]
by auto
lemma normal_rconjby_closed:
"\<lbrakk> Subgroup H; normal H; g\<in>G; h\<in>H \<rbrakk> \<Longrightarrow> rconjby g h \<in> H"
using normal_lconjby_closed[of H "-g" h] uminus_closed[of g] by auto
abbreviation "normal_closure A \<equiv> \<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>"
lemma (in Group) normal_closure:
assumes "A\<subseteq>G"
shows "normal (normal_closure A)"
proof (rule normalI, rule genby_Group)
show "\<forall>x\<in>G. \<forall>h\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>.
\<exists>h'\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. x + h = h' + x"
proof
fix x assume x: "x\<in>G"
show "\<forall>h\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>.
\<exists>h'\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. x + h = h' + x"
proof (rule ballI, erule genby.induct)
show "\<exists>h\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. x + 0 = h + x"
using genby_0_closed by force
next
fix s assume "s \<in> (\<Union>g\<in>G. lconjby g ` A)"
from this obtain g a where ga: "g\<in>G" "a\<in>A" "s = lconjby g a" by fast
from ga(3) have "x + s = lconjby x (lconjby g a) + x"
by (simp add: algebra_simps)
hence "x + s = lconjby (x+g) a + x" by (simp add: lconjby_add)
with x ga(1,2) show "\<exists>h\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. x + s = h + x"
using add_closed by (blast intro: genby_genset_closed)
next
fix w w'
assume w : "w \<in> \<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>"
"\<exists>h \<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. x + w = h + x"
and w': "w'\<in> \<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>"
"\<exists>h'\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. x + w' = h'+ x"
from w(2) w'(2) obtain h h'
where h : "h \<in> \<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>" "x + w = h + x"
and h': "h'\<in> \<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>" "x + w' = h'+ x"
by fast
have "x + (w - w') = x + w - (-x + (x + w'))"
by (simp add: algebra_simps)
also from h(2) h'(2) have "\<dots> = h + x + (-(h' + x) + x)"
by (simp add: algebra_simps)
also have "\<dots> = h + x + (-x + -h') + x"
by (simp add: minus_add add.assoc)
finally have "x + (w-w') = h - h' + x"
using add.assoc[of "h+x" "-x" "-h'"] by simp
with h(1) h'(1)
show "\<exists>h\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. x + (w - w') = h + x"
using genby_diff_closed
by fast
qed
qed
show "\<forall>x\<in>G. \<forall>h\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>.
\<exists>h'\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. h + x = x + h'"
proof
fix x assume x: "x\<in>G"
show "\<forall>h\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>.
\<exists>h'\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. h + x = x + h'"
proof (rule ballI, erule genby.induct)
show "\<exists>h\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. 0 + x = x + h"
using genby_0_closed by force
next
fix s assume "s \<in> (\<Union>g\<in>G. lconjby g ` A)"
from this obtain g a where ga: "g\<in>G" "a\<in>A" "s = lconjby g a" by fast
from ga(3) have "s + x = x + (((-x + g) + a) + -g) + x"
by (simp add: algebra_simps)
also have "\<dots> = x + (-x + g + a + -g + x)" by (simp add: add.assoc)
finally have "s + x = x + lconjby (-x+g) a"
by (simp add: algebra_simps lconjby_add)
with x ga(1,2) show "\<exists>h\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. s + x = x + h"
using uminus_add_closed by (blast intro: genby_genset_closed)
next
fix w w'
assume w : "w \<in> \<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>"
"\<exists>h \<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. w + x = x + h"
and w': "w'\<in> \<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>"
"\<exists>h'\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. w' + x = x + h'"
from w(2) w'(2) obtain h h'
where h : "h \<in> \<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>" "w + x = x + h"
and h': "h'\<in> \<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>" "w' + x = x + h'"
by fast
have "w - w' + x = w + x + (-x + -w') + x" by (simp add: algebra_simps)
also from h(2) h'(2) have "\<dots> = x + h + (-h'+-x) + x"
using minus_add[of w' x] minus_add[of x h'] by simp
finally have "w - w' + x = x + (h - h')" by (simp add: algebra_simps)
with h(1) h'(1) show "\<exists>h\<in>\<langle>\<Union>g\<in>G. lconjby g ` A\<rangle>. w - w' + x = x + h"
using genby_diff_closed by fast
qed
qed
qed
end (* context Group *)
subsubsection \<open>Quotient groups\<close>
text \<open>
Here we use the bridge built by @{const BinOpSetGroup} to make the quotient of a @{const Group}
by a normal subgroup into a @{const Group} itself.
\<close>
context Group
begin
lemma normal_quotient_add_well_defined:
assumes "Subgroup H" "normal H" "g\<in>G" "g'\<in>G"
shows "LCoset_rel H `` {g} + LCoset_rel H `` {g'} = LCoset_rel H `` {g+g'}"
proof (rule seteqI)
fix x assume "x \<in> LCoset_rel H `` {g} + LCoset_rel H `` {g'}"
from this obtain y z
where "y \<in> LCoset_rel H `` {g}" "z \<in> LCoset_rel H `` {g'}" "x = y+z"
unfolding set_plus_def
by fast
with assms show "x \<in> LCoset_rel H `` {g + g'}"
using lcoset_rel_def[of H] normal_lconjby_closed[of H g' "-g'+z"]
Group.add_closed
normal_rconjby_closed[of H g' "-g + y + (z - g')"]
add.assoc[of "-g'" "-g"]
add_closed lcoset_relI[of "g+g'" "y+z"]
by (fastforce simp add: add.assoc minus_add)
next
fix x assume "x \<in> LCoset_rel H `` {g + g'}"
moreover define h where "h \<equiv> -(g+g') + x"
moreover hence "x = g + (g' + h)"
using add.assoc[of "-g'" "-g" x] by (simp add: add.assoc minus_add)
ultimately show "x \<in> LCoset_rel H `` {g} + LCoset_rel H `` {g'}"
using assms(1,3,4) lcoset_rel_def[of H] add_closed
refl_onD[OF subgroup_refl_on_LCoset_rel, of H]
by force
qed
abbreviation "quotient_set H \<equiv> G // LCoset_rel H"
lemma BinOpSetGroup_normal_quotient:
assumes "Subgroup H" "normal H"
shows "BinOpSetGroup (quotient_set H) (+) H"
proof
from assms(1) have H0: "H = LCoset_rel H `` {0}"
using trivial_LCoset by auto
from assms(1) show "H \<in> quotient_set H"
using H0 zero_closed LCoset_rel_quotientI[of 0 H] by simp
fix x assume "x \<in> quotient_set H"
from this obtain gx where gx: "gx\<in>G" "x = LCoset_rel H `` {gx}"
by (fast elim: LCoset_rel_quotientE)
with assms(1,2) show "x+H = x" "H+x = x"
using normal_quotient_add_well_defined[of H gx 0]
normal_quotient_add_well_defined[of H 0 gx]
H0 zero_closed
by auto
from gx(1) have "LCoset_rel H `` {-gx} \<in> quotient_set H"
using uminus_closed by (fast intro: LCoset_rel_quotientI)
moreover from assms(1,2) gx
have "x + LCoset_rel H `` {-gx} = H" "LCoset_rel H `` {-gx} + x = H"
using H0 uminus_closed normal_quotient_add_well_defined
by auto
ultimately show "\<exists>x'\<in>quotient_set H. x + x' = H \<and> x' + x = H" by fast
fix y assume "y \<in> quotient_set H"
from this obtain gy where gy: "gy\<in>G" "y = LCoset_rel H `` {gy}"
by (fast elim: LCoset_rel_quotientE)
with assms gx show "x+y \<in> quotient_set H"
using add_closed normal_quotient_add_well_defined
by (auto intro: LCoset_rel_quotientI)
qed (rule add.assoc)
abbreviation "abs_lcoset_perm H \<equiv>
BinOpSetGroup.Abs_G_perm (quotient_set H) (+)"
abbreviation "abs_lcoset_perm_lift H g \<equiv> abs_lcoset_perm H (LCoset_rel H `` {g})"
abbreviation "abs_lcoset_perm_lift_arg_permutation g H \<equiv> abs_lcoset_perm_lift H g"
notation abs_lcoset_perm_lift_arg_permutation ("\<lceil>_|_\<rceil>" [51,51] 50)
end (* context Group *)
abbreviation "Group_abs_lcoset_perm_lift_arg_permutation G' g H \<equiv>
Group.abs_lcoset_perm_lift_arg_permutation G' g H"
notation Group_abs_lcoset_perm_lift_arg_permutation ("\<lceil>_|_|_\<rceil>" [51,51,51] 50)
context Group
begin
lemmas lcoset_perm_def =
BinOpSetGroup.Abs_G_perm_def[OF BinOpSetGroup_normal_quotient]
lemmas lcoset_perm_comp =
BinOpSetGroup.G_perm_comp[OF BinOpSetGroup_normal_quotient]
lemmas bij_lcoset_perm =
BinOpSetGroup.bij_G_perm[OF BinOpSetGroup_normal_quotient]
lemma trivial_lcoset_perm:
assumes "Subgroup H" "normal H" "h\<in>H"
shows "restrict1 ((+) (LCoset_rel H `` {h})) (quotient_set H) = id"
proof (rule ext, simp, rule impI)
fix x assume x: "x \<in> quotient_set H"
then obtain k where k: "k\<in>G" "x = LCoset_rel H `` {k}"
by (blast elim: LCoset_rel_quotientE)
with x have "LCoset_rel H `` {h} + x = LCoset_rel H `` {h+k}"
using assms normal_quotient_add_well_defined by auto
with assms k show "LCoset_rel H `` {h} + x = x"
using add_closed[of h k] lcoset_relI[of k "h+k" H]
normal_rconjby_closed[of H k h]
eq_equiv_class_iff[OF lcoset_subgroup_rel_equiv, of H]
by (auto simp add: add.assoc)
qed
definition quotient_group :: "'g set \<Rightarrow> 'g set permutation set" where
"quotient_group H \<equiv> BinOpSetGroup.pG (quotient_set H) (+)"
abbreviation "natural_quotient_hom H \<equiv> restrict0 (\<lambda>g. \<lceil>g|H\<rceil>) G"
theorem quotient_group:
"Subgroup H \<Longrightarrow> normal H \<Longrightarrow> Group (quotient_group H)"
unfolding quotient_group_def
using BinOpSetGroup.Group[OF BinOpSetGroup_normal_quotient]
by auto
lemma natural_quotient_hom:
"Subgroup H \<Longrightarrow> normal H \<Longrightarrow> GroupHom G (natural_quotient_hom H)"
using add_closed bij_lcoset_perm lcoset_perm_def supp_restrict0
normal_quotient_add_well_defined[THEN sym]
LCoset_rel_quotientI[of _ H]
by unfold_locales
(force simp add: lcoset_perm_comp plus_permutation_abs_eq)
lemma natural_quotient_hom_image:
"natural_quotient_hom H ` G = quotient_group H"
unfolding quotient_group_def
by (force elim: LCoset_rel_quotientE intro: LCoset_rel_quotientI)
lemma quotient_group_UN: "quotient_group H = (\<lambda>g. \<lceil>g|H\<rceil>) ` G"
using natural_quotient_hom_image by auto
lemma quotient_identity_rule: "\<lbrakk> Subgroup H; normal H; h\<in>H \<rbrakk> \<Longrightarrow> \<lceil>h|H\<rceil> = 0"
using lcoset_perm_def
by (simp add: trivial_lcoset_perm zero_permutation.abs_eq)
lemma quotient_group_lift_to_quotient_set:
"\<lbrakk> Subgroup H; normal H; g\<in>G \<rbrakk> \<Longrightarrow> (\<lceil>g|H\<rceil>) \<rightarrow> H = LCoset_rel H `` {g}"
using LCoset_rel_quotientI
BinOpSetGroup.G_perm_apply_identity[
OF BinOpSetGroup_normal_quotient
]
by simp
end (* context Group *)
subsubsection \<open>The induced homomorphism on a quotient group\<close>
text \<open>
A normal subgroup contained in the kernel of a homomorphism gives rise to a homomorphism on the
quotient group by that subgroup. When the subgroup is the kernel itself (which is always normal),
we obtain an isomorphism on the quotient.
\<close>
context GroupHom
begin
lemma respects_Ker_lcosets: "H \<subseteq> Ker \<Longrightarrow> T respects (LCoset_rel H)"
using uminus_add_in_Ker_eq_eq_im
unfolding lcoset_rel_def
by (blast intro: congruentI)
abbreviation "quotient_hom H \<equiv>
BinOpSetGroup.lift_hom (quotient_set H) (+) (quotientfun T)"
lemmas normal_subgroup_quotientfun_classrep_equality =
quotientfun_classrep_equality[
OF subgroup_refl_on_LCoset_rel, OF _ respects_Ker_lcosets
]
lemma quotient_hom_im:
"\<lbrakk> Subgroup H; normal H; H \<subseteq> Ker; g\<in>G \<rbrakk> \<Longrightarrow> quotient_hom H (\<lceil>g|H\<rceil>) = T g"
using quotient_group_def quotient_group_UN quotient_group_lift_to_quotient_set
BinOpSetGroup.inv_correspondence_conv_apply[
OF BinOpSetGroup_normal_quotient
]
normal_subgroup_quotientfun_classrep_equality
by auto
lemma quotient_hom:
assumes "Subgroup H" "normal H" "H \<subseteq> Ker"
shows "GroupHom (quotient_group H) (quotient_hom H)"
unfolding quotient_group_def
proof (
rule BinOpSetGroup.lift_hom, rule BinOpSetGroup_normal_quotient, rule assms(1),
rule assms(2)
)
from assms
show "\<forall>x \<in> quotient_set H. \<forall>y \<in> quotient_set H.
quotientfun T (x + y) = quotientfun T x + quotientfun T y"
using normal_quotient_add_well_defined normal_subgroup_quotientfun_classrep_equality
add_closed hom
by (fastforce elim: LCoset_rel_quotientE)
qed
end (* context GroupHom *)
subsection \<open>Free groups\<close>
subsubsection \<open>Words in letters of @{type signed} type\<close>
paragraph \<open>Definitions and basic fact\<close>
text \<open>
We pair elements of some type with type @{typ bool}, where the @{typ bool} part of the pair
indicates inversion.
\<close>
abbreviation "pairtrue \<equiv> \<lambda>s. (s,True)"
abbreviation "pairfalse \<equiv> \<lambda>s. (s,False)"
abbreviation flip_signed :: "'a signed \<Rightarrow> 'a signed"
where "flip_signed \<equiv> apsnd (\<lambda>b. \<not>b)"
abbreviation nflipped_signed :: "'a signed \<Rightarrow> 'a signed \<Rightarrow> bool"
where "nflipped_signed x y \<equiv> y \<noteq> flip_signed x"
lemma flip_signed_order2: "flip_signed (flip_signed x) = x"
using apsnd_conv[of "\<lambda>b. \<not>b" "fst x" "snd x"] by simp
abbreviation charpair :: "'a::uminus set \<Rightarrow> 'a \<Rightarrow> 'a signed"
where "charpair S s \<equiv> if s\<in>S then (s,True) else (-s,False)"
lemma map_charpair_uniform:
"ss\<in>lists S \<Longrightarrow> map (charpair S) ss = map pairtrue ss"
by (induct ss) auto
lemma fst_set_map_charpair_un_uminus:
fixes ss :: "'a::group_add list"
shows "ss\<in>lists (S \<union> uminus ` S) \<Longrightarrow> fst ` set (map (charpair S) ss) \<subseteq> S"
by (induct ss) auto
abbreviation apply_sign :: "('a\<Rightarrow>'b::uminus) \<Rightarrow> 'a signed \<Rightarrow> 'b"
where "apply_sign f x \<equiv> (if snd x then f (fst x) else - f (fst x))"
text \<open>
A word in such pairs will be considered proper if it does not contain consecutive letters that
have opposite signs (and so are considered inverse), since such consecutive letters would be
cancelled in a group.
\<close>
abbreviation proper_signed_list :: "'a signed list \<Rightarrow> bool"
where "proper_signed_list \<equiv> binrelchain nflipped_signed"
lemma proper_map_flip_signed:
"proper_signed_list xs \<Longrightarrow> proper_signed_list (map flip_signed xs)"
by (induct xs rule: list_induct_CCons) auto
lemma proper_rev_map_flip_signed:
"proper_signed_list xs \<Longrightarrow> proper_signed_list (rev (map flip_signed xs))"
using proper_map_flip_signed binrelchain_sym_rev[of nflipped_signed] by fastforce
lemma uniform_snd_imp_proper_signed_list:
"snd ` set xs \<subseteq> {b} \<Longrightarrow> proper_signed_list xs"
proof (induct xs rule: list_induct_CCons)
case CCons thus ?case by force
qed auto
lemma proper_signed_list_map_uniform_snd:
"proper_signed_list (map (\<lambda>s. (s,b)) as)"
using uniform_snd_imp_proper_signed_list[of _ b] by force
paragraph \<open>Algebra\<close>
text \<open>
Addition is performed by appending words and recursively removing any newly created adjacent
pairs of inverse letters. Since we will only ever be adding proper words, we only need to care
about newly created adjacent inverse pairs in the middle.
\<close>
function prappend_signed_list :: "'a signed list \<Rightarrow> 'a signed list \<Rightarrow> 'a signed list"
where "prappend_signed_list xs [] = xs"
| "prappend_signed_list [] ys = ys"
| "prappend_signed_list (xs@[x]) (y#ys) = (
if y = flip_signed x then prappend_signed_list xs ys else xs @ x # y # ys
)"
- by (auto, rule two_prod_lists_cases_snoc_Cons)
+ by (auto) (rule two_prod_lists_cases_snoc_Cons)
termination by (relation "measure (\<lambda>(xs,ys). length xs + length ys)") auto
lemma proper_prappend_signed_list:
"proper_signed_list xs \<Longrightarrow> proper_signed_list ys
\<Longrightarrow> proper_signed_list (prappend_signed_list xs ys)"
proof (induct xs ys rule: list_induct2_snoc_Cons)
case (snoc_Cons xs x y ys)
show ?case
proof (cases "y = flip_signed x")
case True with snoc_Cons show ?thesis
using binrelchain_append_reduce1[of nflipped_signed]
binrelchain_Cons_reduce[of nflipped_signed y]
by auto
next
case False with snoc_Cons(2,3) show ?thesis
using binrelchain_join[of nflipped_signed] by simp
qed
qed auto
lemma fully_prappend_signed_list:
"prappend_signed_list (rev (map flip_signed xs)) xs = []"
by (induct xs) auto
lemma prappend_signed_list_single_Cons:
"prappend_signed_list [x] (y#ys) = (if y = flip_signed x then ys else x#y#ys)"
using prappend_signed_list.simps(3)[of "[]" x] by simp
lemma prappend_signed_list_map_uniform_snd:
"prappend_signed_list (map (\<lambda>s. (s,b)) xs) (map (\<lambda>s. (s,b)) ys) =
map (\<lambda>s. (s,b)) xs @ map (\<lambda>s. (s,b)) ys"
by (cases xs ys rule: two_lists_cases_snoc_Cons) auto
lemma prappend_signed_list_assoc_conv_snoc2Cons:
assumes "proper_signed_list (xs@[y])" "proper_signed_list (y#ys)"
shows "prappend_signed_list (xs@[y]) ys = prappend_signed_list xs (y#ys)"
proof (cases xs ys rule: two_lists_cases_snoc_Cons')
case Nil1 with assms(2) show ?thesis
by (simp add: prappend_signed_list_single_Cons)
next
case Nil2 with assms(1) show ?thesis
using binrelchain_append_reduce2 by force
next
case (snoc_Cons as a b bs)
with assms show ?thesis
using prappend_signed_list.simps(3)[of "as@[a]"]
binrelchain_append_reduce2[of nflipped_signed as "[a,y]"]
by simp
qed simp
lemma prappend_signed_list_assoc:
"\<lbrakk> proper_signed_list xs; proper_signed_list ys; proper_signed_list zs \<rbrakk> \<Longrightarrow>
prappend_signed_list (prappend_signed_list xs ys) zs =
prappend_signed_list xs (prappend_signed_list ys zs)"
proof (induct xs ys zs rule: list_induct3_snoc_Conssnoc_Cons_pairwise)
case (snoc_single_Cons xs x y z zs)
thus ?case
using prappend_signed_list.simps(3)[of "[]" y]
prappend_signed_list.simps(3)[of "xs@[x]"]
by (cases "y = flip_signed x" "z = flip_signed y" rule: two_cases)
(auto simp add:
flip_signed_order2 prappend_signed_list_assoc_conv_snoc2Cons
)
next
case (snoc_Conssnoc_Cons xs x y ys w z zs)
thus ?case
using binrelchain_Cons_reduce[of nflipped_signed y "ys@[w]"]
binrelchain_Cons_reduce[of nflipped_signed z zs]
binrelchain_append_reduce1[of nflipped_signed xs]
binrelchain_append_reduce1[of nflipped_signed "y#ys"]
binrelchain_Conssnoc_reduce[of nflipped_signed y ys]
prappend_signed_list.simps(3)[of "y#ys"]
prappend_signed_list.simps(3)[of "xs@x#y#ys"]
by (cases "y = flip_signed x" "z = flip_signed w" rule: two_cases) auto
qed auto
lemma fst_set_prappend_signed_list:
"fst ` set (prappend_signed_list xs ys) \<subseteq> fst ` (set xs \<union> set ys)"
by (induct xs ys rule: list_induct2_snoc_Cons) auto
lemma collapse_flipped_signed:
"prappend_signed_list [(s,b)] [(s,\<not>b)] = []"
using prappend_signed_list.simps(3)[of "[]" "(s,b)"] by simp
subsubsection \<open>The collection of proper signed lists as a type\<close>
text \<open>
Here we create a type out of the collection of proper signed lists. This type will be of class
@{class group_add}, with the empty list as zero, the modified append operation
@{const prappend_signed_list} as addition, and inversion performed by flipping the signs of the
elements in the list and then reversing the order.
\<close>
paragraph \<open>Type definition, instantiations, and instances\<close>
text \<open>Here we define the type and instantiate it with respect to various type classes.\<close>
typedef 'a freeword = "{as::'a signed list. proper_signed_list as}"
morphisms freeword Abs_freeword
using binrelchain.simps(1) by fast
text \<open>
These two functions act as the natural injections of letters and words in the letter type into
the @{type freeword} type.
\<close>
abbreviation Abs_freeletter :: "'a \<Rightarrow> 'a freeword"
where "Abs_freeletter s \<equiv> Abs_freeword [pairtrue s]"
abbreviation Abs_freelist :: "'a list \<Rightarrow> 'a freeword"
where "Abs_freelist as \<equiv> Abs_freeword (map pairtrue as)"
abbreviation Abs_freelistfst :: "'a signed list \<Rightarrow> 'a freeword"
where "Abs_freelistfst xs \<equiv> Abs_freelist (map fst xs)"
setup_lifting type_definition_freeword
instantiation freeword :: (type) zero
begin
lift_definition zero_freeword :: "'a freeword" is "[]::'a signed list" by simp
instance ..
end
instantiation freeword :: (type) plus
begin
lift_definition plus_freeword :: "'a freeword \<Rightarrow> 'a freeword \<Rightarrow> 'a freeword"
is "prappend_signed_list"
using proper_prappend_signed_list
by fast
instance ..
end
instantiation freeword :: (type) uminus
begin
lift_definition uminus_freeword :: "'a freeword \<Rightarrow> 'a freeword"
is "\<lambda>xs. rev (map flip_signed xs)"
by (rule proper_rev_map_flip_signed)
instance ..
end
instantiation freeword :: (type) minus
begin
lift_definition minus_freeword :: "'a freeword \<Rightarrow> 'a freeword \<Rightarrow> 'a freeword"
is "\<lambda>xs ys. prappend_signed_list xs (rev (map flip_signed ys))"
using proper_rev_map_flip_signed proper_prappend_signed_list by fast
instance ..
end
instance freeword :: (type) semigroup_add
proof
fix a b c :: "'a freeword" show "a + b + c = a + (b + c)"
using prappend_signed_list_assoc[of "freeword a" "freeword b" "freeword c"]
by transfer simp
qed
instance freeword :: (type) monoid_add
proof
fix a b c :: "'a freeword"
show "0 + a = a" by transfer simp
show "a + 0 = a" by transfer simp
qed
instance freeword :: (type) group_add
proof
fix a b :: "'a freeword"
show "- a + a = 0"
using fully_prappend_signed_list[of "freeword a"] by transfer simp
show "a + - b = a - b" by transfer simp
qed
paragraph \<open>Basic algebra and transfer facts in the @{type freeword} type\<close>
text \<open>
Here we record basic algebraic manipulations for the @{type freeword} type as well as various
transfer facts for dealing with representations of elements of @{type freeword} type as lists of
signed letters.
\<close>
abbreviation Abs_freeletter_add :: "'a \<Rightarrow> 'a \<Rightarrow> 'a freeword" (infixl "[+]" 65)
where "s [+] t \<equiv> Abs_freeletter s + Abs_freeletter t"
lemma Abs_freeword_Cons:
assumes "proper_signed_list (x#xs)"
shows "Abs_freeword (x#xs) = Abs_freeword [x] + Abs_freeword xs"
proof (cases xs)
case Nil thus ?thesis
using add_0_right[of "Abs_freeword [x]"] by (simp add: zero_freeword.abs_eq)
next
case (Cons y ys)
with assms
have "freeword (Abs_freeword (x#xs)) =
freeword (Abs_freeword [x] + Abs_freeword xs)"
by (simp add:
plus_freeword.rep_eq Abs_freeword_inverse
prappend_signed_list_single_Cons
)
thus ?thesis using freeword_inject by fast
qed
lemma Abs_freelist_Cons: "Abs_freelist (x#xs) = Abs_freeletter x + Abs_freelist xs"
using proper_signed_list_map_uniform_snd[of True "x#xs"] Abs_freeword_Cons
by simp
lemma plus_freeword_abs_eq:
"proper_signed_list xs \<Longrightarrow> proper_signed_list ys \<Longrightarrow>
Abs_freeword xs + Abs_freeword ys = Abs_freeword (prappend_signed_list xs ys)"
using plus_freeword.abs_eq unfolding eq_onp_def by simp
lemma Abs_freeletter_add: "s [+] t = Abs_freelist [s,t]"
using Abs_freelist_Cons[of s "[t]"] by simp
lemma uminus_freeword_Abs_eq:
"proper_signed_list xs \<Longrightarrow>
- Abs_freeword xs = Abs_freeword (rev (map flip_signed xs))"
using uminus_freeword.abs_eq unfolding eq_onp_def by simp
lemma uminus_Abs_freeword_singleton:
"- Abs_freeword [(s,b)] = Abs_freeword [(s,\<not> b)]"
using uminus_freeword_Abs_eq[of "[(s,b)]"] by simp
lemma Abs_freeword_append_uniform_snd:
"Abs_freeword (map (\<lambda>s. (s,b)) (xs@ys)) =
Abs_freeword (map (\<lambda>s. (s,b)) xs) + Abs_freeword (map (\<lambda>s. (s,b)) ys)"
using proper_signed_list_map_uniform_snd[of b xs]
proper_signed_list_map_uniform_snd[of b ys]
plus_freeword_abs_eq prappend_signed_list_map_uniform_snd[of b xs ys]
by force
lemmas Abs_freelist_append = Abs_freeword_append_uniform_snd[of True]
lemma Abs_freelist_append_append:
"Abs_freelist (xs@ys@zs) = Abs_freelist xs + Abs_freelist ys + Abs_freelist zs"
using Abs_freelist_append[of "xs@ys"] Abs_freelist_append by simp
lemma Abs_freelist_inverse: "freeword (Abs_freelist as) = map pairtrue as"
using proper_signed_list_map_uniform_snd Abs_freeword_inverse by fast
lemma Abs_freeword_singleton_conv_apply_sign_freeletter:
"Abs_freeword [x] = apply_sign Abs_freeletter x"
by (cases x) (auto simp add: uminus_Abs_freeword_singleton)
lemma Abs_freeword_conv_freeletter_sum_list:
"proper_signed_list xs \<Longrightarrow>
Abs_freeword xs = (\<Sum>x\<leftarrow>xs. apply_sign Abs_freeletter x)"
proof (induct xs)
case (Cons x xs) thus ?case
using Abs_freeword_Cons[of x] binrelchain_Cons_reduce[of _ x]
by (simp add: Abs_freeword_singleton_conv_apply_sign_freeletter)
qed (simp add: zero_freeword.abs_eq)
lemma freeword_conv_freeletter_sum_list:
"x = (\<Sum>s\<leftarrow>freeword x. apply_sign Abs_freeletter s)"
using Abs_freeword_conv_freeletter_sum_list[of "freeword x"] freeword
by (auto simp add: freeword_inverse)
lemma Abs_freeletter_prod_conv_Abs_freeword:
"snd x \<Longrightarrow> Abs_freeletter (fst x) = Abs_freeword [x]"
using prod_eqI[of x "pairtrue (fst x)"] by simp
subsubsection \<open>Lifts of functions on the letter type\<close>
text \<open>
Here we lift functions on the letter type to type @{type freeword}. In particular, we are
interested in the case where the function being lifted has codomain of class @{class group_add}.
\<close>
paragraph \<open>The universal property\<close>
text \<open>
The universal property for free groups says that every function from the letter type to some
@{class group_add} type gives rise to a unique homomorphism.
\<close>
lemma extend_map_to_freeword_hom':
fixes f :: "'a \<Rightarrow> 'b::group_add"
defines h: "h::'a signed \<Rightarrow> 'b \<equiv> \<lambda>(s,b). if b then f s else - (f s)"
defines g: "g::'a signed list \<Rightarrow> 'b \<equiv> \<lambda>xs. sum_list (map h xs)"
shows "g (prappend_signed_list xs ys) = g xs + g ys"
proof (induct xs ys rule: list_induct2_snoc_Cons)
case (snoc_Cons xs x y ys)
show ?case
proof (cases "y = flip_signed x")
case True
with h have "h y = - h x"
using split_beta'[of "\<lambda>s b. if b then f s else - (f s)"] by simp
with g have "g (xs @ [x]) + g (y # ys) = g xs + g ys"
by (simp add: algebra_simps)
with True snoc_Cons show ?thesis by simp
next
case False with g show ?thesis
using sum_list.append[of "map h (xs@[x])" "map h (y#ys)"] by simp
qed
qed (auto simp add: h g)
lemma extend_map_to_freeword_hom1:
fixes f :: "'a \<Rightarrow> 'b::group_add"
defines "h::'a signed \<Rightarrow> 'b \<equiv> \<lambda>(s,b). if b then f s else - (f s)"
defines "g::'a freeword \<Rightarrow> 'b \<equiv> \<lambda>x. sum_list (map h (freeword x))"
shows "g (Abs_freeletter s) = f s"
using assms
by (simp add: Abs_freeword_inverse)
lemma extend_map_to_freeword_hom2:
fixes f :: "'a \<Rightarrow> 'b::group_add"
defines "h::'a signed \<Rightarrow> 'b \<equiv> \<lambda>(s,b). if b then f s else - (f s)"
defines "g::'a freeword \<Rightarrow> 'b \<equiv> \<lambda>x. sum_list (map h (freeword x))"
shows "UGroupHom g"
using assms
by (
auto intro: UGroupHomI
simp add: plus_freeword.rep_eq extend_map_to_freeword_hom'
)
lemma uniqueness_of_extended_map_to_freeword_hom':
fixes f :: "'a \<Rightarrow> 'b::group_add"
defines h: "h::'a signed \<Rightarrow> 'b \<equiv> \<lambda>(s,b). if b then f s else - (f s)"
defines g: "g::'a signed list \<Rightarrow> 'b \<equiv> \<lambda>xs. sum_list (map h xs)"
assumes singles: "\<And>s. k [(s,True)] = f s"
and adds : "\<And>xs ys. proper_signed_list xs \<Longrightarrow> proper_signed_list ys
\<Longrightarrow> k (prappend_signed_list xs ys) = k xs + k ys"
shows "proper_signed_list xs \<Longrightarrow> k xs = g xs"
proof-
have knil: "k [] = 0" using adds[of "[]" "[]"] add.assoc[of "k []" "k []" "- k []"] by simp
have ksingle: "\<And>x. k [x] = g [x]"
proof-
fix x :: "'a signed"
obtain s b where x: "x = (s,b)" by fastforce
show "k [x] = g [x]"
proof (cases b)
case False
from adds x singles
have "k (prappend_signed_list [x] [(s,True)]) = k [x] + f s"
by simp
moreover have "prappend_signed_list [(s,False)] [(s,True)] = []"
using collapse_flipped_signed[of s False] by simp
ultimately have "- f s = k [x] + f s + - f s" using x False knil by simp
with x False g h show "k [x] = g [x]" by (simp add: algebra_simps)
qed (simp add: x g h singles)
qed
show "proper_signed_list xs \<Longrightarrow> k xs = g xs"
proof (induct xs rule: list_induct_CCons)
case (CCons x y xs)
with g h show ?case
using adds[of "[x]" "y#xs"]
by (simp add:
prappend_signed_list_single_Cons
ksingle extend_map_to_freeword_hom'
)
qed (auto simp add: g h knil ksingle)
qed
lemma uniqueness_of_extended_map_to_freeword_hom:
fixes f :: "'a \<Rightarrow> 'b::group_add"
defines "h::'a signed \<Rightarrow> 'b \<equiv> \<lambda>(s,b). if b then f s else - (f s)"
defines "g::'a freeword \<Rightarrow> 'b \<equiv> \<lambda>x. sum_list (map h (freeword x))"
assumes k: "k \<circ> Abs_freeletter = f" "UGroupHom k"
shows "k = g"
proof
fix x::"'a freeword"
define k' where k': "k' \<equiv> k \<circ> Abs_freeword"
have "k' (freeword x) = g x" unfolding h_def g_def
proof (rule uniqueness_of_extended_map_to_freeword_hom')
from k' k(1) show "\<And>s. k' [pairtrue s] = f s" by auto
show "\<And>xs ys. proper_signed_list xs \<Longrightarrow> proper_signed_list ys
\<Longrightarrow> k' (prappend_signed_list xs ys) = k' xs + k' ys"
proof-
fix xs ys :: "'a signed list"
assume xsys: "proper_signed_list xs" "proper_signed_list ys"
with k'
show "k' (prappend_signed_list xs ys) = k' xs + k' ys"
using UGroupHom.hom[OF k(2), of "Abs_freeword xs" "Abs_freeword ys"]
by (simp add: plus_freeword_abs_eq)
qed
show "proper_signed_list (freeword x)" using freeword by fast
qed
with k' show "k x = g x" using freeword_inverse[of x] by simp
qed
theorem universal_property:
fixes f :: "'a \<Rightarrow> 'b::group_add"
shows "\<exists>!g::'a freeword\<Rightarrow>'b. g \<circ> Abs_freeletter = f \<and> UGroupHom g"
proof
define h where h: "h \<equiv> \<lambda>(s,b). if b then f s else - (f s)"
define g where g: "g \<equiv> \<lambda>x. sum_list (map h (freeword x))"
from g h show "g \<circ> Abs_freeletter = f \<and> UGroupHom g"
using extend_map_to_freeword_hom1[of f] extend_map_to_freeword_hom2
by auto
from g h show "\<And>k. k \<circ> Abs_freeletter = f \<and> UGroupHom k \<Longrightarrow> k = g"
using uniqueness_of_extended_map_to_freeword_hom by auto
qed
paragraph \<open>Properties of homomorphisms afforded by the universal property\<close>
text \<open>
The lift of a function on the letter set is the unique additive function on @{type freeword}
that agrees with the original function on letters.
\<close>
definition freeword_funlift :: "('a \<Rightarrow> 'b::group_add) \<Rightarrow> ('a freeword\<Rightarrow>'b::group_add)"
where "freeword_funlift f \<equiv> (THE g. g \<circ> Abs_freeletter = f \<and> UGroupHom g)"
lemma additive_freeword_funlift: "UGroupHom (freeword_funlift f)"
using theI'[OF universal_property, of f] unfolding freeword_funlift_def by simp
lemma freeword_funlift_Abs_freeletter: "freeword_funlift f (Abs_freeletter s) = f s"
using theI'[OF universal_property, of f]
comp_apply[of "freeword_funlift f" Abs_freeletter]
unfolding freeword_funlift_def
by fastforce
lemmas freeword_funlift_add = UGroupHom.hom [OF additive_freeword_funlift]
lemmas freeword_funlift_0 = UGroupHom.im_zero [OF additive_freeword_funlift]
lemmas freeword_funlift_uminus = UGroupHom.im_uminus [OF additive_freeword_funlift]
lemmas freeword_funlift_diff = UGroupHom.im_diff [OF additive_freeword_funlift]
lemmas freeword_funlift_lconjby = UGroupHom.im_lconjby [OF additive_freeword_funlift]
lemma freeword_funlift_uminus_Abs_freeletter:
"freeword_funlift f (Abs_freeword [(s,False)]) = - f s"
using freeword_funlift_uminus[of f "Abs_freeword [(s,False)]"]
uminus_freeword_Abs_eq[of "[(s,False)]"]
freeword_funlift_Abs_freeletter[of f]
by simp
lemma freeword_funlift_Abs_freeword_singleton:
"freeword_funlift f (Abs_freeword [x]) = apply_sign f x"
proof-
obtain s b where x: "x = (s,b)" by fastforce
thus ?thesis
using freeword_funlift_Abs_freeletter freeword_funlift_uminus_Abs_freeletter
by (cases b) auto
qed
lemma freeword_funlift_Abs_freeword_Cons:
assumes "proper_signed_list (x#xs)"
shows "freeword_funlift f (Abs_freeword (x#xs)) =
apply_sign f x + freeword_funlift f (Abs_freeword xs)"
proof-
from assms
have "freeword_funlift f (Abs_freeword (x#xs)) =
freeword_funlift f (Abs_freeword [x]) +
freeword_funlift f (Abs_freeword xs)"
using Abs_freeword_Cons[of x xs] freeword_funlift_add by simp
thus ?thesis
using freeword_funlift_Abs_freeword_singleton[of f x] by simp
qed
lemma freeword_funlift_Abs_freeword:
"proper_signed_list xs \<Longrightarrow> freeword_funlift f (Abs_freeword xs) =
(\<Sum>x\<leftarrow>xs. apply_sign f x)"
proof (induct xs)
case (Cons x xs) thus ?case
using freeword_funlift_Abs_freeword_Cons[of _ _ f]
binrelchain_Cons_reduce[of _ x xs]
by simp
qed (simp add: zero_freeword.abs_eq[THEN sym] freeword_funlift_0)
lemma freeword_funlift_Abs_freelist:
"freeword_funlift f (Abs_freelist xs) = (\<Sum>x\<leftarrow>xs. f x)"
proof (induct xs)
case (Cons x xs) thus ?case
using Abs_freelist_Cons[of x xs]
by (simp add: freeword_funlift_add freeword_funlift_Abs_freeletter)
qed (simp add: zero_freeword.abs_eq[THEN sym] freeword_funlift_0)
lemma freeword_funlift_im':
"proper_signed_list xs \<Longrightarrow> fst ` set xs \<subseteq> S \<Longrightarrow>
freeword_funlift f (Abs_freeword xs) \<in> \<langle>f`S\<rangle>"
proof (induct xs)
case Nil
have "Abs_freeword ([]::'a signed list) = (0::'a freeword)"
using zero_freeword.abs_eq[THEN sym] by simp
thus "freeword_funlift f (Abs_freeword ([]::'a signed list)) \<in> \<langle>f`S\<rangle>"
using freeword_funlift_0[of f] genby_0_closed by simp
next
case (Cons x xs)
define y where y: "y \<equiv> apply_sign f x"
define z where z: "z \<equiv> freeword_funlift f (Abs_freeword xs)"
from Cons(3) have "fst ` set xs \<subseteq> S" by simp
with z Cons(1,2) have "z \<in> \<langle>f`S\<rangle>" using binrelchain_Cons_reduce by fast
with y Cons(3) have "y + z \<in> \<langle>f`S\<rangle>"
using genby_genset_closed[of _ "f`S"]
genby_uminus_closed genby_add_closed[of y]
by fastforce
with Cons(2) y z show ?case
using freeword_funlift_Abs_freeword_Cons
subst[
OF sym,
of "freeword_funlift f (Abs_freeword (x#xs))" "y+z"
"\<lambda>b. b\<in>\<langle>f`S\<rangle>"
]
by fast
qed
subsubsection \<open>Free groups on a set\<close>
text \<open>
We now take the free group on a set to be the set in the @{type freeword} type with letters
restricted to the given set.
\<close>
paragraph \<open>Definition and basic facts\<close>
text \<open>
Here we define the set of elements of the free group over a set of letters, and record basic
facts about that set.
\<close>
definition FreeGroup :: "'a set \<Rightarrow> 'a freeword set"
where "FreeGroup S \<equiv> {x. fst ` set (freeword x) \<subseteq> S}"
lemma FreeGroupI_transfer:
"proper_signed_list xs \<Longrightarrow> fst ` set xs \<subseteq> S \<Longrightarrow> Abs_freeword xs \<in> FreeGroup S"
using Abs_freeword_inverse unfolding FreeGroup_def by fastforce
lemma FreeGroupD: "x \<in> FreeGroup S \<Longrightarrow> fst ` set (freeword x) \<subseteq> S"
using FreeGroup_def by fast
lemma FreeGroupD_transfer:
"proper_signed_list xs \<Longrightarrow> Abs_freeword xs \<in> FreeGroup S \<Longrightarrow> fst ` set xs \<subseteq> S"
using Abs_freeword_inverse unfolding FreeGroup_def by fastforce
lemma FreeGroupD_transfer':
"Abs_freelist xs \<in> FreeGroup S \<Longrightarrow> xs \<in> lists S"
using proper_signed_list_map_uniform_snd FreeGroupD_transfer by fastforce
lemma FreeGroup_0_closed: "0 \<in> FreeGroup S"
proof-
have "(0::'a freeword) = Abs_freeword []" using zero_freeword.abs_eq by fast
moreover have "Abs_freeword [] \<in> FreeGroup S"
using FreeGroupI_transfer[of "[]"] by simp
ultimately show ?thesis by simp
qed
lemma FreeGroup_diff_closed:
assumes "x \<in> FreeGroup S" "y \<in> FreeGroup S"
shows "x-y \<in> FreeGroup S"
proof-
define xs where xs: "xs \<equiv> freeword x"
define ys where ys: "ys \<equiv> freeword y"
have "freeword (x-y) =
prappend_signed_list (freeword x) (rev (map flip_signed (freeword y)))"
by transfer simp
hence "fst ` set (freeword (x-y)) \<subseteq> fst ` (set (freeword x) \<union> set (freeword y))"
using fst_set_prappend_signed_list by force
with assms show ?thesis unfolding FreeGroup_def by fast
qed
lemma FreeGroup_Group: "Group (FreeGroup S)"
using FreeGroup_0_closed FreeGroup_diff_closed by unfold_locales fast
lemmas FreeGroup_add_closed = Group.add_closed [OF FreeGroup_Group]
lemmas FreeGroup_uminus_closed = Group.uminus_closed [OF FreeGroup_Group]
lemmas FreeGroup_genby_set_lconjby_set_rconjby_closed =
Group.genby_set_lconjby_set_rconjby_closed[OF FreeGroup_Group]
lemma Abs_freelist_in_FreeGroup: "ss \<in> lists S \<Longrightarrow> Abs_freelist ss \<in> FreeGroup S"
using proper_signed_list_map_uniform_snd by (fastforce intro: FreeGroupI_transfer)
lemma Abs_freeletter_in_FreeGroup_iff: "(Abs_freeletter s \<in> FreeGroup S) = (s\<in>S)"
using Abs_freeword_inverse[of "[pairtrue s]"] unfolding FreeGroup_def by simp
paragraph \<open>Lifts of functions from the letter set to some type of class @{class group_add}\<close>
text \<open>
We again obtain a universal property for functions from the (restricted) letter set to some type
of class @{class group_add}.
\<close>
abbreviation "res_freeword_funlift f S \<equiv>
restrict0 (freeword_funlift f) (FreeGroup S)"
lemma freeword_funlift_im: "x \<in> FreeGroup S \<Longrightarrow> freeword_funlift f x \<in> \<langle>f ` S\<rangle>"
using freeword[of x] freeword_funlift_im'[of "freeword x"]
freeword_inverse[of x]
unfolding FreeGroup_def
by auto
lemma freeword_funlift_surj':
"ys \<in> lists (f`S \<union> uminus`f`S) \<Longrightarrow> sum_list ys \<in> freeword_funlift f ` FreeGroup S"
proof (induct ys)
case Nil thus ?case using FreeGroup_0_closed freeword_funlift_0 by fastforce
next
case (Cons y ys)
from this obtain x
where x: "x \<in> FreeGroup S" "sum_list ys = freeword_funlift f x"
by auto
show "sum_list (y#ys) \<in> freeword_funlift f ` FreeGroup S"
proof (cases "y \<in> f`S")
case True
from this obtain s where s: "s\<in>S" "y = f s" by fast
from s(1) x(1) have "Abs_freeletter s + x \<in> FreeGroup S"
using FreeGroupI_transfer[of _ S] FreeGroup_add_closed[of _ S] by force
moreover from s(2) x(2)
have "freeword_funlift f (Abs_freeletter s + x) = sum_list (y#ys)"
using freeword_funlift_add[of f] freeword_funlift_Abs_freeletter
by simp
ultimately show ?thesis by force
next
case False
with Cons(2) obtain s where s: "s\<in>S" "y = - f s" by auto
from s(1) x(1) have "Abs_freeword [(s,False)] + x \<in> FreeGroup S"
using FreeGroupI_transfer[of _ S] FreeGroup_add_closed[of _ S] by force
moreover from s(2) x(2)
have "freeword_funlift f (Abs_freeword [(s,False)] + x) = sum_list (y#ys)"
using freeword_funlift_add[of f] freeword_funlift_uminus_Abs_freeletter
by simp
ultimately show ?thesis by force
qed
qed
lemma freeword_funlift_surj:
fixes f :: "'a \<Rightarrow> 'b::group_add"
shows "freeword_funlift f ` FreeGroup S = \<langle>f`S\<rangle>"
proof (rule seteqI)
show "\<And>a. a \<in> freeword_funlift f ` FreeGroup S \<Longrightarrow> a \<in> \<langle>f`S\<rangle>"
using freeword_funlift_im by auto
next
fix w assume "w\<in>\<langle>f`S\<rangle>"
from this obtain ys where ys: "ys \<in> lists (f`S \<union> uminus`f`S)" "w = sum_list ys"
using genby_eq_sum_lists[of "f`S"] by auto
thus "w \<in> freeword_funlift f ` FreeGroup S" using freeword_funlift_surj' by simp
qed
lemma hom_restrict0_freeword_funlift:
"GroupHom (FreeGroup S) (res_freeword_funlift f S)"
using UGroupHom.restrict0 additive_freeword_funlift FreeGroup_Group
by auto
lemma uniqueness_of_restricted_lift:
assumes "GroupHom (FreeGroup S) T" "\<forall>s\<in>S. T (Abs_freeletter s) = f s"
shows "T = res_freeword_funlift f S"
proof
fix x
define F where "F \<equiv> res_freeword_funlift f S"
define u_Abs where "u_Abs \<equiv> \<lambda>a::'a signed. apply_sign Abs_freeletter a"
show "T x = F x"
proof (cases "x \<in> FreeGroup S")
case True
have 1: "set (map u_Abs (freeword x)) \<subseteq> FreeGroup S"
using u_Abs_def FreeGroupD[OF True]
Abs_freeletter_in_FreeGroup_iff[of _ S]
FreeGroup_uminus_closed
by auto
moreover from u_Abs_def have "x = (\<Sum>a\<leftarrow>freeword x. u_Abs a)"
using freeword_conv_freeletter_sum_list by fast
ultimately
have "T x = (\<Sum>a\<leftarrow>freeword x. T (u_Abs a))"
"F x = (\<Sum>a\<leftarrow>freeword x. F (u_Abs a))"
using F_def
GroupHom.im_sum_list_map[OF assms(1), of u_Abs "freeword x"]
GroupHom.im_sum_list_map[
OF hom_restrict0_freeword_funlift,
of u_Abs "freeword x" S f
]
by auto
moreover have "\<forall>a\<in>set (freeword x). T (u_Abs a) = F (u_Abs a)"
proof
fix a assume "a \<in> set (freeword x)"
moreover define b where "b \<equiv> Abs_freeletter (fst a)"
ultimately show "T (u_Abs a) = F (u_Abs a)"
using F_def u_Abs_def True assms(2) FreeGroupD[of x S]
GroupHom.im_uminus[OF assms(1)]
Abs_freeletter_in_FreeGroup_iff[of "fst a" S]
GroupHom.im_uminus[OF hom_restrict0_freeword_funlift, of b S f]
freeword_funlift_Abs_freeletter[of f]
by auto
qed
ultimately show ?thesis
using F_def
sum_list_map_cong[of "freeword x" "\<lambda>s. T (u_Abs s)" "\<lambda>s. F (u_Abs s)"]
by simp
next
case False
with assms(1) F_def show ?thesis
using hom_restrict0_freeword_funlift GroupHom.supp suppI_contra[of x T]
suppI_contra[of x F]
by fastforce
qed
qed
theorem FreeGroup_universal_property:
fixes f :: "'a \<Rightarrow> 'b::group_add"
shows "\<exists>!T::'a freeword\<Rightarrow>'b. (\<forall>s\<in>S. T (Abs_freeletter s) = f s) \<and>
GroupHom (FreeGroup S) T"
proof (rule ex1I, rule conjI)
show "\<forall>s\<in>S. res_freeword_funlift f S (Abs_freeletter s) = f s"
using Abs_freeletter_in_FreeGroup_iff[of _ S] freeword_funlift_Abs_freeletter
by auto
show "\<And>T. (\<forall>s\<in>S. T (Abs_freeletter s) = f s) \<and>
GroupHom (FreeGroup S) T \<Longrightarrow>
T = restrict0 (freeword_funlift f) (FreeGroup S)"
using uniqueness_of_restricted_lift by auto
qed (rule hom_restrict0_freeword_funlift)
subsubsection \<open>Group presentations\<close>
text \<open>
We now define a group presentation to be the quotient of a free group by the subgroup generated by
all conjugates of a set of relators. We are most concerned with lifting functions on the letter
set to the free group and with the associated induced homomorphisms on the quotient.
\<close>
paragraph \<open>A first group presentation locale and basic facts\<close>
text \<open>
Here we define a locale that provides a way to construct a group by providing sets of generators
and relator words.
\<close>
locale GroupByPresentation =
fixes S :: "'a set" \<comment> \<open>the set of generators\<close>
and P :: "'a signed list set" \<comment> \<open>the set of relator words\<close>
assumes P_S: "ps\<in>P \<Longrightarrow> fst ` set ps \<subseteq> S"
and proper_P: "ps\<in>P \<Longrightarrow> proper_signed_list ps"
begin
abbreviation "P' \<equiv> Abs_freeword ` P" \<comment> \<open>the set of relators\<close>
abbreviation "Q \<equiv> Group.normal_closure (FreeGroup S) P'"
\<comment> \<open>the normal subgroup generated by relators inside the free group\<close>
abbreviation "G \<equiv> Group.quotient_group (FreeGroup S) Q"
lemmas G_UN = Group.quotient_group_UN[OF FreeGroup_Group, of S Q]
lemma P'_FreeS: "P' \<subseteq> FreeGroup S"
using P_S proper_P by (blast intro: FreeGroupI_transfer)
lemma relators: "P' \<subseteq> Q"
using FreeGroup_0_closed genby_genset_subset by fastforce
lemmas lconjby_P'_FreeS =
Group.set_lconjby_subset_closed[
OF FreeGroup_Group _ P'_FreeS, OF basic_monos(1)
]
lemmas Q_FreeS =
Group.genby_closed[OF FreeGroup_Group lconjby_P'_FreeS]
lemmas Q_subgroup_FreeS =
Group.genby_subgroup[OF FreeGroup_Group lconjby_P'_FreeS]
lemmas normal_Q = Group.normal_closure[OF FreeGroup_Group, OF P'_FreeS]
lemmas natural_hom =
Group.natural_quotient_hom[
OF FreeGroup_Group Q_subgroup_FreeS normal_Q
]
lemmas natural_hom_image =
Group.natural_quotient_hom_image[OF FreeGroup_Group, of S Q]
end (* context GroupByPresentation *)
paragraph \<open>Functions on the quotient induced from lifted functions\<close>
text \<open>
A function on the generator set into a type of class @{class group_add} lifts to a unique
homomorphism on the free group. If this lift is trivial on relators, then it factors to a
homomorphism of the group described by the generators and relators.
\<close>
locale GroupByPresentationInducedFun = GroupByPresentation S P
for S :: "'a set"
and P :: "'a signed list set" \<comment> \<open>the set of relator words\<close>
+ fixes f :: "'a \<Rightarrow> 'b::group_add"
assumes lift_f_trivial_P:
"ps\<in>P \<Longrightarrow> freeword_funlift f (Abs_freeword ps) = 0"
begin
abbreviation "lift_f \<equiv> freeword_funlift f"
definition induced_hom :: "'a freeword set permutation \<Rightarrow> 'b"
where "induced_hom \<equiv> GroupHom.quotient_hom (FreeGroup S)
(restrict0 lift_f (FreeGroup S)) Q"
\<comment> \<open>the @{const restrict0} operation is really only necessary to make
@{const GroupByPresentationInducedFun.induced_hom} a @{const GroupHom}\<close>
abbreviation "F \<equiv> induced_hom"
lemma lift_f_trivial_P': "p\<in>P' \<Longrightarrow> lift_f p = 0"
using lift_f_trivial_P by fast
lemma lift_f_trivial_lconjby_P': "p\<in>P' \<Longrightarrow> lift_f (lconjby w p) = 0"
using freeword_funlift_lconjby[of f] lift_f_trivial_P' by simp
lemma lift_f_trivial_Q: "q\<in>Q \<Longrightarrow> lift_f q = 0"
proof (erule genby.induct, rule freeword_funlift_0)
show "\<And>s. s \<in> (\<Union>w \<in> FreeGroup S. lconjby w ` P') \<Longrightarrow> lift_f s = 0"
using lift_f_trivial_lconjby_P' by fast
next
fix w w' :: "'a freeword" assume ww': "lift_f w = 0" "lift_f w' = 0"
have "lift_f (w - w') = lift_f w - lift_f w'"
using freeword_funlift_diff[of f w] by simp
with ww' show "lift_f (w-w') = 0" by simp
qed
lemma lift_f_ker_Q: "Q \<subseteq> ker lift_f"
using lift_f_trivial_Q unfolding ker_def by auto
lemma lift_f_Ker_Q: "Q \<subseteq> GroupHom.Ker (FreeGroup S) lift_f"
using lift_f_ker_Q Q_FreeS by fast
lemma restrict0_lift_f_Ker_Q:
"Q \<subseteq> GroupHom.Ker (FreeGroup S) (restrict0 lift_f (FreeGroup S))"
using lift_f_Ker_Q ker_subset_ker_restrict0 by fast
lemma induced_hom_equality:
"w \<in> FreeGroup S \<Longrightarrow> F (\<lceil>FreeGroup S|w|Q\<rceil>) = lift_f w"
\<comment> \<open>algebraic properties of the induced homomorphism could be proved using its properties as a group
homomorphism, but it's generally easier to prove them using the algebraic properties of the lift
via this lemma\<close>
unfolding induced_hom_def
using GroupHom.quotient_hom_im hom_restrict0_freeword_funlift
Q_subgroup_FreeS normal_Q restrict0_lift_f_Ker_Q
by fastforce
lemma hom_induced_hom: "GroupHom G F"
unfolding induced_hom_def
using GroupHom.quotient_hom hom_restrict0_freeword_funlift
Q_subgroup_FreeS normal_Q restrict0_lift_f_Ker_Q
by fast
lemma induced_hom_Abs_freeletter_equality:
"s\<in>S \<Longrightarrow> F (\<lceil>FreeGroup S|Abs_freeletter s|Q\<rceil>) = f s"
using Abs_freeletter_in_FreeGroup_iff[of s S]
by (simp add: induced_hom_equality freeword_funlift_Abs_freeletter)
lemma uniqueness_of_induced_hom':
defines "q \<equiv> Group.natural_quotient_hom (FreeGroup S) Q"
assumes "GroupHom G T" "\<forall>s\<in>S. T (\<lceil>FreeGroup S|Abs_freeletter s|Q\<rceil>) = f s"
shows "T \<circ> q = F \<circ> q"
proof-
from assms have "T\<circ>q = res_freeword_funlift f S"
using natural_hom natural_hom_image Abs_freeletter_in_FreeGroup_iff[of _ S]
by (force intro: uniqueness_of_restricted_lift GroupHom.comp)
moreover from q_def have "F \<circ> q = res_freeword_funlift f S"
using induced_hom_equality GroupHom.im_zero[OF hom_induced_hom]
by auto
ultimately show ?thesis by simp
qed
lemma uniqueness_of_induced_hom:
assumes "GroupHom G T" "\<forall>s\<in>S. T (\<lceil>FreeGroup S|Abs_freeletter s|Q\<rceil>) = f s"
shows "T = F"
proof
fix x
show "T x = F x"
proof (cases "x\<in>G")
case True
define q where "q \<equiv> Group.natural_quotient_hom (FreeGroup S) Q"
from True obtain w where "w \<in> FreeGroup S" "x = (\<lceil>FreeGroup S|w|Q\<rceil>)"
using G_UN by fast
with q_def have "T x = (T\<circ>q) w" "F x = (F\<circ>q) w" by auto
with assms q_def show ?thesis using uniqueness_of_induced_hom' by simp
next
case False
with assms(1) show ?thesis
using hom_induced_hom GroupHom.supp suppI_contra[of x T]
suppI_contra[of x F]
by fastforce
qed
qed
theorem induced_hom_universal_property:
"\<exists>!F. GroupHom G F \<and> (\<forall>s\<in>S. F (\<lceil>FreeGroup S|Abs_freeletter s|Q\<rceil>) = f s)"
using hom_induced_hom induced_hom_Abs_freeletter_equality
uniqueness_of_induced_hom
by blast
lemma induced_hom_Abs_freelist_conv_sum_list:
"ss\<in>lists S \<Longrightarrow> F (\<lceil>FreeGroup S|Abs_freelist ss|Q\<rceil>) = (\<Sum>s\<leftarrow>ss. f s)"
by (simp add:
Abs_freelist_in_FreeGroup induced_hom_equality freeword_funlift_Abs_freelist
)
lemma induced_hom_surj: "F`G = \<langle>f`S\<rangle>"
proof (rule seteqI)
show "\<And>x. x\<in>F`G \<Longrightarrow> x\<in>\<langle>f`S\<rangle>"
using G_UN induced_hom_equality freeword_funlift_surj[of f S] by auto
next
fix x assume "x\<in>\<langle>f`S\<rangle>"
hence "x \<in> lift_f ` FreeGroup S" using freeword_funlift_surj[of f S] by fast
thus "x \<in> F`G" using induced_hom_equality G_UN by force
qed
end (* context GroupByPresentationInducedFun *)
paragraph \<open>Groups affording a presentation\<close>
text \<open>
The locale @{const GroupByPresentation} allows the construction of a @{const Group} out of any
type from a set of generating letters and a set of relator words in (signed) letters. The
following locale concerns the question of when the @{const Group} generated by a set in class
@{class group_add} is isomorphic to a group presentation.
\<close>
locale GroupWithGeneratorsRelators =
fixes S :: "'g::group_add set" \<comment> \<open>the set of generators\<close>
and R :: "'g list set" \<comment> \<open>the set of relator words\<close>
assumes relators: "rs\<in>R \<Longrightarrow> rs \<in> lists (S \<union> uminus ` S)"
"rs\<in>R \<Longrightarrow> sum_list rs = 0"
"rs\<in>R \<Longrightarrow> proper_signed_list (map (charpair S) rs)"
begin
abbreviation "P \<equiv> map (charpair S) ` R"
abbreviation "P' \<equiv> GroupByPresentation.P' P"
abbreviation "Q \<equiv> GroupByPresentation.Q S P"
abbreviation "G \<equiv> GroupByPresentation.G S P"
abbreviation "relator_freeword rs \<equiv> Abs_freeword (map (charpair S) rs)"
\<comment> \<open>this maps R onto P'\<close>
abbreviation "freeliftid \<equiv> freeword_funlift id"
abbreviation induced_id :: "'g freeword set permutation \<Rightarrow> 'g"
where "induced_id \<equiv> GroupByPresentationInducedFun.induced_hom S P id"
lemma GroupByPresentation_S_P: "GroupByPresentation S P"
proof
show "\<And>ps. ps \<in> P \<Longrightarrow> fst ` set ps \<subseteq> S"
using fst_set_map_charpair_un_uminus relators(1) by fast
show "\<And>ps. ps \<in> P \<Longrightarrow> proper_signed_list ps" using relators(3) by fast
qed
lemmas G_UN = GroupByPresentation.G_UN[OF GroupByPresentation_S_P]
lemmas P'_FreeS = GroupByPresentation.P'_FreeS[OF GroupByPresentation_S_P]
lemma freeliftid_trivial_relator_freeword_R:
"rs\<in>R \<Longrightarrow> freeliftid (relator_freeword rs) = 0"
using relators(2,3) freeword_funlift_Abs_freeword[of "map (charpair S) rs" id]
sum_list_map_cong[of rs "(apply_sign id) \<circ> (charpair S)" id]
by simp
lemma freeliftid_trivial_P: "ps\<in>P \<Longrightarrow> freeliftid (Abs_freeword ps) = 0"
using freeliftid_trivial_relator_freeword_R by fast
lemma GroupByPresentationInducedFun_S_P_id:
"GroupByPresentationInducedFun S P id"
by (
intro_locales, rule GroupByPresentation_S_P,
unfold_locales, rule freeliftid_trivial_P
)
lemma induced_id_Abs_freelist_conv_sum_list:
"ss\<in>lists S \<Longrightarrow> induced_id (\<lceil>FreeGroup S|Abs_freelist ss|Q\<rceil>) = sum_list ss"
by (simp add:
GroupByPresentationInducedFun.induced_hom_Abs_freelist_conv_sum_list[
OF GroupByPresentationInducedFun_S_P_id
]
)
lemma lconj_relator_freeword_R:
"\<lbrakk> rs\<in>R; proper_signed_list xs; fst ` set xs \<subseteq> S \<rbrakk> \<Longrightarrow>
lconjby (Abs_freeword xs) (relator_freeword rs) \<in> Q"
by (blast intro: genby_genset_closed FreeGroupI_transfer)
lemma rconj_relator_freeword:
assumes "rs\<in>R" "proper_signed_list xs" "fst ` set xs \<subseteq> S"
shows "rconjby (Abs_freeword xs) (relator_freeword rs) \<in> Q"
proof (rule genby_genset_closed, rule UN_I)
show "- Abs_freeword xs \<in> FreeGroup S"
using FreeGroupI_transfer[OF assms(2,3)] FreeGroup_uminus_closed by fast
from assms(1)
show "rconjby (Abs_freeword xs) (relator_freeword rs) \<in>
lconjby (- Abs_freeword xs) ` Abs_freeword ` P"
by simp
qed
lemma lconjby_Abs_freelist_relator_freeword:
"\<lbrakk> rs\<in>R; xs\<in>lists S \<rbrakk> \<Longrightarrow> lconjby (Abs_freelist xs) (relator_freeword rs) \<in> Q"
using proper_signed_list_map_uniform_snd by (force intro: lconj_relator_freeword_R)
text \<open>
Here we record that the lift of the identity map to the free group on @{term S} induces a
homomorphic surjection onto the group generated by @{term S} from the group presentation on
@{term S}, subject to the same relations as the elements of @{term S}.
\<close>
theorem induced_id_hom_surj: "GroupHom G induced_id" "induced_id ` G = \<langle>S\<rangle>"
using GroupByPresentationInducedFun.hom_induced_hom[
OF GroupByPresentationInducedFun_S_P_id
]
GroupByPresentationInducedFun.induced_hom_surj[
OF GroupByPresentationInducedFun_S_P_id
]
by auto
end (* context GroupWithGeneratorsRelators *)
locale GroupPresentation = GroupWithGeneratorsRelators S R
for S :: "'g::group_add set" \<comment> \<open>the set of generators\<close>
and R :: "'g list set" \<comment> \<open>the set of relator words\<close>
+ assumes induced_id_inj: "inj_on induced_id G"
begin
abbreviation "inv_induced_id \<equiv> the_inv_into G induced_id"
lemma inv_induced_id_sum_list_S:
"ss \<in> lists S \<Longrightarrow> inv_induced_id (sum_list ss) = (\<lceil>FreeGroup S|Abs_freelist ss|Q\<rceil>)"
using G_UN induced_id_inj induced_id_Abs_freelist_conv_sum_list
Abs_freelist_in_FreeGroup
by (blast intro: the_inv_into_f_eq)
end (* GroupPresentation *)
subsection \<open>Words over a generating set\<close>
text \<open>
Here we gather the necessary constructions and facts for studying a group generated by some set
in terms of words in the generators.
\<close>
context monoid_add
begin
abbreviation "word_for A a as \<equiv> as \<in> lists A \<and> sum_list as = a"
definition reduced_word_for :: "'a set \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> bool"
where "reduced_word_for A a as \<equiv> is_arg_min length (word_for A a) as"
abbreviation "reduced_word A as \<equiv> reduced_word_for A (sum_list as) as"
abbreviation "reduced_words_for A a \<equiv> Collect (reduced_word_for A a)"
abbreviation reduced_letter_set :: "'a set \<Rightarrow> 'a \<Rightarrow> 'a set"
where "reduced_letter_set A a \<equiv> \<Union>( set ` (reduced_words_for A a) )"
\<comment> \<open>will be empty if @{term a} is not in the set generated by @{term A}\<close>
definition word_length :: "'a set \<Rightarrow> 'a \<Rightarrow> nat"
where "word_length A a \<equiv> length (arg_min length (word_for A a))"
lemma reduced_word_forI:
assumes "as \<in> lists A" "sum_list as = a"
"\<And>bs. bs \<in> lists A \<Longrightarrow> sum_list bs = a \<Longrightarrow> length as \<le> length bs"
shows "reduced_word_for A a as"
using assms
unfolding reduced_word_for_def
by (force intro: is_arg_minI)
lemma reduced_word_forI_compare:
"\<lbrakk> reduced_word_for A a as; bs \<in> lists A; sum_list bs = a; length bs = length as \<rbrakk>
\<Longrightarrow> reduced_word_for A a bs"
using reduced_word_for_def is_arg_min_eq[of length] by fast
lemma reduced_word_for_lists: "reduced_word_for A a as \<Longrightarrow> as \<in> lists A"
using reduced_word_for_def is_arg_minD1 by fast
lemma reduced_word_for_sum_list: "reduced_word_for A a as \<Longrightarrow> sum_list as = a"
using reduced_word_for_def is_arg_minD1 by fast
lemma reduced_word_for_minimal:
"\<lbrakk> reduced_word_for A a as; bs \<in> lists A; sum_list bs = a \<rbrakk> \<Longrightarrow>
length as \<le> length bs"
using reduced_word_for_def is_arg_minD2[of length]
by fastforce
lemma reduced_word_for_length:
"reduced_word_for A a as \<Longrightarrow> length as = word_length A a"
unfolding word_length_def reduced_word_for_def is_arg_min_def
by (fastforce intro: arg_min_equality[THEN sym])
lemma reduced_word_for_eq_length:
"reduced_word_for A a as \<Longrightarrow> reduced_word_for A a bs \<Longrightarrow> length as = length bs"
using reduced_word_for_length by simp
lemma reduced_word_for_arg_min:
"as \<in> lists A \<Longrightarrow> sum_list as = a \<Longrightarrow>
reduced_word_for A a (arg_min length (word_for A a))"
using is_arg_min_arg_min_nat[of "word_for A a"]
unfolding reduced_word_for_def
by fast
lemma nil_reduced_word_for_0: "reduced_word_for A 0 []"
by (auto intro: reduced_word_forI)
lemma reduced_word_for_0_imp_nil: "reduced_word_for A 0 as \<Longrightarrow> as = []"
using nil_reduced_word_for_0[of A] reduced_word_for_minimal[of A 0 as]
unfolding reduced_word_for_def is_arg_min_def
by (metis (mono_tags, hide_lams) length_0_conv length_greater_0_conv)
lemma not_reduced_word_for:
"\<lbrakk> bs \<in> lists A; sum_list bs = a; length bs < length as \<rbrakk> \<Longrightarrow>
\<not> reduced_word_for A a as"
using reduced_word_for_minimal by fastforce
lemma reduced_word_for_imp_reduced_word:
"reduced_word_for A a as \<Longrightarrow> reduced_word A as"
unfolding reduced_word_for_def is_arg_min_def
by (fast intro: reduced_word_forI)
lemma sum_list_zero_nreduced:
"as \<noteq> [] \<Longrightarrow> sum_list as = 0 \<Longrightarrow> \<not> reduced_word A as"
using not_reduced_word_for[of "[]"] by simp
lemma order2_nreduced: "a+a=0 \<Longrightarrow> \<not> reduced_word A [a,a]"
using sum_list_zero_nreduced by simp
lemma reduced_word_append_reduce_contra1:
assumes "\<not> reduced_word A as"
shows "\<not> reduced_word A (as@bs)"
proof (cases "as \<in> lists A" "bs \<in> lists A" rule: two_cases)
case both
define cs where cs: "cs \<equiv> ARG_MIN length cs. cs \<in> lists A \<and> sum_list cs = sum_list as"
with both(1) have "reduced_word_for A (sum_list as) cs"
using reduced_word_for_def is_arg_min_arg_min_nat[of "word_for A (sum_list as)"]
by auto
with assms both show ?thesis
using reduced_word_for_lists reduced_word_for_sum_list
reduced_word_for_minimal[of A "sum_list as" cs as]
reduced_word_forI_compare[of A "sum_list as" cs as]
not_reduced_word_for[of "cs@bs" A "sum_list (as@bs)"]
by fastforce
next
case one thus ?thesis using reduced_word_for_lists by fastforce
next
case other thus ?thesis using reduced_word_for_lists by fastforce
next
case neither thus ?thesis using reduced_word_for_lists by fastforce
qed
lemma reduced_word_append_reduce_contra2:
assumes "\<not> reduced_word A bs"
shows "\<not> reduced_word A (as@bs)"
proof (cases "as \<in> lists A" "bs \<in> lists A" rule: two_cases)
case both
define cs where cs: "cs \<equiv> ARG_MIN length cs. cs \<in> lists A \<and> sum_list cs = sum_list bs"
with both(2) have "reduced_word_for A (sum_list bs) cs"
using reduced_word_for_def is_arg_min_arg_min_nat[of "word_for A (sum_list bs)" ]
by auto
with assms both show ?thesis
using reduced_word_for_lists reduced_word_for_sum_list
reduced_word_for_minimal[of A "sum_list bs" cs bs]
reduced_word_forI_compare[of A "sum_list bs" cs bs]
not_reduced_word_for[of "as@cs" A "sum_list (as@bs)"]
by fastforce
next
case one thus ?thesis using reduced_word_for_lists by fastforce
next
case other thus ?thesis using reduced_word_for_lists by fastforce
next
case neither thus ?thesis using reduced_word_for_lists by fastforce
qed
lemma contains_nreduced_imp_nreduced:
"\<not> reduced_word A bs \<Longrightarrow> \<not> reduced_word A (as@bs@cs)"
using reduced_word_append_reduce_contra1 reduced_word_append_reduce_contra2
by fast
lemma contains_order2_nreduced: "a+a=0 \<Longrightarrow> \<not> reduced_word A (as@[a,a]@bs)"
using order2_nreduced contains_nreduced_imp_nreduced by fast
lemma reduced_word_Cons_reduce_contra:
"\<not> reduced_word A as \<Longrightarrow> \<not> reduced_word A (a#as)"
using reduced_word_append_reduce_contra2[of A as "[a]"] by simp
lemma reduced_word_Cons_reduce: "reduced_word A (a#as) \<Longrightarrow> reduced_word A as"
using reduced_word_Cons_reduce_contra by fast
lemma reduced_word_singleton:
assumes "a\<in>A" "a\<noteq>0"
shows "reduced_word A [a]"
proof (rule reduced_word_forI)
from assms(1) show "[a] \<in> lists A" by simp
next
fix bs assume bs: "bs \<in> lists A" "sum_list bs = sum_list [a]"
with assms(2) show "length [a] \<le> length bs" by (cases bs) auto
qed simp
lemma el_reduced:
assumes "0 \<notin> A" "as \<in> lists A" "sum_list as \<in> A" "reduced_word A as"
shows "length as = 1"
proof-
define n where n: "n \<equiv> length as"
from assms(3) obtain a where "[a]\<in>lists A" "sum_list as = sum_list [a]" by auto
with n assms(1,3,4) have "n\<le>1" "n>0"
using reduced_word_for_minimal[of A _ as "[a]"] by auto
hence "n = 1" by simp
with n show ?thesis by fast
qed
lemma reduced_letter_set_0: "reduced_letter_set A 0 = {}"
using reduced_word_for_0_imp_nil by simp
lemma reduced_letter_set_subset: "reduced_letter_set A a \<subseteq> A"
using reduced_word_for_lists by fast
lemma reduced_word_forI_length:
"\<lbrakk> as \<in> lists A; sum_list as = a; length as = word_length A a \<rbrakk> \<Longrightarrow>
reduced_word_for A a as"
using reduced_word_for_arg_min reduced_word_for_length
reduced_word_forI_compare[of A a _ as]
by fastforce
lemma word_length_le:
"as \<in> lists A \<Longrightarrow> sum_list as = a \<Longrightarrow> word_length A a \<le> length as"
using reduced_word_for_arg_min reduced_word_for_length
reduced_word_for_minimal[of A]
by fastforce
lemma reduced_word_forI_length':
"\<lbrakk> as \<in> lists A; sum_list as = a; length as \<le> word_length A a \<rbrakk> \<Longrightarrow>
reduced_word_for A a as"
using word_length_le[of as A] reduced_word_forI_length[of as A] by fastforce
lemma word_length_lt:
"as \<in> lists A \<Longrightarrow> sum_list as = a \<Longrightarrow> \<not> reduced_word_for A a as \<Longrightarrow>
word_length A a < length as"
using reduced_word_forI_length' by fastforce
end (* context monoid_add *)
lemma in_genby_reduced_letter_set:
assumes "as \<in> lists A" "sum_list as = a"
shows "a \<in> \<langle>reduced_letter_set A a\<rangle>"
proof-
define xs where xs: "xs \<equiv> arg_min length (word_for A a)"
with assms have "xs \<in> lists (reduced_letter_set A a)" "sum_list xs = a"
using reduced_word_for_arg_min[of as A] reduced_word_for_sum_list by auto
thus ?thesis using genby_eq_sum_lists by force
qed
lemma reduced_word_for_genby_arg_min:
fixes A :: "'a::group_add set"
defines "B \<equiv> A \<union> uminus ` A"
assumes "a\<in>\<langle>A\<rangle>"
shows "reduced_word_for B a (arg_min length (word_for B a))"
using assms genby_eq_sum_lists[of A] reduced_word_for_arg_min[of _ B a]
by auto
lemma reduced_word_for_genby_sym_arg_min:
assumes "uminus ` A \<subseteq> A" "a\<in>\<langle>A\<rangle>"
shows "reduced_word_for A a (arg_min length (word_for A a))"
proof-
from assms(1) have "A = A \<union> uminus ` A" by auto
with assms(2) show ?thesis
using reduced_word_for_genby_arg_min[of a A] by simp
qed
lemma in_genby_imp_in_reduced_letter_set:
fixes A :: "'a::group_add set"
defines "B \<equiv> A \<union> uminus ` A"
assumes "a \<in> \<langle>A\<rangle>"
shows "a \<in> \<langle>reduced_letter_set B a\<rangle>"
using assms genby_eq_sum_lists[of A] in_genby_reduced_letter_set[of _ B]
by auto
lemma in_genby_sym_imp_in_reduced_letter_set:
"uminus ` A \<subseteq> A \<Longrightarrow> a \<in> \<langle>A\<rangle> \<Longrightarrow> a \<in> \<langle>reduced_letter_set A a\<rangle>"
using in_genby_imp_in_reduced_letter_set by (fastforce simp add: Un_absorb2)
end (* theory *)
diff --git a/thys/Buildings/Prelim.thy b/thys/Buildings/Prelim.thy
--- a/thys/Buildings/Prelim.thy
+++ b/thys/Buildings/Prelim.thy
@@ -1,1742 +1,1742 @@
section \<open>Preliminaries\<close>
text \<open>
In this section, we establish some basic facts about natural numbers, logic, sets, functions and
relations, lists, and orderings and posets, that are either not available in the HOL library or
are in a form not suitable for our purposes.
\<close>
theory Prelim
imports Main "HOL-Library.Set_Algebras"
begin
declare image_cong_simp [cong del]
subsection \<open>Natural numbers\<close>
lemma nat_cases_2Suc [case_names 0 1 SucSuc]:
assumes 0: "n = 0 \<Longrightarrow> P"
and 1: "n = 1 \<Longrightarrow> P"
and SucSuc: "\<And>m. n = Suc (Suc m) \<Longrightarrow> P"
shows "P"
proof (cases n)
case (Suc m) with 1 SucSuc show ?thesis by (cases m) auto
qed (simp add: 0)
lemma nat_even_induct [case_names _ 0 SucSuc]:
assumes even: "even n"
and 0: "P 0"
and SucSuc: "\<And>m. even m \<Longrightarrow> P m \<Longrightarrow> P (Suc (Suc m))"
shows "P n"
proof-
from assms obtain k where "n = 2*k" using evenE by auto
moreover from assms have "P (2*k)" by (induct k) auto
ultimately show ?thesis by fast
qed
lemma nat_induct_step2 [case_names 0 1 SucSuc]:
assumes 0: "P 0"
and 1: "P 1"
and SucSuc: "\<And>m. P m \<Longrightarrow> P (Suc (Suc m))"
shows "P n"
proof (cases "even n")
case True
from this obtain k where "n = 2*k" using evenE by auto
moreover have "P (2*k)" using 0 SucSuc by (induct k) auto
ultimately show ?thesis by fast
next
case False
from this obtain k where "n = 2*k+1" using oddE by blast
moreover have "P (2*k+1)" using 1 SucSuc by (induct k) auto
ultimately show ?thesis by fast
qed
subsection \<open>Logic\<close>
lemma ex1_unique: "\<exists>!x. P x \<Longrightarrow> P a \<Longrightarrow> P b \<Longrightarrow> a=b"
by blast
lemma not_the1:
assumes "\<exists>!x. P x" "y \<noteq> (THE x. P x)"
shows "\<not> P y"
using assms(2) the1_equality[OF assms(1)]
by auto
lemma two_cases [case_names both one other neither]:
assumes both : "P \<Longrightarrow> Q \<Longrightarrow> R"
and one : "P \<Longrightarrow> \<not>Q \<Longrightarrow> R"
and other : "\<not>P \<Longrightarrow> Q \<Longrightarrow> R"
and neither: "\<not>P \<Longrightarrow> \<not>Q \<Longrightarrow> R"
shows "R"
using assms
by fast
subsection \<open>Sets\<close>
lemma bex1_equality: "\<lbrakk> \<exists>!x\<in>A. P x; x\<in>A; P x; y\<in>A; P y \<rbrakk> \<Longrightarrow> x=y"
by blast
lemma prod_ballI: "(\<And>a b. (a,b)\<in>A \<Longrightarrow> P a b) \<Longrightarrow> \<forall>(a,b)\<in>A. P a b"
by fast
lemmas seteqI = set_eqI[OF iffI]
lemma set_decomp_subset:
"\<lbrakk> U = A\<union>B; A\<subseteq>X; B\<subseteq>Y; X\<subseteq>U; X\<inter>Y = {} \<rbrakk> \<Longrightarrow> A = X"
by auto
lemma insert_subset_equality: "\<lbrakk> a\<notin>A; a\<notin>B; insert a A = insert a B \<rbrakk> \<Longrightarrow> A=B"
by auto
lemma insert_compare_element: "a\<notin>A \<Longrightarrow> insert b A = insert a A \<Longrightarrow> b=a"
by auto
lemma card1:
assumes "card A = 1"
shows "\<exists>a. A = {a}"
proof-
from assms obtain a where a: "a \<in> A" by fastforce
with assms show ?thesis using card_ge_0_finite[of A] card_subset_eq[of A "{a}"] by auto
qed
lemma singleton_pow: "a\<in>A \<Longrightarrow> {a}\<in>Pow A"
using Pow_mono Pow_top by fast
definition separated_by :: "'a set set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
where "separated_by w x y \<equiv> \<exists>A B. w={A,B} \<and> x\<in>A \<and> y\<in>B"
lemma separated_byI: "x\<in>A \<Longrightarrow> y\<in>B \<Longrightarrow> separated_by {A,B} x y"
using separated_by_def by fastforce
lemma separated_by_disjoint: "\<lbrakk> separated_by {A,B} x y; A\<inter>B={}; x\<in>A \<rbrakk> \<Longrightarrow> y\<in>B"
unfolding separated_by_def by fast
lemma separated_by_in_other: "separated_by {A,B} x y \<Longrightarrow> x\<notin>A \<Longrightarrow> x\<in>B \<and> y\<in>A"
unfolding separated_by_def by auto
lemma separated_by_not_empty: "separated_by w x y \<Longrightarrow> w\<noteq>{}"
unfolding separated_by_def by fast
lemma not_self_separated_by_disjoint: "A\<inter>B={} \<Longrightarrow> \<not> separated_by {A,B} x x"
unfolding separated_by_def by auto
subsection \<open>Functions and relations\<close>
subsubsection \<open>Miscellaneous\<close>
lemma cong_let: "(let x = y in f x) = f y" by simp
lemma sym_sym: "sym (A\<times>A)" by (fast intro: symI)
lemma trans_sym: "trans (A\<times>A)" by (fast intro: transI)
lemma map_prod_sym: "sym A \<Longrightarrow> sym (map_prod f f ` A)"
using symD[of A] map_prod_def by (fast intro: symI)
abbreviation restrict1 :: "('a\<Rightarrow>'a) \<Rightarrow> 'a set \<Rightarrow> ('a\<Rightarrow>'a)"
where "restrict1 f A \<equiv> (\<lambda>a. if a\<in>A then f a else a)"
lemma restrict1_image: "B\<subseteq>A \<Longrightarrow> restrict1 f A ` B = f`B"
by auto
subsubsection \<open>Equality of functions restricted to a set\<close>
definition "fun_eq_on f g A \<equiv> (\<forall>a\<in>A. f a = g a)"
lemma fun_eq_onI: "(\<And>a. a\<in>A \<Longrightarrow> f a = g a) \<Longrightarrow> fun_eq_on f g A"
using fun_eq_on_def by fast
lemma fun_eq_onD: "fun_eq_on f g A \<Longrightarrow> a \<in> A \<Longrightarrow> f a = g a"
using fun_eq_on_def by fast
lemma fun_eq_on_UNIV: "(fun_eq_on f g UNIV) = (f=g)"
unfolding fun_eq_on_def by fast
lemma fun_eq_on_subset: "fun_eq_on f g A \<Longrightarrow> B\<subseteq>A \<Longrightarrow> fun_eq_on f g B"
unfolding fun_eq_on_def by fast
lemma fun_eq_on_sym: "fun_eq_on f g A \<Longrightarrow> fun_eq_on g f A"
using fun_eq_onD by (fastforce intro: fun_eq_onI)
lemma fun_eq_on_trans: "fun_eq_on f g A \<Longrightarrow> fun_eq_on g h A \<Longrightarrow> fun_eq_on f h A"
using fun_eq_onD fun_eq_onD by (fastforce intro: fun_eq_onI)
lemma fun_eq_on_cong: "fun_eq_on f h A \<Longrightarrow> fun_eq_on g h A \<Longrightarrow> fun_eq_on f g A"
using fun_eq_on_trans fun_eq_on_sym by fastforce
lemma fun_eq_on_im : "fun_eq_on f g A \<Longrightarrow> B\<subseteq>A \<Longrightarrow> f`B = g`B"
using fun_eq_onD by force
lemma fun_eq_on_subset_and_diff_imp_eq_on:
assumes "A\<subseteq>B" "fun_eq_on f g A" "fun_eq_on f g (B-A)"
shows "fun_eq_on f g B"
proof (rule fun_eq_onI)
fix x assume "x\<in>B" with assms(1) show "f x = g x"
using fun_eq_onD[OF assms(2)] fun_eq_onD[OF assms(3)]
by (cases "x\<in>A") auto
qed
lemma fun_eq_on_set_and_comp_imp_eq:
"fun_eq_on f g A \<Longrightarrow> fun_eq_on f g (-A) \<Longrightarrow> f = g"
using fun_eq_on_subset_and_diff_imp_eq_on[of A UNIV]
by (simp add: Compl_eq_Diff_UNIV fun_eq_on_UNIV)
lemma fun_eq_on_bij_betw: "fun_eq_on f g A \<Longrightarrow> bij_betw f A B = bij_betw g A B"
using bij_betw_cong unfolding fun_eq_on_def by fast
lemma fun_eq_on_restrict1: "fun_eq_on (restrict1 f A) f A"
by (auto intro: fun_eq_onI)
abbreviation "fixespointwise f A \<equiv> fun_eq_on f id A"
lemmas fixespointwiseI = fun_eq_onI [of _ _ id]
lemmas fixespointwiseD = fun_eq_onD [of _ id]
lemmas fixespointwise_cong = fun_eq_on_trans [of _ _ _ id]
lemmas fixespointwise_subset = fun_eq_on_subset [of _ id]
lemmas fixespointwise2_imp_eq_on = fun_eq_on_cong [of _ id]
lemmas fixespointwise_subset_and_diff_imp_eq_on =
fun_eq_on_subset_and_diff_imp_eq_on[of _ _ _ id]
lemma id_fixespointwise: "fixespointwise id A"
using fun_eq_on_def by fast
lemma fixespointwise_im: "fixespointwise f A \<Longrightarrow> B\<subseteq>A \<Longrightarrow> f`B = B"
by (auto simp add: fun_eq_on_im)
lemma fixespointwise_comp:
"fixespointwise f A \<Longrightarrow> fixespointwise g A \<Longrightarrow> fixespointwise (g\<circ>f) A"
unfolding fun_eq_on_def by simp
lemma fixespointwise_insert:
assumes "fixespointwise f A" "f ` (insert a A) = insert a A"
shows "fixespointwise f (insert a A)"
using assms(2) insert_compare_element[of a A "f a"]
fixespointwiseD[OF assms(1)] fixespointwise_im[OF assms(1)]
by (cases "a\<in>A") (auto intro: fixespointwiseI)
lemma fixespointwise_restrict1:
"fixespointwise f A \<Longrightarrow> fixespointwise (restrict1 f B) A"
using fixespointwiseD[of f] by (auto intro: fixespointwiseI)
lemma fold_fixespointwise:
"\<forall>x\<in>set xs. fixespointwise (f x) A \<Longrightarrow> fixespointwise (fold f xs) A"
proof (induct xs)
case Nil show ?case using id_fixespointwise subst[of id] by fastforce
next
case (Cons x xs)
hence "fixespointwise (fold f xs \<circ> f x) A"
using fixespointwise_comp[of "f x" A "fold f xs"] by fastforce
moreover have "fold f xs \<circ> f x = fold f (x#xs)" by simp
ultimately show ?case using subst[of _ _ "\<lambda>f. fixespointwise f A"] by fast
qed
lemma funpower_fixespointwise:
assumes "fixespointwise f A"
shows "fixespointwise (f^^n) A"
proof (induct n)
case 0 show ?case using id_fixespointwise subst[of id] by fastforce
next
case (Suc m)
with assms have "fixespointwise (f \<circ> (f^^m)) A"
using fixespointwise_comp by fast
moreover have "f \<circ> (f^^m) = f^^(Suc m)" by simp
ultimately show ?case using subst[of _ _ "\<lambda>f. fixespointwise f A"] by fast
qed
subsubsection \<open>Injectivity, surjectivity, bijectivity, and inverses\<close>
lemma inj_on_to_singleton:
assumes "inj_on f A" "f`A = {b}"
shows "\<exists>a. A = {a}"
proof-
from assms(2) obtain a where a: "a\<in>A" "f a = b" by force
with assms have "A = {a}" using inj_onD[of f A] by blast
thus ?thesis by fast
qed
lemmas inj_inj_on = subset_inj_on[of _ UNIV, OF _ subset_UNIV]
lemma inj_on_eq_image': "\<lbrakk> inj_on f A; X\<subseteq>A; Y\<subseteq>A; f`X\<subseteq>f`Y \<rbrakk> \<Longrightarrow> X\<subseteq>Y"
unfolding inj_on_def by fast
lemma inj_on_eq_image: "\<lbrakk> inj_on f A; X\<subseteq>A; Y\<subseteq>A; f`X=f`Y \<rbrakk> \<Longrightarrow> X=Y"
using inj_on_eq_image'[of f A X Y] inj_on_eq_image'[of f A Y X] by simp
lemmas inj_eq_image = inj_on_eq_image[OF _ subset_UNIV subset_UNIV]
lemma induced_pow_fun_inj_on:
assumes "inj_on f A"
shows "inj_on ((`) f) (Pow A)"
using inj_onD[OF assms] inj_onI[of "Pow A" "(`) f"]
by blast
lemma inj_on_minus_set: "inj_on ((-) A) (Pow A)"
by (fast intro: inj_onI)
lemma induced_pow_fun_surj:
"((`) f) ` (Pow A) = Pow (f`A)"
proof (rule seteqI)
fix X show "X \<in> ((`) f) ` (Pow A) \<Longrightarrow> X \<in> Pow (f`A)" by fast
next
fix Y assume Y: "Y \<in> Pow (f`A)"
moreover hence "Y = f`{a\<in>A. f a \<in> Y}" by fast
ultimately show "Y\<in> ((`) f) ` (Pow A)" by auto
qed
lemma bij_betw_f_the_inv_into_f:
"bij_betw f A B \<Longrightarrow> y\<in>B \<Longrightarrow> f (the_inv_into A f y) = y"
\<comment> \<open>an equivalent lemma appears in the HOL library, but this version avoids the double
@{const bij_betw} premises\<close>
unfolding bij_betw_def by (blast intro: f_the_inv_into_f)
lemma bij_betw_the_inv_into_onto: "bij_betw f A B \<Longrightarrow> the_inv_into A f ` B = A"
unfolding bij_betw_def by force
lemma bij_betw_imp_bij_betw_Pow:
assumes "bij_betw f A B"
shows "bij_betw ((`) f) (Pow A) (Pow B)"
unfolding bij_betw_def
proof (rule conjI, rule inj_onI)
show "\<And>x y. \<lbrakk> x\<in>Pow A; y\<in>Pow A; f`x = f`y \<rbrakk> \<Longrightarrow> x=y"
using inj_onD[OF bij_betw_imp_inj_on, OF assms] by blast
show "(`) f ` Pow A = Pow B"
proof
show "(`) f ` Pow A \<subseteq> Pow B" using bij_betw_imp_surj_on[OF assms] by fast
show "(`) f ` Pow A \<supseteq> Pow B"
proof
fix y assume y: "y\<in>Pow B"
with assms have "y = f ` the_inv_into A f ` y"
using bij_betw_f_the_inv_into_f[THEN sym] by fastforce
moreover from y assms have "the_inv_into A f ` y \<subseteq> A"
using bij_betw_the_inv_into_onto by fastforce
ultimately show "y \<in> (`) f ` Pow A" by auto
qed
qed
qed
lemma comps_fixpointwise_imp_bij_betw:
assumes "f`X\<subseteq>Y" "g`Y\<subseteq>X" "fixespointwise (g\<circ>f) X" "fixespointwise (f\<circ>g) Y"
shows "bij_betw f X Y"
unfolding bij_betw_def
proof
show "inj_on f X"
proof (rule inj_onI)
fix x y show "\<lbrakk> x\<in>X; y\<in>X; f x = f y \<rbrakk> \<Longrightarrow> x=y"
using fixespointwiseD[OF assms(3), of x] fixespointwiseD[OF assms(3), of y]
by simp
qed
from assms(1,2) show "f`X = Y" using fixespointwiseD[OF assms(4)] by force
qed
lemma set_permutation_bij_restrict1:
assumes "bij_betw f A A"
shows "bij (restrict1 f A)"
proof (rule bijI)
have bij_f: "inj_on f A" "f`A = A" using iffD1[OF bij_betw_def, OF assms] by auto
show "inj (restrict1 f A)"
proof (rule injI)
fix x y show "restrict1 f A x = restrict1 f A y \<Longrightarrow> x=y"
using inj_onD bij_f by (cases "x\<in>A" "y\<in>A" rule: two_cases) auto
qed
show "surj (restrict1 f A)"
proof (rule surjI)
fix x
define y where "y \<equiv> restrict1 (the_inv_into A f) A x"
thus "restrict1 f A y = x"
using the_inv_into_into[of f] bij_f f_the_inv_into_f[of f] by (cases "x\<in>A") auto
qed
qed
lemma set_permutation_the_inv_restrict1:
assumes "bij_betw f A A"
shows "the_inv (restrict1 f A) = restrict1 (the_inv_into A f) A"
proof (rule ext, rule the_inv_into_f_eq)
from assms show "inj (restrict1 f A)"
using bij_is_inj set_permutation_bij_restrict1 by fast
next
fix a from assms show "restrict1 f A (restrict1 (the_inv_into A f) A a) = a"
using bij_betw_def[of f] by (simp add: the_inv_into_into f_the_inv_into_f)
qed simp
lemma the_inv_into_the_inv_into:
"inj_on f A \<Longrightarrow> a\<in>A \<Longrightarrow> the_inv_into (f`A) (the_inv_into A f) a = f a"
using inj_on_the_inv_into by (force intro: the_inv_into_f_eq imageI)
lemma the_inv_into_f_im_f_im:
assumes "inj_on f A" "x\<subseteq>A"
shows "the_inv_into A f ` f ` x = x"
using assms(2) the_inv_into_f_f[OF assms(1)]
by force
lemma f_im_the_inv_into_f_im:
assumes "inj_on f A" "x\<subseteq>f`A"
shows "f ` the_inv_into A f ` x = x"
using assms(2) f_the_inv_into_f[OF assms(1)]
by force
lemma the_inv_leftinv: "bij f \<Longrightarrow> the_inv f \<circ> f = id"
using bij_def[of f] the_inv_f_f by fastforce
subsubsection \<open>Induced functions on sets of sets and lists of sets\<close>
text \<open>
Here we create convenience abbreviations for distributing a function over a set of sets and over
a list of sets.
\<close>
abbreviation setsetmapim :: "('a\<Rightarrow>'b) \<Rightarrow> 'a set set \<Rightarrow> 'b set set" (infix "\<turnstile>" 70)
where "f\<turnstile>X \<equiv> ((`) f) ` X"
abbreviation setlistmapim :: "('a\<Rightarrow>'b) \<Rightarrow> 'a set list \<Rightarrow> 'b set list" (infix "\<Turnstile>" 70)
where "f\<Turnstile>Xs \<equiv> map ((`) f) Xs"
lemma setsetmapim_comp: "(f\<circ>g)\<turnstile>A = f\<turnstile>(g\<turnstile>A)"
by (auto simp add: image_comp)
lemma setlistmapim_comp: "(f\<circ>g)\<Turnstile>xs = f\<Turnstile>(g\<Turnstile>xs)"
by auto
lemma setsetmapim_cong_subset:
assumes "fun_eq_on g f (\<Union>A)" "B\<subseteq>A"
shows "g\<turnstile>B \<subseteq> f\<turnstile>B"
proof
fix y assume "y \<in> g\<turnstile>B"
from this obtain x where "x\<in>B" "y = g`x" by fast
with assms(2) show "y \<in> f\<turnstile>B" using fun_eq_on_im[OF assms(1), of x] by fast
qed
lemma setsetmapim_cong:
assumes "fun_eq_on g f (\<Union>A)" "B\<subseteq>A"
shows "g\<turnstile>B = f\<turnstile>B"
using setsetmapim_cong_subset[OF assms]
setsetmapim_cong_subset[OF fun_eq_on_sym, OF assms]
by fast
lemma setsetmapim_restrict1: "B\<subseteq>A \<Longrightarrow> restrict1 f (\<Union>A) \<turnstile> B = f\<turnstile>B"
using setsetmapim_cong[of _ f] fun_eq_on_restrict1[of "\<Union>A" f] by simp
lemma setsetmapim_the_inv_into:
assumes "inj_on f (\<Union>A)"
shows "(the_inv_into (\<Union>A) f) \<turnstile> (f\<turnstile>A) = A"
proof (rule seteqI)
fix x assume "x \<in> (the_inv_into (\<Union>A) f) \<turnstile> (f\<turnstile>A)"
from this obtain y where y: "y \<in> f\<turnstile>A" "x = the_inv_into (\<Union>A) f ` y" by auto
from y(1) obtain z where z: "z\<in>A" "y = f`z" by fast
moreover from z(1) have "the_inv_into (\<Union>A) f ` f ` z = z"
using the_inv_into_f_f[OF assms] by force
ultimately show "x\<in>A" using y(2) the_inv_into_f_im_f_im[OF assms] by simp
next
fix x assume x: "x\<in>A"
moreover hence "the_inv_into (\<Union>A) f ` f ` x = x"
using the_inv_into_f_im_f_im[OF assms, of x] by fast
ultimately show "x \<in> (the_inv_into (\<Union>A) f) \<turnstile> (f\<turnstile>A)" by auto
qed
subsubsection \<open>Induced functions on quotients\<close>
text \<open>
Here we construct the induced function on a quotient for an inducing function that respects the
relation that defines the quotient.
\<close>
lemma respects_imp_unique_image_rel: "f respects r \<Longrightarrow> y\<in>f`r``{a} \<Longrightarrow> y = f a"
using congruentD[of r f] by auto
lemma ex1_class_image:
assumes "refl_on A r" "f respects r" "X\<in>A//r"
shows "\<exists>!b. b\<in>f`X"
proof-
from assms(3) obtain a where a: "a\<in>A" "X = r``{a}" by (auto intro: quotientE)
thus ?thesis
using refl_onD[OF assms(1)] ex1I[of _ "f a"]
respects_imp_unique_image_rel[OF assms(2), of _ a]
by force
qed
definition quotientfun :: "('a\<Rightarrow>'b) \<Rightarrow> 'a set \<Rightarrow> 'b"
where "quotientfun f X = (THE b. b\<in>f`X)"
lemma quotientfun_equality:
assumes "refl_on A r" "f respects r" "X\<in>A//r" "b\<in>f`X"
shows "quotientfun f X = b"
unfolding quotientfun_def
using assms(4) ex1_class_image[OF assms(1-3)]
by (auto intro: the1_equality)
lemma quotientfun_classrep_equality:
"\<lbrakk> refl_on A r; f respects r; a\<in>A \<rbrakk> \<Longrightarrow> quotientfun f (r``{a}) = f a"
using refl_onD by (fastforce intro: quotientfun_equality quotientI)
subsubsection \<open>Support of a function\<close>
definition supp :: "('a \<Rightarrow> 'b::zero) \<Rightarrow> 'a set" where "supp f = {x. f x \<noteq> 0}"
lemma suppI_contra: "x \<notin> supp f \<Longrightarrow> f x = 0"
using supp_def by fast
lemma suppD_contra: "f x = 0 \<Longrightarrow> x \<notin> supp f"
using supp_def by fast
abbreviation restrict0 :: "('a\<Rightarrow>'b::zero) \<Rightarrow> 'a set \<Rightarrow> ('a\<Rightarrow>'b)"
where "restrict0 f A \<equiv> (\<lambda>a. if a \<in> A then f a else 0)"
lemma supp_restrict0 : "supp (restrict0 f A) \<subseteq> A"
proof-
have "\<And>a. a \<notin> A \<Longrightarrow> a \<notin> supp (restrict0 f A)"
using suppD_contra[of "restrict0 f A"] by simp
thus ?thesis by fast
qed
subsection \<open>Lists\<close>
subsubsection \<open>Miscellaneous facts\<close>
lemma snoc_conv_cons: "\<exists>x xs. ys@[y] = x#xs"
by (cases ys) auto
lemma cons_conv_snoc: "\<exists>ys y. x#xs = ys@[y]"
by (cases xs rule: rev_cases) auto
lemma same_length_eq_append:
"length as = length bs \<Longrightarrow> as@cs = bs@ds \<Longrightarrow> as = bs"
by (induct as bs rule: list_induct2) auto
lemma count_list_append:
"count_list (xs@ys) a = count_list xs a + count_list ys a"
by (induct xs) auto
lemma count_list_snoc:
"count_list (xs@[x]) y = (if y=x then Suc (count_list xs y) else count_list xs y)"
by (induct xs) auto
lemma distinct_count_list:
"distinct xs \<Longrightarrow> count_list xs a = (if a \<in> set xs then 1 else 0)"
by (induct xs) auto
lemma map_fst_map_const_snd: "map fst (map (\<lambda>s. (s,b)) xs) = xs"
by (induct xs) auto
lemma inj_on_distinct_setlistmapim:
assumes "inj_on f A"
shows "\<forall>X\<in>set Xs. X \<subseteq> A \<Longrightarrow> distinct Xs \<Longrightarrow> distinct (f\<Turnstile>Xs)"
proof (induct Xs)
case (Cons X Xs)
show ?case
proof (cases "f`X \<in> set (f\<Turnstile>Xs)")
case True
from this obtain Y where Y: "Y\<in>set Xs" "f`X = f`Y" by auto
with assms Y(1) Cons(2,3) show ?thesis
using inj_on_eq_image[of f A X Y] by fastforce
next
case False with Cons show ?thesis by simp
qed
qed simp
subsubsection \<open>Cases\<close>
lemma list_cases_Cons_snoc [case_names Nil Single Cons_snoc]:
assumes Nil: "xs = [] \<Longrightarrow> P"
and Single: "\<And>x. xs = [x] \<Longrightarrow> P"
and Cons_snoc: "\<And>x ys y. xs = x # ys @ [y] \<Longrightarrow> P"
shows "P"
proof (cases xs, rule Nil)
case (Cons x xs) with Single Cons_snoc show ?thesis
by (cases xs rule: rev_cases) auto
qed
lemma two_lists_cases_Cons_Cons [case_names Nil1 Nil2 ConsCons]:
assumes Nil1: "\<And>ys. as = [] \<Longrightarrow> bs = ys \<Longrightarrow> P"
and Nil2: "\<And>xs. as = xs \<Longrightarrow> bs = [] \<Longrightarrow> P"
and ConsCons: "\<And>x xs y ys. as = x # xs \<Longrightarrow> bs = y # ys \<Longrightarrow> P"
shows "P"
proof (cases as)
case Cons with assms(2,3) show ?thesis by (cases bs) auto
qed (simp add: Nil1)
lemma two_lists_cases_snoc_Cons [case_names Nil1 Nil2 snoc_Cons]:
assumes Nil1: "\<And>ys. as = [] \<Longrightarrow> bs = ys \<Longrightarrow> P"
and Nil2: "\<And>xs. as = xs \<Longrightarrow> bs = [] \<Longrightarrow> P"
and snoc_Cons: "\<And>xs x y ys. as = xs @ [x] \<Longrightarrow> bs = y # ys \<Longrightarrow> P"
shows "P"
proof (cases as rule: rev_cases)
case snoc with Nil2 snoc_Cons show ?thesis by (cases bs) auto
qed (simp add: Nil1)
lemma two_lists_cases_snoc_Cons' [case_names both_Nil Nil1 Nil2 snoc_Cons]:
assumes both_Nil: "as = [] \<Longrightarrow> bs = [] \<Longrightarrow> P"
and Nil1: "\<And>y ys. as = [] \<Longrightarrow> bs = y#ys \<Longrightarrow> P"
and Nil2: "\<And>xs x. as = xs@[x] \<Longrightarrow> bs = [] \<Longrightarrow> P"
and snoc_Cons: "\<And>xs x y ys. as = xs @ [x] \<Longrightarrow> bs = y # ys \<Longrightarrow> P"
shows "P"
proof (cases as bs rule: two_lists_cases_snoc_Cons)
case (Nil1 ys) with assms(1,2) show "P" by (cases ys) auto
next
case (Nil2 xs) with assms(1,3) show "P" by (cases xs rule: rev_cases) auto
qed (rule snoc_Cons)
lemma two_prod_lists_cases_snoc_Cons:
- assumes "\<And>xs. as = xs \<and> bs = [] \<Longrightarrow> P" "\<And>ys. as = [] \<and> bs = ys \<Longrightarrow> P"
+ assumes "\<And>xs. as = xs \<Longrightarrow> bs = [] \<Longrightarrow> P" "\<And>ys. as = [] \<Longrightarrow> bs = ys \<Longrightarrow> P"
"\<And>xs aa ba ab bb ys. as = xs @ [(aa, ba)] \<and> bs = (ab, bb) # ys \<Longrightarrow> P"
shows "P"
proof (rule two_lists_cases_snoc_Cons)
from assms
show "\<And>ys. as = [] \<Longrightarrow> bs = ys \<Longrightarrow> P" "\<And>xs. as = xs \<Longrightarrow> bs = [] \<Longrightarrow> P"
by auto
from assms(3) show "\<And>xs x y ys. as = xs @ [x] \<Longrightarrow> bs = y # ys \<Longrightarrow> P"
by fast
qed
lemma three_lists_cases_snoc_mid_Cons
[case_names Nil1 Nil2 Nil3 snoc_single_Cons snoc_mid_Cons]:
assumes Nil1: "\<And>ys zs. as = [] \<Longrightarrow> bs = ys \<Longrightarrow> cs = zs \<Longrightarrow> P"
and Nil2: "\<And>xs zs. as = xs \<Longrightarrow> bs = [] \<Longrightarrow> cs = zs \<Longrightarrow> P"
and Nil3: "\<And>xs ys. as = xs \<Longrightarrow> bs = ys \<Longrightarrow> cs = [] \<Longrightarrow> P"
and snoc_single_Cons:
"\<And>xs x y z zs. as = xs @ [x] \<Longrightarrow> bs = [y] \<Longrightarrow> cs = z # zs \<Longrightarrow> P"
and snoc_mid_Cons:
"\<And>xs x w ys y z zs. as = xs @ [x] \<Longrightarrow> bs = w # ys @ [y] \<Longrightarrow>
cs = z # zs \<Longrightarrow> P"
shows "P"
proof (cases as cs rule: two_lists_cases_snoc_Cons)
case Nil1 with assms(1) show "P" by simp
next
case Nil2 with assms(3) show "P" by simp
next
case snoc_Cons
with Nil2 snoc_single_Cons snoc_mid_Cons show "P"
by (cases bs rule: list_cases_Cons_snoc) auto
qed
subsubsection \<open>Induction\<close>
lemma list_induct_CCons [case_names Nil Single CCons]:
assumes Nil : "P []"
and Single: "\<And>x. P [x]"
and CCons : "\<And>x y xs. P (y#xs) \<Longrightarrow> P (x # y # xs)"
shows "P xs"
proof (induct xs)
case (Cons x xs) with Single CCons show ?case by (cases xs) auto
qed (rule Nil)
lemma list_induct_ssnoc [case_names Nil Single ssnoc]:
assumes Nil : "P []"
and Single: "\<And>x. P [x]"
and ssnoc : "\<And>xs x y. P (xs@[x]) \<Longrightarrow> P (xs@[x,y])"
shows "P xs"
proof (induct xs rule: rev_induct)
case (snoc x xs) with Single ssnoc show ?case by (cases xs rule: rev_cases) auto
qed (rule Nil)
lemma list_induct2_snoc [case_names Nil1 Nil2 snoc]:
assumes Nil1: "\<And>ys. P [] ys"
and Nil2: "\<And>xs. P xs []"
and snoc: "\<And>xs x ys y. P xs ys \<Longrightarrow> P (xs@[x]) (ys@[y])"
shows "P xs ys"
proof (induct xs arbitrary: ys rule: rev_induct, rule Nil1)
case (snoc b bs) with assms(2,3) show ?case by (cases ys rule: rev_cases) auto
qed
lemma list_induct2_snoc_Cons [case_names Nil1 Nil2 snoc_Cons]:
assumes Nil1 : "\<And>ys. P [] ys"
and Nil2 : "\<And>xs. P xs []"
and snoc_Cons: "\<And>xs x y ys. P xs ys \<Longrightarrow> P (xs@[x]) (y#ys)"
shows "P xs ys"
proof (induct ys arbitrary: xs, rule Nil2)
case (Cons y ys) with Nil1 snoc_Cons show ?case
by (cases xs rule: rev_cases) auto
qed
lemma prod_list_induct3_snoc_Conssnoc_Cons_pairwise:
assumes "\<And>ys zs. Q ([],ys,zs)" "\<And>xs zs. Q (xs,[],zs)" "\<And>xs ys. Q (xs,ys,[])"
"\<And>xs x y z zs. Q (xs@[x],[y],z#zs)"
and step:
"\<And>xs x y ys w z zs. Q (xs,ys,zs) \<Longrightarrow> Q (xs,ys@[w],z#zs) \<Longrightarrow>
Q (xs@[x],y#ys,zs) \<Longrightarrow> Q (xs@[x],y#ys@[w],z#zs)"
shows "Q t"
proof (
induct t
taking: "\<lambda>(xs,ys,zs). length xs + length ys + length zs"
rule : measure_induct_rule
)
case (less t)
show ?case
proof (cases t)
case (fields xs ys zs) from assms less fields show ?thesis
by (cases xs ys zs rule: three_lists_cases_snoc_mid_Cons) auto
qed
qed
lemma list_induct3_snoc_Conssnoc_Cons_pairwise
[case_names Nil1 Nil2 Nil3 snoc_single_Cons snoc_Conssnoc_Cons]:
assumes Nil1 : "\<And>ys zs. P [] ys zs"
and Nil2 : "\<And>xs zs. P xs [] zs"
and Nil3 : "\<And>xs ys. P xs ys []"
and snoc_single_Cons : "\<And>xs x y z zs. P (xs@[x]) [y] (z#zs)"
and snoc_Conssnoc_Cons:
"\<And>xs x y ys w z zs. P xs ys zs \<Longrightarrow> P xs (ys@[w]) (z#zs) \<Longrightarrow>
P (xs@[x]) (y#ys) zs \<Longrightarrow> P (xs@[x]) (y#ys@[w]) (z#zs)"
shows "P xs ys zs"
using assms
prod_list_induct3_snoc_Conssnoc_Cons_pairwise[of "\<lambda>(xs,ys,zs). P xs ys zs"]
by auto
subsubsection \<open>Alternating lists\<close>
primrec alternating_list :: "nat \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a list"
where zero: "alternating_list 0 s t = []"
| Suc : "alternating_list (Suc k) s t =
alternating_list k s t @ [if even k then s else t]"
\<comment> \<open>could be defined using Cons, but we want the alternating list to always start with the same
letter as it grows, and it's easier to do that via append\<close>
lemma alternating_list2: "alternating_list 2 s t = [s,t]"
using arg_cong[OF Suc_1, THEN sym, of "\<lambda>n. alternating_list n s t"] by simp
lemma length_alternating_list: "length (alternating_list n s t) = n"
by (induct n) auto
lemma alternating_list_Suc_Cons:
"alternating_list (Suc k) s t = s # alternating_list k t s"
by (induct k) auto
lemma alternating_list_SucSuc_ConsCons:
"alternating_list (Suc (Suc k)) s t = s # t # alternating_list k s t"
using alternating_list_Suc_Cons[of "Suc k" s] alternating_list_Suc_Cons[of k t]
by simp
lemma alternating_list_alternates:
"alternating_list n s t = as@[a,b,c]@bs \<Longrightarrow> a=c"
proof (induct n arbitrary: bs)
case (Suc m) hence prevcase:
"\<And>xs. alternating_list m s t = as @ [a,b,c] @ xs \<Longrightarrow> a = c"
"alternating_list (Suc m) s t = as @ [a,b,c] @ bs"
by auto
show ?case
proof (cases bs rule: rev_cases)
case Nil show ?thesis
proof (cases m)
case 0 with prevcase(2) show ?thesis by simp
next
case (Suc k) with prevcase(2) Nil show ?thesis by (cases k) auto
qed
next
case (snoc ds d) with prevcase show ?thesis by simp
qed
qed simp
lemma alternating_list_split:
"alternating_list (m+n) s t = alternating_list m s t @
(if even m then alternating_list n s t else alternating_list n t s)"
using alternating_list_SucSuc_ConsCons[of _ s]
by (induct n rule: nat_induct_step2) auto
lemma alternating_list_append:
"even m \<Longrightarrow>
alternating_list m s t @ alternating_list n s t = alternating_list (m+n) s t"
"odd m \<Longrightarrow>
alternating_list m s t @ alternating_list n t s = alternating_list (m+n) s t"
using alternating_list_split[THEN sym, of m] by auto
lemma rev_alternating_list:
"rev (alternating_list n s t) =
(if even n then alternating_list n t s else alternating_list n s t)"
using alternating_list_SucSuc_ConsCons[of _ s]
by (induct n rule: nat_induct_step2) auto
lemma set_alternating_list: "set (alternating_list n s t) \<subseteq> {s,t}"
by (induct n) auto
lemma set_alternating_list1:
assumes "n \<ge> 1"
shows "s \<in> set (alternating_list n s t)"
proof (cases n)
case 0 with assms show ?thesis by simp
next
case (Suc m) thus ?thesis using alternating_list_Suc_Cons[of m s] by simp
qed
lemma set_alternating_list2:
"n \<ge> 2 \<Longrightarrow> set (alternating_list n s t) = {s,t}"
proof (induct n rule: nat_induct_step2)
case (SucSuc m) thus ?case
using set_alternating_list alternating_list_SucSuc_ConsCons[of m s t] by fastforce
qed auto
lemma alternating_list_in_lists: "a\<in>A \<Longrightarrow> b\<in>A \<Longrightarrow> alternating_list n a b \<in> lists A"
by (induct n) auto
subsubsection \<open>Binary relation chains\<close>
text \<open>Here we consider lists where each pair of adjacent elements satisfy a given relation.\<close>
fun binrelchain :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool"
where "binrelchain P [] = True"
| "binrelchain P [x] = True"
| "binrelchain P (x # y # xs) = (P x y \<and> binrelchain P (y#xs))"
lemma binrelchain_Cons_reduce: "binrelchain P (x#xs) \<Longrightarrow> binrelchain P xs"
by (induct xs) auto
lemma binrelchain_append_reduce1: "binrelchain P (xs@ys) \<Longrightarrow> binrelchain P xs"
proof (induct xs rule: list_induct_CCons)
case (CCons x y xs) with binrelchain_Cons_reduce show ?case by fastforce
qed auto
lemma binrelchain_append_reduce2:
"binrelchain P (xs@ys) \<Longrightarrow> binrelchain P ys"
proof (induct xs)
case (Cons x xs) with binrelchain_Cons_reduce show ?case by fastforce
qed simp
lemma binrelchain_Conssnoc_reduce:
"binrelchain P (x#xs@[y]) \<Longrightarrow> binrelchain P xs"
using binrelchain_append_reduce1 binrelchain_Cons_reduce by fastforce
lemma binrelchain_overlap_join:
"binrelchain P (xs@[x]) \<Longrightarrow> binrelchain P (x#ys) \<Longrightarrow> binrelchain P (xs@x#ys)"
by (induct xs rule: list_induct_CCons) auto
lemma binrelchain_join:
"\<lbrakk> binrelchain P (xs@[x]); binrelchain P (y#ys); P x y \<rbrakk> \<Longrightarrow>
binrelchain P (xs @ x # y # ys)"
using binrelchain_overlap_join by fastforce
lemma binrelchain_snoc:
"binrelchain P (xs@[x]) \<Longrightarrow> P x y \<Longrightarrow> binrelchain P (xs@[x,y])"
using binrelchain_join by fastforce
lemma binrelchain_sym_rev:
assumes "\<And>x y. P x y \<Longrightarrow> P y x"
shows "binrelchain P xs \<Longrightarrow> binrelchain P (rev xs)"
proof (induct xs rule: list_induct_CCons)
case (CCons x y xs) with assms show ?case by (auto intro: binrelchain_snoc)
qed auto
lemma binrelchain_remdup_adj:
"binrelchain P (xs@[x,x]@ys) \<Longrightarrow> binrelchain P (xs@x#ys)"
by (induct xs rule: list_induct_CCons) auto
abbreviation "proper_binrelchain P xs \<equiv> binrelchain P xs \<and> distinct xs"
lemma binrelchain_obtain_proper:
"x\<noteq>y \<Longrightarrow> binrelchain P (x#xs@[y]) \<Longrightarrow>
\<exists>zs. set zs \<subseteq> set xs \<and> length zs \<le> length xs \<and> proper_binrelchain P (x#zs@[y])"
proof (induct xs arbitrary: x)
case (Cons w ws)
show ?case
proof (cases "w=x" "w=y" rule: two_cases)
case one
from one(1) Cons(3) have "binrelchain P (x#ws@[y])"
using binrelchain_Cons_reduce by simp
with Cons(1,2) obtain zs
where "set zs \<subseteq> set ws" "length zs \<le> length ws" "proper_binrelchain P (x#zs@[y])"
by auto
thus ?thesis by auto
next
case other
with Cons(3) have "proper_binrelchain P (x#[]@[y])"
using binrelchain_append_reduce1 by simp
moreover have "length [] \<le> length (w#ws)" "set [] \<subseteq> set (w#ws)" by auto
ultimately show ?thesis by blast
next
case neither
from Cons(3) have "binrelchain P (w#ws@[y])"
using binrelchain_Cons_reduce by simp
with neither(2) Cons(1) obtain zs
where zs: "set zs \<subseteq> set ws" "length zs \<le> length ws"
"proper_binrelchain P (w#zs@[y])"
by auto
show ?thesis
proof (cases "x\<in>set zs")
case True
from this obtain as bs where asbs: "zs = as@x#bs"
using in_set_conv_decomp[of x] by auto
with zs(3) have "proper_binrelchain P (x#bs@[y])"
using binrelchain_append_reduce2[of P "w#as"] by auto
moreover from zs(1) asbs have "set bs \<subseteq> set (w#ws)" by auto
moreover from asbs zs(2) have "length bs \<le> length (w#ws)" by simp
ultimately show ?thesis by auto
next
case False
with zs(3) neither(1) Cons(2,3) have "proper_binrelchain P (x#(w#zs)@[y])"
by simp
moreover from zs(1) have "set (w#zs) \<subseteq> set (w#ws)" by auto
moreover from zs(2) have "length (w#zs) \<le> length (w#ws)" by simp
ultimately show ?thesis by blast
qed
qed (fastforce simp add: Cons(2))
qed simp
lemma binrelchain_trans_Cons_snoc:
assumes "\<And>x y z. P x y \<Longrightarrow> P y z \<Longrightarrow> P x z"
shows "binrelchain P (x#xs@[y]) \<Longrightarrow> P x y"
proof (induct xs arbitrary: x)
case Cons with assms show ?case using binrelchain_Cons_reduce by auto
qed simp
lemma binrelchain_cong:
assumes "\<And>x y. P x y \<Longrightarrow> Q x y"
shows "binrelchain P xs \<Longrightarrow> binrelchain Q xs"
using assms binrelchain_Cons_reduce
by (induct xs rule: list_induct_CCons) auto
lemma binrelchain_funcong_Cons_snoc:
assumes "\<And>x y. P x y \<Longrightarrow> f y = f x" "binrelchain P (x#xs@[y])"
shows "f y = f x"
using assms binrelchain_cong[of P]
binrelchain_trans_Cons_snoc[of "\<lambda>x y. f y = f x" x xs y]
by auto
lemma binrelchain_funcong_extra_condition_Cons_snoc:
assumes "\<And>x y. Q x \<Longrightarrow> P x y \<Longrightarrow> Q y" "\<And>x y. Q x \<Longrightarrow> P x y \<Longrightarrow> f y = f x"
shows "Q x \<Longrightarrow> binrelchain P (x#zs@[y]) \<Longrightarrow> f y = f x"
proof (induct zs arbitrary: x)
case (Cons z zs) with assms show ?case
using binrelchain_Cons_reduce[of P x "z#zs@[y]"] by fastforce
qed (simp add: assms)
lemma binrelchain_setfuncong_Cons_snoc:
"\<lbrakk> \<forall>x\<in>A. \<forall>y. P x y \<longrightarrow> y\<in>A; \<forall>x\<in>A. \<forall>y. P x y \<longrightarrow> f y = f x; x\<in>A;
binrelchain P (x#zs@[y]) \<rbrakk> \<Longrightarrow> f y = f x"
using binrelchain_funcong_extra_condition_Cons_snoc[of "\<lambda>x. x\<in>A" P f x zs y]
by fast
lemma binrelchain_propcong_Cons_snoc:
assumes "\<And>x y. Q x \<Longrightarrow> P x y \<Longrightarrow> Q y"
shows "Q x \<Longrightarrow> binrelchain P (x#xs@[y]) \<Longrightarrow> Q y"
proof (induct xs arbitrary: x)
case Cons with assms show ?case using binrelchain_Cons_reduce by auto
qed (simp add: assms)
subsubsection \<open>Set of subseqs\<close>
lemma subseqs_Cons: "subseqs (x#xs) = map (Cons x) (subseqs xs) @ (subseqs xs)"
using cong_let[of "subseqs xs" "\<lambda>xss. map (Cons x) xss @ xss"] by simp
abbreviation "ssubseqs xs \<equiv> set (subseqs xs)"
lemma nil_ssubseqs: "[] \<in> ssubseqs xs"
proof (induct xs)
case (Cons x xs) thus ?case using subseqs_Cons[of x] by simp
qed simp
lemma ssubseqs_Cons: "ssubseqs (x#xs) = (Cons x) ` (ssubseqs xs) \<union> ssubseqs xs"
using subseqs_Cons[of x] by simp
lemma ssubseqs_refl: "xs \<in> ssubseqs xs"
proof (induct xs)
case (Cons x xs) thus ?case using ssubseqs_Cons by fast
qed (rule nil_ssubseqs)
lemma ssubseqs_subset: "as \<in> ssubseqs bs \<Longrightarrow> ssubseqs as \<subseteq> ssubseqs bs"
proof (induct bs arbitrary: as)
case (Cons b bs) show ?case
proof (cases "as \<in> set (subseqs bs)")
case True with Cons show ?thesis using ssubseqs_Cons by fastforce
next
case False with Cons show ?thesis
using nil_ssubseqs[of "b#bs"] ssubseqs_Cons[of "hd as"] ssubseqs_Cons[of b]
by (cases as) auto
qed
qed simp
lemma ssubseqs_lists:
"as \<in> lists A \<Longrightarrow> bs \<in> ssubseqs as \<Longrightarrow> bs \<in> lists A"
proof (induct as arbitrary: bs)
case (Cons a as) thus ?case using ssubseqs_Cons[of a] by fastforce
qed simp
lemma delete1_ssubseqs:
"as@bs \<in> ssubseqs (as@[a]@bs)"
proof (induct as)
case Nil show ?case using ssubseqs_refl ssubseqs_Cons[of a bs] by auto
next
case (Cons x xs) thus ?case using ssubseqs_Cons[of x] by simp
qed
lemma delete2_ssubseqs:
"as@bs@cs \<in> ssubseqs (as@[a]@bs@[b]@cs)"
using delete1_ssubseqs[of "as@[a]@bs"] delete1_ssubseqs ssubseqs_subset
by fastforce
subsection \<open>Orders and posets\<close>
text \<open>
We have chosen to work with the @{const ordering} locale instead of the @{class order} class to
more easily facilitate simultaneously working with both an order and its dual.
\<close>
subsubsection \<open>Dual ordering\<close>
context ordering
begin
abbreviation greater_eq :: "'a\<Rightarrow>'a\<Rightarrow>bool" (infix "\<succeq>" 50)
where "greater_eq a b \<equiv> less_eq b a"
abbreviation greater :: "'a\<Rightarrow>'a\<Rightarrow>bool" (infix "\<succ>" 50)
where "greater a b \<equiv> less b a"
lemma dual: "ordering greater_eq greater"
using strict_iff_order refl antisym trans by unfold_locales fastforce
end (* context ordering *)
subsubsection \<open>Morphisms of posets\<close>
locale OrderingSetMap =
domain : ordering less_eq less
+ codomain: ordering less_eq' less'
for less_eq :: "'a\<Rightarrow>'a\<Rightarrow>bool" (infix "\<^bold>\<le>" 50)
and less :: "'a\<Rightarrow>'a\<Rightarrow>bool" (infix "\<^bold><" 50)
and less_eq' :: "'b\<Rightarrow>'b\<Rightarrow>bool" (infix "\<^bold>\<le>*" 50)
and less' :: "'b\<Rightarrow>'b\<Rightarrow>bool" (infix "\<^bold><*" 50)
+ fixes P :: "'a set"
and f :: "'a\<Rightarrow>'b"
assumes ordsetmap: "a\<in>P \<Longrightarrow> b\<in>P \<Longrightarrow> a \<^bold>\<le> b \<Longrightarrow> f a \<^bold>\<le>* f b"
begin
lemma comp:
assumes "OrderingSetMap less_eq' less' less_eq'' less'' Q g" "f`P \<subseteq> Q"
shows "OrderingSetMap less_eq less less_eq'' less'' P (g\<circ>f)"
proof (
intro_locales, rule OrderingSetMap.axioms(2), rule assms(1), unfold_locales
)
from assms(2)
show "\<And>a b. a \<in> P \<Longrightarrow> b \<in> P \<Longrightarrow> a \<^bold>\<le> b \<Longrightarrow> less_eq'' ((g \<circ> f) a) ((g \<circ> f) b)"
using ordsetmap OrderingSetMap.ordsetmap[OF assms(1)]
by force
qed
lemma subset: "Q\<subseteq>P \<Longrightarrow> OrderingSetMap (\<^bold>\<le>) (\<^bold><) (\<^bold>\<le>*) (\<^bold><*) Q f"
using ordsetmap by unfold_locales fast
lemma dual:
"OrderingSetMap domain.greater_eq domain.greater
codomain.greater_eq codomain.greater P f"
proof (intro_locales, rule domain.dual, rule codomain.dual, unfold_locales)
from ordsetmap show "\<And>a b. a\<in>P \<Longrightarrow> b\<in>P \<Longrightarrow> a\<succeq>b \<Longrightarrow> f b \<^bold>\<le>* f a" by fast
qed
end (* context OrderingSetMap *)
locale OrderingSetIso = OrderingSetMap less_eq less less_eq' less' P f
for less_eq :: "'a\<Rightarrow>'a\<Rightarrow>bool" (infix "\<^bold>\<le>" 50)
and less :: "'a\<Rightarrow>'a\<Rightarrow>bool" (infix "\<^bold><" 50)
and less_eq' :: "'b\<Rightarrow>'b\<Rightarrow>bool" (infix "\<^bold>\<le>*" 50)
and less' :: "'b\<Rightarrow>'b\<Rightarrow>bool" (infix "\<^bold><*" 50)
and P :: "'a set"
and f :: "'a\<Rightarrow>'b"
+ assumes inj : "inj_on f P"
and rev_OrderingSetMap:
"OrderingSetMap less_eq' less' less_eq less (f`P) (the_inv_into P f)"
abbreviation "subset_ordering_iso \<equiv> OrderingSetIso (\<subseteq>) (\<subset>) (\<subseteq>) (\<subset>)"
lemma (in OrderingSetMap) isoI:
assumes "inj_on f P" "\<And>a b. a\<in>P \<Longrightarrow> b\<in>P \<Longrightarrow> f a \<^bold>\<le>* f b \<Longrightarrow> a \<^bold>\<le> b"
shows "OrderingSetIso less_eq less less_eq' less' P f"
using assms the_inv_into_f_f[OF assms(1)]
by unfold_locales auto
lemma OrderingSetIsoI_orders_greater2less:
fixes f :: "'a::order \<Rightarrow> 'b::order"
assumes "inj_on f P" "\<And>a b. a \<in> P \<Longrightarrow> b \<in> P \<Longrightarrow> (b\<le>a) = (f a \<le> f b)"
shows "OrderingSetIso (greater_eq::'a\<Rightarrow>'a\<Rightarrow>bool) (greater::'a\<Rightarrow>'a\<Rightarrow>bool)
(less_eq::'b\<Rightarrow>'b\<Rightarrow>bool) (less::'b\<Rightarrow>'b\<Rightarrow>bool) P f"
proof
from assms(2) show "\<And>a b. a \<in> P \<Longrightarrow> b \<in> P \<Longrightarrow> b\<le>a \<Longrightarrow> f a \<le> f b" by auto
from assms(2)
show "\<And>a b. a \<in> f ` P \<Longrightarrow> b \<in> f ` P \<Longrightarrow> b\<le>a \<Longrightarrow>
the_inv_into P f a \<le> the_inv_into P f b"
using the_inv_into_f_f[OF assms(1)]
by force
qed (rule assms(1))
context OrderingSetIso
begin
lemmas ordsetmap = ordsetmap
lemma ordsetmap_strict: "\<lbrakk> a\<in>P; b\<in>P; a\<^bold><b \<rbrakk> \<Longrightarrow> f a \<^bold><* f b"
using domain.strict_iff_order codomain.strict_iff_order ordsetmap inj
inj_on_contraD
by fastforce
lemmas inv_ordsetmap = OrderingSetMap.ordsetmap[OF rev_OrderingSetMap]
lemma rev_ordsetmap: "\<lbrakk> a\<in>P; b\<in>P; f a \<^bold>\<le>* f b \<rbrakk> \<Longrightarrow> a \<^bold>\<le> b"
using inv_ordsetmap the_inv_into_f_f[OF inj] by fastforce
lemma inv_iso: "OrderingSetIso less_eq' less' less_eq less (f`P) (the_inv_into P f)"
using inv_ordsetmap inj_on_the_inv_into[OF inj] the_inv_into_onto[OF inj]
ordsetmap the_inv_into_the_inv_into[OF inj]
by unfold_locales auto
lemmas inv_ordsetmap_strict = OrderingSetIso.ordsetmap_strict[OF inv_iso]
lemma rev_ordsetmap_strict: "\<lbrakk> a\<in>P; b\<in>P; f a \<^bold><* f b \<rbrakk> \<Longrightarrow> a \<^bold>< b"
using inv_ordsetmap_strict the_inv_into_f_f[OF inj] by fastforce
lemma iso_comp:
assumes "OrderingSetIso less_eq' less' less_eq'' less'' Q g" "f`P \<subseteq> Q"
shows "OrderingSetIso less_eq less less_eq'' less'' P (g\<circ>f)"
proof (rule OrderingSetMap.isoI)
from assms show "OrderingSetMap (\<^bold>\<le>) (\<^bold><) less_eq'' less'' P (g \<circ> f)"
using OrderingSetIso.axioms(1) comp by fast
from assms(2) show "inj_on (g \<circ> f) P"
using OrderingSetIso.inj[OF assms(1)]
comp_inj_on[OF inj, OF subset_inj_on]
by fast
next
fix a b
from assms(2) show "\<lbrakk> a\<in>P; b\<in>P; less_eq'' ((g\<circ>f) a) ((g\<circ>f) b) \<rbrakk> \<Longrightarrow> a\<^bold>\<le>b"
using OrderingSetIso.rev_ordsetmap[OF assms(1)] rev_ordsetmap by force
qed
lemma iso_subset:
"Q\<subseteq>P \<Longrightarrow> OrderingSetIso (\<^bold>\<le>) (\<^bold><) (\<^bold>\<le>*) (\<^bold><*) Q f"
using subset[of Q] subset_inj_on[OF inj] rev_ordsetmap
by (blast intro: OrderingSetMap.isoI)
lemma iso_dual:
"OrderingSetIso domain.greater_eq domain.greater
codomain.greater_eq codomain.greater P f"
by (
rule OrderingSetMap.isoI, rule OrderingSetMap.dual, unfold_locales,
rule inj, rule rev_ordsetmap
)
end (* context OrderingSetIso *)
lemma induced_pow_fun_subset_ordering_iso:
assumes "inj_on f A"
shows "subset_ordering_iso (Pow A) ((`) f)"
proof
show "\<And>a b. a \<in> Pow A \<Longrightarrow> b \<in> Pow A \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f ` a \<subseteq> f ` b" by fast
from assms show 2:"inj_on ((`) f) (Pow A)"
using induced_pow_fun_inj_on by fast
show "\<And>a b. a \<in> (`) f ` Pow A \<Longrightarrow> b \<in> (`) f ` Pow A \<Longrightarrow> a \<subseteq> b
\<Longrightarrow> the_inv_into (Pow A) ((`) f) a \<subseteq> the_inv_into (Pow A) ((`) f) b"
proof-
fix Y1 Y2
assume Y: "Y1 \<in> ((`) f) ` Pow A" "Y2 \<in> ((`) f) ` Pow A" "Y1 \<subseteq> Y2"
from Y(1,2) obtain X1 X2 where "X1\<subseteq>A" "X2\<subseteq>A" "Y1 = f`X1" "Y2 = f`X2"
by auto
with assms Y(3)
show "the_inv_into (Pow A) ((`) f) Y1 \<subseteq> the_inv_into (Pow A) ((`) f) Y2"
using inj_onD[OF assms] the_inv_into_f_f[OF 2, of X1]
the_inv_into_f_f[OF 2, of X2]
by blast
qed
qed
subsubsection \<open>More @{const arg_min}\<close>
lemma is_arg_minI:
"\<lbrakk> P x; \<And>y. P y \<Longrightarrow> \<not> m y < m x \<rbrakk> \<Longrightarrow> is_arg_min m P x"
by (simp add: is_arg_min_def)
lemma is_arg_min_linorderI:
"\<lbrakk> P x; \<And>y. P y \<Longrightarrow> m x \<le> (m y::_::linorder) \<rbrakk> \<Longrightarrow> is_arg_min m P x"
by (simp add: is_arg_min_linorder)
lemma is_arg_min_eq:
"\<lbrakk> is_arg_min m P x; P z; m z = m x \<rbrakk> \<Longrightarrow> is_arg_min m P z"
by (metis is_arg_min_def)
lemma is_arg_minD1: "is_arg_min m P x \<Longrightarrow> P x"
unfolding is_arg_min_def by fast
lemma is_arg_minD2: "is_arg_min m P x \<Longrightarrow> P y \<Longrightarrow> \<not> m y < m x"
unfolding is_arg_min_def by fast
lemma is_arg_min_size: fixes m :: "'a \<Rightarrow> 'b::linorder"
shows "is_arg_min m P x \<Longrightarrow> m x = m (arg_min m P)"
by (metis arg_min_equality is_arg_min_linorder)
lemma is_arg_min_size_subprop:
fixes m :: "'a\<Rightarrow>'b::linorder"
assumes "is_arg_min m P x" "Q x" "\<And>y. Q y \<Longrightarrow> P y"
shows "m (arg_min m Q) = m (arg_min m P)"
proof-
have "\<not> is_arg_min m Q x \<Longrightarrow> \<not> is_arg_min m P x"
proof
assume x: "\<not> is_arg_min m Q x"
from assms(2,3) show False
using contrapos_nn[OF x, OF is_arg_minI] is_arg_minD2[OF assms(1)] by auto
qed
with assms(1) show ?thesis
using is_arg_min_size[of m] is_arg_min_size[of m] by fastforce
qed
subsubsection \<open>Bottom of a set\<close>
context ordering
begin
definition has_bottom :: "'a set \<Rightarrow> bool"
where "has_bottom P \<equiv> \<exists>z\<in>P. \<forall>x\<in>P. z \<^bold>\<le> x"
lemma has_bottomI: "z\<in>P \<Longrightarrow> (\<And>x. x\<in>P \<Longrightarrow> z \<^bold>\<le> x) \<Longrightarrow> has_bottom P"
using has_bottom_def by auto
lemma has_uniq_bottom: "has_bottom P \<Longrightarrow> \<exists>!z\<in>P. \<forall>x\<in>P. z\<^bold>\<le>x"
using has_bottom_def antisym by force
definition bottom :: "'a set \<Rightarrow> 'a"
where "bottom P \<equiv> (THE z. z\<in>P \<and> (\<forall>x\<in>P. z\<^bold>\<le>x))"
lemma bottomD:
assumes "has_bottom P"
shows "bottom P \<in> P" "x\<in>P \<Longrightarrow> bottom P \<^bold>\<le> x"
using assms has_uniq_bottom theI'[of "\<lambda>z. z\<in>P \<and> (\<forall>x\<in>P. z\<^bold>\<le>x)"]
unfolding bottom_def
by auto
lemma bottomI: "z\<in>P \<Longrightarrow> (\<And>y. y\<in>P \<Longrightarrow> z \<^bold>\<le> y) \<Longrightarrow> z = bottom P"
using has_bottomI has_uniq_bottom
the1_equality[THEN sym, of "\<lambda>z. z\<in>P \<and> (\<forall>x\<in>P. z\<^bold>\<le>x)"]
unfolding bottom_def
by simp
end (* context ordering *)
lemma has_bottom_pow: "order.has_bottom (Pow A)"
by (fast intro: order.has_bottomI)
lemma bottom_pow: "order.bottom (Pow A) = {}"
proof (rule order.bottomI[THEN sym]) qed auto
context OrderingSetMap
begin
abbreviation "dombot \<equiv> domain.bottom P"
abbreviation "codbot \<equiv> codomain.bottom (f`P)"
lemma im_has_bottom: "domain.has_bottom P \<Longrightarrow> codomain.has_bottom (f`P)"
using domain.bottomD ordsetmap by (fast intro: codomain.has_bottomI)
lemma im_bottom: "domain.has_bottom P \<Longrightarrow> f dombot = codbot"
using domain.bottomD ordsetmap by (auto intro: codomain.bottomI)
end (* context OrderingSetMap *)
lemma (in OrderingSetIso) pullback_has_bottom:
assumes "codomain.has_bottom (f`P)"
shows "domain.has_bottom P"
proof (rule domain.has_bottomI)
from assms show "the_inv_into P f codbot \<in> P"
using codomain.bottomD(1) the_inv_into_into[OF inj] by fast
from assms show "\<And>x. x\<in>P \<Longrightarrow> the_inv_into P f codbot \<^bold>\<le> x"
using codomain.bottomD inv_ordsetmap[of codbot] the_inv_into_f_f[OF inj]
by fastforce
qed
lemma (in OrderingSetIso) pullback_bottom:
"\<lbrakk> domain.has_bottom P; x\<in>P; f x = codomain.bottom (f`P) \<rbrakk> \<Longrightarrow>
x = domain.bottom P"
using im_has_bottom codomain.bottomD(2) rev_ordsetmap
by (auto intro: domain.bottomI)
subsubsection \<open>Minimal and pseudominimal elements in sets\<close>
text \<open>
We will call an element of a poset pseudominimal if the only element below it is the bottom of
the poset.
\<close>
context ordering
begin
definition minimal_in :: "'a set \<Rightarrow> 'a \<Rightarrow> bool"
where "minimal_in P x \<equiv> x\<in>P \<and> (\<forall>z\<in>P. \<not> z \<^bold>< x)"
definition pseudominimal_in :: "'a set \<Rightarrow> 'a \<Rightarrow> bool"
where "pseudominimal_in P x \<equiv> minimal_in (P - {bottom P}) x"
\<comment> \<open>only makes sense for @{term "has_bottom P"}\<close>
lemma minimal_inD1: "minimal_in P x \<Longrightarrow> x\<in>P"
using minimal_in_def by fast
lemma minimal_inD2: "minimal_in P x \<Longrightarrow> z\<in>P \<Longrightarrow> \<not> z \<^bold>< x"
using minimal_in_def by fast
lemma pseudominimal_inD1: "pseudominimal_in P x \<Longrightarrow> x\<in>P"
using pseudominimal_in_def minimal_inD1 by fast
lemma pseudominimal_inD2:
"pseudominimal_in P x \<Longrightarrow> z\<in>P \<Longrightarrow> z\<^bold><x \<Longrightarrow> z = bottom P"
using pseudominimal_in_def minimal_inD2 by fast
lemma pseudominimal_inI:
assumes "x\<in>P" "x \<noteq> bottom P" "\<And>z. z\<in>P \<Longrightarrow> z\<^bold><x \<Longrightarrow> z = bottom P"
shows "pseudominimal_in P x"
using assms
unfolding pseudominimal_in_def minimal_in_def
by fast
lemma pseudominimal_ne_bottom: "pseudominimal_in P x \<Longrightarrow> x \<noteq> bottom P"
using pseudominimal_in_def minimal_inD1 by fast
lemma pseudominimal_comp:
"\<lbrakk> pseudominimal_in P x; pseudominimal_in P y; x\<^bold>\<le>y \<rbrakk> \<Longrightarrow> x = y"
using pseudominimal_inD1 pseudominimal_inD2 pseudominimal_ne_bottom
strict_iff_order[of x y]
by force
end (* context ordering *)
lemma pseudominimal_in_pow:
assumes "order.pseudominimal_in (Pow A) x"
shows "\<exists>a\<in>A. x = {a}"
proof-
from assms obtain a where "{a} \<subseteq> x"
using order.pseudominimal_ne_bottom bottom_pow[of A] by fast
with assms show ?thesis
using order.pseudominimal_inD1 order.pseudominimal_inD2[of _ x "{a}"]
bottom_pow
by fast
qed
lemma pseudominimal_in_pow_singleton:
"a\<in>A \<Longrightarrow> order.pseudominimal_in (Pow A) {a}"
using singleton_pow bottom_pow by (fast intro: order.pseudominimal_inI)
lemma no_pseudominimal_in_pow_is_empty:
"(\<And>x. \<not> order.pseudominimal_in (Pow A) {x}) \<Longrightarrow> A = {}"
using pseudominimal_in_pow_singleton by (fast intro: equals0I)
lemma (in OrderingSetIso) pseudominimal_map:
"domain.has_bottom P \<Longrightarrow> domain.pseudominimal_in P x \<Longrightarrow>
codomain.pseudominimal_in (f`P) (f x)"
using domain.pseudominimal_inD1 pullback_bottom
domain.pseudominimal_ne_bottom rev_ordsetmap_strict
domain.pseudominimal_inD2 im_bottom
by (blast intro: codomain.pseudominimal_inI)
lemma (in OrderingSetIso) pullback_pseudominimal_in:
"\<lbrakk> domain.has_bottom P; x\<in>P; codomain.pseudominimal_in (f`P) (f x) \<rbrakk> \<Longrightarrow>
domain.pseudominimal_in P x"
using im_bottom codomain.pseudominimal_ne_bottom ordsetmap_strict
codomain.pseudominimal_inD2 pullback_bottom
by (blast intro: domain.pseudominimal_inI)
subsubsection \<open>Set of elements below another\<close>
abbreviation (in ordering) below_in :: "'a set \<Rightarrow> 'a \<Rightarrow> 'a set" (infix ".\<^bold>\<le>" 70)
where "P.\<^bold>\<le>x \<equiv> {y\<in>P. y\<^bold>\<le>x}"
abbreviation (in ord) below_in :: "'a set \<Rightarrow> 'a \<Rightarrow> 'a set" (infix ".\<le>" 70)
where "P.\<le>x \<equiv> {y\<in>P. y\<le>x}"
context ordering
begin
lemma below_in_refl: "x\<in>P \<Longrightarrow> x \<in> P.\<^bold>\<le>x"
using refl by fast
lemma below_in_singleton: "x\<in>P \<Longrightarrow> P.\<^bold>\<le>x \<subseteq> {y} \<Longrightarrow> y = x"
using below_in_refl by fast
lemma bottom_in_below_in: "has_bottom P \<Longrightarrow> x\<in>P \<Longrightarrow> bottom P \<in> P.\<^bold>\<le>x"
using bottomD by fast
lemma below_in_singleton_is_bottom:
"\<lbrakk> has_bottom P; x\<in>P; P.\<^bold>\<le>x = {x} \<rbrakk> \<Longrightarrow> x = bottom P"
using bottom_in_below_in by fast
lemma bottom_below_in:
"has_bottom P \<Longrightarrow> x\<in>P \<Longrightarrow> bottom (P.\<^bold>\<le>x) = bottom P"
using bottom_in_below_in by (fast intro: bottomI[THEN sym])
lemma bottom_below_in_relative:
"\<lbrakk> has_bottom (P.\<^bold>\<le>y); x\<in>P; x\<^bold>\<le>y \<rbrakk> \<Longrightarrow> bottom (P.\<^bold>\<le>x) = bottom (P.\<^bold>\<le>y)"
using bottomD trans by (blast intro: bottomI[THEN sym])
lemma has_bottom_pseudominimal_in_below_inI:
assumes "has_bottom P" "x\<in>P" "pseudominimal_in P y" "y\<^bold>\<le>x"
shows "pseudominimal_in (P.\<^bold>\<le>x) y"
using assms(3,4) pseudominimal_inD1[OF assms(3)]
pseudominimal_inD2[OF assms(3)]
bottom_below_in[OF assms(1,2)] pseudominimal_ne_bottom
by (force intro: pseudominimal_inI)
lemma has_bottom_pseudominimal_in_below_in:
assumes "has_bottom P" "x\<in>P" "pseudominimal_in (P.\<^bold>\<le>x) y"
shows "pseudominimal_in P y"
using pseudominimal_inD1[OF assms(3)]
pseudominimal_inD2[OF assms(3)]
pseudominimal_ne_bottom[OF assms(3)]
bottom_below_in[OF assms(1,2)]
strict_implies_order[of _ y] trans[of _ y x]
by (force intro: pseudominimal_inI)
lemma pseudominimal_in_below_in:
assumes "has_bottom (P.\<^bold>\<le>y)" "x\<in>P" "x\<^bold>\<le>y" "pseudominimal_in (P.\<^bold>\<le>x) w"
shows "pseudominimal_in (P.\<^bold>\<le>y) w"
using assms(3) trans[of w x y] trans[of _ w x] strict_iff_order
pseudominimal_inD1[OF assms(4)]
pseudominimal_inD2[OF assms(4)]
pseudominimal_ne_bottom[OF assms(4)]
bottom_below_in_relative[OF assms(1-3)]
by (force intro: pseudominimal_inI)
lemma collect_pseudominimals_below_in_less_eq_top:
assumes "OrderingSetIso less_eq less (\<subseteq>) (\<subset>) (P.\<^bold>\<le>x) f"
"f`(P.\<^bold>\<le>x) = Pow A" "a \<subseteq> {y. pseudominimal_in (P.\<^bold>\<le>x) y}"
defines "w \<equiv> the_inv_into (P.\<^bold>\<le>x) f (\<Union>(f`a))"
shows "w \<^bold>\<le> x"
proof-
from assms(2,3) have "(\<Union>(f`a)) \<in> f`(P.\<^bold>\<le>x)"
using pseudominimal_inD1 by fastforce
with assms(4) show ?thesis
using OrderingSetIso.inj[OF assms(1)] the_inv_into_into[of f "P.\<^bold>\<le>x"] by force
qed
lemma collect_pseudominimals_below_in_poset:
assumes "OrderingSetIso less_eq less (\<subseteq>) (\<subset>) (P.\<^bold>\<le>x) f"
"f`(P.\<^bold>\<le>x) = Pow A"
"a \<subseteq> {y. pseudominimal_in (P.\<^bold>\<le>x) y}"
defines "w \<equiv> the_inv_into (P.\<^bold>\<le>x) f (\<Union>(f`a))"
shows "w \<in> P"
using assms(2-4) OrderingSetIso.inj[OF assms(1)] pseudominimal_inD1
the_inv_into_into[of f "P.\<^bold>\<le>x" "\<Union>(f`a)"]
by force
lemma collect_pseudominimals_below_in_eq:
assumes "x\<in>P" "OrderingSetIso less_eq less (\<subseteq>) (\<subset>) (P.\<^bold>\<le>x) f"
"f`(P.\<^bold>\<le>x) = Pow A" "a \<subseteq> {y. pseudominimal_in (P.\<^bold>\<le>x) y}"
defines w: "w \<equiv> the_inv_into (P.\<^bold>\<le>x) f (\<Union>(f`a))"
shows "a = {y. pseudominimal_in (P.\<^bold>\<le>w) y}"
proof
from assms(3) have has_bot_ltx: "has_bottom (P.\<^bold>\<le>x)"
using has_bottom_pow OrderingSetIso.pullback_has_bottom[OF assms(2)]
by auto
from assms(3,4) have Un_fa: "(\<Union>(f`a)) \<in> f`(P.\<^bold>\<le>x)"
using pseudominimal_inD1 by fastforce
from assms have w_le_x: "w\<^bold>\<le>x" and w_P: "w\<in>P"
using collect_pseudominimals_below_in_less_eq_top
collect_pseudominimals_below_in_poset
by auto
show "a \<subseteq> {y. pseudominimal_in (P.\<^bold>\<le>w) y}"
proof
fix y assume y: "y \<in> a"
show "y \<in> {y. pseudominimal_in (P.\<^bold>\<le>w) y}"
proof (rule CollectI, rule pseudominimal_inI, rule CollectI, rule conjI)
from y assms(4) have y_le_x: "y \<in> P.\<^bold>\<le>x" using pseudominimal_inD1 by fast
thus "y\<in>P" by simp
from y w show "y \<^bold>\<le> w"
using y_le_x Un_fa OrderingSetIso.inv_ordsetmap[OF assms(2)]
the_inv_into_f_f[OF OrderingSetIso.inj, OF assms(2), of y]
by fastforce
from assms(1) y assms(4) show "y \<noteq> bottom (P.\<^bold>\<le>w)"
using w_P w_le_x has_bot_ltx bottom_below_in_relative
pseudominimal_ne_bottom
by fast
next
fix z assume z: "z \<in> P.\<^bold>\<le>w" "z\<^bold><y"
with y assms(4) have "z = bottom (P.\<^bold>\<le>x)"
using w_le_x trans pseudominimal_inD2[ of "P.\<^bold>\<le>x" y z] by fast
moreover from assms(1) have "bottom (P.\<^bold>\<le>w) = bottom (P.\<^bold>\<le>x)"
using has_bot_ltx w_P w_le_x bottom_below_in_relative by fast
ultimately show "z = bottom (P.\<^bold>\<le>w)" by simp
qed
qed
show "a \<supseteq> {y. pseudominimal_in (P.\<^bold>\<le>w) y}"
proof
fix v assume "v \<in> {y. pseudominimal_in (P.\<^bold>\<le>w) y}"
hence "pseudominimal_in (P.\<^bold>\<le>w) v" by fast
moreover hence v_pm_ltx: "pseudominimal_in (P.\<^bold>\<le>x) v"
using has_bot_ltx w_P w_le_x pseudominimal_in_below_in by fast
ultimately
have "f v \<le> (\<Union>(f`a))"
using w pseudominimal_inD1[of _ v] pseudominimal_inD1[of _ v] w_le_x w_P
OrderingSetIso.ordsetmap[OF assms(2), of v w] Un_fa
OrderingSetIso.inj[OF assms(2)]
f_the_inv_into_f
by force
with assms(3) obtain y where "y\<in>a" "f v \<subseteq> f y"
using v_pm_ltx has_bot_ltx pseudominimal_in_pow
OrderingSetIso.pseudominimal_map[OF assms(2)]
by force
with assms(2,4) show "v \<in> a"
using v_pm_ltx pseudominimal_inD1 pseudominimal_comp[of _ v y]
OrderingSetIso.rev_ordsetmap[OF assms(2), of v y]
by fast
qed
qed
end (* context ordering *)
subsubsection \<open>Lower bounds\<close>
context ordering
begin
definition lbound_of :: "'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
where "lbound_of x y b \<equiv> b\<^bold>\<le>x \<and> b\<^bold>\<le>y"
lemma lbound_ofI: "b\<^bold>\<le>x \<Longrightarrow> b\<^bold>\<le>y \<Longrightarrow> lbound_of x y b"
using lbound_of_def by fast
lemma lbound_ofD1: "lbound_of x y b \<Longrightarrow> b\<^bold>\<le>x"
using lbound_of_def by fast
lemma lbound_ofD2: "lbound_of x y b \<Longrightarrow> b\<^bold>\<le>y"
using lbound_of_def by fast
definition glbound_in_of :: "'a set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
where "glbound_in_of P x y b \<equiv>
b\<in>P \<and> lbound_of x y b \<and> (\<forall>a\<in>P. lbound_of x y a \<longrightarrow> a\<^bold>\<le>b)"
lemma glbound_in_ofI:
"\<lbrakk> b\<in>P; lbound_of x y b; \<And>a. a\<in>P \<Longrightarrow> lbound_of x y a \<Longrightarrow> a\<^bold>\<le>b \<rbrakk> \<Longrightarrow>
glbound_in_of P x y b"
using glbound_in_of_def by auto
lemma glbound_in_ofD_in: "glbound_in_of P x y b \<Longrightarrow> b\<in>P"
using glbound_in_of_def by fast
lemma glbound_in_ofD_lbound: "glbound_in_of P x y b \<Longrightarrow> lbound_of x y b"
using glbound_in_of_def by fast
lemma glbound_in_ofD_glbound:
"glbound_in_of P x y b \<Longrightarrow> a\<in>P \<Longrightarrow> lbound_of x y a \<Longrightarrow> a\<^bold>\<le>b"
using glbound_in_of_def by fast
lemma glbound_in_of_less_eq1: "glbound_in_of P x y b \<Longrightarrow> b\<^bold>\<le>x"
using glbound_in_ofD_lbound lbound_ofD1 by fast
lemma glbound_in_of_less_eq2: "glbound_in_of P x y b \<Longrightarrow> b\<^bold>\<le>y"
using glbound_in_ofD_lbound lbound_ofD2 by fast
lemma pseudominimal_in_below_in_less_eq_glbound:
assumes "pseudominimal_in (P.\<^bold>\<le>x) w" "pseudominimal_in (P.\<^bold>\<le>y) w"
"glbound_in_of P x y b"
shows "w \<^bold>\<le> b"
using assms lbound_ofI glbound_in_ofD_glbound
pseudominimal_inD1[of "P.\<^bold>\<le>x"] pseudominimal_inD1[of "P.\<^bold>\<le>y"]
by fast
end (* context ordering *)
subsubsection \<open>Simplex-like posets\<close>
text \<open>Define a poset to be simplex-like if it is isomorphic to the power set of some set.\<close>
context ordering
begin
definition simplex_like :: "'a set \<Rightarrow> bool"
where "simplex_like P \<equiv> finite P \<and>
(\<exists>f A::nat set.
OrderingSetIso less_eq less (\<subseteq>) (\<subset>) P f \<and> f`P = Pow A
)"
lemma simplex_likeI:
assumes "finite P" "OrderingSetIso less_eq less (\<subseteq>) (\<subset>) P f"
"f`P = Pow (A::nat set)"
shows "simplex_like P"
using assms simplex_like_def by auto
lemma simplex_likeD_finite: "simplex_like P \<Longrightarrow> finite P"
using simplex_like_def by simp
lemma simplex_likeD_iso:
"simplex_like P \<Longrightarrow>
\<exists>f A::nat set. OrderingSetIso less_eq less (\<subseteq>) (\<subset>) P f \<and> f`P = Pow A"
using simplex_like_def by simp
lemma simplex_like_has_bottom: "simplex_like P \<Longrightarrow> has_bottom P"
using simplex_likeD_iso has_bottom_pow OrderingSetIso.pullback_has_bottom
by fastforce
lemma simplex_like_no_pseudominimal_imp_singleton:
assumes "simplex_like P" "\<And>x. \<not> pseudominimal_in P x"
shows "\<exists>p. P = {p}"
proof-
obtain f and A::"nat set"
where fA: "OrderingSetIso less_eq less (\<subseteq>) (\<subset>) P f" "f`P = Pow A"
using simplex_likeD_iso[OF assms(1)]
by auto
define e where e: "e \<equiv> {}:: nat set"
with fA(2) have "e \<in> f`P" using Pow_bottom by simp
from this obtain p where "p \<in> P" "f p = e" by fast
have "\<And>x. \<not> order.pseudominimal_in (Pow A) {x}"
proof
fix x::nat assume "order.pseudominimal_in (Pow A) {x}"
moreover with fA(2) have "{x} \<in> f`P"
using order.pseudominimal_inD1 by fastforce
ultimately show False
using assms fA simplex_like_has_bottom
OrderingSetIso.pullback_pseudominimal_in
by fastforce
qed
with e fA(2) show ?thesis
using no_pseudominimal_in_pow_is_empty
inj_on_to_singleton[OF OrderingSetIso.inj, OF fA(1)]
by force
qed
lemma simplex_like_no_pseudominimal_in_below_in_imp_singleton:
"\<lbrakk> x\<in>P; simplex_like (P.\<^bold>\<le>x); \<And>z. \<not> pseudominimal_in (P.\<^bold>\<le>x) z \<rbrakk> \<Longrightarrow>
P.\<^bold>\<le>x = {x}"
using simplex_like_no_pseudominimal_imp_singleton below_in_singleton[of x P]
by fast
lemma pseudo_simplex_like_has_bottom:
"OrderingSetIso less_eq less (\<subseteq>) (\<subset>) P f \<Longrightarrow> f`P = Pow A \<Longrightarrow>
has_bottom P"
using has_bottom_pow OrderingSetIso.pullback_has_bottom by fastforce
lemma pseudo_simplex_like_above_pseudominimal_is_top:
assumes "OrderingSetIso less_eq less (\<subseteq>) (\<subset>) P f" "f`P = Pow A" "t\<in>P"
"\<And>x. pseudominimal_in P x \<Longrightarrow> x \<^bold>\<le> t"
shows "f t = A"
proof
from assms(2,3) show "f t \<subseteq> A" by fast
show "A \<subseteq> f t"
proof
fix a assume "a\<in>A"
moreover with assms(2) have "{a} \<in> f`P" by simp
ultimately show "a \<in> f t"
using assms pseudominimal_in_pow_singleton[of a A]
pseudo_simplex_like_has_bottom[of P f]
OrderingSetIso.pullback_pseudominimal_in[OF assms(1)]
OrderingSetIso.ordsetmap[OF assms(1), of _ t]
by force
qed
qed
lemma pseudo_simplex_like_below_in_above_pseudominimal_is_top:
assumes "x\<in>P" "OrderingSetIso less_eq less (\<subseteq>) (\<subset>) (P.\<^bold>\<le>x) f"
"f`(P.\<^bold>\<le>x) = Pow A" "t \<in> P.\<^bold>\<le>x"
"\<And>y. pseudominimal_in (P.\<^bold>\<le>x) y \<Longrightarrow> y \<^bold>\<le> t"
shows "t = x"
using assms(1,3-5)
pseudo_simplex_like_above_pseudominimal_is_top[OF assms(2)]
below_in_refl[of x P] OrderingSetIso.ordsetmap[OF assms(2), of t x]
inj_onD[OF OrderingSetIso.inj[OF assms(2)], of t x]
by auto
lemma simplex_like_below_in_above_pseudominimal_is_top:
assumes "x\<in>P" "simplex_like (P.\<^bold>\<le>x)" "t \<in> P.\<^bold>\<le>x"
"\<And>y. pseudominimal_in (P.\<^bold>\<le>x) y \<Longrightarrow> y \<^bold>\<le> t"
shows "t = x"
using assms simplex_likeD_iso
pseudo_simplex_like_below_in_above_pseudominimal_is_top[of x P _ _ t]
by blast
end (* context ordering *)
lemma (in OrderingSetIso) simplex_like_map:
assumes "domain.simplex_like P"
shows "codomain.simplex_like (f`P)"
proof-
obtain g::"'a \<Rightarrow> nat set" and A::"nat set"
where gA: "OrderingSetIso (\<^bold>\<le>) (\<^bold><) (\<subseteq>) (\<subset>) P g" "g`P = Pow A"
using domain.simplex_likeD_iso[OF assms]
by auto
from gA(1) inj
have "OrderingSetIso (\<^bold>\<le>*) (\<^bold><*) (\<subseteq>) (\<subset>) (f`P)
(g \<circ> (the_inv_into P f))"
using OrderingSetIso.iso_comp[OF inv_iso] the_inv_into_onto
by fast
moreover from gA(2) inj have "(g \<circ> (the_inv_into P f)) ` (f`P) = Pow A"
using the_inv_into_onto by (auto simp add: image_comp[THEN sym])
moreover from assms have "finite (f`P)"
using domain.simplex_likeD_finite by fast
ultimately show ?thesis by (auto intro: codomain.simplex_likeI)
qed
lemma (in OrderingSetIso) pullback_simplex_like:
assumes "finite P" "codomain.simplex_like (f`P)"
shows "domain.simplex_like P"
proof-
obtain g::"'b \<Rightarrow> nat set" and A::"nat set"
where gA: "OrderingSetIso (\<^bold>\<le>*) (\<^bold><*) (\<subseteq>) (\<subset>) (f`P) g"
"g`(f`P) = Pow A"
using codomain.simplex_likeD_iso[OF assms(2)]
by auto
from assms(1) gA(2) show ?thesis
using iso_comp[OF gA(1)]
by (auto intro: domain.simplex_likeI simp add: image_comp)
qed
lemma simplex_like_pow:
assumes "finite A"
shows "order.simplex_like (Pow A)"
proof-
from assms obtain f::"'a\<Rightarrow>nat" where "inj_on f A"
using finite_imp_inj_to_nat_seg[of A] by auto
hence "subset_ordering_iso (Pow A) ((`) f)"
using induced_pow_fun_subset_ordering_iso by fast
with assms show ?thesis using induced_pow_fun_surj
by (blast intro: order.simplex_likeI)
qed
subsubsection \<open>The superset ordering\<close>
abbreviation "supset_has_bottom \<equiv> ordering.has_bottom (\<supseteq>)"
abbreviation "supset_bottom \<equiv> ordering.bottom (\<supseteq>)"
abbreviation "supset_lbound_of \<equiv> ordering.lbound_of (\<supseteq>)"
abbreviation "supset_glbound_in_of \<equiv> ordering.glbound_in_of (\<supseteq>)"
abbreviation "supset_simplex_like \<equiv> ordering.simplex_like (\<supseteq>) (\<supset>)"
abbreviation "supset_pseudominimal_in \<equiv>
ordering.pseudominimal_in (\<supseteq>) (\<supset>)"
abbreviation supset_below_in :: "'a set set \<Rightarrow> 'a set \<Rightarrow> 'a set set" (infix ".\<supseteq>" 70)
where "P.\<supseteq>A \<equiv> ordering.below_in (\<supseteq>) P A"
lemma supset_poset: "ordering (\<supseteq>) (\<supset>)" ..
lemmas supset_bottomI = ordering.bottomI [OF supset_poset]
lemmas supset_pseudominimal_inI = ordering.pseudominimal_inI [OF supset_poset]
lemmas supset_pseudominimal_inD1 = ordering.pseudominimal_inD1 [OF supset_poset]
lemmas supset_pseudominimal_inD2 = ordering.pseudominimal_inD2 [OF supset_poset]
lemmas supset_lbound_ofI = ordering.lbound_ofI [OF supset_poset]
lemmas supset_lbound_of_def = ordering.lbound_of_def [OF supset_poset]
lemmas supset_glbound_in_ofI = ordering.glbound_in_ofI [OF supset_poset]
lemmas supset_pseudominimal_ne_bottom =
ordering.pseudominimal_ne_bottom[OF supset_poset]
lemmas supset_has_bottom_pseudominimal_in_below_inI =
ordering.has_bottom_pseudominimal_in_below_inI[OF supset_poset]
lemmas supset_has_bottom_pseudominimal_in_below_in =
ordering.has_bottom_pseudominimal_in_below_in[OF supset_poset]
lemma OrderingSetIso_pow_complement:
"OrderingSetIso (\<supseteq>) (\<supset>) (\<subseteq>) (\<subset>) (Pow A) ((-) A)"
using inj_on_minus_set by (fast intro: OrderingSetIsoI_orders_greater2less)
lemma simplex_like_pow_above_in:
assumes "finite A" "X\<subseteq>A"
shows "supset_simplex_like ((Pow A).\<supseteq>X)"
proof (
rule OrderingSetIso.pullback_simplex_like, rule OrderingSetIso.iso_subset,
rule OrderingSetIso_pow_complement
)
from assms(1) show "finite ((Pow A).\<supseteq>X)" by simp
from assms(1) have "finite (Pow (A-X))" by fast
moreover from assms(2) have "((-) A) ` ((Pow A).\<supseteq>X) = Pow (A-X)"
by auto
ultimately
show "ordering.simplex_like (\<subseteq>) (\<subset>) ( ((-) A) ` ((Pow A).\<supseteq>X))"
using simplex_like_pow
by fastforce
qed fast
end (* theory *)
diff --git a/thys/BytecodeLogicJmlTypes/Cachera.thy b/thys/BytecodeLogicJmlTypes/Cachera.thy
--- a/thys/BytecodeLogicJmlTypes/Cachera.thy
+++ b/thys/BytecodeLogicJmlTypes/Cachera.thy
@@ -1,1415 +1,1417 @@
(*File: Cachera.thy
Author: L Beringer & M Hofmann, LMU Munich
Date: 05/12/2008
Purpose: Strong interpretation of, and derived proof system for,
heap analysis a la Cachera/Jensen/Pichardie/Schneider,
using invariants.
*)
(*<*)
theory Cachera imports Logic begin
(*>*)
section\<open>A derived logic for a strong type system\<close>
text\<open>In this section we consider a system of derived assertions, for
a type system for bounded heap consumption. The type system arises by
reformulating the analysis of Cachera, Jensen, Pichardie, and
Schneider \cite{CaJePiSc05MemoryUsage} for a high-level functional
language. The original approach of Cachera et al.~consists of
formalising the correctness proof of a certain analysis technique in
Coq. Consequently, the verification of a program requires the
execution of the analysis algorithm inside the theorem prover, which
involves the computation of the (method) call graph and fixed point
iterations. In contrast, our approach follows the proof-carrying code
paradigm more closely: the analysis amounts to a type inference which
is left unformalised and can thus be carried out outside the trusted
code base. Only the result of the analysis is communicated to the code
recipient. The recipient verifies the validity of the certificate by
a largely syntax-directed single-pass traversal of the (low-level)
code using a domain-specific program logic. This approach to
proof-carrying code was already explored in the MRG project, with
respect to program logics of partial correctness
\cite{BeringerHofmannMomiglianoShkaravska:LPAR2004} and a type system
for memory consumption by Hofmann and
Jost~\cite{HofmannJost:POPL2003}. In order to obtain
syntax-directedness of the proof rules, these had to be formulated at
the granularity of typing judgements. In contrast, the present proof
system admits proof rules for individual JVM instructions.
Having derived proof rules for individual JVM instructions, we
introduce a type system for a small functional language, and a
compilation into bytecode. The type system associates a natural
number $n$ to an expression $e$, in a typing context
$\Sigma$. Informally, the interpretation of a typing judgement
$\Sigma \rhd e:n$ is that the evaluation of $e$ (which may include
the invocation of functions whose resource behaviour is specified in
$\Sigma$) does not perform more than $n$ allocations. The type system
is then formally proven sound, using the derived logic for bytecode.
By virtue of the invariants, the guarantee given by the present system
is stronger than the one given by our encoding of the Hofmann-Jost
system, as even non-terminating programs can be verified in a
meaningful way.\<close>
subsection\<open>Syntax and semantics of judgements\<close>
text\<open>The formal interpretation at JVM level of a type \<open>n\<close> is
given by a triple $$\mathit{Cachera}(n) = (A, B, I)$$ consisting of a
(trivial) precondition, a post-condition, and a strong invariant.\<close>
definition Cachera::"nat \<Rightarrow> (Assn \<times> Post \<times> Inv)" where
"Cachera n = (\<lambda> s0 s . True,
\<lambda> s0 (ops,s,h) (k,v) . |k| \<le> |h| + n,
\<lambda> s0 (ops,s,h) k. |k| \<le> |h| + n)"
text\<open>This definition is motivated by the expectation that $\rhd
\lbrace A \rbrace\; \ulcorner\, e \urcorner\, \lbrace B \rbrace\, I$
should be derivable whenever the type judgement $\Sigma \rhd e:n$
holds, where $ \ulcorner e \urcorner$ is the translation of compiling
the expression $e$ into JVML, and the specification table \<open>MST\<close>
contains the interpretations of the entries in $\Sigma$.\<close>
text\<open>We abbreviate the above construction of judgements by a
predicate \<open>deriv\<close>.\<close>
definition deriv::"CTXT \<Rightarrow> Class \<Rightarrow> Method \<Rightarrow> Label \<Rightarrow>
(Assn \<times> Post \<times> Inv) \<Rightarrow> bool" where
"deriv G C m l (ABI) = (let (A,B,I) = ABI in (G \<rhd> \<lbrace> A \<rbrace> C,m,l \<lbrace> B \<rbrace> I))"
text\<open>Thus, the intended interpretation of a typing judgement $\Sigma
\rhd e: n$ is $$\mathit{deriv}\; C\; m\; l\; (\mathit{Cachera}\; n)$$
if $e$ translates to a code block whose first instruction is at
$C.m.l$.\<close>
text\<open>We also define a judgement of the auxiliary form of
sequents.\<close>
definition derivAssum::"CTXT \<Rightarrow> Class \<Rightarrow> Method \<Rightarrow> Label \<Rightarrow>
(Assn \<times> Post \<times> Inv) \<Rightarrow> bool" where
"derivAssum G C m l (ABI) = (let (A,B,I) = ABI in G \<rhd> \<langle> A \<rangle> C,m,l \<langle> B \<rangle> I)"
text\<open>The following operation converts a derived judgement into the
syntactical form of method specifications.\<close>
definition mkSPEC::"(Assn \<times> Post \<times> Inv) \<Rightarrow> ANNO \<Rightarrow>
(MethSpec \<times> MethInv \<times> ANNO)" where
"mkSPEC (ABI) Anno = (let (A,B,I) = ABI in
(\<lambda> s0 t . B s0 (mkState s0) t, \<lambda> s0 h . I s0 (mkState s0) h, Anno))"
text\<open>This enables the interpretation of typing contexts $\Sigma$ as a
set of constraints on the specification table \<open>MST\<close>.\<close>
subsection\<open>Derived proof rules\<close>
(*<*)
declare Let_def[simp]
(*>*)
text\<open>We are now ready to prove derived rules, i.e.~proof rules where
assumptions as well as conclusions are of the restricted assertion
form. While their justification unfolds the definition of the
predicate \<open>deriv\<close>, their application will not. We first give
syntax-directed proof rules for all JVM instructions:\<close>
lemma CACH_NEW:
"\<lbrakk> ins_is C m l (new c); MST\<down>(C,m)=Some(Mspec,Minv,Anno);
Anno\<down>(l) = None; n = k + 1; derivAssum G C m (l+1) (Cachera k) \<rbrakk>
\<Longrightarrow> deriv G C m l (Cachera n)"
(*<*)
apply (simp add: ins_is_def Cachera_def deriv_def derivAssum_def, clarsimp)
apply (rule INSTR) apply assumption+ apply simp apply (simp add: heap_def)
apply fast
apply (erule CONSEQ)
apply (simp add: SP_pre_def)
apply clarsimp apply (simp add: SP_post_def) apply clarsimp
apply (drule NewElim1, fastforce) apply clarsimp
apply (subgoal_tac "ba\<down>(nextLoc ba) = None")
apply (simp add: AL_Size_UpdateSuc)
apply (rule nextLoc_fresh)
apply clarsimp
apply (simp add: SP_inv_def)
apply clarsimp
apply (drule NewElim1, fastforce) apply clarsimp
apply (subgoal_tac "ba\<down>(nextLoc ba) = None")
apply (simp add: AL_Size_UpdateSuc)
apply (rule nextLoc_fresh)
done
(*>*)
lemma CACH_INSTR:
"\<lbrakk> ins_is C m l I;
I \<in> { const c, dup, pop, swap, load x, store x, binop f,
unop g, getfield d F, putfield d F, checkcast d};
MST\<down>(C,m)=Some(Mspec,Minv,Anno); Anno\<down>(l) = None;
derivAssum G C m (l+1) (Cachera n) \<rbrakk>
\<Longrightarrow> deriv G C m l (Cachera n)"
(*<*)
apply (simp add: ins_is_def Cachera_def deriv_def derivAssum_def, clarsimp)
apply (rule INSTR) apply assumption+ apply simp apply (simp add: heap_def)
apply fast
apply (erule CONSEQ)
apply (simp add: SP_pre_def)
apply (simp add: SP_post_def) apply clarsimp
apply safe
apply (drule ConstElim1, fastforce) apply clarsimp
apply (drule DupElim1, fastforce) apply clarsimp
apply (drule PopElim1, fastforce) apply clarsimp
apply (drule SwapElim1, fastforce) apply clarsimp
apply (drule LoadElim1, fastforce) apply clarsimp
apply (drule StoreElim1, fastforce) apply clarsimp
apply (drule BinopElim1, fastforce) apply clarsimp
apply (drule UnopElim1, fastforce) apply clarsimp
apply (drule GetElim1, fastforce) apply clarsimp
apply (drule PutElim1, fastforce) apply clarsimp
apply (simp add: updSize)
apply (drule CastElim1, fastforce) apply clarsimp
apply (simp_all add: SP_inv_def)
apply safe
apply (drule ConstElim1, fastforce) apply clarsimp
apply (drule DupElim1, fastforce) apply clarsimp
apply (drule PopElim1, fastforce) apply clarsimp
apply (drule SwapElim1, fastforce) apply clarsimp
apply (drule LoadElim1, fastforce) apply clarsimp
apply (drule StoreElim1, fastforce) apply clarsimp
apply (drule BinopElim1, fastforce) apply clarsimp
apply (drule UnopElim1, fastforce) apply clarsimp
apply (drule GetElim1, fastforce) apply clarsimp
apply (drule PutElim1, fastforce) apply clarsimp
apply (simp add: updSize)
apply (drule CastElim1, fastforce) apply clarsimp
done
(*>*)
lemma CACH_RET:
"\<lbrakk> ins_is C m l vreturn; MST\<down>(C,m)=Some(Mspec,Minv,Anno);
Anno\<down>(l) = None \<rbrakk>
\<Longrightarrow> deriv G C m l (Cachera 0)"
(*<*)
apply (simp add: ins_is_def Cachera_def deriv_def derivAssum_def, clarsimp)
apply (rule VRET) apply assumption+ apply simp apply (simp add: heap_def)
apply clarsimp
done
(*>*)
lemma CACH_GOTO:
"\<lbrakk> ins_is C m l (goto pc); MST\<down>(C,m)=Some(Mspec,Minv,Anno);
Anno\<down>(l) = None; derivAssum G C m pc (Cachera n) \<rbrakk>
\<Longrightarrow> deriv G C m l (Cachera n)"
(*<*)
apply (simp add: ins_is_def Cachera_def deriv_def derivAssum_def, clarsimp)
apply (rule GOTO) apply assumption+ apply simp apply (simp add: heap_def)
apply (erule CONSEQ)
apply (simp add: SP_pre_def)
apply (simp add: SP_post_def) apply clarsimp apply (drule GotoElim1, fastforce)
apply clarsimp
apply (simp add: SP_inv_def) apply clarsimp apply (drule GotoElim1, fastforce)
apply clarsimp
done
(*>*)
lemma CACH_IF:
"\<lbrakk> ins_is C m l (iftrue pc); MST\<down>(C,m)=Some(Mspec,Minv,Anno);
Anno\<down>(l) = None; derivAssum G C m pc (Cachera n);
derivAssum G C m (l+1) (Cachera n) \<rbrakk>
\<Longrightarrow> deriv G C m l (Cachera n)"
(*<*)
apply (simp add: ins_is_def Cachera_def deriv_def derivAssum_def, clarsimp)
apply (rule IF) apply assumption+ apply (simp, simp add: heap_def)
apply (erule CONSEQ)
apply (simp add: SP_pre_def)
apply (simp add: SP_post_def) apply clarsimp apply (drule IfElim1, fastforce) apply clarsimp
apply (simp add: SP_inv_def) apply clarsimp apply (drule IfElim1, fastforce) apply clarsimp
apply clarsimp
apply (erule CONSEQ)
apply (simp add: SP_pre_def)
apply (simp add: SP_post_def) apply clarsimp apply (drule IfElim1, fastforce) apply clarsimp
apply (simp add: SP_inv_def) apply clarsimp apply (drule IfElim1, fastforce) apply clarsimp
done
(*>*)
lemma CACH_INVS:
"\<lbrakk> ins_is C m l (invokeS D m'); mbody_is D m' (par,code,l0);
MST\<down>(C,m)=Some(Mspec,Minv,Anno); Anno\<down>(l) = None;
MST\<down>(D, m') = Some(mkSPEC (Cachera k) Anno2);
nk = n+k; derivAssum G C m (l+1) (Cachera n)\<rbrakk>
\<Longrightarrow> deriv G C m l (Cachera nk)"
(*<*)
apply (simp add: ins_is_def Cachera_def deriv_def derivAssum_def, clarsimp)
apply (rule INVS) apply assumption+
apply (simp add: mkSPEC_def) apply fastforce apply simp apply (simp add: heap_def)
apply (simp add: mkState_def)
apply (erule CONSEQ)
apply clarsimp
apply clarsimp apply (simp add: SINV_post_def mkState_def)
apply clarsimp apply (simp add: SINV_inv_def mkState_def)
done
(*>*)
text\<open>In addition, we have two rules for subtyping\<close>
lemma CACH_SUB:
"\<lbrakk> deriv G C m l (Cachera n); n \<le> k\<rbrakk> \<Longrightarrow> deriv G C m l (Cachera k)"
(*<*)
apply (simp add: deriv_def derivAssum_def Cachera_def)
apply (rule CONSEQ, assumption+)
apply simp
apply simp
apply simp
done
(*>*)
lemma CACHAssum_SUB:
"\<lbrakk> derivAssum G C m l (Cachera n); n \<le> k\<rbrakk>
\<Longrightarrow> derivAssum G C m l (Cachera k)"
(*<*)
apply (simp add: derivAssum_def Cachera_def)
apply (rule CONSEQ, assumption+)
apply simp
apply simp
apply simp
done
(*>*)
text\<open>and specialised forms of the axiom rule and the injection rule.\<close>
lemma CACH_AX:
"\<lbrakk> G\<down>(C,m,l) = Some (Cachera n); MST\<down>(C,m)=Some(Mspec,Minv,Anno);
Anno\<down>(l) = None\<rbrakk>
\<Longrightarrow> derivAssum G C m l (Cachera n)"
(*<*)
apply (simp add: derivAssum_def Cachera_def)
apply (drule AX) apply assumption apply simp apply (simp add: heap_def)
apply (rule CONSEQ) apply assumption+ apply simp apply simp apply simp
done
(*>*)
lemma CACH_INJECT:
"deriv G C m l (Cachera n) \<Longrightarrow> derivAssum G C m l (Cachera n)"
(*<*)
apply (simp add: deriv_def derivAssum_def Cachera_def)
apply (erule INJECT)
done
(*>*)
text\<open>Finally, a verified-program rule relates specifications to
judgements for the method bodies. Thus, even the method specifications
may be given as derived assertions (modulo the \<open>mkSPEC\<close>-conversion).\<close>
lemma CACH_VP:
"\<lbrakk> \<forall> c m par code l0. mbody_is c m (par, code, l0) \<longrightarrow>
(\<exists> n Anno . MST\<down>(c,m) = Some(mkSPEC(Cachera n) Anno) \<and>
deriv G c m l0 (Cachera n));
\<forall> c m l A B I. G\<down>(c,m,l) = Some(A,B,I) \<longrightarrow>
(\<exists> n . (A,B,I) = Cachera n \<and> deriv G c m l (Cachera n))\<rbrakk>
\<Longrightarrow> VP"
(*<*)
apply (simp add: VP_def) apply (rule_tac x=G in exI, simp add: VP_G_def)
apply safe
(*G*)
apply (erule thin_rl)
apply (erule_tac x=C in allE, erule_tac x=m in allE, erule_tac x=l in allE, clarsimp)
apply (simp add: deriv_def)
(*MST*)
apply (rotate_tac 1, erule thin_rl)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply(erule_tac x=par in allE, erule_tac x=code in allE, erule_tac x=l0 in allE, clarsimp)
apply (simp add: Cachera_def mkSPEC_def deriv_def, clarsimp)
apply (rule CONSEQ) apply assumption+
apply simp
apply (simp add: mkPost_def mkState_def)
apply (simp add: mkState_def mkInv_def)
done
(*>*)
subsection\<open>Soundness of high-level type system\<close>
text\<open>We define a first-order functional language where expressions
are stratified into primitive expressions and general expressions. The
language supports the construction of lists using constructors
$\mathit{NilPrim}$ and $\mathit{ConsPrim}\; h\; t$, and includes a
corresponding pattern match operation. In order to simplify the
compilation, function identifiers are taken to be pairs of class names
and method names.\<close>
type_synonym Fun = "Class \<times> Method"
datatype Prim =
IntPrim int
| UnPrim "Val \<Rightarrow> Val" Var
| BinPrim "Val \<Rightarrow> Val \<Rightarrow> Val" Var Var
| NilPrim
| ConsPrim Var Var
| CallPrim Fun "Var list"
datatype Expr =
PrimE Prim
| LetE Var Prim Expr
| CondE Var Expr Expr
| MatchE Var Expr Var Var Expr
type_synonym FunProg = "(Fun,Var list \<times> Expr) AssList"
text\<open>The type system uses contexts that associate a type (natural
number) to function identifiers.\<close>
type_synonym TP_Sig = "(Fun, nat) AssList"
text\<open>We first give the rules for primitive expressions.\<close>
inductive_set TP_prim::"(TP_Sig \<times> Prim \<times> nat)set"
where
TP_int: "(\<Sigma>,IntPrim i,0) : TP_prim"
|
TP_un: "(\<Sigma>,UnPrim f x,0) : TP_prim"
|
TP_bin: "(\<Sigma>,BinPrim f x y,0) : TP_prim"
|
TP_nil: "(\<Sigma>,NilPrim,0) : TP_prim"
|
TP_cons: "(\<Sigma>,ConsPrim x y,1) : TP_prim"
|
TP_Call: "\<lbrakk>\<Sigma>\<down>f = Some n\<rbrakk> \<Longrightarrow> (\<Sigma>,CallPrim f args,n) : TP_prim"
text\<open>Next, the rules for general expressions.\<close>
inductive_set TP_expr::"(TP_Sig \<times> Expr \<times> nat) set"
where
TP_sub: "\<lbrakk>(\<Sigma>,e,m):TP_expr; m \<le> n\<rbrakk> \<Longrightarrow> (\<Sigma>,e,n):TP_expr"
|
TP_prim:"\<lbrakk>(\<Sigma>,p,n):TP_prim\<rbrakk> \<Longrightarrow> (\<Sigma>,PrimE p,n) : TP_expr"
|
TP_let: "\<lbrakk>(\<Sigma>,p,k):TP_prim; (\<Sigma>,e,m):TP_expr; n = k+m\<rbrakk>
\<Longrightarrow> (\<Sigma>,LetE x p e,n) : TP_expr"
|
TP_Cond:"\<lbrakk>(\<Sigma>,e1,n):TP_expr; (\<Sigma>,e2,n):TP_expr\<rbrakk>
\<Longrightarrow>(\<Sigma>,CondE x e1 e2,n) : TP_expr"
|
TP_Match:"\<lbrakk>(\<Sigma>,e1,n):TP_expr; (\<Sigma>,e2,n):TP_expr \<rbrakk>
\<Longrightarrow> (\<Sigma>,MatchE x e1 h t e2,n):TP_expr"
text\<open>A functional program is well-typed if its domain agrees with
that of some context such that each function body validates the
context entry.\<close>
definition TP::"TP_Sig \<Rightarrow> FunProg \<Rightarrow> bool" where
"TP \<Sigma> F = ((\<forall> f . (\<Sigma>\<down>f = None) = (F\<down>f = None)) \<and>
(\<forall> f n par e . \<Sigma>\<down>f = Some n \<longrightarrow> F\<down>f = Some (par,e) \<longrightarrow> (\<Sigma>,e,n):TP_expr))"
text\<open>For the translation into bytecode, we introduce identifiers for
a class of lists, the expected field names, and a temporary (reserved)
variable name.\<close>
axiomatization
LIST::Class and
HD::Field and
TL::Field and
tmp::Var
text\<open>The compilation of primitive expressions extends a code block by
a sequence of JVM instructions that leave a value on the top of the
operand stack.\<close>
inductive_set compilePrim::
"(Label \<times> (Label,Instr) AssList \<times> Prim \<times> ((Label,Instr) AssList \<times> Label)) set"
where
compileInt: "(l, code, IntPrim i, (code[l\<mapsto>(const (IVal i))],l+1)) : compilePrim"
|
compileUn:
"(l, code, UnPrim f x, (code[l\<mapsto>(load x)][(l+1)\<mapsto>(unop f)],l+2)) : compilePrim"
|
compileBin:
"(l, code, BinPrim f x y,
(code[l\<mapsto>(load x)][(l+1)\<mapsto>(load y)][(l+2)\<mapsto>(binop f)],l+3)) : compilePrim"
|
compileNil:
"(l, code, NilPrim, (code[l\<mapsto>(const (RVal Nullref))],l+1)) : compilePrim"
|
compileCons:
"(l, code, ConsPrim x y,
(code[l\<mapsto>(load y)][(l+1)\<mapsto>(load x)]
[(l+2)\<mapsto>(new LIST)][(l+3)\<mapsto>store tmp]
[(l+4)\<mapsto>load tmp][(l+5)\<mapsto>(putfield LIST HD)]
[(l+6)\<mapsto>load tmp][(l+7)\<mapsto>(putfield LIST TL)]
[(l+8)\<mapsto>(load tmp)], l+9)) : compilePrim"
|
compileCall_Nil:
"(l, code, CallPrim f [],(code[l\<mapsto>invokeS (fst f) (snd f)],l+1)): compilePrim"
|
compileCall_Cons:
"\<lbrakk> (l+1,code[l\<mapsto>load x], CallPrim f args, OUT) : compilePrim\<rbrakk>
\<Longrightarrow> (l, code, CallPrim f (x#args), OUT): compilePrim"
text\<open>The following lemma shows that the resulting code is an
extension of the code submitted as an argument, and that the
new instructions define a contiguous block.\<close>
lemma compilePrim_Prop1[rule_format]:
"(l, code, p, OUT) : compilePrim \<Longrightarrow>
(\<forall> code1l1 . OUT = (code1, l1) \<longrightarrow>
(l < l1 \<and> (\<forall> ll . ll < l \<longrightarrow> code1\<down>ll = code\<down>ll) \<and>
(\<forall> ll . l \<le> ll \<longrightarrow> ll < l1 \<longrightarrow> (\<exists> ins . code1\<down>ll = Some ins))))"
(*<*)
apply (erule compilePrim.induct)
apply clarsimp
apply rule
apply clarsimp apply (rule AL_update2) apply simp
apply (rule, rule AL_update1)
apply clarsimp
apply rule
apply clarsimp
apply (rule AL_update5)
apply (rule AL_update5)
apply (simp, simp, simp)
apply clarsimp
apply (case_tac "ll=l", clarsimp, rule)
apply (rule AL_update5)
apply (simp add: AL_update1)
apply simp
apply (case_tac "ll=l+1", clarsimp, rule)
apply (simp add: AL_update1)
apply simp
(*BIN*)
apply clarsimp
apply rule
apply clarsimp
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (simp, simp, simp, simp)
apply clarsimp
apply (case_tac "ll=l", clarsimp, rule)
apply (rule AL_update5)
apply (rule AL_update5)
apply (simp add: AL_update1)
apply (simp, simp)
apply (case_tac "ll=l+1", clarsimp, rule)
apply (rule AL_update5)
apply (simp add: AL_update1)
apply simp
apply (case_tac "ll=l+2", clarsimp, rule)
apply (simp add: AL_update1)
apply simp
apply clarsimp
apply rule
apply clarsimp apply (rule AL_update2) apply simp
apply (rule, rule AL_update1) apply simp
apply clarsimp
apply rule apply clarsimp
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (simp, simp, simp, simp)
apply (simp, simp, simp, simp)
apply (simp, simp)
apply clarsimp
apply (case_tac "ll=l", clarsimp, rule)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (simp add: AL_update1)
apply (simp, simp, simp,simp)
apply (simp, simp, simp,simp)
apply (case_tac "ll=l+1", clarsimp, rule)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (simp add: AL_update1)
apply (simp, simp, simp,simp)
apply (simp, simp, simp)
apply (case_tac "ll=l+2", clarsimp, rule)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (simp add: AL_update1)
apply (simp, simp, simp,simp)
apply (simp, simp)
apply (case_tac "ll=l+3", clarsimp, rule)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update1)
apply (simp, simp, simp,simp)
apply simp
apply (case_tac "ll=l+4", clarsimp, rule)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update1)
apply (simp, simp, simp,simp)
apply (case_tac "ll=l+5", clarsimp, rule)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update1)
apply (simp, simp, simp)
apply (case_tac "ll=l+6", clarsimp, rule)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update1)
apply (simp, simp)
apply (case_tac "ll=l+7", clarsimp, rule)
apply (rule AL_update5)
apply (rule AL_update1)
apply (simp)
apply (case_tac "ll=l+8", clarsimp, rule)
apply (rule AL_update1)
apply simp
(*CALL-NIL*)
apply clarsimp
apply (rule, clarsimp)
apply (rule AL_update5) apply (simp ,simp)
apply (rule, rule AL_update1)
(*CALL-CONS*)
apply clarsimp
apply (rule, clarsimp)
apply (rule AL_update5) apply simp apply simp
apply clarsimp
apply (erule_tac x=ll in allE)+ apply clarsimp
apply (case_tac "l=ll", clarsimp) apply (rule, rule AL_update1)
apply clarsimp
done
(*>*)
text\<open>A signature corresponds to a method specification table if all
context entries are represented as \<open>MST\<close> entries and method
names that are defined in the global program \<open>P\<close>.\<close>
definition Sig_good::"TP_Sig \<Rightarrow> bool" where
"Sig_good \<Sigma> =
(\<forall> C m n. \<Sigma>\<down>(C,m) = Some n \<longrightarrow>
(MST\<down>(C, m) = Some (mkSPEC (Cachera n) emp) \<and>
(\<exists> par code l0 . mbody_is C m (par,code,l0))))"
text\<open>This definition requires \<open>MST\<close> to associate the
specification $$\mathit{mkSPEC}\; (\mathit{Cachera}\; n)\;
\mathit{emp}$$ to each method to which the type signature associates
the type $n$. In particular, this requires the annotation table of
such a method to be empty. Additionally, the global program $P$ is
required to contain a method definition for each method
(i.e.~function) name occurring in the domain of the signature.\<close>
text\<open>An auxiliary abbreviation that captures when a block of code has
trivial annotations and only comprises defined program labels.\<close>
definition Segment::
"Class \<Rightarrow> Method \<Rightarrow> Label \<Rightarrow> Label \<Rightarrow> (Label,Instr)AssList \<Rightarrow> bool"
where
"Segment C m l l1 code =
(\<exists> Mspec Minv Anno . MST\<down>(C,m) = Some(Mspec,Minv,Anno) \<and>
(\<forall>ll. l \<le> ll \<longrightarrow> ll < l1 \<longrightarrow>
Anno\<down>(ll) = None \<and> (\<exists>ins. ins_is C m ll ins \<and> code\<down>ll = Some ins)))"
(*<*)
lemma Segment_triv:
"\<lbrakk>Segment C m l l1 code; MST\<down>(C,m) = Some(Mspec,Minv,Anno); l \<le> ll; ll < l1\<rbrakk>
\<Longrightarrow> (Anno\<down>(ll) = None \<and> (\<exists>ins. ins_is C m ll ins \<and> code\<down>ll = Some ins))"
by (simp add: Segment_def)
lemma Segment_triv1:
"\<lbrakk>Segment C m l l1 code; MST\<down>(C,m) = Some(Mspec,Minv,Anno); l \<le> ll; ll < l1\<rbrakk> \<Longrightarrow> Anno\<down>(ll) = None"
by (simp add: Segment_def)
lemma Segment_triv2:
"\<lbrakk>Segment C m l l1 code; l \<le> ll; ll < l1\<rbrakk> \<Longrightarrow> (\<exists>ins. ins_is C m ll ins \<and> code\<down>ll = Some ins)"
apply (simp add: Segment_def) apply clarsimp done
lemma Segment_A:
"\<lbrakk>Segment C m l l1 code; l \<le> ll; ll < l1\<rbrakk> \<Longrightarrow> Segment C m ll l1 code" by (simp add: Segment_def, clarsimp)
(*>*)
text\<open>The soundness of (the translation of) a function call is proven
by induction on the list of arguments.\<close>
lemma Call_SoundAux[rule_format]:
"\<Sigma>\<down>f = Some n \<longrightarrow>
MST\<down>(fst f,snd f) = Some(mkSPEC (Cachera n) Anno2) \<longrightarrow>
(\<exists> par body l0 . mbody_is (fst f) (snd f) (par,body,l0)) \<longrightarrow>
(\<forall>l code code1 l1 G C m T MI k.
(l, code, CallPrim f args, code1, l1) \<in> compilePrim \<longrightarrow>
MST\<down>(C, m) = Some (T, MI,Anno) \<longrightarrow> Segment C m l l1 code1 \<longrightarrow>
derivAssum G C m l1 (Cachera k) \<longrightarrow>
deriv G C m l (Cachera (n+k)))"
(*<*)
+supply [[simproc del: defined_all]]
apply (induct args)
apply clarsimp
apply (erule compilePrim.cases) apply (simp,simp,simp, simp, simp, clarsimp)
apply (drule Segment_triv) apply assumption apply (subgoal_tac "la \<le> la", assumption, simp) apply clarsimp
apply (erule conjE)+ apply (erule exE)+ apply (erule conjE)+
apply (rule CACH_INVS)
apply (simp add: AL_update1) apply clarsimp apply fast
apply assumption
apply assumption
apply assumption
apply assumption
apply simp
apply simp
apply simp
(*CONS*)
apply clarsimp
apply (erule compilePrim.cases) apply (simp, simp, simp, simp, simp, simp) apply clarsimp
apply (erule impE) apply fast
apply (erule_tac x="la+1" in allE, rotate_tac -1)
apply (erule_tac x="codea[la\<mapsto>load x]" in allE, rotate_tac -1)
apply (erule_tac x=ab in allE, rotate_tac -1)
apply (erule_tac x=ba in allE, clarsimp)
apply (erule_tac x=G in allE, rotate_tac -1)
apply (erule_tac x=C in allE, rotate_tac -1)
apply (erule_tac x=m in allE, clarsimp)
apply (frule compilePrim_Prop1) apply fastforce apply clarsimp
apply (erule impE, erule Segment_A) apply simp apply simp
apply (drule Segment_triv) apply assumption apply (subgoal_tac "la \<le> la", assumption, simp) apply clarsimp
apply (erule conjE)+ apply (erule exE)+ apply (erule conjE)+
apply (rule CACH_INSTR) apply assumption
apply (rotate_tac -5, erule_tac x=la in allE, clarsimp) apply (simp add: AL_update1)
apply clarsimp apply fast
apply assumption
apply assumption
apply (rule CACH_INJECT)
apply simp
done
(*>*)
lemma Call_Sound:
"\<lbrakk> Sig_good \<Sigma>; \<Sigma>\<down>f = Some n;
(l, code, CallPrim f args, code1, l1) \<in> compilePrim;
MST\<down>(C,m) = Some (T, MI,Anno); Segment C m l l1 code1;
derivAssum G C m l1 (Cachera nn); k = n+nn\<rbrakk>
\<Longrightarrow> deriv G C m l (Cachera k)"
(*<*)
apply (case_tac f, clarsimp)
apply (rule Call_SoundAux)
apply assumption apply (simp add: Sig_good_def) apply (simp add: Sig_good_def)
apply assumption apply assumption apply assumption apply assumption
done
(*>*)
text\<open>The definition of basic instructions.\<close>
definition basic::"Instr \<Rightarrow> bool" where
"basic ins = ((\<exists> c . ins = const c) \<or> ins = dup \<or>
ins= pop \<or> ins= swap \<or> (\<exists> x. ins= load x) \<or>
(\<exists> y. ins= store y) \<or> (\<exists> f. ins= binop f) \<or>
(\<exists> g. ins= unop g) \<or> (\<exists> c1 F1. ins= getfield c1 F1) \<or>
(\<exists> c2 F2. ins= putfield c2 F2) \<or>
(\<exists> c3. ins= checkcast c3))"
text\<open>Next, we prove the soundness of basic instructions. The
hypothesis refers to instructions located at the program
continuation.\<close>
lemma Basic_Sound:
"\<lbrakk> Segment C m l ll code; MST\<down>(C,m) = Some (T, MI,Anno); l\<le>l1; l1 < ll;
l2=l1+1; code\<down>l1 = Some ins; basic ins; derivAssum G C m l2 (Cachera n)\<rbrakk>
\<Longrightarrow> deriv G C m l1 (Cachera n)"
(*<*)
apply (drule Segment_triv) apply assumption+ apply clarsimp apply (simp add: basic_def)
apply (erule disjE, clarsimp) apply (erule CACH_INSTR) apply fast apply assumption apply assumption apply simp
apply (erule disjE, clarsimp) apply (erule CACH_INSTR) apply fast apply assumption apply assumption apply simp
apply (erule disjE, clarsimp) apply (erule CACH_INSTR) apply fast apply assumption apply assumption apply simp
apply (erule disjE, clarsimp) apply (erule CACH_INSTR) apply fast apply assumption apply assumption apply simp
apply (erule disjE, clarsimp) apply (erule CACH_INSTR) apply fast apply assumption apply assumption apply simp
apply (erule disjE, clarsimp) apply (erule CACH_INSTR) apply fast apply assumption apply assumption apply simp
apply (erule disjE, clarsimp) apply (erule CACH_INSTR) apply fast apply assumption apply assumption apply simp
apply (erule disjE, clarsimp) apply (erule CACH_INSTR) apply fast apply assumption apply assumption apply simp
apply (erule disjE, clarsimp) apply (erule CACH_INSTR) apply fast apply assumption apply assumption apply simp
apply (erule disjE, clarsimp) apply (erule CACH_INSTR) apply fast apply assumption apply assumption apply simp
apply clarsimp apply (erule CACH_INSTR) apply fast apply assumption apply assumption apply simp
done
(*>*)
text\<open>Following this, the soundness of the type system for primitive
expressions. The proof proceeds by induction on the typing
judgement.\<close>
lemma TP_prim_Sound[rule_format]:
"(\<Sigma>,p,n):TP_prim \<Longrightarrow>
Sig_good \<Sigma> \<longrightarrow>
(\<forall> l code code1 l1 G C m T MI Anno nn k.
(l, code, p, (code1,l1)) : compilePrim \<longrightarrow>
MST\<down>(C,m) = Some (T,MI,Anno) \<longrightarrow>
Segment C m l l1 code1 \<longrightarrow> derivAssum G C m l1 (Cachera nn) \<longrightarrow>
k = n+nn \<longrightarrow> deriv G C m l (Cachera k))"
(*<*)
apply (erule TP_prim.induct)
(*INT*)
apply clarsimp apply (erule thin_rl)
apply (erule compilePrim.cases, simp_all, clarsimp)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp,simp) apply simp
apply (simp add: AL_update1)
apply (simp add: basic_def)
apply assumption
(*UN*)
apply clarsimp apply (erule thin_rl)
apply (erule compilePrim.cases, simp_all, clarsimp)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp)
apply simp
apply (rule AL_update5) apply (simp add: AL_update1) apply simp
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp)
apply simp
apply (simp add: AL_update1)
apply (simp add: basic_def)
apply (subgoal_tac "la+2=2+la", clarsimp,simp)
(*BIN*)
apply clarsimp apply (erule thin_rl)
apply (erule compilePrim.cases, simp_all, clarsimp)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp) apply simp
apply (rule AL_update5) apply (rule AL_update5) apply (simp add: AL_update1) apply (simp, simp)
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp)
apply simp
apply (rule AL_update5) apply (simp add: AL_update1) apply simp
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp)
apply simp
apply (rule AL_update1a) apply simp
apply (simp add: basic_def)
apply (subgoal_tac "la+3=3+la", clarsimp, simp)
(*Nil*)
apply clarsimp apply (erule thin_rl)
apply (erule compilePrim.cases, simp_all, clarsimp)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp,simp) apply simp
apply (simp add: AL_update1)
apply (simp add: basic_def)
apply assumption
(*CONS*)
apply clarsimp apply (erule thin_rl)
apply (erule compilePrim.cases, simp_all, clarsimp)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp) apply simp
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (simp add: AL_update1) apply (simp,simp,simp,simp) apply (simp,simp,simp,simp)
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp) apply simp
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (simp add: AL_update1) apply (simp,simp,simp,simp) apply (simp,simp,simp)
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (frule Segment_triv) apply assumption apply (subgoal_tac "la \<le> la+2", assumption, simp)
apply simp
apply (subgoal_tac "la+2=2+la", clarsimp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp) apply (simp add: AL_update1) apply clarsimp
apply (rule CACH_NEW) apply assumption+
apply (simp, simp)
apply (rule CACH_INJECT)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp) apply simp
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update1a) apply simp apply (simp,simp,simp,simp) apply simp
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp) apply simp
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update1a) apply simp apply (simp,simp,simp,simp)
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp) apply simp
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update1a) apply simp apply (simp,simp,simp)
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp) apply simp
apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update1a) apply simp apply (simp,simp)
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp) apply simp
apply (rule AL_update5)
apply (rule AL_update1a) apply simp apply simp
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound) apply assumption apply assumption apply (simp, simp) apply simp
apply (rule AL_update1a) apply simp
apply (simp add: basic_def)
apply (subgoal_tac "la+9=9+la", clarsimp, simp)
apply simp
(*Call*)
apply clarsimp
apply (rule Call_Sound) apply assumption+ apply simp
done
(*>*)
text\<open>The translation of general expressions is defined similarly, but
no code continuation is required.\<close>
inductive_set compileExpr::
"(Label \<times> (Label,Instr) AssList \<times> Expr \<times> ((Label,Instr) AssList \<times> Label)) set"
where
compilePrimE:
"\<lbrakk>(l, code, p, (code1,l1)) : compilePrim; OUT = (code1[l1\<mapsto>vreturn],l1+1)\<rbrakk>
\<Longrightarrow> (l, code, PrimE p, OUT):compileExpr"
|
compileLetE:
"\<lbrakk>(l, code, p, (code1,l1)) : compilePrim; (code2,l2) = (code1[l1\<mapsto>(store x)],l1+1);
(l2, code2, e, OUT) : compileExpr\<rbrakk>
\<Longrightarrow> (l, code, LetE x p e, OUT) : compileExpr"
|
compileCondE:
"\<lbrakk>(l+2, code, e2, (codeElse,XXX)) : compileExpr;
(XXX, codeElse, e1, (codeThen,YYY) ) : compileExpr ;
OUT = (codeThen[l\<mapsto>load x][(l+1)\<mapsto>(iftrue XXX)], YYY)\<rbrakk>
\<Longrightarrow> (l, code, CondE x e1 e2, OUT): compileExpr"
|
compileMatchE:
"\<lbrakk>(l+9, code, e2, (codeCons,lNil)) : compileExpr;
(lNil, codeCons, e1, (codeNil,lRes) ) : compileExpr ;
OUT = (codeNil[l\<mapsto>(load x)]
[(l+1)\<mapsto>(unop (\<lambda> v . if v = RVal Nullref
then TRUE else FALSE))]
[(l+2)\<mapsto>(iftrue lNil)]
[(l+3)\<mapsto>(load x)]
[(l+4)\<mapsto>(getfield LIST HD)]
[(l+5)\<mapsto>(store h)]
[(l+6)\<mapsto>(load x)]
[(l+7)\<mapsto>(getfield LIST TL)]
[(l+8)\<mapsto>(store t)], lRes) \<rbrakk>
\<Longrightarrow> (l, code, MatchE x e1 h t e2, OUT): compileExpr"
text\<open>Again, we prove an auxiliary result on the emitted code, by
induction on the compilation judgement.\<close>
lemma compileExpr_Prop1[rule_format]:
"(l,code,e,OUT) : compileExpr \<Longrightarrow>
(\<forall> code1 l1 . OUT = (code1, l1) \<longrightarrow>
(l < l1 \<and>
(\<forall> ll . ll < l \<longrightarrow> code1\<down>ll = code\<down>ll) \<and>
(\<forall> ll . l \<le> ll \<longrightarrow> ll < l1 \<longrightarrow> (\<exists> ins . code1\<down>ll = Some ins))))"
(*<*)
apply (erule compileExpr.induct)
(*PRIM*)
apply clarsimp apply (drule compilePrim_Prop1) apply fastforce apply clarsimp
apply rule apply clarsimp apply (erule_tac x=ll in allE, clarsimp)
apply (erule AL_update5) apply simp
apply clarsimp apply (case_tac "ll < l1", clarsimp)
apply (rotate_tac 1, erule_tac x=ll in allE, clarsimp)
apply (rule, erule AL_update5) apply simp
apply (subgoal_tac "ll= l1", clarsimp) apply (rule, simp add: AL_update1) apply simp
(*LET*)
apply clarsimp
apply (drule compilePrim_Prop1) apply fastforce apply clarsimp
apply rule apply clarsimp
apply (rule AL_update5) apply simp apply simp
apply clarsimp apply (erule_tac x=ll in allE)+ apply clarsimp
apply (case_tac "ll <l1", clarsimp) apply (rule, erule AL_update5) apply simp
apply (subgoal_tac "ll=l1", clarsimp)
apply (simp add: AL_update1) apply simp
(*COND*)
apply clarsimp
apply (rule, clarsimp)
apply (rule AL_update5)
apply (rule AL_update5) apply (simp ,simp)
apply simp
apply clarsimp
apply (case_tac "ll=l+1", clarsimp, rule) apply(rule AL_update1)
apply (case_tac "ll=l", clarsimp, rule)
apply (rule AL_update5) apply(simp add: AL_update1) apply simp
apply (case_tac "XXX \<le> ll", clarsimp)
apply (rotate_tac -3, erule_tac x=ll in allE, clarsimp)
apply rule apply (rule AL_update5) apply (erule AL_update5) apply (simp, simp)
apply clarsimp
apply (rotate_tac -6, erule_tac x=ll in allE, clarsimp)
apply (rotate_tac -2, erule_tac x=ll in allE, clarsimp)
apply rule apply (rule AL_update5) apply (erule AL_update5) apply (simp, simp)
(*Match*)
apply clarsimp
apply (rule, clarsimp)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5) apply (simp ,simp, simp) apply (simp ,simp, simp) apply (simp ,simp, simp)
apply simp
apply clarsimp
apply (case_tac "ll=l+8", clarsimp, rule) apply(simp add: AL_update1)
apply (case_tac "ll=l+7", clarsimp, rule) apply (rule AL_update5) apply(simp add: AL_update1) apply simp
apply (case_tac "ll=l+6", clarsimp, rule) apply (rule AL_update5) apply (rule AL_update5)
apply(simp add: AL_update1) apply simp apply simp
apply (case_tac "ll=l+5", clarsimp, rule) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply(simp add: AL_update1) apply (simp,simp,simp)
apply (case_tac "ll=l+4", clarsimp, rule) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply(simp add: AL_update1) apply (simp,simp,simp,simp)
apply (case_tac "ll=l+3", clarsimp, rule) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5) apply(simp add: AL_update1)
apply (simp,simp,simp,simp,simp)
apply (case_tac "ll=l+2", clarsimp, rule) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply(simp add: AL_update1)
apply (simp,simp,simp,simp,simp, simp)
apply (case_tac "ll=l+1", clarsimp, rule) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply(simp add: AL_update1)
apply (simp,simp,simp,simp,simp, simp, simp)
apply (case_tac "ll=l", clarsimp, rule) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply(simp add: AL_update1) apply (simp, simp)
apply (simp,simp,simp,simp,simp, simp)
apply (case_tac "lNil \<le> ll", clarsimp)
apply (rotate_tac -3, erule_tac x=ll in allE, clarsimp)
apply rule apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply (erule AL_update5)
apply (simp, simp, simp)
apply (simp, simp, simp)
apply (simp, simp, simp)
(* apply simp*)
apply (rotate_tac 5) apply( erule_tac x=ll in allE, erule impE) apply (erule thin_rl)
apply (erule thin_rl) apply (rotate_tac -1, erule thin_rl)
apply (rotate_tac -1, erule thin_rl)
apply (rotate_tac -1, erule thin_rl)
apply (rotate_tac -1, erule thin_rl)
apply (rotate_tac -1, erule thin_rl) apply simp
apply (subgoal_tac "ll < lNil")
prefer 2 apply( erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, rotate_tac 1, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl) apply simp
apply (erule impE, assumption)
apply (erule_tac x=ll in allE, erule impE, assumption)
apply (erule exE)
apply rule apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (rule AL_update5) apply (rule AL_update5) apply (rule AL_update5)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl)
apply (erule thin_rl, erule thin_rl) apply simp
apply fast+
done
(*>*)
text\<open>Then, soundness of the epxression type system is proven by
induction on the typing judgement.\<close>
lemma TP_epxr_Sound[rule_format]:
"(\<Sigma>,e,n):TP_expr \<Longrightarrow> Sig_good \<Sigma> \<longrightarrow>
(\<forall> l code code1 l1 G C m T MI Anno.
(l, code, e, (code1,l1)):compileExpr \<longrightarrow>
MST\<down>(C,m) = Some (T,MI,Anno) \<longrightarrow>
Segment C m l l1 code1 \<longrightarrow> deriv G C m l (Cachera n))"
(*<*)
+supply [[simproc del: defined_all]]
apply (erule TP_expr.induct)
(*SUB*)
apply clarsimp
apply (rotate_tac 2, erule thin_rl, rotate_tac -2)
apply (erule_tac x=l in allE, erule_tac x=code in allE)
apply (erule_tac x=code1 in allE, rotate_tac -1, erule_tac x=l1 in allE, clarsimp)
apply (erule_tac x=G in allE, erule_tac x=C in allE)
apply (erule_tac x=ma in allE, clarsimp)
apply (erule CACH_SUB) apply assumption
(*PRIM*)
apply clarsimp
apply (erule compileExpr.cases)
prefer 2 apply simp
prefer 2 apply simp
prefer 2 apply simp
apply clarsimp
apply (erule TP_prim_Sound) apply fast apply assumption+
apply (simp add: Segment_def, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (drule AL_update3) apply simp
apply (rule, rule, assumption, assumption)
prefer 2 apply simp
apply (simp add: Segment_def)
apply (erule_tac x=l1a in allE, simp) apply (erule impE)
apply (drule compilePrim_Prop1) apply fastforce apply simp
apply clarsimp
apply (simp add: AL_update1, clarsimp)
apply (rule CACH_INJECT)
apply (erule CACH_RET) apply assumption apply assumption
(*LET*)
apply clarsimp
apply (erule compileExpr.cases)
apply simp
prefer 2 apply simp
prefer 2 apply simp
apply clarsimp
apply (frule compilePrim_Prop1) apply fastforce
apply (frule compileExpr_Prop1) apply fastforce apply clarsimp
apply (erule_tac x="l1a+1" in allE, erule_tac x="code1a[l1a\<mapsto>store xa]" in allE,
erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=b in allE, clarsimp)
apply (erule_tac x=G in allE, erule_tac x=C in allE, erule_tac x=ma in allE, clarsimp)
apply (erule impE) apply (erule Segment_A) apply (simp,simp)
apply (rule TP_prim_Sound) apply assumption apply assumption
apply assumption apply assumption
apply (simp add: Segment_def) apply clarsimp apply (erule_tac x=ll in allE, clarsimp) apply (rule, rule, assumption)
apply(drule AL_update3) apply simp apply assumption
prefer 2 apply simp
apply (rule CACH_INJECT)
apply (rule Basic_Sound) prefer 2 apply assumption prefer 4 apply simp prefer 2 apply (subgoal_tac "l1a \<le> l1a", assumption, simp)
prefer 3 apply (subgoal_tac "a\<down>l1a = Some(store xa)", assumption)
apply (rotate_tac -3, erule_tac x=l1a in allE, clarsimp) apply (simp add: AL_update1)
apply (subgoal_tac "Segment C ma l1a b a", assumption) apply (erule Segment_A) apply (simp,simp)
apply simp
apply (simp add: basic_def)
apply (erule CACH_INJECT)
(*Cond*)
apply clarsimp
apply (erule compileExpr.cases)
apply simp apply simp prefer 2 apply simp
apply clarsimp
apply (erule_tac x=XXX in allE, erule_tac x= codeElse in allE,
erule_tac x=codeThen in allE, rotate_tac -1)
apply (erule_tac x= YYY in allE, clarsimp)
apply (erule_tac x=G in allE, rotate_tac -1, erule_tac x=C in allE,
erule_tac x=m in allE, clarsimp)
apply (erule_tac x="la+2" in allE, erule_tac x= codea in allE,
erule_tac x=codeElse in allE, rotate_tac -1)
apply (erule_tac x= XXX in allE, clarsimp)
apply (erule_tac x=G in allE, rotate_tac -1, erule_tac x=C in allE,
erule_tac x=m in allE, clarsimp)
apply (drule compileExpr_Prop1) apply fastforce apply clarsimp
apply (drule compileExpr_Prop1) apply fastforce apply clarsimp
apply (erule impE) apply (rotate_tac 5, erule thin_rl, simp add: Segment_def,clarsimp)
apply (rotate_tac -3, erule_tac x=ll in allE, clarsimp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp) apply (rule, rule, assumption, assumption)
apply (erule impE) apply (simp add: Segment_def,clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp) apply clarsimp
apply (frule Segment_triv) apply assumption apply (subgoal_tac "la \<le> la", assumption, simp) apply simp
apply clarsimp
apply (drule AL_update3) apply simp apply (simp add: AL_update1) apply clarsimp
apply (rule CACH_INSTR)
apply assumption
apply fast
apply assumption
apply assumption
apply (rule CACH_INJECT)
apply (frule Segment_triv) apply assumption apply (subgoal_tac "la \<le> la+1", assumption, simp) apply simp
apply clarsimp
apply (simp add: AL_update1) apply clarsimp
apply (frule Segment_triv) apply assumption apply (subgoal_tac "la \<le> XXX", assumption, simp) apply simp
apply clarsimp
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (rule CACH_IF)
apply assumption
apply assumption
apply assumption
apply (erule CACH_INJECT)
apply (subgoal_tac "la+2=2+la", clarsimp) apply (erule CACH_INJECT) apply simp
(*Match*)
apply clarsimp
apply (erule compileExpr.cases)
apply simp apply simp apply simp
apply clarsimp
apply (erule_tac x=lNil in allE, erule_tac x= codeCons in allE,
erule_tac x=codeNil in allE, rotate_tac -1)
apply (erule_tac x=lRes in allE, clarsimp)
apply (erule_tac x=G in allE, rotate_tac -1, erule_tac x=C in allE,
erule_tac x=m in allE, clarsimp)
apply (erule_tac x="la+9" in allE, erule_tac x=codea in allE,
erule_tac x=codeCons in allE, rotate_tac -1)
apply (erule_tac x=lNil in allE, clarsimp)
apply (erule_tac x=G in allE, rotate_tac -1, erule_tac x=C in allE,
erule_tac x=m in allE, clarsimp)
apply (drule compileExpr_Prop1) apply fastforce apply clarsimp
apply (drule compileExpr_Prop1) apply fastforce apply clarsimp
apply (erule impE) apply (rotate_tac 5, erule thin_rl)
apply (simp add: Segment_def, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp) apply clarsimp
apply (erule impE)
apply (simp add: Segment_def, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=ll in allE, clarsimp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp) apply clarsimp
apply (rule Basic_Sound) prefer 2 apply assumption prefer 4 apply simp prefer 2 apply (subgoal_tac "la \<le> la", assumption, simp)
apply (erule Segment_A) apply simp apply simp apply simp
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (simp add: AL_update1) apply (simp, simp, simp) apply (simp, simp, simp) apply (simp, simp)
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound) prefer 2 apply assumption prefer 4 apply simp prefer 2 apply (subgoal_tac "la +1 \<le> la+1", assumption, simp)
apply (erule Segment_A) apply simp apply simp apply simp
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (simp add: AL_update1) apply (simp, simp, simp) apply (simp, simp, simp) apply simp
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (frule Segment_triv) apply assumption apply (subgoal_tac "la \<le> lNil", assumption, simp, simp)
apply (frule Segment_triv) apply assumption apply (subgoal_tac "la \<le> la+3", assumption, simp, simp)
apply clarsimp
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (subgoal_tac "la+3=3+la", clarsimp) prefer 2 apply simp
apply (simp add: AL_update1) apply clarsimp
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (frule Segment_triv) apply assumption apply (subgoal_tac "la \<le> 2+la", assumption, simp, simp,clarsimp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (drule AL_update3, simp)
apply (simp add: AL_update1a, clarsimp)
apply (rule CACH_IF)
apply assumption
apply assumption
apply assumption
apply (erule CACH_INJECT)
apply simp
apply (rule CACH_INJECT)
apply (rule Basic_Sound)
prefer 2 apply assumption
prefer 2 apply (subgoal_tac "3+la \<le> 3+la", assumption, simp)
apply (erule Segment_A) apply (simp,simp)
apply simp
apply simp
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (simp add: AL_update1)
apply (simp,simp,simp)
apply (simp,simp)
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound)
prefer 2 apply assumption
prefer 2 apply (subgoal_tac "4+la \<le> 4+la", assumption, simp)
apply (erule Segment_A) apply (simp,simp)
apply simp
apply simp
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update1a) apply simp
apply (simp,simp,simp)
apply simp
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound)
prefer 2 apply assumption
prefer 2 apply (subgoal_tac "5+la \<le> 5+la", assumption, simp)
apply (erule Segment_A) apply (simp,simp)
apply simp
apply simp
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update1a) apply simp
apply (simp,simp,simp)
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound)
prefer 2 apply assumption
prefer 2 apply (subgoal_tac "6+la \<le> 6+la", assumption, simp)
apply (erule Segment_A) apply (simp,simp)
apply simp
apply simp
apply (rule AL_update5)
apply (rule AL_update5)
apply (rule AL_update1a) apply simp
apply (simp,simp)
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound)
prefer 2 apply assumption
prefer 2 apply (subgoal_tac "7+la \<le> 7+la", assumption, simp)
apply (erule Segment_A) apply (simp,simp)
apply simp
apply simp
apply (rule AL_update5)
apply (rule AL_update1a) apply simp
apply simp
apply (simp add: basic_def)
apply (rule CACH_INJECT)
apply (rule Basic_Sound)
prefer 2 apply assumption
prefer 2 apply (subgoal_tac "8+la \<le> 8+la", assumption, simp)
apply (erule Segment_A) apply (simp,simp)
apply simp
apply simp
apply (rule AL_update1a) apply simp
apply (simp add: basic_def)
apply (rule CACH_INJECT) apply (subgoal_tac "la+9=9+la", clarsimp)
apply simp
done
(*>*)
text\<open>The full translation of a functional program into a bytecode
program is defined as follows.\<close>
definition compileProg::"FunProg \<Rightarrow> bool" where
"compileProg F =
((\<forall> C m par e. F\<down>(C,m) = Some(par,e) \<longrightarrow>
(\<exists> code l0 l. mbody_is C m (rev par,code,l0) \<and>
(l0,[],e,(code,l)):compileExpr)) \<and>
(\<forall> C m. (\<exists> M. mbody_is C m M) = (\<exists> fdecl . F\<down>(C,m) = Some fdecl)))"
text\<open>The final condition relating a typing context to a method
specification table.\<close>
definition TP_MST::"TP_Sig \<Rightarrow> bool" where
"TP_MST \<Sigma> =
(\<forall> C m . case (MST\<down>(C,m)) of
None \<Rightarrow> \<Sigma>\<down>(C,m) = None
| Some(T,MI,Anno) \<Rightarrow> Anno = emp \<and>
(\<exists> n . \<Sigma>\<down>(C,m)=Some n \<and>
(T,MI,Anno) = mkSPEC (Cachera n) emp))"
text\<open>For well-typed programs, this property implies the earlier
condition on signatures.\<close>
lemma translation_good: "\<lbrakk>compileProg F; TP_MST \<Sigma>; TP \<Sigma> F\<rbrakk> \<Longrightarrow> Sig_good \<Sigma>"
(*<*)
apply (simp add: compileProg_def TP_MST_def Sig_good_def TP_def, clarsimp)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (case_tac "MST\<down>(C,m)", clarsimp,clarsimp)
done
(*>*)
text\<open>We can thus prove that well-typed function bodies satisfy the
specifications asserted by the typing context.\<close>
lemma CACH_BodiesDerivable[rule_format]:
"\<lbrakk> mbody_is C m (par, code, l); compileProg F; TP_MST \<Sigma>; TP \<Sigma> F\<rbrakk>
\<Longrightarrow> \<exists> n . MST\<down>(C,m) = Some(mkSPEC(Cachera n) emp) \<and>
deriv [] C m l (Cachera n)"
(*<*)
apply (subgoal_tac "(\<forall> C m par e. F\<down>(C,m) = Some(par,e) \<longrightarrow>
(\<exists> code l0 l. mbody_is C m (rev par,code,l0) \<and>
(l0,[],e,(code,l)):compileExpr)) \<and>
(\<forall> C m . (\<exists> M. mbody_is C m M) = (\<exists> fdecl . F\<down>(C,m) = Some fdecl))")
prefer 2 apply (simp add: compileProg_def, clarsimp)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (erule_tac x=C in allE, erule_tac x=m in allE, auto)
apply (simp add: mbody_is_def, clarsimp)
apply (subgoal_tac "((\<Sigma>\<down>(C,m) = None) = (F\<down>(C,m) = None)) \<and>
(\<forall> n par e . \<Sigma>\<down>(C,m) = Some n \<longrightarrow> F\<down>(C,m) = Some (par,e) \<longrightarrow> (\<Sigma>,e,n):TP_expr)")
prefer 2 apply (simp add: TP_def, clarsimp)
apply (subgoal_tac "MST\<down>(C, m) = Some (mkSPEC (Cachera y) emp)", clarsimp)
prefer 2 apply (simp add: TP_MST_def)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (case_tac "MST\<down>(C, m)", clarsimp, clarsimp)
apply (rule_tac x=y in exI, simp)
apply (rule TP_epxr_Sound) apply assumption
apply (erule translation_good) apply assumption+
apply (simp add: mkSPEC_def Cachera_def)
apply (drule compileExpr_Prop1) apply fastforce apply clarsimp
apply (simp add: Segment_def)
apply (rule, rule, rule, rule, simp add: mkSPEC_def Cachera_def)
apply clarsimp apply (rule, rule AL_emp1)
apply (erule_tac x=ll in allE)+
apply clarsimp
apply (simp add: mbody_is_def get_ins_def ins_is_def)
done
(*>*)
text\<open>From this, the overall soundness result follows easily.\<close>
theorem CACH_VERIFIED: "\<lbrakk>TP \<Sigma> F; TP_MST \<Sigma>; compileProg F\<rbrakk> \<Longrightarrow> VP"
(*<*)
apply (rule CACH_VP)
apply clarsimp apply (drule CACH_BodiesDerivable) apply assumption+ apply fast
apply clarsimp
done
(*>*)
(*<*)
end
(*>*)
diff --git a/thys/BytecodeLogicJmlTypes/Language.thy b/thys/BytecodeLogicJmlTypes/Language.thy
--- a/thys/BytecodeLogicJmlTypes/Language.thy
+++ b/thys/BytecodeLogicJmlTypes/Language.thy
@@ -1,515 +1,516 @@
(*File: Language.thy
Author: L Beringer & M Hofmann, LMU Munich
Date: 05/12/2008
Purpose: Syntax and operational semantics of subset of JVML
*)
(*<*)
theory Language imports AssocLists begin
(*>*)
section\<open>Language \label{sec:language}\<close>
subsection\<open>Syntax\<close>
text\<open>We have syntactic classes of (local) variables, class names,
field names, and method names. Naming restrictions, namespaces, long
Java names etc.~are not modelled.\<close>
typedecl Var
typedecl Class
typedecl Field
typedecl Method
text\<open>Since arithmetic operations are modelled as unimplemented
functions, we introduce the type of values in this section. The domain
of heap locations is arbitrary.\<close>
typedecl Addr
text\<open>A reference is either null or an address.\<close>
datatype Ref = Nullref | Loc Addr
text\<open>Values are either integer numbers or references.\<close>
datatype Val = RVal Ref | IVal int
text\<open>The type of (instruction) labels is fixed, since the operational
semantics increments the program counter after each instruction.\<close>
type_synonym Label = int
text\<open>Regarding the instructions, we support basic operand-stack
manipulations, object creation, field modifications, casts, static
method invocations, conditional and unconditional jumps, and a return
instruction.
For every (Isabelle) function \<open>f : Val\<Rightarrow>Val\<Rightarrow>Val\<close> we have an
instruction \<open>binop f\<close> whose semantics is to invoke \<open>f\<close>
on the two topmost values on the operand stack and replace them with
the result. Similarly for \<open>unop f\<close>.\<close>
datatype Instr =
const Val
| dup
| pop
| swap
| load Var
| store Var
| binop "Val \<Rightarrow> Val \<Rightarrow> Val"
| unop "Val \<Rightarrow> Val"
| new Class
| getfield Class Field
| putfield Class Field
| checkcast Class
| invokeS Class Method
| goto Label
| iftrue Label
| vreturn
text\<open>Method body declarations contain a list of formal parameters, a
mapping from instruction labels to instructions, and a start
label. The operational semantics assumes that instructions are
labelled consecutively\footnote{In the paper, we slightly abstract
from this by including a successor functions on labels}.\<close>
type_synonym Mbody = "Var list \<times> (Label, Instr) AssList \<times> Label"
text\<open>A class definition associates method bodies to method names.\<close>
type_synonym Classdef = "(Method, Mbody) AssList"
text\<open>Finally, a program consists of classes.\<close>
type_synonym Prog = "(Class, Classdef) AssList"
text\<open>Taken together, the three types \<open>Prog\<close>, \<open>Classdef\<close>,
and \<open>Mbody\<close> represent an abstract model of the virtual machine
environment. In our opinion, it would be desirable to avoid modelling
this environment at a finer level, at least for the purpose of the
program logic. For example, we prefer not to consider in detail the
representation of the constant pool.\<close>
subsection\<open>Dynamic semantics\<close>
subsubsection\<open>Semantic components\<close>
text\<open>An object consists of the identifier of its dynamic class and a
map from field names to values. Currently, we do not model
type-correctness, nor do we require that all (or indeed any) of the
fields stem from the static definition of the class, or a super-class.
Note, however, that type correctness can be expressed in the logic.\<close>
type_synonym Object = "Class \<times> (Field, Val) AssList"
text\<open>The heap is represented as a map from addresses to values. The
JVM specification does not prescribe any particular object layout. The
proposed type reflects this indeterminacy, but allows one to calculate
the byte-correct size of a heap only after a layout scheme has been
supplied. Alternative heap models would be the store-less semantics in
the sense of Jonkers~\cite{Jonkers1981} and
Deutsch~\cite{Deutsch1992}, (where the heap is modelled as a partial
equivalence relation on access paths), or object-based semantics in
the sense of Reddy~\cite{Reddy1996}, where the heap is represented as
a history of update operations. H\"ahnle et al.~use a variant of the
latter in their dynamic logic for a {\sc
JavaCard}~~\cite{HaehnleM:Cassis2005}.\<close>
type_synonym Heap = "(Addr, Object) AssList"
text\<open>Later, one might extend heaps by a component for static fields.\<close>
text\<open>The types of the (register) store and the operand stack are as
expected.\<close>
type_synonym Store = "(Var, Val) AssList"
type_synonym OpStack = "Val list"
text\<open>States contain an operand stack, a store, and a heap.\<close>
type_synonym State = "OpStack \<times> Store \<times> Heap"
definition heap::"State \<Rightarrow> Heap"
where "heap s = snd(snd s)"
text\<open>The operational semantics and the program logic are defined
relative to a fixed program \<open>P\<close>. Alternatively, the type of the
operational semantics (and proof judgements) could be extended by a
program component. We also define the constant value \<open>TRUE\<close>,
the representation of which does not matter for the current
formalisation.\<close>
axiomatization P::Prog and TRUE::Val
text\<open>In order to obtain more readable rules, we define operations
for extracting method bodies and instructions from the program.\<close>
definition mbody_is::"Class \<Rightarrow> Method \<Rightarrow> Mbody \<Rightarrow> bool"
where "mbody_is C m M = (\<exists> CD . P\<down>C = Some CD \<and> CD\<down>m = Some M)"
definition get_ins::"Mbody \<Rightarrow> Label \<Rightarrow> Instr option"
where "get_ins M l = (fst(snd M))\<down>l"
definition ins_is::"Class \<Rightarrow> Method \<Rightarrow> Label \<Rightarrow> Instr \<Rightarrow> bool"
where "ins_is C m l ins = (\<exists> M . mbody_is C m M \<and> get_ins M l = Some ins)"
text\<open>The transfer of method arguments from the caller's operand stack
to the formal parameters of an invoked method is modelled by the
predicate\<close>
inductive_set Frame::"(OpStack \<times> (Var list) \<times> Store \<times> OpStack) set"
where
FrameNil: "\<lbrakk>oo=ops\<rbrakk> \<Longrightarrow> (ops,[],emp,oo) : Frame"
|
Frame_cons: "\<lbrakk>(oo,par,S,ops) : Frame; R =S[x\<mapsto>v]\<rbrakk>
\<Longrightarrow> (v # oo, x # par,R,ops):Frame"
(*<*)
lemma Frame_deterministic[rule_format]:
"(ops, par, S, os) \<in> Frame \<Longrightarrow>
(\<forall> R opsa . (ops, par, R, opsa) \<in> Frame \<longrightarrow> R=S \<and> opsa = os)"
apply (erule Frame.induct, clarsimp)
apply (erule Frame.cases, clarsimp, clarsimp)
apply (erule thin_rl, clarsimp)
apply (erule Frame.cases, clarsimp, clarsimp)
done
(*>*)
text\<open>In order to obtain a deterministic semantics, we assume the
existence of a function, with the obvious freshness axiom for this
construction.\<close>
axiomatization nextLoc::"Heap \<Rightarrow> Addr"
where nextLoc_fresh: "h\<down>(nextLoc h) = None"
subsubsection\<open>Operational judgements\<close>
text\<open>Similar to Bannwart-M\"uller~\cite{BannwartMueller05}, we define
two operational judgements: a one-step relation and a relation that
represents the transitive closure of the former until the end of the
current method invocation. These relations are mutually recursive,
since the method invocation rule contracts the execution of the
invoked method to a single step. The one-step relation associates a
state to its immediate successor state, where the program counter is
interpreted with respect to the current method body. The transitive
closure ignores the bottom part of the operand stack and the store of
the final configuration. It simply returns the heap and the result of
the method invocation, where the latter is given by the topmost value
on the operand stack. In contrast to~\cite{BannwartMueller05}, we do
not use an explicit \<open>return\<close> variable. Both relations take an
additional index of type \<open>nat\<close> that monitors the derivation
height. This is useful in the proof of soundness of the program
logic.\<close>
text\<open>Intuitively, \<open>(M,l,s,n,l',s'):Step\<close> means that method
(body) \<open>M\<close> evolves in one step from state \<open>s\<close> to state
\<open>s'\<close>, while statement \<open>(M,s,n,h,v):Exec\<close> indicates that
executing from \<open>s\<close> in method \<open>M\<close> leads eventually to a
state whose final value is \<open>h\<close>, where precisely the last step in
this sequence is a \<open>vreturn\<close> instruction and the return value is
\<open>v\<close>.\<close>
text\<open>Like Bannwart and M\"uller, we define a "frame-less"
semantics. i.e.~the execution of a method body is modelled by a
transitive closure of the basic step-relation, which results in a
one-step reduction at the invocation site. Arguably, an operational
semantics with an explicit frame stack is closer to the real JVM. It
should not be difficult to verify the operational soundness of the
present system w.r.t.~such a finer model, or to modify the
semantics.\<close>
inductive_set
Step::"(Mbody \<times> Label \<times> State \<times> nat \<times> Label \<times> State) set"
and
Exec::"(Mbody \<times> Label \<times> State \<times> nat \<times> Heap \<times> Val) set"
where
Const:"\<lbrakk>get_ins M l = Some (const v); NEXT = (v # os,s,h); ll=l+1\<rbrakk>
\<Longrightarrow> (M,l,(os,s,h), 1, ll, NEXT) : Step"
|
Dup: "\<lbrakk>get_ins M l = Some dup; NEXT = (v # v # os,s,h); ll =l+1\<rbrakk>
\<Longrightarrow> (M,l,(v # os,s,h), 1, ll, NEXT) : Step"
|
Pop: "\<lbrakk>get_ins M l = Some pop; NEXT = (os,s,h); ll=l+1\<rbrakk>
\<Longrightarrow> (M,l,(v # os,s,h), 1, ll, NEXT) : Step"
|
Swap: "\<lbrakk>get_ins M l = Some swap; NEXT = (w # (v # os),s,h); ll= l+1\<rbrakk>
\<Longrightarrow> (M,l,(v # (w # os),s,h), 1, ll, NEXT) : Step"
|
Load: "\<lbrakk>get_ins M l = Some (load x); s\<down>x = Some v;
NEXT = (v # os,s,h); ll=l+1\<rbrakk>
\<Longrightarrow> (M,l,(os,s,h), 1, ll,NEXT) : Step"
|
Store:"\<lbrakk>get_ins M l = Some (store x); NEXT = (os,s[x\<mapsto>v],h); ll= l+1\<rbrakk>
\<Longrightarrow> (M,l,(v # os,s,h), 1, ll, NEXT) : Step"
|
Binop:"\<lbrakk>get_ins M l = Some (binop f); NEXT = ((f v w) # os,s,h); ll=l+1\<rbrakk>
\<Longrightarrow> (M,l,(v # (w # os),s,h), 1, ll,NEXT) : Step"
|
Unop: "\<lbrakk>get_ins M l = Some (unop f); NEXT = ((f v) # os,s,h);ll=l+1\<rbrakk>
\<Longrightarrow> (M,l,(v # os,s,h), 1, ll, NEXT) : Step"
|
New: "\<lbrakk>get_ins M l = Some (new d); newobj = (d, emp); a=nextLoc h;
NEXT = ((RVal (Loc a)) # os,s,h[a\<mapsto>newobj]); ll = l+1\<rbrakk>
\<Longrightarrow> (M,l,(os,s,h), 1, ll,NEXT) : Step"
|
Get: "\<lbrakk>get_ins M l = Some (getfield d F); h\<down>a = Some (d, Flds);
Flds\<down>F = Some v; NEXT = (v # os,s,h); ll=l+1\<rbrakk>
\<Longrightarrow> (M,l,((RVal (Loc a)) # os,s,h), 1, ll,NEXT) : Step"
|
Put: "\<lbrakk>get_ins M l = Some (putfield d F); h\<down>a = Some (d, Flds);
newobj = (d, Flds[F\<mapsto>v]); NEXT = (os,s,h[a\<mapsto>newobj]); ll=l+1\<rbrakk>
\<Longrightarrow> (M,l,(v # ((RVal (Loc a)) # os),s,h), 1, ll, NEXT) : Step"
|
Cast: "\<lbrakk>get_ins M l = Some (checkcast d); h\<down>a = Some (d, Flds);
NEXT = ((RVal (Loc a)) # os,s,h); ll=l+1\<rbrakk>
\<Longrightarrow> (M,l,((RVal (Loc a)) # os,s,h), 1, ll,NEXT) : Step"
|
Goto: "\<lbrakk>get_ins M l = Some (goto pc)\<rbrakk> \<Longrightarrow> (M,l,S, 1, pc,S) : Step"
|
IfT: "\<lbrakk>get_ins M l = Some (iftrue pc); NEXT = (os,s,h)\<rbrakk>
\<Longrightarrow> (M,l,(TRUE # os,s,h), 1, pc, NEXT) : Step"
|
IfF: "\<lbrakk>get_ins M l = Some (iftrue pc); v \<noteq> TRUE; NEXT = (os,s,h); ll=l+1\<rbrakk>
\<Longrightarrow> (M,l,(v # os,s,h), 1, ll, NEXT) : Step"
|
InvS: "\<lbrakk>get_ins M l = Some (invokeS C m); mbody_is C m (par,code,l0);
((par,code,l0),l0,([], S,h), n, hh, v): Exec;
(ops,par,S,os) : Frame; NEXT = (v # os,s,hh); ll = l+1\<rbrakk>
\<Longrightarrow> (M,l,(ops,s,h), Suc n, ll, NEXT) : Step"
|
Vret: "\<lbrakk>get_ins M l = Some vreturn\<rbrakk> \<Longrightarrow> (M,l,(v # os,s,h), 1, h, v) : Exec"
|
Run: "\<lbrakk>(M,l,s,n,ll,t):Step; (M,ll,t,m,h,v):Exec; k = (max n m) +1 \<rbrakk>
\<Longrightarrow> (M,l,s,k,h,v) : Exec"
text\<open>A big-step operational judgement that abstracts from the
derivation height is easily defined.\<close>
definition Opsem::"Mbody \<Rightarrow> Label \<Rightarrow> State \<Rightarrow> Heap \<Rightarrow> Val \<Rightarrow> bool"
where "Opsem M l s h v = (\<exists> n . (M,l,s,n,h,v):Exec)"
subsection \<open>Basic properties\<close>
text \<open>We provide elimination lemmas for the inductively defined
relations\<close>
inductive_cases eval_cases:
"(M,l,s,n,ll,t) : Step"
"(M,l,s,n,h,v) : Exec"
(*<*)
lemma no_zero_height_derivsAux[rule_format]:
"\<forall>n . ((M,l,s,n,ll,t) : Step \<longrightarrow> (n=0 \<longrightarrow> False)) \<and> ((MM,lll,ss,m,h,v):Exec \<longrightarrow> (m=0 \<longrightarrow> False))"
by (rule allI, rule Step_Exec.induct, simp_all)
lemma no_zero_height_derivsAux2: "((M,l,s,0,ll,t):Step \<longrightarrow> False) \<and> ((MM,lll,ss,0,h,v):Exec \<longrightarrow> False)"
by (insert no_zero_height_derivsAux, fast)
(*>*)
text \<open>and observe that no derivations of height 0 exist.\<close>
lemma no_zero_height_Step_derivs: "(M,l,s,0,ll,t):Step \<Longrightarrow> False"
(*<*)by (insert no_zero_height_derivsAux2, fast)(*>*)
(*<*)
lemma no_zero_height_Step_derivs1: "(M,l,(os,S,H),0,ll,t):Step \<Longrightarrow> False"
by (insert no_zero_height_derivsAux2, fast)
(*>*)
lemma no_zero_height_Exec_derivs: "(M,l,s,0,h,v):Exec \<Longrightarrow> False"
(*<*)by (insert no_zero_height_derivsAux2, fast)(*>*)
(*<*)
lemma no_zero_height_Exec_derivs1: "(M,l,(os,S,H),0,h,v):Exec \<Longrightarrow> False"
by (insert no_zero_height_derivsAux2, fast)
(*>*)
(*<*)
(*Elimination rules*)
lemma ConstElim1:"\<lbrakk>(M, l, (os, S, h), n, ll,t) \<in> Step; get_ins M l = Some (const v)\<rbrakk>
\<Longrightarrow> n = Suc 0 \<and> t = (v # os, S, h) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma DupElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some dup\<rbrakk>
\<Longrightarrow> \<exists> v ops . os = v # ops \<and> n = Suc 0 \<and> t = (v # os, S, h) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma PopElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some pop\<rbrakk>
\<Longrightarrow> \<exists> v ops . os = v # ops \<and> n = Suc 0 \<and> t = (ops, S, h) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma SwapElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some swap\<rbrakk>
\<Longrightarrow> \<exists> v w ops . os = v # w # ops \<and> n = Suc 0 \<and> t = (w # v # ops, S, h) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma LoadElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (load x)\<rbrakk>
\<Longrightarrow> \<exists> v . S\<down>x = Some v \<and> n = Suc 0 \<and> t = (v # os, S, h) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma StoreElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (store x)\<rbrakk>
\<Longrightarrow> \<exists> v ops . os = v # ops \<and> n = Suc 0 \<and> t = (ops, S[x\<mapsto>v], h) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma BinopElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (binop f)\<rbrakk>
\<Longrightarrow> \<exists> v w ops . os = v # w # ops \<and> n = Suc 0 \<and> t = (f v w # ops, S, h) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma UnopElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (unop f)\<rbrakk>
\<Longrightarrow> \<exists> v ops . os = v # ops \<and> n = Suc 0 \<and> t = (f v # ops, S, h) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma NewElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (new d)\<rbrakk>
\<Longrightarrow> \<exists> a . a = nextLoc h \<and> n = Suc 0 \<and> t = (RVal (Loc a) # os, S, h[a\<mapsto>(d, emp)]) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma GetElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (getfield d F)\<rbrakk>
\<Longrightarrow> \<exists> a Flds v ops. h\<down>a = Some (d, Flds) \<and> Flds\<down>F = Some v \<and>
os = RVal (Loc a) # ops \<and> n = Suc 0 \<and> t = (v # ops, S, h) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma PutElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (putfield d F)\<rbrakk>
\<Longrightarrow> \<exists> a Flds v ops. h\<down>a = Some (d, Flds) \<and> os = v # RVal (Loc a) # ops \<and>
n = Suc 0 \<and> t = (ops, S, h[a\<mapsto>(d, Flds[F\<mapsto>v])]) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma CastElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (checkcast d)\<rbrakk>
\<Longrightarrow> \<exists> a Flds ops . h\<down>a = Some (d, Flds) \<and> os = RVal (Loc a) # ops \<and> n = Suc 0 \<and>
t = (RVal (Loc a) # ops, S, h) \<and> ll = l+1"
by (erule eval_cases, simp_all)
lemma GotoElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (goto pc)\<rbrakk>
\<Longrightarrow> n = Suc 0 \<and> t = (os, S, h) \<and> ll = pc"
by (erule eval_cases, simp_all)
lemma IfElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (iftrue pc)\<rbrakk>
\<Longrightarrow> (\<exists> ops . os = TRUE # ops \<and> n = Suc 0 \<and> t = (ops, S, h) \<and> ll = pc) \<or>
(\<exists> v ops . v \<noteq> TRUE \<and> os = v # ops \<and> n = Suc 0 \<and> t = (ops, S, h) \<and> ll = l+1)"
by (erule eval_cases, simp_all)
lemma InvokeElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (invokeS C m)\<rbrakk>
\<Longrightarrow> \<exists> code l0 v ops hh u k R par.
mbody_is C m (par,code, l0) \<and> (os,par,R,ops):Frame \<and>
((par,code,l0), l0, ([], R, h), k, hh, v) \<in> Exec \<and>
n = Suc k \<and> t = (v # ops, S, hh) \<and> ll = l+1"
by (erule eval_cases, simp_all, fastforce)
lemma InvokeElim: "\<lbrakk>(M, l, s, n, ll, t) \<in> Step; get_ins M l = Some (invokeS C m)\<rbrakk>
\<Longrightarrow> \<exists> code l0 v ops hh u k R par.
mbody_is C m (par,code, l0) \<and> (fst s,par,R,ops):Frame \<and>
((par,code,l0), l0, ([], R, snd(snd s)), k, hh, v) \<in> Exec \<and>
n = Suc k \<and> t = (v # ops, fst (snd s), hh) \<and> ll = l+1"
by (erule eval_cases, simp_all, fastforce)
lemma RetElim1: "\<lbrakk>(M, l, (os, S, h), n, ll, t) \<in> Step; get_ins M l = Some (vreturn)\<rbrakk> \<Longrightarrow> False"
by (erule eval_cases, simp_all)
lemma ExecElim1: "\<lbrakk>(M,l,(os,S,H),k,h,v):Exec\<rbrakk>
\<Longrightarrow> (get_ins M l = Some vreturn \<and> (\<exists> ops . os = v # ops \<and> k = Suc 0 \<and> h=H)) \<or>
(\<exists> n m t ll. (M, l,(os,S,H), n, ll,t) \<in> Step \<and> (M, ll, t, m, h, v) \<in> Exec \<and> k = Suc (max n m))"
apply (erule eval_cases, simp_all)
apply (rule disjI2)
apply clarsimp
apply (rule, rule, rule, rule) apply (rule, rule, rule) apply assumption
apply (rule, assumption) apply simp
done
lemma InstrElimNext:
"\<lbrakk>(M, l, s, n, ll, t) \<in> Step;
get_ins M l = Some I;
I = const c \<or> I = dup \<or> I = pop \<or> I = swap \<or> I = load x \<or>
I = store x \<or> I = binop f \<or> I = unop g \<or> I = new d \<or>
I = getfield d F \<or> I = putfield d F \<or> I = checkcast d\<rbrakk>
\<Longrightarrow> ll = l+1"
apply (drule eval_cases, simp_all)
apply clarsimp
apply clarsimp
done
(*>*)
text\<open>By induction on the derivation system one can show
determinism.\<close>
(*<*)
lemma StepExec_determ_Aux[rule_format]:
"(\<forall> n1 l M s l1 t . n1 \<le> n \<longrightarrow> (M, l, s, n1, l1, t) \<in> Step \<longrightarrow>
(\<forall> n2 l2 r. (M,l,s,n2,l2,r):Step \<longrightarrow> (n1=n2 \<and> t=r \<and> l1=l2))) \<and>
(\<forall> n1 l M s h v . n1 \<le> n \<longrightarrow> (M, l, s, n1, h, v) \<in> Exec \<longrightarrow>
(\<forall> n2 k w . (M,l,s,n2,k,w):Exec \<longrightarrow> (n1=n2 \<and> h=k \<and> v=w)))"
+supply [[simproc del: defined_all]]
apply (induct n)
apply clarsimp apply rule apply clarsimp apply (drule no_zero_height_Step_derivs1, simp)
apply clarsimp apply (drule no_zero_height_Exec_derivs1, simp)
apply clarsimp
apply rule
apply clarsimp
apply (erule Step.cases)
apply clarsimp apply (drule ConstElim1) apply simp apply clarsimp
apply clarsimp apply (drule DupElim1) apply simp apply clarsimp
apply clarsimp apply (drule PopElim1) apply simp apply clarsimp
apply clarsimp apply (drule SwapElim1) apply simp apply clarsimp
apply clarsimp apply (drule LoadElim1) apply simp apply clarsimp
apply clarsimp apply (drule StoreElim1) apply simp apply clarsimp
apply clarsimp apply (drule BinopElim1) apply simp apply clarsimp
apply clarsimp apply (drule UnopElim1) apply simp apply clarsimp
apply clarsimp apply (drule NewElim1) apply simp apply clarsimp
apply clarsimp apply (drule GetElim1) apply fastforce apply clarsimp
apply clarsimp apply (drule PutElim1) apply fastforce apply clarsimp
apply clarsimp apply (drule CastElim1) apply simp apply clarsimp
apply clarsimp apply (drule GotoElim1) apply simp apply clarsimp
apply clarsimp apply (drule IfElim1) apply simp apply clarsimp
apply clarsimp apply (drule IfElim1) apply simp apply clarsimp
apply clarsimp apply (drule InvokeElim1) apply simp apply clarsimp
apply (erule thin_rl)
apply (simp add: mbody_is_def, clarsimp)
apply (drule Frame_deterministic, assumption, clarsimp)
apply clarsimp
apply (erule Exec.cases)
apply clarsimp apply (erule Exec.cases)
apply clarsimp
apply clarsimp apply (drule RetElim1, simp, simp)
apply clarsimp apply (drule ExecElim1)
apply clarsimp
apply (erule disjE, clarsimp) apply (drule RetElim1, simp, simp)
apply clarsimp
apply (erule_tac x=na in allE, rotate_tac -1, clarsimp)
apply (erule_tac x=la in allE, rotate_tac -1)
apply (erule_tac x=ad in allE, rotate_tac -1)
apply (erule_tac x=ae in allE, rotate_tac -1)
apply (erule_tac x=bb in allE, rotate_tac -1)
apply (erule_tac x=af in allE, rotate_tac -1)
apply (erule_tac x=ag in allE, rotate_tac -1)
apply (erule_tac x=bc in allE, rotate_tac -1)
apply (erule_tac x=ll in allE, rotate_tac -1)
apply (erule_tac x=ah in allE, rotate_tac -1)
apply (erule_tac x=ai in allE, rotate_tac -1)
apply (erule_tac x=bd in allE, clarsimp)
apply (rotate_tac -1)
apply (erule_tac x=nb in allE, rotate_tac -1)
apply (erule_tac x=lla in allE, rotate_tac -1)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=aa in allE, rotate_tac -1)
apply (erule_tac x=b in allE, clarsimp)
apply (erule_tac x=m in allE, rotate_tac -1, clarsimp)
apply (erule_tac x=lla in allE, rotate_tac -1)
apply (erule_tac x=ad in allE, rotate_tac -1)
apply (erule_tac x=ae in allE, rotate_tac -1)
apply (erule_tac x=bb in allE, rotate_tac -1)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=aa in allE, rotate_tac -1)
apply (erule_tac x=b in allE, rotate_tac -1)
apply (erule_tac x=ha in allE, rotate_tac -1)
apply (erule_tac x=va in allE, rotate_tac -1, clarsimp)
done
(*>*)
lemma Step_determ:
"\<lbrakk>(M,l,s,n,l1,t) \<in> Step; (M,l,s,m,l2,r):Step\<rbrakk> \<Longrightarrow> n=m \<and> t=r \<and> l1=l2"
(*<*)
apply (insert StepExec_determ_Aux[of n])
apply (erule conjE)
apply (rotate_tac -1, erule thin_rl)
apply fast
done
(*>*)
lemma Exec_determ:
"\<lbrakk>(M,l,s,n,h,v) \<in> Exec; (M,l,s,m,k,w):Exec\<rbrakk> \<Longrightarrow> n=m \<and> h=k \<and> v=w"
(*<*)
apply (insert StepExec_determ_Aux[of n])
apply (erule conjE)
apply (rotate_tac -2, erule thin_rl)
apply fast
done
(*>*)
(*<*)
end
(*>*)
diff --git a/thys/BytecodeLogicJmlTypes/Sound.thy b/thys/BytecodeLogicJmlTypes/Sound.thy
--- a/thys/BytecodeLogicJmlTypes/Sound.thy
+++ b/thys/BytecodeLogicJmlTypes/Sound.thy
@@ -1,1076 +1,1078 @@
(*File: Sound.thy
Author: L Beringer & M Hofmann, LMU Munich
Date: 05/12/2008
Purpose: Interpretation of judgements, and soundness proof,
of the program logic with invariants, pre-
and post-conditions, and local assertions, but no
exploitation of sucessful bytecode verification.
*)
(*<*)
theory Sound imports Logic MultiStep Reachability begin
(*>*)
section\<open>Soundness\<close>
text\<open>This section contains the soundness proof of the program logic.
In the first subsection, we define our notion of validity, thus
formalising our intuitive explanation of the terms preconditions,
specifications, and invariants. The following two subsections contain
the details of the proof and can easily be skipped during a first pass
through the document.\<close>
subsection\<open>Validity\<close>
text\<open>A judgement is valid at the program point \<open>C.m.l\<close>
(i.e.~at label \<open>l\<close> in method \<open>m\<close> of class \<open>C\<close>),
written $\mathit{valid}\; C\; m\; l\; A\; B\; I$ or, in symbols, $$\vDash\,
\lbrace A \rbrace\, C,m,l\, \lbrace B \rbrace\, I,$$ if $A$ is a
precondition for $B$ and for all local annotations following $l$ in an
execution of \<open>m\<close>, and all reachable states in the current frame
or yet-to-be created subframes satisfy $I$. More precisely,
whenever an execution of the method starting in an initial state $s_0$
reaches the label \<open>l\<close> with state \<open>s\<close>, the following
properties are implied by $A(s_0,s)$.
\begin{enumerate}
\item If the continued execution from \<open>s\<close> reaches a final
state \<open>t\<close> (i.e.~the method terminates), then that final state
\<open>t\<close> satisfies $B(s_0,s,t)$.
\item Any state $s'$ visited in the current frame during the remaining
program execution whose label carries an annotation \<open>Q\<close> will
satisfy $Q(s_0,s')$, even if the execution of the frame does not
terminate.
\item Any state $s'$ visited in the current frame or a subframe of
the current frame will satisfy $I(s_0,s,\mathit{heap}(s'))$, again even if the
execution does not terminate.
\end{enumerate}
Formally, this interpretation is expressed as follows.
\<close>
definition valid::"Class \<Rightarrow> Method \<Rightarrow> Label \<Rightarrow> Assn \<Rightarrow> Post \<Rightarrow> Inv \<Rightarrow> bool" where
"valid C m l A B I =
(\<forall> M. mbody_is C m M \<longrightarrow>
(\<forall> Mspec Minv Anno . MST\<down>(C,m) = Some(Mspec,Minv,Anno) \<longrightarrow>
(\<forall> par code l0 . M = (par,code,l0) \<longrightarrow>
(\<forall> s0 s . MS M l0 (mkState s0) l s \<longrightarrow> A s0 s \<longrightarrow>
((\<forall> h v. Opsem M l s h v \<longrightarrow> B s0 s (h,v)) \<and>
(\<forall> ll r . (MS M l s ll r \<longrightarrow> (\<forall> Q . Anno\<down>(ll) = Some Q \<longrightarrow> Q s0 r)) \<and>
(Reach M l s r \<longrightarrow> I s0 s (heap r))))))))"
abbreviation valid_syntax :: "Assn \<Rightarrow> Class \<Rightarrow> Method \<Rightarrow>
Label \<Rightarrow> Post \<Rightarrow> Inv \<Rightarrow> bool"
(" \<Turnstile> \<lbrace> _ \<rbrace> _ , _ , _ \<lbrace> _ \<rbrace> _" [200,200,200,200,200,200] 200)
where "valid_syntax A C m l B I == valid C m l A B I"
text\<open>This notion of validity extends that of Bannwart-M\"uller by
allowing the post-condition to differ from method specification and to
refer to the initial state, and by including invariants. In
the logic of Bannwart and M\"uller, the validity of a method
specification is given by a partial correctness (Hoare-style)
interpretation, while the validity of preconditions of
individual instructions is such that a precondition at $l$ implies the
preconditions of its immediate control flow successors.\<close>
text\<open>Validity us lifted to contexts and the method specification
table. In the case of the former, we simply require that all entries
be valid.\<close>
definition G_valid::"CTXT \<Rightarrow> bool" where
"G_valid G = (\<forall> C m l A B I. G\<down>(C,m,l) = Some (A,B,I)\<longrightarrow>
\<Turnstile> \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I)"
text\<open>Regarding the specification table, we require that the initial
label of each method satisfies an assertion that ties the method
precondition to the current state.\<close>
definition MST_valid ::bool where
"MST_valid = (\<forall> C m par code l0 T MI Anno.
mbody_is C m (par,code,l0) \<longrightarrow> MST\<down>(C, m) = Some (T,MI,Anno) \<longrightarrow>
\<Turnstile> \<lbrace>(\<lambda> s0 s. s = mkState s0)\<rbrace> C, m, l0 \<lbrace>(mkPost T)\<rbrace> (mkInv MI))"
definition Prog_valid::bool where
"Prog_valid = (\<exists> G . G_valid G \<and> MST_valid)"
text\<open>The remainder of this section contains a proof of soundness,
i.e.~of the property $$\<open>VP\<close> \Longrightarrow \<open>Prog_valid\<close>,$$ and is structured into two parts. The first step
(Section \ref{SoundnessUnderValidContexts}) establishes a soundness
result where the \<open>VP\<close> property is replaced by validity
assumptions regarding the method specification table and the
context. In the second step (Section \ref{SectionContextElimination}),
we show that these validity assumptions are satisfied by verified
programs, which implies the overall soundness theorem.\<close>
subsection\<open>Soundness under valid contexts
\label{SoundnessUnderValidContexts}\<close>
text\<open>The soundness proof proceeds by induction on the axiomatic
semantics, based on an auxiliary lemma for method invocations that is
proven by induction on the derivation height of the operational
semantics. For the latter induction, relativised notions of validity
are employed that restrict the derivation height of the program
continuations affected by an assertion. The appropriate definitions of
relativised validity for judgements, for the precondition table, and
for the method specification table are as follows.\<close>
definition validn::
"nat \<Rightarrow> Class \<Rightarrow> Method \<Rightarrow> Label \<Rightarrow> Assn \<Rightarrow> Post \<Rightarrow> Inv \<Rightarrow> bool" where
"validn K C m l A B I =
(\<forall> M. mbody_is C m M \<longrightarrow>
(\<forall> Mspec Minv Anno . MST\<down>(C,m) = Some(Mspec,Minv,Anno) \<longrightarrow>
(\<forall> par code l0 . M = (par,code,l0) \<longrightarrow>
(\<forall> s0 s . MS M l0 (mkState s0) l s \<longrightarrow> A s0 s \<longrightarrow>
(\<forall> k . k \<le> K \<longrightarrow>
((\<forall> h v. (M,l,s,k,h,v):Exec \<longrightarrow> B s0 s (h,v)) \<and>
(\<forall> ll r . ((M,l,s,k,ll,r):MStep \<longrightarrow>
(\<forall> Q . Anno\<down>(ll) = Some Q \<longrightarrow> Q s0 r)) \<and>
((M,l,s,k,r):Reachable \<longrightarrow> I s0 s (heap r)))))))))"
abbreviation validn_syntax :: "nat \<Rightarrow> Assn \<Rightarrow> Class \<Rightarrow> Method \<Rightarrow>
Label \<Rightarrow> Post \<Rightarrow> Inv \<Rightarrow> bool"
("\<Turnstile>\<^sub>_ \<lbrace> _ \<rbrace> _ , _ , _ \<lbrace> _ \<rbrace> _ " [200,200,200,200,200,200,200] 200)
where "validn_syntax K A C m l B I == validn K C m l A B I"
definition G_validn::"nat \<Rightarrow> CTXT \<Rightarrow> bool" where
"G_validn K G = (\<forall> C m l A B I. G\<down>(C,m,l) = Some (A,B,I) \<longrightarrow>
\<Turnstile>\<^sub>K \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I)"
definition MST_validn::"nat \<Rightarrow> bool" where
"MST_validn K = (\<forall> C m par code l0 T MI Anno.
mbody_is C m (par,code,l0) \<longrightarrow> MST\<down>(C, m) = Some (T,MI,Anno) \<longrightarrow>
\<Turnstile>\<^sub>K \<lbrace>(\<lambda> s0 s. s = mkState s0)\<rbrace> C, m, l0 \<lbrace>(mkPost T)\<rbrace> (mkInv MI))"
definition Prog_validn::"nat \<Rightarrow> bool" where
"Prog_validn K = (\<exists> G . G_validn K G \<and> MST_validn K)"
text\<open>The relativised notions are related to each other, and to the
native notions of validity as follows.\<close>
lemma valid_validn: "\<Turnstile> \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I \<Longrightarrow> \<Turnstile>\<^sub>K \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I"
(*<*)
apply (simp add: valid_def validn_def Opsem_def MS_def Reach_def)
apply clarsimp
apply (erule_tac x=a in allE)
apply (erule_tac x=aa in allE)
apply (erule_tac x=b in allE, clarsimp)
apply (erule_tac x=ab in allE)
apply (erule_tac x=ba in allE)
apply (erule_tac x=ac in allE)
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE, erule impE) apply fast
apply fastforce
done
(*>*)
lemma validn_valid: "\<lbrakk>\<forall> K . \<Turnstile>\<^sub>K \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I\<rbrakk> \<Longrightarrow> \<Turnstile> \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I"
(*<*)
apply (unfold valid_def validn_def)(*apply ( Opsem_def MS_def Reach_def)*)
apply (rule, rule, rule)
apply (rule, rule, rule)
apply (rule, rule, rule)
apply (rule, rule, rule)
apply (rule, rule)
apply (rule,rule,rule, rule)
apply (unfold Opsem_def, erule exE)
apply (erule_tac x=n in allE)
apply (erule_tac x=M in allE, erule impE, assumption)
apply (erule_tac x=Mspec in allE, erule_tac x=Minv in allE, erule_tac x=Anno in allE, erule impE, assumption)
apply (erule_tac x=par in allE, erule_tac x=code in allE)
apply (erule_tac x=l0 in allE, erule impE, assumption)
apply (erule_tac x=s0 in allE, erule_tac x=s in allE)
apply (erule impE, assumption)
apply (erule impE, assumption)
apply (erule_tac x=n in allE, erule impE, simp)
apply fast
apply (rule, rule, rule)
apply (rule, rule, rule)
apply (unfold MS_def, erule exE, erule exE)
apply (erule_tac x=ka in allE)
apply (erule_tac x=M in allE, erule impE, assumption)
apply (erule_tac x=Mspec in allE, erule_tac x=Minv in allE, erule_tac x=Anno in allE, erule impE, assumption)
apply (erule_tac x=par in allE, erule_tac x=code in allE)
apply (erule_tac x=l0 in allE, erule impE, assumption)
apply (erule_tac x=s0 in allE, erule_tac x=s in allE)
apply (erule impE, fast)
apply (erule impE, assumption)
apply (erule_tac x=ka in allE, erule impE, simp)
apply fast
apply rule
apply (unfold Reach_def, erule exE, erule exE)
apply (erule_tac x=ka in allE)
apply (erule_tac x=M in allE, erule impE, assumption)
apply (erule_tac x=Mspec in allE, erule_tac x=Minv in allE, erule_tac x=Anno in allE, erule impE, assumption)
apply (erule_tac x=par in allE, erule_tac x=code in allE)
apply (erule_tac x=l0 in allE, erule impE, assumption)
apply (erule_tac x=s0 in allE, erule_tac x=s in allE)
apply (erule impE, fast)
apply (erule impE, assumption)
apply (erule_tac x=ka in allE, erule impE, simp)
apply fast
done
(*>*)
lemma validn_lower:
"\<lbrakk>\<Turnstile>\<^sub>K \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I; L \<le> K\<rbrakk> \<Longrightarrow> \<Turnstile>\<^sub>L \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I"
(*<*)
apply (unfold validn_def)
apply (rule, rule, rule)
apply (rule, rule, rule)
apply (rule, rule, rule)
apply (rule, rule, rule)
apply (rule, rule, rule)
apply rule
apply (erule_tac x=M in allE, erule impE, assumption)
apply (erule_tac x=Mspec in allE, erule_tac x=Minv in allE, erule_tac x=Anno in allE, erule impE, assumption)
apply (erule_tac x=par in allE)
apply (erule_tac x=code in allE)
apply (erule_tac x=l0 in allE, erule impE, assumption)
apply (erule_tac x=s0 in allE)
apply (erule_tac x=s in allE, erule impE, assumption, erule impE, assumption)
apply (erule_tac x=k in allE, erule impE, simp)
apply assumption
done
(*>*)
lemma G_valid_validn: "G_valid G \<Longrightarrow> G_validn K G"
(*<*)
apply (simp add: G_valid_def G_validn_def, clarsimp)
apply (rule valid_validn) apply fast
done
(*>*)
lemma G_validn_valid:"\<lbrakk>\<forall> K . G_validn K G\<rbrakk> \<Longrightarrow> G_valid G"
(*<*)
apply (simp add: G_valid_def G_validn_def, clarsimp)
apply (rule validn_valid) apply clarsimp
done
(*>*)
lemma G_validn_lower: "\<lbrakk>G_validn K G; L \<le> K\<rbrakk> \<Longrightarrow> G_validn L G"
(*<*)
apply (simp add: G_validn_def, clarsimp)
apply (rule validn_lower) apply fast apply assumption
done
(*>*)
lemma MST_validn_valid:"\<lbrakk>\<forall> K. MST_validn K\<rbrakk> \<Longrightarrow> MST_valid"
(*<*)
apply (simp add: MST_validn_def MST_valid_def, clarsimp)
apply (rule validn_valid, clarsimp)
done
(*>*)
lemma MST_valid_validn:"MST_valid \<Longrightarrow> MST_validn K"
(*<*)
apply (unfold MST_validn_def MST_valid_def)
apply (rule, rule, rule)
apply (rule, rule, rule)
apply (rule, rule, rule)
apply rule
apply (rule valid_validn)
apply fast
done
(*>*)
lemma MST_validn_lower: "\<lbrakk>MST_validn K; L \<le> K\<rbrakk> \<Longrightarrow> MST_validn L"
(*<*)
apply (simp add: MST_validn_def, clarsimp)
apply (erule_tac x=C in allE)
apply (erule_tac x=m in allE)
apply (erule_tac x=par in allE)
apply (erule_tac x=code in allE)
apply (erule_tac x=l0 in allE, erule impE, assumption)
apply (erule_tac x=T in allE)
apply (erule_tac x=MI in allE, clarsimp)
apply (erule validn_lower) apply assumption
done
(*>*)
text\<open>We define an abbreviation for the side conditions of the rule
for static method invocations\ldots\<close>
definition INVS_SC::
"Class \<Rightarrow> Method \<Rightarrow> Label \<Rightarrow> Class \<Rightarrow> Method \<Rightarrow> MethSpec \<Rightarrow> MethInv \<Rightarrow>
ANNO \<Rightarrow> ANNO \<Rightarrow> Mbody \<Rightarrow> Assn \<Rightarrow> Inv \<Rightarrow> bool" where
"INVS_SC C m l D m' T MI Anno Anno2 M' A I = (\<exists> M par code l0 T1 MI1.
mbody_is C m M \<and> get_ins M l = Some (invokeS D m') \<and>
MST\<down>(C,m) = Some (T1,MI1,Anno) \<and>
MST\<down>(D,m') = Some (T,MI,Anno2) \<and>
mbody_is D m' M' \<and> M'=(par,code,l0) \<and>
(\<forall> Q . Anno\<down>(l) = Some Q \<longrightarrow> (\<forall> s0 s . A s0 s \<longrightarrow> Q s0 s)) \<and>
(\<forall> s0 s . A s0 s \<longrightarrow> I s0 s (heap s)) \<and>
(\<forall> s0 ops1 ops2 S R h t . (ops1,par,R,ops2) : Frame \<longrightarrow>
A s0 (ops1,S,h) \<longrightarrow> MI (R,h) t \<longrightarrow> I s0 (ops1,S,h) t))"
text\<open>\ldots and another abbreviation for the soundness property of
the same rule.\<close>
definition INVS_soundK::
"nat \<Rightarrow> CTXT \<Rightarrow> Class \<Rightarrow> Method \<Rightarrow> Label \<Rightarrow> Class \<Rightarrow> Method \<Rightarrow>
MethSpec \<Rightarrow> MethInv \<Rightarrow> ANNO \<Rightarrow> ANNO \<Rightarrow> Mbody \<Rightarrow> Assn \<Rightarrow>
Post \<Rightarrow> Inv \<Rightarrow> bool" where
"INVS_soundK K G C m l D m' T MI Anno Anno2 M' A B I =
(INVS_SC C m l D m' T MI Anno Anno2 M' A I \<longrightarrow>
G_validn K G \<longrightarrow> MST_validn K \<longrightarrow>
\<Turnstile>\<^sub>K \<lbrace>(SINV_pre (fst M') T A)\<rbrace> C,m,(l+1)
\<lbrace>(SINV_post (fst M') T B)\<rbrace> (SINV_inv (fst M') T I)
\<longrightarrow> \<Turnstile>\<^sub>(K+1) \<lbrace> A \<rbrace> C,m,l \<lbrace> B \<rbrace> I)"
text\<open>The proof that this property holds for all $K$ proceeds by
induction on $K$.\<close>
lemma INVS_soundK_all:
"INVS_soundK K G C m l D m' T MI Anno Anno2 M' A B I"
(*<*)
+supply [[simproc del: defined_all]]
apply (induct K)
(*K=0*)
apply (simp add: INVS_soundK_def , clarsimp)
apply (unfold validn_def)
apply (rule, rule) apply (erule_tac x=M in allE, erule impE, assumption)
apply (rule, rule, rule, rule)
apply (erule_tac x=Mspec in allE, erule_tac x=Minv in allE, erule_tac x=Annoa in allE, erule impE, assumption)
apply (rule, rule, rule, rule)
apply (erule_tac x=par in allE, erule_tac x=code in allE, erule_tac x=l0 in allE, erule impE, assumption)
apply (rule, rule, rule)
apply (rule, rule, rule)
apply rule
apply (rule, rule, rule) apply (case_tac k, clarsimp) apply (drule no_zero_height_Exec_derivs, simp) apply clarsimp
apply (erule eval_cases) apply (simp add: INVS_SC_def mbody_is_def,clarsimp) apply clarsimp
apply (drule no_zero_height_Exec_derivs, simp)
apply (rule, rule, rule, rule)
apply clarsimp
apply (case_tac k, clarsimp) apply (drule ZeroHeightMultiElim, clarsimp)
apply (simp add: INVS_SC_def mbody_is_def,clarsimp)
apply clarsimp
apply (drule MultiSplit, simp, clarsimp) apply (drule no_zero_height_Step_derivs, simp)
apply rule apply (case_tac k, clarsimp) apply (drule ZeroHeightReachableElim, clarsimp)
apply (simp add: INVS_SC_def mbody_is_def,clarsimp)
apply clarsimp apply (drule ReachableSplit, simp, clarsimp)
apply (erule disjE)
apply clarsimp apply (drule no_zero_height_Step_derivs, simp)
apply clarsimp apply (drule ZeroHeightReachableElim, clarsimp)
apply (rotate_tac 5, erule thin_rl)
apply (simp add: INVS_SC_def mbody_is_def,clarsimp)
apply (simp add: MST_validn_def)
apply (erule_tac x=D in allE, rotate_tac -1)
apply (erule_tac x=m' in allE, rotate_tac -1)
apply (erule_tac x=para in allE, rotate_tac -1)
apply (erule_tac x=codea in allE, rotate_tac -1)
apply (erule_tac x=l0a in allE, rotate_tac -1,erule impE) apply (simp add: mbody_is_def)
apply (rotate_tac -1, erule_tac x=T in allE)
apply (rotate_tac -1, erule_tac x=MI in allE, clarsimp)
apply (simp add: validn_def)
apply (simp add: mbody_is_def)
apply (simp add: heap_def)
(*k>0*)
apply (simp add: INVS_soundK_def , clarsimp)
apply (rotate_tac -1, erule thin_rl)
apply (unfold validn_def)
apply (rule, rule) apply (erule_tac x=M in allE, erule impE, assumption)
apply (rule, rule) apply (rule, rule) apply (rule, rule) apply (rule, rule) apply (rule, rule)
apply (erule_tac x=Mspec in allE, erule_tac x=Minv in allE, erule_tac x=Annoa in allE, erule impE, assumption)
apply (erule_tac x=par in allE, erule_tac x=code in allE)
apply (erule_tac x=l0 in allE, erule impE, assumption)
apply (rule, rule)
apply (rule, rule)
apply (erule_tac x=s0 in allE)
apply rule
(*Exec*)
apply (rule, rule, rule)
apply (erule eval_cases) apply clarsimp apply (simp add: INVS_SC_def mbody_is_def) apply clarsimp
apply (erule_tac x=t in allE, erule impE)
apply (frule InvokeElim) apply (simp add: INVS_SC_def mbody_is_def) apply clarsimp apply fastforce
apply (simp add: MS_def, erule exE, rule) apply clarsimp apply (erule MultiApp) apply assumption
apply (erule impE, drule InvokeElim, simp add: INVS_SC_def mbody_is_def) apply clarsimp apply fastforce
apply clarsimp apply (simp add: INVS_SC_def mbody_is_def SINV_pre_def) apply clarsimp apply (rule, rule)
apply (rule, rule, rule, assumption) apply (rule, rule) prefer 2 apply (rule, assumption) apply simp
apply (simp add: MST_validn_def) apply (erule_tac x=D in allE, erule_tac x=m' in allE)
apply (erule_tac x=parb in allE, erule_tac x=codeb in allE)
apply (erule_tac x=l0b in allE, simp add: mbody_is_def)
apply (simp add: validn_def)
apply (erule_tac x=parb in allE, erule_tac x=codeb in allE)
apply (erule_tac x=l0b in allE, simp add: mbody_is_def)
apply (rotate_tac -1, erule_tac x=R in allE, erule_tac x=bb in allE)
apply (rotate_tac -1, erule_tac x="[]" in allE, rotate_tac -1, erule_tac x=R in allE, erule_tac x=bb in allE)
apply (simp add: mkState_def)
apply (erule impE) apply (simp add: MS_def, rule) apply (rule MS_zero) apply simp apply simp apply simp
apply (erule_tac x=ka in allE, erule impE, simp)
apply (simp add: mkPost_def, erule conjE)
apply (erule_tac x=hh in allE, erule_tac x=va in allE, simp add: mkState_def)
apply (erule_tac x=ma in allE, erule impE) apply (simp add: max_def) apply (case_tac "n \<le> ma") apply clarsimp apply clarsimp
apply (erule conjE) apply (erule_tac x=h in allE, erule_tac x=v in allE, erule impE)
apply (drule InvokeElim) apply (simp add: INVS_SC_def mbody_is_def) apply clarsimp apply fastforce
apply clarsimp
apply (simp add: SINV_post_def) apply (simp add: INVS_SC_def mbody_is_def, clarsimp)
apply (drule InvokeElim, clarsimp) apply fastforce apply clarsimp
apply (simp add: mbody_is_def, clarsimp)
apply (erule_tac x=ac in allE, erule_tac x=ops in allE, rotate_tac -1)
apply (erule_tac x=ad in allE, erule_tac x=R in allE, rotate_tac -1, erule impE, assumption)
apply (erule_tac x=bb in allE, simp, erule mp)
apply (simp add: MST_validn_def)
apply (erule_tac x=D in allE, erule_tac x=m' in allE, rotate_tac -1)
apply (erule_tac x=par in allE, erule_tac x=code in allE, erule_tac x=l0 in allE, simp add: mbody_is_def)
apply (simp add: validn_def)
apply (rotate_tac -1) apply (erule thin_rl) apply (rotate_tac -1)
apply (simp add: mbody_is_def)
apply (erule_tac x=R in allE, rotate_tac -1, erule_tac x=bb in allE, rotate_tac -1)
apply (erule_tac x="[]" in allE, rotate_tac -1, erule_tac x=R in allE, rotate_tac -1)
apply (erule_tac x=bb in allE, erule impE)
apply (simp add: MS_def, rule, rule MS_zero) apply (simp, simp add: mkState_def, simp)
apply (simp add: mkState_def)
apply (erule_tac x=k in allE, clarsimp)
apply (erule_tac x=bc in allE, erule_tac x=va in allE, simp add: mkPost_def mkState_def)
(*MStep*)
apply (rule, rule)
apply (rule, rule)
apply (rule, rule)
apply (case_tac k, clarsimp) apply (drule ZeroHeightMultiElim, clarsimp) apply (simp add: INVS_SC_def) apply clarsimp
apply clarsimp
apply (frule MultiSplit) apply clarsimp apply clarsimp
apply (frule InvokeElim) apply (simp add: INVS_SC_def mbody_is_def) apply clarsimp apply fastforce
apply clarsimp
apply (erule_tac x="v # ops" in allE, erule_tac x=ad in allE, erule_tac x=b in allE, erule impE)
apply (simp add: MS_def, erule exE, rule) apply (erule MultiApp) apply assumption
apply (erule impE, simp add: SINV_pre_def INVS_SC_def mbody_is_def) apply clarsimp
apply (rule, rule, rule, rule, rule, assumption) apply (rule, rule)
prefer 2 apply (rule, assumption) apply simp
apply (simp add: MST_validn_def) apply (erule_tac x=D in allE, erule_tac x=m' in allE)
apply (erule_tac x=parb in allE, erule_tac x=codeb in allE)
apply (erule_tac x=l0b in allE, simp add: mbody_is_def)
apply (simp add: validn_def)
apply (erule_tac x=parb in allE, erule_tac x=codeb in allE)
apply (erule_tac x=l0b in allE, simp add: mbody_is_def)
apply (rotate_tac -1, erule_tac x=R in allE, erule_tac x=bb in allE)
apply (rotate_tac -1, erule_tac x="[]" in allE, rotate_tac -1, erule_tac x=R in allE, erule_tac x=bb in allE)
apply (simp add: mkState_def)
apply (erule impE) apply (simp add: MS_def, rule) apply (rule MS_zero) apply simp apply simp apply simp
apply (erule_tac x=k in allE, erule impE, simp)
apply (simp add: mkPost_def, erule conjE)
apply (erule_tac x=hh in allE, erule_tac x=va in allE, simp add: mkState_def)
apply (erule_tac x=ma in allE, erule impE) apply simp
apply (erule conjE) apply (erule_tac x=ll in allE, erule_tac x=ae in allE)
apply (erule_tac x=af in allE, rotate_tac -1, erule_tac x=bc in allE, clarsimp)
(*Reach*)
apply rule
apply (case_tac k, clarsimp) apply (drule ZeroHeightReachableElim, clarsimp)
apply (simp add: INVS_SC_def mbody_is_def, clarsimp)
apply clarsimp
apply (drule ReachableSplit) apply simp apply clarsimp
apply (erule disjE)
apply clarsimp
apply (frule InvokeElim) apply (simp add: INVS_SC_def mbody_is_def) apply clarsimp apply fastforce
apply clarsimp
apply (erule_tac x="v#ops" in allE, erule_tac x=ad in allE, erule_tac x=b in allE, erule impE)
apply (simp add: MS_def, clarsimp, rule) apply (erule MultiApp) apply assumption
apply (erule impE) apply (simp add: SINV_pre_def) apply (simp add: INVS_SC_def mbody_is_def, clarsimp)
apply (rule, rule, rule, rule, rule) apply assumption
apply (rule, rule) prefer 2 apply (rule, assumption, simp)
apply (simp add: MST_validn_def)
apply (erule_tac x=D in allE, erule_tac x=m' in allE, rotate_tac -1)
apply (erule_tac x=parb in allE, erule_tac x=codeb in allE, erule_tac x=l0b in allE, simp add: mbody_is_def)
apply (simp add: validn_def)
apply (simp add: mbody_is_def) apply (rotate_tac -1)
apply (erule_tac x=R in allE, rotate_tac -1, erule_tac x=bb in allE, rotate_tac -1)
apply (erule_tac x="[]" in allE, rotate_tac -1, erule_tac x=R in allE, rotate_tac -1)
apply (erule_tac x=bb in allE, erule impE)
apply (simp add: MS_def, rule, rule MS_zero) apply (simp, simp add: mkState_def, simp)
apply (simp add: mkState_def)
apply (erule_tac x=k in allE, clarsimp)
apply (erule_tac x=b in allE, erule_tac x=v in allE, simp add: mkPost_def mkState_def)
apply (erule_tac x=ma in allE, clarsimp)
apply (rotate_tac -1)
apply (simp add: SINV_inv_def)
apply (erule_tac x="l+1" in allE, erule_tac x=ae in allE)
apply (erule_tac x=af in allE, rotate_tac -1, erule_tac x=bc in allE, clarsimp)
apply (rotate_tac -2, erule thin_rl)
apply (simp add: SINV_inv_def INVS_SC_def mbody_is_def) apply clarsimp
apply (rotate_tac -1, erule_tac x=ac in allE)
apply (rotate_tac -1, erule_tac x=ops in allE)
apply (rotate_tac -1, erule_tac x=ad in allE)
apply (rotate_tac -1, erule_tac x=R in allE, clarsimp)
apply (rotate_tac -1, erule_tac x=bb in allE, erule mp)
apply (erule thin_rl)
apply (simp add: MST_validn_def)
apply (erule_tac x=D in allE, erule_tac x=m' in allE, rotate_tac -1)
apply (erule_tac x=parb in allE, erule_tac x=codeb in allE, erule_tac x=l0b in allE, simp add: mbody_is_def)
apply (simp add: validn_def)
apply (simp add: mbody_is_def) apply (rotate_tac -1)
apply (erule_tac x=R in allE, rotate_tac -1, erule_tac x=bb in allE, rotate_tac -1)
apply (erule_tac x="[]" in allE, rotate_tac -1, erule_tac x=R in allE, rotate_tac -1)
apply (erule_tac x=bb in allE, erule impE)
apply (simp add: MS_def, rule, rule MS_zero) apply (simp, simp add: mkState_def, simp)
apply (simp add: mkState_def)
apply (erule_tac x=k in allE, clarsimp)
apply (erule_tac x=b in allE, erule_tac x=v in allE, simp add: mkPost_def mkState_def)
apply clarsimp
apply (simp add: INVS_SC_def mbody_is_def, clarsimp)
apply (rotate_tac -1, erule_tac x=ab in allE)
apply (rotate_tac -1, erule_tac x=ba in allE)
apply (rotate_tac -1, erule_tac x=ac in allE)
apply (rotate_tac -1, erule_tac x=ops1 in allE)
apply (rotate_tac -1, erule_tac x=ad in allE)
apply (rotate_tac -1, erule_tac x=R in allE, clarsimp)
apply (rotate_tac -1, erule_tac x=bb in allE, clarsimp)
apply (rotate_tac -1, erule_tac x="heap (ae,af,bc)" in allE, erule mp)
apply (rotate_tac 6) apply (erule thin_rl)
apply (simp add: MST_validn_def)
apply (erule_tac x=D in allE, erule_tac x=m' in allE)
apply (simp add: mbody_is_def)
apply (simp add: validn_def)
apply (simp add: mbody_is_def mkInv_def)
apply (rotate_tac -1)
apply (erule_tac x=R in allE, erule_tac x=bb in allE)
apply (rotate_tac -1) apply (erule_tac x="[]" in allE)
apply (rotate_tac -1)
apply (erule_tac x=R in allE, erule_tac x=bb in allE, simp add: mkState_def, erule impE)
apply (simp add: MS_def, rule, rule MS_zero) apply (simp, simp, simp)
apply (erule_tac x=nat in allE, simp)
done
(*>*)
text\<open>The heart of the soundness proof - the induction on the
axiomatic semantics.\<close>
lemma SOUND_Aux[rule_format]:
"(b,G,C,m,l,A,B,I):SP_Judgement \<Longrightarrow> G_validn K G \<longrightarrow> MST_validn K \<longrightarrow>
((b \<longrightarrow> \<Turnstile>\<^sub>K \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I) \<and>
((\<not> b) \<longrightarrow> \<Turnstile>\<^sub>(Suc K) \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I))"
(*<*)
+supply [[simproc del: defined_all]]
apply (erule SP_Judgement.induct)
(*INSTR*)
apply clarsimp
apply (rotate_tac -4) apply (erule thin_rl)
apply (simp add: mbody_is_def validn_def Opsem_def MS_def Reach_def, clarsimp)
apply rule
apply clarsimp apply (erule eval_cases) apply simp
apply clarsimp apply (frule InstrElimNext, simp, assumption, clarsimp)
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=aa in allE)
apply (erule_tac x=b in allE, erule impE)
apply rule
apply (erule MultiApp)
apply assumption
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule, rule, assumption)
apply (rule, rule,assumption)
apply (erule_tac x=ma in allE, clarsimp)
apply (erule_tac x=h in allE, erule_tac x=v in allE, clarsimp)
apply (simp add: SP_post_def)
apply (rotate_tac -1)
apply (erule_tac x=ae in allE, rotate_tac -1)
apply (erule_tac x=af in allE, rotate_tac -1)
apply (erule_tac x=bc in allE, erule mp)
apply (rule, rule, assumption)
apply clarsimp
apply rule
apply clarsimp
apply (case_tac ka) apply clarsimp apply (drule ZeroHeightMultiElim) apply clarsimp
apply clarsimp
apply (rotate_tac -2, drule MultiSplit) apply simp apply clarsimp
apply (frule InstrElimNext, simp, assumption, clarsimp)
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=ag in allE, rotate_tac -1)
apply (erule_tac x=ah in allE)
apply (erule_tac x=bd in allE, erule impE)
apply rule
apply (erule MultiApp)
apply assumption
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule, rule, assumption)
apply (rule, rule, assumption)
apply (erule_tac x=ma in allE, clarsimp)
apply (erule_tac x=ll in allE, rotate_tac -1)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=aa in allE, rotate_tac -1)
apply (erule_tac x=b in allE, clarsimp)
apply clarsimp
apply (case_tac ka) apply clarsimp apply (drule ZeroHeightReachableElim) apply clarsimp
apply clarsimp
apply (rotate_tac -2, drule ReachableSplit) apply simp apply clarsimp
apply (rotate_tac -1, erule disjE)
apply clarsimp
apply (frule InstrElimNext, simp, assumption, clarsimp)
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=ag in allE, rotate_tac -1)
apply (erule_tac x=ah in allE)
apply (erule_tac x=bd in allE, erule impE)
apply rule
apply (erule MultiApp)
apply assumption
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule, rule, assumption)
apply (rule, rule, assumption)
apply (erule_tac x=ma in allE, clarsimp)
apply (rotate_tac -1, erule_tac x=ll in allE)
apply (rotate_tac -1, erule_tac x=a in allE)
apply (rotate_tac -1, erule_tac x=aa in allE)
apply (rotate_tac -1, erule_tac x=b in allE, clarsimp)
apply (simp add: SP_inv_def)
(* apply (rotate_tac -1, erule_tac x="l+1" in allE)*)
apply (rotate_tac -1, erule_tac x=ae in allE)
apply (rotate_tac -1, erule_tac x=af in allE)
apply (rotate_tac -1, erule_tac x=bc in allE, erule mp) apply (rule, rule, assumption)
apply clarsimp
(*GOTO*)
apply clarsimp
apply (rotate_tac 5) apply (erule thin_rl)
apply (simp add: mbody_is_def validn_def Opsem_def MS_def Reach_def, clarsimp)
apply rule
apply clarsimp apply (erule eval_cases) apply simp
apply clarsimp apply (drule GotoElim1) apply (simp, clarsimp)
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=ae in allE, rotate_tac- 1)
apply (erule_tac x=af in allE)
apply (erule_tac x=bc in allE, erule impE)
apply rule
apply (erule MultiApp)
apply (erule Goto)
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule, rule, assumption)
apply (rule, rule, erule Goto)
apply (erule_tac x=ma in allE, clarsimp)
apply (erule_tac x=h in allE, erule_tac x=v in allE, clarsimp)
apply(simp add: SP_post_def, rotate_tac -1)
apply (erule_tac x=ae in allE, rotate_tac -1)
apply (erule_tac x=af in allE, rotate_tac -1)
apply (erule_tac x=bc in allE, erule mp)
apply (rule, rule, erule Goto)
apply clarsimp
apply rule
apply clarsimp
apply (case_tac ka) apply clarsimp apply (drule ZeroHeightMultiElim) apply clarsimp
apply clarsimp
apply (rotate_tac -2, drule MultiSplit) apply simp apply clarsimp
apply (drule GotoElim1) apply simp apply clarsimp
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=ae in allE,rotate_tac -1)
apply (erule_tac x=af in allE)
apply (erule_tac x=bc in allE, erule impE)
apply rule
apply (erule MultiApp)
apply (erule Goto)
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule, rule, assumption)
apply (rule, rule, erule Goto)
apply (erule_tac x=ma in allE, clarsimp)
apply (erule_tac x=ll in allE, rotate_tac -1)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=aa in allE)
apply (erule_tac x=b in allE, clarsimp)
apply clarsimp
apply (case_tac ka) apply clarsimp apply (drule ZeroHeightReachableElim) apply clarsimp
apply clarsimp
apply (rotate_tac -2, drule ReachableSplit) apply simp apply clarsimp
apply (drule GotoElim1) apply simp apply clarsimp
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=ae in allE,rotate_tac -1)
apply (erule_tac x=af in allE)
apply (erule_tac x=bc in allE, erule impE)
apply rule
apply (erule MultiApp)
apply (erule Goto)
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule, rule, assumption)
apply (rule, rule, erule Goto)
apply (erule_tac x=ma in allE, clarsimp)
apply (rotate_tac -1, erule_tac x=ll in allE)
apply (rotate_tac -1, erule_tac x=a in allE)
apply (rotate_tac -1, erule_tac x=aa in allE)
apply (rotate_tac -1, erule_tac x=b in allE, clarsimp)
apply (simp add: SP_inv_def)
(* apply (rotate_tac -1, erule_tac x=pc in allE)*)
apply (rotate_tac -1, erule_tac x=ae in allE)
apply (rotate_tac -1, erule_tac x=af in allE)
apply (rotate_tac -1, erule_tac x=bc in allE, erule mp)
apply (rule, rule)
apply (erule Goto)
(*If*)
apply clarsimp
apply (rotate_tac 5) apply (erule thin_rl,erule thin_rl)
apply (simp add: mbody_is_def validn_def Opsem_def MS_def Reach_def, clarsimp)
apply rule
apply clarsimp apply (erule eval_cases) apply simp
apply clarsimp apply (drule IfElim1) apply (simp, clarsimp)
apply (erule disjE)
apply clarsimp
apply (rotate_tac 3) apply (erule thin_rl)
apply (rotate_tac -1)
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=af in allE)
apply (erule_tac x=bc in allE, erule impE)
apply rule
apply (erule MultiApp)
apply (erule IfT, simp)
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule, rule) prefer 2
apply (rule, rule, erule IfT, simp) apply fastforce
apply (erule_tac x=ma in allE, clarsimp)
apply (erule_tac x=h in allE, erule_tac x=v in allE, clarsimp)
apply (simp add: SP_post_def)
apply (rotate_tac -1)
apply (erule_tac x="TRUE # a" in allE, erule_tac x=af in allE, erule_tac x=bc in allE, erule impE)
apply (rule, rule, rule IfT, simp,simp)
apply clarsimp
apply clarsimp
apply (rotate_tac 2) apply (erule thin_rl)
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=af in allE)
apply (erule_tac x=bc in allE, erule impE)
apply rule
apply (erule MultiApp)
apply (rule IfF) apply (simp, assumption, simp) apply simp
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule, rule) prefer 2
apply (rule, rule, rule IfF) apply (simp, assumption) apply fastforce
apply simp
apply clarsimp
apply (erule_tac x=ma in allE, clarsimp)
apply (erule_tac x=h in allE, erule_tac x=v in allE, clarsimp)
apply (simp add: SP_post_def)
apply (rotate_tac -1)
apply (erule_tac x="va # a" in allE, erule_tac x=af in allE, erule_tac x=bc in allE, erule impE)
apply (rule, rule, rule IfF) apply (simp, assumption) apply fastforce apply simp
apply clarsimp
apply clarsimp
apply rule
apply clarsimp
apply (case_tac ka) apply clarsimp apply (drule ZeroHeightMultiElim) apply clarsimp
apply clarsimp
apply (rotate_tac -2, drule MultiSplit) apply simp apply clarsimp
apply (drule IfElim1) apply simp apply clarsimp
apply (erule disjE)
apply clarsimp
apply (rotate_tac 4) apply (erule thin_rl)
apply (rotate_tac -1)
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=ag in allE, rotate_tac -1)
apply (erule_tac x=af in allE)
apply (erule_tac x=bc in allE, erule impE)
apply rule
apply (erule MultiApp)
apply (rule IfT) apply (simp, fastforce)
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule,rule) prefer 2
apply (rule, rule, rule IfT) apply simp apply fastforce
apply clarsimp
apply (erule_tac x=ma in allE, clarsimp)
apply (erule_tac x=ll in allE, rotate_tac -1)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=aa in allE, rotate_tac -1)
apply (erule_tac x=b in allE, rotate_tac -1, clarsimp)
apply clarsimp
apply (rotate_tac 3) apply (erule thin_rl)
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=ag in allE, rotate_tac -1)
apply (erule_tac x=af in allE)
apply (erule_tac x=bc in allE, erule impE)
apply rule
apply (erule MultiApp)
apply (rule IfF) apply (simp, assumption, simp) apply simp
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule, rule) prefer 2
apply (rule, rule, rule IfF) apply (simp, assumption) apply fastforce apply simp
apply clarsimp
apply (erule_tac x=ma in allE, clarsimp)
apply (erule_tac x=ll in allE, rotate_tac -1)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=aa in allE, rotate_tac -1)
apply (erule_tac x=b in allE, rotate_tac -1, clarsimp)
apply clarsimp
apply (case_tac ka) apply clarsimp apply (drule ZeroHeightReachableElim) apply clarsimp
apply clarsimp
apply (rotate_tac -2, drule ReachableSplit) apply simp apply clarsimp
apply (drule IfElim1) apply simp apply clarsimp
apply (erule disjE)
apply clarsimp
apply (rotate_tac 4) apply (erule thin_rl)
apply (rotate_tac -1)
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=ag in allE, rotate_tac -1)
apply (erule_tac x=af in allE)
apply (erule_tac x=bc in allE, erule impE)
apply rule
apply (erule MultiApp)
apply (rule IfT) apply (simp, fastforce)
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule, rule) prefer 2
apply (rule, rule, rule IfT) apply simp apply fastforce
apply clarsimp
apply (erule_tac x=ma in allE, clarsimp)
apply (rotate_tac -1, erule_tac x=ll in allE)
apply (rotate_tac -1, erule_tac x=a in allE)
apply (rotate_tac -1, erule_tac x=aa in allE)
apply (rotate_tac -1, erule_tac x=b in allE, clarsimp)
apply (simp add: SP_inv_def)
(* apply (rotate_tac -1, erule_tac x=pc in allE)*)
apply (rotate_tac -1, erule_tac x="TRUE#ag" in allE)
apply (rotate_tac -1, erule_tac x=af in allE)
apply (rotate_tac -1, erule_tac x=bc in allE)
apply (rotate_tac -1, erule impE)
apply (rule, rule, erule IfT) apply simp
apply clarsimp
apply clarsimp
apply (rotate_tac 3) apply (erule thin_rl)
apply (erule_tac x=ad in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=ag in allE, rotate_tac -1)
apply (erule_tac x=af in allE)
apply (erule_tac x=bc in allE, erule impE)
apply rule
apply (erule MultiApp)
apply (erule IfF) apply (assumption, simp, simp)
apply (erule impE, simp add: SP_pre_def) apply (rule, rule, rule, rule) prefer 2
apply (rule, rule, erule IfF) apply (assumption, fastforce,simp)
apply clarsimp
apply (erule_tac x=ma in allE, clarsimp)
apply (rotate_tac -1, erule_tac x=ll in allE)
apply (rotate_tac -1, erule_tac x=a in allE)
apply (rotate_tac -1, erule_tac x=aa in allE)
apply (rotate_tac -1, erule_tac x=b in allE, clarsimp)
apply (simp add: SP_inv_def)
(* apply (rotate_tac -1, erule_tac x="Suc l" in allE)*)
apply (rotate_tac -1, erule_tac x="v#ag" in allE)
apply (rotate_tac -1, erule_tac x=af in allE)
apply (rotate_tac -1, erule_tac x=bc in allE)
apply (rotate_tac -1, erule impE)
apply (rule, rule)
apply (erule IfF) apply (assumption, simp, simp)
apply clarsimp
(*RET*)
apply clarsimp
apply (simp add: mbody_is_def validn_def Opsem_def MS_def Reach_def, clarsimp)
apply rule
apply clarsimp apply (erule eval_cases) apply clarsimp
apply (drule RetElim1) apply simp apply simp
apply clarsimp
apply rule
apply clarsimp
apply (case_tac ka) apply clarsimp apply (drule ZeroHeightMultiElim) apply clarsimp
apply clarsimp
apply (rotate_tac -2, drule MultiSplit) apply simp apply clarsimp
apply (drule RetElim1, clarsimp) apply simp
apply clarsimp
apply (case_tac ka) apply clarsimp apply (drule ZeroHeightReachableElim) apply clarsimp
apply clarsimp
apply (rotate_tac -2, drule ReachableSplit) apply simp apply clarsimp
apply (drule RetElim1, clarsimp) apply simp
(*INVS*)
apply clarsimp
apply (subgoal_tac "INVS_soundK K G C m l D m' T MI Anno Anno2 (par,code,l0) A B I")
apply (simp add: INVS_soundK_def)
apply (erule impE) apply (simp add: INVS_SC_def) apply (rule, rule, rule, rule) apply assumption apply (rule, assumption)
apply assumption
apply assumption
apply (rule INVS_soundK_all)
(*CONSEQ*)
apply clarsimp
apply rule
apply clarsimp
apply (simp add: validn_def,clarsimp)
apply (erule thin_rl)
apply (rotate_tac 5, erule thin_rl)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=aa in allE, rotate_tac -1)
apply (erule_tac x=ba in allE, clarsimp, rotate_tac -1)
apply (erule_tac x=ab in allE)
apply (erule_tac x=bb in allE)
apply (rotate_tac -1)
apply (erule_tac x=ac in allE, rotate_tac -1)
apply (erule_tac x=ad in allE, rotate_tac -1)
apply (erule_tac x=bc in allE, clarsimp)
apply (erule_tac x=k in allE, clarsimp)
apply (rotate_tac -1)
apply (erule_tac x=ll in allE, rotate_tac -1)
apply (erule_tac x=ae in allE, rotate_tac -1)
apply (erule_tac x=af in allE)
apply (erule_tac x=bd in allE, clarsimp)
apply clarsimp
apply (simp add: validn_def,clarsimp)
apply (erule thin_rl)
apply (rotate_tac 5, erule thin_rl)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=aa in allE, rotate_tac -1)
apply (erule_tac x=ba in allE, clarsimp, rotate_tac -1)
apply (erule_tac x=ab in allE)
apply (erule_tac x=bb in allE)
apply (rotate_tac -1)
apply (erule_tac x=ac in allE, rotate_tac -1)
apply (erule_tac x=ad in allE, rotate_tac -1)
apply (erule_tac x=bc in allE, clarsimp)
apply (erule_tac x=k in allE, clarsimp)
apply (rotate_tac -1)
apply (erule_tac x=ll in allE, rotate_tac -1)
apply (erule_tac x=ae in allE, rotate_tac -1)
apply (erule_tac x=af in allE)
apply (erule_tac x=bd in allE, clarsimp)
(*INJECT*)
apply clarsimp apply (rule validn_lower) apply fast apply simp
(*AX*)
apply clarsimp apply (simp add: G_validn_def)
done
(*>*)
text\<open>The statement of this lemma gives a semantic interpretation of
the two judgement forms, as \<open>SP_Assum\<close>-judgements enjoy validity
up to execution height $K$, while \<open>SP_Deriv\<close>-judgements are
valid up to level $K+1$.\<close>
(*<*)
lemma SOUND_K:
"\<lbrakk> G \<rhd> \<lbrace>A\<rbrace> C,m,l \<lbrace>B\<rbrace> I; G_validn K G ; MST_validn K\<rbrakk>
\<Longrightarrow> \<Turnstile>\<^sub>(Suc K) \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I"
apply (drule SOUND_Aux) apply assumption+ apply simp
done
(*>*)
text\<open>From this, we obtain a soundness result that still involves
context validity.\<close>
theorem SOUND_in_CTXT:
"\<lbrakk>G \<rhd> \<lbrace>A\<rbrace> C,m,l \<lbrace>B\<rbrace> I; G_valid G; MST_valid\<rbrakk> \<Longrightarrow> \<Turnstile> \<lbrace>A\<rbrace> C, m, l \<lbrace>B\<rbrace> I"
(*<*)
apply (rule validn_valid)
apply clarsimp
apply (rule validn_lower)
apply (erule SOUND_K)
prefer 3 apply (subgoal_tac "K \<le> Suc K", assumption) apply simp
apply (erule G_valid_validn)
apply (erule MST_valid_validn)
done
(*>*)
text\<open>We will now show that the two semantic assumptions can be replaced by
the verified-program property.\<close>
subsection\<open>Soundness of verified programs \label{SectionContextElimination}\<close>
text\<open>In order to obtain a soundness result that does not require
validity assumptions of the context or the specification table,
we show that the \<open>VP\<close> property implies context validity.
First, the elimination of contexts. By induction on
\<open>k\<close> we prove\<close>
lemma VPG_MSTn_Gn[rule_format]:
"VP_G G \<longrightarrow> MST_validn k \<longrightarrow> G_validn k G"
(*<*)
apply (induct k)
(*k=0*)
apply clarsimp
apply (simp add: VP_G_def, clarsimp)
apply (simp add: G_validn_def, clarsimp)
apply (simp add: validn_def) apply clarsimp
apply rule
apply clarsimp apply (drule no_zero_height_Exec_derivs) apply simp
apply clarsimp
apply rule
apply clarsimp apply (drule ZeroHeightMultiElim) apply clarsimp
apply (rotate_tac 2)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (erule_tac x=a in allE, erule_tac x=aa in allE, erule_tac x=b in allE, clarsimp)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (erule_tac x=l in allE, erule_tac x=A in allE, erule_tac x=B in allE, clarsimp)
apply (rule AssertionsImplyAnnoInvariants)
prefer 3 apply assumption
apply assumption+
apply clarsimp apply (drule ZeroHeightReachableElim) apply clarsimp
apply (rotate_tac 2)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (erule_tac x=a in allE, erule_tac x=aa in allE, erule_tac x=b in allE, clarsimp)
apply (erule_tac x=C in allE, erule_tac x=m in allE,
erule_tac x=l in allE, erule_tac x=A in allE, erule_tac x=B in allE, clarsimp)
apply (erule AssertionsImplyMethInvariants, assumption+)
(*k>0*)
apply clarsimp apply (simp add: VP_G_def)
apply (simp (no_asm) add: G_validn_def, clarsimp)
apply (subgoal_tac "MST_validn k", clarsimp)
apply (rule SOUND_K) apply fast
apply assumption
apply assumption
apply (erule MST_validn_lower) apply simp
done
(*>*)
text\<open>which implies\<close>
lemma VPG_MST_G: "\<lbrakk>VP_G G; MST_valid\<rbrakk> \<Longrightarrow> G_valid G"
(*<*)
apply (rule G_validn_valid, clarsimp)
apply (erule VPG_MSTn_Gn)
apply (erule MST_valid_validn)
done
(*>*)
text\<open>Next, the elimination of \<open>MST_valid\<close>. Again by induction on
\<open>k\<close>, we prove\<close>
lemma VPG_MSTn[rule_format]: "VP_G G \<longrightarrow> MST_validn k"
(*<*)
apply (induct k)
apply (simp add: MST_validn_def, clarsimp)
apply (simp add: validn_def, clarsimp)
apply rule
apply clarsimp apply (drule no_zero_height_Exec_derivs) apply simp
apply clarsimp
apply rule
apply clarsimp apply (drule ZeroHeightMultiElim) apply clarsimp
apply (simp add: VP_G_def, clarsimp, rotate_tac -1)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (erule_tac x=par in allE, erule_tac x=code in allE, erule_tac x=l0 in allE, clarsimp)
apply (rule AssertionsImplyAnnoInvariants) apply fast apply assumption+ apply simp
apply clarsimp apply (drule ZeroHeightReachableElim) apply clarsimp
apply (simp add: VP_G_def, clarsimp, rotate_tac -1)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (erule_tac x=par in allE, erule_tac x=code in allE, erule_tac x=l0 in allE, clarsimp)
apply (frule AssertionsImplyMethInvariants) apply assumption apply (simp add: mkState_def)
apply clarsimp
apply (frule VPG_MSTn_Gn, assumption)
apply (simp add: VP_G_def)
apply (simp (no_asm) add: MST_validn_def, clarsimp)
apply (rule SOUND_K)
apply (rotate_tac 3)
apply (erule_tac x=C in allE, erule_tac x=m in allE)
apply (erule_tac x=par in allE, erule_tac x=code in allE, erule_tac x=l0 in allE, clarsimp) apply assumption
apply assumption
apply assumption
done
(*>*)
text\<open>which yields\<close>
lemma VPG_MST:"VP_G G \<Longrightarrow> MST_valid"
(*<*)
apply (rule MST_validn_valid, clarsimp)
apply (erule VPG_MSTn)
done
(*>*)
text\<open>Combining these two results, and unfolding the definition of
program validity yields the final soundness result.\<close>
theorem VP_VALID: "VP \<Longrightarrow> Prog_valid"
(*<*)
apply (simp add: VP_def Prog_valid_def, clarsimp)
apply (frule VPG_MST, simp)
apply (drule VPG_MST_G, assumption) apply fast
done
(*>*)
(*<*)
text \<open>In particular, the $\mathit{VP}$ property implies that all
method specifications are honoured by their respective method
implementations.\<close>
theorem "VP \<Longrightarrow> MST_valid"
(*<*)
by (drule VP_VALID, simp add: Prog_valid_def)
(*>*)
end
(*>*)
diff --git a/thys/CRDT/Network.thy b/thys/CRDT/Network.thy
--- a/thys/CRDT/Network.thy
+++ b/thys/CRDT/Network.thy
@@ -1,586 +1,588 @@
(* Victor B. F. Gomes, University of Cambridge
Martin Kleppmann, University of Cambridge
Dominic P. Mulligan, University of Cambridge
Alastair R. Beresford, University of Cambridge
*)
section\<open>Axiomatic network models\<close>
text\<open>In this section we develop a formal definition of an \emph{asynchronous unreliable causal broadcast network}.
We choose this model because it satisfies the causal delivery requirements of many operation-based
CRDTs~\cite{Almeida:2015fc,Baquero:2014ed}. Moreover, it is suitable for use in decentralised settings,
as motivated in the introduction, since it does not require waiting for communication with
a central server or a quorum of nodes.\<close>
theory
Network
imports
Convergence
begin
subsection\<open>Node histories\<close>
text\<open>We model a distributed system as an unbounded number of communicating nodes.
We assume nothing about the communication pattern of nodes---we assume only that each node is
uniquely identified by a natural number, and that the flow of execution at each node consists
of a finite, totally ordered sequence of execution steps (events).
We call that sequence of events at node $i$ the \emph{history} of that node.
For convenience, we assume that every event or execution step is unique within a node's history.\<close>
locale node_histories =
fixes history :: "nat \<Rightarrow> 'evt list"
assumes histories_distinct [intro!, simp]: "distinct (history i)"
lemma (in node_histories) history_finite:
shows "finite (set (history i))"
by auto
definition (in node_histories) history_order :: "'evt \<Rightarrow> nat \<Rightarrow> 'evt \<Rightarrow> bool" ("_/ \<sqsubset>\<^sup>_/ _" [50,1000,50]50) where
"x \<sqsubset>\<^sup>i z \<equiv> \<exists>xs ys zs. xs@x#ys@z#zs = history i"
lemma (in node_histories) node_total_order_trans:
assumes "e1 \<sqsubset>\<^sup>i e2"
and "e2 \<sqsubset>\<^sup>i e3"
shows "e1 \<sqsubset>\<^sup>i e3"
proof -
obtain xs1 xs2 ys1 ys2 zs1 zs2 where *: "xs1 @ e1 # ys1 @ e2 # zs1 = history i"
"xs2 @ e2 # ys2 @ e3 # zs2 = history i"
using history_order_def assms by auto
hence "xs1 @ e1 # ys1 = xs2 \<and> zs1 = ys2 @ e3 # zs2"
by(rule_tac xs="history i" and ys="[e2]" in pre_suf_eq_distinct_list) auto
thus ?thesis
by(clarsimp simp: history_order_def) (metis "*"(2) append.assoc append_Cons)
qed
lemma (in node_histories) local_order_carrier_closed:
assumes "e1 \<sqsubset>\<^sup>i e2"
shows "{e1,e2} \<subseteq> set (history i)"
using assms by (clarsimp simp add: history_order_def)
(metis in_set_conv_decomp Un_iff Un_subset_iff insert_subset list.simps(15)
set_append set_subset_Cons)+
lemma (in node_histories) node_total_order_irrefl:
shows "\<not> (e \<sqsubset>\<^sup>i e)"
by(clarsimp simp add: history_order_def)
(metis Un_iff histories_distinct distinct_append distinct_set_notin
list.set_intros(1) set_append)
lemma (in node_histories) node_total_order_antisym:
assumes "e1 \<sqsubset>\<^sup>i e2"
and "e2 \<sqsubset>\<^sup>i e1"
shows "False"
using assms node_total_order_irrefl node_total_order_trans by blast
lemma (in node_histories) node_order_is_total:
assumes "e1 \<in> set (history i)"
and "e2 \<in> set (history i)"
and "e1 \<noteq> e2"
shows "e1 \<sqsubset>\<^sup>i e2 \<or> e2 \<sqsubset>\<^sup>i e1"
using assms unfolding history_order_def by(metis list_split_two_elems histories_distinct)
definition (in node_histories) prefix_of_node_history :: "'evt list \<Rightarrow> nat \<Rightarrow> bool" (infix "prefix of" 50) where
"xs prefix of i \<equiv> \<exists>ys. xs@ys = history i"
lemma (in node_histories) carriers_head_lt:
assumes "y#ys = history i"
shows "\<not>(x \<sqsubset>\<^sup>i y)"
using assms
apply(clarsimp simp add: history_order_def)
apply(rename_tac xs1 ys1 zs1)
apply (subgoal_tac "xs1 @ x # ys1 = [] \<and> zs1 = ys")
apply clarsimp
apply (rule_tac xs="history i" and ys="[y]" in pre_suf_eq_distinct_list)
apply auto
done
lemma (in node_histories) prefix_of_ConsD [dest]:
assumes "x # xs prefix of i"
shows "[x] prefix of i"
using assms by(auto simp: prefix_of_node_history_def)
lemma (in node_histories) prefix_of_appendD [dest]:
assumes "xs @ ys prefix of i"
shows "xs prefix of i"
using assms by(auto simp: prefix_of_node_history_def)
lemma (in node_histories) prefix_distinct:
assumes "xs prefix of i"
shows "distinct xs"
using assms by(clarsimp simp: prefix_of_node_history_def) (metis histories_distinct distinct_append)
lemma (in node_histories) prefix_to_carriers [intro]:
assumes "xs prefix of i"
shows "set xs \<subseteq> set (history i)"
using assms by(clarsimp simp: prefix_of_node_history_def) (metis Un_iff set_append)
lemma (in node_histories) prefix_elem_to_carriers:
assumes "xs prefix of i"
and "x \<in> set xs"
shows "x \<in> set (history i)"
using assms by(clarsimp simp: prefix_of_node_history_def) (metis Un_iff set_append)
lemma (in node_histories) local_order_prefix_closed:
assumes "x \<sqsubset>\<^sup>i y"
and "xs prefix of i"
and "y \<in> set xs"
shows "x \<in> set xs"
proof -
obtain ys where "xs @ ys = history i"
using assms prefix_of_node_history_def by blast
moreover obtain as bs cs where "as @ x # bs @ y # cs = history i"
using assms history_order_def by blast
moreover obtain pre suf where *: "xs = pre @ y # suf"
using assms split_list by fastforce
ultimately have "pre = as @ x # bs \<and> suf @ ys = cs"
by (rule_tac xs="history i" and ys="[y]" in pre_suf_eq_distinct_list) auto
thus ?thesis
using assms * by clarsimp
qed
lemma (in node_histories) local_order_prefix_closed_last:
assumes "x \<sqsubset>\<^sup>i y"
and "xs@[y] prefix of i"
shows "x \<in> set xs"
proof -
have "x \<in> set (xs @ [y])"
using assms by (force dest: local_order_prefix_closed)
thus ?thesis
using assms by(force simp add: node_total_order_irrefl prefix_to_carriers)
qed
lemma (in node_histories) events_before_exist:
assumes "x \<in> set (history i)"
shows "\<exists>pre. pre @ [x] prefix of i"
proof -
have "\<exists>idx. idx < length (history i) \<and> (history i) ! idx = x"
using assms by(simp add: set_elem_nth)
thus ?thesis
by(metis append_take_drop_id take_Suc_conv_app_nth prefix_of_node_history_def)
qed
lemma (in node_histories) events_in_local_order:
assumes "pre @ [e2] prefix of i"
and "e1 \<in> set pre"
shows "e1 \<sqsubset>\<^sup>i e2"
using assms split_list unfolding history_order_def prefix_of_node_history_def by fastforce
subsection\<open>Asynchronous broadcast networks\<close>
text\<open>We define a new locale $\isa{network}$ containing three axioms that define how broadcast
and deliver events may interact, with these axioms defining the properties of our network model.\<close>
datatype 'msg event
= Broadcast 'msg
| Deliver 'msg
locale network = node_histories history for history :: "nat \<Rightarrow> 'msg event list" +
fixes msg_id :: "'msg \<Rightarrow> 'msgid"
(* Broadcast/Deliver interaction *)
assumes delivery_has_a_cause: "\<lbrakk> Deliver m \<in> set (history i) \<rbrakk> \<Longrightarrow>
\<exists>j. Broadcast m \<in> set (history j)"
and deliver_locally: "\<lbrakk> Broadcast m \<in> set (history i) \<rbrakk> \<Longrightarrow>
Broadcast m \<sqsubset>\<^sup>i Deliver m"
and msg_id_unique: "\<lbrakk> Broadcast m1 \<in> set (history i);
Broadcast m2 \<in> set (history j);
msg_id m1 = msg_id m2 \<rbrakk> \<Longrightarrow> i = j \<and> m1 = m2"
text\<open>
The axioms can be understood as follows:
\begin{description}
\item[delivery-has-a-cause:] If some message $\isa{m}$ was delivered at some node, then there exists some node on which $\isa{m}$ was broadcast.
With this axiom, we assert that messages are not created ``out of thin air'' by the network itself, and that the only source of messages are the nodes.
\item[deliver-locally:] If a node broadcasts some message $\isa{m}$, then the same node must subsequently also deliver $\isa{m}$ to itself.
Since $\isa{m}$ does not actually travel over the network, this local delivery is always possible, even if the network is interrupted.
Local delivery may seem redundant, since the effect of the delivery could also be implemented by the broadcast event itself; however, it is standard practice in the description of broadcast protocols that the sender of a message also sends it to itself, since this property simplifies the definition of algorithms built on top of the broadcast abstraction \cite{Cachin:2011wt}.
\item[msg-id-unique:] We do not assume that the message type $\isacharprime\isa{msg}$ has any particular structure; we only assume the existence of a function $\isa{msg-id} \mathbin{\isacharcolon\isacharcolon} \isacharprime\isa{msg} \mathbin{\isasymRightarrow} \isacharprime\isa{msgid}$ that maps every message to some globally unique identifier of type $\isacharprime\isa{msgid}$.
We assert this uniqueness by stating that if $\isa{m1}$ and $\isa{m2}$ are any two messages broadcast by any two nodes, and their $\isa{msg-id}$s are the same, then they were in fact broadcast by the same node and the two messages are identical.
In practice, these globally unique IDs can by implemented using unique node identifiers, sequence numbers or timestamps.
\end{description}
\<close>
lemma (in network) broadcast_before_delivery:
assumes "Deliver m \<in> set (history i)"
shows "\<exists>j. Broadcast m \<sqsubset>\<^sup>j Deliver m"
using assms deliver_locally delivery_has_a_cause by blast
lemma (in network) broadcasts_unique:
assumes "i \<noteq> j"
and "Broadcast m \<in> set (history i)"
shows "Broadcast m \<notin> set (history j)"
using assms msg_id_unique by blast
text\<open>Based on the well-known definition by \cite{Lamport:1978jq}, we say that
$\isa{m1}\prec\isa{m2}$ if any of the following is true:
\begin{enumerate}
\item $\isa{m1}$ and $\isa{m2}$ were broadcast by the same node, and $\isa{m1}$ was broadcast before $\isa{m2}$.
\item The node that broadcast $\isa{m2}$ had delivered $\isa{m1}$ before it broadcast $\isa{m2}$.
\item There exists some operation $\isa{m3}$ such that $\isa{m1} \prec \isa{m3}$ and $\isa{m3} \prec \isa{m2}$.
\end{enumerate}\<close>
inductive (in network) hb :: "'msg \<Rightarrow> 'msg \<Rightarrow> bool" where
hb_broadcast: "\<lbrakk> Broadcast m1 \<sqsubset>\<^sup>i Broadcast m2 \<rbrakk> \<Longrightarrow> hb m1 m2" |
hb_deliver: "\<lbrakk> Deliver m1 \<sqsubset>\<^sup>i Broadcast m2 \<rbrakk> \<Longrightarrow> hb m1 m2" |
hb_trans: "\<lbrakk> hb m1 m2; hb m2 m3 \<rbrakk> \<Longrightarrow> hb m1 m3"
inductive_cases (in network) hb_elim: "hb x y"
definition (in network) weak_hb :: "'msg \<Rightarrow> 'msg \<Rightarrow> bool" where
"weak_hb m1 m2 \<equiv> hb m1 m2 \<or> m1 = m2"
locale causal_network = network +
assumes causal_delivery: "Deliver m2 \<in> set (history j) \<Longrightarrow> hb m1 m2 \<Longrightarrow> Deliver m1 \<sqsubset>\<^sup>j Deliver m2"
lemma (in causal_network) causal_broadcast:
assumes "Deliver m2 \<in> set (history j)"
and "Deliver m1 \<sqsubset>\<^sup>i Broadcast m2"
shows "Deliver m1 \<sqsubset>\<^sup>j Deliver m2"
using assms causal_delivery hb.intros(2) by blast
lemma (in network) hb_broadcast_exists1:
assumes "hb m1 m2"
shows "\<exists>i. Broadcast m1 \<in> set (history i)"
using assms
apply(induction rule: hb.induct)
apply(meson insert_subset node_histories.local_order_carrier_closed node_histories_axioms)
apply(meson delivery_has_a_cause insert_subset local_order_carrier_closed)
apply simp
done
lemma (in network) hb_broadcast_exists2:
assumes "hb m1 m2"
shows "\<exists>i. Broadcast m2 \<in> set (history i)"
using assms
apply(induction rule: hb.induct)
apply(meson insert_subset node_histories.local_order_carrier_closed node_histories_axioms)
apply(meson delivery_has_a_cause insert_subset local_order_carrier_closed)
apply simp
done
subsection\<open>Causal networks\<close>
lemma (in causal_network) hb_has_a_reason:
assumes "hb m1 m2"
and "Broadcast m2 \<in> set (history i)"
shows "Deliver m1 \<in> set (history i) \<or> Broadcast m1 \<in> set (history i)"
using assms apply (induction rule: hb.induct)
apply(metis insert_subset local_order_carrier_closed network.broadcasts_unique network_axioms)
apply(metis insert_subset local_order_carrier_closed network.broadcasts_unique network_axioms)
using hb_trans causal_delivery local_order_carrier_closed apply blast
done
lemma (in causal_network) hb_cross_node_delivery:
assumes "hb m1 m2"
and "Broadcast m1 \<in> set (history i)"
and "Broadcast m2 \<in> set (history j)"
and "i \<noteq> j"
shows "Deliver m1 \<in> set (history j)"
using assms
apply(induction rule: hb.induct)
apply(metis broadcasts_unique insert_subset local_order_carrier_closed)
apply(metis insert_subset local_order_carrier_closed network.broadcasts_unique network_axioms)
using broadcasts_unique hb.intros(3) hb_has_a_reason apply blast
done
lemma (in causal_network) hb_irrefl:
assumes "hb m1 m2"
shows "m1 \<noteq> m2"
using assms proof(induction rule: hb.induct)
case (hb_broadcast m1 i m2) thus ?case
using node_total_order_antisym by blast
next
case (hb_deliver m1 i m2) thus ?case
by(meson causal_broadcast insert_subset local_order_carrier_closed node_total_order_irrefl)
next
case (hb_trans m1 m2 m3)
then obtain i j where "Broadcast m3 \<in> set (history i)" "Broadcast m2 \<in> set (history j)"
using hb_broadcast_exists2 by blast
then show ?case
using assms hb_trans by (meson causal_network.causal_delivery causal_network_axioms
deliver_locally insert_subset network.hb.intros(3) network_axioms
node_histories.local_order_carrier_closed assms hb_trans
node_histories_axioms node_total_order_irrefl)
qed
lemma (in causal_network) hb_broadcast_broadcast_order:
assumes "hb m1 m2"
and "Broadcast m1 \<in> set (history i)"
and "Broadcast m2 \<in> set (history i)"
shows "Broadcast m1 \<sqsubset>\<^sup>i Broadcast m2"
using assms proof(induction rule: hb.induct)
case (hb_broadcast m1 i m2) thus ?case
by(metis insertI1 local_order_carrier_closed network.broadcasts_unique network_axioms subsetCE)
next
case (hb_deliver m1 i m2) thus ?case
by(metis broadcasts_unique insert_subset local_order_carrier_closed
network.broadcast_before_delivery network_axioms node_total_order_trans)
next
case (hb_trans m1 m2 m3)
then show ?case
proof (cases "Broadcast m2 \<in> set (history i)")
case True thus ?thesis
using hb_trans node_total_order_trans by blast
next
case False hence "Deliver m2 \<in> set (history i)" "m1 \<noteq> m2" "m2 \<noteq> m3"
using hb_has_a_reason hb_trans by auto
thus ?thesis
by(metis hb_trans event.inject(1) hb.intros(1) hb_irrefl network.hb.intros(3) network_axioms node_order_is_total hb_irrefl)
qed
qed
lemma (in causal_network) hb_antisym:
assumes "hb x y"
and "hb y x"
shows "False"
using assms proof(induction rule: hb.induct)
fix m1 i m2
assume "hb m2 m1" and "Broadcast m1 \<sqsubset>\<^sup>i Broadcast m2"
thus False
apply - proof(erule hb_elim)
show "\<And>ia. Broadcast m1 \<sqsubset>\<^sup>i Broadcast m2 \<Longrightarrow> Broadcast m2 \<sqsubset>\<^sup>ia Broadcast m1 \<Longrightarrow> False"
by(metis broadcasts_unique insert_subset local_order_carrier_closed node_total_order_irrefl node_total_order_trans)
next
show "\<And>ia. Broadcast m1 \<sqsubset>\<^sup>i Broadcast m2 \<Longrightarrow> Deliver m2 \<sqsubset>\<^sup>ia Broadcast m1 \<Longrightarrow> False"
by(metis broadcast_before_delivery broadcasts_unique insert_subset local_order_carrier_closed node_total_order_irrefl node_total_order_trans)
next
show "\<And>m2a. Broadcast m1 \<sqsubset>\<^sup>i Broadcast m2 \<Longrightarrow> hb m2 m2a \<Longrightarrow> hb m2a m1 \<Longrightarrow> False"
using assms(1) assms(2) hb.intros(3) hb_irrefl by blast
qed
next
fix m1 i m2
assume "hb m2 m1"
and "Deliver m1 \<sqsubset>\<^sup>i Broadcast m2"
thus "False"
apply - proof(erule hb_elim)
show "\<And>ia. Deliver m1 \<sqsubset>\<^sup>i Broadcast m2 \<Longrightarrow> Broadcast m2 \<sqsubset>\<^sup>ia Broadcast m1 \<Longrightarrow> False"
by (metis broadcast_before_delivery broadcasts_unique insert_subset local_order_carrier_closed node_total_order_irrefl node_total_order_trans)
next
show "\<And>ia. Deliver m1 \<sqsubset>\<^sup>i Broadcast m2 \<Longrightarrow> Deliver m2 \<sqsubset>\<^sup>ia Broadcast m1 \<Longrightarrow> False"
by (meson causal_network.causal_delivery causal_network_axioms hb.intros(2) hb.intros(3) insert_subset local_order_carrier_closed node_total_order_irrefl)
next
show "\<And>m2a. Deliver m1 \<sqsubset>\<^sup>i Broadcast m2 \<Longrightarrow> hb m2 m2a \<Longrightarrow> hb m2a m1 \<Longrightarrow> False"
by (meson causal_delivery hb.intros(2) insert_subset local_order_carrier_closed network.hb.intros(3) network_axioms node_total_order_irrefl)
qed
next
fix m1 m2 m3
assume "hb m1 m2" "hb m2 m3" "hb m3 m1"
and "(hb m2 m1 \<Longrightarrow> False)" "(hb m3 m2 \<Longrightarrow> False)"
thus "False"
using hb.intros(3) by blast
qed
definition (in network) node_deliver_messages :: "'msg event list \<Rightarrow> 'msg list" where
"node_deliver_messages cs \<equiv> List.map_filter (\<lambda>e. case e of Deliver m \<Rightarrow> Some m | _ \<Rightarrow> None) cs"
lemma (in network) node_deliver_messages_empty [simp]:
shows "node_deliver_messages [] = []"
by(auto simp add: node_deliver_messages_def List.map_filter_simps)
lemma (in network) node_deliver_messages_Cons:
shows "node_deliver_messages (x#xs) = (node_deliver_messages [x])@(node_deliver_messages xs)"
by(auto simp add: node_deliver_messages_def map_filter_def)
lemma (in network) node_deliver_messages_append:
shows "node_deliver_messages (xs@ys) = (node_deliver_messages xs)@(node_deliver_messages ys)"
by(auto simp add: node_deliver_messages_def map_filter_def)
lemma (in network) node_deliver_messages_Broadcast [simp]:
shows "node_deliver_messages [Broadcast m] = []"
by(clarsimp simp: node_deliver_messages_def map_filter_def)
lemma (in network) node_deliver_messages_Deliver [simp]:
shows "node_deliver_messages [Deliver m] = [m]"
by(clarsimp simp: node_deliver_messages_def map_filter_def)
lemma (in network) prefix_msg_in_history:
assumes "es prefix of i"
and "m \<in> set (node_deliver_messages es)"
shows "Deliver m \<in> set (history i)"
using assms prefix_to_carriers by(fastforce simp: node_deliver_messages_def map_filter_def split: event.split_asm)
lemma (in network) prefix_contains_msg:
assumes "es prefix of i"
and "m \<in> set (node_deliver_messages es)"
shows "Deliver m \<in> set es"
using assms by(auto simp: node_deliver_messages_def map_filter_def split: event.split_asm)
lemma (in network) node_deliver_messages_distinct:
assumes "xs prefix of i"
shows "distinct (node_deliver_messages xs)"
using assms proof(induction xs rule: rev_induct)
case Nil thus ?case by simp
next
case (snoc x xs)
{ fix y assume *: "y \<in> set (node_deliver_messages xs)" "y \<in> set (node_deliver_messages [x])"
moreover have "distinct (xs @ [x])"
using assms snoc prefix_distinct by blast
ultimately have "False"
using assms apply(case_tac x; clarsimp simp add: map_filter_def node_deliver_messages_def)
using * prefix_contains_msg snoc.prems by blast
} thus ?case
using snoc by(fastforce simp add: node_deliver_messages_append node_deliver_messages_def map_filter_def)
qed
lemma (in network) drop_last_message:
assumes "evts prefix of i"
and "node_deliver_messages evts = msgs @ [last_msg]"
shows "\<exists>pre. pre prefix of i \<and> node_deliver_messages pre = msgs"
proof -
have "Deliver last_msg \<in> set evts"
using assms network.prefix_contains_msg network_axioms by force
then obtain idx where *: "idx < length evts" "evts ! idx = Deliver last_msg"
by (meson set_elem_nth)
then obtain pre suf where "evts = pre @ (evts ! idx) # suf"
using id_take_nth_drop by blast
hence **: "evts = pre @ (Deliver last_msg) # suf"
using assms * by auto
moreover hence "distinct (node_deliver_messages ([Deliver last_msg] @ suf))"
by (metis assms(1) assms(2) distinct_singleton node_deliver_messages_Cons node_deliver_messages_Deliver
node_deliver_messages_append node_deliver_messages_distinct not_Cons_self2 pre_suf_eq_distinct_list)
ultimately have "node_deliver_messages ([Deliver last_msg] @ suf) = [last_msg] @ []"
by (metis append_self_conv assms(1) assms(2) node_deliver_messages_Cons node_deliver_messages_Deliver
node_deliver_messages_append node_deliver_messages_distinct not_Cons_self2 pre_suf_eq_distinct_list)
thus ?thesis
using assms * ** by (metis append1_eq_conv append_Cons append_Nil node_deliver_messages_append prefix_of_appendD)
qed
locale network_with_ops = causal_network history fst
for history :: "nat \<Rightarrow> ('msgid \<times> 'op) event list" +
fixes interp :: "'op \<Rightarrow> 'state \<rightharpoonup> 'state"
and initial_state :: "'state"
context network_with_ops begin
definition interp_msg :: "'msgid \<times> 'op \<Rightarrow> 'state \<rightharpoonup> 'state" where
"interp_msg msg state \<equiv> interp (snd msg) state"
sublocale hb: happens_before weak_hb hb interp_msg
proof
fix x y :: "'msgid \<times> 'op"
show "hb x y = (weak_hb x y \<and> \<not> weak_hb y x)"
unfolding weak_hb_def using hb_antisym by blast
next
fix x
show "weak_hb x x"
using weak_hb_def by blast
next
fix x y z
assume "weak_hb x y" "weak_hb y z"
thus "weak_hb x z"
using weak_hb_def by (metis network.hb.intros(3) network_axioms)
qed
end
definition (in network_with_ops) apply_operations :: "('msgid \<times> 'op) event list \<rightharpoonup> 'state" where
"apply_operations es \<equiv> hb.apply_operations (node_deliver_messages es) initial_state"
definition (in network_with_ops) node_deliver_ops :: "('msgid \<times> 'op) event list \<Rightarrow> 'op list" where
"node_deliver_ops cs \<equiv> map snd (node_deliver_messages cs)"
lemma (in network_with_ops) apply_operations_empty [simp]:
shows "apply_operations [] = Some initial_state"
by(auto simp add: apply_operations_def)
lemma (in network_with_ops) apply_operations_Broadcast [simp]:
shows "apply_operations (xs @ [Broadcast m]) = apply_operations xs"
by(auto simp add: apply_operations_def node_deliver_messages_def map_filter_def)
lemma (in network_with_ops) apply_operations_Deliver [simp]:
shows "apply_operations (xs @ [Deliver m]) = (apply_operations xs \<bind> interp_msg m)"
by(auto simp add: apply_operations_def node_deliver_messages_def map_filter_def kleisli_def)
lemma (in network_with_ops) hb_consistent_technical:
assumes "\<And>m n. m < length cs \<Longrightarrow> n < m \<Longrightarrow> cs ! n \<sqsubset>\<^sup>i cs ! m"
shows "hb.hb_consistent (node_deliver_messages cs)"
using assms proof (induction cs rule: rev_induct)
case Nil thus ?case
by(simp add: node_deliver_messages_def hb.hb_consistent.intros(1) map_filter_simps(2))
next
case (snoc x xs)
hence *: "(\<And>m n. m < length xs \<Longrightarrow> n < m \<Longrightarrow> xs ! n \<sqsubset>\<^sup>i xs ! m)"
by(-, erule_tac x=m in meta_allE, erule_tac x=n in meta_allE, clarsimp simp add: nth_append)
then show ?case
proof (cases x)
case (Broadcast x1) thus ?thesis
using snoc * by (simp add: node_deliver_messages_append)
next
case (Deliver x2) thus ?thesis
- using snoc * apply(clarsimp simp add: node_deliver_messages_def map_filter_def map_filter_append)
+ using snoc * [[simproc del: defined_all]]
+ apply (clarsimp simp add: node_deliver_messages_def map_filter_def map_filter_append)
apply (rename_tac m m1 m2)
apply (case_tac m; clarsimp)
- apply(drule set_elem_nth, erule exE, erule conjE)
- apply(erule_tac x="length xs" in meta_allE)
- apply(clarsimp simp add: nth_append)
- by (metis causal_delivery insert_subset node_histories.local_order_carrier_closed
- node_histories_axioms node_total_order_antisym)
- qed
+ apply (drule set_elem_nth, erule exE, erule conjE)
+ apply (erule_tac x="length xs" in meta_allE)
+ apply (clarsimp simp add: nth_append)
+ apply (metis causal_delivery insert_subset local_order_carrier_closed
+ node_total_order_antisym)
+ done
+ qed
qed
corollary (in network_with_ops)
shows "hb.hb_consistent (node_deliver_messages (history i))"
by (metis hb_consistent_technical history_order_def less_one linorder_neqE_nat list_nth_split zero_order(3))
lemma (in network_with_ops) hb_consistent_prefix:
assumes "xs prefix of i"
shows "hb.hb_consistent (node_deliver_messages xs)"
using assms proof (clarsimp simp: prefix_of_node_history_def, rule_tac i=i in hb_consistent_technical)
fix m n ys assume *: "xs @ ys = history i" "m < length xs" "n < m"
consider (a) "xs = []" | (b) "\<exists>c. xs = [c]" | (c) "Suc 0 < length (xs)"
by (metis Suc_pred length_Suc_conv length_greater_0_conv zero_less_diff)
thus "xs ! n \<sqsubset>\<^sup>i xs ! m"
proof (cases)
case a thus ?thesis
using * by clarsimp
next
case b thus ?thesis
using assms * by clarsimp
next
case c thus ?thesis
using assms * apply clarsimp
apply(drule list_nth_split, assumption, clarsimp simp: c)
apply (metis append.assoc append.simps(2) history_order_def)
done
qed
qed
locale network_with_constrained_ops = network_with_ops +
fixes valid_msg :: "'c \<Rightarrow> ('a \<times> 'b) \<Rightarrow> bool"
assumes broadcast_only_valid_msgs: "pre @ [Broadcast m] prefix of i \<Longrightarrow>
\<exists>state. apply_operations pre = Some state \<and> valid_msg state m"
lemma (in network_with_constrained_ops) broadcast_is_valid:
assumes "Broadcast m \<in> set (history i)"
shows "\<exists>state. valid_msg state m"
using assms broadcast_only_valid_msgs events_before_exist by blast
lemma (in network_with_constrained_ops) deliver_is_valid:
assumes "Deliver m \<in> set (history i)"
shows "\<exists>j pre state. pre @ [Broadcast m] prefix of j \<and> apply_operations pre = Some state \<and> valid_msg state m"
using assms apply (clarsimp dest!: delivery_has_a_cause)
using broadcast_only_valid_msgs events_before_exist apply blast
done
lemma (in network_with_constrained_ops) deliver_in_prefix_is_valid:
assumes "xs prefix of i"
and "Deliver m \<in> set xs"
shows "\<exists>state. valid_msg state m"
by (meson assms network_with_constrained_ops.deliver_is_valid network_with_constrained_ops_axioms prefix_elem_to_carriers)
subsection\<open>Dummy network models\<close>
interpretation trivial_node_histories: node_histories "\<lambda>m. []"
by standard auto
interpretation trivial_network: network "\<lambda>m. []" id
by standard auto
interpretation trivial_causal_network: causal_network "\<lambda>m. []" id
by standard auto
interpretation trivial_network_with_ops: network_with_ops "\<lambda>m. []" "(\<lambda>x y. Some y)" 0
by standard auto
interpretation trivial_network_with_constrained_ops: network_with_constrained_ops "\<lambda>m. []" "(\<lambda>x y. Some y)" 0 "\<lambda>x y. True"
by standard (simp add: trivial_node_histories.prefix_of_node_history_def)
end
diff --git a/thys/CYK/CYK.thy b/thys/CYK/CYK.thy
--- a/thys/CYK/CYK.thy
+++ b/thys/CYK/CYK.thy
@@ -1,1066 +1,1064 @@
(* Title: A formalisation of the Cocke-Younger-Kasami algorithm
Author: Maksym Bortin <Maksym.Bortin@nicta.com.au>
*)
theory CYK
imports Main
begin
text \<open>The theory is structured as follows. First section deals with modelling
of grammars, derivations, and the language semantics. Then the basic
properties are proved. Further, CYK is abstractly specified and its
underlying recursive relationship proved. The final section contains a
prototypical implementation accompanied by a proof of its correctness.\<close>
section "Basic modelling"
subsection "Grammars in Chomsky normal form"
text "A grammar in Chomsky normal form is here simply modelled
by a list of production rules (the type CNG below), each having a non-terminal
symbol on the lhs and either two non-terminals or one terminal
symbol on the rhs."
datatype ('n, 't) RHS = Branch 'n 'n
| Leaf 't
type_synonym ('n, 't) CNG = "('n \<times> ('n, 't) RHS) list"
text "Abbreviating the list append symbol for better readability"
abbreviation list_append :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "\<cdot>" 65)
where "xs \<cdot> ys \<equiv> xs @ ys"
subsection "Derivation by grammars"
text\<open>A \emph{word form} (or sentential form) may be built of both non-terminal and terminal
symbols, as opposed to a \emph{word} that contains only terminals. By the usage of disjoint
union, non-terminals are injected into a word form by @{term "Inl"} whereas terminals --
by @{term "Inr"}.\<close>
type_synonym ('n, 't) word_form = "('n + 't) list"
type_synonym 't word = "'t list"
text "A single step derivation relation on word forms is induced by a grammar in the standard way,
replacing a non-terminal within a word form in accordance to the production rules."
definition DSTEP :: "('n, 't) CNG \<Rightarrow> (('n, 't) word_form \<times> ('n, 't) word_form) set"
where "DSTEP G = {(l \<cdot> [Inl N] \<cdot> r, x) | l N r rhs x. (N, rhs) \<in> set G \<and>
(case rhs of
Branch A B \<Rightarrow> x = l \<cdot> [Inl A, Inl B] \<cdot> r
| Leaf t \<Rightarrow> x = l \<cdot> [Inr t] \<cdot> r)}"
abbreviation DSTEP' :: "('n, 't) word_form \<Rightarrow> ('n, 't) CNG \<Rightarrow> ('n, 't) word_form \<Rightarrow> bool" ("_ -_\<rightarrow> _" [60, 61, 60] 61)
where "w -G\<rightarrow> w' \<equiv> (w, w') \<in> DSTEP G"
abbreviation DSTEP_reflc :: "('n, 't) word_form \<Rightarrow> ('n, 't) CNG \<Rightarrow> ('n, 't) word_form \<Rightarrow> bool" ("_ -_\<rightarrow>\<^sup>= _" [60, 61, 60] 61)
where "w -G\<rightarrow>\<^sup>= w' \<equiv> (w, w') \<in> (DSTEP G)\<^sup>="
abbreviation DSTEP_transc :: "('n, 't) word_form \<Rightarrow> ('n, 't) CNG \<Rightarrow> ('n, 't) word_form \<Rightarrow> bool" ("_ -_\<rightarrow>\<^sup>+ _" [60, 61, 60] 61)
where "w -G\<rightarrow>\<^sup>+ w' \<equiv> (w, w') \<in> (DSTEP G)\<^sup>+"
abbreviation DSTEP_rtransc :: "('n, 't) word_form \<Rightarrow> ('n, 't) CNG \<Rightarrow> ('n, 't) word_form \<Rightarrow> bool" ("_ -_\<rightarrow>\<^sup>* _" [60, 61, 60] 61)
where "w -G\<rightarrow>\<^sup>* w' \<equiv> (w, w') \<in> (DSTEP G)\<^sup>*"
subsection "The generated language semantics"
text "The language generated by a grammar from a non-terminal symbol
comprises all words that can be derived from the non-terminal
in one or more steps.
Notice that by the presented grammar modelling, languages containing
the empty word cannot be generated. Hence in rare situations when such
languages are required, the empty word case should be treated separately."
definition Lang :: "('n, 't) CNG \<Rightarrow> 'n \<Rightarrow> 't word set"
where "Lang G S = {w. [Inl S] -G\<rightarrow>\<^sup>+ map Inr w }"
text\<open>So, for instance, a grammar generating the language $a^nb^n$
from the non-terminal @{term "''S''"} might look as follows.\<close>
definition "G_anbn =
[(''S'', Branch ''A'' ''T''),
(''S'', Branch ''A'' ''B''),
(''T'', Branch ''S'' ''B''),
(''A'', Leaf ''a''),
(''B'', Leaf ''b'')]"
text\<open>Now the term @{term "Lang G_anbn ''S''"} denotes the set of words of
the form $a^nb^n$ with $n > 0$. This is intuitively clear, but not
straight forward to show, and a lengthy proof for that is out of scope.\<close>
section "Basic properties"
lemma prod_into_DSTEP1 :
"(S, Branch A B) \<in> set G \<Longrightarrow>
L \<cdot> [Inl S] \<cdot> R -G\<rightarrow> L \<cdot> [Inl A, Inl B] \<cdot> R"
by(simp add: DSTEP_def, rule_tac x="L" in exI, force)
lemma prod_into_DSTEP2 :
"(S, Leaf a) \<in> set G \<Longrightarrow>
L \<cdot> [Inl S] \<cdot> R -G\<rightarrow> L \<cdot> [Inr a] \<cdot> R"
by(simp add: DSTEP_def, rule_tac x="L" in exI, force)
lemma DSTEP_D :
"s -G\<rightarrow> t \<Longrightarrow>
\<exists>L N R rhs. s = L \<cdot> [Inl N] \<cdot> R \<and> (N, rhs) \<in> set G \<and>
(\<forall>A B. rhs = Branch A B \<longrightarrow> t = L \<cdot> [Inl A, Inl B] \<cdot> R) \<and>
(\<forall>x. rhs = Leaf x \<longrightarrow> t = L \<cdot> [Inr x] \<cdot> R)"
by(unfold DSTEP_def, clarsimp, simp split: RHS.split_asm, blast+)
lemma DSTEP_append :
assumes a: "s -G\<rightarrow> t"
shows "L \<cdot> s \<cdot> R -G\<rightarrow> L \<cdot> t \<cdot> R"
proof -
from a have "\<exists>l N r rhs. s = l \<cdot> [Inl N] \<cdot> r \<and> (N, rhs) \<in> set G \<and>
(\<forall>A B. rhs = Branch A B \<longrightarrow> t = l \<cdot> [Inl A, Inl B] \<cdot> r) \<and>
(\<forall>x. rhs = Leaf x \<longrightarrow> t = l \<cdot> [Inr x] \<cdot> r)" (is "\<exists>l N r rhs. ?P l N r rhs")
by(rule DSTEP_D)
then obtain l N r rhs where "?P l N r rhs" by blast
thus ?thesis
by(simp add: DSTEP_def, rule_tac x="L \<cdot> l" in exI,
rule_tac x=N in exI, rule_tac x="r \<cdot> R" in exI,
simp, rule_tac x=rhs in exI, simp split: RHS.split)
qed
lemma DSTEP_star_mono :
"s -G\<rightarrow>\<^sup>* t \<Longrightarrow> length s \<le> length t"
proof(erule rtrancl_induct, simp)
fix t u
assume "s -G\<rightarrow>\<^sup>* t"
assume a: "t -G\<rightarrow> u"
assume b: "length s \<le> length t"
show "length s \<le> length u"
proof -
from a have "\<exists>L N R rhs. t = L \<cdot> [Inl N] \<cdot> R \<and> (N, rhs) \<in> set G \<and>
(\<forall>A B. rhs = Branch A B \<longrightarrow> u = L \<cdot> [Inl A, Inl B] \<cdot> R) \<and>
(\<forall>x. rhs = Leaf x \<longrightarrow> u = L \<cdot> [Inr x] \<cdot> R)" (is "\<exists>L N R rhs. ?P L N R rhs")
by(rule DSTEP_D)
then obtain L N R rhs where "?P L N R rhs" by blast
with b show ?thesis
by(case_tac rhs, clarsimp+)
qed
qed
lemma DSTEP_comp :
assumes a: "l \<cdot> r -G\<rightarrow> t"
shows "\<exists>l' r'. l -G\<rightarrow>\<^sup>= l' \<and> r -G\<rightarrow>\<^sup>= r' \<and> t = l' \<cdot> r'"
proof -
from a have "\<exists>L N R rhs. l \<cdot> r = L \<cdot> [Inl N] \<cdot> R \<and> (N, rhs) \<in> set G \<and>
(\<forall>A B. rhs = Branch A B \<longrightarrow> t = L \<cdot> [Inl A, Inl B] \<cdot> R) \<and>
(\<forall>x. rhs = Leaf x \<longrightarrow> t = L \<cdot> [Inr x] \<cdot> R)" (is "\<exists>L N R rhs. ?T L N R rhs")
by(rule DSTEP_D)
then obtain L N R rhs where b: "?T L N R rhs" by blast
hence "l \<cdot> r = L \<cdot> Inl N # R" by simp
hence "\<exists>u. (l = L \<cdot> u \<and> u \<cdot> r = Inl N # R) \<or> (l \<cdot> u = L \<and> r = u \<cdot> Inl N # R)" by(rule append_eq_append_conv2[THEN iffD1])
then obtain xs where c: "l = L \<cdot> xs \<and> xs \<cdot> r = Inl N # R \<or> l \<cdot> xs = L \<and> r = xs \<cdot> Inl N # R" (is "?C1 \<or> ?C2") by blast
show ?thesis
proof(cases rhs)
case (Leaf x)
with b have d: "t = L \<cdot> [Inr x] \<cdot> R \<and> (N, Leaf x) \<in> set G" by simp
from c show ?thesis
proof
assume e: "?C1"
show ?thesis
proof(cases xs)
case Nil with d and e show ?thesis
by(clarsimp, rule_tac x=L in exI, simp add: DSTEP_def, simp split: RHS.split, blast)
next
case (Cons z zs) with d and e show ?thesis
by(rule_tac x="L \<cdot> Inr x # zs" in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split, blast)
qed
next
assume e: "?C2"
show ?thesis
proof(cases xs)
case Nil with d and e show ?thesis
by(rule_tac x=L in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split, blast)
next
case (Cons z zs) with d and e show ?thesis
by(rule_tac x="l" in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split,
rule_tac x="z#zs" in exI, rule_tac x=N in exI, rule_tac x=R in exI, simp, rule_tac x="Leaf x" in exI, simp)
qed
qed
next
case (Branch A B)
with b have d: "t = L \<cdot> [Inl A, Inl B] \<cdot> R \<and> (N, Branch A B) \<in> set G" by simp
from c show ?thesis
proof
assume e: "?C1"
show ?thesis
proof(cases xs)
case Nil with d and e show ?thesis
by(clarsimp, rule_tac x=L in exI, simp add: DSTEP_def, simp split: RHS.split, blast)
next
case (Cons z zs) with d and e show ?thesis
by(rule_tac x="L \<cdot> [Inl A, Inl B] \<cdot> zs" in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split, blast)
qed
next
assume e: "?C2"
show ?thesis
proof(cases xs)
case Nil with d and e show ?thesis
by(rule_tac x=L in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split, blast)
next
case (Cons z zs) with d and e show ?thesis
by(rule_tac x="l" in exI, clarsimp, simp add: DSTEP_def, simp split: RHS.split,
rule_tac x="z#zs" in exI, rule_tac x=N in exI, rule_tac x=R in exI, simp, rule_tac x="Branch A B" in exI, simp)
qed
qed
qed
qed
theorem DSTEP_star_comp1 :
assumes A: "l \<cdot> r -G\<rightarrow>\<^sup>* t"
shows "\<exists>l' r'. l -G\<rightarrow>\<^sup>* l' \<and> r -G\<rightarrow>\<^sup>* r' \<and> t = l' \<cdot> r'"
proof -
have "\<And>s. s -G\<rightarrow>\<^sup>* t \<Longrightarrow>
\<forall>l r. s = l \<cdot> r \<longrightarrow> (\<exists>l' r'. l -G\<rightarrow>\<^sup>* l' \<and> r -G\<rightarrow>\<^sup>* r' \<and> t = l' \<cdot> r')" (is "\<And>s. ?P s t \<Longrightarrow> ?Q s t")
proof(erule rtrancl_induct, force)
fix s t u
assume "?P s t"
assume a: "t -G\<rightarrow> u"
assume b: "?Q s t"
show "?Q s u"
proof(clarify)
fix l r
assume "s = l \<cdot> r"
with b have "\<exists>l' r'. l -G\<rightarrow>\<^sup>* l' \<and> r -G\<rightarrow>\<^sup>* r' \<and> t = l' \<cdot> r'" by simp
then obtain l' r' where c: "l -G\<rightarrow>\<^sup>* l' \<and> r -G\<rightarrow>\<^sup>* r' \<and> t = l' \<cdot> r'" by blast
with a have "l' \<cdot> r' -G\<rightarrow> u" by simp
hence "\<exists>l'' r''. l' -G\<rightarrow>\<^sup>= l'' \<and> r' -G\<rightarrow>\<^sup>= r'' \<and> u = l'' \<cdot> r''" by(rule DSTEP_comp)
then obtain l'' r'' where "l' -G\<rightarrow>\<^sup>= l'' \<and> r' -G\<rightarrow>\<^sup>= r'' \<and> u = l'' \<cdot> r''" by blast
hence "l' -G\<rightarrow>\<^sup>* l'' \<and> r' -G\<rightarrow>\<^sup>* r'' \<and> u = l'' \<cdot> r''" by blast
with c show "\<exists>l' r'. l -G\<rightarrow>\<^sup>* l' \<and> r -G\<rightarrow>\<^sup>* r' \<and> u = l' \<cdot> r'"
by(rule_tac x=l'' in exI, rule_tac x=r'' in exI, force)
qed
qed
with A show ?thesis by force
qed
theorem DSTEP_star_comp2 :
assumes A: "l -G\<rightarrow>\<^sup>* l'"
and B: "r -G\<rightarrow>\<^sup>* r'"
shows "l \<cdot> r -G\<rightarrow>\<^sup>* l' \<cdot> r'"
proof -
have "l -G\<rightarrow>\<^sup>* l' \<Longrightarrow>
\<forall>r r'. r -G\<rightarrow>\<^sup>* r' \<longrightarrow> l \<cdot> r -G\<rightarrow>\<^sup>* l' \<cdot> r'" (is "?P l l' \<Longrightarrow> ?Q l l'")
proof(erule rtrancl_induct)
show "?Q l l"
proof(clarify, erule rtrancl_induct, simp)
fix r s t
assume a: "s -G\<rightarrow> t"
assume b: "l \<cdot> r -G\<rightarrow>\<^sup>* l \<cdot> s"
show "l \<cdot> r -G\<rightarrow>\<^sup>* l \<cdot> t"
proof -
from a have "l \<cdot> s -G\<rightarrow> l \<cdot> t" by(drule_tac L=l and R="[]" in DSTEP_append, simp)
with b show ?thesis by simp
qed
qed
next
fix s t
assume a: "s -G\<rightarrow> t"
assume b: "?Q l s"
show "?Q l t"
proof(clarsimp)
fix r r'
assume "r -G\<rightarrow>\<^sup>* r'"
with b have c: "l \<cdot> r -G\<rightarrow>\<^sup>* s \<cdot> r'" by simp
from a have "s \<cdot> r' -G\<rightarrow> t \<cdot> r'" by(drule_tac L="[]" and R=r' in DSTEP_append, simp)
with c show "l \<cdot> r -G\<rightarrow>\<^sup>* t \<cdot> r'" by simp
qed
qed
with A and B show ?thesis by simp
qed
lemma DSTEP_trancl_term :
assumes A: "[Inl S] -G\<rightarrow>\<^sup>+ t"
and B: "Inr x \<in> set t"
shows "\<exists>N. (N, Leaf x) \<in> set G"
proof -
have "[Inl S] -G\<rightarrow>\<^sup>+ t \<Longrightarrow>
\<forall>x. Inr x \<in> set t \<longrightarrow> (\<exists>N. (N, Leaf x) \<in> set G)" (is "?P t \<Longrightarrow> ?Q t")
proof(erule trancl_induct)
fix t
assume a: "[Inl S] -G\<rightarrow> t"
show "?Q t"
proof -
from a have "\<exists>rhs. (S, rhs) \<in> set G \<and>
(\<forall>A B. rhs = Branch A B \<longrightarrow> t = [Inl A, Inl B]) \<and>
(\<forall>x. rhs = Leaf x \<longrightarrow> t = [Inr x])" (is "\<exists>rhs. ?P rhs")
by(simp add: DSTEP_def, clarsimp, simp split: RHS.split_asm, case_tac l, force, simp,
clarsimp, simp split: RHS.split_asm, case_tac l, force, simp)
then obtain rhs where "?P rhs" by blast
thus ?thesis
by(case_tac rhs, clarsimp, force)
qed
next
fix s t
assume a: "s -G\<rightarrow> t"
assume b: "?Q s"
show "?Q t"
proof -
from a have "\<exists>L N R rhs. s = L \<cdot> [Inl N] \<cdot> R \<and> (N, rhs) \<in> set G \<and>
(\<forall>A B. rhs = Branch A B \<longrightarrow> t = L \<cdot> [Inl A, Inl B] \<cdot> R) \<and>
(\<forall>x. rhs = Leaf x \<longrightarrow> t = L \<cdot> [Inr x] \<cdot> R)" (is "\<exists>L N R rhs. ?P L N R rhs")
by(rule DSTEP_D)
then obtain L N R rhs where "?P L N R rhs" by blast
with b show ?thesis
by(case_tac rhs, clarsimp, force)
qed
qed
with A and B show ?thesis by simp
qed
subsection "Properties of generated languages"
lemma Lang_no_Nil :
"w \<in> Lang G S \<Longrightarrow> w \<noteq> []"
by(simp add: Lang_def, drule trancl_into_rtrancl, drule DSTEP_star_mono, force)
lemma Lang_rtrancl_eq :
"(w \<in> Lang G S) = [Inl S] -G\<rightarrow>\<^sup>* map Inr w" (is "?L = (?p \<in> ?R\<^sup>*)")
proof(simp add: Lang_def, rule iffI, erule trancl_into_rtrancl)
assume "?p \<in> ?R\<^sup>*"
hence "?p \<in> (?R\<^sup>+)\<^sup>=" by(subst rtrancl_trancl_reflcl[THEN sym], assumption)
hence "[Inl S] = map Inr w \<or> ?p \<in> ?R\<^sup>+" by force
thus "?p \<in> ?R\<^sup>+" by(case_tac w, simp_all)
qed
lemma Lang_term :
"w \<in> Lang G S \<Longrightarrow>
\<forall>x \<in> set w. \<exists>N. (N, Leaf x) \<in> set G"
by(clarsimp simp add: Lang_def, drule DSTEP_trancl_term,
simp, erule imageI, assumption)
lemma Lang_eq1 :
"([x] \<in> Lang G S) = ((S, Leaf x) \<in> set G)"
proof(simp add: Lang_def, rule iffI, subst (asm) trancl_unfold_left, clarsimp)
fix t
assume a: "[Inl S] -G\<rightarrow> t"
assume b: "t -G\<rightarrow>\<^sup>* [Inr x]"
note DSTEP_star_mono[OF b, simplified]
hence c: "length t \<le> 1" by simp
have "\<exists>z. t = [z]"
proof(cases t)
assume "t = []"
with b have d: "[] -G\<rightarrow>\<^sup>* [Inr x]" by simp
have "\<And>s. ([], s) \<in> (DSTEP G)\<^sup>* \<Longrightarrow> s = []"
by(erule rtrancl_induct, simp_all, drule DSTEP_D, clarsimp)
note this[OF d]
thus ?thesis by simp
next
fix z zs
assume "t = z#zs"
with c show ?thesis by force
qed
with a have "\<exists>z. (S, Leaf z) \<in> set G \<and> t = [Inr z]"
by(clarsimp simp add: DSTEP_def, simp split: RHS.split_asm, case_tac l, simp_all)
with b show "(S, Leaf x) \<in> set G"
proof(clarsimp)
fix z
assume c: "(S, Leaf z) \<in> set G"
assume "[Inr z] -G\<rightarrow>\<^sup>* [Inr x]"
hence "([Inr z], [Inr x]) \<in> ((DSTEP G)\<^sup>+)\<^sup>=" by simp
hence "[Inr z] = [Inr x] \<or> [Inr z] -G\<rightarrow>\<^sup>+ [Inr x]" by force
hence "x = z"
proof
assume "[Inr z] = [Inr x]" thus ?thesis by simp
next
assume "[Inr z] -G\<rightarrow>\<^sup>+ [Inr x]"
hence "\<exists>u. [Inr z] -G\<rightarrow> u \<and> u -G\<rightarrow>\<^sup>* [Inr x]" by(subst (asm) trancl_unfold_left, force)
then obtain u where "[Inr z] -G\<rightarrow> u" by blast
thus ?thesis by(clarsimp simp add: DSTEP_def, case_tac l, simp_all)
qed
with c show ?thesis by simp
qed
next
assume a: "(S, Leaf x) \<in> set G"
show "[Inl S] -G\<rightarrow>\<^sup>+ [Inr x]"
by(rule r_into_trancl, simp add: DSTEP_def, rule_tac x="[]" in exI,
rule_tac x="S" in exI, rule_tac x="[]" in exI, simp, rule_tac x="Leaf x" in exI,
simp add: a)
qed
theorem Lang_eq2 :
"(w \<in> Lang G S \<and> 1 < length w) =
(\<exists>A B. (S, Branch A B) \<in> set G \<and> (\<exists>l r. w = l \<cdot> r \<and> l \<in> Lang G A \<and> r \<in> Lang G B))"
(is "?L = ?R")
proof(rule iffI, clarify, subst (asm) Lang_def, simp, subst (asm) trancl_unfold_left, clarsimp)
have map_Inr_split : "\<And>xs. \<forall>zs w. map Inr w = xs \<cdot> zs \<longrightarrow>
(\<exists>u v. w = u \<cdot> v \<and> xs = map Inr u \<and> zs = map Inr v)"
by(induct_tac xs, simp, force)
fix t
assume a: "Suc 0 < length w"
assume b: "[Inl S] -G\<rightarrow> t"
assume c: "t -G\<rightarrow>\<^sup>* map Inr w"
from b have "\<exists>A B. (S, Branch A B) \<in> set G \<and> t = [Inl A, Inl B]"
proof(simp add: DSTEP_def, clarify, case_tac l, simp_all, simp split: RHS.split_asm, clarify)
fix x
assume "t = [Inr x]"
with c have d: "[Inr x] -G\<rightarrow>\<^sup>* map Inr w"by simp
have "\<And>x s. [Inr x] -G\<rightarrow>\<^sup>* s \<Longrightarrow> s = [Inr x]"
by(erule rtrancl_induct, simp_all, drule DSTEP_D, clarsimp, case_tac L, simp_all)
note this[OF d]
hence "w = [x]" by(case_tac w, simp_all)
with a show "False" by simp
qed
then obtain A B where d: "(S, Branch A B) \<in> set G \<and> t = [Inl A, Inl B]" by blast
with c have e: "[Inl A] \<cdot> [Inl B] -G\<rightarrow>\<^sup>* map Inr w" by simp
note DSTEP_star_comp1[OF e]
then obtain l' r' where e: "[Inl A] -G\<rightarrow>\<^sup>* l' \<and> [Inl B] -G\<rightarrow>\<^sup>* r' \<and>
map Inr w = l' \<cdot> r'" by blast
note map_Inr_split[rule_format, OF e[THEN conjunct2, THEN conjunct2]]
then obtain u v where f: "w = u \<cdot> v \<and> l' = map Inr u \<and> r' = map Inr v" by blast
with e have g: "[Inl A] -G\<rightarrow>\<^sup>* map Inr u \<and> [Inl B] -G\<rightarrow>\<^sup>* map Inr v" by simp
show "?R"
by(rule_tac x=A in exI, rule_tac x=B in exI, simp add: d,
rule_tac x=u in exI, rule_tac x=v in exI, simp add: f,
(subst Lang_rtrancl_eq)+, rule g)
next
assume "?R"
then obtain A B l r where a: "(S, Branch A B) \<in> set G \<and> w = l \<cdot> r \<and> l \<in> Lang G A \<and> r \<in> Lang G B" by blast
have "[Inl A] \<cdot> [Inl B] -G\<rightarrow>\<^sup>* map Inr l \<cdot> map Inr r"
by(rule DSTEP_star_comp2, subst Lang_rtrancl_eq[THEN sym], simp add: a,
subst Lang_rtrancl_eq[THEN sym], simp add: a)
hence b: "[Inl A] \<cdot> [Inl B] -G\<rightarrow>\<^sup>* map Inr w" by(simp add: a)
have c: "w \<in> Lang G S"
by(simp add: Lang_def, subst trancl_unfold_left, rule_tac b="[Inl A] \<cdot> [Inl B]" in relcompI,
simp add: DSTEP_def, rule_tac x="[]" in exI, rule_tac x="S" in exI, rule_tac x="[]" in exI,
simp, rule_tac x="Branch A B" in exI, simp add: a[THEN conjunct1], rule b)
thus "?L"
proof
show "1 < length w"
proof(simp add: a, rule ccontr, drule leI)
assume "length l + length r \<le> Suc 0"
hence "l = [] \<or> r = []" by(case_tac l, simp_all)
thus "False"
proof
assume "l = []"
with a have "[] \<in> Lang G A" by simp
note Lang_no_Nil[OF this]
thus ?thesis by simp
next
assume "r = []"
with a have "[] \<in> Lang G B" by simp
note Lang_no_Nil[OF this]
thus ?thesis by simp
qed
qed
qed
qed
section "Abstract specification of CYK"
text "A subword of a word $w$, starting at the position $i$
(first element is at the position $0$) and having the length $j$, is defined as follows."
definition "subword w i j = take j (drop i w)"
text "Thus, to any subword of the given word $w$ CYK assigns all non-terminals
from which this subword is derivable by the grammar $G$."
definition "CYK G w i j = {S. subword w i j \<in> Lang G S}"
subsection \<open>Properties of @{term "subword"}\<close>
lemma subword_length :
"i + j \<le> length w \<Longrightarrow> length(subword w i j) = j"
by(simp add: subword_def)
lemma subword_nth1 :
"i + j \<le> length w \<Longrightarrow> k < j \<Longrightarrow>
(subword w i j)!k = w!(i + k)"
by(simp add: subword_def)
lemma subword_nth2 :
assumes A: "i + 1 \<le> length w"
shows "subword w i 1 = [w!i]"
proof -
note subword_length[OF A]
hence "\<exists>x. subword w i 1 = [x]" by(case_tac "subword w i 1", simp_all)
then obtain x where a:"subword w i 1 = [x]" by blast
note subword_nth1[OF A, where k="(0 :: nat)", simplified]
with a have "x = w!i" by simp
with a show ?thesis by simp
qed
lemma subword_self :
"subword w 0 (length w) = w"
by(simp add: subword_def)
lemma take_split[rule_format] :
"\<forall>n m. n \<le> length xs \<longrightarrow> n \<le> m \<longrightarrow>
take n xs \<cdot> take (m - n) (drop n xs) = take m xs"
by(induct_tac xs, clarsimp+, case_tac n, simp_all, case_tac m, simp_all)
lemma subword_split :
"i + j \<le> length w \<Longrightarrow> 0 < k \<Longrightarrow> k < j \<Longrightarrow>
subword w i j = subword w i k \<cdot> subword w (i + k) (j - k)"
by(simp add: subword_def, subst take_split[where n=k, THEN sym], simp_all,
rule_tac f="\<lambda>x. take (j - k) (drop x w)" in arg_cong, simp)
lemma subword_split2 :
assumes A: "subword w i j = l \<cdot> r"
and B: "i + j \<le> length w"
and C: "0 < length l"
and D: "0 < length r"
shows "l = subword w i (length l) \<and> r = subword w (i + length l) (j - length l)"
proof -
have a: "length(subword w i j) = j" by(rule subword_length, rule B)
note arg_cong[where f=length, OF A]
with a and D have b: "length l < j" by force
with B have c: "i + length l \<le> length w" by force
have "subword w i j = subword w i (length l) \<cdot> subword w (i + length l) (j - length l)"
by(rule subword_split, rule B, rule C, rule b)
with A have d: "l \<cdot> r = subword w i (length l) \<cdot> subword w (i + length l) (j - length l)" by simp
show ?thesis
by(rule append_eq_append_conv[THEN iffD1], subst subword_length, rule c, simp, rule d)
qed
subsection \<open>Properties of @{term "CYK"}\<close>
lemma CYK_Lang :
"(S \<in> CYK G w 0 (length w)) = (w \<in> Lang G S)"
by(simp add: CYK_def subword_self)
lemma CYK_eq1 :
"i + 1 \<le> length w \<Longrightarrow>
CYK G w i 1 = {S. (S, Leaf (w!i)) \<in> set G}"
by(simp add: CYK_def, subst subword_nth2[simplified], assumption,
subst Lang_eq1, rule refl)
theorem CYK_eq2 :
assumes A: "i + j \<le> length w"
and B: "1 < j"
shows "CYK G w i j = {X | X A B k. (X, Branch A B) \<in> set G \<and> A \<in> CYK G w i k \<and> B \<in> CYK G w (i + k) (j - k) \<and> 1 \<le> k \<and> k < j}"
proof(rule set_eqI, rule iffI, simp_all add: CYK_def)
fix X
assume a: "subword w i j \<in> Lang G X"
show "\<exists>A B. (X, Branch A B) \<in> set G \<and> (\<exists>k. subword w i k \<in> Lang G A \<and> subword w (i + k) (j - k) \<in> Lang G B \<and> Suc 0 \<le> k \<and> k < j)"
proof -
have b: "1 < length(subword w i j)" by(subst subword_length, rule A, rule B)
note Lang_eq2[THEN iffD1, OF conjI, OF a b]
then obtain A B l r where c: "(X, Branch A B) \<in> set G \<and> subword w i j = l \<cdot> r \<and> l \<in> Lang G A \<and> r \<in> Lang G B" by blast
note Lang_no_Nil[OF c[THEN conjunct2, THEN conjunct2, THEN conjunct1]]
hence d: "0 < length l" by(case_tac l, simp_all)
note Lang_no_Nil[OF c[THEN conjunct2, THEN conjunct2, THEN conjunct2]]
hence e: "0 < length r" by(case_tac r, simp_all)
note subword_split2[OF c[THEN conjunct2, THEN conjunct1], OF A, OF d, OF e]
with c show ?thesis
proof(rule_tac x=A in exI, rule_tac x=B in exI, simp,
rule_tac x="length l" in exI, simp)
show "Suc 0 \<le> length l \<and> length l < j" (is "?A \<and> ?B")
proof
from d show "?A" by(case_tac l, simp_all)
next
note arg_cong[where f=length, OF c[THEN conjunct2, THEN conjunct1], THEN sym]
also have "length(subword w i j) = j" by(rule subword_length, rule A)
finally have "length l + length r = j" by simp
with e show ?B by force
qed
qed
qed
next
fix X
assume "\<exists>A B. (X, Branch A B) \<in> set G \<and> (\<exists>k. subword w i k \<in> Lang G A \<and> subword w (i + k) (j - k) \<in> Lang G B \<and> Suc 0 \<le> k \<and> k < j)"
then obtain A B k where a: "(X, Branch A B) \<in> set G \<and> subword w i k \<in> Lang G A \<and> subword w (i + k) (j - k) \<in> Lang G B \<and> Suc 0 \<le> k \<and> k < j" by blast
show "subword w i j \<in> Lang G X"
proof(rule Lang_eq2[THEN iffD2, THEN conjunct1], rule_tac x=A in exI, rule_tac x=B in exI, simp add: a,
rule_tac x="subword w i k" in exI, rule_tac x="subword w (i + k) (j - k)" in exI, simp add: a,
rule subword_split, rule A)
from a show "0 < k" by force
next
from a show "k < j" by simp
qed
qed
section "Implementation"
text "One of the particularly interesting features of CYK implementation
is that it follows the principles of dynamic programming, constructing a
table of solutions for sub-problems in the bottom-up style reusing already
stored results."
subsection "Main cycle"
text "This is an auxiliary implementation of the membership test on lists."
fun mem :: "'a \<Rightarrow> 'a list \<Rightarrow> bool"
where
"mem a [] = False" |
"mem a (x#xs) = (x = a \<or> mem a xs)"
lemma mem[simp] :
"mem x xs = (x \<in> set xs)"
by(induct_tac xs, simp, force)
text "The purpose of the following is to collect non-terminals that appear on the lhs of a production
such that the first non-terminal on its rhs appears in the first of two given lists and the second
non-terminal -- in the second list."
fun match_prods :: "('n, 't) CNG \<Rightarrow> 'n list \<Rightarrow> 'n list \<Rightarrow> 'n list"
where "match_prods [] ls rs = []" |
"match_prods ((X, Branch A B)#ps) ls rs =
(if mem A ls \<and> mem B rs then X # match_prods ps ls rs
else match_prods ps ls rs)" |
"match_prods ((X, Leaf a)#ps) ls rs = match_prods ps ls rs"
lemma match_prods :
"(X \<in> set(match_prods G ls rs)) =
(\<exists>A \<in> set ls. \<exists>B \<in> set rs. (X, Branch A B) \<in> set G)"
by(induct_tac G, clarsimp+, rename_tac l r ps, case_tac r, force+)
text "The following function is the inner cycle of the algorithm. The parameters $i$ and $j$
identify a subword starting at $i$ with the length $j$, whereas $k$ is used to iterate through
its splits (which are of course subwords as well) all having the length greater $0$ but less than $j$.
The parameter $T$ represents a table containing CYK solutions for those splits."
function inner :: "('n, 't) CNG \<Rightarrow> (nat \<times> nat \<Rightarrow> 'n list) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'n list"
where "inner G T i k j =
(if k < j then match_prods G (T(i, k)) (T(i + k, j - k)) @ inner G T i (k + 1) j
else [])"
by pat_completeness auto
termination
by(relation "measure(\<lambda>(a, b, c, d, e). e - d)", rule wf_measure, simp)
declare inner.simps[simp del]
lemma inner :
"(X \<in> set(inner G T i k j)) =
(\<exists>l. k \<le> l \<and> l < j \<and> X \<in> set(match_prods G (T(i, l)) (T(i + l, j - l))))"
(is "?L G T i k j = ?R G T i k j")
proof(induct_tac G T i k j rule: inner.induct)
fix G T i k j
assume a: "k < j \<Longrightarrow> ?L G T i (k + 1) j = ?R G T i (k + 1) j"
show "?L G T i k j = ?R G T i k j"
proof(case_tac "k < j")
assume b: "k < j"
with a have c: "?L G T i (k + 1) j = ?R G T i (k + 1) j" by simp
show ?thesis
proof(subst inner.simps, simp add: b, rule iffI, erule disjE, rule_tac x=k in exI, simp add: b)
assume "X \<in> set(inner G T i (Suc k) j)"
with c have "?R G T i (k + 1) j" by simp
thus "?R G T i k j" by(clarsimp, rule_tac x=l in exI, simp)
next
assume "?R G T i k j"
then obtain l where d: "k \<le> l \<and> l < j \<and> X \<in> set(match_prods G (T(i, l)) (T(i + l, j - l)))" by blast
show "X \<in> set(match_prods G (T(i, k)) (T(i + k, j - k))) \<or> ?L G T i (Suc k) j"
proof(case_tac "Suc k \<le> l", rule disjI2, subst c[simplified], rule_tac x=l in exI, simp add: d,
rule disjI1)
assume "\<not> Suc k \<le> l"
with d have "l = k" by force
with d show "X \<in> set(match_prods G (T(i, k)) (T(i + k, j - k)))" by simp
qed
qed
next
assume "\<not> k < j"
thus ?thesis by(subst inner.simps, simp)
qed
qed
text\<open>Now the main part of the algorithm just iterates through all subwords up to the given length $len$,
calls @{term "inner"} on these, and stores the results in the table $T$. The length $j$ is supposed to
be greater than $1$ -- the subwords of length $1$ will be handled in the initialisation phase below.\<close>
function main :: "('n, 't) CNG \<Rightarrow> (nat \<times> nat \<Rightarrow> 'n list) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (nat \<times> nat \<Rightarrow> 'n list)"
where "main G T len i j = (let T' = T((i, j) := inner G T i 1 j) in
if i + j < len then main G T' len (i + 1) j
else if j < len then main G T' len 0 (j + 1)
else T')"
by pat_completeness auto
termination
by(relation "inv_image (less_than <*lex*> less_than) (\<lambda>(a, b, c, d, e). (c - e, c - (d + e)))", rule wf_inv_image, rule wf_lex_prod, (rule wf_less_than)+, simp_all)
declare main.simps[simp del]
lemma main :
assumes "1 < j"
and "i + j \<le> length w"
and "\<And>i' j'. j' < j \<Longrightarrow> 1 \<le> j' \<Longrightarrow> i' + j' \<le> length w \<Longrightarrow> set(T(i', j')) = CYK G w i' j'"
and "\<And>i'. i' < i \<Longrightarrow> i' + j \<le> length w \<Longrightarrow> set(T(i', j)) = CYK G w i' j"
and "1 \<le> j'"
and "i' + j' \<le> length w"
shows "set((main G T (length w) i j)(i', j')) = CYK G w i' j'"
proof -
have "\<forall>len T' w. main G T len i j = T' \<longrightarrow> length w = len \<longrightarrow> 1 < j \<longrightarrow> i + j \<le> len \<longrightarrow>
(\<forall>j' < j. \<forall>i'. 1 \<le> j' \<longrightarrow> i' + j' \<le> len \<longrightarrow> set(T(i', j')) = CYK G w i' j') \<longrightarrow>
(\<forall>i' < i. i' + j \<le> len \<longrightarrow> set(T(i', j)) = CYK G w i' j) \<longrightarrow>
(\<forall>j' \<ge> 1. \<forall>i'. i' + j' \<le> len \<longrightarrow> set(T'(i', j')) = CYK G w i' j')" (is "\<forall>len. ?P G T len i j")
proof(rule allI, induct_tac G T len i j rule: main.induct, (drule meta_spec, drule meta_mp, rule refl)+, clarify)
fix G T i j i' j'
fix w :: "'a list"
assume a: "i + j < length w \<Longrightarrow> ?P G (T((i, j) := inner G T i 1 j)) (length w) (i + 1) j"
assume b: "\<not> i + j < length w \<Longrightarrow> j < length w \<Longrightarrow> ?P G (T((i, j) := inner G T i 1 j)) (length w) 0 (j + 1)"
assume c: "1 < j"
assume d: "i + j \<le> length w"
assume e: "(1::nat) \<le> j'"
assume f: "i' + j' \<le> length w"
assume g: "\<forall>j' < j. \<forall>i'. 1 \<le> j' \<longrightarrow> i' + j' \<le> length w \<longrightarrow> set(T(i', j')) = CYK G w i' j'"
assume h: "\<forall>i' < i. i' + j \<le> length w \<longrightarrow> set(T(i', j)) = CYK G w i' j"
have inner: "set (inner G T i (Suc 0) j) = CYK G w i j"
proof(rule set_eqI, subst inner, subst match_prods, subst CYK_eq2, rule d, rule c, simp)
fix X
show "(\<exists>l\<ge>Suc 0. l < j \<and> (\<exists>A \<in> set(T(i, l)). \<exists>B \<in> set(T(i + l, j - l)). (X, Branch A B) \<in> set G)) =
(\<exists>A B. (X, Branch A B) \<in> set G \<and> (\<exists>k. A \<in> CYK G w i k \<and> B \<in> CYK G w (i + k) (j - k) \<and> Suc 0 \<le> k \<and> k < j))" (is "?L = ?R")
proof
assume "?L"
thus "?R"
proof(clarsimp, rule_tac x=A in exI, rule_tac x=B in exI, simp, rule_tac x=l in exI, simp)
fix l A B
assume i: "Suc 0 \<le> l"
assume j: "l < j"
assume k: "A \<in> set(T(i, l))"
assume l: "B \<in> set(T(i + l, j - l))"
note g[rule_format, where i'=i and j'=l]
with d i j have A: "set(T(i, l)) = CYK G w i l" by force
note g[rule_format, where i'="i + l" and j'="j - l"]
with d i j have "set(T(i + l, j - l)) = CYK G w (i + l) (j - l)" by force
with k l A show "A \<in> CYK G w i l \<and> B \<in> CYK G w (i + l) (j - l)" by simp
qed
next
assume "?R"
thus "?L"
proof(clarsimp, rule_tac x=k in exI, simp)
fix A B k
assume i: "Suc 0 \<le> k"
assume j: "k < j"
assume k: "A \<in> CYK G w i k"
assume l: "B \<in> CYK G w (i + k) (j - k)"
assume m: "(X, Branch A B) \<in> set G"
note g[rule_format, where i'=i and j'=k]
with d i j have A: "CYK G w i k = set(T(i, k))" by force
note g[rule_format, where i'="i + k" and j'="j - k"]
with d i j have "CYK G w (i + k) (j - k) = set(T(i + k, j - k))" by force
with k l A have "A \<in> set(T(i, k)) \<and> B \<in> set(T(i + k, j - k))" by simp
with m show "\<exists>A \<in> set(T(i, k)). \<exists>B \<in> set(T(i + k, j - k)). (X, Branch A B) \<in> set G" by force
qed
qed
qed (* inner *)
show "set((main G T (length w) i j)(i', j')) = CYK G w i' j'"
proof(case_tac "i + j = length w")
assume i: "i + j = length w"
show ?thesis
proof(case_tac "j < length w")
assume j: "j < length w"
show ?thesis
proof(subst main.simps, simp add: Let_def i j,
rule b[rule_format, where w=w and i'=i' and j'=j', OF _ _ refl, simplified],
simp_all add: inner)
from i show "\<not> i + j < length w" by simp
next
from c show "0 < j" by simp
next
from j show "Suc j \<le> length w" by simp
next
from e show "Suc 0 \<le> j'" by simp
next
from f show "i' + j' \<le> length w" by assumption
next
fix i'' j''
assume k: "j'' < Suc j"
assume l: "Suc 0 \<le> j''"
assume m: "i'' + j'' \<le> length w"
show "(i'' = i \<longrightarrow> j'' \<noteq> j) \<longrightarrow> set(T(i'',j'')) = CYK G w i'' j''"
proof(case_tac "j'' = j", simp_all, clarify)
assume n: "j'' = j"
assume "i'' \<noteq> i"
with i m n have "i'' < i" by simp
with n m h show "set(T(i'', j)) = CYK G w i'' j" by simp
next
assume "j'' \<noteq> j"
with k have "j'' < j" by simp
with l m g show "set(T(i'', j'')) = CYK G w i'' j''" by simp
qed
qed
next
assume "\<not> j < length w"
with i have j: "i = 0 \<and> j = length w" by simp
show ?thesis
proof(subst main.simps, simp add: Let_def j, intro conjI, clarify)
from j and inner show "set (inner G T 0 (Suc 0) (length w)) = CYK G w 0 (length w)" by simp
next
show "0 < i' \<longrightarrow> set(T(i', j')) = CYK G w i' j'"
proof
assume "0 < i'"
with j and f have "j' < j" by simp
with e g f show "set(T(i', j')) = CYK G w i' j'" by simp
qed
next
show "j' \<noteq> length w \<longrightarrow> set(T(i', j')) = CYK G w i' j'"
proof
assume "j' \<noteq> length w "
with j and f have "j' < j" by simp
with e g f show "set(T(i', j')) = CYK G w i' j'" by simp
qed
qed
qed
next
assume "i + j \<noteq> length w"
with d have i: "i + j < length w" by simp
show ?thesis
proof(subst main.simps, simp add: Let_def i,
rule a[rule_format, where w=w and i'=i' and j'=j', OF i, OF refl, simplified])
from c show "Suc 0 < j" by simp
next
from i show "Suc(i + j) \<le> length w" by simp
next
from e show "Suc 0 \<le> j'" by simp
next
from f show "i' + j' \<le> length w" by assumption
next
fix i'' j''
assume "j'' < j"
and "Suc 0 \<le> j''"
and "i'' + j'' \<le> length w"
with g show "set(T(i'', j'')) = CYK G w i'' j''" by simp
next
fix i'' assume j: "i'' < Suc i"
show "set(if i'' = i then inner G T i (Suc 0) j else T(i'', j)) = CYK G w i'' j"
proof(simp split: if_split, rule conjI, clarify, rule inner, clarify)
assume "i'' \<noteq> i"
with j have "i'' < i" by simp
with d h show "set(T(i'', j)) = CYK G w i'' j" by simp
qed
qed
qed
qed
with assms show ?thesis by force
qed
subsection "Initialisation phase"
text\<open>Similarly to @{term "match_prods"} above, here we collect non-terminals from which
the given terminal symbol can be derived.\<close>
fun init_match :: "('n, 't) CNG \<Rightarrow> 't \<Rightarrow> 'n list"
where "init_match [] t = []" |
"init_match ((X, Branch A B)#ps) t = init_match ps t" |
"init_match ((X, Leaf a)#ps) t = (if a = t then X # init_match ps t
else init_match ps t)"
lemma init_match :
"(X \<in> set(init_match G a)) =
((X, Leaf a) \<in> set G)"
by(induct_tac G a rule: init_match.induct, simp_all)
text "Representing the empty table."
definition "emptyT = (\<lambda>(i, j). [])"
text "The following function initialises the empty table for subwords of
length $1$, i.e. each symbol occurring in the given word."
fun init' :: "('n, 't) CNG \<Rightarrow> 't list \<Rightarrow> nat \<Rightarrow> nat \<times> nat \<Rightarrow> 'n list"
where "init' G [] k = emptyT" |
"init' G (t#ts) k = (init' G ts (k + 1))((k, 1) := init_match G t)"
lemma init' :
assumes "i + 1 \<le> length w"
shows "set(init' G w 0 (i, 1)) = CYK G w i 1"
proof -
have "\<forall>i. Suc i \<le> length w \<longrightarrow>
(\<forall>k. set(init' G w k (k + i, Suc 0)) = CYK G w i (Suc 0))" (is "\<forall>i. ?P i w \<longrightarrow> (\<forall>k. ?Q i k w)")
proof(induct_tac w, clarsimp+, rule conjI, clarsimp, rule set_eqI, subst init_match)
fix x w S
show "((S, Leaf x) \<in> set G) = (S \<in> CYK G (x#w) 0 (Suc 0))" by(subst CYK_eq1[simplified], simp_all)
next
fix x w i
assume a: "\<forall>i. ?P i w \<longrightarrow> (\<forall>k. ?Q i k w)"
assume b: "i \<le> length w"
show "0 < i \<longrightarrow> (\<forall>k. set(init' G w (Suc k) (k + i, Suc 0)) = CYK G (x#w) i (Suc 0))"
proof(clarify, case_tac i, simp_all, subst CYK_eq1[simplified], simp, erule subst, rule b, simp)
fix k j
assume c: "i = Suc j"
note a[rule_format, where i=j and k="Suc k"]
with b and c have "set(init' G w (Suc k) (Suc k + j, Suc 0)) = CYK G w j (Suc 0)" by simp
also with b and c have "... = {S. (S, Leaf (w ! j)) \<in> set G}" by(subst CYK_eq1[simplified], simp_all)
finally show "set(init' G w (Suc k) (Suc (k + j), Suc 0)) = {S. (S, Leaf (w ! j)) \<in> set G}" by simp
qed
qed
with assms have "\<forall>k. ?Q i k w" by simp
note this[rule_format, where k=0]
thus ?thesis by simp
qed
text\<open>The next version of initialization refines @{term "init'"} in that
it takes additional account of the cases when the given word is
empty or contains a terminal symbol that does not have any matching
production (that is, @{term "init_match"} is an empty list). No initial
table is then needed as such words can immediately be rejected.\<close>
fun init :: "('n, 't) CNG \<Rightarrow> 't list \<Rightarrow> nat \<Rightarrow> (nat \<times> nat \<Rightarrow> 'n list) option"
where "init G [] k = None" |
"init G [t] k = (case (init_match G t) of
[] \<Rightarrow> None
| xs \<Rightarrow> Some(emptyT((k, 1) := xs)))" |
"init G (t#ts) k = (case (init_match G t) of
[] \<Rightarrow> None
| xs \<Rightarrow> (case (init G ts (k + 1)) of
None \<Rightarrow> None
| Some T \<Rightarrow> Some(T((k, 1) := xs))))"
-lemma init1[rule_format] :
-"\<forall>T. init G w k = Some T \<longrightarrow>
- init' G w k = T"
-by(induct_tac G w k rule: init.induct, clarsimp+, simp split: list.split_asm, rule ext, clarsimp+,
- simp split: list.split_asm option.split_asm, rule ext, clarsimp, force)
-
+lemma init1:
+ \<open>init' G w k = T\<close> if \<open>init G w k = Some T\<close>
+ using that by (induction G w k arbitrary: T rule: init.induct)
+ (simp_all split: list.splits option.splits)
lemma init2 :
"(init G w k = None) =
(w = [] \<or> (\<exists>a \<in> set w. init_match G a = []))"
by(induct_tac G w k rule: init.induct, simp, simp split: list.split,
simp split: list.split option.split, force)
subsection \<open>The overall procedure\<close>
definition "cyk G S w = (case init G w 0 of
None \<Rightarrow> False
| Some T \<Rightarrow> let len = length w in
if len = 1 then mem S (T(0, 1))
else let T' = main G T len 0 2 in
mem S (T'(0, len)))"
theorem cyk :
"cyk G S w = (w \<in> Lang G S)"
proof(simp add: cyk_def split: option.split, simp_all add: Let_def,
rule conjI, subst init2, simp, rule conjI)
show "w = [] \<longrightarrow> [] \<notin> Lang G S" by(clarify, drule Lang_no_Nil, clarify)
next
show "(\<exists>x\<in>set w. init_match G x = []) \<longrightarrow> w \<notin> Lang G S" by(clarify, drule Lang_term, subst (asm) init_match[THEN sym], force)
next
show "\<forall>T. init G w 0 = Some T \<longrightarrow>
((length w = Suc 0 \<longrightarrow> S \<in> set(T(0, Suc 0))) \<and>
(length w \<noteq> Suc 0 \<longrightarrow> S \<in> set(main G T (length w) 0 2 (0, length w)))) =
(w \<in> Lang G S)" (is "\<forall>T. ?P T \<longrightarrow> ?L T = ?R")
proof clarify
fix T
assume a: "?P T"
hence b: "init' G w 0 = T" by(rule init1)
note init2[THEN iffD2, OF disjI1]
have c: "w \<noteq> []" by(clarify, drule init2[where G=G and k=0, THEN iffD2, OF disjI1], simp add: a)
have "?L (init' G w 0) = ?R"
proof(case_tac "length w = 1", simp_all)
assume d: "length w = Suc 0"
show "S \<in> set(init' G w 0 (0, Suc 0)) = ?R"
by(subst init'[simplified], simp add: d, subst CYK_Lang[THEN sym], simp add: d)
next
assume "length w \<noteq> Suc 0"
with c have "1 < length w" by(case_tac w, simp_all)
hence d: "Suc(Suc 0) \<le> length w" by simp
show "(S \<in> set(main G (init' G w 0) (length w) 0 2 (0, length w))) = (w \<in> Lang G S)"
proof(subst main, simp_all, rule d)
fix i' j'
assume "j' < 2" and "Suc 0 \<le> j'"
hence e: "j' = 1" by simp
assume "i' + j' \<le> length w"
with e have f: "i' + 1 \<le> length w" by simp
have "set(init' G w 0 (i', 1)) = CYK G w i' 1" by(rule init', rule f)
with e show "set(init' G w 0 (i', j')) = CYK G w i' j'" by simp
next
from d show "Suc 0 \<le> length w" by simp
next
show "(S \<in> CYK G w 0 (length w)) = (w \<in> Lang G S)" by(rule CYK_Lang)
qed
qed
with b show "?L T = ?R" by simp
qed
qed
value [code]
"let G = [(0::int, Branch 1 2), (0, Branch 2 3),
(1, Branch 2 1), (1, Leaf ''a''),
(2, Branch 3 3), (2, Leaf ''b''),
(3, Branch 1 2), (3, Leaf ''a'')]
in map (cyk G 0)
[[''b'',''a'',''a'',''b'',''a''],
[''b'',''a'',''b'',''a'']]"
end
diff --git a/thys/CakeML_Codegen/Compiler/Composition.thy b/thys/CakeML_Codegen/Compiler/Composition.thy
--- a/thys/CakeML_Codegen/Compiler/Composition.thy
+++ b/thys/CakeML_Codegen/Compiler/Composition.thy
@@ -1,1082 +1,1081 @@
section \<open>Composition of correctness results\<close>
theory Composition
imports "../Backend/CakeML_Correctness"
begin
hide_const (open) sem_env.v
text \<open>@{typ term} \<open>\<longrightarrow>\<close> @{typ nterm} \<open>\<longrightarrow>\<close> @{typ pterm} \<open>\<longrightarrow>\<close> @{typ sterm}\<close>
subsection \<open>Reflexive-transitive closure of @{thm [source=true] irules.compile_correct}.\<close>
lemma (in prules) prewrite_closed:
assumes "rs \<turnstile>\<^sub>p t \<longrightarrow> t'" "closed t"
shows "closed t'"
using assms proof induction
case (step name rhs)
thus ?case
using all_rules by force
next
case (beta c)
obtain pat rhs where "c = (pat, rhs)" by (cases c) auto
with beta have "closed_except rhs (frees pat)"
by (auto simp: closed_except_simps)
show ?case
apply (rule rewrite_step_closed[OF _ beta(2)[unfolded \<open>c = _\<close>]])
using \<open>closed_except rhs (frees pat)\<close> beta by (auto simp: closed_except_def)
qed (auto simp: closed_except_def)
corollary (in prules) prewrite_rt_closed:
assumes "rs \<turnstile>\<^sub>p t \<longrightarrow>* t'" "closed t"
shows "closed t'"
using assms
by induction (auto intro: prewrite_closed)
corollary (in irules) compile_correct_rt:
assumes "Rewriting_Pterm.compile rs \<turnstile>\<^sub>p t \<longrightarrow>* t'" "finished rs"
shows "rs \<turnstile>\<^sub>i t \<longrightarrow>* t'"
using assms proof (induction rule: rtranclp_induct)
case step
thus ?case
by (meson compile_correct rtranclp.simps)
qed auto
subsection \<open>Reflexive-transitive closure of @{thm [source=true] prules.compile_correct}.\<close>
lemma (in prules) compile_correct_rt:
assumes "Rewriting_Sterm.compile rs \<turnstile>\<^sub>s u \<longrightarrow>* u'"
shows "rs \<turnstile>\<^sub>p sterm_to_pterm u \<longrightarrow>* sterm_to_pterm u'"
using assms proof induction
case step
thus ?case
by (meson compile_correct rtranclp.simps)
qed auto
lemma srewrite_stepD:
assumes "srewrite_step rs name t"
shows "(name, t) \<in> set rs"
using assms by induct auto
lemma (in srules) srewrite_wellformed:
assumes "rs \<turnstile>\<^sub>s t \<longrightarrow> t'" "wellformed t"
shows "wellformed t'"
using assms proof induction
case (step name rhs)
hence "(name, rhs) \<in> set rs"
by (auto dest: srewrite_stepD)
thus ?case
using all_rules by (auto simp: list_all_iff)
next
case (beta cs t t')
then obtain pat rhs env where "(pat, rhs) \<in> set cs" "match pat t = Some env" "t' = subst rhs env"
by (elim rewrite_firstE)
show ?case
unfolding \<open>t' = _\<close>
proof (rule subst_wellformed)
show "wellformed rhs"
using \<open>(pat, rhs) \<in> set cs\<close> beta by (auto simp: list_all_iff)
next
show "wellformed_env env"
using \<open>match pat t = Some env\<close> beta
by (auto intro: wellformed.match)
qed
qed auto
lemma (in srules) srewrite_wellformed_rt:
assumes "rs \<turnstile>\<^sub>s t \<longrightarrow>* t'" "wellformed t"
shows "wellformed t'"
using assms
by induction (auto intro: srewrite_wellformed)
lemma vno_abs_value_to_sterm: "no_abs (value_to_sterm v) \<longleftrightarrow> vno_abs v" for v
by (induction v) (auto simp: no_abs.list_comb list_all_iff)
subsection \<open>Reflexive-transitive closure of @{thm [source=true] rules.compile_correct}.\<close>
corollary (in rules) compile_correct_rt:
assumes "compile \<turnstile>\<^sub>n u \<longrightarrow>* u'" "closed u"
shows "rs \<turnstile> nterm_to_term' u \<longrightarrow>* nterm_to_term' u'"
using assms
proof induction
case (step u' u'')
hence "rs \<turnstile> nterm_to_term' u \<longrightarrow>* nterm_to_term' u'"
by auto
also have "rs \<turnstile> nterm_to_term' u' \<longrightarrow> nterm_to_term' u''"
using step by (auto dest: rewrite_rt_closed intro!: compile_correct simp: closed_except_def)
finally show ?case .
qed auto
subsection \<open>Reflexive-transitive closure of @{thm [source=true] irules.transform_correct}.\<close>
corollary (in irules) transform_correct_rt:
assumes "transform_irule_set rs \<turnstile>\<^sub>i u \<longrightarrow>* u''" "t \<approx>\<^sub>p u" "closed t"
obtains t'' where "rs \<turnstile>\<^sub>i t \<longrightarrow>* t''" "t'' \<approx>\<^sub>p u''"
using assms proof (induction arbitrary: thesis t)
case (step u' u'')
obtain t' where "rs \<turnstile>\<^sub>i t \<longrightarrow>* t'" "t' \<approx>\<^sub>p u'"
using step by blast
obtain t'' where "rs \<turnstile>\<^sub>i t' \<longrightarrow>* t''" "t'' \<approx>\<^sub>p u''"
apply (rule transform_correct)
apply (rule \<open>transform_irule_set rs \<turnstile>\<^sub>i u' \<longrightarrow> u''\<close>)
apply (rule \<open>t' \<approx>\<^sub>p u'\<close>)
apply (rule irewrite_rt_closed)
apply (rule \<open>rs \<turnstile>\<^sub>i t \<longrightarrow>* t'\<close>)
apply (rule \<open>closed t\<close>)
apply blast
done
show ?case
apply (rule step.prems)
apply (rule rtranclp_trans)
apply fact+
done
qed blast
corollary (in irules) transform_correct_rt_no_abs:
assumes "transform_irule_set rs \<turnstile>\<^sub>i t \<longrightarrow>* u" "closed t" "no_abs u"
shows "rs \<turnstile>\<^sub>i t \<longrightarrow>* u"
proof -
have "t \<approx>\<^sub>p t" by (rule prelated_refl)
obtain t' where "rs \<turnstile>\<^sub>i t \<longrightarrow>* t'" "t' \<approx>\<^sub>p u"
apply (rule transform_correct_rt)
apply (rule assms)
apply (rule \<open>t \<approx>\<^sub>p t\<close>)
apply (rule assms)
apply blast
done
thus ?thesis
using assms by (metis prelated_no_abs_right)
qed
corollary transform_correct_rt_n_no_abs0:
assumes "irules C rs" "(transform_irule_set ^^ n) rs \<turnstile>\<^sub>i t \<longrightarrow>* u" "closed t" "no_abs u"
shows "rs \<turnstile>\<^sub>i t \<longrightarrow>* u"
using assms(1,2) proof (induction n arbitrary: rs)
case (Suc n)
interpret irules C rs by fact
show ?case
apply (rule transform_correct_rt_no_abs)
apply (rule Suc.IH)
apply (rule rules_transform)
using Suc(3) apply (simp add: funpow_swap1)
apply fact+
done
qed auto
corollary (in irules) transform_correct_rt_n_no_abs:
assumes "(transform_irule_set ^^ n) rs \<turnstile>\<^sub>i t \<longrightarrow>* u" "closed t" "no_abs u"
shows "rs \<turnstile>\<^sub>i t \<longrightarrow>* u"
by (rule transform_correct_rt_n_no_abs0) (rule irules_axioms assms)+
hide_fact transform_correct_rt_n_no_abs0
subsection \<open>Iterated application of @{const transform_irule_set}.\<close>
definition max_arity :: "irule_set \<Rightarrow> nat" where
"max_arity rs = fMax ((arity \<circ> snd) |`| rs)"
lemma rules_transform_iter0:
assumes "irules C_info rs"
shows "irules C_info ((transform_irule_set ^^ n) rs)"
using assms
by (induction n) (auto intro: irules.rules_transform del: irulesI)
lemma (in irules) rules_transform_iter: "irules C_info ((transform_irule_set ^^ n) rs)"
by (rule rules_transform_iter0) (rule irules_axioms)
lemma transform_irule_set_n_heads: "fst |`| ((transform_irule_set ^^ n) rs) = fst |`| rs"
by (induction n) (auto simp: transform_irule_set_heads)
hide_fact rules_transform_iter0
definition transform_irule_set_iter :: "irule_set \<Rightarrow> irule_set" where
"transform_irule_set_iter rs = (transform_irule_set ^^ max_arity rs) rs"
lemma transform_irule_set_iter_heads: "fst |`| transform_irule_set_iter rs = fst |`| rs"
unfolding transform_irule_set_iter_def by (simp add: transform_irule_set_n_heads)
lemma (in irules) finished_alt_def: "finished rs \<longleftrightarrow> max_arity rs = 0"
proof
assume "max_arity rs = 0"
hence "\<not> fBex ((arity \<circ> snd) |`| rs) (\<lambda>x. 0 < x)"
using nonempty
unfolding max_arity_def
by (metis fBex_fempty fmax_ex_gr not_less0)
thus "finished rs"
unfolding finished_def
by force
next
assume "finished rs"
have "fMax ((arity \<circ> snd) |`| rs) \<le> 0"
proof (rule fMax_le)
show "fBall ((arity \<circ> snd) |`| rs) (\<lambda>x. x \<le> 0)"
using \<open>finished rs\<close> unfolding finished_def by force
next
show "(arity \<circ> snd) |`| rs \<noteq> {||}"
using nonempty by force
qed
thus "max_arity rs = 0"
unfolding max_arity_def by simp
qed
lemma (in irules) transform_finished_id: "finished rs \<Longrightarrow> transform_irule_set rs = rs"
unfolding transform_irule_set_def finished_def transform_irules_def map_prod_def id_apply
by (rule fset_map_snd_id) (auto simp: fmember.rep_eq elim!: fBallE)
lemma (in irules) max_arity_decr: "max_arity (transform_irule_set rs) = max_arity rs - 1"
proof (cases "finished rs")
case True
thus ?thesis
by (auto simp: transform_finished_id finished_alt_def)
next
case False
have "(arity \<circ> snd) |`| transform_irule_set rs = (\<lambda>x. x - 1) |`| (arity \<circ> snd) |`| rs"
unfolding transform_irule_set_def fset.map_comp
proof (rule fset.map_cong0, safe, unfold o_apply map_prod_simp id_apply snd_conv)
fix name irs
assume "(name, irs) \<in> fset rs"
hence "(name, irs) |\<in>| rs"
by (simp add: fmember.rep_eq)
hence "arity_compatibles irs" "irs \<noteq> {||}"
using nonempty inner by (blast dest: fpairwiseD)+
thus "arity (transform_irules irs) = arity irs - 1"
by (simp add: arity_transform_irules)
qed
hence "max_arity (transform_irule_set rs) = fMax ((\<lambda>x. x - 1) |`| (arity \<circ> snd) |`| rs)"
unfolding max_arity_def by simp
also have "\<dots> = fMax ((arity \<circ> snd) |`| rs) - 1"
proof (rule fmax_decr)
show "fBex ((arity \<circ> snd) |`| rs) ((\<le>) 1)"
using False unfolding finished_def by force
qed
finally show ?thesis
unfolding max_arity_def
by simp
qed
lemma max_arity_decr'0:
assumes "irules C rs"
shows "max_arity ((transform_irule_set ^^ n) rs) = max_arity rs - n"
proof (induction n)
case (Suc n)
show ?case
apply simp
apply (subst irules.max_arity_decr)
using Suc assms by (auto intro: irules.rules_transform_iter del: irulesI)
qed auto
lemma (in irules) max_arity_decr': "max_arity ((transform_irule_set ^^ n) rs) = max_arity rs - n"
by (rule max_arity_decr'0) (rule irules_axioms)
hide_fact max_arity_decr'0
lemma (in irules) transform_finished: "finished (transform_irule_set_iter rs)"
unfolding transform_irule_set_iter_def
by (subst irules.finished_alt_def)
(auto simp: max_arity_decr' intro: rules_transform_iter del: Rewriting_Pterm_Elim.irulesI)
text \<open>Trick as described in \<open>\<section>7.1\<close> in the locale manual.\<close>
locale irules' = irules
sublocale irules' \<subseteq> irules'_as_irules: irules C_info "transform_irule_set_iter rs"
unfolding transform_irule_set_iter_def by (rule rules_transform_iter)
sublocale crules \<subseteq> crules_as_irules': irules' C_info "Rewriting_Pterm_Elim.compile rs"
unfolding irules'_def by (fact compile_rules)
sublocale irules' \<subseteq> irules'_as_prules: prules C_info "Rewriting_Pterm.compile (transform_irule_set_iter rs)"
by (rule irules'_as_irules.compile_rules) (rule transform_finished)
subsection \<open>Big-step semantics\<close>
context srules begin
definition global_css :: "(name, sclauses) fmap" where
"global_css = fmap_of_list (map (map_prod id clauses) rs)"
lemma fmdom_global_css: "fmdom global_css = fst |`| fset_of_list rs"
unfolding global_css_def by simp
definition as_vrules :: "vrule list" where
"as_vrules = map (\<lambda>(name, _). (name, Vrecabs global_css name fmempty)) rs"
lemma as_vrules_fst[simp]: "fst |`| fset_of_list as_vrules = fst |`| fset_of_list rs"
unfolding as_vrules_def
apply simp
apply (rule fset.map_cong)
apply (rule refl)
by auto
lemma as_vrules_fst'[simp]: "map fst as_vrules = map fst rs"
unfolding as_vrules_def
by auto
lemma list_all_as_vrulesI:
assumes "list_all (\<lambda>(_, t). P fmempty (clauses t)) rs"
assumes "R (fst |`| fset_of_list rs)"
shows "list_all (\<lambda>(_, t). value_pred.pred P Q R t) as_vrules"
proof (rule list_allI, safe)
fix name rhs
assume "(name, rhs) \<in> set as_vrules"
hence "rhs = Vrecabs global_css name fmempty"
unfolding as_vrules_def by auto
moreover have "fmpred (\<lambda>_. P fmempty) global_css"
unfolding global_css_def list.pred_map
using assms by (auto simp: list_all_iff intro!: fmpred_of_list)
moreover have "name |\<in>| fmdom global_css"
unfolding global_css_def
apply auto
using \<open>(name, rhs) \<in> set as_vrules\<close> unfolding as_vrules_def
including fset.lifting apply transfer'
by force
moreover have "R (fmdom global_css)"
using assms unfolding global_css_def
by auto
ultimately show "value_pred.pred P Q R rhs"
by (simp add: value_pred.pred_alt_def)
qed
lemma srules_as_vrules: "vrules C_info as_vrules"
proof (standard, unfold as_vrules_fst)
have "list_all (\<lambda>(_, t). vwellformed t) as_vrules"
unfolding vwellformed_def
apply (rule list_all_as_vrulesI)
apply (rule list.pred_mono_strong)
apply (rule all_rules)
apply (auto elim: clausesE)
done
moreover have "list_all (\<lambda>(_, t). vclosed t) as_vrules"
unfolding vclosed_def
apply (rule list_all_as_vrulesI)
apply auto
apply (rule list.pred_mono_strong)
apply (rule all_rules)
apply (auto elim: clausesE simp: Sterm.closed_except_simps)
done
moreover have "list_all (\<lambda>(_, t). \<not> is_Vconstr t) as_vrules"
unfolding as_vrules_def
by (auto simp: list_all_iff)
ultimately show "list_all vrule as_vrules"
unfolding list_all_iff by fastforce
next
show "distinct (map fst as_vrules)"
using distinct by auto
next
show "fdisjnt (fst |`| fset_of_list rs) C"
using disjnt by simp
next
show "list_all (\<lambda>(_, rhs). not_shadows_vconsts rhs) as_vrules"
unfolding not_shadows_vconsts_def
apply (rule list_all_as_vrulesI)
apply auto
apply (rule list.pred_mono_strong)
apply (rule not_shadows)
by (auto simp: list_all_iff list_ex_iff all_consts_def elim!: clausesE)
next
show "vconstructor_value_rs as_vrules"
unfolding vconstructor_value_rs_def
apply (rule conjI)
unfolding vconstructor_value_def
apply (rule list_all_as_vrulesI)
apply (simp add: list_all_iff)
apply simp
apply simp
using disjnt by simp
next
show "list_all (\<lambda>(_, rhs). vwelldefined rhs) as_vrules"
unfolding vwelldefined_def
apply (rule list_all_as_vrulesI)
apply auto
apply (rule list.pred_mono_strong)
apply (rule swelldefined_rs)
apply auto
apply (erule clausesE)
apply hypsubst_thin
apply (subst (asm) welldefined_sabs)
by simp
next
show "distinct all_constructors"
by (fact distinct_ctr)
qed
sublocale srules_as_vrules: vrules C_info as_vrules
by (fact srules_as_vrules)
lemma rs'_rs_eq: "srules_as_vrules.rs' = rs"
unfolding srules_as_vrules.rs'_def
unfolding as_vrules_def
apply (subst map_prod_def)
apply simp
unfolding comp_def
apply (subst case_prod_twice)
apply (rule list_map_snd_id)
unfolding global_css_def
using all_rules map
apply (auto simp: list_all_iff map_of_is_map map_of_map map_prod_def fmap_of_list.rep_eq)
subgoal for a b
by (erule ballE[where x = "(a, b)"], cases b, auto simp: is_abs_def term_cases_def)
done
lemma veval_correct:
fixes v
assumes "as_vrules, fmempty \<turnstile>\<^sub>v t \<down> v" "wellformed t" "closed t"
shows "rs, fmempty \<turnstile>\<^sub>s t \<down> value_to_sterm v"
using assms
by (rule srules_as_vrules.veval_correct[unfolded rs'_rs_eq])
end
subsection \<open>ML-style semantics\<close>
context srules begin
lemma as_vrules_mk_rec_env: "fmap_of_list as_vrules = mk_rec_env global_css fmempty"
apply (subst global_css_def)
apply (subst as_vrules_def)
apply (subst mk_rec_env_def)
apply (rule fmap_ext)
apply (subst fmlookup_fmmap_keys)
apply (subst fmap_of_list.rep_eq)
apply (subst fmap_of_list.rep_eq)
apply (subst map_of_map_keyed)
apply (subst (2) map_prod_def)
apply (subst id_apply)
apply (subst map_of_map)
apply simp
apply (subst option.map_comp)
apply (rule option.map_cong)
apply (rule refl)
apply simp
apply (subst global_css_def)
apply (rule refl)
done
abbreviation (input) "vrelated \<equiv> srules_as_vrules.vrelated"
notation srules_as_vrules.vrelated ("\<turnstile>\<^sub>v/ _ \<approx> _" [0, 50] 50)
lemma vrecabs_global_css_refl:
assumes "name |\<in>| fmdom global_css"
shows "\<turnstile>\<^sub>v Vrecabs global_css name fmempty \<approx> Vrecabs global_css name fmempty"
using assms
proof (coinduction arbitrary: name)
case vrelated
have "rel_option (\<lambda>v\<^sub>1 v\<^sub>2. (\<exists>name. v\<^sub>1 = Vrecabs global_css name fmempty \<and> v\<^sub>2 = Vrecabs global_css name fmempty \<and> name |\<in>| fmdom global_css) \<or> \<turnstile>\<^sub>v v\<^sub>1 \<approx> v\<^sub>2) (fmlookup (fmap_of_list as_vrules) y) (fmlookup (mk_rec_env global_css fmempty) y)" for y
apply (subst as_vrules_mk_rec_env)
apply (rule option.rel_refl_strong)
apply (rule disjI1)
apply (simp add: mk_rec_env_def)
apply (elim conjE exE)
- apply (intro exI conjI)
by (auto intro: fmdomI)
with vrelated show ?case
by fastforce
qed
lemma as_vrules_refl_rs: "fmrel_on_fset (fst |`| fset_of_list as_vrules) vrelated (fmap_of_list as_vrules) (fmap_of_list as_vrules)"
apply rule
apply (subst (2) as_vrules_def)
apply (subst (2) as_vrules_def)
apply (simp add: fmap_of_list.rep_eq)
apply (rule rel_option_reflI)
apply simp
apply (drule map_of_SomeD)
apply auto
apply (rule vrecabs_global_css_refl)
unfolding global_css_def
by (auto simp: fset_of_list_elem intro: rev_fimage_eqI)
lemma as_vrules_refl_C: "fmrel_on_fset C vrelated (fmap_of_list as_vrules) (fmap_of_list as_vrules)"
proof
fix c
assume "c |\<in>| C"
hence "c |\<notin>| fset_of_list (map fst as_vrules)"
using srules_as_vrules.vconstructor_value_rs
unfolding vconstructor_value_rs_def fdisjnt_alt_def
by auto
hence "c |\<notin>| fmdom (fmap_of_list as_vrules)"
by simp
hence "fmlookup (fmap_of_list as_vrules) c = None"
by (metis fmdom_notD)
thus "rel_option vrelated (fmlookup (fmap_of_list as_vrules) c) (fmlookup (fmap_of_list as_vrules) c)"
by simp
qed
lemma veval'_correct'':
fixes t v
assumes "fmap_of_list as_vrules \<turnstile>\<^sub>v t \<down> v"
assumes "wellformed t"
assumes "\<not> shadows_consts t"
assumes "welldefined t"
assumes "closed t"
assumes "vno_abs v"
shows "as_vrules, fmempty \<turnstile>\<^sub>v t \<down> v"
proof -
obtain v\<^sub>1 where "as_vrules, fmempty \<turnstile>\<^sub>v t \<down> v\<^sub>1" "\<turnstile>\<^sub>v v\<^sub>1 \<approx> v"
using \<open>fmap_of_list as_vrules \<turnstile>\<^sub>v t \<down> v\<close>
proof (rule srules_as_vrules.veval'_correct', unfold as_vrules_fst)
show "wellformed t" "\<not> shadows_consts t" "closed t" "consts t |\<subseteq>| all_consts"
by fact+
next
show "wellformed_venv (fmap_of_list as_vrules)"
apply rule
using srules_as_vrules.all_rules
apply (auto simp: list_all_iff)
done
next
show "not_shadows_vconsts_env (fmap_of_list as_vrules) "
apply rule
using srules_as_vrules.not_shadows
apply (auto simp: list_all_iff)
done
next
have "fmrel_on_fset (fst |`| fset_of_list as_vrules |\<union>| C) vrelated (fmap_of_list as_vrules) (fmap_of_list as_vrules)"
apply (rule fmrel_on_fset_unionI)
apply (rule as_vrules_refl_rs)
apply (rule as_vrules_refl_C)
done
show "fmrel_on_fset (consts t) vrelated (fmap_of_list as_vrules) (fmap_of_list as_vrules)"
apply (rule fmrel_on_fsubset)
apply fact+
using assms by (auto simp: all_consts_def)
qed
thus ?thesis
using assms by (metis srules_as_vrules.vrelated.eq_right)
qed
end
subsection \<open>CakeML\<close>
context srules begin
definition as_sem_env :: "v sem_env \<Rightarrow> v sem_env" where
"as_sem_env env = \<lparr> sem_env.v = build_rec_env (mk_letrec_body all_consts rs) env nsEmpty, sem_env.c = nsEmpty \<rparr>"
lemma compile_sem_env:
"evaluate_dec ck mn env state (compile_group all_consts rs) (state, Rval (as_sem_env env))"
unfolding compile_group_def as_sem_env_def
apply (rule evaluate_dec.dletrec1)
unfolding mk_letrec_body_def Let_def
apply (simp add:comp_def case_prod_twice)
using name_as_string.fst_distinct[OF distinct]
by auto
lemma compile_sem_env':
"fun_evaluate_decs mn state env [(compile_group all_consts rs)] = (state, Rval (as_sem_env env))"
unfolding compile_group_def as_sem_env_def mk_letrec_body_def Let_def
apply (simp add: comp_def case_prod_twice)
using name_as_string.fst_distinct[OF distinct]
by auto
lemma compile_prog[unfolded combine_dec_result.simps, simplified]:
"evaluate_prog ck env state (compile rs) (state, combine_dec_result (as_sem_env env) (Rval \<lparr> sem_env.v = nsEmpty, sem_env.c = nsEmpty \<rparr>))"
unfolding compile_def
apply (rule evaluate_prog.cons1)
apply rule
apply (rule evaluate_top.tdec1)
apply (rule compile_sem_env)
apply (rule evaluate_prog.empty)
done
lemma compile_prog'[unfolded combine_dec_result.simps, simplified]:
"fun_evaluate_prog state env (compile rs) = (state, combine_dec_result (as_sem_env env) (Rval \<lparr> sem_env.v = nsEmpty, sem_env.c = nsEmpty \<rparr>))"
unfolding compile_def fun_evaluate_prog_def no_dup_mods_def no_dup_top_types_def prog_to_mods_def prog_to_top_types_def decs_to_types_def
using compile_sem_env' compile_group_def by simp
definition sem_env :: "v sem_env" where
"sem_env \<equiv> extend_dec_env (as_sem_env empty_sem_env) empty_sem_env"
(* FIXME introduce lemma: is_cupcake_all_env extend_dec_env *)
(* FIXME introduce lemma: is_cupcake_all_env empty_sem_env *)
lemma cupcake_sem_env: "is_cupcake_all_env sem_env"
unfolding as_sem_env_def sem_env_def
apply (rule is_cupcake_all_envI)
apply (simp add: extend_dec_env_def empty_sem_env_def nsEmpty_def)
apply (rule cupcake_nsAppend_preserve)
apply (rule cupcake_build_rec_preserve)
apply (simp add: empty_sem_env_def)
apply (simp add: nsEmpty_def)
apply (rule mk_letrec_cupcake)
apply simp
apply (simp add: empty_sem_env_def)
done
lemma sem_env_refl: "fmrel related_v (fmap_of_list as_vrules) (fmap_of_ns (sem_env.v sem_env))"
proof
fix name
show "rel_option related_v (fmlookup (fmap_of_list as_vrules) name) (fmlookup (fmap_of_ns (sem_env.v sem_env)) name)"
apply (simp add:
as_sem_env_def build_rec_env_fmap cake_mk_rec_env_def sem_env_def
fmap_of_list.rep_eq map_of_map_keyed option.rel_map
as_vrules_def mk_letrec_body_def comp_def case_prod_twice)
apply (rule option.rel_refl_strong)
apply (rule related_v.rec_closure)
apply auto[]
apply (simp add:
fmmap_of_list[symmetric, unfolded apsnd_def map_prod_def id_def] fmap.rel_map
global_css_def Let_def map_prod_def comp_def case_prod_twice)
apply (thin_tac "map_of rs name = _")
apply (rule fmap.rel_refl_strong)
apply simp
subgoal premises prems for rhs
proof -
obtain name where "(name, rhs) \<in> set rs"
using prems
including fmap.lifting
by transfer' (auto dest: map_of_SomeD)
hence "is_abs rhs" "closed rhs" "welldefined rhs"
using all_rules swelldefined_rs by (auto simp add: list_all_iff)
then obtain cs where "clauses rhs = cs" "rhs = Sabs cs" "wellformed_clauses cs"
using \<open>(name, rhs) \<in> set rs\<close> all_rules
by (cases rhs) (auto simp: list_all_iff is_abs_def term_cases_def)
show ?thesis
unfolding related_fun_alt_def \<open>clauses rhs = cs\<close>
proof (intro conjI)
show "list_all2 (rel_prod related_pat related_exp) cs (map (\<lambda>(pat, t). (mk_ml_pat (mk_pat pat), mk_con (frees pat |\<union>| all_consts) t)) cs)"
unfolding list.rel_map
apply (rule list.rel_refl_strong)
apply (rename_tac z, case_tac z, hypsubst_thin)
apply simp
subgoal premises prems for pat t
proof (rule mk_exp_correctness)
have "\<not> shadows_consts rhs"
using \<open>(name, rhs) \<in> set rs\<close> not_shadows
by (auto simp: list_all_iff all_consts_def)
thus "\<not> shadows_consts t"
unfolding \<open>rhs = Sabs cs\<close> using prems
by (auto simp: list_all_iff list_ex_iff)
next
have "frees t |\<subseteq>| frees pat"
using \<open>closed rhs\<close> prems unfolding \<open>rhs = _\<close>
apply (auto simp: list_all_iff Sterm.closed_except_simps)
apply (erule ballE[where x = "(pat, t)"])
apply (auto simp: closed_except_def)
done
moreover have "consts t |\<subseteq>| all_consts"
using \<open>welldefined rhs\<close> prems unfolding \<open>rhs = _\<close> welldefined_sabs
by (auto simp: list_all_iff all_consts_def)
ultimately show "ids t |\<subseteq>| frees pat |\<union>| all_consts"
unfolding ids_def by auto
qed (auto simp: all_consts_def)
done
next
have 1: "frees (Sabs cs) = {||}"
using \<open>closed rhs\<close> unfolding \<open>rhs = Sabs cs\<close>
by (auto simp: closed_except_def)
have 2: "welldefined rhs"
using swelldefined_rs \<open>(name, rhs) \<in> set rs\<close>
by (auto simp: list_all_iff)
show "fresh_fNext all_consts |\<notin>| ids (Sabs cs)"
apply (rule fNext_not_member_subset)
unfolding ids_def 1
using 2 \<open>rhs = _\<close> by (simp add: all_consts_def del: consts_sterm.simps)
next
show "fresh_fNext all_consts |\<notin>| all_consts"
by (rule fNext_not_member)
qed
qed
done
qed
lemma semantic_correctness':
assumes "cupcake_evaluate_single sem_env (mk_con all_consts t) (Rval ml_v)"
assumes "welldefined t" "closed t" "\<not> shadows_consts t" "wellformed t"
obtains v where "fmap_of_list as_vrules \<turnstile>\<^sub>v t \<down> v" "related_v v ml_v"
using assms(1) proof (rule semantic_correctness)
show "is_cupcake_all_env sem_env"
by (fact cupcake_sem_env)
next
show "related_exp t (mk_con all_consts t)"
apply (rule mk_exp_correctness)
using assms
unfolding ids_def closed_except_def by (auto simp: all_consts_def)
next
show "wellformed t" "\<not> shadows_consts t" by fact+
next
show "closed_except t (fmdom (fmap_of_list as_vrules))"
using \<open>closed t\<close> by (auto simp: closed_except_def)
next
show "closed_venv (fmap_of_list as_vrules)"
apply (rule fmpred_of_list)
using srules_as_vrules.all_rules
by (auto simp: list_all_iff)
show "wellformed_venv (fmap_of_list as_vrules)"
apply (rule fmpred_of_list)
using srules_as_vrules.all_rules
by (auto simp: list_all_iff)
next
have 1: "fmpred (\<lambda>_. list_all (\<lambda>(pat, t). consts t |\<subseteq>| C |\<union>| fmdom global_css)) global_css"
apply (subst (2) global_css_def)
apply (rule fmpred_of_list)
apply (auto simp: map_prod_def)
subgoal premises prems for pat t
proof -
from prems obtain cs where "t = Sabs cs"
by (elim clausesE)
have "welldefined t"
using swelldefined_rs prems
by (auto simp: list_all_iff fmdom_global_css)
show ?thesis
using \<open>welldefined t\<close>
unfolding \<open>t = _\<close> welldefined_sabs
by (auto simp: all_consts_def list_all_iff fmdom_global_css)
qed
done
show "fmpred (\<lambda>_. vwelldefined') (fmap_of_list as_vrules)"
apply (rule fmpred_of_list)
unfolding as_vrules_def
apply simp
apply (erule imageE)
apply (auto split: prod.splits)
apply (subst fdisjnt_alt_def)
apply simp
apply (rule 1)
apply (subst global_css_def)
apply simp
subgoal for x1 x2
apply (rule fimage_eqI[where x = "(x1, x2)"])
by (auto simp: fset_of_list_elem)
subgoal
using disjnt by (auto simp: fdisjnt_alt_def fmdom_global_css)
done
next
show "not_shadows_vconsts_env (fmap_of_list as_vrules)"
apply (rule fmpred_of_list)
using srules_as_vrules.not_shadows
unfolding list_all_iff
by auto
next
show "fdisjnt C (fmdom (fmap_of_list as_vrules))"
using disjnt by (auto simp: fdisjnt_alt_def)
next
show "fmrel_on_fset (ids t) related_v (fmap_of_list as_vrules) (fmap_of_ns (sem_env.v sem_env))"
unfolding fmrel_on_fset_fmrel_restrict
apply (rule fmrel_restrict_fset)
apply (rule sem_env_refl)
done
next
show "consts t |\<subseteq>| fmdom (fmap_of_list as_vrules) |\<union>| C"
apply (subst fmdom_fmap_of_list)
apply (subst as_vrules_fst')
apply simp
using assms by (auto simp: all_consts_def)
qed blast
end
fun cake_to_value :: "v \<Rightarrow> value" where
"cake_to_value (Conv (Some (name, _)) vs) = Vconstr (Name name) (map cake_to_value vs)"
context cakeml' begin
lemma cake_to_value_abs_free:
assumes "is_cupcake_value v" "cake_no_abs v"
shows "vno_abs (cake_to_value v)"
using assms by (induction v) (auto elim: is_cupcake_value.elims simp: list_all_iff)
lemma cake_to_value_related:
assumes "cake_no_abs v" "is_cupcake_value v"
shows "related_v (cake_to_value v) v"
using assms proof (induction v)
case (Conv c vs)
then obtain name tid where "c = Some ((as_string name), TypeId (Short tid))"
apply (elim is_cupcake_value.elims)
subgoal
by (metis name.sel v.simps(2))
by auto
show ?case
unfolding \<open>c = _\<close>
apply simp
apply (rule related_v.conv)
apply (simp add: list.rel_map)
apply (rule list.rel_refl_strong)
apply (rule Conv)
using Conv unfolding \<open>c = _\<close>
by (auto simp: list_all_iff)
qed auto
lemma related_v_abs_free_uniq:
assumes "related_v v\<^sub>1 ml_v" "related_v v\<^sub>2 ml_v" "cake_no_abs ml_v"
shows "v\<^sub>1 = v\<^sub>2"
using assms proof (induction arbitrary: v\<^sub>2)
case (conv vs\<^sub>1 ml_vs name)
then obtain vs\<^sub>2 where "v\<^sub>2 = Vconstr name vs\<^sub>2" "list_all2 related_v vs\<^sub>2 ml_vs"
by (auto elim: related_v.cases simp: name.expand)
moreover have "list_all cake_no_abs ml_vs"
using conv by simp
have "list_all2 (=) vs\<^sub>1 vs\<^sub>2"
using \<open>list_all2 _ vs\<^sub>1 _\<close> \<open>list_all2 _ vs\<^sub>2 _\<close> \<open>list_all cake_no_abs ml_vs\<close>
by (induction arbitrary: vs\<^sub>2 rule: list.rel_induct) (auto simp: list_all2_Cons2)
thus ?case
unfolding \<open>v\<^sub>2 = _\<close>
by (simp add: list.rel_eq)
qed auto
corollary related_v_abs_free_cake_to_value:
assumes "related_v v ml_v" "cake_no_abs ml_v" "is_cupcake_value ml_v"
shows "v = cake_to_value ml_v"
using assms by (metis cake_to_value_related related_v_abs_free_uniq)
end
context srules begin
lemma cupcake_sem_env_preserve:
assumes "cupcake_evaluate_single sem_env (mk_con S t) (Rval ml_v)" "wellformed t"
shows "is_cupcake_value ml_v"
apply (rule cupcake_single_preserve[OF assms(1)])
apply (rule cupcake_sem_env)
apply (rule mk_exp_cupcake)
apply fact
done
lemma semantic_correctness'':
assumes "cupcake_evaluate_single sem_env (mk_con all_consts t) (Rval ml_v)"
assumes "welldefined t" "closed t" "\<not> shadows_consts t" "wellformed t"
assumes "cake_no_abs ml_v"
shows "fmap_of_list as_vrules \<turnstile>\<^sub>v t \<down> cake_to_value ml_v"
using assms
by (metis cupcake_sem_env_preserve semantic_correctness' related_v_abs_free_cake_to_value)
end
subsection \<open>Composition\<close>
context rules begin
abbreviation term_to_nterm where
"term_to_nterm t \<equiv> fresh_frun (Term_to_Nterm.term_to_nterm [] t) all_consts"
abbreviation sterm_to_cake where
"sterm_to_cake \<equiv> rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.mk_con all_consts"
abbreviation "term_to_cake t \<equiv> sterm_to_cake (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
abbreviation "cake_to_term t \<equiv> (convert_term (value_to_sterm (cake_to_value t)) :: term)"
abbreviation "cake_sem_env \<equiv> rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.sem_env"
definition "compiled \<equiv> rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.as_vrules"
lemma fmdom_compiled: "fmdom (fmap_of_list compiled) = heads_of rs"
unfolding compiled_def
by (simp add:
rules_as_nrules.crules_as_irules'.irules'_as_prules.compile_heads
Rewriting_Pterm.compile_heads transform_irule_set_iter_heads
Rewriting_Pterm_Elim.compile_heads
compile_heads consts_of_heads)
lemma cake_semantic_correctness:
assumes "cupcake_evaluate_single cake_sem_env (sterm_to_cake t) (Rval ml_v)"
assumes "welldefined t" "closed t" "\<not> shadows_consts t" "wellformed t"
assumes "cake_no_abs ml_v"
shows "fmap_of_list compiled \<turnstile>\<^sub>v t \<down> cake_to_value ml_v"
unfolding compiled_def
apply (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.semantic_correctness'')
using assms
by (simp_all add:
rules_as_nrules.crules_as_irules'.irules'_as_prules.compile_heads
Rewriting_Pterm.compile_heads transform_irule_set_iter_heads
Rewriting_Pterm_Elim.compile_heads
compile_heads consts_of_heads all_consts_def)
text \<open>Lo and behold, this is the final correctness theorem!\<close>
theorem compiled_correct:
\<comment> \<open>If CakeML evaluation of a term succeeds ...\<close>
assumes "\<exists>k. Evaluate_Single.evaluate cake_sem_env (s \<lparr> clock := k \<rparr>) (term_to_cake t) = (s', Rval ml_v)"
\<comment> \<open>... producing a constructor term without closures ...\<close>
assumes "cake_no_abs ml_v"
\<comment> \<open>... and some syntactic properties of the involved terms hold ...\<close>
assumes "closed t" "\<not> shadows_consts t" "welldefined t" "wellformed t"
\<comment> \<open>... then this evaluation can be reproduced in the term-rewriting semantics\<close>
shows "rs \<turnstile> t \<longrightarrow>* cake_to_term ml_v"
proof -
let ?heads = "fst |`| fset_of_list rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.as_vrules"
have "?heads = heads_of rs"
using fmdom_compiled unfolding compiled_def by simp
have "wellformed (nterm_to_pterm (term_to_nterm t))"
by auto
hence "wellformed (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
by (auto intro: pterm_to_sterm_wellformed)
have "is_cupcake_all_env cake_sem_env"
by (rule rules_as_nrules.nrules_as_crules.crules_as_irules'.irules'_as_prules.prules_as_srules.cupcake_sem_env)
have "is_cupcake_exp (term_to_cake t)"
by (rule rules_as_nrules.nrules_as_crules.crules_as_irules'.irules'_as_prules.prules_as_srules.srules_as_cake.mk_exp_cupcake) fact
obtain k where "Evaluate_Single.evaluate cake_sem_env (s \<lparr> clock := k \<rparr>) (term_to_cake t) = (s', Rval ml_v)"
using assms by blast
then have "Big_Step_Unclocked_Single.evaluate cake_sem_env (s \<lparr> clock := (clock s') \<rparr>) (term_to_cake t) (s', Rval ml_v)"
using unclocked_single_fun_eq by fastforce
have "cupcake_evaluate_single cake_sem_env (sterm_to_cake (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))) (Rval ml_v)"
apply (rule cupcake_single_complete)
apply fact+
done
hence "is_cupcake_value ml_v"
apply (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.cupcake_sem_env_preserve)
by (auto intro: pterm_to_sterm_wellformed)
hence "vno_abs (cake_to_value ml_v)"
using \<open>cake_no_abs _\<close>
by (metis rules_as_nrules.nrules_as_crules.crules_as_irules'.irules'_as_prules.prules_as_srules.srules_as_cake.cake_to_value_abs_free)
hence "no_abs (value_to_sterm (cake_to_value ml_v))"
by (metis vno_abs_value_to_sterm)
hence "no_abs (sterm_to_pterm (value_to_sterm (cake_to_value ml_v)))"
by (metis sterm_to_pterm convert_term_no_abs)
have "welldefined (term_to_nterm t)"
unfolding term_to_nterm'_def
apply (subst fresh_frun_def)
apply (subst pred_stateD[OF term_to_nterm_consts])
apply (subst surjective_pairing)
apply (rule refl)
apply fact
done
have "welldefined (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
apply (subst pterm_to_sterm_consts)
apply fact
apply (subst consts_nterm_to_pterm)
apply fact+
done
have "\<not> shadows_consts t"
using assms unfolding shadows_consts_def fdisjnt_alt_def
by auto
hence "\<not> shadows_consts (term_to_nterm t)"
unfolding shadows_consts_def shadows_consts_def
apply auto
using term_to_nterm_all_vars[folded wellformed_term_def]
by (metis assms(6) fdisjnt_swap sup_idem)
have "\<not> shadows_consts (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
apply (subst pterm_to_sterm_shadows[symmetric])
apply fact
apply (subst shadows_nterm_to_pterm)
unfolding shadows_consts_def
apply simp
apply (rule term_to_nterm_all_vars[where T = "fempty", simplified, THEN fdisjnt_swap])
apply (fold wellformed_term_def)
apply fact
using \<open>closed t\<close> unfolding closed_except_def by (auto simp: fdisjnt_alt_def)
have "closed (term_to_nterm t)"
using assms unfolding closed_except_def
using term_to_nterm_vars unfolding wellformed_term_def by blast
hence "closed (nterm_to_pterm (term_to_nterm t))"
using closed_nterm_to_pterm unfolding closed_except_def
by auto
have "closed (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
unfolding closed_except_def
apply (subst pterm_to_sterm_frees)
apply fact
using \<open>closed (term_to_nterm t)\<close> closed_nterm_to_pterm unfolding closed_except_def
by auto
have "fmap_of_list compiled \<turnstile>\<^sub>v pterm_to_sterm (nterm_to_pterm (term_to_nterm t)) \<down> cake_to_value ml_v"
by (rule cake_semantic_correctness) fact+
hence "fmap_of_list rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.as_vrules \<turnstile>\<^sub>v pterm_to_sterm (nterm_to_pterm (term_to_nterm t)) \<down> cake_to_value ml_v"
using assms unfolding compiled_def by simp
hence "rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.as_vrules, fmempty \<turnstile>\<^sub>v pterm_to_sterm (nterm_to_pterm (term_to_nterm t)) \<down> cake_to_value ml_v"
proof (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.veval'_correct'')
show "\<not> rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.shadows_consts (pterm_to_sterm (nterm_to_pterm (term_to_nterm t)))"
using \<open>\<not> shadows_consts (_::sterm)\<close> \<open>?heads = heads_of rs\<close> by auto
next
show "consts (pterm_to_sterm (nterm_to_pterm (term_to_nterm t))) |\<subseteq>| rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.all_consts"
using \<open>welldefined (pterm_to_sterm _)\<close> \<open>?heads = _\<close> by auto
qed fact+
hence "Rewriting_Sterm.compile (Rewriting_Pterm.compile (transform_irule_set_iter (Rewriting_Pterm_Elim.compile (consts_of compile)))), fmempty \<turnstile>\<^sub>s pterm_to_sterm (nterm_to_pterm (term_to_nterm t)) \<down> value_to_sterm (cake_to_value ml_v)"
by (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.veval_correct) fact+
hence "Rewriting_Sterm.compile (Rewriting_Pterm.compile (transform_irule_set_iter (Rewriting_Pterm_Elim.compile (consts_of compile)))) \<turnstile>\<^sub>s pterm_to_sterm (nterm_to_pterm (term_to_nterm t)) \<longrightarrow>* value_to_sterm (cake_to_value ml_v)"
by (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.prules_as_srules.seval_correct) fact
hence "Rewriting_Pterm.compile (transform_irule_set_iter (Rewriting_Pterm_Elim.compile (consts_of compile))) \<turnstile>\<^sub>p sterm_to_pterm (pterm_to_sterm (nterm_to_pterm (term_to_nterm t))) \<longrightarrow>* sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
by (rule rules_as_nrules.crules_as_irules'.irules'_as_prules.compile_correct_rt)
hence "Rewriting_Pterm.compile (transform_irule_set_iter (Rewriting_Pterm_Elim.compile (consts_of compile))) \<turnstile>\<^sub>p nterm_to_pterm (term_to_nterm t) \<longrightarrow>* sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
by (subst (asm) pterm_to_sterm_sterm_to_pterm) fact
hence "transform_irule_set_iter (Rewriting_Pterm_Elim.compile (consts_of compile)) \<turnstile>\<^sub>i nterm_to_pterm (term_to_nterm t) \<longrightarrow>* sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
by (rule rules_as_nrules.crules_as_irules'.irules'_as_irules.compile_correct_rt)
(rule rules_as_nrules.crules_as_irules.transform_finished)
have "Rewriting_Pterm_Elim.compile (consts_of compile) \<turnstile>\<^sub>i nterm_to_pterm (term_to_nterm t) \<longrightarrow>* sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
apply (rule rules_as_nrules.crules_as_irules.transform_correct_rt_n_no_abs)
using \<open>transform_irule_set_iter _ \<turnstile>\<^sub>i _ \<longrightarrow>* _\<close> unfolding transform_irule_set_iter_def
apply simp
apply fact+
done
then obtain t' where "compile \<turnstile>\<^sub>n term_to_nterm t \<longrightarrow>* t'" "t' \<approx>\<^sub>i sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
using \<open>closed (term_to_nterm t)\<close>
by (metis rules_as_nrules.compile_correct_rt)
hence "no_abs t'"
using \<open>no_abs (sterm_to_pterm _)\<close>
by (metis irelated_no_abs)
have "rs \<turnstile> nterm_to_term' (term_to_nterm t) \<longrightarrow>* nterm_to_term' t'"
by (rule compile_correct_rt) fact+
hence "rs \<turnstile> t \<longrightarrow>* nterm_to_term' t'"
apply (subst (asm) fresh_frun_def)
apply (subst (asm) term_to_nterm_nterm_to_term[where S = "fempty" and t = t, simplified])
apply (fold wellformed_term_def)
apply fact
using assms unfolding closed_except_def by auto
have "nterm_to_pterm t' = sterm_to_pterm (value_to_sterm (cake_to_value ml_v))"
using \<open>t' \<approx>\<^sub>i _\<close>
by auto
hence "(convert_term t' :: pterm) = convert_term (value_to_sterm (cake_to_value ml_v))"
apply (subst (asm) nterm_to_pterm)
apply fact
apply (subst (asm) sterm_to_pterm)
apply fact
apply assumption
done
hence "nterm_to_term' t' = convert_term (value_to_sterm (cake_to_value ml_v))"
apply (subst nterm_to_term')
apply (rule \<open>no_abs t'\<close>)
apply (rule convert_term_inj)
subgoal premises
apply (rule convert_term_no_abs)
apply fact
done
subgoal premises
apply (rule convert_term_no_abs)
apply fact
done
apply (subst convert_term_idem)
apply (rule \<open>no_abs t'\<close>)
apply (subst convert_term_idem)
apply (rule \<open>no_abs (value_to_sterm (cake_to_value ml_v))\<close>)
apply assumption
done
thus ?thesis
using \<open>rs \<turnstile> t \<longrightarrow>* nterm_to_term' t'\<close> by simp
qed
end
end
\ No newline at end of file
diff --git a/thys/Circus/Circus_Actions.thy b/thys/Circus/Circus_Actions.thy
--- a/thys/Circus/Circus_Actions.thy
+++ b/thys/Circus/Circus_Actions.thy
@@ -1,511 +1,512 @@
section \<open>Circus actions\<close>
theory Circus_Actions
imports HOLCF CSP_Processes
begin
text \<open>In this section, we introduce definitions for Circus actions with
some useful theorems and lemmas.\<close>
default_sort type
subsection \<open>Definitions\<close>
text \<open>The Circus actions type is defined as the set of all the CSP healthy reactive processes.\<close>
typedef ('\<theta>::"ev_eq",'\<sigma>) "action" = "{p::('\<theta>,'\<sigma>) relation_rp. is_CSP_process p}"
morphisms relation_of action_of
proof -
have "R (false \<turnstile> true) \<in> {p :: ('\<theta>,'\<sigma>) relation_rp. is_CSP_process p}"
by (auto intro: rd_is_CSP)
thus ?thesis by auto
qed
print_theorems
text \<open>The type-definition introduces a new type by stating a set. In our case,
it is the set of reactive processes that satisfy the healthiness-conditions
for CSP-processes, isomorphic to the new type.
Technically, this construct introduces two constants (morphisms) definitions $relation\_of$
and $action\_of$ as well as the usual axioms expressing the bijection @{thm relation_of_inverse}
and @{thm action_of_inverse}.\<close>
lemma relation_of_CSP: "is_CSP_process (relation_of x)"
proof -
have "(relation_of x) :{p. is_CSP_process p}" by (rule relation_of)
then show "is_CSP_process (relation_of x)" ..
qed
lemma relation_of_CSP1: "(relation_of x) is CSP1 healthy"
by (rule CSP_is_CSP1[OF relation_of_CSP])
lemma relation_of_CSP2: "(relation_of x) is CSP2 healthy"
by (rule CSP_is_CSP2[OF relation_of_CSP])
lemma relation_of_R: "(relation_of x) is R healthy"
by (rule CSP_is_R[OF relation_of_CSP])
subsection \<open>Proofs\<close>
text \<open>In the following, Circus actions are proved to be an instance of the $Complete\_Lattice$ class.\<close>
lemma relation_of_spec_f_f:
"\<forall>a b. (relation_of y \<longrightarrow> relation_of x) (a, b) \<Longrightarrow>
(relation_of y)\<^sup>f\<^sub>f (a\<lparr>tr := []\<rparr>, b) \<Longrightarrow>
(relation_of x)\<^sup>f\<^sub>f (a\<lparr>tr := []\<rparr>, b)"
by (auto simp: spec_def)
lemma relation_of_spec_t_f:
"\<forall>a b. (relation_of y \<longrightarrow> relation_of x) (a, b) \<Longrightarrow>
(relation_of y)\<^sup>t\<^sub>f (a\<lparr>tr := []\<rparr>, b) \<Longrightarrow>
(relation_of x)\<^sup>t\<^sub>f (a\<lparr>tr := []\<rparr>, b)"
by (auto simp: spec_def)
instantiation "action"::(ev_eq, type) below
begin
definition ref_def : "P \<sqsubseteq> Q \<equiv> [(relation_of Q) \<longrightarrow> (relation_of P)]"
instance ..
end
instance "action" :: (ev_eq, type) po
proof
fix x y z::"('a, 'b) action"
{
show "x \<sqsubseteq> x" by (simp add: ref_def utp_defs)
next
assume "x \<sqsubseteq> y" and "y \<sqsubseteq> z" then show " x \<sqsubseteq> z"
by (simp add: ref_def utp_defs)
next
assume A:"x \<sqsubseteq> y" and B:"y \<sqsubseteq> x" then show " x = y"
by (auto simp add: ref_def relation_of_inject[symmetric] fun_eq_iff)
}
qed
instantiation "action" :: (ev_eq, type) lattice
begin
definition inf_action : "(inf P Q \<equiv> action_of ((relation_of P) \<sqinter> (relation_of Q)))"
definition sup_action : "(sup P Q \<equiv> action_of ((relation_of P) \<squnion> (relation_of Q)))"
definition less_eq_action : "(less_eq (P::('a, 'b) action) Q \<equiv> P \<sqsubseteq> Q)"
definition less_action : "(less (P::('a, 'b) action) Q \<equiv> P \<sqsubseteq> Q \<and> \<not> Q \<sqsubseteq> P)"
instance
proof
fix x y z::"('a, 'b) action"
{
show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
by (simp add: less_action less_eq_action)
next
show "(x \<le> x)" by (simp add: less_eq_action)
next
assume "x \<le> y" and "y \<le> z"
then show " x \<le> z"
by (simp add: less_eq_action ref_def utp_defs)
next
assume "x \<le> y" and "y \<le> x"
then show " x = y"
by (auto simp add: less_eq_action ref_def relation_of_inject[symmetric] utp_defs)
next
show "inf x y \<le> x"
apply (auto simp add: less_eq_action inf_action ref_def
csp_defs design_defs rp_defs)
apply (subst action_of_inverse, simp add: Healthy_def)
apply (insert relation_of_CSP[where x="x"])
apply (insert relation_of_CSP[where x="y"])
apply (simp_all add: CSP_join)
apply (simp add: utp_defs)
done
next
show "inf x y \<le> y"
apply (auto simp add: less_eq_action inf_action ref_def csp_defs)
apply (subst action_of_inverse, simp add: Healthy_def)
apply (insert relation_of_CSP[where x="x"])
apply (insert relation_of_CSP[where x="y"])
apply (simp_all add: CSP_join)
apply (simp add: utp_defs)
done
next
assume "x \<le> y" and "x \<le> z"
then show "x \<le> inf y z"
apply (auto simp add: less_eq_action inf_action ref_def impl_def csp_defs)
apply (erule_tac x="a" in allE, erule_tac x="a" in allE)
apply (erule_tac x="b" in allE)+
apply (subst (asm) action_of_inverse)
apply (simp add: Healthy_def)
apply (insert relation_of_CSP[where x="z"])
apply (insert relation_of_CSP[where x="y"])
apply (auto simp add: CSP_join)
done
next
show "x \<le> sup x y"
apply (auto simp add: less_eq_action sup_action ref_def
impl_def csp_defs)
apply (subst (asm) action_of_inverse)
apply (simp add: Healthy_def)
apply (insert relation_of_CSP[where x="x"])
apply (insert relation_of_CSP[where x="y"])
apply (auto simp add: CSP_meet)
done
next
show "y \<le> sup x y"
apply (auto simp add: less_eq_action sup_action ref_def
impl_def csp_defs)
apply (subst (asm) action_of_inverse)
apply (simp add: Healthy_def)
apply (insert relation_of_CSP[where x="x"])
apply (insert relation_of_CSP[where x="y"])
apply (auto simp add: CSP_meet)
done
next
assume "y \<le> x" and "z \<le> x"
then show "sup y z \<le> x"
apply (auto simp add: less_eq_action sup_action ref_def impl_def csp_defs)
apply (erule_tac x="a" in allE)
apply (erule_tac x="a" in allE)
apply (erule_tac x="b" in allE)+
apply (subst action_of_inverse)
apply (simp add: Healthy_def)
apply (insert relation_of_CSP[where x="z"])
apply (insert relation_of_CSP[where x="y"])
apply (auto simp add: CSP_meet)
done
}
qed
end
lemma bot_is_action: "R (false \<turnstile> true) \<in> {p. is_CSP_process p}"
by (auto intro: rd_is_CSP)
lemma bot_eq_true: "R (false \<turnstile> true) = R true"
by (auto simp: fun_eq_iff design_defs rp_defs split: cond_splits)
instantiation "action" :: (ev_eq, type) bounded_lattice
begin
definition bot_action : "(bot::('a, 'b) action) \<equiv> action_of (R(false \<turnstile> true))"
definition top_action : "(top::('a, 'b) action) \<equiv> action_of (R(true \<turnstile> false))"
instance
proof
fix x::"('a, 'b) action"
{
show "bot \<le> x"
unfolding bot_action
apply (auto simp add: less_action less_eq_action ref_def bot_action)
apply (subst action_of_inverse) apply (rule bot_is_action)
apply (subst bot_eq_true)
apply (subst (asm) CSP_is_rd)
apply (rule relation_of_CSP)
apply (auto simp add: csp_defs rp_defs fun_eq_iff split: cond_splits)
done
next
show "x \<le> top"
apply (auto simp add: less_action less_eq_action ref_def top_action)
apply (subst (asm) action_of_inverse)
apply (simp)
apply (rule rd_is_CSP)
apply auto
apply (subst action_of_cases[where x=x], simp_all)
apply (subst action_of_inverse, simp_all)
apply (subst CSP_is_rd[where P=y], simp_all)
apply (auto simp: rp_defs design_defs fun_eq_iff split: cond_splits)
done
}
qed
end
lemma relation_of_top: "relation_of top = R(true \<turnstile> false)"
apply (simp add: top_action)
apply (subst action_of_inverse)
apply (simp)
apply (rule rd_is_CSP)
apply (auto simp add: utp_defs design_defs rp_defs)
done
lemma relation_of_bot: "relation_of bot = R true"
apply (simp add: bot_action)
apply (subst action_of_inverse)
apply (simp add: bot_is_action[simplified], rule bot_eq_true)
done
lemma non_emptyE: assumes "A \<noteq> {}" obtains x where "x : A"
using assms by (auto simp add: ex_in_conv [symmetric])
lemma CSP1_Inf:
assumes *:"A \<noteq> {}"
shows "(\<Sqinter> relation_of ` A) is CSP1 healthy"
proof -
have "(\<Sqinter> relation_of ` A) = CSP1 (\<Sqinter> relation_of ` A)"
proof
fix P
note * then
show "(P \<in> \<Union>{{p. P p} |P. P \<in> relation_of ` A}) = CSP1 (\<lambda>Aa. Aa \<in> \<Union>{{p. P p} |P. P \<in> relation_of ` A}) P"
apply (intro iffI)
apply (simp_all add: csp_defs)
apply (rule disj_introC, simp)
apply (erule disj_elim, simp_all)
apply (cases P, simp_all)
apply (erule non_emptyE)
apply (rule_tac x="Collect (relation_of x)" in exI, simp)
apply (rule conjI)
apply (rule_tac x="(relation_of x)" in exI, simp)
apply (subst CSP_is_rd, simp add: relation_of_CSP)
apply (auto simp add: csp_defs design_defs rp_defs fun_eq_iff split: cond_splits)
done
qed
then show "(\<Sqinter> relation_of ` A) is CSP1 healthy" by (simp add: design_defs)
qed
lemma CSP2_Inf:
assumes *:"A \<noteq> {}"
shows "(\<Sqinter> relation_of ` A) is CSP2 healthy"
proof -
have "(\<Sqinter> relation_of ` A) = CSP2 (\<Sqinter> relation_of ` A)"
proof
fix P
note * then
show "(P \<in> \<Union>{{p. P p} |P. P \<in> relation_of ` A}) = CSP2 (\<lambda>Aa. Aa \<in> \<Union>{{p. P p} |P. P \<in> relation_of ` A}) P"
apply (intro iffI)
apply (simp_all add: csp_defs)
apply (cases P, simp_all)
apply (erule exE)
apply (rule_tac b=b in comp_intro, simp_all)
apply (rule_tac x=x in exI, simp)
apply (erule comp_elim, simp_all)
apply (erule exE | erule conjE)+
apply (simp_all)
apply (rule_tac x="Collect Pa" in exI, simp)
apply (rule conjI)
apply (rule_tac x="Pa" in exI, simp)
apply (erule Set.imageE, simp add: relation_of)
apply (subst CSP_is_rd, simp add: relation_of_CSP)
apply (subst (asm) CSP_is_rd, simp add: relation_of_CSP)
apply (auto simp add: csp_defs rp_defs prefix_def design_defs fun_eq_iff split: cond_splits)
apply (subgoal_tac "b\<lparr>tr := zs, ok := False\<rparr> = c\<lparr>tr := zs, ok := False\<rparr>", auto)
apply (subgoal_tac "b\<lparr>tr := zs, ok := False\<rparr> = c\<lparr>tr := zs, ok := False\<rparr>", auto)
apply (subgoal_tac "b\<lparr>tr := zs, ok := False\<rparr> = c\<lparr>tr := zs, ok := False\<rparr>", auto)
apply (subgoal_tac "b\<lparr>tr := zs, ok := False\<rparr> = c\<lparr>tr := zs, ok := False\<rparr>", auto)
apply (subgoal_tac "b\<lparr>tr := zs, ok := False\<rparr> = c\<lparr>tr := zs, ok := False\<rparr>", auto)
apply (subgoal_tac "b\<lparr>tr := zs, ok := False\<rparr> = c\<lparr>tr := zs, ok := False\<rparr>", auto)
apply (subgoal_tac "b\<lparr>tr := zs, ok := True\<rparr> = c\<lparr>tr := zs, ok := True\<rparr>", auto)
apply (subgoal_tac "b\<lparr>tr := zs, ok := True\<rparr> = c\<lparr>tr := zs, ok := True\<rparr>", auto)
done
qed
then show "(\<Sqinter> relation_of ` A) is CSP2 healthy" by (simp add: design_defs)
qed
lemma R_Inf:
assumes *:"A \<noteq> {}"
shows "(\<Sqinter> relation_of ` A) is R healthy"
proof -
have "(\<Sqinter> relation_of ` A) = R (\<Sqinter> relation_of ` A)"
proof
fix P
show "(P \<in> \<Union>{{p. P p} |P. P \<in> relation_of ` A}) = R (\<lambda>Aa. Aa \<in> \<Union>{{p. P p} |P. P \<in> relation_of ` A}) P"
apply (cases P, simp_all)
apply (rule)
apply (simp_all add: csp_defs rp_defs split: cond_splits)
apply (erule exE)
apply (erule exE | erule conjE)+
apply (simp_all)
apply (erule Set.imageE, simp add: relation_of)
apply (subst (asm) CSP_is_rd, simp add: relation_of_CSP)
apply (simp add: csp_defs prefix_def design_defs rp_defs fun_eq_iff split: cond_splits)
apply (rule_tac x="x" in exI, simp)
apply (rule conjI)
apply (rule_tac x="relation_of xa" in exI, simp)
apply (subst CSP_is_rd, simp add: relation_of_CSP)
apply (simp add: csp_defs prefix_def design_defs rp_defs fun_eq_iff split: cond_splits)
apply (insert *)
apply (erule non_emptyE)
apply (rule_tac x="Collect (relation_of x)" in exI, simp)
apply (rule conjI)
apply (rule_tac x="(relation_of x)" in exI, simp)
apply (subst CSP_is_rd, simp add: relation_of_CSP)
apply (simp add: csp_defs prefix_def design_defs rp_defs fun_eq_iff split: cond_splits)
apply (erule exE | erule conjE)+
apply (simp_all)
apply (erule Set.imageE, simp add: relation_of)
apply (rule_tac x="x" in exI, simp)
apply (rule conjI)
apply (rule_tac x="(relation_of xa)" in exI, simp)
apply (subst CSP_is_rd, simp add: relation_of_CSP)
apply (simp add: csp_defs prefix_def design_defs rp_defs fun_eq_iff split: cond_splits)
apply (subst (asm) CSP_is_rd, simp add: relation_of_CSP)
apply (simp add: csp_defs prefix_def design_defs rp_defs fun_eq_iff split: cond_splits)
done
qed
then show "(\<Sqinter> relation_of ` A) is R healthy" by (simp add: design_defs)
qed
lemma CSP_Inf:
assumes "A \<noteq> {}"
shows "is_CSP_process (\<Sqinter> relation_of ` A)"
unfolding is_CSP_process_def
using assms CSP1_Inf CSP2_Inf R_Inf
by auto
lemma Inf_is_action: "A \<noteq> {} \<Longrightarrow> \<Sqinter> relation_of ` A \<in> {p. is_CSP_process p}"
by (auto dest!: CSP_Inf)
lemma CSP1_Sup: "A \<noteq> {} \<Longrightarrow> (\<Squnion> relation_of ` A) is CSP1 healthy"
apply (auto simp add: design_defs csp_defs fun_eq_iff)
apply (subst CSP_is_rd, simp add: relation_of_CSP)
apply (simp add: csp_defs prefix_def design_defs rp_defs split: cond_splits)
done
lemma CSP2_Sup: "A \<noteq> {} \<Longrightarrow> (\<Squnion> relation_of ` A) is CSP2 healthy"
+ supply [[simproc del: defined_all]]
apply (simp add: design_defs csp_defs fun_eq_iff)
apply (rule allI)+
apply (rule)
apply (rule_tac b=b in comp_intro, simp_all)
apply (erule comp_elim, simp_all)
apply (rule allI)
apply (erule_tac x=x in allE)
apply (rule impI)
apply (case_tac "(\<exists>P. x = Collect P & P \<in> relation_of ` A)", simp_all)
apply (erule exE | erule conjE)+
apply (simp_all)
apply (erule Set.imageE, simp add: relation_of)
apply (subst (asm) CSP_is_rd, simp add: relation_of_CSP, subst CSP_is_rd, simp add: relation_of_CSP)
apply (auto simp add: csp_defs design_defs rp_defs split: cond_splits)
apply (subgoal_tac "ba\<lparr>tr := tr c - tr aa, ok := False\<rparr> = c\<lparr>tr := tr c - tr aa, ok := False\<rparr>", simp_all)
apply (subgoal_tac "ba\<lparr>tr := tr c - tr aa, ok := False\<rparr> = c\<lparr>tr := tr c - tr aa, ok := False\<rparr>", simp_all)
apply (subgoal_tac "ba\<lparr>tr := tr c - tr aa, ok := False\<rparr> = c\<lparr>tr := tr c - tr aa, ok := False\<rparr>", simp_all)
apply (subgoal_tac "ba\<lparr>tr := tr c - tr aa, ok := False\<rparr> = c\<lparr>tr := tr c - tr aa, ok := False\<rparr>", simp_all)
apply (subgoal_tac "ba\<lparr>tr := tr c - tr aa, ok := False\<rparr> = c\<lparr>tr := tr c - tr aa, ok := False\<rparr>", simp_all)
apply (subgoal_tac "ba\<lparr>tr := tr c - tr aa, ok := False\<rparr> = c\<lparr>tr := tr c - tr aa, ok := False\<rparr>", simp_all)
apply (subgoal_tac "ba\<lparr>tr := tr c - tr aa, ok := True\<rparr> = c\<lparr>tr := tr c - tr aa, ok := True\<rparr>", simp_all)
apply (subgoal_tac "ba\<lparr>tr := tr c - tr aa, ok := True\<rparr> = c\<lparr>tr := tr c - tr aa, ok := True\<rparr>", simp_all)
done
lemma R_Sup: "A \<noteq> {} \<Longrightarrow> (\<Squnion> relation_of ` A) is R healthy"
apply (simp add: rp_defs design_defs csp_defs fun_eq_iff)
apply (rule allI)+
apply (rule)
apply (simp split: cond_splits)
apply (case_tac "wait a", simp_all)
apply (erule non_emptyE)
apply (erule_tac x="Collect (relation_of x)" in allE, simp_all)
apply (case_tac "relation_of x (a, b)", simp_all)
apply (subst (asm) CSP_is_rd, simp add: relation_of_CSP)
apply (simp add: csp_defs design_defs rp_defs split: cond_splits)
apply (erule_tac x="(relation_of x)" in allE, simp_all)
apply (rule conjI)
apply (rule allI)
apply (erule_tac x=x in allE)
apply (rule impI)
apply (case_tac "(\<exists>P. x = Collect P & P \<in> relation_of ` A)", simp_all)
apply (erule exE | erule conjE)+
apply (simp_all)
apply (erule Set.imageE, simp add: relation_of)
apply (subst (asm) CSP_is_rd, simp add: relation_of_CSP, subst CSP_is_rd, simp add: relation_of_CSP)
apply (simp add: csp_defs design_defs rp_defs split: cond_splits)
apply (erule non_emptyE)
apply (erule_tac x="Collect (relation_of x)" in allE, simp_all)
apply (case_tac "relation_of x (a, b)", simp_all)
apply (subst (asm) CSP_is_rd, simp add: relation_of_CSP)
apply (simp add: csp_defs design_defs rp_defs split: cond_splits)
apply (erule_tac x="(relation_of x)" in allE, simp_all)
apply (simp split: cond_splits)
apply (rule allI)
apply (rule impI)
apply (erule exE | erule conjE)+
apply (simp_all)
apply (erule Set.imageE, simp add: relation_of)
apply (subst (asm) CSP_is_rd, simp add: relation_of_CSP, subst CSP_is_rd, simp add: relation_of_CSP)
apply (simp add: csp_defs design_defs rp_defs split: cond_splits)
apply (rule allI)
apply (rule impI)
apply (erule exE | erule conjE)+
apply (simp_all)
apply (erule Set.imageE, simp add: relation_of)
apply (erule_tac x="x" in allE, simp_all)
apply (case_tac "relation_of xa (a\<lparr>tr := []\<rparr>, b\<lparr>tr := tr b - tr a\<rparr>)", simp_all)
apply (subst (asm) CSP_is_rd) back back back
apply (simp add: relation_of_CSP, subst CSP_is_rd, simp add: relation_of_CSP)
apply (simp add: csp_defs design_defs rp_defs split: cond_splits)
apply (erule_tac x="P" in allE, simp_all)
done
lemma CSP_Sup: "A \<noteq> {} \<Longrightarrow> is_CSP_process (\<Squnion> relation_of ` A)"
unfolding is_CSP_process_def using CSP1_Sup CSP2_Sup R_Sup by auto
lemma Sup_is_action: "A \<noteq> {} \<Longrightarrow> \<Squnion> relation_of ` A \<in> {p. is_CSP_process p}"
by (auto dest!: CSP_Sup)
lemma relation_of_Sup:
"A \<noteq> {} \<Longrightarrow> relation_of (action_of \<Squnion> relation_of ` A) = \<Squnion> relation_of ` A"
by (auto simp: action_of_inverse dest!: Sup_is_action)
instantiation "action" :: (ev_eq, type) complete_lattice
begin
definition Sup_action :
"(Sup (S:: ('a, 'b) action set) \<equiv> if S={} then bot else action_of \<Squnion> (relation_of ` S))"
definition Inf_action :
"(Inf (S:: ('a, 'b) action set) \<equiv> if S={} then top else action_of \<Sqinter> (relation_of ` S))"
instance
proof
fix A::"('a, 'b) action set" and z::"('a, 'b) action"
{
fix x::"('a, 'b) action"
assume "x \<in> A"
then show "Inf A \<le> x"
apply (auto simp add: less_action less_eq_action Inf_action ref_def)
apply (subst (asm) action_of_inverse)
apply (auto intro: Inf_is_action[simplified])
done
} note rule1 = this
{
assume *: "\<And>x. x \<in> A \<Longrightarrow> z \<le> x"
show "z \<le> Inf A"
proof (cases "A = {}")
case True
then show ?thesis by (simp add: Inf_action)
next
case False
show ?thesis
using *
apply (auto simp add: Inf_action)
using \<open>A \<noteq> {}\<close>
apply (simp add: less_eq_action Inf_action ref_def)
apply (subst (asm) action_of_inverse)
apply (subst (asm) ex_in_conv[symmetric])
apply (erule exE)
apply (auto intro: Inf_is_action[simplified])
done
qed
}{
fix x::"('a, 'b) action"
assume "x \<in> A"
then show "x \<le> (Sup A)"
apply (auto simp add: less_action less_eq_action Sup_action ref_def)
apply (subst (asm) action_of_inverse)
apply (auto intro: Sup_is_action[simplified])
done
} note rule2 = this
{
assume "\<And>x. x \<in> A \<Longrightarrow> x \<le> z"
then show "Sup A \<le> z"
apply (auto simp add: Sup_action)
apply atomize
apply (case_tac "A = {}", simp_all)
apply (insert rule2)
apply (auto simp add: less_action less_eq_action Sup_action ref_def)
apply (subst (asm) action_of_inverse)
apply (auto intro: Sup_is_action[simplified])
apply (subst (asm) action_of_inverse)
apply (auto intro: Sup_is_action[simplified])
done
}
{ show "Inf ({}::('a, 'b) action set) = top" by(simp add:Inf_action) }
{ show "Sup ({}::('a, 'b) action set) = bot" by(simp add:Sup_action) }
qed
end
end
diff --git a/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy b/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy
--- a/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy
+++ b/thys/Complex_Geometry/Unit_Circle_Preserving_Moebius.thy
@@ -1,1202 +1,1202 @@
(* ---------------------------------------------------------------------------- *)
section \<open>Unit circle preserving Möbius transformations\<close>
(* ---------------------------------------------------------------------------- *)
text \<open>In this section we shall examine Möbius transformations that map the unit circle onto itself.
We shall say that they fix or preserve the unit circle (although, they do not need to fix each of
its points).\<close>
theory Unit_Circle_Preserving_Moebius
imports Unitary11_Matrices Moebius Oriented_Circlines
begin
(* ---------------------------------------------------------------------------- *)
subsection \<open>Möbius transformations that fix the unit circle\<close>
(* ---------------------------------------------------------------------------- *)
text \<open>We define Möbius transformations that preserve unit circle as transformations represented by
generalized unitary matrices with the $1-1$ signature (elements of the gruop $GU_{1,1}(2,
\mathbb{C})$, defined earlier in the theory Unitary11Matrices).\<close>
lift_definition unit_circle_fix_mmat :: "moebius_mat \<Rightarrow> bool" is unitary11_gen
done
lift_definition unit_circle_fix :: "moebius \<Rightarrow> bool" is unit_circle_fix_mmat
apply transfer
apply (auto simp del: mult_sm.simps)
apply (simp del: mult_sm.simps add: unitary11_gen_mult_sm)
apply (simp del: mult_sm.simps add: unitary11_gen_div_sm)
done
text \<open>Our algebraic characterisation (by matrices) is geometrically correct.\<close>
lemma unit_circle_fix_iff:
shows "unit_circle_fix M \<longleftrightarrow>
moebius_circline M unit_circle = unit_circle" (is "?rhs = ?lhs")
proof
assume ?lhs
thus ?rhs
proof (transfer, transfer)
fix M :: complex_mat
assume "mat_det M \<noteq> 0"
assume "circline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat"
then obtain k where "k \<noteq> 0" "(1, 0, 0, -1) = cor k *\<^sub>s\<^sub>m congruence (mat_inv M) (1, 0, 0, -1)"
by auto
hence "(1/cor k, 0, 0, -1/cor k) = congruence (mat_inv M) (1, 0, 0, -1)"
using mult_sm_inv_l[of "cor k" "congruence (mat_inv M) (1, 0, 0, -1)" ]
by simp
hence "congruence M (1/cor k, 0, 0, -1/cor k) = (1, 0, 0, -1)"
using \<open>mat_det M \<noteq> 0\<close> mat_det_inv[of M]
using congruence_inv[of "mat_inv M" "(1, 0, 0, -1)" "(1/cor k, 0, 0, -1/cor k)"]
by simp
hence "congruence M (1, 0, 0, -1) = cor k *\<^sub>s\<^sub>m (1, 0, 0, -1)"
using congruence_scale_m[of "M" "1/cor k" "(1, 0, 0, -1)"]
using mult_sm_inv_l[of "1/ cor k" "congruence M (1, 0, 0, -1)" "(1, 0, 0, -1)"] \<open>k \<noteq> 0\<close>
by simp
thus "unitary11_gen M"
using \<open>k \<noteq> 0\<close>
unfolding unitary11_gen_def
by simp
qed
next
assume ?rhs
thus ?lhs
proof (transfer, transfer)
fix M :: complex_mat
assume "mat_det M \<noteq> 0"
assume "unitary11_gen M"
hence "unitary11_gen (mat_inv M)"
using \<open>mat_det M \<noteq> 0\<close>
using unitary11_gen_mat_inv
by simp
thus " circline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat"
unfolding unitary11_gen_real
by auto (rule_tac x="1/k" in exI, simp)
qed
qed
lemma circline_set_fix_iff_circline_fix:
assumes "circline_set H' \<noteq> {}"
shows "circline_set (moebius_circline M H) = circline_set H' \<longleftrightarrow>
moebius_circline M H = H'"
using assms
by auto (rule inj_circline_set, auto)
lemma unit_circle_fix_iff_unit_circle_set:
shows "unit_circle_fix M \<longleftrightarrow> moebius_pt M ` unit_circle_set = unit_circle_set"
proof-
have "circline_set unit_circle \<noteq> {}"
using one_in_unit_circle_set
by auto
thus ?thesis
using unit_circle_fix_iff[of M] circline_set_fix_iff_circline_fix[of unit_circle M unit_circle]
by (simp add: unit_circle_set_def)
qed
text \<open>Unit circle preserving Möbius transformations form a group. \<close>
lemma unit_circle_fix_id_moebius [simp]:
shows "unit_circle_fix id_moebius"
by (transfer, transfer, simp add: unitary11_gen_def mat_adj_def mat_cnj_def)
lemma unit_circle_fix_moebius_add [simp]:
assumes "unit_circle_fix M1" and "unit_circle_fix M2"
shows "unit_circle_fix (M1 + M2)"
using assms
unfolding unit_circle_fix_iff
by auto
lemma unit_circle_fix_moebius_comp [simp]:
assumes "unit_circle_fix M1" and "unit_circle_fix M2"
shows "unit_circle_fix (moebius_comp M1 M2)"
using unit_circle_fix_moebius_add[OF assms]
by simp
lemma unit_circle_fix_moebius_uminus [simp]:
assumes "unit_circle_fix M"
shows "unit_circle_fix (-M)"
using assms
unfolding unit_circle_fix_iff
by (metis moebius_circline_comp_inv_left uminus_moebius_def)
lemma unit_circle_fix_moebius_inv [simp]:
assumes "unit_circle_fix M"
shows "unit_circle_fix (moebius_inv M)"
using unit_circle_fix_moebius_uminus[OF assms]
by simp
text \<open>Unit circle fixing transforms preserve inverse points.\<close>
lemma unit_circle_fix_moebius_pt_inversion [simp]:
assumes "unit_circle_fix M"
shows "moebius_pt M (inversion z) = inversion (moebius_pt M z)"
using assms
using symmetry_principle[of z "inversion z" unit_circle M]
using unit_circle_fix_iff[of M, symmetric]
using circline_symmetric_inv_homo_disc[of z]
using circline_symmetric_inv_homo_disc'[of "moebius_pt M z" "moebius_pt M (inversion z)"]
by metis
(* -------------------------------------------------------------------------- *)
subsection \<open>Möbius transformations that fix the imaginary unit circle\<close>
(* -------------------------------------------------------------------------- *)
text \<open>Only for completeness we show that Möbius transformations that preserve the imaginary unit
circle are exactly those characterised by generalized unitary matrices (with the (2, 0) signature).\<close>
lemma imag_unit_circle_fixed_iff_unitary_gen:
assumes "mat_det (A, B, C, D) \<noteq> 0"
shows "moebius_circline (mk_moebius A B C D) imag_unit_circle = imag_unit_circle \<longleftrightarrow>
unitary_gen (A, B, C, D)" (is "?lhs = ?rhs")
proof
assume ?lhs
thus ?rhs
using assms
proof (transfer, transfer)
fix A B C D :: complex
let ?M = "(A, B, C, D)" and ?E = "(1, 0, 0, 1)"
assume "circline_eq_cmat (moebius_circline_cmat_cmat (mk_moebius_cmat A B C D) imag_unit_circle_cmat) imag_unit_circle_cmat"
"mat_det ?M \<noteq> 0"
then obtain k where "k \<noteq> 0" "?E = cor k *\<^sub>s\<^sub>m congruence (mat_inv ?M) ?E"
by auto
hence "unitary_gen (mat_inv ?M)"
using mult_sm_inv_l[of "cor k" "congruence (mat_inv ?M) ?E" "?E"]
unfolding unitary_gen_def
by (metis congruence_def divide_eq_0_iff eye_def mat_eye_r of_real_eq_0_iff one_neq_zero)
thus "unitary_gen ?M"
using unitary_gen_inv[of "mat_inv ?M"] \<open>mat_det ?M \<noteq> 0\<close>
by (simp del: mat_inv.simps)
qed
next
assume ?rhs
thus ?lhs
using assms
proof (transfer, transfer)
fix A B C D :: complex
let ?M = "(A, B, C, D)" and ?E = "(1, 0, 0, 1)"
assume "unitary_gen ?M" "mat_det ?M \<noteq> 0"
hence "unitary_gen (mat_inv ?M)"
using unitary_gen_inv[of ?M]
by simp
then obtain k where "k \<noteq> 0" "mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m (mat_inv ?M) = cor k *\<^sub>s\<^sub>m eye"
using unitary_gen_real[of "mat_inv ?M"] mat_det_inv[of ?M]
by auto
hence *: "?E = (1 / cor k) *\<^sub>s\<^sub>m (mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m (mat_inv ?M))"
using mult_sm_inv_l[of "cor k" eye "mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m (mat_inv ?M)"]
by simp
have "\<exists>k. k \<noteq> 0 \<and>
(1, 0, 0, 1) = cor k *\<^sub>s\<^sub>m (mat_adj (mat_inv (A, B, C, D)) *\<^sub>m\<^sub>m (1, 0, 0, 1) *\<^sub>m\<^sub>m mat_inv (A, B, C, D))"
using \<open>mat_det ?M \<noteq> 0\<close> \<open>k \<noteq> 0\<close>
by (metis "*" Im_complex_of_real Re_complex_of_real \<open>mat_adj (mat_inv ?M) *\<^sub>m\<^sub>m mat_inv ?M = cor k *\<^sub>s\<^sub>m eye\<close> complex_of_real_Re eye_def mat_eye_l mult_mm_assoc mult_mm_sm mult_sm_eye_mm of_real_1 of_real_divide of_real_eq_1_iff zero_eq_1_divide_iff)
thus "circline_eq_cmat (moebius_circline_cmat_cmat (mk_moebius_cmat A B C D) imag_unit_circle_cmat) imag_unit_circle_cmat"
using \<open>mat_det ?M \<noteq> 0\<close> \<open>k \<noteq> 0\<close>
by (simp del: mat_inv.simps)
qed
qed
(* -------------------------------------------------------------------------- *)
subsection \<open>Möbius transformations that fix the oriented unit circle and the unit disc\<close>
(* -------------------------------------------------------------------------- *)
text \<open>Möbius transformations that fix the unit circle either map the unit disc onto itself or
exchange it with its exterior. The transformations that fix the unit disc can be recognized from
their matrices -- they have the form as before, but additionally it must hold that $|a|^2 > |b|^2$.\<close>
definition unit_disc_fix_cmat :: "complex_mat \<Rightarrow> bool" where
[simp]: "unit_disc_fix_cmat M \<longleftrightarrow>
(let (A, B, C, D) = M
in unitary11_gen (A, B, C, D) \<and> (B = 0 \<or> Re ((A*D)/(B*C)) > 1))"
lift_definition unit_disc_fix_mmat :: "moebius_mat \<Rightarrow> bool" is unit_disc_fix_cmat
done
lift_definition unit_disc_fix :: "moebius \<Rightarrow> bool" is unit_disc_fix_mmat
proof transfer
fix M M' :: complex_mat
assume det: "mat_det M \<noteq> 0" "mat_det M' \<noteq> 0"
assume "moebius_cmat_eq M M'"
then obtain k where *: "k \<noteq> 0" "M' = k *\<^sub>s\<^sub>m M"
by auto
hence **: "unitary11_gen M \<longleftrightarrow> unitary11_gen M'"
using unitary11_gen_mult_sm[of k M] unitary11_gen_div_sm[of k M]
by auto
obtain A B C D where MM: "(A, B, C, D) = M"
by (cases M) auto
obtain A' B' C' D' where MM': "(A', B', C', D') = M'"
by (cases M') auto
show "unit_disc_fix_cmat M = unit_disc_fix_cmat M'"
using * ** MM MM'
by auto
qed
text \<open>Transformations that fix the unit disc also fix the unit circle.\<close>
lemma unit_disc_fix_unit_circle_fix [simp]:
assumes "unit_disc_fix M"
shows "unit_circle_fix M"
using assms
by (transfer, transfer, auto)
text \<open>Transformations that preserve the unit disc preserve the orientation of the unit circle.\<close>
lemma unit_disc_fix_iff_ounit_circle:
shows "unit_disc_fix M \<longleftrightarrow>
moebius_ocircline M ounit_circle = ounit_circle" (is "?rhs \<longleftrightarrow> ?lhs")
proof
assume *: ?lhs
have "moebius_circline M unit_circle = unit_circle"
apply (subst moebius_circline_ocircline[of M unit_circle])
apply (subst of_circline_unit_circle)
apply (subst *)
by simp
hence "unit_circle_fix M"
by (simp add: unit_circle_fix_iff)
thus ?rhs
using *
proof (transfer, transfer)
fix M :: complex_mat
assume "mat_det M \<noteq> 0"
let ?H = "(1, 0, 0, -1)"
obtain A B C D where MM: "(A, B, C, D) = M"
by (cases M) auto
assume "unitary11_gen M" "ocircline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat"
then obtain k where "0 < k" "?H = cor k *\<^sub>s\<^sub>m congruence (mat_inv M) ?H"
by auto
hence "congruence M ?H = cor k *\<^sub>s\<^sub>m ?H"
using congruence_inv[of "mat_inv M" "?H" "(1/cor k) *\<^sub>s\<^sub>m ?H"] \<open>mat_det M \<noteq> 0\<close>
using mult_sm_inv_l[of "cor k" "congruence (mat_inv M) ?H" "?H"]
using mult_sm_inv_l[of "1/cor k" "congruence M ?H"]
using congruence_scale_m[of M "1/cor k" "?H"]
using \<open>\<And>B. \<lbrakk>1 / cor k \<noteq> 0; (1 / cor k) *\<^sub>s\<^sub>m congruence M (1, 0, 0, - 1) = B\<rbrakk> \<Longrightarrow> congruence M (1, 0, 0, - 1) = (1 / (1 / cor k)) *\<^sub>s\<^sub>m B\<close>
by (auto simp add: mat_det_inv)
then obtain a b k' where "k' \<noteq> 0" "M = k' *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)" "sgn (Re (mat_det (a, b, cnj b, cnj a))) = 1"
using unitary11_sgn_det_orientation'[of M k] \<open>k > 0\<close>
by auto
moreover
have "mat_det (a, b, cnj b, cnj a) \<noteq> 0"
using \<open>sgn (Re (mat_det (a, b, cnj b, cnj a))) = 1\<close>
by (smt sgn_0 zero_complex.simps(1))
ultimately
show "unit_disc_fix_cmat M"
using unitary11_sgn_det[of k' a b M A B C D]
using MM[symmetric] \<open>k > 0\<close> \<open>unitary11_gen M\<close>
by (simp add: sgn_1_pos split: if_split_asm)
qed
next
assume ?rhs
thus ?lhs
proof (transfer, transfer)
fix M :: complex_mat
assume "mat_det M \<noteq> 0"
obtain A B C D where MM: "(A, B, C, D) = M"
by (cases M) auto
assume "unit_disc_fix_cmat M"
hence "unitary11_gen M" "B = 0 \<or> 1 < Re (A * D / (B * C))"
using MM[symmetric]
by auto
have "sgn (if B = 0 then 1 else sgn (Re (A * D / (B * C)) - 1)) = 1"
using \<open>B = 0 \<or> 1 < Re (A * D / (B * C))\<close>
by auto
then obtain k' where "k' > 0" "congruence M (1, 0, 0, -1) = cor k' *\<^sub>s\<^sub>m (1, 0, 0, -1)"
using unitary11_orientation[OF \<open>unitary11_gen M\<close> MM[symmetric]]
by (auto simp add: sgn_1_pos)
thus "ocircline_eq_cmat (moebius_circline_cmat_cmat M unit_circle_cmat) unit_circle_cmat"
using congruence_inv[of M "(1, 0, 0, -1)" "cor k' *\<^sub>s\<^sub>m (1, 0, 0, -1)"] \<open>mat_det M \<noteq> 0\<close>
using congruence_scale_m[of "mat_inv M" "cor k'" "(1, 0, 0, -1)"]
by auto
qed
qed
text \<open>Our algebraic characterisation (by matrices) is geometrically correct.\<close>
lemma unit_disc_fix_iff [simp]:
assumes "unit_disc_fix M"
shows "moebius_pt M ` unit_disc = unit_disc"
using assms
using unit_disc_fix_iff_ounit_circle[of M]
unfolding unit_disc_def
by (subst disc_moebius_ocircline[symmetric], simp)
lemma unit_disc_fix_discI [simp]:
assumes "unit_disc_fix M" and "u \<in> unit_disc"
shows "moebius_pt M u \<in> unit_disc"
using unit_disc_fix_iff assms
by blast
text \<open>Unit disc preserving transformations form a group.\<close>
lemma unit_disc_fix_id_moebius [simp]:
shows "unit_disc_fix id_moebius"
by (transfer, transfer, simp add: unitary11_gen_def mat_adj_def mat_cnj_def)
lemma unit_disc_fix_moebius_add [simp]:
assumes "unit_disc_fix M1" and "unit_disc_fix M2"
shows "unit_disc_fix (M1 + M2)"
using assms
unfolding unit_disc_fix_iff_ounit_circle
by auto
lemma unit_disc_fix_moebius_comp [simp]:
assumes "unit_disc_fix M1" and "unit_disc_fix M2"
shows "unit_disc_fix (moebius_comp M1 M2)"
using unit_disc_fix_moebius_add[OF assms]
by simp
lemma unit_disc_fix_moebius_uminus [simp]:
assumes "unit_disc_fix M"
shows "unit_disc_fix (-M)"
using assms
unfolding unit_disc_fix_iff_ounit_circle
by (metis moebius_ocircline_comp_inv_left uminus_moebius_def)
lemma unit_disc_fix_moebius_inv [simp]:
assumes "unit_disc_fix M"
shows "unit_disc_fix (moebius_inv M)"
using unit_disc_fix_moebius_uminus[OF assms]
by simp
(* -------------------------------------------------------------------------- *)
subsection \<open>Rotations are unit disc preserving transformations\<close>
(* -------------------------------------------------------------------------- *)
lemma unit_disc_fix_rotation [simp]:
shows "unit_disc_fix (moebius_rotation \<phi>)"
unfolding moebius_rotation_def moebius_similarity_def
by (transfer, transfer, simp add: unitary11_gen_def mat_adj_def mat_cnj_def cis_mult)
lemma moebius_rotation_unit_circle_fix [simp]:
shows "moebius_pt (moebius_rotation \<phi>) u \<in> unit_circle_set \<longleftrightarrow> u \<in> unit_circle_set"
using moebius_pt_moebius_inv_in_set unit_circle_fix_iff_unit_circle_set
by fastforce
lemma ex_rotation_mapping_u_to_positive_x_axis:
assumes "u \<noteq> 0\<^sub>h" and "u \<noteq> \<infinity>\<^sub>h"
shows "\<exists> \<phi>. moebius_pt (moebius_rotation \<phi>) u \<in> positive_x_axis"
proof-
from assms obtain c where *: "u = of_complex c"
using inf_or_of_complex
by blast
have "is_real (cis (- arg c) * c)" "Re (cis (-arg c) * c) > 0"
using "*" assms is_real_rot_to_x_axis positive_rot_to_x_axis of_complex_zero_iff
by blast+
thus ?thesis
using *
by (rule_tac x="-arg c" in exI) (simp add: positive_x_axis_def circline_set_x_axis)
qed
lemma ex_rotation_mapping_u_to_positive_y_axis:
assumes "u \<noteq> 0\<^sub>h" and "u \<noteq> \<infinity>\<^sub>h"
shows "\<exists> \<phi>. moebius_pt (moebius_rotation \<phi>) u \<in> positive_y_axis"
proof-
from assms obtain c where *: "u = of_complex c"
using inf_or_of_complex
by blast
have "is_imag (cis (pi/2 - arg c) * c)" "Im (cis (pi/2 - arg c) * c) > 0"
using "*" assms is_real_rot_to_x_axis positive_rot_to_x_axis of_complex_zero_iff
by - (simp, simp, simp add: field_simps)
thus ?thesis
using *
by (rule_tac x="pi/2-arg c" in exI) (simp add: positive_y_axis_def circline_set_y_axis)
qed
lemma wlog_rotation_to_positive_x_axis:
assumes in_disc: "u \<in> unit_disc" and not_zero: "u \<noteq> 0\<^sub>h"
assumes preserving: "\<And>\<phi> u. \<lbrakk>u \<in> unit_disc; u \<noteq> 0\<^sub>h; P (moebius_pt (moebius_rotation \<phi>) u)\<rbrakk> \<Longrightarrow> P u"
assumes x_axis: "\<And>x. \<lbrakk>is_real x; 0 < Re x; Re x < 1\<rbrakk> \<Longrightarrow> P (of_complex x)"
shows "P u"
proof-
from in_disc obtain \<phi> where *:
"moebius_pt (moebius_rotation \<phi>) u \<in> positive_x_axis"
using ex_rotation_mapping_u_to_positive_x_axis[of u]
using inf_notin_unit_disc not_zero
by blast
let ?Mu = "moebius_pt (moebius_rotation \<phi>) u"
have "P ?Mu"
proof-
let ?x = "to_complex ?Mu"
have "?Mu \<in> unit_disc" "?Mu \<noteq> 0\<^sub>h" "?Mu \<noteq> \<infinity>\<^sub>h"
using \<open>u \<in> unit_disc\<close> \<open>u \<noteq> 0\<^sub>h\<close>
by auto
hence "is_real (to_complex ?Mu)" "0 < Re ?x" "Re ?x < 1"
using *
unfolding positive_x_axis_def circline_set_x_axis
by (auto simp add: cmod_eq_Re)
thus ?thesis
using x_axis[of ?x] \<open>?Mu \<noteq> \<infinity>\<^sub>h\<close>
by simp
qed
thus ?thesis
using preserving[OF in_disc] not_zero
by simp
qed
lemma wlog_rotation_to_positive_x_axis':
assumes not_zero: "u \<noteq> 0\<^sub>h" and not_inf: "u \<noteq> \<infinity>\<^sub>h"
assumes preserving: "\<And>\<phi> u. \<lbrakk>u \<noteq> 0\<^sub>h; u \<noteq> \<infinity>\<^sub>h; P (moebius_pt (moebius_rotation \<phi>) u)\<rbrakk> \<Longrightarrow> P u"
assumes x_axis: "\<And>x. \<lbrakk>is_real x; 0 < Re x\<rbrakk> \<Longrightarrow> P (of_complex x)"
shows "P u"
proof-
from not_zero and not_inf obtain \<phi> where *:
"moebius_pt (moebius_rotation \<phi>) u \<in> positive_x_axis"
using ex_rotation_mapping_u_to_positive_x_axis[of u]
using inf_notin_unit_disc
by blast
let ?Mu = "moebius_pt (moebius_rotation \<phi>) u"
have "P ?Mu"
proof-
let ?x = "to_complex ?Mu"
have "?Mu \<noteq> 0\<^sub>h" "?Mu \<noteq> \<infinity>\<^sub>h"
using \<open>u \<noteq> \<infinity>\<^sub>h\<close> \<open>u \<noteq> 0\<^sub>h\<close>
by auto
hence "is_real (to_complex ?Mu)" "0 < Re ?x"
using *
unfolding positive_x_axis_def circline_set_x_axis
by (auto simp add: cmod_eq_Re)
thus ?thesis
using x_axis[of ?x] \<open>?Mu \<noteq> \<infinity>\<^sub>h\<close>
by simp
qed
thus ?thesis
using preserving[OF not_zero not_inf]
by simp
qed
lemma wlog_rotation_to_positive_y_axis:
assumes in_disc: "u \<in> unit_disc" and not_zero: "u \<noteq> 0\<^sub>h"
assumes preserving: "\<And>\<phi> u. \<lbrakk>u \<in> unit_disc; u \<noteq> 0\<^sub>h; P (moebius_pt (moebius_rotation \<phi>) u)\<rbrakk> \<Longrightarrow> P u"
assumes y_axis: "\<And>x. \<lbrakk>is_imag x; 0 < Im x; Im x < 1\<rbrakk> \<Longrightarrow> P (of_complex x)"
shows "P u"
proof-
from in_disc and not_zero obtain \<phi> where *:
"moebius_pt (moebius_rotation \<phi>) u \<in> positive_y_axis"
using ex_rotation_mapping_u_to_positive_y_axis[of u]
using inf_notin_unit_disc
by blast
let ?Mu = "moebius_pt (moebius_rotation \<phi>) u"
have "P ?Mu"
proof-
let ?y = "to_complex ?Mu"
have "?Mu \<in> unit_disc" "?Mu \<noteq> 0\<^sub>h" "?Mu \<noteq> \<infinity>\<^sub>h"
using \<open>u \<in> unit_disc\<close> \<open>u \<noteq> 0\<^sub>h\<close>
by auto
hence "is_imag (to_complex ?Mu)" "0 < Im ?y" "Im ?y < 1"
using *
unfolding positive_y_axis_def circline_set_y_axis
by (auto simp add: cmod_eq_Im)
thus ?thesis
using y_axis[of ?y] \<open>?Mu \<noteq> \<infinity>\<^sub>h\<close>
by simp
qed
thus ?thesis
using preserving[OF in_disc not_zero]
by simp
qed
(* ---------------------------------------------------------------------------- *)
subsection \<open>Blaschke factors are unit disc preserving transformations\<close>
(* ---------------------------------------------------------------------------- *)
text \<open>For a given point $a$, Blaschke factor transformations are of the form $k \cdot
\left(\begin{array}{cc}1 & -a\\ -\overline{a} & 1\end{array}\right)$. It is a disc preserving
Möbius transformation that maps the point $a$ to zero (by the symmetry principle, it must map the
inverse point of $a$ to infinity).\<close>
definition blaschke_cmat :: "complex \<Rightarrow> complex_mat" where
[simp]: "blaschke_cmat a = (if cmod a \<noteq> 1 then (1, -a, -cnj a, 1) else eye)"
lift_definition blaschke_mmat :: "complex \<Rightarrow> moebius_mat" is blaschke_cmat
by simp
lift_definition blaschke :: "complex \<Rightarrow> moebius" is blaschke_mmat
done
lemma blaschke_0_id [simp]: "blaschke 0 = id_moebius"
by (transfer, transfer, simp)
lemma blaschke_a_to_zero [simp]:
assumes "cmod a \<noteq> 1"
shows "moebius_pt (blaschke a) (of_complex a) = 0\<^sub>h"
using assms
by (transfer, transfer, simp)
lemma blaschke_inv_a_inf [simp]:
assumes "cmod a \<noteq> 1"
shows "moebius_pt (blaschke a) (inversion (of_complex a)) = \<infinity>\<^sub>h"
using assms
unfolding inversion_def
by (transfer, transfer) (simp add: vec_cnj_def, rule_tac x="1/(1 - a*cnj a)" in exI, simp)
lemma blaschke_inf [simp]:
assumes "cmod a < 1" and "a \<noteq> 0"
shows "moebius_pt (blaschke a) \<infinity>\<^sub>h = of_complex (- 1 / cnj a)"
using assms
by (transfer, transfer, simp add: complex_mod_sqrt_Re_mult_cnj)
lemma blaschke_0_minus_a [simp]:
assumes "cmod a \<noteq> 1"
shows "moebius_pt (blaschke a) 0\<^sub>h = ~\<^sub>h (of_complex a)"
using assms
by (transfer, transfer, simp)
lemma blaschke_unit_circle_fix [simp]:
assumes "cmod a \<noteq> 1"
shows "unit_circle_fix (blaschke a)"
using assms
by (transfer, transfer) (simp add: unitary11_gen_def mat_adj_def mat_cnj_def)
lemma blaschke_unit_disc_fix [simp]:
assumes "cmod a < 1"
shows "unit_disc_fix (blaschke a)"
using assms
proof (transfer, transfer)
fix a
assume *: "cmod a < 1"
show "unit_disc_fix_cmat (blaschke_cmat a)"
proof (cases "a = 0")
case True
thus ?thesis
by (simp add: unitary11_gen_def mat_adj_def mat_cnj_def)
next
case False
hence "Re (a * cnj a) < 1"
using *
by (metis complex_mod_sqrt_Re_mult_cnj real_sqrt_lt_1_iff)
hence "1 / Re (a * cnj a) > 1"
using False
by (smt complex_div_gt_0 less_divide_eq_1_pos one_complex.simps(1) right_inverse_eq)
hence "Re (1 / (a * cnj a)) > 1"
by (simp add: complex_is_Real_iff)
thus ?thesis
by (simp add: unitary11_gen_def mat_adj_def mat_cnj_def)
qed
qed
lemma blaschke_unit_circle_fix':
assumes "cmod a \<noteq> 1"
shows "moebius_circline (blaschke a) unit_circle = unit_circle"
using assms
using blaschke_unit_circle_fix unit_circle_fix_iff
by simp
lemma blaschke_ounit_circle_fix':
assumes "cmod a < 1"
shows "moebius_ocircline (blaschke a) ounit_circle = ounit_circle"
proof-
have "Re (a * cnj a) < 1"
using assms
by (metis complex_mod_sqrt_Re_mult_cnj real_sqrt_lt_1_iff)
thus ?thesis
using assms
using blaschke_unit_disc_fix unit_disc_fix_iff_ounit_circle
by simp
qed
lemma moebius_pt_blaschke [simp]:
assumes "cmod a \<noteq> 1" and "z \<noteq> 1 / cnj a"
shows "moebius_pt (blaschke a) (of_complex z) = of_complex ((z - a) / (1 - cnj a * z))"
using assms
proof (cases "a = 0")
case True
thus ?thesis
by auto
next
case False
thus ?thesis
using assms
apply (transfer, transfer)
apply (simp add: complex_mod_sqrt_Re_mult_cnj)
apply (rule_tac x="1 / (1 - cnj a * z)" in exI)
apply (simp add: field_simps)
done
qed
(* -------------------------------------------------------------------------- *)
subsubsection \<open>Blaschke factors for a real point $a$\<close>
(* -------------------------------------------------------------------------- *)
text \<open>If the point $a$ is real, the Blaschke factor preserve x-axis and the upper and the lower
halfplane.\<close>
lemma blaschke_real_preserve_x_axis [simp]:
assumes "is_real a" and "cmod a < 1"
shows "moebius_pt (blaschke a) z \<in> circline_set x_axis \<longleftrightarrow> z \<in> circline_set x_axis"
proof (cases "a = 0")
case True
thus ?thesis
by simp
next
case False
have "cmod a \<noteq> 1"
using assms
by linarith
let ?a = "of_complex a"
let ?i = "inversion ?a"
let ?M = "moebius_pt (blaschke a)"
have *: "?M ?a = 0\<^sub>h" "?M ?i = \<infinity>\<^sub>h" "?M 0\<^sub>h = of_complex (-a)"
using \<open>cmod a \<noteq> 1\<close> blaschke_a_to_zero[of a] blaschke_inv_a_inf[of a] blaschke_0_minus_a[of a]
by auto
let ?Mx = "moebius_circline (blaschke a) x_axis"
have "?a \<in> circline_set x_axis" "?i \<in> circline_set x_axis" "0\<^sub>h \<in> circline_set x_axis"
using \<open>is_real a\<close> \<open>a \<noteq> 0\<close> eq_cnj_iff_real[of a]
by auto
hence "0\<^sub>h \<in> circline_set ?Mx" "\<infinity>\<^sub>h \<in> circline_set ?Mx" "of_complex (-a) \<in> circline_set ?Mx"
using *
apply -
apply (force simp add: image_iff)+
apply (simp add: image_iff, rule_tac x="0\<^sub>h" in bexI, simp_all)
done
moreover
have "0\<^sub>h \<in> circline_set x_axis" "\<infinity>\<^sub>h \<in> circline_set x_axis" "of_complex (-a) \<in> circline_set x_axis"
using \<open>is_real a\<close>
by auto
moreover
have "of_complex (-a) \<noteq> 0\<^sub>h"
using \<open>a \<noteq> 0\<close>
by simp
hence "0\<^sub>h \<noteq> of_complex (-a)"
by metis
hence "\<exists>!H. 0\<^sub>h \<in> circline_set H \<and> \<infinity>\<^sub>h \<in> circline_set H \<and> of_complex (- a) \<in> circline_set H"
using unique_circline_set[of "0\<^sub>h" "\<infinity>\<^sub>h" "of_complex (-a)"] \<open>a \<noteq> 0\<close>
by simp
ultimately
have "moebius_circline (blaschke a) x_axis = x_axis"
by auto
thus ?thesis
by (metis circline_set_moebius_circline_iff)
qed
lemma blaschke_real_preserve_sgn_Im [simp]:
assumes "is_real a" and "cmod a < 1" and "z \<noteq> \<infinity>\<^sub>h" and "z \<noteq> inversion (of_complex a)"
shows "sgn (Im (to_complex (moebius_pt (blaschke a) z))) = sgn (Im (to_complex z))"
proof (cases "a = 0")
case True
thus ?thesis
by simp
next
case False
obtain z' where z': "z = of_complex z'"
using inf_or_of_complex[of z] \<open>z \<noteq> \<infinity>\<^sub>h\<close>
by auto
have "z' \<noteq> 1 / cnj a"
using assms z' \<open>a \<noteq> 0\<close>
by (auto simp add: of_complex_inj)
moreover
have "a * cnj a \<noteq> 1"
using \<open>cmod a < 1\<close>
by auto (simp add: complex_mod_sqrt_Re_mult_cnj)
moreover
have "sgn (Im ((z' - a) / (1 - a * z'))) = sgn (Im z')"
proof-
have "a * z' \<noteq> 1"
using \<open>is_real a\<close> \<open>z' \<noteq> 1 / cnj a\<close> \<open>a \<noteq> 0\<close> eq_cnj_iff_real[of a]
by (simp add: field_simps)
moreover
have "Re (1 - a\<^sup>2) > 0"
using \<open>is_real a\<close> \<open>cmod a < 1\<close>
by (smt Re_power2 minus_complex.simps(1) norm_complex_def one_complex.simps(1) power2_less_0 real_sqrt_lt_1_iff)
moreover
have "Im ((z' - a) / (1 - a * z')) = Re (((1 - a\<^sup>2) * Im z') / (cmod (1 - a*z'))\<^sup>2)"
proof-
have "1 - a * cnj z' \<noteq> 0"
using \<open>z' \<noteq> 1 / cnj a\<close>
by (metis Im_complex_div_eq_0 complex_cnj_zero_iff diff_eq_diff_eq diff_numeral_special(9) eq_divide_imp is_real_div mult_not_zero one_complex.simps(2) zero_neq_one)
hence "Im ((z' - a) / (1 - a * z')) = Im (((z' - a) * (1 - a * cnj z')) / ((1 - a * z') * cnj (1 - a * z')))"
using \<open>is_real a\<close> eq_cnj_iff_real[of a]
by simp
also have "... = Im ((z' - a - a * z' * cnj z' + a\<^sup>2 * cnj z') / (cmod (1 - a*z'))\<^sup>2)"
unfolding complex_mult_cnj_cmod
by (simp add: power2_eq_square field_simps)
finally show ?thesis
using \<open>is_real a\<close>
by (simp add: field_simps)
qed
moreover
have "0 < (1 - (Re a)\<^sup>2) * Im z' / (cmod (1 - a * z'))\<^sup>2 \<Longrightarrow> Im z' > 0"
using `is_real a` `0 < Re (1 - a\<^sup>2)`
by (smt Re_power_real divide_le_0_iff minus_complex.simps(1) not_sum_power2_lt_zero one_complex.simps(1) zero_less_mult_pos)
ultimately
show ?thesis
unfolding sgn_real_def
using \<open>cmod a < 1\<close> \<open>a * z' \<noteq> 1\<close> \<open>is_real a\<close>
by (auto simp add: cmod_eq_Re)
qed
ultimately
show ?thesis
using assms z' moebius_pt_blaschke[of a z'] \<open>is_real a\<close> eq_cnj_iff_real[of a]
by simp
qed
lemma blaschke_real_preserve_sgn_arg [simp]:
assumes "is_real a" and "cmod a < 1" and "z \<notin> circline_set x_axis"
shows "sgn (arg (to_complex (moebius_pt (blaschke a) z))) = sgn (arg (to_complex z))"
proof-
have "z \<noteq> \<infinity>\<^sub>h"
using assms
using special_points_on_x_axis''(3) by blast
moreover
have "z \<noteq> inversion (of_complex a)"
using assms
by (metis calculation circline_equation_x_axis circline_set_x_axis_I conjugate_of_complex inversion_of_complex inversion_sym is_real_inversion o_apply of_complex_zero reciprocal_zero to_complex_of_complex)
ultimately
show ?thesis
using blaschke_real_preserve_sgn_Im[OF assms(1) assms(2), of z]
by (smt arg_Im_sgn assms(3) circline_set_x_axis_I norm_sgn of_complex_to_complex)
qed
(* -------------------------------------------------------------------------- *)
subsubsection \<open>Inverse Blaschke transform\<close>
(* -------------------------------------------------------------------------- *)
definition inv_blaschke_cmat :: "complex \<Rightarrow> complex_mat" where
[simp]: "inv_blaschke_cmat a = (if cmod a \<noteq> 1 then (1, a, cnj a, 1) else eye)"
lift_definition inv_blaschke_mmat :: "complex \<Rightarrow> moebius_mat" is inv_blaschke_cmat
by simp
lift_definition inv_blaschke :: "complex \<Rightarrow> moebius" is inv_blaschke_mmat
done
lemma inv_blaschke_neg [simp]: "inv_blaschke a = blaschke (-a)"
by (transfer, transfer) simp
lemma inv_blaschke:
assumes "cmod a \<noteq> 1"
shows "blaschke a + inv_blaschke a = 0"
apply simp
apply (transfer, transfer)
by auto (rule_tac x="1/(1 - a*cnj a)" in exI, simp)
lemma ex_unit_disc_fix_mapping_u_to_zero:
assumes "u \<in> unit_disc"
shows "\<exists> M. unit_disc_fix M \<and> moebius_pt M u = 0\<^sub>h"
proof-
from assms obtain c where *: "u = of_complex c"
by (metis inf_notin_unit_disc inf_or_of_complex)
hence "cmod c < 1"
using assms unit_disc_iff_cmod_lt_1
by simp
thus ?thesis
using *
by (rule_tac x="blaschke c" in exI)
(smt blaschke_a_to_zero blaschke_ounit_circle_fix' unit_disc_fix_iff_ounit_circle)
qed
lemma wlog_zero:
assumes in_disc: "u \<in> unit_disc"
assumes preserving: "\<And> a u. \<lbrakk>u \<in> unit_disc; cmod a < 1; P (moebius_pt (blaschke a) u)\<rbrakk> \<Longrightarrow> P u"
assumes zero: "P 0\<^sub>h"
shows "P u"
proof-
have *: "moebius_pt (blaschke (to_complex u)) u = 0\<^sub>h"
by (smt blaschke_a_to_zero in_disc inf_notin_unit_disc of_complex_to_complex unit_disc_iff_cmod_lt_1)
thus ?thesis
using preserving[of u "to_complex u"] in_disc zero
using inf_or_of_complex[of u]
by auto
qed
lemma wlog_real_zero:
assumes in_disc: "u \<in> unit_disc" and real: "is_real (to_complex u)"
assumes preserving: "\<And> a u. \<lbrakk>u \<in> unit_disc; is_real a; cmod a < 1; P (moebius_pt (blaschke a) u)\<rbrakk> \<Longrightarrow> P u"
assumes zero: "P 0\<^sub>h"
shows "P u"
proof-
have *: "moebius_pt (blaschke (to_complex u)) u = 0\<^sub>h"
by (smt blaschke_a_to_zero in_disc inf_notin_unit_disc of_complex_to_complex unit_disc_iff_cmod_lt_1)
thus ?thesis
using preserving[of u "to_complex u"] in_disc zero real
using inf_or_of_complex[of u]
by auto
qed
lemma unit_disc_fix_transitive:
assumes in_disc: "u \<in> unit_disc" and "u' \<in> unit_disc"
shows "\<exists> M. unit_disc_fix M \<and> moebius_pt M u = u'"
proof-
have "\<forall> u \<in> unit_disc. \<exists> M. unit_disc_fix M \<and> moebius_pt M u = u'" (is "?P u'")
proof (rule wlog_zero)
show "u' \<in> unit_disc" by fact
next
show "?P 0\<^sub>h"
by (simp add: ex_unit_disc_fix_mapping_u_to_zero)
next
fix a u
assume "cmod a < 1" and *: "?P (moebius_pt (blaschke a) u)"
show "?P u"
proof
fix u'
assume "u' \<in> unit_disc"
then obtain M' where "unit_disc_fix M'" "moebius_pt M' u' = moebius_pt (blaschke a) u"
using *
by auto
thus "\<exists>M. unit_disc_fix M \<and> moebius_pt M u' = u"
using \<open>cmod a < 1\<close> blaschke_unit_disc_fix[of a]
using unit_disc_fix_moebius_comp[of "- blaschke a" "M'"]
using unit_disc_fix_moebius_inv[of "blaschke a"]
by (rule_tac x="(- (blaschke a)) + M'" in exI, simp)
qed
qed
thus ?thesis
using assms
by auto
qed
(* -------------------------------------------------------------------------- *)
subsection \<open>Decomposition of unit disc preserving Möbius transforms\<close>
(* -------------------------------------------------------------------------- *)
text \<open>Each transformation preserving unit disc can be decomposed to a rotation around the origin and
a Blaschke factors that maps a point within the unit disc to zero.\<close>
lemma unit_disc_fix_decompose_blaschke_rotation:
assumes "unit_disc_fix M"
shows "\<exists> k \<phi>. cmod k < 1 \<and> M = moebius_rotation \<phi> + blaschke k"
using assms
unfolding moebius_rotation_def moebius_similarity_def
proof (simp, transfer, transfer)
fix M
assume *: "mat_det M \<noteq> 0" "unit_disc_fix_cmat M"
then obtain k a b :: complex where
**: "k \<noteq> 0" "mat_det (a, b, cnj b, cnj a) \<noteq> 0" "M = k *\<^sub>s\<^sub>m (a, b, cnj b, cnj a)"
using unitary11_gen_iff[of M]
by auto
have "a \<noteq> 0"
using * **
by auto
then obtain a' k' \<phi>
where ***: "k' \<noteq> 0 \<and> a' * cnj a' \<noteq> 1 \<and> M = k' *\<^sub>s\<^sub>m (cis \<phi>, 0, 0, 1) *\<^sub>m\<^sub>m (1, - a', - cnj a', 1)"
using ** unitary11_gen_cis_blaschke[of k M a b]
- by auto
+ by auto blast
have "a' = 0 \<or> 1 < 1 / (cmod a')\<^sup>2"
using * *** complex_mult_cnj_cmod[of a']
by simp
hence "cmod a' < 1"
by (smt less_divide_eq_1_pos norm_zero one_less_power one_power2 pos2)
thus "\<exists>k. cmod k < 1 \<and>
(\<exists>\<phi>. moebius_cmat_eq M (moebius_comp_cmat (mk_moebius_cmat (cis \<phi>) 0 0 1) (blaschke_cmat k)))"
using ***
apply (rule_tac x=a' in exI)
apply simp
apply (rule_tac x=\<phi> in exI)
apply simp
apply (rule_tac x="1/k'" in exI)
by auto
qed
lemma wlog_unit_disc_fix:
assumes "unit_disc_fix M"
assumes b: "\<And> k. cmod k < 1 \<Longrightarrow> P (blaschke k)"
assumes r: "\<And> \<phi>. P (moebius_rotation \<phi>)"
assumes comp: "\<And>M1 M2. \<lbrakk>unit_disc_fix M1; P M1; unit_disc_fix M2; P M2\<rbrakk> \<Longrightarrow> P (M1 + M2)"
shows "P M"
using assms
using unit_disc_fix_decompose_blaschke_rotation[OF assms(1)]
using blaschke_unit_disc_fix
by auto
lemma ex_unit_disc_fix_to_zero_positive_x_axis:
assumes "u \<in> unit_disc" and "v \<in> unit_disc" and "u \<noteq> v"
shows "\<exists> M. unit_disc_fix M \<and>
moebius_pt M u = 0\<^sub>h \<and> moebius_pt M v \<in> positive_x_axis"
proof-
from assms obtain B where
*: "unit_disc_fix B" "moebius_pt B u = 0\<^sub>h"
using ex_unit_disc_fix_mapping_u_to_zero
by blast
let ?v = "moebius_pt B v"
have "?v \<in> unit_disc"
using \<open>v \<in> unit_disc\<close> *
by auto
hence "?v \<noteq> \<infinity>\<^sub>h"
using inf_notin_unit_disc by auto
have "?v \<noteq> 0\<^sub>h"
using \<open>u \<noteq> v\<close> *
by (metis moebius_pt_invert)
obtain R where
"unit_disc_fix R"
"moebius_pt R 0\<^sub>h = 0\<^sub>h" "moebius_pt R ?v \<in> positive_x_axis"
using ex_rotation_mapping_u_to_positive_x_axis[of ?v] \<open>?v \<noteq> 0\<^sub>h\<close> \<open>?v \<noteq> \<infinity>\<^sub>h\<close>
using moebius_pt_rotation_inf_iff moebius_pt_moebius_rotation_zero unit_disc_fix_rotation
by blast
thus ?thesis
using * moebius_comp[of R B, symmetric]
using unit_disc_fix_moebius_comp
by (rule_tac x="R + B" in exI) (simp add: comp_def)
qed
lemma wlog_x_axis:
assumes in_disc: "u \<in> unit_disc" "v \<in> unit_disc"
assumes preserved: "\<And> M u v. \<lbrakk>unit_disc_fix M; u \<in> unit_disc; v \<in> unit_disc; P (moebius_pt M u) (moebius_pt M v)\<rbrakk> \<Longrightarrow> P u v"
assumes axis: "\<And> x. \<lbrakk>is_real x; 0 \<le> Re x; Re x < 1\<rbrakk> \<Longrightarrow> P 0\<^sub>h (of_complex x)"
shows "P u v"
proof (cases "u = v")
case True
have "P u u" (is "?Q u")
proof (rule wlog_zero[where P="?Q"])
show "u \<in> unit_disc"
by fact
next
show "?Q 0\<^sub>h"
using axis[of 0]
by simp
next
fix a u
assume "u \<in> unit_disc" "cmod a < 1" "?Q (moebius_pt (blaschke a) u)"
thus "?Q u"
using preserved[of "blaschke a" u u]
using blaschke_unit_disc_fix[of a]
by simp
qed
thus ?thesis
using True
by simp
next
case False
from in_disc obtain M where
*: "unit_disc_fix M" "moebius_pt M u = 0\<^sub>h" "moebius_pt M v \<in> positive_x_axis"
using ex_unit_disc_fix_to_zero_positive_x_axis False
by auto
then obtain x where **: "moebius_pt M v = of_complex x" "is_real x"
unfolding positive_x_axis_def circline_set_x_axis
by auto
moreover
have "of_complex x \<in> unit_disc"
using \<open>unit_disc_fix M\<close> \<open>v \<in> unit_disc\<close> **
using unit_disc_fix_discI
by fastforce
hence "0 < Re x" "Re x < 1"
using \<open>moebius_pt M v \<in> positive_x_axis\<close> **
by (auto simp add: positive_x_axis_def cmod_eq_Re)
ultimately
have "P 0\<^sub>h (of_complex x)"
using \<open>is_real x\<close> axis
by auto
thus ?thesis
using preserved[OF *(1) assms(1-2)] *(2) **(1)
by simp
qed
lemma wlog_positive_x_axis:
assumes in_disc: "u \<in> unit_disc" "v \<in> unit_disc" "u \<noteq> v"
assumes preserved: "\<And> M u v. \<lbrakk>unit_disc_fix M; u \<in> unit_disc; v \<in> unit_disc; u \<noteq> v; P (moebius_pt M u) (moebius_pt M v)\<rbrakk> \<Longrightarrow> P u v"
assumes axis: "\<And> x. \<lbrakk>is_real x; 0 < Re x; Re x < 1\<rbrakk> \<Longrightarrow> P 0\<^sub>h (of_complex x)"
shows "P u v"
proof-
have "u \<noteq> v \<longrightarrow> P u v" (is "?Q u v")
proof (rule wlog_x_axis)
show "u \<in> unit_disc" "v \<in> unit_disc"
by fact+
next
fix M u v
assume "unit_disc_fix M" "u \<in> unit_disc" "v \<in> unit_disc"
"?Q (moebius_pt M u) (moebius_pt M v)"
thus "?Q u v"
using preserved[of M u v]
using moebius_pt_invert
by blast
next
fix x
assume "is_real x" "0 \<le> Re x" "Re x < 1"
thus "?Q 0\<^sub>h (of_complex x)"
using axis[of x] of_complex_zero_iff[of x] complex.expand[of x 0]
by fastforce
qed
thus ?thesis
using \<open>u \<noteq> v\<close>
by simp
qed
(* -------------------------------------------------------------------------- *)
subsection \<open>All functions that fix the unit disc\<close>
(* -------------------------------------------------------------------------- *)
text \<open>It can be proved that continuous functions that fix the unit disc are either actions of
Möbius transformations that fix the unit disc (homographies), or are compositions of actions of
Möbius transformations that fix the unit disc and the conjugation (antihomographies). We postulate
this as a definition, but it this characterisation could also be formally shown (we do not need this
for our further applications).\<close>
definition unit_disc_fix_f where
"unit_disc_fix_f f \<longleftrightarrow>
(\<exists> M. unit_disc_fix M \<and> (f = moebius_pt M \<or> f = moebius_pt M \<circ> conjugate))"
text \<open>Unit disc fixing functions really fix unit disc.\<close>
lemma unit_disc_fix_f_unit_disc:
assumes "unit_disc_fix_f M"
shows "M ` unit_disc = unit_disc"
using assms
unfolding unit_disc_fix_f_def
using image_comp
by force
text \<open>Actions of unit disc fixing Möbius transformations (unit disc fixing homographies) are unit
disc fixing functions.\<close>
lemma unit_disc_fix_f_moebius_pt [simp]:
assumes "unit_disc_fix M"
shows "unit_disc_fix_f (moebius_pt M)"
using assms
unfolding unit_disc_fix_f_def
by auto
text \<open>Compositions of unit disc fixing Möbius transformations and conjugation (unit disc fixing
antihomographies) are unit disc fixing functions.\<close>
lemma unit_disc_fix_conjugate_moebius [simp]:
assumes "unit_disc_fix M"
shows "unit_disc_fix (conjugate_moebius M)"
proof-
have "\<And>a aa ab b. \<lbrakk>1 < Re (a * b / (aa * ab)); \<not> 1 < Re (cnj a * cnj b / (cnj aa * cnj ab))\<rbrakk> \<Longrightarrow> aa = 0"
by (metis cnj.simps(1) complex_cnj_divide complex_cnj_mult)
thus ?thesis
using assms
by (transfer, transfer)
(auto simp add: mat_cnj_def unitary11_gen_def mat_adj_def field_simps)
qed
lemma unit_disc_fix_conjugate_comp_moebius [simp]:
assumes "unit_disc_fix M"
shows "unit_disc_fix_f (conjugate \<circ> moebius_pt M)"
using assms
apply (subst conjugate_moebius)
apply (simp add: unit_disc_fix_f_def)
apply (rule_tac x="conjugate_moebius M" in exI, simp)
done
text \<open>Uniti disc fixing functions form a group under function composition.\<close>
lemma unit_disc_fix_f_comp [simp]:
assumes "unit_disc_fix_f f1" and "unit_disc_fix_f f2"
shows "unit_disc_fix_f (f1 \<circ> f2)"
using assms
apply (subst (asm) unit_disc_fix_f_def)
apply (subst (asm) unit_disc_fix_f_def)
proof safe
fix M M'
assume "unit_disc_fix M" "unit_disc_fix M'"
thus "unit_disc_fix_f (moebius_pt M \<circ> moebius_pt M')"
unfolding unit_disc_fix_f_def
by (rule_tac x="M + M'" in exI) auto
next
fix M M'
assume "unit_disc_fix M" "unit_disc_fix M'"
thus "unit_disc_fix_f (moebius_pt M \<circ> (moebius_pt M' \<circ> conjugate))"
unfolding unit_disc_fix_f_def
by (subst comp_assoc[symmetric])+
(rule_tac x="M + M'" in exI, auto)
next
fix M M'
assume "unit_disc_fix M" "unit_disc_fix M'"
thus "unit_disc_fix_f ((moebius_pt M \<circ> conjugate) \<circ> moebius_pt M')"
unfolding unit_disc_fix_f_def
by (subst comp_assoc, subst conjugate_moebius, subst comp_assoc[symmetric])+
(rule_tac x="M + conjugate_moebius M'" in exI, auto)
next
fix M M'
assume "unit_disc_fix M" "unit_disc_fix M'"
thus "unit_disc_fix_f ((moebius_pt M \<circ> conjugate) \<circ> (moebius_pt M' \<circ> conjugate))"
apply (subst comp_assoc[symmetric], subst comp_assoc)
apply (subst conjugate_moebius, subst comp_assoc, subst comp_assoc)
apply (simp add: unit_disc_fix_f_def)
apply (rule_tac x="M + conjugate_moebius M'" in exI, auto)
done
qed
lemma unit_disc_fix_f_inv:
assumes "unit_disc_fix_f M"
shows "unit_disc_fix_f (inv M)"
using assms
apply (subst (asm) unit_disc_fix_f_def)
proof safe
fix M
assume "unit_disc_fix M"
have "inv (moebius_pt M) = moebius_pt (-M)"
by (rule ext) (simp add: moebius_inv)
thus "unit_disc_fix_f (inv (moebius_pt M))"
using \<open>unit_disc_fix M\<close>
unfolding unit_disc_fix_f_def
by (rule_tac x="-M" in exI, simp)
next
fix M
assume "unit_disc_fix M"
have "inv (moebius_pt M \<circ> conjugate) = conjugate \<circ> inv (moebius_pt M)"
by (subst o_inv_distrib, simp_all)
also have "... = conjugate \<circ> (moebius_pt (-M))"
using moebius_inv
by auto
also have "... = moebius_pt (conjugate_moebius (-M)) \<circ> conjugate"
by (simp add: conjugate_moebius)
finally
show "unit_disc_fix_f (inv (moebius_pt M \<circ> conjugate))"
using \<open>unit_disc_fix M\<close>
unfolding unit_disc_fix_f_def
by (rule_tac x="conjugate_moebius (-M)" in exI, simp)
qed
(* -------------------------------------------------------------------------- *)
subsubsection \<open>Action of unit disc fixing functions on circlines\<close>
(* -------------------------------------------------------------------------- *)
definition unit_disc_fix_f_circline where
"unit_disc_fix_f_circline f H =
(if \<exists> M. unit_disc_fix M \<and> f = moebius_pt M then
moebius_circline (THE M. unit_disc_fix M \<and> f = moebius_pt M) H
else if \<exists> M. unit_disc_fix M \<and> f = moebius_pt M \<circ> conjugate then
(moebius_circline (THE M. unit_disc_fix M \<and> f = moebius_pt M \<circ> conjugate) \<circ> conjugate_circline) H
else
H)"
lemma unique_moebius_pt_conjugate:
assumes "moebius_pt M1 \<circ> conjugate = moebius_pt M2 \<circ> conjugate"
shows "M1 = M2"
proof-
from assms have "moebius_pt M1 = moebius_pt M2"
using conjugate_conjugate_comp rewriteL_comp_comp2 by fastforce
thus ?thesis
using unique_moebius_pt
by auto
qed
lemma unit_disc_fix_f_circline_direct:
assumes "unit_disc_fix M" and "f = moebius_pt M"
shows "unit_disc_fix_f_circline f H = moebius_circline M H"
proof-
have "M = (THE M. unit_disc_fix M \<and> f = moebius_pt M)"
using assms
using theI_unique[of "\<lambda> M. unit_disc_fix M \<and> f = moebius_pt M" M]
using unique_moebius_pt[of M]
by auto
thus ?thesis
using assms
unfolding unit_disc_fix_f_circline_def
by auto
qed
lemma unit_disc_fix_f_circline_indirect:
assumes "unit_disc_fix M" and "f = moebius_pt M \<circ> conjugate"
shows "unit_disc_fix_f_circline f H = ((moebius_circline M) \<circ> conjugate_circline) H"
proof-
have "\<not> (\<exists> M. unit_disc_fix M \<and> f = moebius_pt M)"
using assms homography_antihomography_exclusive[of f]
unfolding is_homography_def is_antihomography_def is_moebius_def
by auto
moreover
have "M = (THE M. unit_disc_fix M \<and> f = moebius_pt M \<circ> conjugate)"
using assms
using theI_unique[of "\<lambda> M. unit_disc_fix M \<and> f = moebius_pt M \<circ> conjugate" M]
using unique_moebius_pt_conjugate[of M]
by auto
ultimately
show ?thesis
using assms
unfolding unit_disc_fix_f_circline_def
by metis
qed
text \<open>Disc automorphisms - it would be nice to show that there are no disc automorphisms other than
unit disc fixing homographies and antihomographies, but this part of the theory is not yet
developed.\<close>
definition is_disc_aut where "is_disc_aut f \<longleftrightarrow> bij_betw f unit_disc unit_disc"
end
\ No newline at end of file
diff --git a/thys/Complx/OG_Soundness.thy b/thys/Complx/OG_Soundness.thy
--- a/thys/Complx/OG_Soundness.thy
+++ b/thys/Complx/OG_Soundness.thy
@@ -1,2029 +1,2036 @@
(*
* Copyright 2016, Data61, CSIRO
*
* This software may be distributed and modified according to the terms of
* the BSD 2-Clause license. Note that NO WARRANTY is provided.
* See "LICENSE_BSD2.txt" for details.
*
* @TAG(DATA61_BSD)
*)
section \<open>Soundness proof of Owicki-Gries w.r.t.
COMPLX small-step semantics\<close>
theory OG_Soundness
imports
OG_Hoare
SeqCatch_decomp
begin
lemma pre_weaken_pre:
" x \<in> pre P \<Longrightarrow> x \<in> pre (weaken_pre P P')"
by (induct P, clarsimp+)
lemma oghoare_Skip[rule_format, OF _ refl]:
"\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P c Q,A \<Longrightarrow> c = Skip \<longrightarrow>
(\<exists>P'. P = AnnExpr P' \<and> P' \<subseteq> Q)"
apply (induct rule: oghoare_induct, simp_all)
apply clarsimp
apply (rename_tac \<Gamma> \<Theta> F P Q A P' Q' A' P'')
apply(case_tac P, simp_all)
by force
lemma oghoare_Throw[rule_format, OF _ refl]:
"\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P c Q,A \<Longrightarrow> c = Throw \<longrightarrow>
(\<exists>P'. P = AnnExpr P' \<and> P' \<subseteq> A)"
apply (induct rule: oghoare_induct, simp_all)
apply clarsimp
apply (rename_tac \<Gamma> \<Theta> F P Q A P' Q' A' P'')
apply(case_tac P, simp_all)
by force
lemma oghoare_Basic[rule_format, OF _ refl]:
"\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P c Q,A \<Longrightarrow> c = Basic f \<longrightarrow>
(\<exists>P'. P = AnnExpr P' \<and> P' \<subseteq> {x. f x \<in> Q})"
apply (induct rule: oghoare_induct, simp_all)
apply clarsimp
apply (rename_tac \<Gamma> \<Theta> F P Q A P' Q' A' P'')
apply(case_tac P, simp_all)
by force
lemma oghoare_Spec[rule_format, OF _ refl]:
"\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P c Q,A \<Longrightarrow> c = Spec r \<longrightarrow>
(\<exists>P'. P = AnnExpr P' \<and> P' \<subseteq> {s. (\<forall>t. (s, t) \<in> r \<longrightarrow> t \<in> Q) \<and> (\<exists>t. (s, t) \<in> r)})"
apply (induct rule: oghoare_induct, simp_all)
apply clarsimp
apply (rename_tac \<Gamma> \<Theta> F P Q A P' Q' A' P'')
apply(case_tac P, simp_all)
by force
lemma oghoare_DynCom[rule_format, OF _ refl]:
"\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P c Q,A \<Longrightarrow> c = (DynCom c') \<longrightarrow>
(\<exists>r ad. P = (AnnRec r ad) \<and> r \<subseteq> pre ad \<and> (\<forall>s\<in>r. \<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> ad (c' s) Q,A))"
apply (induct rule: oghoare_induct, simp_all)
apply clarsimp
apply clarsimp
apply (rename_tac \<Gamma> \<Theta> F P Q A P' Q' A' P'' x)
apply(case_tac P, simp_all)
apply clarsimp
apply (rename_tac P s)
apply (drule_tac x=s in bspec, simp)
apply (rule Conseq)
apply (rule_tac x="{}" in exI)
apply (fastforce)
done
lemma oghoare_Guard[rule_format, OF _ refl]:
"\<Gamma>,\<Theta>\<turnstile>\<^bsub>/F\<^esub> P c Q,A \<Longrightarrow> c = Guard f g d \<longrightarrow>
(\<exists>P' r . P = AnnRec r P' \<and>
\<Gamma>,\<Theta>\<turnstile>\<^bsub>/F\<^esub> P' d Q,A \<and>
r \<inter> g \<subseteq> pre P' \<and>
(r \<inter> -g \<noteq> {} \<longrightarrow> f \<in> F))"
apply (induct rule: oghoare_induct, simp_all)
apply force
apply clarsimp
apply (rename_tac \<Gamma> \<Theta> F P Q A P' Q' A' P'' r)
apply (case_tac P, simp_all)
apply clarsimp
apply (rename_tac r)
apply(rule conjI)
apply (rule Conseq)
apply (rule_tac x="{}" in exI)
apply (rule_tac x="Q'" in exI)
apply (rule_tac x="A'" in exI)
apply (clarsimp)
apply (case_tac "(r \<union> P') \<inter> g \<noteq> {}")
apply fast
apply clarsimp
apply(drule equalityD1, force)
done
lemma oghoare_Await[rule_format, OF _ refl]:
"\<Gamma>, \<Theta>\<turnstile>\<^bsub>/F\<^esub> P x Q,A \<Longrightarrow> \<forall>b c. x = Await b c \<longrightarrow>
(\<exists>r P' Q' A'. P = AnnRec r P' \<and> \<Gamma>, \<Theta>\<tturnstile>\<^bsub>/F\<^esub>(r \<inter> b) P' c Q',A' \<and> atom_com c
\<and> Q' \<subseteq> Q \<and> A' \<subseteq> A)"
+ supply [[simproc del: defined_all]]
apply (induct rule: oghoare_induct, simp_all)
apply (rename_tac \<Gamma> \<Theta> F r P Q A)
apply (rule_tac x=Q in exI)
apply (rule_tac x=A in exI)
apply clarsimp
apply (rename_tac \<Gamma> \<Theta> F P c Q A)
apply clarsimp
apply(case_tac P, simp_all)
apply (rename_tac P'' Q'' A'' x y)
apply (rule_tac x=Q'' in exI)
apply (rule_tac x=A'' in exI)
apply clarsimp
apply (rule conjI[rotated])
apply blast
apply (erule SeqConseq[rotated])
apply simp
apply simp
apply blast
done
lemma oghoare_Seq[rule_format, OF _ refl]:
"\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P x Q,A \<Longrightarrow> \<forall>p1 p2. x = Seq p1 p2 \<longrightarrow>
(\<exists> P\<^sub>1 P\<^sub>2 P' Q' A'. P = AnnComp P\<^sub>1 P\<^sub>2 \<and> \<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P\<^sub>1 p1 P', A' \<and> P' \<subseteq> pre P\<^sub>2 \<and>
\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P\<^sub>2 p2 Q',A' \<and>
Q' \<subseteq> Q \<and> A' \<subseteq> A)"
apply (induct rule: oghoare_induct, simp_all)
apply blast
apply (rename_tac \<Gamma> \<Theta> F P c Q A)
apply clarsimp
apply (rename_tac P'' Q'' A'')
apply(case_tac P, simp_all)
apply clarsimp
apply (rule_tac x="P''" in exI)
apply (rule_tac x="Q''" in exI)
apply (rule_tac x="A''" in exI)
apply clarsimp
apply (rule conjI)
apply (rule Conseq)
apply (rule_tac x="P'" in exI)
apply (rule_tac x="P''" in exI)
apply (rule_tac x="A''" in exI)
apply simp
apply fastforce
done
lemma oghoare_Catch[rule_format, OF _ refl]:
"\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P x Q,A \<Longrightarrow> \<forall>p1 p2. x = Catch p1 p2 \<longrightarrow>
(\<exists> P\<^sub>1 P\<^sub>2 P' Q' A'. P = AnnComp P\<^sub>1 P\<^sub>2 \<and> \<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P\<^sub>1 p1 Q', P' \<and> P' \<subseteq> pre P\<^sub>2 \<and>
\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P\<^sub>2 p2 Q',A' \<and>
Q' \<subseteq> Q \<and> A' \<subseteq> A)"
apply (induct rule: oghoare_induct, simp_all)
apply blast
apply (rename_tac \<Gamma> \<Theta> F P c Q A)
apply clarsimp
apply(case_tac P, simp_all)
apply clarsimp
apply (rename_tac P'' Q'' A'' x)
apply (rule_tac x="P''" in exI)
apply (rule_tac x="Q''" in exI)
apply clarsimp
apply (rule conjI)
apply (rule Conseq)
apply (rule_tac x=P' in exI)
apply fastforce
apply (rule_tac x="A''" in exI)
apply clarsimp
apply fastforce
done
lemma oghoare_Cond[rule_format, OF _ refl]:
"\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P x Q,A \<Longrightarrow>
\<forall>c\<^sub>1 c\<^sub>2 b. x = (Cond b c\<^sub>1 c\<^sub>2) \<longrightarrow>
(\<exists>P' P\<^sub>1 P\<^sub>2 Q' A'. P = (AnnBin P' P\<^sub>1 P\<^sub>2) \<and>
P' \<subseteq> {s. (s\<in>b \<longrightarrow> s\<in>pre P\<^sub>1) \<and> (s\<notin>b \<longrightarrow> s\<in>pre P\<^sub>2)} \<and>
\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q',A' \<and>
\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q',A' \<and> Q' \<subseteq> Q \<and> A' \<subseteq> A)"
apply (induct rule: oghoare_induct, simp_all)
apply (rule conjI)
apply fastforce
apply (rename_tac Q A P\<^sub>2 c\<^sub>2 r b)
apply (rule_tac x=Q in exI)
apply (rule_tac x=A in exI)
apply fastforce
apply (rename_tac \<Gamma> \<Theta> F P c Q A)
apply clarsimp
apply (case_tac P, simp_all)
apply fastforce
done
lemma oghoare_While[rule_format, OF _ refl]:
"\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P x Q,A \<Longrightarrow>
\<forall> b c. x = While b c \<longrightarrow>
(\<exists> r i P' A' Q'. P = AnnWhile r i P' \<and>
\<Gamma>, \<Theta>\<turnstile>\<^bsub>/F\<^esub> P' c i,A' \<and>
r \<subseteq> i \<and>
i \<inter> b \<subseteq> pre P' \<and>
i \<inter> -b \<subseteq> Q' \<and>
Q' \<subseteq> Q \<and> A' \<subseteq> A)"
apply (induct rule: oghoare_induct, simp_all)
apply blast
apply (rename_tac \<Gamma> \<Theta> F P c Q A)
apply clarsimp
apply (rename_tac P' Q' A' b ca r i P'' A'' Q'')
apply (case_tac P; simp)
apply (rule_tac x= A'' in exI)
apply (rule conjI)
apply blast
apply clarsimp
apply (rule_tac x= "Q'" in exI)
by fast
lemma oghoare_Call:
"\<Gamma>,\<Theta>\<turnstile>\<^bsub>/F\<^esub> P x Q,A \<Longrightarrow>
\<forall>p. x = Call p \<longrightarrow>
(\<exists>r n.
P = (AnnCall r n) \<and>
(\<exists>as P' f b.
\<Theta> p = Some as \<and>
(as ! n) = P' \<and>
r \<subseteq> pre P' \<and>
\<Gamma> p = Some b \<and>
n < length as \<and>
\<Gamma>,\<Theta> \<turnstile>\<^bsub>/F\<^esub> P' b Q,A))"
apply (induct rule: oghoare_induct, simp_all)
apply (rename_tac \<Gamma> \<Theta> F P c Q A)
apply clarsimp
apply (case_tac P, simp_all)
apply clarsimp
apply (rule Conseq)
apply (rule_tac x="{}" in exI)
apply force
done
lemma oghoare_Parallel[rule_format, OF _ refl]:
"\<Gamma>, \<Theta>\<turnstile>\<^bsub>/F\<^esub> P x Q,A \<Longrightarrow> \<forall>cs. x = Parallel cs \<longrightarrow>
(\<exists>as. P = AnnPar as \<and>
length as = length cs \<and>
\<Inter>(set (map postcond as)) \<subseteq> Q \<and>
\<Union>(set (map abrcond as)) \<subseteq> A \<and>
(\<forall>i<length cs. \<exists>Q' A'. \<Gamma>,\<Theta>\<turnstile>\<^bsub>/F\<^esub> (pres (as!i)) (cs!i) Q', A' \<and>
Q' \<subseteq> postcond (as!i) \<and> A' \<subseteq> abrcond (as!i)) \<and>
interfree \<Gamma> \<Theta> F as cs)"
apply (induct rule: oghoare_induct, simp_all)
apply clarsimp
apply (drule_tac x=i in spec)
apply fastforce
apply clarsimp
apply (case_tac P, simp_all)
apply blast
done
lemma ann_matches_weaken[OF _ refl]:
" ann_matches \<Gamma> \<Theta> X c \<Longrightarrow> X = (weaken_pre P P') \<Longrightarrow> ann_matches \<Gamma> \<Theta> P c"
apply (induct arbitrary: P P' rule: ann_matches.induct)
apply (case_tac P, simp_all, fastforce simp add: ann_matches.intros)+
done
lemma oghoare_seq_imp_ann_matches:
" \<Gamma>,\<Theta>\<tturnstile>\<^bsub>/F\<^esub> P a c Q,A \<Longrightarrow> ann_matches \<Gamma> \<Theta> a c"
apply (induct rule: oghoare_seq_induct, simp_all add: ann_matches.intros)
apply (clarsimp, erule ann_matches_weaken)+
done
lemma oghoare_imp_ann_matches:
" \<Gamma>,\<Theta>\<turnstile>\<^bsub>/F\<^esub> a c Q,A \<Longrightarrow> ann_matches \<Gamma> \<Theta> a c"
apply (induct rule: oghoare_induct, simp_all add: ann_matches.intros oghoare_seq_imp_ann_matches)
apply (clarsimp, erule ann_matches_weaken)+
done
(* intros *)
lemma SkipRule: "P \<subseteq> Q \<Longrightarrow> \<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> (AnnExpr P) Skip Q, A"
apply (rule Conseq, simp)
apply (rule exI, rule exI, rule exI)
apply (rule conjI, rule Skip, auto)
done
lemma ThrowRule: "P \<subseteq> A \<Longrightarrow> \<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> (AnnExpr P) Throw Q, A"
apply (rule Conseq, simp)
apply (rule exI, rule exI, rule exI)
apply (rule conjI, rule Throw, auto)
done
lemma BasicRule: "P \<subseteq> {s. (f s) \<in> Q} \<Longrightarrow> \<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> (AnnExpr P) (Basic f) Q, A"
apply (rule Conseq, simp, rule exI[where x="{s. (f s) \<in> Q}"])
apply (rule exI[where x=Q], rule exI[where x=A])
apply simp
apply (subgoal_tac "(P \<union> {s. f s \<in> Q}) = {s. f s \<in> Q}")
apply (auto intro: Basic)
done
lemma SpecRule:
"P \<subseteq> {s. (\<forall>t. (s, t) \<in> r \<longrightarrow> t \<in> Q) \<and> (\<exists>t. (s, t) \<in> r)}
\<Longrightarrow> \<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> (AnnExpr P) (Spec r) Q, A"
apply (rule Conseq, simp, rule exI[where x="{s. (\<forall>t. (s, t) \<in> r \<longrightarrow> t \<in> Q) \<and> (\<exists>t. (s, t) \<in> r) }"])
apply (rule exI[where x=Q], rule exI[where x=A])
apply simp
apply (subgoal_tac "(P \<union> {s. (\<forall>t. (s, t) \<in> r \<longrightarrow> t \<in> Q) \<and> (\<exists>t. (s, t) \<in> r)}) = {s. (\<forall>t. (s, t) \<in> r \<longrightarrow> t \<in> Q) \<and> (\<exists>t. (s, t) \<in> r)}")
apply (auto intro: Spec)
done
lemma CondRule:
"\<lbrakk> P \<subseteq> {s. (s\<in>b \<longrightarrow> s\<in>pre P\<^sub>1) \<and> (s\<notin>b \<longrightarrow> s\<in>pre P\<^sub>2)};
\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q,A;
\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q,A \<rbrakk>
\<Longrightarrow> \<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> (AnnBin P P\<^sub>1 P\<^sub>2) (Cond b c\<^sub>1 c\<^sub>2) Q,A"
by (auto intro: Cond)
lemma WhileRule: "\<lbrakk> r \<subseteq> I; I \<inter> b \<subseteq> pre P; (I \<inter> -b) \<subseteq> Q;
\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P c I,A \<rbrakk>
\<Longrightarrow> \<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> (AnnWhile r I P) (While b c) Q,A"
by (simp add: While)
lemma AwaitRule:
"\<lbrakk>atom_com c ; \<Gamma>, \<Theta> \<tturnstile>\<^bsub>/F\<^esub>P a c Q,A ; (r \<inter> b) \<subseteq> P\<rbrakk> \<Longrightarrow>
\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> (AnnRec r a) (Await b c) Q,A"
apply (erule Await[rotated])
apply (erule (1) SeqConseq, simp+)
done
lemma rtranclp_1n_induct [consumes 1, case_names base step]:
"\<lbrakk>r\<^sup>*\<^sup>* a b; P a; \<And>y z. \<lbrakk>r y z; r\<^sup>*\<^sup>* z b; P y\<rbrakk> \<Longrightarrow> P z\<rbrakk> \<Longrightarrow> P b"
by (induct rule: rtranclp.induct)
(simp add: rtranclp.rtrancl_into_rtrancl)+
lemmas rtranclp_1n_induct2[consumes 1, case_names base step] =
rtranclp_1n_induct[of _ "(ax,ay)" "(bx,by)", split_rule]
lemma oghoare_seq_valid:
" \<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P c\<^sub>1 R,A \<Longrightarrow>
\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> R c\<^sub>2 Q,A \<Longrightarrow>
\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P Seq c\<^sub>1 c\<^sub>2 Q,A"
apply (clarsimp simp add: valid_def)
apply (rename_tac t c' s)
apply (case_tac t)
apply simp
apply (drule (1) Seq_decomp_star)
apply (erule disjE)
apply fastforce
apply clarsimp
apply (rename_tac s1 t')
apply (drule_tac x="Normal s" and y="Normal t'" in spec2)
apply (erule_tac x="Skip" in allE)
apply (fastforce simp: final_def)
apply (clarsimp simp add: final_def)
apply (drule Seq_decomp_star_Fault)
apply (erule disjE)
apply (rename_tac s2)
apply (drule_tac x="Normal s" and y="Fault s2" in spec2)
apply (erule_tac x="Skip" in allE)
apply fastforce
apply clarsimp
apply (rename_tac s s2 s')
apply (drule_tac x="Normal s" and y="Normal s'" in spec2)
apply (erule_tac x="Skip" in allE)
apply clarsimp
apply (drule_tac x="Normal s'" and y="Fault s2" in spec2)
apply (erule_tac x="Skip" in allE)
apply clarsimp
apply clarsimp
apply (simp add: final_def)
apply (drule Seq_decomp_star_Stuck)
apply (erule disjE)
apply fastforce
apply clarsimp
apply fastforce
done
lemma oghoare_if_valid:
"\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q,A \<Longrightarrow>
\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q,A \<Longrightarrow>
r \<inter> b \<subseteq> P\<^sub>1 \<Longrightarrow> r \<inter> - b \<subseteq> P\<^sub>2 \<Longrightarrow>
\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> r Cond b c\<^sub>1 c\<^sub>2 Q,A"
apply (simp (no_asm) add: valid_def)
apply (clarsimp)
apply (erule converse_rtranclpE)
apply (clarsimp simp: final_def)
apply (erule step_Normal_elim_cases)
apply (fastforce simp: valid_def[where c=c\<^sub>1])
apply (fastforce simp: valid_def[where c=c\<^sub>2])
done
lemma Skip_normal_steps_end:
"\<Gamma> \<turnstile> (Skip, Normal s) \<rightarrow>\<^sup>* (c, s') \<Longrightarrow> c = Skip \<and> s' = Normal s"
apply (erule converse_rtranclpE)
apply simp
apply (erule step_Normal_elim_cases)
done
lemma Throw_normal_steps_end:
"\<Gamma> \<turnstile> (Throw, Normal s) \<rightarrow>\<^sup>* (c, s') \<Longrightarrow> c = Throw \<and> s' = Normal s"
apply (erule converse_rtranclpE)
apply simp
apply (erule step_Normal_elim_cases)
done
lemma while_relpower_induct:
"\<And>t c' x .
\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P c i,A \<Longrightarrow>
i \<inter> b \<subseteq> P \<Longrightarrow>
i \<inter> - b \<subseteq> Q \<Longrightarrow>
final (c', t) \<Longrightarrow>
x \<in> i \<Longrightarrow>
t \<notin> Fault ` F \<Longrightarrow>
c' = Throw \<longrightarrow> t \<notin> Normal ` A \<Longrightarrow>
(step \<Gamma> ^^ n) (While b c, Normal x) (c', t) \<Longrightarrow> c' = Skip \<and> t \<in> Normal ` Q"
apply (induct n rule:less_induct)
apply (rename_tac n t c' x)
apply (case_tac n)
apply (clarsimp simp: final_def)
apply clarify
apply (simp only: relpowp.simps)
apply (subst (asm) relpowp_commute[symmetric])
apply clarsimp
apply (erule step_Normal_elim_cases)
apply clarsimp
apply (rename_tac t c' x v)
apply(case_tac "t")
apply clarsimp
apply(drule Seq_decomp_relpow)
apply(simp add: final_def)
apply(erule disjE, erule conjE)
apply clarify
apply(drule relpowp_imp_rtranclp)
apply (simp add: valid_def)
apply (rename_tac x n t' n1)
apply (drule_tac x="Normal x" in spec)
apply (drule_tac x="Normal t'" in spec)
apply (drule spec[where x=Throw])
apply (fastforce simp add: final_def)
apply clarsimp
apply (rename_tac c' x n t' t n1 n2)
apply (drule_tac x=n2 and y="Normal t'" in meta_spec2)
apply (drule_tac x=c' and y="t" in meta_spec2)
apply (erule meta_impE, fastforce)
apply (erule meta_impE, fastforce)
apply (erule meta_impE)
apply(drule relpowp_imp_rtranclp)
apply (simp add: valid_def)
apply (drule_tac x="Normal x" and y="Normal t" in spec2)
apply (drule spec[where x=Skip])
apply (fastforce simp add: final_def)
apply assumption
apply clarsimp
apply (rename_tac c' s n f)
apply (subgoal_tac "c' = Skip", simp)
prefer 2
apply (case_tac c'; fastforce simp: final_def)
apply (drule Seq_decomp_relpowp_Fault)
apply (erule disjE)
apply (clarsimp simp: valid_def)
apply (drule_tac x="Normal s" and y="Fault f" in spec2)
apply (drule spec[where x=Skip])
apply(drule relpowp_imp_rtranclp)
apply (fastforce simp: final_def)
apply clarsimp
apply (rename_tac t n1 n2)
apply (subgoal_tac "t \<in> i")
prefer 2
apply (fastforce dest:relpowp_imp_rtranclp simp: final_def valid_def)
apply (drule_tac x=n2 in meta_spec)
apply (drule_tac x="Fault f" in meta_spec)
apply (drule meta_spec[where x=Skip])
apply (drule_tac x=t in meta_spec)
apply (fastforce simp: final_def)
apply clarsimp
apply (rename_tac c' s n)
apply (subgoal_tac "c' = Skip", simp)
prefer 2
apply (case_tac c'; fastforce simp: final_def)
apply (drule Seq_decomp_relpowp_Stuck)
apply clarsimp
apply (erule disjE)
apply (simp add:valid_def)
apply (drule_tac x="Normal s" and y="Stuck" in spec2)
apply clarsimp
apply (drule spec[where x=Skip])
apply(drule relpowp_imp_rtranclp)
apply (fastforce)
apply clarsimp
apply (rename_tac t n1 n2)
apply (subgoal_tac "t \<in> i")
prefer 2
apply (fastforce dest:relpowp_imp_rtranclp simp: final_def valid_def)
apply (drule_tac x=n2 in meta_spec)
apply (drule meta_spec[where x=Stuck])
apply (drule meta_spec[where x=Skip])
apply (drule_tac x=t in meta_spec)
apply (fastforce simp: final_def)
apply clarsimp
apply (drule relpowp_imp_rtranclp)
apply (drule Skip_normal_steps_end)
apply fastforce
done
lemma valid_weaken_pre:
"\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P c Q,A \<Longrightarrow>
P' \<subseteq> P \<Longrightarrow> \<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P' c Q,A"
by (fastforce simp: valid_def)
lemma valid_strengthen_post:
"\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P c Q,A \<Longrightarrow>
Q \<subseteq> Q' \<Longrightarrow> \<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P c Q',A"
by (fastforce simp: valid_def)
lemma valid_strengthen_abr:
"\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P c Q,A \<Longrightarrow>
A \<subseteq> A' \<Longrightarrow> \<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P c Q,A'"
by (fastforce simp: valid_def)
lemma oghoare_while_valid:
"\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P c i,A \<Longrightarrow>
i \<inter> b \<subseteq> P \<Longrightarrow>
i \<inter> - b \<subseteq> Q \<Longrightarrow>
\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> i While b c Q,A"
apply (simp (no_asm) add: valid_def)
apply (clarsimp simp add: )
apply (drule rtranclp_imp_relpowp)
apply (clarsimp)
by (erule while_relpower_induct)
lemma oghoare_catch_valid:
"\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q,P\<^sub>2 \<Longrightarrow>
\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q,A \<Longrightarrow>
\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> P\<^sub>1 Catch c\<^sub>1 c\<^sub>2 Q,A"
apply (clarsimp simp add: valid_def[where c="Catch _ _"])
apply (rename_tac t c' s)
apply (case_tac t)
apply simp
apply (drule Catch_decomp_star)
apply (fastforce simp: final_def)
apply clarsimp
apply (erule disjE)
apply (clarsimp simp add: valid_def[where c="c\<^sub>1"])
apply (rename_tac s x t)
apply (drule_tac x="Normal s" in spec)
apply (drule_tac x="Normal t" in spec)
apply (drule_tac x=Throw in spec)
apply (fastforce simp: final_def valid_def)
apply (clarsimp simp add: valid_def[where c="c\<^sub>1"])
apply (rename_tac s t)
apply (drule_tac x="Normal s" in spec)
apply (drule_tac x="Normal t" in spec)
apply (drule_tac x=Skip in spec)
apply (fastforce simp: final_def)
apply (rename_tac c' s t)
apply (simp add: final_def)
apply (drule Catch_decomp_star_Fault)
apply clarsimp
apply (erule disjE)
apply (clarsimp simp: valid_def[where c=c\<^sub>1] final_def)
apply (fastforce simp: valid_def final_def)
apply (simp add: final_def)
apply (drule Catch_decomp_star_Stuck)
apply clarsimp
apply (erule disjE)
apply (clarsimp simp: valid_def[where c=c\<^sub>1] final_def)
apply (fastforce simp: valid_def final_def)
done
lemma ann_matches_imp_assertionsR:
"ann_matches \<Gamma> \<Theta> a c \<Longrightarrow> \<not> pre_par a \<Longrightarrow>
assertionsR \<Gamma> \<Theta> Q A a c (pre a)"
by (induct arbitrary: Q A rule: ann_matches.induct , (fastforce intro: assertionsR.intros )+)
lemma ann_matches_imp_assertionsR':
"ann_matches \<Gamma> \<Theta> a c \<Longrightarrow> a' \<in> pre_set a \<Longrightarrow>
assertionsR \<Gamma> \<Theta> Q A a c a'"
apply (induct arbitrary: Q A rule: ann_matches.induct)
apply ((fastforce intro: assertionsR.intros )+)[12]
apply simp
apply (erule bexE)
apply (simp only: in_set_conv_nth)
apply (erule exE)
apply (drule_tac x=i in spec)
apply clarsimp
apply (erule AsParallelExprs)
apply simp
done
lemma rtranclp_conjD:
"(\<lambda>x1 x2. r1 x1 x2 \<and> r2 x1 x2)\<^sup>*\<^sup>* x1 x2 \<Longrightarrow>
r1\<^sup>*\<^sup>* x1 x2 \<and> r2\<^sup>*\<^sup>* x1 x2"
by (metis (no_types, lifting) rtrancl_mono_proof)
lemma rtranclp_mono' :
"r\<^sup>*\<^sup>* a b \<Longrightarrow> r \<le> s \<Longrightarrow> s\<^sup>*\<^sup>* a b"
by (metis rtrancl_mono_proof sup.orderE sup2CI)
lemma state_upd_in_atomicsR[rule_format, OF _ refl refl]:
"\<Gamma>\<turnstile> cf \<rightarrow> cf' \<Longrightarrow>
cf = (c, Normal s) \<Longrightarrow>
cf' = (c', Normal t) \<Longrightarrow>
s \<noteq> t \<Longrightarrow>
ann_matches \<Gamma> \<Theta> a c \<Longrightarrow>
s \<in> pre a \<Longrightarrow>
(\<exists>p cm x. atomicsR \<Gamma> \<Theta> a c (p, cm) \<and> s \<in> p \<and>
\<Gamma> \<turnstile> (cm, Normal s) \<rightarrow> (x, Normal t) \<and> final (x, Normal t))"
+supply [[simproc del: defined_all]]
apply (induct arbitrary: c c' s t a rule: step.induct, simp_all)
apply clarsimp
apply (erule ann_matches.cases, simp_all)
apply (rule exI)+
apply (rule conjI)
apply (rule atomicsR.intros)
apply clarsimp
apply (rule_tac x="Skip" in exI)
apply (simp add: final_def)
apply (rule step.Basic)
apply clarsimp
apply (erule ann_matches.cases, simp_all)
apply (rule exI)+
apply (rule conjI)
apply (rule atomicsR.intros)
apply clarsimp
apply (rule_tac x="Skip" in exI)
apply (simp add: final_def)
apply (erule step.Spec)
apply clarsimp
apply (erule ann_matches.cases, simp_all)
apply clarsimp
apply (drule meta_spec)+
apply (erule meta_impE, rule conjI, (rule refl)+)+
apply clarsimp
apply (erule (1) meta_impE)
apply (erule meta_impE, fastforce)
apply clarsimp
apply (rule exI)+
apply (rule conjI)
apply (erule AtSeqExpr1)
apply fastforce
apply clarsimp
apply (erule ann_matches.cases, simp_all)
apply clarsimp
apply (drule meta_spec)+
apply (erule meta_impE, rule conjI, (rule refl)+)+
apply clarsimp
apply (erule (1) meta_impE)
apply (erule meta_impE, fastforce)
apply clarsimp
apply (rule exI)+
apply (rule conjI)
apply (erule AtCatchExpr1)
apply fastforce
apply (erule ann_matches.cases, simp_all)
apply clarsimp
apply (drule meta_spec)+
apply (erule meta_impE, rule conjI, (rule refl)+)+
apply clarsimp
apply (erule meta_impE)
apply fastforce
apply (erule meta_impE)
apply (case_tac "i=0"; fastforce)
apply clarsimp
apply (rule exI)+
apply (rule conjI)
apply (erule AtParallelExprs)
apply fastforce
apply (drule_tac x=i in spec)
apply clarsimp
apply fastforce
apply (erule ann_matches.cases, simp_all)
apply clarsimp
apply (rule exI)+
apply (rule conjI)
apply (rule AtAwait)
apply clarsimp
apply (rename_tac c' sa t aa e r ba)
apply (rule_tac x=c' in exI)
apply (rule conjI)
apply (erule step.Await)
apply (erule rtranclp_mono')
apply clarsimp
apply assumption+
apply (simp add: final_def)
done
lemma oghoare_atom_com_sound:
"\<Gamma>, \<Theta> \<tturnstile>\<^bsub>/F\<^esub>P a c Q,A \<Longrightarrow> atom_com c \<Longrightarrow> \<Gamma> \<Turnstile>\<^bsub>/F\<^esub> P c Q, A"
unfolding valid_def
proof (induct rule: oghoare_seq_induct)
case SeqSkip thus ?case
by (fastforce
elim: converse_rtranclpE step_Normal_elim_cases(1))
next
case SeqThrow thus ?case
by (fastforce
elim: converse_rtranclpE step_Normal_elim_cases)
next
case SeqBasic thus ?case
by (fastforce
elim: converse_rtranclpE step_Normal_elim_cases
simp: final_def)
next
case (SeqSpec \<Gamma> \<Theta> F r Q A) thus ?case
apply clarsimp
apply (erule converse_rtranclpE)
apply (clarsimp simp: final_def)
apply (erule step_Normal_elim_cases)
apply (fastforce elim!: converse_rtranclpE step_Normal_elim_cases)
by clarsimp
next
case (SeqSeq \<Gamma> \<Theta> F P\<^sub>1 c\<^sub>1 P\<^sub>2 A c\<^sub>2 Q) show ?case
using SeqSeq
by (fold valid_def)
(fastforce intro: oghoare_seq_valid simp: valid_weaken_pre)
next
case (SeqCatch \<Gamma> \<Theta> F P\<^sub>1 c\<^sub>1 Q P\<^sub>2 c\<^sub>2 A) thus ?case
apply (fold valid_def)
apply simp
apply (fastforce elim: oghoare_catch_valid)+
done
next
case (SeqCond \<Gamma> \<Theta> F P b c1 Q A c2) thus ?case
by (fold valid_def)
(fastforce intro:oghoare_if_valid)
next
case (SeqWhile \<Gamma> \<Theta> F P c A b) thus ?case
by (fold valid_def)
(fastforce elim: valid_weaken_pre[rotated] oghoare_while_valid)
next
case (SeqGuard \<Gamma> \<Theta> F P c Q A r g f) thus ?case
apply (fold valid_def)
apply (simp (no_asm) add: valid_def)
apply clarsimp
apply (erule converse_rtranclpE)
apply (fastforce simp: final_def)
apply clarsimp
apply (erule step_Normal_elim_cases)
apply (case_tac "r \<inter> - g \<noteq> {}")
apply clarsimp
apply (fastforce simp: valid_def)
apply clarsimp
apply (fastforce simp: valid_def)
apply clarsimp
apply (case_tac "r \<inter> - g \<noteq> {}")
apply (fastforce dest: no_steps_final simp:final_def)
apply (fastforce dest: no_steps_final simp:final_def)
done
next
case (SeqCall \<Gamma> p f \<Theta> F P Q A) thus ?case
by simp
next
case (SeqDynCom r fa \<Gamma> \<Theta> F P c Q A) thus ?case
apply -
apply clarsimp
apply (erule converse_rtranclpE)
apply (clarsimp simp: final_def)
apply clarsimp
apply (erule step_Normal_elim_cases)
apply clarsimp
apply (rename_tac t c' x)
apply (drule_tac x=x in spec)
apply (drule_tac x=x in bspec, fastforce)
apply clarsimp
apply (drule_tac x="Normal x" in spec)
apply (drule_tac x="t" in spec)
apply (drule_tac x="c'" in spec)
apply fastforce+
done
next
case (SeqConseq \<Gamma> \<Theta> F P c Q A) thus ?case
apply (clarsimp)
apply (rename_tac t c' x)
apply (erule_tac x="Normal x" in allE)
apply (erule_tac x="t" in allE)
apply (erule_tac x="c'" in allE)
apply (clarsimp simp: pre_weaken_pre)
apply force
done
qed simp_all
lemma ParallelRuleAnn:
" length as = length cs \<Longrightarrow>
\<forall>i<length cs. \<Gamma>,\<Theta> \<turnstile>\<^bsub>/F \<^esub>(pres (as ! i)) (cs ! i) (postcond (as ! i)),(abrcond (as ! i)) \<Longrightarrow>
interfree \<Gamma> \<Theta> F as cs \<Longrightarrow>
\<Inter>(set (map postcond as)) \<subseteq> Q \<Longrightarrow>
\<Union>(set (map abrcond as)) \<subseteq> A \<Longrightarrow> \<Gamma>,\<Theta> \<turnstile>\<^bsub>/F \<^esub>(AnnPar as) (Parallel cs) Q,A"
apply (erule (3) Parallel)
apply auto
done
lemma oghoare_step[rule_format, OF _ refl refl]:
shows
"\<Gamma> \<turnstile> cf \<rightarrow> cf' \<Longrightarrow> cf = (c, Normal s) \<Longrightarrow> cf' = (c', Normal t) \<Longrightarrow>
\<Gamma>,\<Theta>\<turnstile>\<^bsub>/F \<^esub>a c Q,A \<Longrightarrow>
s \<in> pre a \<Longrightarrow>
\<exists>a'. \<Gamma>,\<Theta>\<turnstile>\<^bsub>/F \<^esub>a' c' Q,A \<and> t \<in> pre a' \<and>
(\<forall>as. assertionsR \<Gamma> \<Theta> Q A a' c' as \<longrightarrow> assertionsR \<Gamma> \<Theta> Q A a c as) \<and>
(\<forall>pm cm. atomicsR \<Gamma> \<Theta> a' c' (pm, cm) \<longrightarrow> atomicsR \<Gamma> \<Theta> a c (pm, cm))"
proof (induct arbitrary:c c' s a t Q A rule: step.induct)
- case (Parallel i cs s c' s' ca c'a sa a t Q A) thus ?case
+ case (Parallel i cs s c' s' ca c'a sa a t Q A) thus ?case
+ supply [[simproc del: defined_all]]
apply (clarsimp simp:)
apply (drule oghoare_Parallel)
apply clarsimp
apply (rename_tac as)
apply (frule_tac x=i in spec, erule (1) impE)
apply (elim exE conjE)
apply (drule meta_spec)+
apply (erule meta_impE, rule conjI, (rule refl)+)+
apply (erule meta_impE)
apply (rule_tac P="(pres (as ! i))" in Conseq)
apply (rule exI[where x="{}"])
apply (rule_tac x="Q'" in exI)
apply (rule_tac x="A'" in exI)
apply (simp)
apply (erule meta_impE, simp)
apply clarsimp
apply (rule_tac x="AnnPar (as[i:=(a',postcond(as!i), abrcond(as!i))])" in exI)
apply (rule conjI)
apply (rule ParallelRuleAnn, simp)
apply clarsimp
apply (rename_tac j)
apply (drule_tac x=j in spec)
apply clarsimp
apply (case_tac "i = j")
apply (clarsimp simp: )
apply (rule Conseq)
apply (rule exI[where x="{}"])
apply (fastforce)
apply (simp add: )
apply (clarsimp simp: interfree_def)
apply (rename_tac i' j')
apply (drule_tac x=i' and y=j' in spec2)
apply clarsimp
apply (case_tac "i = i'")
apply clarsimp
apply (simp add: interfree_aux_def prod.case_eq_if )
apply clarsimp
apply (case_tac "j' = i")
apply (clarsimp)
apply (simp add: interfree_aux_def prod.case_eq_if)
apply clarsimp
apply (clarsimp)
apply (erule subsetD)
apply (clarsimp simp: in_set_conv_nth)
apply (rename_tac a' x a b c i')
apply (case_tac "i' = i")
apply clarsimp
apply (drule_tac x="(a', b, c)" in bspec, simp)
apply (fastforce simp add: in_set_conv_nth)
apply fastforce
apply (drule_tac x="(a, b, c)" in bspec, simp)
apply (simp add: in_set_conv_nth)
apply (rule_tac x=i' in exI)
apply clarsimp
apply fastforce
apply clarsimp
apply (erule_tac A="(\<Union>x\<in>set as. abrcond x) " in subsetD)
apply (clarsimp simp: in_set_conv_nth)
apply (rename_tac a b c j)
apply (case_tac "j = i")
apply clarsimp
apply (rule_tac x="as!i" in bexI)
apply simp
apply clarsimp
apply clarsimp
apply (rule_tac x="(a,b,c)" in bexI)
apply simp
apply (clarsimp simp:in_set_conv_nth)
apply (rule_tac x=j in exI)
apply fastforce
apply (rule conjI)
apply (case_tac "s = Normal t")
apply clarsimp
apply (clarsimp simp: in_set_conv_nth)
apply (rename_tac a b c j)
apply (case_tac "j = i")
apply clarsimp
apply clarsimp
apply (drule_tac x="as!j" in bspec)
apply (clarsimp simp add: in_set_conv_nth)
apply (rule_tac x=j in exI)
apply fastforce
apply clarsimp
apply (frule state_upd_in_atomicsR, simp)
apply (erule oghoare_imp_ann_matches)
apply (clarsimp simp: in_set_conv_nth)
apply fastforce
apply (clarsimp simp: in_set_conv_nth)
apply (rename_tac j)
apply (case_tac "j = i")
apply clarsimp
apply clarsimp
apply (thin_tac "\<Gamma>,\<Theta> \<turnstile>\<^bsub>/F \<^esub>a' c' (postcond (as ! i)),(abrcond (as ! i))")
apply (simp add: interfree_def interfree_aux_def)
apply (drule_tac x="j" and y=i in spec2)
apply (simp add: prod.case_eq_if)
apply (drule spec2, drule (1) mp)
apply clarsimp
apply (case_tac "pre_par a")
apply (subst pre_set)
apply clarsimp
apply (drule_tac x="as!j" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x=j in exI)
apply fastforce
apply clarsimp
apply (frule (1) pre_imp_pre_set)
apply (rename_tac as Q' A' a' a b c p cm x j X)
apply (drule_tac x="X" in spec, erule_tac P="assertionsR \<Gamma> \<Theta> b c a (cs ! j) X" in impE)
apply (rule ann_matches_imp_assertionsR')
apply (drule_tac x=j in spec, clarsimp)
apply (erule (1) oghoare_imp_ann_matches)
apply (rename_tac a b c p cm x j X )
apply (thin_tac "\<Gamma>\<Turnstile>\<^bsub>/F\<^esub> (b \<inter> p) cm b,b")
apply (thin_tac " \<Gamma>\<Turnstile>\<^bsub>/F\<^esub> (c \<inter> p) cm c,c")
apply (simp add: valid_def)
apply (drule_tac x="Normal sa" in spec)
apply (drule_tac x="Normal t" in spec)
apply (drule_tac x=x in spec)
apply (erule impE, fastforce)
apply force
apply (drule_tac x=j in spec)
apply clarsimp
apply (rename_tac a b c p cm x j Q'' A'')
apply (drule_tac x="pre a" in spec,erule impE, rule ann_matches_imp_assertionsR)
apply (erule (1) oghoare_imp_ann_matches)
apply (thin_tac " \<Gamma>\<Turnstile>\<^bsub>/F\<^esub> (b \<inter> p) cm b,b")
apply (thin_tac " \<Gamma>\<Turnstile>\<^bsub>/F\<^esub> (c \<inter> p) cm c,c")
apply (simp add: valid_def)
apply (drule_tac x="Normal sa" in spec)
apply (drule_tac x="Normal t" in spec)
apply (drule_tac x=x in spec)
apply (erule impE, fastforce)
apply clarsimp
apply (drule_tac x="as ! j" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x=j in exI, fastforce)
apply clarsimp
apply fastforce
apply (rule conjI)
apply (clarsimp simp: )
apply (erule assertionsR.cases ; simp)
apply (clarsimp simp: )
apply (rename_tac j a)
apply (case_tac "j = i")
apply clarsimp
apply (drule_tac x=a in spec, erule (1) impE)
apply (erule (1) AsParallelExprs)
apply (subst (asm) nth_list_update_neq, simp)
apply (erule_tac i=j in AsParallelExprs)
apply fastforce
apply clarsimp
apply (rule AsParallelSkips)
apply (clarsimp simp:)
apply (rule equalityI)
apply (clarsimp simp: in_set_conv_nth)
apply (rename_tac a' x a b c j)
apply (case_tac "j = i")
apply (thin_tac "\<forall>a\<in>set as. sa \<in> precond a")
apply clarsimp
apply (drule_tac x="(a', b, c)" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x="i" in exI)
apply fastforce
apply fastforce
apply (drule_tac x="as ! j" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x=j in exI)
apply fastforce
apply clarsimp
apply (drule_tac x=" as ! j" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x=j in exI, fastforce)
apply fastforce
apply (clarsimp simp: in_set_conv_nth)
apply (rename_tac x a b c j)
apply (thin_tac "\<forall>a\<in>set as. sa \<in> precond a")
apply (case_tac "j = i")
apply clarsimp
apply (drule_tac x="as!i" in bspec, fastforce)
apply fastforce
apply clarsimp
apply (drule_tac x="as!j" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x=j in exI, fastforce)
apply fastforce
apply clarsimp
apply (erule atomicsR.cases ; simp)
apply clarsimp
apply (rename_tac j atc atp)
apply (case_tac "j = i")
apply clarsimp
apply (drule_tac x=atc and y=atp in spec2, erule impE)
apply (clarsimp)
apply (erule (1) AtParallelExprs)
apply (subst (asm) nth_list_update_neq, simp)
apply (erule_tac i=j in AtParallelExprs)
apply (clarsimp)
done
next
case (Basic f s c c' sa a t Q A) thus ?case
apply clarsimp
apply (drule oghoare_Basic)
apply clarsimp
apply (rule_tac x="AnnExpr Q" in exI)
apply clarsimp
apply (rule conjI)
apply (rule SkipRule)
apply fastforce
apply (rule conjI)
apply fastforce
apply clarsimp
apply (drule assertionsR.cases, simp_all)
apply (rule assertionsR.AsBasicSkip)
done
next
case (Spec s t r c c' sa a ta Q A) thus ?case
apply clarsimp
apply (drule oghoare_Spec)
apply clarsimp
apply (rule_tac x="AnnExpr Q" in exI)
apply clarsimp
apply (rule conjI)
apply (rule SkipRule)
apply fastforce
apply (rule conjI)
apply force
apply clarsimp
apply (erule assertionsR.cases, simp_all)
apply clarsimp
apply (rule assertionsR.AsSpecSkip)
done
next
case (Guard s g f c ca c' sa a t Q A) thus ?case
apply -
apply clarsimp
apply (drule oghoare_Guard)
apply clarsimp
apply (rule exI, rule conjI, assumption)
by (fastforce dest: oghoare_Guard
intro:assertionsR.intros atomicsR.intros)
next
case (GuardFault s g f c ca c' sa a t Q A) thus ?case
by (fastforce dest: oghoare_Guard
intro:assertionsR.intros atomicsR.intros)
next
case (Seq c\<^sub>1 s c\<^sub>1' s' c\<^sub>2 c c' sa a t A Q) thus ?case
+ supply [[simproc del: defined_all]]
apply (clarsimp simp:)
apply (drule oghoare_Seq)
apply clarsimp
apply (drule meta_spec)+
apply (erule meta_impE, rule conjI, (rule refl)+)+
apply (erule meta_impE)
apply (rule Conseq)
apply (rule exI[where x="{}"])
apply (rule exI)+
apply (rule conjI)
apply (simp)
apply (erule (1) conjI)
apply clarsimp
apply (rule_tac x="AnnComp a' P\<^sub>2" in exI, rule conjI)
apply (rule oghoare_oghoare_seq.Seq)
apply (rule Conseq)
apply (rule_tac x="{}" in exI)
apply (fastforce)
apply (rule Conseq)
apply (rule_tac x="{}" in exI)
apply (fastforce)
apply clarsimp
apply (rule conjI)
apply clarsimp
apply (erule assertionsR.cases, simp_all)
apply clarsimp
apply (drule_tac x=a in spec, simp)
apply (erule AsSeqExpr1)
apply clarsimp
apply (erule AsSeqExpr2)
apply clarsimp
apply (erule atomicsR.cases, simp_all)
apply clarsimp
apply (drule_tac x="a" and y=b in spec2, simp)
apply (erule AtSeqExpr1)
apply clarsimp
apply (erule AtSeqExpr2)
done
next
case (SeqSkip c\<^sub>2 s c c' sa a t Q A) thus ?case
apply clarsimp
apply (drule oghoare_Seq)
apply clarsimp
apply (rename_tac P\<^sub>1 P\<^sub>2 P' Q' A')
apply (rule_tac x=P\<^sub>2 in exI)
apply (rule conjI, rule Conseq)
apply (rule_tac x="{}" in exI)
apply (fastforce)
apply (rule conjI)
apply (drule oghoare_Skip)
apply fastforce
apply (rule conjI)
apply clarsimp
apply (erule assertionsR.AsSeqExpr2)
apply clarsimp
apply (fastforce intro: atomicsR.intros)
done
next
case (SeqThrow c\<^sub>2 s c c' sa a t Q A) thus ?case
apply clarsimp
apply (drule oghoare_Seq)
apply clarsimp
apply (rename_tac P\<^sub>1 P\<^sub>2 P' Q' A')
apply (rule_tac x=P\<^sub>1 in exI)
apply (drule oghoare_Throw)
apply clarsimp
apply (rename_tac P'')
apply (rule conjI, rule Conseq)
apply (rule_tac x="{}" in exI)
apply (rule_tac x="Q'" in exI)
apply (rule_tac x="P''" in exI)
apply (fastforce intro: Throw)
apply clarsimp
apply (erule assertionsR.cases, simp_all)
apply clarsimp
apply (rule AsSeqExpr1)
apply (rule AsThrow)
done
next
case (CondTrue s b c\<^sub>1 c\<^sub>2 c sa c' s' ann) thus ?case
apply (clarsimp)
apply (drule oghoare_Cond)
apply clarsimp
apply (rename_tac P' P\<^sub>1 P\<^sub>2 Q' A')
apply (rule_tac x= P\<^sub>1 in exI)
apply (rule conjI)
apply (rule Conseq, rule_tac x="{}" in exI, fastforce)
apply (rule conjI, fastforce)
apply (rule conjI)
apply (fastforce elim: assertionsR.cases intro: AsCondExpr1)
apply (fastforce elim: atomicsR.cases intro: AtCondExpr1)
done
next
case (CondFalse s b c\<^sub>1 c\<^sub>2 c sa c' s' ann) thus ?case
apply (clarsimp)
apply (drule oghoare_Cond)
apply clarsimp
apply (rename_tac P' P\<^sub>1 P\<^sub>2 Q' A')
apply (rule_tac x= P\<^sub>2 in exI)
apply (rule conjI)
apply (rule Conseq, rule_tac x="{}" in exI, fastforce)
apply (rule conjI, fastforce)
apply (rule conjI)
apply (fastforce elim: assertionsR.cases intro: AsCondExpr2)
apply (fastforce elim: atomicsR.cases intro: AtCondExpr2)
done
next
case (WhileTrue s b c ca sa c' s' ann) thus ?case
apply clarsimp
apply (frule oghoare_While)
apply clarsimp
apply (rename_tac r i P' A' Q')
apply (rule_tac x="AnnComp P' (AnnWhile i i P')" in exI)
apply (rule conjI)
apply (rule Seq)
apply (rule Conseq)
apply (rule_tac x="{}" in exI)
apply (rule_tac x="i" in exI)
apply (rule_tac x="A'" in exI)
apply (subst weaken_pre_empty)
apply clarsimp
apply (rule While)
apply (rule Conseq)
apply (rule_tac x="{}" in exI)
apply (rule_tac x="i" in exI)
apply (rule_tac x="A'" in exI)
apply (subst weaken_pre_empty)
apply clarsimp
apply clarsimp
apply force
apply simp
apply simp
apply (rule conjI)
apply blast
apply (rule conjI)
apply clarsimp
apply (erule assertionsR.cases, simp_all)
apply clarsimp
apply (rule AsWhileExpr)
apply clarsimp
apply (erule assertionsR.cases,simp_all)
apply clarsimp
apply (erule AsWhileExpr)
apply clarsimp
apply (rule AsWhileInv)
apply clarsimp
apply (rule AsWhileInv)
apply clarsimp
apply (rule AsWhileSkip)
apply clarsimp
apply (erule atomicsR.cases, simp_all)
apply clarsimp
apply (rule AtWhileExpr)
apply clarsimp+
apply (erule atomicsR.cases, simp_all)
apply clarsimp
apply (erule AtWhileExpr)
done
next
case (WhileFalse s b c ca sa c' ann s' Q A) thus ?case
apply clarsimp
apply (drule oghoare_While)
apply clarsimp
apply (rule_tac x="AnnExpr Q" in exI)
apply (rule conjI)
apply (rule SkipRule)
apply blast
apply (rule conjI)
apply fastforce
apply clarsimp
apply (erule assertionsR.cases, simp_all)
apply (drule sym, simp, rule AsWhileSkip)
done
next
case (Call p bs s c c' sa a t Q A) thus ?case
apply clarsimp
apply (drule oghoare_Call)
apply clarsimp
apply (rename_tac n as)
apply (rule_tac x="as ! n" in exI)
apply clarsimp
apply (rule conjI, fastforce)
apply (rule conjI)
apply clarsimp
apply (erule (2) AsCallExpr)
apply fastforce
apply clarsimp
apply (erule (2) AtCallExpr)
apply simp
done
next
case (DynCom c s ca c' sa a t Q A) thus ?case
apply -
apply clarsimp
apply (drule oghoare_DynCom)
apply clarsimp
apply (drule_tac x=t in bspec, assumption)
apply (rule exI)
apply (erule conjI)
apply (rule conjI, fastforce)
apply (rule conjI)
apply clarsimp
apply (erule (1) AsDynComExpr)
apply (clarsimp)
apply (erule (1) AtDynCom)
done
next
case (Catch c\<^sub>1 s c\<^sub>1' s' c\<^sub>2 c c' sa a t Q A) thus ?case
+ supply [[simproc del: defined_all]]
apply (clarsimp simp:)
apply (drule oghoare_Catch)
apply clarsimp
apply (drule meta_spec)+
apply (erule meta_impE, rule conjI, (rule refl)+)+
apply (erule meta_impE)
apply (rule Conseq)
apply (rule exI[where x="{}"])
apply (rule exI)+
apply (rule conjI)
apply (simp)
apply (erule (1) conjI)
apply clarsimp
apply (rename_tac P\<^sub>1 P\<^sub>2 P' Q' A' a')
apply (rule_tac x="AnnComp a' P\<^sub>2" in exI, rule conjI)
apply (rule oghoare_oghoare_seq.Catch)
apply (rule Conseq)
apply (rule_tac x="{}" in exI)
apply (fastforce)
apply (rule Conseq)
apply (rule_tac x="{}" in exI)
apply (fastforce)
apply clarsimp
apply (rule conjI)
apply clarsimp
apply (erule assertionsR.cases, simp_all)
apply clarsimp
apply (rename_tac a)
apply (drule_tac x=a in spec, simp)
apply (erule AsCatchExpr1)
apply clarsimp
apply (erule AsCatchExpr2)
apply clarsimp
apply (erule atomicsR.cases, simp_all)
apply clarsimp
apply (rename_tac a b a2)
apply (drule_tac x="a" and y=b in spec2, simp)
apply (erule AtCatchExpr1)
apply clarsimp
apply (erule AtCatchExpr2)
done
next
case (CatchSkip c\<^sub>2 s c c' sa a t Q A) thus ?case
apply clarsimp
apply (drule oghoare_Catch, clarsimp)
apply (rename_tac P\<^sub>1 P\<^sub>2 P' Q' A')
apply (rule_tac x=P\<^sub>1 in exI)
apply clarsimp
apply (rule conjI)
apply (rule Conseq)
apply (rule_tac x="{}" in exI)
apply (drule oghoare_Skip)
apply clarsimp
apply (rule_tac x=Q' in exI)
apply (rule_tac x=A' in exI)
apply (rule conjI, erule SkipRule)
apply clarsimp
apply clarsimp
apply (rule AsCatchExpr1)
apply (erule assertionsR.cases, simp_all)
apply (rule assertionsR.AsSkip)
done
next
case (CatchThrow c\<^sub>2 s c c' sa a t Q A) thus ?case
apply clarsimp
apply (drule oghoare_Catch, clarsimp)
apply (rename_tac P\<^sub>1 P\<^sub>2 P' Q' A')
apply (rule_tac x=P\<^sub>2 in exI)
apply (rule conjI)
apply (rule Conseq)
apply (rule_tac x="{}" in exI)
apply (fastforce )
apply (rule conjI)
apply (drule oghoare_Throw)
apply clarsimp
apply fastforce
apply (rule conjI)
apply (clarsimp)
apply (erule AsCatchExpr2)
apply clarsimp
apply (erule AtCatchExpr2)
done
next
case (ParSkip cs s c c' sa a t Q A) thus ?case
apply clarsimp
apply (drule oghoare_Parallel)
apply clarsimp
apply (rename_tac as)
apply (rule_tac x="AnnExpr (\<Inter>x\<in>set as. postcond x)" in exI)
apply (rule conjI, rule SkipRule)
apply blast
apply (rule conjI)
apply simp
apply clarsimp
apply (simp only: in_set_conv_nth)
apply clarsimp
apply (drule_tac x="i" in spec)
apply clarsimp
apply (drule_tac x="cs!i" in bspec)
apply clarsimp
apply clarsimp
apply (drule oghoare_Skip)
apply clarsimp
apply (drule_tac x="as!i" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x=i in exI, fastforce)
apply clarsimp
apply blast
apply clarsimp
apply (erule assertionsR.cases; simp)
apply clarsimp
apply (rule AsParallelSkips; clarsimp)
done
next
case (ParThrow cs s c c' sa a t Q A) thus ?case
apply clarsimp
apply (drule oghoare_Parallel)
apply (clarsimp simp: in_set_conv_nth)
apply (drule_tac x=i in spec)
apply clarsimp
apply (drule oghoare_Throw)
apply clarsimp
apply (rename_tac i as Q' A' P')
apply (rule_tac x="AnnExpr P'" in exI)
apply (rule conjI)
apply (rule ThrowRule)
apply clarsimp
apply (erule_tac A="(\<Union>x\<in>set as. abrcond x)" in subsetD[where B=A], force)
apply (rule conjI)
apply (drule_tac x="as!i" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x=i in exI, fastforce)
apply (fastforce)
apply clarsimp
apply (erule AsParallelExprs)
apply clarsimp
apply (erule assertionsR.cases, simp_all)
apply (rule AsThrow)
done
next
case (Await x b c c' x' c'' c''' x'' a x''' Q A) thus ?case
apply (clarsimp)
apply (drule oghoare_Await)
apply clarsimp
apply (drule rtranclp_conjD)
apply clarsimp
apply (erule disjE)
apply clarsimp
apply (rename_tac P' Q' A')
apply (rule_tac x="AnnExpr Q" in exI)
apply clarsimp
apply (rule conjI)
apply (rule Skip)
apply (rule conjI)
apply (drule (1) oghoare_atom_com_sound)
apply (fastforce simp: final_def valid_def)
apply clarsimp
apply (erule assertionsR.cases, simp_all)
apply clarsimp
apply (rule AsAwaitSkip)
apply (rule_tac x="AnnExpr A" in exI)
apply clarsimp
apply (rule conjI)
apply (rule Throw)
apply (rule conjI)
apply (drule (1) oghoare_atom_com_sound)
apply (fastforce simp: final_def valid_def)
apply clarsimp
apply (erule assertionsR.cases, simp_all)
apply clarsimp
apply (rule AsAwaitThrow)
done
qed simp_all
lemma oghoare_steps[rule_format, OF _ refl refl]:
"\<Gamma> \<turnstile> cf \<rightarrow>\<^sup>* cf' \<Longrightarrow> cf = (c, Normal s) \<Longrightarrow> cf' = (c', Normal t) \<Longrightarrow>
\<Gamma>,\<Theta>\<turnstile>\<^bsub>/F \<^esub>a c Q,A \<Longrightarrow>
s \<in> pre a \<Longrightarrow>
\<exists>a'. \<Gamma>,\<Theta>\<turnstile>\<^bsub>/F \<^esub>a' c' Q,A \<and> t \<in> pre a' \<and>
(\<forall>as. assertionsR \<Gamma> \<Theta> Q A a' c' as \<longrightarrow> assertionsR \<Gamma> \<Theta> Q A a c as) \<and>
(\<forall>pm cm. atomicsR \<Gamma> \<Theta> a' c' (pm, cm) \<longrightarrow> atomicsR \<Gamma> \<Theta> a c (pm, cm))"
apply (induct arbitrary: a c s c' t rule: converse_rtranclp_induct)
+ supply [[simproc del: defined_all]]
apply fastforce
apply clarsimp
apply (frule Normal_pre_star)
apply clarsimp
apply (drule (2) oghoare_step)
apply clarsimp
apply ((drule meta_spec)+, (erule meta_impE, rule conjI, (rule refl)+)+)
apply (erule (1) meta_impE)+
apply clarsimp
apply (rule exI)
apply (rule conjI, fastforce)+
apply force
done
lemma oghoare_sound_Parallel_Normal_case[rule_format, OF _ refl refl]:
"\<Gamma> \<turnstile> (c, s) \<rightarrow>\<^sup>* (c', t) \<Longrightarrow>
\<forall>P x y cs. c = Parallel cs \<longrightarrow> s = Normal x \<longrightarrow>
t = Normal y \<longrightarrow>
\<Gamma>,\<Theta>\<turnstile>\<^bsub>/F \<^esub>P c Q,A \<longrightarrow> final (c', t) \<longrightarrow>
x \<in> pre P \<longrightarrow> t \<notin> Fault ` F \<longrightarrow> (c' = Throw \<and> t \<in> Normal ` A) \<or> (c' = Skip \<and> t \<in> Normal ` Q)"
apply(erule converse_rtranclp_induct2)
apply (clarsimp simp: final_def)
apply(erule step.cases, simp_all)
\<comment> \<open>Parallel\<close>
apply clarsimp
apply (frule Normal_pre_star)
apply (drule oghoare_Parallel)
apply clarsimp
apply (rename_tac i cs c1' x y s' as)
apply (subgoal_tac "\<Gamma>\<turnstile> (Parallel cs, Normal x) \<rightarrow> (Parallel (cs[i := c1']), Normal s')")
apply (frule_tac c="Parallel cs" and
a="AnnPar as" and
Q="(\<Inter>x\<in>set as. postcond x)" and A ="(\<Union>x\<in>set as. abrcond x)"
in oghoare_step[where \<Theta>=\<Theta> and F=F])
apply(rule Parallel, simp)
apply clarsimp
apply (rule Conseq, rule exI[where x="{}"], fastforce)
apply clarsimp
apply force
apply force
apply clarsimp
apply clarsimp
apply (rename_tac a')
apply (drule_tac x=a' in spec)
apply (drule mp, rule Conseq)
apply (rule_tac x="{}" in exI)
apply (rule_tac x="(\<Inter>x\<in>set as. postcond x)" in exI)
apply (rule_tac x="(\<Union>x\<in>set as. abrcond x)" in exI)
apply (simp)
apply clarsimp
apply(erule (1) step.Parallel)
\<comment> \<open>ParSkip\<close>
apply (frule no_steps_final, simp add: final_def)
apply clarsimp
apply (drule oghoare_Parallel)
apply clarsimp
apply (rule imageI)
apply (erule subsetD)
apply clarsimp
apply (clarsimp simp: in_set_conv_nth)
apply (rename_tac i)
apply (frule_tac x="i" in spec)
apply clarsimp
apply (frule_tac x="cs!i" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x="i" in exI)
apply clarsimp
apply clarsimp
apply (drule_tac x="as ! i" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply fastforce
apply (drule oghoare_Skip)
apply fastforce
\<comment> \<open>ParThrow\<close>
apply clarsimp
apply (frule no_steps_final, simp add: final_def)
apply clarsimp
apply (drule oghoare_Parallel)
apply (clarsimp simp: in_set_conv_nth)
apply (drule_tac x=i in spec)
apply clarsimp
apply (drule oghoare_Throw)
apply clarsimp
apply (rename_tac i as Q' A' P')
apply (drule_tac x="as ! i" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x=i in exI, fastforce)
apply clarsimp
apply (rule imageI)
apply (erule_tac A="(\<Union>x\<in>set as. abrcond x)" in subsetD)
apply clarsimp
apply (rule_tac x="as!i" in bexI)
apply blast
apply clarsimp
done
lemma oghoare_step_Fault[rule_format, OF _ refl refl]:
"\<Gamma>\<turnstile> cf \<rightarrow> cf' \<Longrightarrow>
cf = (c, Normal x) \<Longrightarrow>
cf' = (c', Fault f) \<Longrightarrow>
x \<in> pre P \<Longrightarrow>
\<Gamma>,\<Theta>\<turnstile>\<^bsub>/F \<^esub>P c Q,A \<Longrightarrow> f \<in> F"
+ supply [[simproc del: defined_all]]
apply (induct arbitrary: x c c' P Q A f rule: step.induct, simp_all)
apply clarsimp
apply (drule oghoare_Guard)
apply clarsimp
apply blast
apply clarsimp
apply (drule oghoare_Seq)
apply clarsimp
apply clarsimp
apply (drule oghoare_Catch)
apply clarsimp
apply clarsimp
apply (rename_tac i cs c' x P Q A f)
apply (drule oghoare_Parallel)
apply clarsimp
apply (rename_tac i cs c' x Q A f as)
apply (drule_tac x="i" in spec)
apply clarsimp
apply (drule meta_spec)+
apply (erule meta_impE, rule conjI, (rule refl)+)+
apply (drule_tac x="as!i" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x="i" in exI, fastforce)
apply (erule (1) meta_impE)
apply (erule (2) meta_impE)
apply clarsimp
apply (drule rtranclp_conjD[THEN conjunct1])
apply (drule oghoare_Await)
apply clarsimp
apply (rename_tac b c c' x Q A f r P' Q' A')
apply (drule (1) oghoare_atom_com_sound)
apply (simp add: valid_def)
apply (drule_tac x="Normal x" in spec)
apply (drule_tac x="Fault f" in spec)
apply (drule_tac x=Skip in spec)
apply clarsimp
apply (erule impE)
apply (cut_tac f=f and c=c' in steps_Fault[where \<Gamma>=\<Gamma>])
apply fastforce
apply (fastforce simp: final_def steps_Fault)
done
lemma oghoare_step_Stuck[rule_format, OF _ refl refl]:
"\<Gamma>\<turnstile> cf \<rightarrow> cf' \<Longrightarrow>
cf = (c, Normal x) \<Longrightarrow>
cf' = (c', Stuck) \<Longrightarrow>
x \<in> pre P \<Longrightarrow>
\<Gamma>,\<Theta>\<turnstile>\<^bsub>/F \<^esub>P c Q,A \<Longrightarrow> P'"
apply (induct arbitrary: x c c' P Q A rule: step.induct, simp_all)
apply clarsimp
apply (drule oghoare_Spec)
apply force
apply clarsimp
apply (drule oghoare_Seq)
apply clarsimp
apply clarsimp
apply (drule oghoare_Call)
apply clarsimp
apply clarsimp
apply (drule oghoare_Catch)
apply clarsimp
apply clarsimp
apply (drule oghoare_Parallel)
apply clarsimp
apply (rename_tac i cs c' x Q A as)
apply (drule_tac x="i" in spec)
apply clarsimp
apply (drule meta_spec)+
- apply (erule meta_impE, rule conjI, (rule refl)+)+
apply (drule_tac x="as!i" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x="i" in exI, fastforce)
apply (erule meta_impE[OF _ refl])
apply (erule (1) meta_impE)
apply (erule (2) meta_impE)
apply clarsimp
apply (drule rtranclp_conjD[THEN conjunct1])
apply (drule oghoare_Await)
apply clarsimp
apply (rename_tac b c c' x Q A r P'' Q' A')
apply (drule (1) oghoare_atom_com_sound)
apply (simp add: valid_def)
apply (drule_tac x="Normal x" in spec)
apply (drule_tac x=Stuck in spec)
apply (drule_tac x=Skip in spec)
apply clarsimp
apply (erule impE)
apply (cut_tac c=c' in steps_Stuck[where \<Gamma>=\<Gamma>])
apply fastforce
apply (fastforce simp: final_def steps_Fault)
apply clarsimp
apply (drule oghoare_Await)
apply clarsimp
done
lemma oghoare_steps_Fault[rule_format, OF _ refl refl]:
"\<Gamma>\<turnstile> cf \<rightarrow>\<^sup>* cf' \<Longrightarrow>
cf = (c, Normal x) \<Longrightarrow>
cf' = (c', Fault f) \<Longrightarrow>
x \<in> pre P \<Longrightarrow>
\<Gamma>,\<Theta>\<turnstile>\<^bsub>/F \<^esub>P c Q,A \<Longrightarrow> f \<in> F"
apply (induct arbitrary: x c c' f rule: rtranclp_induct)
+ supply [[simproc del: defined_all]]
apply fastforce
apply clarsimp
apply (rename_tac b x c c' f)
apply (case_tac b)
apply clarsimp
apply (drule (2) oghoare_steps)
apply clarsimp
apply (drule (3) oghoare_step_Fault)
apply clarsimp
apply (drule meta_spec)+
apply (erule meta_impE, (rule conjI, (rule refl)+))+
apply simp
apply (drule step_Fault_prop ; simp)
apply simp
apply clarsimp
apply (drule step_Stuck_prop ; simp)
done
lemma oghoare_steps_Stuck[rule_format, OF _ refl refl]:
"\<Gamma>\<turnstile> cf \<rightarrow>\<^sup>* cf' \<Longrightarrow>
cf = (c, Normal x) \<Longrightarrow>
cf' = (c', Stuck) \<Longrightarrow>
x \<in> pre P \<Longrightarrow>
\<Gamma>,\<Theta>\<turnstile>\<^bsub>/F \<^esub>P c Q,A \<Longrightarrow> P'"
apply (induct arbitrary: x c c' rule: rtranclp_induct)
apply fastforce
apply clarsimp
apply (rename_tac b x c c')
apply (case_tac b)
apply clarsimp
apply (drule (2) oghoare_steps)
apply clarsimp
apply (drule (3) oghoare_step_Stuck)
apply clarsimp
apply (drule step_Fault_prop ; simp)
apply simp
done
lemma oghoare_sound_Parallel_Fault_case[rule_format, OF _ refl refl]:
"\<Gamma> \<turnstile> (c, s) \<rightarrow>\<^sup>* (c', t) \<Longrightarrow>
\<forall>P x f cs. c = Parallel cs \<longrightarrow> s = Normal x \<longrightarrow>
x \<in> pre P \<longrightarrow> t = Fault f \<longrightarrow>
\<Gamma>,\<Theta>\<turnstile>\<^bsub>/F \<^esub>P c Q,A \<longrightarrow> final (c', t) \<longrightarrow>
f \<in> F"
apply(erule converse_rtranclp_induct2)
apply (clarsimp simp: final_def)
apply clarsimp
apply (rename_tac c s P x f cs)
apply (case_tac s)
apply clarsimp
apply(erule step.cases, simp_all)
apply (clarsimp simp: final_def)
apply (drule oghoare_Parallel)
apply clarsimp
apply (rename_tac x f s' i cs c1' as)
apply (subgoal_tac "\<Gamma>\<turnstile> (Parallel cs, Normal x) \<rightarrow> (Parallel (cs[i := c1']), Normal s')")
apply (frule_tac c="Parallel cs" and a="AnnPar as" and
Q="(\<Inter>x\<in>set as. postcond x)" and A="(\<Union>x\<in>set as. abrcond x)"
in oghoare_step[where \<Theta>=\<Theta> and F=F])
apply(rule Parallel)
apply simp
apply clarsimp
apply (rule Conseq, rule exI[where x="{}"], fastforce)
apply assumption
apply clarsimp
apply clarsimp
apply simp
apply clarsimp
apply (rename_tac a')
apply (drule_tac x=a' in spec)
apply clarsimp
apply (erule notE[where P="oghoare _ _ _ _ _ _ _"])
apply (rule Conseq, rule exI[where x="{}"])
apply (clarsimp)
apply (rule_tac x="(\<Inter>x\<in>set as. postcond x)" in exI)
apply (rule_tac x="(\<Union>x\<in>set as. abrcond x)" in exI ; simp)
apply(erule (1) step.Parallel)
apply clarsimp
apply (fastforce dest: no_steps_final simp: final_def)+
apply (clarsimp simp: final_def)
apply (drule oghoare_Parallel)
apply (erule step_Normal_elim_cases, simp_all)
apply clarsimp
apply (rename_tac f cs f' i c1' as)
apply (drule_tac x="i" in spec)
apply (erule impE, fastforce)
apply clarsimp
apply (drule_tac x="as!i" in bspec)
apply (clarsimp simp: in_set_conv_nth)
apply (rule_tac x="i" in exI, fastforce)
apply (drule_tac P="pres (as ! i)" in oghoare_step_Fault[where \<Theta>=\<Theta> and F=F])
apply assumption+
apply (drule steps_Fault_prop ; simp)
apply simp
apply (drule steps_Stuck_prop ;simp)
done
lemma oghoare_soundness:
"(\<Gamma>, \<Theta> \<turnstile>\<^bsub>/F\<^esub> P c Q,A \<longrightarrow> \<Gamma> \<Turnstile>\<^bsub>/F\<^esub> (pre P) c Q, A) \<and>
(\<Gamma>, \<Theta> \<tturnstile>\<^bsub>/F\<^esub>P' P c Q,A \<longrightarrow> \<Gamma> \<Turnstile>\<^bsub>/F\<^esub> P' c Q, A)"
unfolding valid_def
proof (induct rule: oghoare_oghoare_seq.induct)
case SeqSkip thus ?case
by (fastforce
elim: converse_rtranclpE step_Normal_elim_cases(1))
next
case SeqThrow thus ?case
by (fastforce
elim: converse_rtranclpE step_Normal_elim_cases)
next
case SeqBasic thus ?case
by (fastforce
elim: converse_rtranclpE step_Normal_elim_cases
simp: final_def)
next
case (SeqSpec \<Gamma> \<Theta> F r Q A) thus ?case
apply clarsimp
apply (erule converse_rtranclpE)
apply (clarsimp simp: final_def)
apply (erule step_Normal_elim_cases)
apply (fastforce elim!: converse_rtranclpE step_Normal_elim_cases)
by clarsimp
next
case (SeqSeq \<Gamma> \<Theta> F P\<^sub>1 c\<^sub>1 P\<^sub>2 A c\<^sub>2 Q) show ?case
using SeqSeq
by (fold valid_def)
(fastforce intro: oghoare_seq_valid simp: valid_weaken_pre)
next
case (SeqCatch \<Gamma> \<Theta> F P\<^sub>1 c\<^sub>1 Q P\<^sub>2 c\<^sub>2 A) thus ?case
by (fold valid_def)
(fastforce elim: oghoare_catch_valid)+
next
case (SeqCond \<Gamma> \<Theta> F P b c1 Q A c2) thus ?case
by (fold valid_def)
(fastforce intro:oghoare_if_valid)
next
case (SeqWhile \<Gamma> \<Theta> F P c A b) thus ?case
by (fold valid_def)
(fastforce elim: valid_weaken_pre[rotated] oghoare_while_valid)
next
case (SeqGuard \<Gamma> \<Theta> F P c Q A r g f) thus ?case
apply (fold valid_def)
apply (simp (no_asm) add: valid_def)
apply clarsimp
apply (erule converse_rtranclpE)
apply (fastforce simp: final_def)
apply clarsimp
apply (erule step_Normal_elim_cases)
apply (case_tac "r \<inter> - g \<noteq> {}")
apply clarsimp
apply (fastforce simp: valid_def)
apply clarsimp
apply (fastforce simp: valid_def)
apply clarsimp
apply (case_tac "r \<inter> - g \<noteq> {}")
apply (fastforce dest: no_steps_final simp:final_def)
apply (fastforce dest: no_steps_final simp:final_def)
done
next
case (SeqCall \<Gamma> p f \<Theta> F P Q A) thus ?case
apply clarsimp
apply (erule converse_rtranclpE)
apply (clarsimp simp add: final_def)
apply (erule step_Normal_elim_cases)
apply (clarsimp simp: final_def)
apply fastforce
apply fastforce
done
next
case (SeqDynCom r P fa \<Gamma> \<Theta> F c Q A) thus ?case
apply -
apply clarsimp
apply (erule converse_rtranclpE)
apply (clarsimp simp: final_def)
apply clarsimp
apply (erule step_Normal_elim_cases)
apply clarsimp
apply (rename_tac t c' x)
apply (drule_tac x=x in bspec, fastforce)
apply clarsimp
apply (drule_tac x="Normal x" in spec)
apply (drule_tac x="t" in spec)
apply (drule_tac x="c'" in spec)
apply fastforce+
done
next
case (SeqConseq \<Gamma> \<Theta> F P c Q A) thus ?case
apply (clarsimp)
apply (rename_tac t c' x)
apply (erule_tac x="Normal x" in allE)
apply (erule_tac x="t" in allE)
apply (erule_tac x="c'" in allE)
apply (clarsimp simp: pre_weaken_pre)
apply force
done
next
case (SeqParallel as P \<Gamma> \<Theta> F cs Q A) thus ?case
by (fold valid_def)
(erule (1) valid_weaken_pre)
next
case (Call \<Theta> p as n P Q A r \<Gamma> f F) thus ?case
apply clarsimp
apply (erule converse_rtranclpE)
apply (clarsimp simp add: final_def)
apply (erule step_Normal_elim_cases)
apply (clarsimp simp: final_def)
apply (erule disjE)
apply clarsimp
apply fastforce
apply fastforce
apply fastforce
done
next
case (Await \<Gamma> \<Theta> F P c Q A r b) thus ?case
apply clarsimp
apply (erule converse_rtranclpE)
apply (clarsimp simp add: final_def)
apply (erule step_Normal_elim_cases)
apply (erule converse_rtranclpE)
apply (fastforce simp add: final_def )
apply (force dest!:no_step_final simp: final_def)
apply clarsimp
apply (rename_tac x c'')
apply (drule_tac x="Normal x" in spec)
apply (drule_tac x="Stuck" in spec)
apply (drule_tac x="Skip" in spec)
apply (clarsimp simp: final_def)
apply (erule impE[where P="rtranclp _ _ _"])
apply (cut_tac c="c''" in steps_Stuck[where \<Gamma>=\<Gamma>])
apply fastforce
apply fastforce
apply clarsimp
apply (rename_tac x c'' f)
apply (drule_tac x="Normal x" in spec)
apply (drule_tac x="Fault f" in spec)
apply (drule_tac x="Skip" in spec)
apply (erule impE[where P="rtranclp _ _ _"])
apply (cut_tac c="c''" and f=f in steps_Fault[where \<Gamma>=\<Gamma>])
apply fastforce
apply clarsimp
apply (erule converse_rtranclpE)
apply fastforce
apply (erule step_elim_cases)
apply (fastforce)
done
next
case (Parallel as cs \<Gamma> \<Theta> F Q A ) thus ?case
apply (fold valid_def)
apply (simp only:pre.simps)
apply (simp (no_asm) only: valid_def)
apply clarsimp
apply (rename_tac t c' x')
apply (case_tac t)
apply clarsimp
apply (drule oghoare_sound_Parallel_Normal_case[where \<Theta>=\<Theta> and Q=Q and A=A and F=F and P="AnnPar as", OF _ refl])
apply (rule oghoare_oghoare_seq.Parallel)
apply simp
apply clarsimp
apply assumption
apply (clarsimp)
apply clarsimp
apply (clarsimp simp: final_def)
apply (clarsimp )
apply clarsimp
apply clarsimp
apply (drule oghoare_sound_Parallel_Fault_case[where \<Theta>=\<Theta> and Q=Q and A=A and F=F and P="AnnPar as", OF _ ])
apply clarsimp
apply assumption
apply (rule oghoare_oghoare_seq.Parallel)
apply simp
apply clarsimp
apply assumption
apply clarsimp
apply clarsimp
apply (simp add: final_def)
apply (fastforce simp add: final_def)
apply (clarsimp simp: final_def)
apply (erule oghoare_steps_Stuck[where \<Theta>=\<Theta> and F=F and Q=Q and A=A and P="AnnPar as"])
apply simp
apply (rule oghoare_oghoare_seq.Parallel)
apply simp
apply simp
apply simp
apply clarsimp
apply clarsimp
done
next
case Skip thus ?case
by (fastforce
elim: converse_rtranclpE step_Normal_elim_cases(1))
next
case Basic thus ?case
by (fastforce
elim: converse_rtranclpE step_Normal_elim_cases
simp: final_def)
next
case (Spec \<Gamma> \<Theta> F r Q A) thus ?case
apply clarsimp
apply (erule converse_rtranclpE)
apply (clarsimp simp: final_def)
apply (erule step_Normal_elim_cases)
apply (fastforce elim!: converse_rtranclpE step_Normal_elim_cases)
by clarsimp
next
case (Seq \<Gamma> \<Theta> F P\<^sub>1 c\<^sub>1 P\<^sub>2 A c\<^sub>2 Q) show ?case
using Seq
by (fold valid_def)
(fastforce intro: oghoare_seq_valid simp: valid_weaken_pre)
next
case (Cond \<Gamma> \<Theta> F P\<^sub>1 c\<^sub>1 Q A P\<^sub>2 c\<^sub>2 r b) thus ?case
by (fold valid_def)
(fastforce intro:oghoare_if_valid)
next
case (While \<Gamma> \<Theta> F P c i A b Q r) thus ?case
by (fold valid_def)
(fastforce elim: valid_weaken_pre[rotated] oghoare_while_valid)
next
case Throw thus ?case
by (fastforce
elim: converse_rtranclpE step_Normal_elim_cases)
next
case (Catch \<Gamma> \<Theta> F P\<^sub>1 c\<^sub>1 Q P\<^sub>2 c\<^sub>2 A) thus ?case
apply (fold valid_def)
apply (fastforce elim: oghoare_catch_valid)+
done
next
case (Guard \<Gamma> \<Theta> F P c Q A r g f) thus ?case
apply (fold valid_def)
apply (simp)
apply (frule (1) valid_weaken_pre[rotated])
apply (simp (no_asm) add: valid_def)
apply clarsimp
apply (erule converse_rtranclpE)
apply (fastforce simp: final_def)
apply clarsimp
apply (erule step_Normal_elim_cases)
apply (case_tac "r \<inter> - g \<noteq> {}")
apply clarsimp
apply (fastforce simp: valid_def)
apply clarsimp
apply (fastforce simp: valid_def)
apply clarsimp
apply (case_tac "r \<inter> - g \<noteq> {}")
apply clarsimp
apply (fastforce dest: no_steps_final simp:final_def)
apply (clarsimp simp: ex_in_conv[symmetric])
done
next
case (DynCom r \<Gamma> \<Theta> F P c Q A) thus ?case
apply clarsimp
apply (erule converse_rtranclpE)
apply (clarsimp simp: final_def)
apply clarsimp
apply (erule step_Normal_elim_cases)
apply clarsimp
apply (rename_tac t c' x)
apply (erule_tac x=x in ballE)
apply clarsimp
apply (drule_tac x="Normal x" in spec)
apply (drule_tac x="t" in spec)
apply (drule_tac x="c'" in spec)
apply fastforce+
done
next
case (Conseq \<Gamma> \<Theta> F P c Q A) thus ?case
apply (clarsimp)
apply (rename_tac P' Q' A' t c' x)
apply (erule_tac x="Normal x" in allE)
apply (erule_tac x="t" in allE)
apply (erule_tac x="c'" in allE)
apply (clarsimp simp: pre_weaken_pre)
apply force
done
qed
lemmas oghoare_sound = oghoare_soundness[THEN conjunct1, rule_format]
lemmas oghoare_seq_sound = oghoare_soundness[THEN conjunct2, rule_format]
end
diff --git a/thys/ConcurrentGC/Handshakes.thy b/thys/ConcurrentGC/Handshakes.thy
--- a/thys/ConcurrentGC/Handshakes.thy
+++ b/thys/ConcurrentGC/Handshakes.thy
@@ -1,583 +1,585 @@
(*<*)
(*
* Copyright 2015, NICTA
*
* This software may be distributed and modified according to the terms of
* the BSD 2-Clause license. Note that NO WARRANTY is provided.
* See "LICENSE_BSD2.txt" for details.
*
* @TAG(NICTA_BSD)
*)
theory Handshakes
imports
TSO
begin
(*>*)
subsection\<open>Handshake phases\<close>
text\<open>
The mutators can be at most one step behind the garbage collector (and
system). If any mutator is behind then the GC is stalled on a pending
handshake. Unfortunately this is a complicated by needing to consider
the handshake type due to \<open>get_work\<close>. This relation is very
precise.
\<close>
definition hp_step_rel :: "(bool \<times> handshake_type \<times> handshake_phase \<times> handshake_phase) set" where
"hp_step_rel \<equiv>
{ True } \<times> ({ (ht_NOOP, hp, hp) |hp. hp \<in> {hp_Idle, hp_IdleInit, hp_InitMark, hp_Mark} }
\<union> { (ht_GetRoots, hp_IdleMarkSweep, hp_IdleMarkSweep)
, (ht_GetWork, hp_IdleMarkSweep, hp_IdleMarkSweep) })
\<union> { False } \<times> { (ht_NOOP, hp_Idle, hp_IdleMarkSweep)
, (ht_NOOP, hp_IdleInit, hp_Idle)
, (ht_NOOP, hp_InitMark, hp_IdleInit)
, (ht_NOOP, hp_Mark, hp_InitMark)
, (ht_GetRoots, hp_IdleMarkSweep, hp_Mark)
, (ht_GetWork, hp_IdleMarkSweep, hp_IdleMarkSweep) }"
definition handshake_phase_inv :: "('field, 'mut, 'ref) lsts_pred" where
"handshake_phase_inv = (\<^bold>\<forall>m.
(sys_ghost_handshake_in_sync m \<^bold>\<otimes> sys_handshake_type
\<^bold>\<otimes> sys_ghost_handshake_phase \<^bold>\<otimes> mut_m.mut_ghost_handshake_phase m) \<^bold>\<in> \<langle>hp_step_rel\<rangle>
\<^bold>\<and> (sys_handshake_pending m \<^bold>\<longrightarrow> \<^bold>\<not>(sys_ghost_handshake_in_sync m)))"
(*<*)
(* Sanity *)
lemma handshake_step_inv:
"hp' = handshake_step hp \<Longrightarrow> \<exists>in' ht. (in', ht, hp', hp) \<in> hp_step_rel"
by (cases hp) (auto simp: hp_step_rel_def)
(* Sanity *)
lemma handshake_step_invD:
"(False, ht, hp', hp) \<in> hp_step_rel \<Longrightarrow> hp' = hp_step ht hp"
by (cases ht) (auto simp: hp_step_rel_def)
lemma (in mut_m) handshake_phase_invD:
"handshake_phase_inv s \<Longrightarrow> (sys_ghost_handshake_in_sync m s, sys_handshake_type s, sys_ghost_handshake_phase s, mut_ghost_handshake_phase s) \<in> hp_step_rel
\<and> (sys_handshake_pending m s \<longrightarrow> \<not>sys_ghost_handshake_in_sync m s)"
by (simp add: handshake_phase_inv_def)
lemma handshake_in_syncD:
"\<lbrakk> All (ghost_handshake_in_sync (s sys)); handshake_phase_inv s \<rbrakk>
\<Longrightarrow> \<forall>m'. mut_m.mut_ghost_handshake_phase m' s = sys_ghost_handshake_phase s"
by clarsimp (auto simp: hp_step_rel_def dest!: mut_m.handshake_phase_invD)
lemma (in sys) handshake_phase_inv[intro]:
"\<lbrace> LSTP handshake_phase_inv \<rbrace> sys"
by (vcg_jackhammer simp: handshake_phase_inv_def)
(*>*)
text\<open>
Connect @{const "sys_ghost_handshake_phase"} with locations in the GC.
\<close>
locset_definition "hp_Idle_locs =
(prefixed ''idle_noop'' - { ''idle_noop_mfence'', ''idle_noop_init_type'' })
\<union> { ''idle_read_fM'', ''idle_invert_fM'', ''idle_write_fM'', ''idle_flip_noop_mfence'', ''idle_flip_noop_init_type'' }"
locset_definition "hp_IdleInit_locs =
(prefixed ''idle_flip_noop'' - { ''idle_flip_noop_mfence'', ''idle_flip_noop_init_type'' })
\<union> { ''idle_phase_init'', ''init_noop_mfence'', ''init_noop_init_type'' }"
locset_definition "hp_InitMark_locs =
(prefixed ''init_noop'' - { ''init_noop_mfence'', ''init_noop_init_type'' })
\<union> { ''init_phase_mark'', ''mark_read_fM'', ''mark_write_fA'', ''mark_noop_mfence'', ''mark_noop_init_type'' }"
locset_definition "hp_IdleMarkSweep_locs =
{ ''idle_noop_mfence'', ''idle_noop_init_type'', ''mark_end'' }
\<union> sweep_locs
\<union> (mark_loop_locs - { ''mark_loop_get_roots_init_type'' })"
locset_definition "hp_Mark_locs =
(prefixed ''mark_noop'' - { ''mark_noop_mfence'', ''mark_noop_init_type'' })
\<union> { ''mark_loop_get_roots_init_type'' }"
abbreviation
"hs_noop_prefixes \<equiv> {''idle_noop'', ''idle_flip_noop'', ''init_noop'', ''mark_noop'' }"
locset_definition "hs_noop_locs =
(\<Union>l \<in> hs_noop_prefixes. prefixed l - (suffixed ''_noop_mfence'' \<union> suffixed ''_noop_init_type''))"
locset_definition "hs_get_roots_locs =
prefixed ''mark_loop_get_roots'' - {''mark_loop_get_roots_init_type''}"
locset_definition "hs_get_work_locs =
prefixed ''mark_loop_get_work'' - {''mark_loop_get_work_init_type''}"
abbreviation "hs_prefixes \<equiv>
hs_noop_prefixes \<union> { ''mark_loop_get_roots'', ''mark_loop_get_work'' }"
locset_definition "hs_init_loop_locs = (\<Union>l \<in> hs_prefixes. prefixed (l @ ''_init_loop''))"
locset_definition "hs_done_loop_locs = (\<Union>l \<in> hs_prefixes. prefixed (l @ ''_done_loop''))"
locset_definition "hs_done_locs = (\<Union>l \<in> hs_prefixes. prefixed (l @ ''_done''))"
locset_definition "hs_none_pending_locs = - (hs_init_loop_locs \<union> hs_done_locs)"
locset_definition "hs_in_sync_locs =
(- ( (\<Union>l \<in> hs_prefixes. prefixed (l @ ''_init'')) \<union> hs_done_locs ))
\<union> (\<Union>l \<in> hs_prefixes. {l @ ''_init_type''})"
locset_definition "hs_out_of_sync_locs =
(\<Union>l \<in> hs_prefixes. {l @ ''_init_muts''})"
locset_definition "hs_mut_in_muts_locs =
(\<Union>l \<in> hs_prefixes. {l @ ''_init_loop_set_pending'', l @ ''_init_loop_done''})"
locset_definition "hs_init_loop_done_locs =
(\<Union>l \<in> hs_prefixes. {l @ ''_init_loop_done''})"
locset_definition "hs_init_loop_not_done_locs =
(hs_init_loop_locs - (\<Union>l \<in> hs_prefixes. {l @ ''_init_loop_done''}))"
inv_definition (in gc) handshake_invL :: "('field, 'mut, 'ref) gc_pred" where
"handshake_invL =
(atS_gc hs_noop_locs (sys_handshake_type \<^bold>= \<langle>ht_NOOP\<rangle>)
\<^bold>\<and> atS_gc hs_get_roots_locs (sys_handshake_type \<^bold>= \<langle>ht_GetRoots\<rangle>)
\<^bold>\<and> atS_gc hs_get_work_locs (sys_handshake_type \<^bold>= \<langle>ht_GetWork\<rangle>)
\<^bold>\<and> atS_gc hs_mut_in_muts_locs (gc_mut \<^bold>\<in> gc_muts)
\<^bold>\<and> atS_gc hs_init_loop_locs (\<^bold>\<forall>m. \<^bold>\<not>(\<langle>m\<rangle> \<^bold>\<in> gc_muts) \<^bold>\<longrightarrow> sys_handshake_pending m
\<^bold>\<or> sys_ghost_handshake_in_sync m)
\<^bold>\<and> atS_gc hs_init_loop_not_done_locs (\<^bold>\<forall>m. \<langle>m\<rangle> \<^bold>\<in> gc_muts \<^bold>\<longrightarrow> \<^bold>\<not>(sys_handshake_pending m)
\<^bold>\<and> \<^bold>\<not>(sys_ghost_handshake_in_sync m))
\<^bold>\<and> atS_gc hs_init_loop_done_locs ( (sys_handshake_pending \<^bold>$ gc_mut
\<^bold>\<or> sys_ghost_handshake_in_sync \<^bold>$ gc_mut)
\<^bold>\<and> (\<^bold>\<forall>m. \<langle>m\<rangle> \<^bold>\<in> gc_muts \<^bold>\<and> \<langle>m\<rangle> \<^bold>\<noteq> gc_mut
\<^bold>\<longrightarrow> \<^bold>\<not>(sys_handshake_pending m)
\<^bold>\<and> \<^bold>\<not>(sys_ghost_handshake_in_sync m)) )
\<^bold>\<and> atS_gc hs_done_locs (\<^bold>\<forall>m. sys_handshake_pending m \<^bold>\<or> sys_ghost_handshake_in_sync m)
\<^bold>\<and> atS_gc hs_done_loop_locs (\<^bold>\<forall>m. \<^bold>\<not>(\<langle>m\<rangle> \<^bold>\<in> gc_muts) \<^bold>\<longrightarrow> \<^bold>\<not>(sys_handshake_pending m))
\<^bold>\<and> atS_gc hs_none_pending_locs (\<^bold>\<forall>m. \<^bold>\<not>(sys_handshake_pending m))
\<^bold>\<and> atS_gc hs_in_sync_locs (\<^bold>\<forall>m. sys_ghost_handshake_in_sync m)
\<^bold>\<and> atS_gc hs_out_of_sync_locs (\<^bold>\<forall>m. \<^bold>\<not>(sys_handshake_pending m)
\<^bold>\<and> \<^bold>\<not>(sys_ghost_handshake_in_sync m))
\<^bold>\<and> atS_gc hp_Idle_locs (sys_ghost_handshake_phase \<^bold>= \<langle>hp_Idle\<rangle>)
\<^bold>\<and> atS_gc hp_IdleInit_locs (sys_ghost_handshake_phase \<^bold>= \<langle>hp_IdleInit\<rangle>)
\<^bold>\<and> atS_gc hp_InitMark_locs (sys_ghost_handshake_phase \<^bold>= \<langle>hp_InitMark\<rangle>)
\<^bold>\<and> atS_gc hp_IdleMarkSweep_locs (sys_ghost_handshake_phase \<^bold>= \<langle>hp_IdleMarkSweep\<rangle>)
\<^bold>\<and> atS_gc hp_Mark_locs (sys_ghost_handshake_phase \<^bold>= \<langle>hp_Mark\<rangle>))"
(*<*)
lemma hs_get_roots_locs_subseteq_hp_IdleMarkSweep_locs:
"hs_get_roots_locs \<subseteq> hp_IdleMarkSweep_locs"
by (auto simp: hs_get_roots_locs_def hp_IdleMarkSweep_locs_def mark_loop_locs_def
intro: append_prefixD)
lemma hs_get_work_locs_subseteq_hp_IdleMarkSweep_locs:
"hs_get_work_locs \<subseteq> hp_IdleMarkSweep_locs"
apply (simp add: hs_get_work_locs_def hp_IdleMarkSweep_locs_def mark_loop_locs_def)
apply clarsimp
apply (drule mp)
apply (auto intro: append_prefixD)[1]
apply auto
done
lemma gc_handshake_invL_eq_imp:
"eq_imp (\<lambda>(_::unit) s. (AT s gc, s\<down> gc, sys_ghost_handshake_phase s\<down>, handshake_pending (s\<down> sys), ghost_handshake_in_sync (s\<down> sys), sys_handshake_type s\<down>))
gc.handshake_invL"
by (simp add: gc.handshake_invL_def eq_imp_def)
lemmas gc_handshake_invL_niE[nie] =
iffD1[OF gc_handshake_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode], rotated -1]
lemma (in gc) handshake_invL[intro]:
"\<lbrace> handshake_invL \<rbrace> gc"
by vcg_jackhammer_ff
lemma (in sys) gc_handshake_invL[intro]:
notes gc.handshake_invL_def[inv]
shows "\<lbrace> gc.handshake_invL \<rbrace> sys"
by vcg_ni
lemma (in gc) handshake_phase_inv[intro]:
"\<lbrace> handshake_invL \<^bold>\<and> LSTP handshake_phase_inv \<rbrace> gc \<lbrace> LSTP handshake_phase_inv \<rbrace>"
by (vcg_jackhammer_ff simp: handshake_phase_inv_def hp_step_rel_def)
(*>*)
text\<open>
Local handshake phase invariant for the mutators.
\<close>
locset_definition "mut_no_pending_mutates_locs =
(prefixed ''hs_noop'' - { ''hs_noop'', ''hs_noop_mfence'' })
\<union> (prefixed ''hs_get_roots'' - { ''hs_get_roots'', ''hs_get_roots_mfence'' })
\<union> (prefixed ''hs_get_work'' - { ''hs_get_work'', ''hs_get_work_mfence'' })"
inv_definition (in mut_m) handshake_invL :: "('field, 'mut, 'ref) gc_pred" where
"handshake_invL =
(atS_mut (prefixed ''hs_noop_'') (sys_handshake_type \<^bold>= \<langle>ht_NOOP\<rangle> \<^bold>\<and> sys_handshake_pending m)
\<^bold>\<and> atS_mut (prefixed ''hs_get_roots_'') (sys_handshake_type \<^bold>= \<langle>ht_GetRoots\<rangle> \<^bold>\<and> sys_handshake_pending m)
\<^bold>\<and> atS_mut (prefixed ''hs_get_work_'') (sys_handshake_type \<^bold>= \<langle>ht_GetWork\<rangle> \<^bold>\<and> sys_handshake_pending m)
\<^bold>\<and> atS_mut mut_no_pending_mutates_locs (LIST_NULL (tso_pending_mutate (mutator m))))"
(*<*)
lemma (in mut_m) handshake_invL_eq_imp:
"eq_imp (\<lambda>(_::unit) s. (AT s (mutator m), s\<down> (mutator m), sys_handshake_type s\<down>, handshake_pending (s\<down> sys), mem_write_buffers (s\<down> sys) (mutator m)))
handshake_invL"
by (simp add: eq_imp_def handshake_invL_def)
lemmas mut_m_handshake_invL_niE[nie] =
iffD1[OF mut_m.handshake_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode], rotated -1]
lemma (in mut_m) handshake_invL[intro]:
"\<lbrace> handshake_invL \<rbrace> mutator m"
by vcg_jackhammer
lemma (in mut_m') handshake_invL[intro]:
"\<lbrace> handshake_invL \<rbrace> mutator m'"
by vcg_nihe vcg_ni+
lemma (in gc) mut_handshake_invL[intro]:
notes mut_m.handshake_invL_def[inv]
shows "\<lbrace> handshake_invL \<^bold>\<and> mut_m.handshake_invL m \<rbrace> gc \<lbrace> mut_m.handshake_invL m \<rbrace>"
by vcg_nihe vcg_ni+
lemma (in sys) mut_handshake_invL[intro]:
notes mut_m.handshake_invL_def[inv]
shows "\<lbrace> mut_m.handshake_invL m \<rbrace> sys"
by (vcg_ni split: if_splits)
lemma (in mut_m) gc_handshake_invL[intro]:
notes gc.handshake_invL_def[inv]
shows "\<lbrace> handshake_invL \<^bold>\<and> gc.handshake_invL \<rbrace> mutator m \<lbrace> gc.handshake_invL \<rbrace>"
by vcg_nihe vcg_ni+
lemma (in mut_m) handshake_phase_inv[intro]:
"\<lbrace> handshake_invL \<^bold>\<and> LSTP handshake_phase_inv \<rbrace> mutator m \<lbrace> LSTP handshake_phase_inv \<rbrace>"
by (vcg_jackhammer simp: handshake_phase_inv_def) (auto simp: hp_step_rel_def)
(*>*)
text\<open>
Relate @{const "sys_ghost_handshake_phase"}, @{const "gc_phase"},
@{const "sys_phase"} and writes to the phase in the GC's TSO buffer.
The first relation treats the case when the GC's TSO buffer does not
contain any writes to the phase.
The second relation exhibits the data race on the phase variable: we
need to precisely track the possible states of the GC's TSO buffer.
\<close>
definition handshake_phase_rel :: "handshake_phase \<Rightarrow> bool \<Rightarrow> gc_phase \<Rightarrow> bool" where
"handshake_phase_rel hp in_sync ph \<equiv>
case hp of
hp_Idle \<Rightarrow> ph = ph_Idle
| hp_IdleInit \<Rightarrow> ph = ph_Idle \<or> (in_sync \<and> ph = ph_Init)
| hp_InitMark \<Rightarrow> ph = ph_Init \<or> (in_sync \<and> ph = ph_Mark)
| hp_Mark \<Rightarrow> ph = ph_Mark
| hp_IdleMarkSweep \<Rightarrow> ph = ph_Mark \<or> (in_sync \<and> ph \<in> { ph_Idle, ph_Sweep })"
definition phase_rel :: "(bool \<times> handshake_phase \<times> gc_phase \<times> gc_phase \<times> ('field, 'ref) mem_write_action list) set" where
"phase_rel \<equiv>
{ (in_sync, hp, ph, ph, []) |in_sync hp ph. handshake_phase_rel hp in_sync ph }
\<union> ({True} \<times> { (hp_IdleInit, ph_Init, ph_Idle, [mw_Phase ph_Init]),
(hp_InitMark, ph_Mark, ph_Init, [mw_Phase ph_Mark]),
(hp_IdleMarkSweep, ph_Sweep, ph_Mark, [mw_Phase ph_Sweep]),
(hp_IdleMarkSweep, ph_Idle, ph_Mark, [mw_Phase ph_Sweep, mw_Phase ph_Idle]),
(hp_IdleMarkSweep, ph_Idle, ph_Sweep, [mw_Phase ph_Idle]) })"
definition phase_rel_inv :: "('field, 'mut, 'ref) lsts_pred" where
"phase_rel_inv = ((\<^bold>\<forall>m. sys_ghost_handshake_in_sync m) \<^bold>\<otimes> sys_ghost_handshake_phase \<^bold>\<otimes> gc_phase \<^bold>\<otimes> sys_phase \<^bold>\<otimes> tso_pending_phase gc \<^bold>\<in> \<langle>phase_rel\<rangle>)"
(*<*)
simps_of_case handshake_phase_rel_simps[simp]: handshake_phase_rel_def (splits: handshake_phase.split)
lemma phase_rel_invD:
"phase_rel_inv s \<Longrightarrow> (\<forall>m. sys_ghost_handshake_in_sync m s, sys_ghost_handshake_phase s, gc_phase s, sys_phase s, tso_pending_phase gc s) \<in> phase_rel"
by (simp add: phase_rel_inv_def)
lemma phases_rel_Id[iff]:
"(\<forall>m. sys_ghost_handshake_in_sync m s, sys_ghost_handshake_phase s, gc_phase s, sys_phase s, tso_pending_phase gc s) \<in> phase_rel
\<Longrightarrow> (\<forall>ph. mw_Phase ph \<notin> set (sys_mem_write_buffers gc s)) \<longleftrightarrow> sys_phase s = gc_phase s"
by (auto simp: phase_rel_def filter_empty_conv filter_eq_Cons_iff)
(*>*)
text\<open>
Tie the garbage collector's control location to the value of @{const
"gc_phase"}.
\<close>
locset_definition no_pending_phase_locs :: "location set" where
"no_pending_phase_locs \<equiv>
(idle_locs - { ''idle_noop_mfence'' })
\<union> (init_locs - { ''init_noop_mfence'' })
\<union> (mark_locs - { ''mark_read_fM'', ''mark_write_fA'', ''mark_noop_mfence'' })"
inv_definition (in gc) phase_invL :: "('field, 'mut, 'ref) gc_pred" where
"phase_invL =
(atS_gc idle_locs (gc_phase \<^bold>= \<langle>ph_Idle\<rangle>)
\<^bold>\<and> atS_gc init_locs (gc_phase \<^bold>= \<langle>ph_Init\<rangle>)
\<^bold>\<and> atS_gc mark_locs (gc_phase \<^bold>= \<langle>ph_Mark\<rangle>)
\<^bold>\<and> atS_gc sweep_locs (gc_phase \<^bold>= \<langle>ph_Sweep\<rangle>)
\<^bold>\<and> atS_gc no_pending_phase_locs (LIST_NULL (tso_pending_phase gc)))"
(*<*)
lemma (in gc) phase_invL_eq_imp:
"eq_imp (\<lambda>r (s :: ('field, 'mut, 'ref) gc_pred_state). (AT s gc, s\<down> gc, tso_pending_phase gc s\<down>))
phase_invL"
by (clarsimp simp: eq_imp_def inv)
lemmas gc_phase_invL_niE[nie] =
iffD1[OF gc.phase_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode], rotated -1]
lemma (in gc) phase_invL[intro]:
"\<lbrace> phase_invL \<^bold>\<and> LSTP phase_rel_inv \<rbrace> gc \<lbrace> phase_invL \<rbrace>"
by (vcg_jackhammer dest!: phase_rel_invD simp: phase_rel_def)
lemma (in sys) gc_phase_invL[intro]:
notes gc.phase_invL_def[inv]
shows "\<lbrace> gc.phase_invL \<^bold>\<and> LSTP tso_writes_inv \<rbrace> sys \<lbrace> gc.phase_invL \<rbrace>"
by (vcg_ni split: if_splits)
lemma (in mut_m) gc_phase_invL[intro]:
notes gc.phase_invL_def[inv]
shows "\<lbrace> gc.phase_invL \<rbrace> mutator m"
by vcg_nihe
lemma (in gc) phase_rel_inv[intro]:
notes phase_rel_inv_def[inv]
shows "\<lbrace> handshake_invL \<^bold>\<and> phase_invL \<^bold>\<and> LSTP phase_rel_inv \<rbrace> gc \<lbrace> LSTP phase_rel_inv \<rbrace>"
by (vcg_jackhammer_ff dest!: phase_rel_invD simp: phase_rel_def)
lemma (in sys) phase_rel_inv[intro]:
notes gc.phase_invL_def[inv] phase_rel_inv_def[inv]
shows "\<lbrace> LSTP (phase_rel_inv \<^bold>\<and> tso_writes_inv) \<rbrace> sys \<lbrace> LSTP phase_rel_inv \<rbrace>"
apply vcg_jackhammer
apply (simp add: phase_rel_def p_not_sys split: if_splits)
apply (elim disjE; auto split: if_splits)
done
lemma (in mut_m) phase_rel_inv[intro]:
"\<lbrace> handshake_invL \<^bold>\<and> LSTP (handshake_phase_inv \<^bold>\<and> phase_rel_inv) \<rbrace>
mutator m
\<lbrace> LSTP phase_rel_inv \<rbrace>"
apply (vcg_jackhammer simp: phase_rel_inv_def)
subgoal by (auto dest!: handshake_phase_invD
simp: handshake_phase_rel_def phase_rel_def hp_step_rel_def
split: handshake_phase.splits)
subgoal by (auto dest!: handshake_phase_invD
simp: handshake_phase_rel_def phase_rel_def hp_step_rel_def
split: handshake_phase.splits)
subgoal by (auto dest!: handshake_phase_invD
simp: handshake_phase_rel_def phase_rel_def hp_step_rel_def
split: handshake_phase.splits)
done
(*>*)
text\<open>
Validity of @{const "sys_fM"} wrt @{const "gc_fM"} and the handshake
phase. Effectively we use @{const "gc_fM"} as ghost state. We also
include the TSO lock to rule out the GC having any pending marks
during the @{const "hp_Idle"} handshake phase.
\<close>
definition fM_rel :: "(bool \<times> handshake_phase \<times> gc_mark \<times> gc_mark \<times> ('field, 'ref) mem_write_action list \<times> bool) set" where
"fM_rel =
{ (in_sync, hp, fM, fM, [], l) |fM hp in_sync l. hp = hp_Idle \<longrightarrow> \<not>in_sync }
\<union> { (in_sync, hp_Idle, fM, fM', [], l) |fM fM' in_sync l. in_sync }
\<union> { (in_sync, hp_Idle, \<not>fM, fM, [mw_fM (\<not>fM)], False) |fM in_sync. in_sync }"
definition fM_rel_inv :: "('field, 'mut, 'ref) lsts_pred" where
"fM_rel_inv = ((\<^bold>\<forall>m. sys_ghost_handshake_in_sync m) \<^bold>\<otimes> sys_ghost_handshake_phase \<^bold>\<otimes> gc_fM \<^bold>\<otimes> sys_fM \<^bold>\<otimes> tso_pending_fM gc \<^bold>\<otimes> (sys_mem_lock \<^bold>= \<langle>Some gc\<rangle>) \<^bold>\<in> \<langle>fM_rel\<rangle>)"
definition fA_rel :: "(bool \<times> handshake_phase \<times> gc_mark \<times> gc_mark \<times> ('field, 'ref) mem_write_action list) set" where
"fA_rel =
{ (in_sync, hp_Idle, fA, fM, []) |fA fM in_sync. \<not>in_sync \<longrightarrow> fA = fM }
\<union> { (in_sync, hp_IdleInit, fA, \<not>fA, []) |fA in_sync. True }
\<union> { (in_sync, hp_InitMark, fA, \<not>fA, [mw_fA (\<not>fA)]) |fA in_sync. in_sync }
\<union> { (in_sync, hp_InitMark, fA, fM, []) |fA fM in_sync. \<not>in_sync \<longrightarrow> fA \<noteq> fM }
\<union> { (in_sync, hp_Mark, fA, fA, []) |fA in_sync. True }
\<union> { (in_sync, hp_IdleMarkSweep, fA, fA, []) |fA in_sync. True }"
definition fA_rel_inv :: "('field, 'mut, 'ref) lsts_pred" where
"fA_rel_inv = ((\<^bold>\<forall>m. sys_ghost_handshake_in_sync m) \<^bold>\<otimes> sys_ghost_handshake_phase \<^bold>\<otimes> sys_fA \<^bold>\<otimes> gc_fM \<^bold>\<otimes> tso_pending_fA gc \<^bold>\<in> \<langle>fA_rel\<rangle>)"
locset_definition "fM_eq_locs = (- { ''idle_write_fM'', ''idle_flip_noop_mfence'' })"
locset_definition "fM_tso_empty_locs = (- { ''idle_flip_noop_mfence'' })"
locset_definition "fA_tso_empty_locs = (- { ''mark_noop_mfence'' })"
locset_definition
"fA_eq_locs \<equiv> { ''idle_read_fM'', ''idle_invert_fM'' }
\<union> prefixed ''idle_noop''
\<union> (mark_locs - { ''mark_read_fM'', ''mark_write_fA'', ''mark_noop_mfence'' })
\<union> sweep_locs"
locset_definition
"fA_neq_locs \<equiv> { ''idle_phase_init'', ''idle_write_fM'', ''mark_read_fM'', ''mark_write_fA'' }
\<union> prefixed ''idle_flip_noop''
\<union> init_locs"
inv_definition (in gc) fM_fA_invL :: "('field, 'mut, 'ref) gc_pred" where
"fM_fA_invL =
(atS_gc fM_eq_locs (sys_fM \<^bold>= gc_fM)
\<^bold>\<and> at_gc ''idle_write_fM'' (sys_fM \<^bold>\<noteq> gc_fM)
\<^bold>\<and> at_gc ''idle_flip_noop_mfence'' (sys_fM \<^bold>\<noteq> gc_fM \<^bold>\<longrightarrow> (\<^bold>\<not>(LIST_NULL (tso_pending_fM gc))))
\<^bold>\<and> atS_gc fM_tso_empty_locs (LIST_NULL (tso_pending_fM gc))
\<^bold>\<and> atS_gc fA_eq_locs (sys_fA \<^bold>= gc_fM)
\<^bold>\<and> atS_gc fA_neq_locs (sys_fA \<^bold>\<noteq> gc_fM)
\<^bold>\<and> at_gc ''mark_noop_mfence'' (sys_fA \<^bold>\<noteq> gc_fM \<^bold>\<longrightarrow> (\<^bold>\<not>(LIST_NULL (tso_pending_fA gc))))
\<^bold>\<and> atS_gc fA_tso_empty_locs (LIST_NULL (tso_pending_fA gc)))"
(*<*)
lemmas fM_rel_invD = iffD1[OF fun_cong[OF fM_rel_inv_def[simplified atomize_eq]]]
lemma do_write_action_fM[simp]:
"\<lbrakk> sys_mem_write_buffers p s = w # ws; tso_writes_inv s; handshake_phase_inv s; fM_rel_inv s;
ghost_handshake_phase (s (mutator m)) \<noteq> hp_Idle \<or> (sys_ghost_handshake_phase s = hp_IdleMarkSweep \<and> All (ghost_handshake_in_sync (s sys)));
p \<noteq> sys \<rbrakk>
\<Longrightarrow> fM (do_write_action w (s sys)) = fM (s sys)"
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (clarsimp simp: do_write_action_def p_not_sys
split: mem_write_action.splits)
apply (simp add: hp_step_rel_def)
apply (elim disjE, auto simp: fM_rel_inv_def fM_rel_def)
done
lemmas fA_rel_invD = iffD1[OF fun_cong[OF fA_rel_inv_def[simplified atomize_eq]]]
lemma gc_fM_fA_invL_eq_imp:
"eq_imp (\<lambda>(_::unit) s. (AT s gc, s\<down> gc, sys_fA s\<down>, sys_fM s\<down>, sys_mem_write_buffers gc s\<down>))
gc.fM_fA_invL"
by (simp add: gc.fM_fA_invL_def eq_imp_def)
lemmas gc_fM_fA_invL_niE[nie] =
iffD1[OF gc_fM_fA_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode], rotated -1]
lemma (in gc) fM_fA_invL[intro]:
"\<lbrace> fM_fA_invL \<rbrace> gc"
by vcg_jackhammer
lemma (in gc) fM_rel_inv[intro]:
"\<lbrace> fM_fA_invL \<^bold>\<and> handshake_invL \<^bold>\<and> tso_lock_invL \<^bold>\<and> LSTP fM_rel_inv \<rbrace>
gc
\<lbrace> LSTP fM_rel_inv \<rbrace>"
by (vcg_jackhammer simp: fM_rel_inv_def fM_rel_def)
lemma (in gc) fA_rel_inv[intro]:
"\<lbrace> fM_fA_invL \<^bold>\<and> handshake_invL \<^bold>\<and> LSTP fA_rel_inv \<rbrace>
gc
\<lbrace> LSTP fA_rel_inv \<rbrace>"
by (vcg_jackhammer simp: fA_rel_inv_def; auto simp: fA_rel_def)
lemma (in mut_m) gc_fM_fA_invL[intro]:
notes gc.fM_fA_invL_def[inv]
shows "\<lbrace> gc.fM_fA_invL \<rbrace> mutator m"
by vcg_nihe
lemma (in mut_m) fM_rel_inv[intro]:
"\<lbrace> LSTP fM_rel_inv \<rbrace> mutator m"
by (vcg_jackhammer simp: fM_rel_inv_def fM_rel_def; elim disjE; auto split: if_splits)
(* FIXME trivial but eta reduction plays merry hell *)
lemma (in mut_m) fA_rel_inv[intro]:
"\<lbrace> LSTP fA_rel_inv \<rbrace> mutator m"
by (vcg_jackhammer simp: fA_rel_inv_def; simp add: fA_rel_def; elim disjE; auto split: if_splits)
lemma fA_neq_locs_diff_fA_tso_empty_locs[simp]:
"fA_neq_locs - fA_tso_empty_locs = {}"
by (simp add: fA_neq_locs_def fA_tso_empty_locs_def loc)
lemma (in sys) gc_fM_fA_invL[intro]:
notes gc.fM_fA_invL_def[inv]
shows
"\<lbrace> gc.fM_fA_invL \<^bold>\<and> LSTP (fA_rel_inv \<^bold>\<and> fM_rel_inv \<^bold>\<and> tso_writes_inv) \<rbrace>
sys
\<lbrace> gc.fM_fA_invL \<rbrace>"
apply (vcg_ni simp: p_not_sys)
apply (erule disjE)
apply (clarsimp simp: fM_rel_inv_def fM_rel_def split: if_splits)
apply (clarsimp simp: fM_rel_inv_def fM_rel_def filter_empty_conv)
apply (erule disjE; clarsimp)
apply (rule conjI; clarsimp simp: fM_rel_inv_def fM_rel_def split: if_splits)
apply (clarsimp split: if_splits)
apply (erule disjE)
apply (clarsimp simp: fA_rel_inv_def fA_rel_def)
apply clarsimp
apply (erule disjE)
apply (clarsimp simp: fA_rel_inv_def fA_rel_def fM_rel_inv_def fM_rel_def)
apply (drule (1) atS_dests(3), fastforce simp: atS_simps)
apply clarsimp
apply (rule conjI)
apply (clarsimp simp: fA_rel_inv_def fA_rel_def split: if_splits)
apply clarsimp
apply (clarsimp split: if_splits)
done
lemma (in sys) fM_rel_inv[intro]:
"\<lbrace> LSTP (fM_rel_inv \<^bold>\<and> tso_writes_inv) \<rbrace> sys \<lbrace> LSTP fM_rel_inv \<rbrace>"
apply (vcg_ni simp: p_not_sys)
apply (fastforce simp: do_write_action_def fM_rel_inv_def fM_rel_def
split: mem_write_action.splits)
done
lemma (in sys) fA_rel_inv[intro]:
"\<lbrace> LSTP (fA_rel_inv \<^bold>\<and> tso_writes_inv) \<rbrace> sys \<lbrace> LSTP fA_rel_inv \<rbrace>"
apply (vcg_ni simp: p_not_sys)
apply (fastforce simp: do_write_action_def fA_rel_inv_def fA_rel_def
split: mem_write_action.splits)
done
lemma mut_m_get_roots_no_fM_write:
"\<lbrakk> atS (mutator m) (prefixed ''hs_get_roots_'') s; mut_m.handshake_invL m s; handshake_phase_inv s\<down>; fM_rel_inv s\<down>; tso_writes_inv s\<down>; p \<noteq> sys \<rbrakk>
\<Longrightarrow> \<not>sys_mem_write_buffers p s\<down> = mw_fM fl # ws"
unfolding mut_m.handshake_invL_def
apply (elim conjE)
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (drule fM_rel_invD)
apply (fastforce simp: hp_step_rel_def fM_rel_def loc filter_empty_conv p_not_sys)
done
lemma mut_m_get_roots_no_phase_write:
"\<lbrakk> atS (mutator m) (prefixed ''hs_get_roots_'') s; mut_m.handshake_invL m s; handshake_phase_inv s\<down>; phase_rel_inv s\<down>; tso_writes_inv s\<down>; p \<noteq> sys \<rbrakk>
\<Longrightarrow> \<not>sys_mem_write_buffers p s\<down> = mw_Phase ph # ws"
unfolding mut_m.handshake_invL_def
apply (elim conjE)
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (drule phase_rel_invD)
apply (fastforce simp: hp_step_rel_def phase_rel_def loc filter_empty_conv p_not_sys)
done
lemma mut_m_not_idle_no_fM_write:
"\<lbrakk> ghost_handshake_phase (s (mutator m)) \<noteq> hp_Idle; fM_rel_inv s; handshake_phase_inv s; tso_writes_inv s; p \<noteq> sys \<rbrakk>
\<Longrightarrow> \<not>sys_mem_write_buffers p s = mw_fM fl # ws"
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (drule fM_rel_invD)
apply (fastforce simp: hp_step_rel_def fM_rel_def filter_empty_conv p_not_sys)
done
lemma (in mut_m) mut_ghost_handshake_phase_idle:
"\<lbrakk> mut_ghost_handshake_phase s = hp_Idle; handshake_phase_inv s; phase_rel_inv s \<rbrakk>
\<Longrightarrow> sys_phase s = ph_Idle"
-by (auto dest!: phase_rel_invD handshake_phase_invD
- simp: phase_rel_def hp_step_rel_def)
+ apply (drule phase_rel_invD)
+ apply (drule handshake_phase_invD)
+ apply (auto simp: phase_rel_def hp_step_rel_def)
+ done
(*>*)
(*<*)
end
(*>*)
diff --git a/thys/ConcurrentGC/StrongTricolour.thy b/thys/ConcurrentGC/StrongTricolour.thy
--- a/thys/ConcurrentGC/StrongTricolour.thy
+++ b/thys/ConcurrentGC/StrongTricolour.thy
@@ -1,3398 +1,3400 @@
(*<*)
(*
* Copyright 2015, NICTA
*
* This software may be distributed and modified according to the terms of
* the BSD 2-Clause license. Note that NO WARRANTY is provided.
* See "LICENSE_BSD2.txt" for details.
*
* @TAG(NICTA_BSD)
*)
theory StrongTricolour
imports
MarkObject
begin
(*>*)
subsection\<open>The strong-tricolour invariant \label{sec:strong-tricolour-invariant} \<close>
text\<open>
As the GC algorithm uses both insertion and deletion barriers, it
preserves the \emph{strong tricolour-invariant}:
\<close>
abbreviation points_to_white :: "'ref \<Rightarrow> 'ref \<Rightarrow> ('field, 'mut, 'ref) lsts_pred" (infix "points'_to'_white" 51) where
"x points_to_white y \<equiv> x points_to y \<^bold>\<and> white y"
definition strong_tricolour_inv :: "('field, 'mut, 'ref) lsts_pred" where
"strong_tricolour_inv = (\<^bold>\<forall>b w. black b \<^bold>\<longrightarrow> \<^bold>\<not>(b points_to_white w))"
text\<open>
Intuitively this invariant says that there are no pointers from
completely processed objects to the unexplored space; i.e., the grey
references properly separate the two. In contrast the weak tricolour
invariant allows such pointers, provided there is a grey reference
that protects the unexplored object.
\<close>
definition has_white_path_to :: "'ref \<Rightarrow> 'ref \<Rightarrow> ('field, 'mut, 'ref) lsts_pred" (infix "has'_white'_path'_to" 51) where
"x has_white_path_to y \<equiv> \<lambda>s. (\<lambda>x y. (x points_to_white y) s)\<^sup>*\<^sup>* x y"
definition grey_protects_white :: "'ref \<Rightarrow> 'ref \<Rightarrow> ('field, 'mut, 'ref) lsts_pred" (infix "grey'_protects'_white" 51) where
"g grey_protects_white w = (grey g \<^bold>\<and> g has_white_path_to w)"
definition weak_tricolour_inv :: "('field, 'mut, 'ref) lsts_pred" where
"weak_tricolour_inv =
(\<^bold>\<forall>b w. black b \<^bold>\<and> b points_to_white w \<^bold>\<longrightarrow> (\<^bold>\<exists>g. g grey_protects_white w))"
lemma "strong_tricolour_inv s \<Longrightarrow> weak_tricolour_inv s"
(*<*)
by (clarsimp simp: strong_tricolour_inv_def weak_tricolour_inv_def grey_protects_white_def)
(*>*)
text\<open>
The key invariant that the mutators establish as they perform \<open>get_roots\<close>: they protect their white-reachable references with grey
objects.
\<close>
definition in_snapshot :: "'ref \<Rightarrow> ('field, 'mut, 'ref) lsts_pred" where
"in_snapshot r = (black r \<^bold>\<or> (\<^bold>\<exists>g. g grey_protects_white r))"
definition (in mut_m) reachable_snapshot_inv :: "('field, 'mut, 'ref) lsts_pred" where
"reachable_snapshot_inv = (\<^bold>\<forall>r. reachable r \<^bold>\<longrightarrow> in_snapshot r)"
text\<open>
Note that it is not easy to specify precisely when the snapshot (of
objects the GC will retain) is taken due to the raggedness of the
initialisation.
In some phases we need to know that the insertion and deletion
barriers are installed, in order to preserve the snapshot. These can
ignore TSO effects as marks hit system memory in a timely way.
\<close>
abbreviation marked_insertion :: "('field, 'ref) mem_write_action \<Rightarrow> ('field, 'mut, 'ref) lsts_pred" where
"marked_insertion w \<equiv> \<lambda>s. case w of mw_Mutate r f (Some r') \<Rightarrow> marked r' s | _ \<Rightarrow> True"
definition (in mut_m) marked_insertions :: "('field, 'mut, 'ref) lsts_pred" where
"marked_insertions = (\<^bold>\<forall>w. tso_pending_write (mutator m) w \<^bold>\<longrightarrow> marked_insertion w)"
abbreviation marked_deletion :: "('field, 'ref) mem_write_action \<Rightarrow> ('field, 'mut, 'ref) lsts_pred" where
"marked_deletion w \<equiv> \<lambda>s. case w of mw_Mutate r f opt_r' \<Rightarrow> obj_at_field_on_heap (\<lambda>r'. marked r' s) r f s | _ \<Rightarrow> True"
definition (in mut_m) marked_deletions :: "('field, 'mut, 'ref) lsts_pred" where
"marked_deletions = (\<^bold>\<forall>w. tso_pending_write (mutator m) w \<^bold>\<longrightarrow> marked_deletion w)"
text\<open>
Finally, in some phases the heap is somewhat monochrome.
\<close>
definition black_heap :: "('field, 'mut, 'ref) lsts_pred" where
"black_heap = (\<^bold>\<forall>r. valid_ref r \<^bold>\<longrightarrow> black r)"
definition white_heap :: "('field, 'mut, 'ref) lsts_pred" where
"white_heap = (\<^bold>\<forall>r. valid_ref r \<^bold>\<longrightarrow> white r)"
definition no_black_refs :: "('field, 'mut, 'ref) lsts_pred" where
"no_black_refs = (\<^bold>\<forall>r. \<^bold>\<not>(black r))"
definition no_grey_refs :: "('field, 'mut, 'ref) lsts_pred" where
"no_grey_refs = (\<^bold>\<forall>r. \<^bold>\<not>(grey r))"
(*<*)
lemma no_black_refsD:
"no_black_refs s \<Longrightarrow> \<not>black r s"
unfolding no_black_refs_def by simp
lemma has_white_path_to_eq_imp:
"eq_imp (\<lambda>(_::unit). sys_fM \<^bold>\<otimes> sys_heap)
(x has_white_path_to y)"
by (clarsimp simp: eq_imp_def has_white_path_to_def obj_at_def cong: option.case_cong)
lemmas has_white_path_to_fun_upd[simp] = eq_imp_fun_upd[OF has_white_path_to_eq_imp, simplified eq_imp_simps, rule_format]
lemma has_white_path_toD[dest]:
"(x has_white_path_to y) s \<Longrightarrow> white y s \<or> x = y"
unfolding has_white_path_to_def by (fastforce elim: rtranclp.cases)
lemma has_white_path_toI[iff]:
"(x has_white_path_to x) s"
by (simp add: has_white_path_to_def)
lemma has_white_path_toE[elim!]:
"\<lbrakk> (x points_to y) s; white y s \<rbrakk> \<Longrightarrow> (x has_white_path_to y) s"
"\<lbrakk> (x has_white_path_to y) s; (y points_to z) s; white z s \<rbrakk> \<Longrightarrow> (x has_white_path_to z) s"
by (auto simp: has_white_path_to_def
elim: rtranclp.intros(2))
lemma has_white_path_to_reaches[elim]:
"(x has_white_path_to y) s \<Longrightarrow> (x reaches y) s"
unfolding has_white_path_to_def
by (induct rule: rtranclp.induct)
(auto intro: rtranclp.intros(2))
lemma has_white_path_to_blacken[simp]:
"(x has_white_path_to w) (s(gc := s gc\<lparr> W := gc_W s - rs \<rparr>)) \<longleftrightarrow> (x has_white_path_to w) s"
by (simp add: has_white_path_to_def)
text\<open>WL\<close>
lemma WL_mo_co_mark[simp]:
"ghost_honorary_grey (s p) = {}
\<Longrightarrow> WL p' (s(p := s p\<lparr> ghost_honorary_grey := rs \<rparr>)) = WL p' s \<union> { r |r. p' = p \<and> r \<in> rs}"
by (auto simp: WL_def)
lemma WL_blacken[simp]:
"gc_ghost_honorary_grey s = {}
\<Longrightarrow> WL p (s(gc := s gc\<lparr> W := gc_W s - rs \<rparr>)) = WL p s - { r |r. p = gc \<and> r \<in> rs }"
by (auto simp: WL_def)
lemma WL_hs_done[simp]:
"ghost_honorary_grey (s (mutator m)) = {}
\<Longrightarrow> WL p (s(mutator m := s (mutator m)\<lparr> W := {}, ghost_handshake_phase := hp' \<rparr>,
sys := s sys\<lparr> handshake_pending := hsp', W := sys_W s \<union> W (s (mutator m)),
ghost_handshake_in_sync := in' \<rparr>))
= (case p of gc \<Rightarrow> WL gc s | mutator m' \<Rightarrow> (if m' = m then {} else WL (mutator m') s) | sys \<Rightarrow> WL sys s \<union> WL (mutator m) s)"
"ghost_honorary_grey (s (mutator m)) = {}
\<Longrightarrow> WL p (s(mutator m := s (mutator m)\<lparr> W := {} \<rparr>,
sys := s sys\<lparr> handshake_pending := hsp', W := sys_W s \<union> W (s (mutator m)),
ghost_handshake_in_sync := in' \<rparr>))
= (case p of gc \<Rightarrow> WL gc s | mutator m' \<Rightarrow> (if m' = m then {} else WL (mutator m') s) | sys \<Rightarrow> WL sys s \<union> WL (mutator m) s)"
by (auto simp: WL_def split: process_name.splits)
lemma colours_load_W[iff]:
"gc_W s = {} \<Longrightarrow> black r (s(gc := (s gc)\<lparr>W := W (s sys)\<rparr>, sys := (s sys)\<lparr>W := {}\<rparr>)) \<longleftrightarrow> black r s"
"gc_W s = {} \<Longrightarrow> grey r (s(gc := (s gc)\<lparr>W := W (s sys)\<rparr>, sys := (s sys)\<lparr>W := {}\<rparr>)) \<longleftrightarrow> grey r s"
apply (simp_all add: black_def grey_def WL_def)
apply safe
apply (case_tac [!] x)
apply blast+
done
lemma colours_sweep_loop_free[iff]:
"black r (s(sys := s sys\<lparr>heap := (heap (s sys))(r' := None)\<rparr>)) \<longleftrightarrow> (black r s \<and> r \<noteq> r')"
"grey r (s(sys := s sys\<lparr>heap := (heap (s sys))(r' := None)\<rparr>)) \<longleftrightarrow> (grey r s)"
"white r (s(sys := s sys\<lparr>heap := (heap (s sys))(r' := None)\<rparr>)) \<longleftrightarrow> (white r s \<and> r \<noteq> r')"
by (auto simp: black_def grey_def split: obj_at_splits)
lemma colours_get_work_done[simp]:
"black r (s(mutator m := (s (mutator m))\<lparr>W := {}\<rparr>,
sys := (s sys)\<lparr> handshake_pending := hp', W := W (s sys) \<union> W (s (mutator m)),
ghost_handshake_in_sync := his' \<rparr>)) \<longleftrightarrow> black r s"
"grey r (s(mutator m := (s (mutator m))\<lparr>W := {}\<rparr>,
sys := (s sys)\<lparr> handshake_pending := hp', W := W (s sys) \<union> W (s (mutator m)),
ghost_handshake_in_sync := his' \<rparr>)) \<longleftrightarrow> grey r s"
"white r (s(mutator m := (s (mutator m))\<lparr>W := {}\<rparr>,
sys := (s sys)\<lparr> handshake_pending := hp', W := W (s sys) \<union> W (s (mutator m)),
ghost_handshake_in_sync := his' \<rparr>)) \<longleftrightarrow> white r s"
apply (simp_all add: black_def grey_def WL_def split: obj_at_splits)
apply auto
apply (metis (full_types) process_name.distinct(3))
by metis
lemma colours_get_roots_done[simp]:
"black r (s(mutator m := (s (mutator m))\<lparr> W := {}, ghost_handshake_phase := hs' \<rparr>,
sys := (s sys)\<lparr> handshake_pending := hp', W := W (s sys) \<union> W (s (mutator m)),
ghost_handshake_in_sync := his' \<rparr>)) \<longleftrightarrow> black r s"
"grey r (s(mutator m := (s (mutator m))\<lparr> W := {}, ghost_handshake_phase := hs' \<rparr>,
sys := (s sys)\<lparr> handshake_pending := hp', W := W (s sys) \<union> W (s (mutator m)),
ghost_handshake_in_sync := his' \<rparr>)) \<longleftrightarrow> grey r s"
"white r (s(mutator m := (s (mutator m))\<lparr> W := {}, ghost_handshake_phase := hs' \<rparr>,
sys := (s sys)\<lparr> handshake_pending := hp', W := W (s sys) \<union> W (s (mutator m)),
ghost_handshake_in_sync := his' \<rparr>)) \<longleftrightarrow> white r s"
apply (simp_all add: black_def grey_def WL_def split: obj_at_splits)
apply auto
apply (metis process_name.distinct(3))
by metis
lemma colours_mark[simp]:
"\<lbrakk> ghost_honorary_grey (s p) = {} \<rbrakk> \<Longrightarrow> black b (s(p := s p\<lparr>ghost_honorary_grey := {r}\<rparr>)) \<longleftrightarrow> black b s \<and> b \<noteq> r"
"\<lbrakk> ghost_honorary_grey (s p) = {} \<rbrakk> \<Longrightarrow> grey g (s(p := (s p)\<lparr>ghost_honorary_grey := {r}\<rparr>)) \<longleftrightarrow> grey g s \<or> g = r"
"\<lbrakk> ghost_honorary_grey (s p) = {} \<rbrakk> \<Longrightarrow> white w (s(p := s p\<lparr>ghost_honorary_grey := {r}\<rparr>)) \<longleftrightarrow> white w s"
by (auto simp: black_def grey_def)
lemma colours_flip_fM[simp]:
"fl \<noteq> sys_fM s \<Longrightarrow> black b (s(sys := (s sys)\<lparr>fM := fl, mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>)) \<longleftrightarrow> white b s \<and> \<not>grey b s"
by (simp_all add: black_def)
lemma colours_alloc[simp]:
"heap (s sys) r' = None
\<Longrightarrow> black r (s(mutator m := (s (mutator m))\<lparr> roots := roots' \<rparr>, sys := (s sys)\<lparr>heap := heap (s sys)(r' \<mapsto> \<lparr>obj_mark = fl, obj_fields = Map.empty\<rparr>)\<rparr>))
\<longleftrightarrow> black r s \<or> (r' = r \<and> fl = sys_fM s \<and> \<not>grey r' s)"
"grey r (s(mutator m := (s (mutator m))\<lparr> roots := roots' \<rparr>, sys := (s sys)\<lparr>heap := heap (s sys)(r' \<mapsto> \<lparr>obj_mark = fl, obj_fields = Map.empty\<rparr>)\<rparr>))
\<longleftrightarrow> grey r s"
"heap (s sys) r' = None
\<Longrightarrow> white r (s(mutator m := (s (mutator m))\<lparr> roots := roots' \<rparr>, sys := (s sys)\<lparr>heap := heap (s sys)(r' \<mapsto> \<lparr>obj_mark = fl, obj_fields = Map.empty\<rparr>)\<rparr>))
\<longleftrightarrow> white r s \<or> (r' = r \<and> fl \<noteq> sys_fM s)"
by (auto simp: black_def split: obj_at_splits dest!: valid_refs_invD)
lemma colours_blacken[simp]:
"valid_W_inv s \<Longrightarrow> black b (s(gc := s gc\<lparr>W := gc_W s - {r}\<rparr>)) \<longleftrightarrow> black b s \<or> (r \<in> gc_W s \<and> b = r)"
"\<lbrakk> r \<in> gc_W s; valid_W_inv s \<rbrakk> \<Longrightarrow> grey g (s(gc := s gc\<lparr>W := gc_W s - {r}\<rparr>)) \<longleftrightarrow> (grey g s \<and> g \<noteq> r)"
"white w (s(gc := s gc\<lparr>W := gc_W s - {r}\<rparr>)) \<longleftrightarrow> white w s"
apply (auto simp: black_def grey_def WL_def split: obj_at_splits)
apply metis+
done
lemma colours_dequeue[simp]:
"black b (s(sys := (s sys)\<lparr> heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\<lambda>_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws) \<rparr>))
\<longleftrightarrow> (black b s \<and> b \<noteq> r) \<or> (b = r \<and> fl = sys_fM s \<and> valid_ref r s \<and> \<not>grey r s)"
by (auto simp: black_def split: obj_at_splits)
lemma W_load_W[simp]:
"gc_W s = {} \<Longrightarrow> (\<Union>p. W (if p = sys then (s sys)\<lparr>W := {}\<rparr> else if p = gc then (s gc)\<lparr>W := sys_W s\<rparr> else s p)) = (\<Union>p. W (s p))"
apply (rule equalityI)
apply (fastforce split: if_splits)
apply clarsimp
apply (rename_tac x p)
apply (case_tac p)
apply force+
done
lemma WL_load_W[simp]:
"gc_W s = {}
\<Longrightarrow> (WL p (s(gc := (s gc)\<lparr>W := sys_W s\<rparr>, sys := (s sys)\<lparr>W := {}\<rparr>)))
= (case p of gc \<Rightarrow> WL gc s \<union> sys_W s | mutator m \<Rightarrow> WL (mutator m) s | sys \<Rightarrow> sys_ghost_honorary_grey s)"
by (auto simp: WL_def split: process_name.splits)
lemma colours_mo_co_W[simp]:
"ghost_honorary_grey (s p') = {r}
\<Longrightarrow> (WL p (s(p' := (s p')\<lparr>W := insert r (W (s p')), ghost_honorary_grey := {}\<rparr>))) = (WL p s)"
"ghost_honorary_grey (s p') = {r}
\<Longrightarrow> grey g (s(p' := (s p')\<lparr>W := insert r (W (s p')), ghost_honorary_grey := {}\<rparr>)) \<longleftrightarrow> grey g s"
by (force simp: grey_def WL_def split: process_name.splits if_splits)+
lemma no_grey_refs_eq_imp:
"eq_imp (\<lambda>(_::unit). (\<lambda>s. \<Union>p. WL p s))
no_grey_refs"
by (auto simp add: eq_imp_def grey_def no_grey_refs_def set_eq_iff)
lemmas no_grey_refs_fun_upd[simp] = eq_imp_fun_upd[OF no_grey_refs_eq_imp, simplified eq_imp_simps, rule_format]
lemma no_grey_refs_no_pending_marks:
"\<lbrakk> no_grey_refs s; valid_W_inv s \<rbrakk> \<Longrightarrow> tso_no_pending_marks s"
by (auto intro!: filter_False simp: no_grey_refs_def)
lemma (in mut_m) reachable_blackD:
"\<lbrakk> no_grey_refs s; reachable_snapshot_inv s; reachable r s \<rbrakk> \<Longrightarrow> black r s"
by (simp add: no_grey_refs_def reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def)
lemma (in mut_m) no_grey_refs_not_reachable:
"\<lbrakk> no_grey_refs s; reachable_snapshot_inv s; white r s \<rbrakk> \<Longrightarrow> \<not>reachable r s"
by (fastforce simp: no_grey_refs_def reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def
split: obj_at_splits)
lemma (in mut_m) no_grey_refs_not_rootD:
"\<lbrakk> no_grey_refs s; reachable_snapshot_inv s; white r s \<rbrakk>
\<Longrightarrow> r \<notin> mut_roots s \<and> r \<notin> mut_ghost_honorary_root s \<and> r \<notin> tso_write_refs s"
apply (drule (2) no_grey_refs_not_reachable)
apply (force simp: reachable_def)
done
lemma no_grey_refs_not_grey_reachableD:
"no_grey_refs s \<Longrightarrow> \<not>grey_reachable x s"
by (clarsimp simp: no_grey_refs_def grey_reachable_def)
lemma (in mut_m) reachable_snapshot_inv_white_root:
"\<lbrakk> white w s; w \<in> mut_roots s \<or> w \<in> mut_ghost_honorary_root s; reachable_snapshot_inv s \<rbrakk> \<Longrightarrow> \<exists>g. (g grey_protects_white w) s"
by (auto simp: reachable_snapshot_inv_def in_snapshot_def
reachable_def grey_protects_white_def)
lemma no_grey_refsD:
"no_grey_refs s \<Longrightarrow> r \<notin> W (s p)"
"no_grey_refs s \<Longrightarrow> r \<notin> WL p s"
"no_grey_refs s \<Longrightarrow> r \<notin> ghost_honorary_grey (s p)"
by (auto simp: no_grey_refs_def)
lemma no_grey_refs_marked[dest]:
"\<lbrakk> marked r s; no_grey_refs s \<rbrakk> \<Longrightarrow> black r s"
by (auto simp: no_grey_refs_def black_def)
lemma no_grey_refs_bwD[dest]:
"\<lbrakk> heap (s sys) r = Some obj; no_grey_refs s \<rbrakk> \<Longrightarrow> black r s \<or> white r s"
by (clarsimp simp: black_def grey_def no_grey_refs_def split: obj_at_splits)
text\<open>tso write refs\<close>
lemma tso_write_refs_simps[simp]:
"mut_m.tso_write_refs m (s(mutator m' := s (mutator m')\<lparr>roots := roots'\<rparr>))
= mut_m.tso_write_refs m s"
"mut_m.tso_write_refs m (s(mutator m' := s (mutator m')\<lparr>ghost_honorary_root := {}\<rparr>,
sys := s sys\<lparr>mem_write_buffers := (mem_write_buffers (s sys))(mutator m' := sys_mem_write_buffers (mutator m') s @ [mw_Mutate r f opt_r'])\<rparr>))
= mut_m.tso_write_refs m s \<union> (if m' = m then write_refs (mw_Mutate r f opt_r') else {})"
"mut_m.tso_write_refs m (s(sys := s sys\<lparr>heap := (sys_heap s)(r' := None)\<rparr>))
= mut_m.tso_write_refs m s"
"mut_m.tso_write_refs m (s(mutator m' := s (mutator m')\<lparr>roots := insert r (roots (s (mutator m')))\<rparr>, sys := s sys\<lparr>heap := sys_heap s(r \<mapsto> obj)\<rparr>))
= mut_m.tso_write_refs m s"
"mut_m.tso_write_refs m (s(mutator m' := s (mutator m')\<lparr>ghost_honorary_root := Option.set_option opt_r', ref := opt_r'\<rparr>))
= mut_m.tso_write_refs m s"
"mut_m.tso_write_refs m (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (\<lambda>obj. obj\<lparr>obj_fields := (obj_fields obj)(f := opt_r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))
= (if p = mutator m then \<Union>w \<in> set ws. write_refs w else mut_m.tso_write_refs m s)"
"sys_mem_write_buffers p s = mw_Mark r fl # ws
\<Longrightarrow> mut_m.tso_write_refs m (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\<lambda>_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))
= mut_m.tso_write_refs m s"
by (auto simp: mut_m.tso_write_refs_def)
lemma mutator_reachable_tso:
"sys_mem_write_buffers (mutator m) s = mw_Mutate r f opt_r' # ws
\<Longrightarrow> mut_m.reachable m r s \<and> (\<forall>r'. opt_r' = Some r' \<longrightarrow> mut_m.reachable m r' s)"
by (auto simp: mut_m.tso_write_refs_def)
text\<open>coloured heaps\<close>
lemma black_heap_eq_imp:
"eq_imp (\<lambda>(_::unit). (\<lambda>s. \<Union>p. WL p s) \<^bold>\<otimes> sys_fM \<^bold>\<otimes> (\<lambda>s. Option.map_option obj_mark \<circ> sys_heap s))
black_heap"
apply (clarsimp simp add: eq_imp_def black_heap_def black_def grey_def all_conj_distrib fun_eq_iff split: option.splits)
apply (rename_tac s s')
apply (subgoal_tac "\<forall>x. marked x s \<longleftrightarrow> marked x s'")
apply (subgoal_tac "\<forall>x. valid_ref x s \<longleftrightarrow> valid_ref x s'")
apply (subgoal_tac "\<forall>x. (\<forall>p. x \<notin> WL p s) \<longleftrightarrow> (\<forall>p. x \<notin> WL p s')")
apply clarsimp
apply (auto simp: set_eq_iff)[1]
apply clarsimp
apply (rename_tac x)
apply (rule eq_impD[OF obj_at_eq_imp])
apply (drule_tac x=x in spec)
apply (drule_tac f="Option.map_option \<langle>True\<rangle>" in arg_cong)
apply (clarsimp simp: map_option.compositionality o_def)
apply (clarsimp simp: map_option.compositionality o_def)
apply (rule eq_impD[OF obj_at_eq_imp])
apply clarsimp
apply (rename_tac x)
apply (drule_tac x=x in spec)
apply (drule_tac f="Option.map_option (\<lambda>fl. fl = sys_fM s)" in arg_cong)
apply (simp add: map_option.compositionality o_def)
done
lemma white_heap_eq_imp:
"eq_imp (\<lambda>(_::unit). sys_fM \<^bold>\<otimes> (\<lambda>s. Option.map_option obj_mark \<circ> sys_heap s))
white_heap"
apply (clarsimp simp: all_conj_distrib eq_imp_def white_heap_def obj_at_def fun_eq_iff
split: option.splits)
apply (rule iffI)
apply (metis (hide_lams, no_types) map_option_eq_Some)+
done
lemma no_black_refs_eq_imp:
"eq_imp (\<lambda>(_::unit). (\<lambda>s. (\<Union>p. WL p s)) \<^bold>\<otimes> sys_fM \<^bold>\<otimes> (\<lambda>s. Option.map_option obj_mark \<circ> sys_heap s))
no_black_refs"
apply (clarsimp simp add: eq_imp_def no_black_refs_def black_def grey_def all_conj_distrib fun_eq_iff set_eq_iff split: option.splits)
apply (rename_tac s s')
apply (subgoal_tac "\<forall>x. marked x s \<longleftrightarrow> marked x s'")
apply clarsimp
apply clarsimp
apply (rename_tac x)
apply ( (drule_tac x=x in spec)+ )[1]
apply (auto split: obj_at_splits)
done
lemmas black_heap_fun_upd[simp] = eq_imp_fun_upd[OF black_heap_eq_imp, simplified eq_imp_simps, rule_format]
lemmas white_heap_fun_upd[simp] = eq_imp_fun_upd[OF white_heap_eq_imp, simplified eq_imp_simps, rule_format]
lemmas no_black_refs_fun_upd[simp] = eq_imp_fun_upd[OF no_black_refs_eq_imp, simplified eq_imp_simps, rule_format]
lemma white_heap_imp_no_black_refs[elim!]:
"white_heap s \<Longrightarrow> no_black_refs s"
apply (clarsimp simp: white_heap_def no_black_refs_def black_def)
apply (rename_tac x)
apply (drule_tac x=x in spec)
apply (clarsimp split: obj_at_splits)
done
lemma (in sys) no_black_refs_dequeue[simp]:
"\<lbrakk> sys_mem_write_buffers p s = mw_Mark r fl # ws; no_black_refs s; valid_W_inv s \<rbrakk>
\<Longrightarrow> no_black_refs (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\<lambda>_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))"
"\<lbrakk> sys_mem_write_buffers p s = mw_Mutate r f r' # ws; no_black_refs s \<rbrakk>
\<Longrightarrow> no_black_refs (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (\<lambda>obj. obj\<lparr>obj_fields := (obj_fields obj)(f := r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))"
by (auto simp: no_black_refs_def map_option.compositionality o_def)
lemma black_heap_no_greys[elim]:
"\<lbrakk> black_heap s; valid_refs_inv s \<rbrakk> \<Longrightarrow> no_grey_refs s"
"\<lbrakk> no_grey_refs s; \<forall>r. marked r s \<or> \<not>valid_ref r s \<rbrakk> \<Longrightarrow> black_heap s"
by (auto simp: black_def black_heap_def no_grey_refs_def dest: valid_refs_invD)
lemma heap_colours_alloc[simp]:
"\<lbrakk> heap (s sys) r' = None; valid_refs_inv s \<rbrakk>
\<Longrightarrow> black_heap (s(mutator m := s (mutator m)\<lparr>roots := insert r' (roots (s (mutator m)))\<rparr>, sys := s sys\<lparr>heap := sys_heap s(r' \<mapsto> \<lparr>obj_mark = fl, obj_fields = Map.empty\<rparr>)\<rparr>))
\<longleftrightarrow> black_heap s \<and> fl = sys_fM s"
"heap (s sys) r' = None
\<Longrightarrow> white_heap (s(mutator m := s (mutator m)\<lparr>roots := insert r' (roots (s (mutator m)))\<rparr>, sys := s sys\<lparr>heap := sys_heap s(r' \<mapsto> \<lparr>obj_mark = fl, obj_fields = Map.empty\<rparr>)\<rparr>))
\<longleftrightarrow> white_heap s \<and> fl \<noteq> sys_fM s"
apply (simp_all add: black_heap_def white_heap_def split: obj_at_splits)
apply (rule iffI)
apply (intro allI conjI impI)
apply (rename_tac x)
apply (drule_tac x=x in spec)
apply clarsimp
apply (drule spec[where x=r'], auto dest!: valid_refs_invD split: obj_at_splits)[2]
apply (rule iffI)
apply (intro allI conjI impI)
apply (rename_tac x obj)
apply (drule_tac x=x in spec)
apply clarsimp
apply (drule spec[where x=r'], auto dest!: valid_refs_invD split: obj_at_splits)[2]
done
lemma heap_colours_colours:
"black_heap s \<Longrightarrow> \<not>white r s"
"white_heap s \<Longrightarrow> \<not>black r s"
by (auto simp: black_heap_def white_heap_def
dest!: spec[where x=r]
split: obj_at_splits)
lemma heap_colours_marked:
"\<lbrakk> black_heap s; obj_at P r s \<rbrakk> \<Longrightarrow> marked r s"
"\<lbrakk> white_heap s; obj_at P r s \<rbrakk> \<Longrightarrow> white r s"
by (auto simp: black_heap_def white_heap_def
dest!: spec[where x=r]
split: obj_at_splits)
lemma (in gc) obj_fields_marked_inv_blacken:
"\<lbrakk> gc_field_set s = {}; obj_fields_marked_inv s; (gc_tmp_ref s points_to w) s; white w s; tso_writes_inv s \<rbrakk> \<Longrightarrow> False"
by (simp add: obj_fields_marked_inv_def obj_at_field_on_heap_def ran_def split: option.splits obj_at_splits)
lemma (in gc) obj_fields_marked_inv_has_white_path_to_blacken:
"\<lbrakk> gc_field_set s = {}; gc_tmp_ref s \<in> gc_W s; (gc_tmp_ref s has_white_path_to w) s; obj_fields_marked_inv s; valid_W_inv s; tso_writes_inv s \<rbrakk> \<Longrightarrow> w = gc_tmp_ref s"
apply (drule (1) valid_W_invD)
apply (clarsimp simp: has_white_path_to_def)
apply (erule converse_rtranclpE)
apply (clarsimp split: obj_at_splits)
apply clarsimp
apply (simp add: obj_fields_marked_inv_def obj_at_field_on_heap_def ran_def split: option.splits obj_at_splits)
done
lemma fold_writes_points_to[rule_format, simplified conj_explode]:
"heap (fold_writes ws (s sys)) r = Some obj \<and> obj_fields obj f = Some r'
\<longrightarrow> (r points_to r') s \<or> (\<exists>w \<in> set ws. r' \<in> write_refs w)" (is "?P (fold_writes ws) obj")
unfolding fold_writes_def
apply (rule spec[OF fold_invariant[where P="\<lambda>fr. \<forall>obj. ?P fr obj" and Q="\<lambda>w. w \<in> set ws"]])
apply simp
apply (fastforce simp: ran_def split: obj_at_splits)
apply clarsimp
apply (drule (1) bspec)
apply (clarsimp split: mem_write_action.split_asm if_splits)
done
lemma (in mut_m) reachable_load[simp]:
assumes "sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref r'"
assumes "r \<in> mut_roots s"
shows "mut_m.reachable m' y (s(mutator m := s (mutator m)\<lparr> roots := mut_roots s \<union> Option.set_option r' \<rparr>)) \<longleftrightarrow> mut_m.reachable m' y s" (is "?lhs = ?rhs")
proof(cases "m' = m")
case True show ?thesis
proof(rule iffI)
assume ?lhs with assms True show ?rhs
apply (clarsimp simp: sys_read_def)
apply (clarsimp simp: reachable_def tso_write_refs_def)
apply (clarsimp simp: sys_read_def fold_writes_def)
apply (elim disjE)
apply blast
defer
apply blast
apply blast
apply (fold fold_writes_def)
apply clarsimp
apply (drule (1) fold_writes_points_to)
apply (erule disjE)
apply (fastforce elim!: converse_rtranclp_into_rtranclp[rotated] split: obj_at_splits intro!: ranI)
apply clarsimp
apply (case_tac w, simp_all)
apply (erule disjE)
apply (rule_tac x=x in exI)
apply (fastforce elim!: converse_rtranclp_into_rtranclp[rotated] split: obj_at_splits intro!: ranI)
apply (rule_tac x=x in exI)
apply (fastforce elim!: converse_rtranclp_into_rtranclp[rotated] split: obj_at_splits intro!: ranI)
done (* filthy *)
next
assume ?rhs with True show ?lhs by (fastforce simp: reachable_def)
qed
qed simp
lemma (in mut_m) reachable_deref_del[simp]:
"\<lbrakk> sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref opt_r'; r \<in> mut_roots s; mut_ghost_honorary_root s = {} \<rbrakk>
\<Longrightarrow> mut_m.reachable m' y (s(mutator m := s (mutator m)\<lparr> ghost_honorary_root := Option.set_option opt_r', ref := opt_r' \<rparr>))
\<longleftrightarrow> mut_m.reachable m' y s"
apply (clarsimp simp: mut_m.reachable_def sys_read_def)
apply (rule iffI)
apply clarsimp
apply (elim disjE)
apply metis
apply (erule option_bind_invE; auto dest!: fold_writes_points_to)
apply (auto elim!: converse_rtranclp_into_rtranclp[rotated]
simp: tso_write_refs_def)
done
text\<open>strong tricolour\<close>
lemma strong_tricolour_invD:
"\<lbrakk> black x s; (x points_to y) s; valid_ref y s; strong_tricolour_inv s \<rbrakk>
\<Longrightarrow> marked y s"
apply (clarsimp simp: strong_tricolour_inv_def)
apply (drule spec[where x=x])
apply (auto split: obj_at_splits)
done
text\<open>grey protects white\<close>
lemma grey_protects_whiteD[dest]:
"(g grey_protects_white w) s \<Longrightarrow> grey g s \<and> (g = w \<or> white w s)"
by (auto simp: grey_protects_white_def)
lemma grey_protects_whiteI[iff]:
"grey g s \<Longrightarrow> (g grey_protects_white g) s"
by (simp add: grey_protects_white_def)
lemma grey_protects_whiteE[elim!]:
"\<lbrakk> (g points_to w) s; grey g s; white w s \<rbrakk> \<Longrightarrow> (g grey_protects_white w) s"
"\<lbrakk> (g grey_protects_white y) s; (y points_to w) s; white w s \<rbrakk> \<Longrightarrow> (g grey_protects_white w) s"
by (auto simp: grey_protects_white_def)
lemma grey_protects_white_reaches[elim]:
"(g grey_protects_white w) s \<Longrightarrow> (g reaches w) s"
by (auto simp: grey_protects_white_def)
lemma grey_protects_white_hs_done[simp]:
"(g grey_protects_white w) (s(mutator m := s (mutator m)\<lparr> W := {}, ghost_handshake_phase := hs' \<rparr>,
sys := s sys\<lparr> handshake_pending := hp', W := sys_W s \<union> W (s (mutator m)),
ghost_handshake_in_sync := his' \<rparr>))
\<longleftrightarrow> (g grey_protects_white w) s"
by (simp add: grey_protects_white_def)
lemma grey_protects_white_alloc[simp]:
"\<lbrakk> fl = sys_fM s; sys_heap s r = None \<rbrakk>
\<Longrightarrow> (g grey_protects_white w) (s(mutator m := s (mutator m)\<lparr>roots := insert r (roots (s (mutator m)))\<rparr>, sys := s sys\<lparr>heap := sys_heap s(r \<mapsto> \<lparr>obj_mark = fl, obj_fields = Map.empty\<rparr>)\<rparr>))
\<longleftrightarrow> (g grey_protects_white w) s"
by (clarsimp simp: grey_protects_white_def has_white_path_to_def)
text\<open>reachable snapshot inv\<close>
lemma (in mut_m) reachable_snapshot_invI[intro]:
"(\<And>y. reachable y s \<Longrightarrow> in_snapshot y s) \<Longrightarrow> reachable_snapshot_inv s"
by (simp add: reachable_snapshot_inv_def)
lemma (in mut_m) reachable_snapshot_inv_eq_imp:
"eq_imp (\<lambda>r. mut_roots \<^bold>\<otimes> mut_ghost_honorary_root \<^bold>\<otimes> (\<lambda>s. r \<in> (\<Union>p. WL p s)) \<^bold>\<otimes> sys_fM \<^bold>\<otimes> sys_heap \<^bold>\<otimes> tso_pending_mutate (mutator m))
reachable_snapshot_inv"
apply (clarsimp simp: eq_imp_def mut_m.reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def)
apply (rename_tac s s')
apply (clarsimp simp: black_def grey_def obj_at_def cong: option.case_cong)
apply (subst eq_impD[OF has_white_path_to_eq_imp])
defer
apply (subst eq_impD[OF reachable_eq_imp])
defer
apply (subgoal_tac "\<forall>x. (\<forall>p. x \<notin> WL p s) \<longleftrightarrow> (\<forall>p. x \<notin> WL p s')")
apply force
apply blast
apply simp
apply simp
done
lemmas reachable_snapshot_fun_upd[simp] = eq_imp_fun_upd[OF mut_m.reachable_snapshot_inv_eq_imp, simplified eq_imp_simps, rule_format]
lemma in_snapshotI[intro]:
"black r s \<Longrightarrow> in_snapshot r s"
"grey r s \<Longrightarrow> in_snapshot r s"
"\<lbrakk> white w s; (g grey_protects_white w) s \<rbrakk> \<Longrightarrow> in_snapshot w s"
by (auto simp: in_snapshot_def)
lemma in_snapshot_colours:
"in_snapshot r s \<Longrightarrow> black r s \<or> grey r s \<or> white r s"
by (auto simp: in_snapshot_def)
lemma in_snapshot_valid_ref:
"\<lbrakk> in_snapshot r s; valid_refs_inv s \<rbrakk> \<Longrightarrow> valid_ref r s"
by (fastforce dest: in_snapshot_colours)
lemma (in mut_m) reachable_snapshot_inv_sweep_loop_free:
fixes s :: "('field, 'mut, 'ref) lsts"
assumes nmr: "white r s"
assumes ngs: "no_grey_refs s"
assumes rsi: "reachable_snapshot_inv s"
shows "reachable_snapshot_inv (s(sys := (s sys)\<lparr>heap := (heap (s sys))(r := None)\<rparr>))" (is "reachable_snapshot_inv ?s'")
proof
fix y :: "'ref"
assume rx: "reachable y ?s'"
then have "black y s \<and> y \<noteq> r"
proof(induct rule: reachable_induct)
case (root x) with ngs nmr rsi show ?case
by (auto dest: reachable_blackD)
next
case (ghost_honorary_root x) with ngs nmr rsi show ?case
apply -
apply (frule (1) reachable_blackD[where r=x])
apply (auto simp: reachable_def)
done
next
case (tso_root x) with ngs nmr rsi show ?case
apply -
apply (frule (1) reachable_blackD[where r=x])
apply (auto simp: reachable_def tso_write_refs_def)
done
next
case (reaches x y) with ngs nmr rsi show ?case
apply (clarsimp simp: reachable_def)
apply (drule predicate2D[OF rtranclp_mono[where s="\<lambda>x y. (x points_to y) s", OF predicate2I], rotated])
apply (clarsimp split: obj_at_splits if_splits)
apply (rule conjI)
apply (rule reachable_blackD, assumption, assumption)
apply (simp add: reachable_def)
apply (blast intro: rtranclp.intros(2))
apply clarsimp
apply (frule (1) reachable_blackD[where r=r])
apply (simp add: reachable_def)
apply (blast intro: rtranclp.intros(2))
apply auto
done
qed
then show "in_snapshot y ?s'"
unfolding in_snapshot_def by simp
qed
lemma reachableI2[intro]:
"x \<in> mut_m.mut_ghost_honorary_root m s \<Longrightarrow> mut_m.reachable m x s"
by (auto simp: mut_m.reachable_def)
lemma reachable_alloc[simp]:
assumes rn: "sys_heap s r = None"
shows "mut_m.reachable m r' (s(mutator m' := (s (mutator m'))\<lparr>roots := insert r (roots (s (mutator m')))\<rparr>, sys := (s sys)\<lparr>heap := (sys_heap s)(r \<mapsto> \<lparr>obj_mark = fl, obj_fields = Map.empty\<rparr>)\<rparr>))
\<longleftrightarrow> mut_m.reachable m r' s \<or> (m' = m \<and> r' = r)" (is "?lhs \<longleftrightarrow> ?rhs")
proof(rule iffI)
assume ?lhs from this assms show ?rhs
proof(induct rule: reachable_induct)
case (reaches x y) then show ?case by clarsimp (fastforce simp: mut_m.reachable_def elim: rtranclp.intros(2) split: obj_at_splits)
qed (auto split: if_splits)
next
assume ?rhs then show ?lhs
proof(rule disjE)
assume "mut_m.reachable m r' s" then show ?thesis
proof(induct rule: reachable_induct)
case (tso_root x) then show ?case
by (fastforce simp: mut_m.reachable_def simp del: fun_upd_apply)
next
case (reaches x y) with rn show ?case
by (fastforce simp: mut_m.reachable_def simp del: fun_upd_apply elim: rtranclp.intros(2))
qed auto
next
assume "m' = m \<and> r' = r" with rn show ?thesis
by (fastforce simp: mut_m.reachable_def)
qed
qed
lemma (in mut_m) reachable_snapshot_inv_alloc[simp]:
fixes s :: "('field, 'mut, 'ref) lsts"
assumes rn: "sys_heap s r = None"
assumes fl: "fl = sys_fM s"
assumes vri: "valid_refs_inv s"
assumes rsi: "reachable_snapshot_inv s"
shows "reachable_snapshot_inv (s(mutator m' := (s (mutator m'))\<lparr>roots := insert r (roots (s (mutator m')))\<rparr>, sys := (s sys)\<lparr>heap := (sys_heap s)(r \<mapsto> \<lparr>obj_mark = fl, obj_fields = Map.empty\<rparr>)\<rparr>))" (is "reachable_snapshot_inv ?s'")
using assms unfolding reachable_snapshot_inv_def in_snapshot_def
by (auto simp del: reachable_fun_upd)
lemma (in mut_m) reachable_snapshot_inv_discard_roots[simp]:
"\<lbrakk> reachable_snapshot_inv s; roots' \<subseteq> roots (s (mutator m)) \<rbrakk>
\<Longrightarrow> reachable_snapshot_inv (s(mutator m := (s (mutator m))\<lparr>roots := roots'\<rparr>))"
unfolding reachable_snapshot_inv_def reachable_def in_snapshot_def grey_protects_white_def by auto
lemma grey_protects_white_mark[simp]:
assumes ghg: "ghost_honorary_grey (s p) = {}"
shows "(\<exists>g. (g grey_protects_white w) (s(p := s p\<lparr> ghost_honorary_grey := {r} \<rparr>)))
\<longleftrightarrow> (\<exists>g'. (g' grey_protects_white w) s) \<or> (r has_white_path_to w) s" (is "?lhs \<longleftrightarrow> ?rhs")
proof
assume ?lhs
then obtain g where "(g grey_protects_white w) (s(p := s p\<lparr>ghost_honorary_grey := {r}\<rparr>))" by blast
from this ghg show ?rhs
apply (clarsimp simp: grey_protects_white_def has_white_path_to_def if_distrib cong: if_cong)
apply (rotate_tac 2)
apply (induct rule: rtranclp.induct)
apply (auto intro: rtranclp.intros(2))
done
next
assume ?rhs then show ?lhs
proof(safe)
fix g assume "(g grey_protects_white w) s"
with ghg show ?thesis
apply (clarsimp simp: grey_protects_white_def has_white_path_to_def)
apply (rotate_tac 2)
apply (induct rule: rtranclp.induct)
apply (auto intro: rtranclp.intros(2))
done
next
assume "(r has_white_path_to w) s" with ghg show ?thesis
by (auto simp: grey_protects_white_def has_white_path_to_def)
qed
qed
(*>*)
subsection\<open>Invariants\<close>
text (in mut_m) \<open>
We need phase invariants in terms of both @{const
"mut_ghost_handshake_phase"} and @{const "sys_ghost_handshake_phase"}
which respectively track what the mutators and GC know by virtue of
the synchronisation structure of the system.
Read the following as ``when mutator \<open>m\<close> is past the specified
handshake, and has yet to reach the next one, ... holds.''
\<close>
primrec (in mut_m) mutator_phase_inv_aux :: "handshake_phase \<Rightarrow> ('field, 'mut, 'ref) lsts_pred" where
"mutator_phase_inv_aux hp_Idle = \<langle>True\<rangle>"
| "mutator_phase_inv_aux hp_IdleInit = no_black_refs"
| "mutator_phase_inv_aux hp_InitMark = marked_insertions"
| "mutator_phase_inv_aux hp_Mark = (marked_insertions \<^bold>\<and> marked_deletions)"
| "mutator_phase_inv_aux hp_IdleMarkSweep = (marked_insertions \<^bold>\<and> marked_deletions \<^bold>\<and> reachable_snapshot_inv)"
abbreviation (in mut_m) mutator_phase_inv :: "('field, 'mut, 'ref) lsts_pred" where
"mutator_phase_inv s \<equiv> mutator_phase_inv_aux (mut_ghost_handshake_phase s) s"
abbreviation mutators_phase_inv :: "('field, 'mut, 'ref) lsts_pred" where
"mutators_phase_inv \<equiv> (\<^bold>\<forall>m. mut_m.mutator_phase_inv m)"
text\<open>
This is what the GC guarantees. Read this as ``when the GC is at or
past the specified handshake, ... holds.''
\<close>
primrec sys_phase_inv_aux :: "handshake_phase \<Rightarrow> ('field, 'mut, 'ref) lsts_pred" where
"sys_phase_inv_aux hp_Idle = ( (If sys_fA \<^bold>= sys_fM Then black_heap Else white_heap) \<^bold>\<and> no_grey_refs )"
| "sys_phase_inv_aux hp_IdleInit = no_black_refs"
| "sys_phase_inv_aux hp_InitMark = (sys_fA \<^bold>\<noteq> sys_fM \<^bold>\<longrightarrow> no_black_refs)"
| "sys_phase_inv_aux hp_Mark = \<langle>True\<rangle>"
| "sys_phase_inv_aux hp_IdleMarkSweep = ( (sys_phase \<^bold>= \<langle>ph_Idle\<rangle> \<^bold>\<or> tso_pending_write gc (mw_Phase ph_Idle)) \<^bold>\<longrightarrow> no_grey_refs )"
abbreviation sys_phase_inv :: "('field, 'mut, 'ref) lsts_pred" where
"sys_phase_inv s \<equiv> sys_phase_inv_aux (sys_ghost_handshake_phase s) s"
(*<*)
declare mut_m.mutator_phase_inv_aux.simps[simp]
case_of_simps mutator_phase_inv_aux_case: mut_m.mutator_phase_inv_aux.simps
case_of_simps sys_phase_inv_aux_case: sys_phase_inv_aux.simps
lemma tso_pending_mutate_cong:
"\<lbrakk> filter is_mw_Mutate (sys_mem_write_buffers p s) = filter is_mw_Mutate (sys_mem_write_buffers p s'); \<And>r f r'. P r f r' \<longleftrightarrow> Q r f r' \<rbrakk>
\<Longrightarrow> (\<forall>r f r'. mw_Mutate r f r' \<in> set (sys_mem_write_buffers p s) \<longrightarrow> P r f r') =
(\<forall>r f r'. mw_Mutate r f r' \<in> set (sys_mem_write_buffers p s') \<longrightarrow> Q r f r')"
apply (intro iff_allI)
apply (auto dest!: arg_cong[where f=set])
done
lemma (in mut_m) marked_insertions_eq_imp:
"eq_imp (\<lambda>(_::unit). sys_fM \<^bold>\<otimes> (\<lambda>s. Option.map_option obj_mark \<circ> sys_heap s) \<^bold>\<otimes> tso_pending_mutate (mutator m))
marked_insertions"
apply (clarsimp simp: eq_imp_def marked_insertions_def obj_at_def fun_eq_iff
split: mem_write_action.splits)
apply (erule tso_pending_mutate_cong)
apply (clarsimp split: option.splits obj_at_splits)
apply (rename_tac s s' opt x)
apply (drule_tac x=x in spec)
apply auto
done
lemmas marked_insertions_fun_upd[simp] = eq_imp_fun_upd[OF mut_m.marked_insertions_eq_imp, simplified eq_imp_simps, rule_format]
lemma marked_insertionD[elim!]:
"\<lbrakk> sys_mem_write_buffers (mutator m) s = mw_Mutate r f (Some r') # ws; mut_m.marked_insertions m s \<rbrakk>
\<Longrightarrow> marked r' s"
by (auto simp: mut_m.marked_insertions_def)
lemma (in mut_m) marked_insertions_store_buffer_empty[intro]:
"tso_pending_mutate (mutator m) s = [] \<Longrightarrow> marked_insertions s"
by (auto simp: marked_insertions_def filter_empty_conv
split: mem_write_action.splits option.splits)
lemma (in mut_m) marked_insertions_blacken[simp]:
"marked_insertions (s(gc := (s gc)\<lparr> W := gc_W s - {r} \<rparr>)) \<longleftrightarrow> marked_insertions s"
by (clarsimp simp: marked_insertions_def
split: mem_write_action.splits option.splits)
lemma (in mut_m) marked_insertions_alloc[simp]:
"\<lbrakk> heap (s sys) r' = None; valid_refs_inv s \<rbrakk>
\<Longrightarrow> marked_insertions (s(mutator m' := s (mutator m')\<lparr>roots := roots'\<rparr>, sys := s sys\<lparr>heap := sys_heap s(r' \<mapsto> obj')\<rparr>))
\<longleftrightarrow> marked_insertions s"
apply (clarsimp simp: mut_m.marked_insertions_def split: mem_write_action.splits option.splits)
apply (rule iffI)
apply clarsimp
apply (rename_tac ref field x)
apply (drule_tac x=ref in spec, drule_tac x=field in spec, drule_tac x=x in spec, clarsimp)
apply (drule valid_refs_invD(6)[where x=r' and y=r'], simp_all)
done
lemma marked_insertions_free[simp]:
"\<lbrakk> mut_m.marked_insertions m s; white r s \<rbrakk>
\<Longrightarrow> mut_m.marked_insertions m (s(sys := (s sys)\<lparr>heap := (heap (s sys))(r := None)\<rparr>))"
by (fastforce simp: mut_m.marked_insertions_def split: mem_write_action.splits obj_at_splits option.splits)
lemma (in mut_m) marked_deletions_eq_imp:
"eq_imp (\<lambda>(_::unit). sys_fM \<^bold>\<otimes> sys_heap \<^bold>\<otimes> tso_pending_mutate (mutator m))
marked_deletions"
apply (clarsimp simp: eq_imp_def fun_eq_iff marked_deletions_def obj_at_field_on_heap_def)
apply (rule iffI)
apply (clarsimp split: mem_write_action.splits)
apply (rename_tac s s' ref field option)
apply (drule_tac x="mw_Mutate ref field option" in spec)
apply (drule mp)
apply (metis (lifting, full_types) filter_set member_filter)
apply clarsimp
apply (subst eq_impD[OF obj_at_eq_imp])
prefer 2
apply (fastforce cong: option.case_cong)
apply clarsimp
(* opposite dir *)
apply (clarsimp split: mem_write_action.splits)
apply (rename_tac s s' ref field option)
apply (drule_tac x="mw_Mutate ref field option" in spec)
apply (drule mp)
apply (metis (lifting, full_types) filter_set member_filter)
apply clarsimp
apply (subst eq_impD[OF obj_at_eq_imp])
prefer 2
apply (fastforce cong: option.case_cong)
apply clarsimp
done
lemmas marked_deletions_fun_upd[simp] = eq_imp_fun_upd[OF mut_m.marked_deletions_eq_imp, simplified eq_imp_simps, rule_format]
lemma (in mut_m) marked_deletions_store_buffer_empty[intro]:
"tso_pending_mutate (mutator m) s = [] \<Longrightarrow> marked_deletions s"
by (auto simp: marked_deletions_def filter_empty_conv
split: mem_write_action.splits)
lemma (in mut_m) marked_deletions_alloc[simp]:
"\<lbrakk> marked_deletions s; heap (s sys) r' = None; valid_refs_inv s \<rbrakk>
\<Longrightarrow> marked_deletions (s(mutator m' := s (mutator m')\<lparr>roots := roots'\<rparr>, sys := s sys\<lparr>heap := sys_heap s(r' \<mapsto> obj')\<rparr>))"
apply (clarsimp simp: marked_deletions_def split: mem_write_action.splits)
apply (rename_tac ref field option)
apply (drule_tac x="mw_Mutate ref field option" in spec)
apply clarsimp
apply (case_tac "ref = r'")
apply (auto simp: obj_at_field_on_heap_def split: option.splits)
done
(*>*)
subsection\<open> Mutator proofs \<close>
lemma (in mut_m) reachable_snapshot_inv_hs_get_roots_done[simp]:
assumes sti: "strong_tricolour_inv s"
assumes m: "\<forall>r \<in> mut_roots s. marked r s"
assumes ghr: "mut_ghost_honorary_root s = {}"
assumes t: "tso_pending_mutate (mutator m) s = []"
assumes vri: "valid_refs_inv s"
shows "reachable_snapshot_inv
(s(mutator m := s (mutator m)\<lparr>W := {}, ghost_handshake_phase := ghp'\<rparr>,
sys := s sys\<lparr>handshake_pending := hp', W := sys_W s \<union> mut_W s, ghost_handshake_in_sync := in'\<rparr>))"
(is "reachable_snapshot_inv ?s'")
proof(rule, clarsimp)
fix r assume "reachable r s"
then show "in_snapshot r ?s'"
proof (induct rule: reachable_induct)
case (root x) with m show ?case
apply (clarsimp simp: in_snapshot_def simp del: fun_upd_apply) (* FIXME intro rules *)
apply (auto dest: marked_imp_black_or_grey)
done
next
case (ghost_honorary_root x) with ghr show ?case by simp
next
case (tso_root x) with t show ?case
apply (clarsimp simp: filter_empty_conv tso_write_refs_def)
apply (force split: mem_write_action.splits)
done
next
case (reaches x y)
from reaches vri have "valid_ref x s" "valid_ref y s" by auto
with reaches sti vri show ?case
apply (clarsimp simp: in_snapshot_def simp del: fun_upd_apply)
apply (elim disjE)
apply (clarsimp simp: strong_tricolour_inv_def)
apply (drule spec[where x=x])
apply clarsimp
apply (auto dest!: marked_imp_black_or_grey)[1]
apply (cases "white y s")
apply (auto dest: grey_protects_whiteE
dest!: marked_imp_black_or_grey)
done
qed
qed
lemma (in mut_m) reachable_snapshot_inv_hs_get_work_done[simp]:
"reachable_snapshot_inv s
\<Longrightarrow> reachable_snapshot_inv
(s(mutator m := s (mutator m)\<lparr>W := {}\<rparr>,
sys := s sys\<lparr>handshake_pending := (handshake_pending (s sys))(m := False), W := sys_W s \<union> mut_W s,
ghost_handshake_in_sync := (ghost_handshake_in_sync (s sys))(m := True)\<rparr>))"
by (simp add: reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def)
lemma (in mut_m) reachable_write_enqueue[simp]:
"reachable r (s(sys := s sys\<lparr>mem_write_buffers := (mem_write_buffers (s sys))(mutator m := sys_mem_write_buffers (mutator m) s @ [w])\<rparr>))
\<longleftrightarrow> reachable r s \<or> (\<exists>x. x \<in> write_refs w \<and> (x reaches r) s)"
by (auto simp: reachable_def tso_write_refs_def)
lemma no_black_refs_mo_co_mark[simp]:
"\<lbrakk> ghost_honorary_grey (s p) = {}; white r s \<rbrakk>
\<Longrightarrow> no_black_refs (s(p := s p\<lparr>ghost_honorary_grey := {r}\<rparr>)) \<longleftrightarrow> no_black_refs s"
by (auto simp: no_black_refs_def)
lemma (in mut_m) reachable_snapshot_inv_mo_co_mark[simp]:
"\<lbrakk> ghost_honorary_grey (s p) = {}; reachable_snapshot_inv s \<rbrakk>
\<Longrightarrow> reachable_snapshot_inv (s(p := s p\<lparr> ghost_honorary_grey := {r} \<rparr>))"
by (auto simp: in_snapshot_def reachable_snapshot_inv_def)
lemma (in mut_m) no_black_refs_alloc[simp]:
"\<lbrakk> heap (s sys) r' = None; no_black_refs s \<rbrakk>
\<Longrightarrow> no_black_refs (s(mutator m' := s (mutator m')\<lparr>roots := roots'\<rparr>, sys := s sys\<lparr>heap := sys_heap s(r' \<mapsto> \<lparr>obj_mark = fl, obj_fields = Map.empty\<rparr>)\<rparr>))
\<longleftrightarrow> fl \<noteq> sys_fM s \<or> grey r' s"
by (auto simp: no_black_refs_def)
lemma (in mut_m) reachable_snapshot_inv_load[simp]:
"\<lbrakk> reachable_snapshot_inv s; sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref r'; r \<in> mut_roots s \<rbrakk>
\<Longrightarrow> reachable_snapshot_inv (s(mutator m := s (mutator m)\<lparr> roots := mut_roots s \<union> Option.set_option r' \<rparr>))"
by (simp add: reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def)
lemma (in mut_m) reachable_snapshot_inv_store_ins[simp]:
"\<lbrakk> reachable_snapshot_inv s; r \<in> mut_roots s; (\<exists>r'. opt_r' = Some r') \<longrightarrow> the opt_r' \<in> mut_roots s \<rbrakk>
\<Longrightarrow> reachable_snapshot_inv (s(mutator m := s (mutator m)\<lparr>ghost_honorary_root := {}\<rparr>,
sys := s sys\<lparr> mem_write_buffers := (mem_write_buffers (s sys))(mutator m := sys_mem_write_buffers (mutator m) s @ [mw_Mutate r f opt_r']) \<rparr>))"
apply (clarsimp simp: reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def)
apply (drule_tac x=x in spec)
apply (auto simp: reachable_def)
done
lemma (in mut_m) marked_insertions_store_ins[simp]:
"\<lbrakk> marked_insertions s; (\<exists>r'. opt_r' = Some r') \<longrightarrow> marked (the opt_r') s \<rbrakk>
\<Longrightarrow> marked_insertions
(s(mutator m := s (mutator m)\<lparr>ghost_honorary_root := {}\<rparr>,
sys := s sys
\<lparr>mem_write_buffers := (mem_write_buffers (s sys))(mutator m := sys_mem_write_buffers (mutator m) s @ [mw_Mutate r f opt_r'])\<rparr>))"
by (auto simp: marked_insertions_def
split: mem_write_action.splits option.splits)
lemma (in mut_m) marked_deletions_store_ins[simp]:
"\<lbrakk> marked_deletions s; obj_at_field_on_heap (\<lambda>r'. marked r' s) r f s \<rbrakk>
\<Longrightarrow> marked_deletions
(s(mutator m := s (mutator m)\<lparr>ghost_honorary_root := {}\<rparr>,
sys := s sys
\<lparr>mem_write_buffers := (mem_write_buffers (s sys))(mutator m := sys_mem_write_buffers (mutator m) s @ [mw_Mutate r f opt_r'])\<rparr>))"
by (auto simp: marked_deletions_def
split: mem_write_action.splits option.splits)
lemma (in mut_m) reachable_snapshot_inv_deref_del[simp]:
"\<lbrakk> reachable_snapshot_inv s; sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref opt_r'; r \<in> mut_roots s; mut_ghost_honorary_root s = {} \<rbrakk>
\<Longrightarrow> reachable_snapshot_inv (s(mutator m := s (mutator m)\<lparr>ghost_honorary_root := Option.set_option opt_r', ref := opt_r'\<rparr>))"
by (clarsimp simp: reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def)
lemma (in mut_m) mutator_phase_inv[intro]:
"\<lbrace> handshake_invL
\<^bold>\<and> mark_object_invL
\<^bold>\<and> mut_get_roots.mark_object_invL m
\<^bold>\<and> mut_store_del.mark_object_invL m
\<^bold>\<and> mut_store_ins.mark_object_invL m
\<^bold>\<and> LSTP (fA_rel_inv \<^bold>\<and> fM_rel_inv \<^bold>\<and> handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> phase_rel_inv \<^bold>\<and> strong_tricolour_inv \<^bold>\<and> sys_phase_inv \<^bold>\<and> valid_refs_inv \<^bold>\<and> valid_W_inv) \<rbrace>
mutator m
\<lbrace> LSTP mutator_phase_inv \<rbrace>"
apply (vcg_jackhammer dest!: handshake_phase_invD simp: fA_rel_inv_def fM_rel_inv_def;
simp add: mutator_phase_inv_aux_case split: handshake_phase.splits if_splits)
subgoal
apply (drule_tac x=m in spec)
apply (clarsimp simp: fM_rel_def hp_step_rel_def)
apply (intro conjI impI; simp)
apply (elim disjE; force simp: fA_rel_def) (* FIXME annoying: unfolding fA_rel early leads to non-termination *)
apply (rule reachable_snapshot_inv_alloc, simp_all)
apply (elim disjE; force simp: fA_rel_def) (* FIXME annoying: unfolding fA_rel early leads to non-termination *)
done
subgoal for s s'
apply (drule_tac x=m in spec)
apply (intro conjI impI)
apply clarsimp
apply (rule marked_deletions_store_ins, assumption) (* FIXME shuffle the following into this lemma *)
apply (cases "(\<forall>opt_r'. mw_Mutate (mut_tmp_ref s\<down>) (mut_field s\<down>) opt_r' \<notin> set (sys_mem_write_buffers (mutator m) s\<down>))")
apply force
apply (force simp: marked_deletions_def)
apply clarsimp
apply (erule marked_insertions_store_ins)
apply (drule phase_rel_invD)
apply (clarsimp simp: phase_rel_def hp_step_rel_def; elim disjE; fastforce dest: reachable_blackD elim: blackD)
apply clarsimp
apply (rule marked_deletions_store_ins, assumption) (* FIXME as above *)
apply clarsimp
apply (erule disjE)
apply (drule phase_rel_invD)
apply (clarsimp simp: phase_rel_def)
apply (elim disjE, simp_all)[1]
apply (clarsimp simp: hp_step_rel_def)
apply (clarsimp simp: hp_step_rel_def)
apply (case_tac "sys_ghost_handshake_phase s\<down>", simp_all)[1] (* FIXME invert handshake_phase_rel *)
apply (clarsimp simp: fA_rel_def fM_rel_def)
apply (elim disjE, simp_all)[1]
apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits)
apply (rule conjI)
apply fast
apply clarsimp
apply (frule_tac r=x2a in blackD(1)[OF reachable_blackD], simp_all)[1]
apply (rule_tac x="mut_tmp_ref s\<down>" in reachableE; auto simp: ran_def split: obj_at_splits; fail)
apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits)
apply (rule conjI)
apply fast
apply clarsimp
apply (frule_tac r=x2a in blackD(1)[OF reachable_blackD], simp_all)[1]
apply (rule_tac x="mut_tmp_ref s\<down>" in reachableE; auto simp: ran_def split: obj_at_splits; fail)
apply (force simp: marked_deletions_def)
done
(* hs_noop_done *)
subgoal for s s'
apply (drule_tac x=m in spec)
apply (simp add: fA_rel_def fM_rel_def hp_step_rel_def)
apply (cases "mut_ghost_handshake_phase s\<down>", simp_all)[1] (* FIXME invert handshake_step *)
apply (erule marked_insertions_store_buffer_empty) (* FIXME simp? *)
apply (erule marked_deletions_store_buffer_empty) (* FIXME simp? *)
done
(* hs_get_roots_done *)
subgoal
apply (drule_tac x=m in spec)
apply (simp add: fA_rel_def fM_rel_def hp_step_rel_def)
done
(* hs_get_work_done *)
subgoal
apply (drule_tac x=m in spec)
apply (simp add: fA_rel_def fM_rel_def hp_step_rel_def)
done
done
lemma (in mut_m') mutator_phase_inv[intro]:
notes mut_m.mark_object_invL_def[inv]
notes mut_m.handshake_invL_def[inv]
shows
"\<lbrace> handshake_invL \<^bold>\<and> mut_m.handshake_invL m'
\<^bold>\<and> mut_m.mark_object_invL m'
\<^bold>\<and> mut_get_roots.mark_object_invL m'
\<^bold>\<and> mut_store_del.mark_object_invL m'
\<^bold>\<and> mut_store_ins.mark_object_invL m'
\<^bold>\<and> LSTP (fA_rel_inv \<^bold>\<and> fM_rel_inv \<^bold>\<and> handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> valid_refs_inv) \<rbrace>
mutator m'
\<lbrace> LSTP mutator_phase_inv \<rbrace>"
apply (vcg_jackhammer simp: fA_rel_inv_def fM_rel_inv_def dest!: handshake_phase_invD)
apply (simp_all add: mutator_phase_inv_aux_case split: handshake_phase.splits)
apply (drule spec[where x=m])
apply (intro conjI impI)
apply (clarsimp simp: fA_rel_def fM_rel_def hp_step_rel_def)
apply (elim disjE, auto)[1]
apply (rule reachable_snapshot_inv_alloc, simp_all)[1]
apply (clarsimp simp: fA_rel_def fM_rel_def hp_step_rel_def)
apply (elim disjE, auto)[1]
apply (drule spec[where x=m])
apply (clarsimp simp: no_black_refs_def)
apply (clarsimp simp: reachable_snapshot_inv_def in_snapshot_def)
apply (drule spec[where x=m])
apply (clarsimp simp: no_black_refs_def reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def)
done
lemma (in sys) grey_protects_white_dequeue_mark[simp]:
assumes fl: "fl = sys_fM s"
assumes "r \<in> ghost_honorary_grey (s p)"
shows "(\<exists>g. (g grey_protects_white w) (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\<lambda>_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>)))
\<longleftrightarrow> (\<exists>g. (g grey_protects_white w) s)" (is "(\<exists>g. (g grey_protects_white w) ?s') \<longleftrightarrow> ?rhs")
proof
assume "\<exists>g. (g grey_protects_white w) ?s'"
then obtain g where "(g grey_protects_white w) ?s'" by blast
with assms show ?rhs
apply (clarsimp simp: grey_protects_white_def has_white_path_to_def)
apply (rotate_tac -1)
apply (induct rule: rtranclp.induct)
apply fastforce
apply clarsimp
apply (rename_tac a b c g)
apply (case_tac "white c s")
apply (rule_tac x=g in exI)
apply clarsimp
apply (erule rtranclp.intros)
apply clarsimp
apply (auto split: obj_at_splits if_splits)
done
next
assume ?rhs
then obtain g' where "(g' grey_protects_white w) s" by blast
with assms show "\<exists>g. (g grey_protects_white w) ?s'"
apply (clarsimp simp: grey_protects_white_def has_white_path_to_def)
apply (rotate_tac -1)
apply (induct rule: rtranclp.induct)
apply (fastforce simp: grey_def)
apply clarsimp
apply (rename_tac a b c g)
apply (case_tac "c = r")
apply (clarsimp simp: grey_def)
apply blast
apply (rule_tac x=g in exI)
apply clarsimp
apply (erule rtranclp.intros)
apply clarsimp
done
qed
lemma (in sys) reachable_snapshot_inv_dequeue_mark[simp]:
"\<lbrakk> sys_mem_write_buffers p s = mw_Mark r fl # ws; mut_m.reachable_snapshot_inv m s; valid_W_inv s \<rbrakk>
\<Longrightarrow> mut_m.reachable_snapshot_inv m (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\<lambda>_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))"
apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def)
apply (rename_tac x)
apply (subst (asm) reachable_fun_upd, simp_all)[1]
apply auto[1]
apply (drule_tac x=x in spec)
apply clarsimp
apply (subst (asm) arg_cong[where f=Not, OF grey_protects_white_dequeue_mark, simplified], simp_all)
apply blast
apply blast
apply blast
done
lemma (in sys) marked_insertions_dequeue_mark[simp]:
"\<lbrakk> sys_mem_write_buffers p s = mw_Mark r fl # ws; mut_m.marked_insertions m s; tso_writes_inv s; valid_W_inv s \<rbrakk>
\<Longrightarrow> mut_m.marked_insertions m (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\<lambda>_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))"
apply (clarsimp simp: mut_m.marked_insertions_def)
apply (cases "mutator m = p")
apply clarsimp
apply (rename_tac x)
apply (drule_tac x=x in spec)
apply (auto split: mem_write_action.splits option.splits obj_at_splits)[1]
apply clarsimp
apply (rename_tac x)
apply (drule_tac x=x in spec)
apply (auto split: mem_write_action.splits option.splits obj_at_splits)[1]
done
lemma (in sys) marked_insertions_dequeue_ref[simp]:
"\<lbrakk> sys_mem_write_buffers p s = mw_Mutate r f r' # ws; mut_m.marked_insertions m s \<rbrakk>
\<Longrightarrow> mut_m.marked_insertions m (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (\<lambda>obj. obj\<lparr>obj_fields := (obj_fields obj)(f := r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))"
apply (clarsimp simp: mut_m.marked_insertions_def)
apply (cases "mutator m = p")
apply clarsimp
apply (rename_tac x)
apply (drule_tac x=x in spec)
apply (auto split: mem_write_action.splits option.splits obj_at_splits)[1]
apply clarsimp
apply (rename_tac x)
apply (drule_tac x=x in spec)
apply (auto split: mem_write_action.splits option.splits obj_at_splits)[1]
done
(* redundant to fit with other rules. Perhaps want points_to with explicit witness for f? *)
lemma points_to_mw_Mutate:
"(x points_to y)
(s(sys := (s sys)\<lparr> heap := (sys_heap s)(r := Option.map_option (\<lambda>obj :: ('field, 'ref) object. obj\<lparr>obj_fields := (obj_fields obj)(f := opt_r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))(p := ws) \<rparr>))
\<longleftrightarrow> (r \<noteq> x \<and> (x points_to y) s) \<or> (r = x \<and> valid_ref r s \<and> (opt_r' = Some y \<or> ( (x points_to y) s \<and> obj_at (\<lambda>obj. \<exists>f'. obj_fields obj f' = Some y \<and> f \<noteq> f') r s)))"
by (auto simp: ran_def split: obj_at_splits)
(* shows the snapshot is preserved by the two marks. *)
lemma (in sys) grey_protects_white_dequeue_ref[simp]:
assumes sb: "sys_mem_write_buffers (mutator m) s = mw_Mutate r f opt_r' # ws"
assumes mi: "mut_m.marked_insertions m s"
assumes md: "mut_m.marked_deletions m s"
notes map_option.compositionality[simp] o_def[simp]
shows "(\<exists>g. (g grey_protects_white w) (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (\<lambda>obj. obj\<lparr>obj_fields := (obj_fields obj)(f := opt_r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))(mutator m := ws)\<rparr>)))
\<longleftrightarrow> (\<exists>g. (g grey_protects_white w) s)" (is "(\<exists>g. (g grey_protects_white w) ?s') \<longleftrightarrow> ?rhs")
proof
assume "(\<exists>g. (g grey_protects_white w) ?s')"
then obtain g where "(g grey_protects_white w) ?s'" by blast
with mi sb show ?rhs
apply (clarsimp simp: grey_protects_white_def has_white_path_to_def)
apply (rotate_tac -1)
apply (induct rule: rtranclp.induct)
apply fastforce
apply (auto simp: points_to_mw_Mutate elim: rtranclp.intros(2))
apply (rename_tac a c g)
apply (clarsimp simp: mut_m.marked_insertions_def) (* FIXME rule *)
apply (drule_tac x="mw_Mutate r f (Some c)" in spec)
apply (simp split: obj_at_splits)
done
next
assume ?rhs then show "(\<exists>g. (g grey_protects_white w) ?s')"
proof(clarsimp)
fix g assume "(g grey_protects_white w) s"
with md sb show ?thesis
apply (clarsimp simp: grey_protects_white_def has_white_path_to_def)
apply (rotate_tac -1)
apply (induct rule: rtranclp.induct)
apply fastforce
apply clarsimp
apply (rename_tac a b c g)
apply (case_tac "b = r")
defer
apply (auto simp: points_to_mw_Mutate elim: rtranclp.intros(2))[1]
apply clarsimp
apply (subst (asm) (3) obj_at_def) (* FIXME rule: witness field for r points_to c *)
apply (clarsimp simp: ran_def split: option.splits)
apply (rename_tac a c g x2 aa)
apply (case_tac "aa = f")
defer
apply (rule_tac x=g in exI)
apply clarsimp
apply (erule rtranclp.intros)
apply (auto split: obj_at_splits)[1]
apply clarsimp
apply (clarsimp simp: mut_m.marked_deletions_def) (* FIXME rule *)
apply (drule spec[where x="mw_Mutate r f opt_r'"])
apply (clarsimp simp: obj_at_field_on_heap_def)
apply (simp split: obj_at_splits)
done
qed
qed
(* write barrier installed but not all mutators are necessarily past get_roots *)
lemma (in sys) reachable_snapshot_inv_dequeue_ref[simp]:
fixes s :: "('field, 'mut, 'ref) lsts"
assumes sb: "sys_mem_write_buffers (mutator m') s = mw_Mutate r f opt_r' # ws"
assumes mi: "mut_m.marked_insertions m' s"
assumes md: "mut_m.marked_deletions m' s"
assumes rsi: "mut_m.reachable_snapshot_inv m s"
assumes sti: "strong_tricolour_inv s"
assumes vri: "valid_refs_inv s"
notes map_option.compositionality[simp] o_def[simp]
shows "mut_m.reachable_snapshot_inv m (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (\<lambda>obj. obj\<lparr>obj_fields := (obj_fields obj)(f := opt_r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))(mutator m' := ws)\<rparr>))" (is "mut_m.reachable_snapshot_inv m ?s'")
proof(rule mut_m.reachable_snapshot_invI)
fix y assume y: "mut_m.reachable m y ?s'"
then have "(mut_m.reachable m y s \<or> mut_m.reachable m' y s) \<and> in_snapshot y ?s'"
proof(induct rule: reachable_induct)
case (root x) with mi md rsi sb show ?case
apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def
simp del: fun_upd_apply)
apply auto
done
next
case (ghost_honorary_root x) with mi md rsi sb show ?case
apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def
simp del: fun_upd_apply)
apply auto
done
next
case (tso_root x) with mi md rsi sb show ?case
apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def
simp del: fun_upd_apply)
apply (clarsimp split: if_splits)
apply (rename_tac xa)
apply (case_tac xa, simp_all)[1]
apply (rename_tac ref field option)
apply (clarsimp simp: mut_m.marked_deletions_def mut_m.marked_insertions_def)
apply (drule_tac x="mw_Mutate ref field option" in spec)
apply (drule_tac x="mw_Mutate ref field option" in spec)
apply clarsimp
apply (frule spec[where x=x])
apply (subgoal_tac "mut_m.reachable m x s")
apply force
apply (rule reachableI(2))
apply (force simp: mut_m.tso_write_refs_def)
apply auto
done
next
case (reaches x y)
from reaches sb have y: "mut_m.reachable m y s \<or> mut_m.reachable m' y s"
apply (clarsimp simp: points_to_mw_Mutate mut_m.reachable_snapshot_inv_def in_snapshot_def)
apply (elim disjE, (force dest!: reachableE mutator_reachable_tso)+)[1]
done
moreover
from y vri have "valid_ref y s" by auto
with reaches mi md rsi sb sti y have "(black y s \<or> (\<exists>x. (x grey_protects_white y) s))"
apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def
simp del: fun_upd_apply)
apply clarsimp
apply (drule spec[where x=y])
apply (clarsimp simp: points_to_mw_Mutate mut_m.marked_insertions_def mut_m.marked_deletions_def) (* FIXME lemmas *)
apply (drule spec[where x="mw_Mutate r f opt_r'"])+
apply clarsimp
apply (elim disjE, simp_all) (* FIXME probably want points_to_mw_Mutate as an elim rule to make this robust, reduce duplication *)
apply (force dest!: reachableE)
apply (force dest!: reachableE)
apply clarsimp
apply (drule (3) strong_tricolour_invD)
apply (metis (no_types) grey_protects_whiteI marked_imp_black_or_grey(1))
apply clarsimp
apply (cases "white y s") (* FIXME lemma *)
apply (drule (2) grey_protects_whiteE)
apply blast
apply (force simp: black_def split: obj_at_splits)
apply clarsimp
apply (elim disjE, simp_all)
apply (force simp: black_def)
apply clarsimp
apply (drule (3) strong_tricolour_invD)
apply (force simp: black_def)
apply clarsimp
apply (elim disjE, simp_all)
apply (force simp: black_def)
apply clarsimp
apply (cases "white y s") (* FIXME lemma *)
apply (drule (2) grey_protects_whiteE)
apply blast
apply (force simp: black_def split: obj_at_splits)
apply clarsimp
apply (elim disjE, simp_all)
apply (force simp: black_def)
apply clarsimp
apply (drule (3) strong_tricolour_invD)
apply (force simp: black_def)
apply clarsimp
apply (elim disjE, simp_all)
apply (force simp: black_def)
apply clarsimp
apply (cases "white y s") (* FIXME lemma *)
apply (drule (2) grey_protects_whiteE)
apply blast
apply (force simp: black_def split: obj_at_splits)
done
moreover note mi md rsi sb
ultimately show ?case
apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def
simp del: fun_upd_apply)
apply clarsimp
done
qed
then show "in_snapshot y ?s'" by blast
qed
lemma valid_refs_invI:
"\<lbrakk> \<And>m x y. \<lbrakk> (x reaches y) s; x \<in> mut_m.mut_roots m s \<or> x \<in> mut_m.mut_ghost_honorary_root m s \<or> x \<in> mut_m.tso_write_refs m s \<or> grey x s \<rbrakk> \<Longrightarrow> valid_ref y s
\<rbrakk> \<Longrightarrow> valid_refs_inv s"
by (auto simp: valid_refs_inv_def mut_m.reachable_def grey_reachable_def)
lemma black_heap_reachable:
assumes "mut_m.reachable m y s"
assumes bh: "black_heap s"
assumes vri: "valid_refs_inv s"
shows "black y s"
using assms
apply (induct rule: reachable_induct)
apply (simp_all add: black_heap_def valid_refs_invD)
apply (metis obj_at_weakenE reachableE valid_refs_inv_def)
done
lemma valid_refs_inv_dequeue_ref[simp]:
notes map_option.compositionality[simp] o_def[simp]
fixes s :: "('field, 'mut, 'ref) lsts"
assumes vri: "valid_refs_inv s"
assumes sb: "sys_mem_write_buffers (mutator m') s = mw_Mutate r f opt_r' # ws"
shows "valid_refs_inv (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (\<lambda>obj. obj\<lparr>obj_fields := (obj_fields obj)(f := opt_r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))(mutator m' := ws)\<rparr>))" (is "valid_refs_inv ?s'")
proof(rule valid_refs_invI)
fix m
let ?root = "\<lambda>m x (s::('field, 'mut, 'ref) lsts). x \<in> mut_m.mut_roots m s \<or> x \<in> mut_m.mut_ghost_honorary_root m s \<or> x \<in> mut_m.tso_write_refs m s \<or> grey x s"
fix x y assume xy: "(x reaches y) ?s'" and x: "?root m x ?s'"
from xy have "(\<exists>m x. ?root m x s \<and> (x reaches y) s) \<and> valid_ref y ?s'"
proof induct
case base with x sb vri show ?case
apply -
apply (subst obj_at_fun_upd)
apply (auto simp: mut_m.tso_write_refs_def split: if_splits intro: valid_refs_invD(5)[where m=m])
apply (metis list.set_intros(2) rtranclp.rtrancl_refl)
done (* FIXME rules *)
next
case (step y z)
with sb vri show ?case
apply -
apply (subst obj_at_fun_upd, clarsimp)
apply (subst (asm) obj_at_fun_upd, fastforce)
apply (clarsimp simp: points_to_mw_Mutate simp del: fun_upd_apply)
apply (fastforce elim: rtranclp.intros(2) simp: mut_m.tso_write_refs_def intro: exI[where x=m'] valid_refs_invD(5)[where m=m'])
done
qed
then show "valid_ref y ?s'" by blast
qed
declare map_option.compositionality[simp] o_def[simp]
lemma (in sys) reachable_snapshot_inv_black_heap_no_grey_refs_dequeue_ref[simp]:
assumes sb: "sys_mem_write_buffers (mutator m') s = mw_Mutate r f opt_r' # ws"
assumes bh: "black_heap s"
assumes ngr: "no_grey_refs s"
assumes vri: "valid_refs_inv s"
shows "mut_m.reachable_snapshot_inv m (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (\<lambda>obj. obj\<lparr>obj_fields := (obj_fields obj)(f := opt_r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))(mutator m' := ws)\<rparr>))" (is "mut_m.reachable_snapshot_inv m ?s'")
apply (rule mut_m.reachable_snapshot_invI)
apply (rule in_snapshotI)
apply (erule black_heap_reachable)
using bh vri
apply (simp add: black_heap_def)
using bh ngr sb vri
apply (subst valid_refs_inv_def)
apply clarsimp
apply (simp add: no_grey_refs_def grey_reachable_def)
apply clarsimp
apply (drule black_heap_reachable)
apply (simp add: black_heap_def)
apply clarsimp
apply auto
done
lemma (in sys) marked_deletions_dequeue_mark[simp]:
"\<lbrakk> sys_mem_write_buffers p s = mw_Mark r fl # ws; mut_m.marked_deletions m s; tso_writes_inv s; valid_W_inv s \<rbrakk>
\<Longrightarrow> mut_m.marked_deletions m (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\<lambda>_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))"
apply (clarsimp simp: mut_m.marked_deletions_def)
apply (cases "mutator m = p")
apply clarsimp
apply (rename_tac x)
apply (drule_tac x=x in spec)
apply (clarsimp split: mem_write_action.splits)
apply (simp add: obj_at_field_on_heap_def cong: option.case_cong)
apply (auto split: option.splits)[1]
apply clarsimp
apply (rename_tac x)
apply (drule_tac x=x in spec)
apply (clarsimp split: mem_write_action.splits)
apply (simp add: obj_at_field_on_heap_def cong: option.case_cong)
apply (auto split: option.splits)[1]
done
lemma (in sys) marked_deletions_dequeue_ref[simp]:
"\<lbrakk> sys_mem_write_buffers (mutator m') s = mw_Mutate r f opt_r' # ws; mut_m.marked_deletions m s; mut_m.marked_insertions m' s \<rbrakk>
\<Longrightarrow> mut_m.marked_deletions m (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (\<lambda>obj. obj\<lparr>obj_fields := (obj_fields obj)(f := opt_r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))((mutator m') := ws)\<rparr>))"
+supply [[simproc del: defined_all]]
apply (clarsimp simp: mut_m.marked_deletions_def)
apply (cases "m = m'")
apply clarsimp
apply (rename_tac x)
apply (drule_tac x=x in spec)
apply (fastforce simp: obj_at_field_on_heap_def mut_m.marked_insertions_def split: mem_write_action.splits obj_at_splits option.splits)
apply clarsimp
apply (rename_tac x)
apply (drule_tac x=x in spec)
apply (fastforce simp: obj_at_field_on_heap_def mut_m.marked_insertions_def split: mem_write_action.splits obj_at_splits option.splits)
done
lemma (in sys) black_heap_marked_insertions_dequeue[simp]:
"\<lbrakk> black_heap s; valid_refs_inv s \<rbrakk> \<Longrightarrow> mut_m.marked_insertions m s"
by (auto simp: mut_m.marked_insertions_def black_heap_def black_def
split: mem_write_action.splits option.splits
dest: valid_refs_invD)
lemma (in sys) mutator_phase_inv[intro]:
"\<lbrace> LSTP (fA_rel_inv \<^bold>\<and> fM_rel_inv \<^bold>\<and> handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> strong_tricolour_inv \<^bold>\<and> sys_phase_inv \<^bold>\<and> tso_writes_inv \<^bold>\<and> valid_refs_inv \<^bold>\<and> valid_W_inv) \<rbrace>
sys
\<lbrace> LSTP (mut_m.mutator_phase_inv m) \<rbrace>"
apply vcg_nihe
apply vcg_ni
apply (clarsimp simp: fA_rel_inv_def fM_rel_inv_def mutator_phase_inv_aux_case p_not_sys hp_step_rel_def do_write_action_def
split: handshake_phase.splits mem_write_action.splits)
(* fM *)
prefer 2
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (erule disjE)
apply (clarsimp simp: fM_rel_def hp_step_rel_def)
apply clarsimp
(* FIXME mess *)
apply (frule mut_m.handshake_phase_invD[where m=m])
apply (intro allI conjI impI)
apply (erule disjE)
apply auto[1]
apply clarsimp
apply (rename_tac s s' ws ref field option ma)
apply (rule marked_deletions_dequeue_ref, simp_all)
apply (drule_tac x=ma in spec)
apply (frule_tac m=ma in mut_m.handshake_phase_invD, clarsimp simp: hp_step_rel_def)
apply (elim disjE, simp_all)[1]
apply (erule disjE)
apply auto[1]
apply clarsimp
apply (rule marked_deletions_dequeue_ref, simp_all)
apply (drule_tac x=ma in spec)
apply (frule_tac m=ma in mut_m.handshake_phase_invD, clarsimp simp: hp_step_rel_def)
apply (elim disjE, simp_all)[1]
apply (clarsimp simp: fA_rel_def fM_rel_def if_splits)
apply (elim disjE, simp_all)[1]
apply (erule disjE, simp)
apply (clarsimp simp: hp_step_rel_def)
apply (frule_tac m=ma in mut_m.handshake_phase_invD, clarsimp simp: hp_step_rel_def)
apply (elim disjE, simp_all split: if_splits)[1]
apply (clarsimp simp: fA_rel_def fM_rel_def, blast)
done
(*>*)
subsection\<open>The infamous termination argument.\<close>
text\<open>
We need to know that if the GC does not receive any further work to do
at \<open>get_roots\<close> and \<open>get_work\<close>, then there are no grey
objects left. Essentially this encodes the stability property that
grey objects must exist for mutators to create grey objects.
Note that this is not invariant across the scan: it is possible for
the GC to hold all the grey references. The two handshakes transform
the GC's local knowledge that it has no more work to do into a global
property, or gives it more work.
\<close>
definition (in mut_m) gc_W_empty_mut_inv :: "('field, 'mut, 'ref) lsts_pred" where
"gc_W_empty_mut_inv =
((EMPTY sys_W \<^bold>\<and> sys_ghost_handshake_in_sync m \<^bold>\<and> \<^bold>\<not>(EMPTY (WL (mutator m))))
\<^bold>\<longrightarrow> (\<^bold>\<exists>m'. \<^bold>\<not>(sys_ghost_handshake_in_sync m') \<^bold>\<and> \<^bold>\<not>(EMPTY (WL (mutator m')))))"
locset_definition (in -) gc_W_empty_locs :: "location set" where
"gc_W_empty_locs \<equiv>
idle_locs \<union> init_locs \<union> sweep_locs \<union> { ''mark_read_fM'', ''mark_write_fA'', ''mark_end'' }
\<union> prefixed ''mark_noop''
\<union> prefixed ''mark_loop_get_roots''
\<union> prefixed ''mark_loop_get_work''"
locset_definition "black_heap_locs = { ''sweep_idle'', ''idle_noop_mfence'', ''idle_noop_init_type'' }"
locset_definition "no_grey_refs_locs = black_heap_locs \<union> sweep_locs \<union> {''mark_end''}"
inv_definition (in gc) gc_W_empty_invL :: "('field, 'mut, 'ref) gc_pred" where
"gc_W_empty_invL =
(atS_gc (hs_get_roots_locs \<union> hs_get_work_locs) (\<^bold>\<forall>m. mut_m.gc_W_empty_mut_inv m)
\<^bold>\<and> at_gc ''mark_loop_get_roots_load_W'' (EMPTY sys_W \<^bold>\<longrightarrow> no_grey_refs)
\<^bold>\<and> at_gc ''mark_loop_get_work_load_W'' (EMPTY sys_W \<^bold>\<longrightarrow> no_grey_refs)
\<^bold>\<and> at_gc ''mark_loop'' (EMPTY gc_W \<^bold>\<longrightarrow> no_grey_refs)
\<^bold>\<and> atS_gc no_grey_refs_locs no_grey_refs
\<^bold>\<and> atS_gc gc_W_empty_locs (EMPTY gc_W))"
(*<*)
lemma (in mut_m) gc_W_empty_mut_inv_eq_imp:
"eq_imp (\<lambda>m'. sys_W \<^bold>\<otimes> WL (mutator m') \<^bold>\<otimes> sys_ghost_handshake_in_sync m')
gc_W_empty_mut_inv"
by (simp add: eq_imp_def gc_W_empty_mut_inv_def)
lemmas gc_W_empty_mut_inv_fun_upd[simp] = eq_imp_fun_upd[OF mut_m.gc_W_empty_mut_inv_eq_imp, simplified eq_imp_simps, rule_format]
lemma (in gc) gc_W_empty_invL_eq_imp:
"eq_imp (\<lambda>(m', p) s. (AT s gc, s\<down> gc, sys_W s\<down>, WL p s\<down>, sys_ghost_handshake_in_sync m' s\<down>))
gc_W_empty_invL"
by (simp add: eq_imp_def gc_W_empty_invL_def mut_m.gc_W_empty_mut_inv_def no_grey_refs_def grey_def)
lemmas gc_W_empty_invL_niE[nie] =
iffD1[OF gc.gc_W_empty_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode, rule_format], rotated -1]
lemma get_roots_get_work_subseteq_gc_W_empty_locs:
"hs_get_roots_locs \<union> hs_get_work_locs \<subseteq> gc_W_empty_locs"
by (auto simp: hs_get_roots_locs_def hs_get_work_locs_def gc_W_empty_locs_def)
lemma (in gc) gc_W_empty_mut_inv_handshake_init[iff]:
"mut_m.gc_W_empty_mut_inv m (s(sys := s sys\<lparr>handshake_type := ht, ghost_handshake_in_sync := \<langle>False\<rangle>\<rparr>))"
"mut_m.gc_W_empty_mut_inv m (s(sys := s sys\<lparr>handshake_type := ht, ghost_handshake_in_sync := \<langle>False\<rangle>, ghost_handshake_phase := hp' \<rparr>))"
by (simp_all add: mut_m.gc_W_empty_mut_inv_def)
lemma gc_W_empty_mut_inv_load_W:
"\<lbrakk> \<forall>m. mut_m.gc_W_empty_mut_inv m s; \<forall>m. sys_ghost_handshake_in_sync m s; WL gc s = {}; WL sys s = {} \<rbrakk>
\<Longrightarrow> no_grey_refs s"
apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def no_grey_refs_def grey_def)
apply (rename_tac x xa)
apply (case_tac xa)
apply (simp_all add: WL_def)
done
lemma (in gc) gc_W_empty_invL[intro]:
"\<lbrace> handshake_invL \<^bold>\<and> obj_fields_marked_invL \<^bold>\<and> gc_W_empty_invL \<^bold>\<and> LSTP valid_W_inv \<rbrace>
gc
\<lbrace> gc_W_empty_invL \<rbrace>"
apply vcg_jackhammer
apply (auto elim: gc_W_empty_mut_inv_load_W simp: WL_def)
done
lemma (in sys) gc_gc_W_empty_invL[intro]:
"\<lbrace> gc.gc_W_empty_invL \<rbrace> sys"
by vcg_nihe
lemma (in gc) handshake_get_rootsD:
"\<lbrakk> atS gc hs_get_roots_locs s; handshake_invL s \<rbrakk> \<Longrightarrow> sys_ghost_handshake_phase s\<down> = hp_IdleMarkSweep \<and> sys_handshake_type s\<down> = ht_GetRoots"
apply (simp add: handshake_invL_def)
apply (elim conjE)
apply (drule mp, erule atS_mono[OF _ hs_get_roots_locs_subseteq_hp_IdleMarkSweep_locs])
apply simp
done
lemma (in gc) handshake_get_workD:
"\<lbrakk> atS gc hs_get_work_locs s; handshake_invL s \<rbrakk> \<Longrightarrow> sys_ghost_handshake_phase s\<down> = hp_IdleMarkSweep \<and> sys_handshake_type s\<down> = ht_GetWork"
apply (simp add: handshake_invL_def)
apply (elim conjE)
apply (drule mp, erule atS_mono[OF _ hs_get_work_locs_subseteq_hp_IdleMarkSweep_locs])
apply simp
done
lemma (in gc) handshake_get_roots_get_workD:
"\<lbrakk> atS gc (hs_get_roots_locs \<union> hs_get_work_locs) s; handshake_invL s \<rbrakk> \<Longrightarrow> sys_ghost_handshake_phase s\<down> = hp_IdleMarkSweep \<and> sys_handshake_type s\<down> \<in> {ht_GetWork, ht_GetRoots}"
apply (simp add: handshake_invL_def)
apply (elim conjE)
apply (drule mp, rule atS_mono[OF _ iffD2[OF Un_subset_iff, unfolded conj_explode, OF hs_get_roots_locs_subseteq_hp_IdleMarkSweep_locs hs_get_work_locs_subseteq_hp_IdleMarkSweep_locs]], assumption)
apply (auto simp: atS_un)
done
lemma no_grey_refs_locs_subseteq_hs_in_sync_locs:
"no_grey_refs_locs \<subseteq> hs_in_sync_locs"
by (auto simp: no_grey_refs_locs_def black_heap_locs_def hs_in_sync_locs_def hs_done_locs_def sweep_locs_def
dest: prefix_same_cases)
lemma (in mut_m) handshake_sweep_mark_endD:
"\<lbrakk> atS gc no_grey_refs_locs s; gc.handshake_invL s; handshake_phase_inv s\<down> \<rbrakk>
\<Longrightarrow> mut_ghost_handshake_phase s\<down> = hp_IdleMarkSweep \<and> All (ghost_handshake_in_sync (s\<down> sys))"
apply (simp add: gc.handshake_invL_def)
apply (elim conjE)
apply (drule mp, erule atS_mono[OF _ no_grey_refs_locs_subseteq_hs_in_sync_locs])
apply (drule handshake_phase_invD)
apply (simp only: no_grey_refs_locs_def cong del: atS_state_cong)
apply (clarsimp simp: atS_un)
apply (elim disjE)
apply (drule mp, erule atS_mono[where ls'="hp_IdleMarkSweep_locs"])
apply (clarsimp simp: black_heap_locs_def loc)
apply (clarsimp simp: hp_step_rel_def)
apply blast
apply (drule mp, erule atS_mono[where ls'="hp_IdleMarkSweep_locs"])
apply (clarsimp simp: hp_IdleMarkSweep_locs_def hp_step_rel_def)
apply (clarsimp simp: hp_step_rel_def)
apply blast
apply (clarsimp simp: atS_simps loc hp_step_rel_def)
apply blast
done
lemma empty_WL_GC:
"\<lbrakk> atS gc (hs_get_roots_locs \<union> hs_get_work_locs) s; gc.obj_fields_marked_invL s \<rbrakk> \<Longrightarrow> gc_ghost_honorary_grey s\<down> = {}"
by (clarsimp simp: gc.obj_fields_marked_invL_def
dest!: atS_mono[OF _ get_roots_get_work_subseteq_ghost_honorary_grey_empty_locs])
(* think about showing gc_W_empty_invL instead *)
lemma (in mut_m) gc_W_empty_mut_mo_co_mark[simp]:
"\<lbrakk> \<forall>x. mut_m.gc_W_empty_mut_inv x s\<down>; mutators_phase_inv s\<down>;
mut_ghost_honorary_grey s\<down> = {};
r \<in> mut_roots s\<down> \<union> mut_ghost_honorary_root s\<down>; white r s\<down>;
atS gc (hs_get_roots_locs \<union> hs_get_work_locs) s; gc.handshake_invL s; gc.obj_fields_marked_invL s;
atS gc gc_W_empty_locs s \<longrightarrow> gc_W s\<down> = {};
handshake_phase_inv s\<down>; valid_W_inv s\<down> \<rbrakk>
\<Longrightarrow> mut_m.gc_W_empty_mut_inv m' (s\<down>(mutator m := s\<down> (mutator m)\<lparr>ghost_honorary_grey := {r}\<rparr>))"
apply (frule (1) gc.handshake_get_roots_get_workD)
apply (frule handshake_phase_invD)
apply (clarsimp simp: hp_step_rel_def simp del: Un_iff)
apply (elim disjE, simp_all)
(* before get work *)
apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def)
apply blast
(* past get work *)
apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def)
apply (frule spec[where x=m], clarsimp)
apply (frule (2) reachable_snapshot_inv_white_root)
apply clarsimp
apply (drule grey_protects_whiteD)
apply (clarsimp simp: grey_def)
apply (rename_tac g x)
apply (case_tac x, simp_all)
(* Can't be the GC *)
prefer 2
apply (frule (1) empty_WL_GC)
apply (drule mp, erule atS_mono[OF _ get_roots_get_work_subseteq_gc_W_empty_locs])
apply (clarsimp simp: WL_def)
(* Can't be sys *)
prefer 2
apply (clarsimp simp: WL_def)
apply clarsimp
apply (rename_tac g mut)
apply (case_tac "sys_ghost_handshake_in_sync mut s\<down>")
apply (drule mp, rule_tac x=mut in exI, clarsimp)
apply blast
apply (rule_tac x=mut in exI)
apply clarsimp
(* before get roots *)
apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def)
apply blast
(* after get roots *)
apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def)
apply (frule spec[where x=m], clarsimp)
apply (frule (2) reachable_snapshot_inv_white_root)
apply clarsimp
apply (drule grey_protects_whiteD)
apply (clarsimp simp: grey_def)
apply (rename_tac g x)
apply (case_tac x, simp_all)
(* Can't be the GC *)
prefer 2
apply (frule (1) empty_WL_GC)
apply (drule mp, erule atS_mono[OF _ get_roots_get_work_subseteq_gc_W_empty_locs])
apply (clarsimp simp: WL_def)
(* Can't be sys *)
prefer 2
apply (clarsimp simp: WL_def)
apply clarsimp
apply (rename_tac g mut)
apply (case_tac "sys_ghost_handshake_in_sync mut s\<down>")
apply (drule mp, rule_tac x=mut in exI, clarsimp)
apply blast
apply (rule_tac x=mut in exI)
apply clarsimp
done (* FIXME common up *)
lemma (in mut_m) no_grey_refs_mo_co_mark[simp]:
"\<lbrakk> mutators_phase_inv s\<down>;
no_grey_refs s\<down>;
gc.handshake_invL s;
at gc ''mark_loop'' s \<or> at gc ''mark_loop_get_roots_load_W'' s \<or> at gc ''mark_loop_get_work_load_W'' s \<or> atS gc no_grey_refs_locs s;
r \<in> mut_roots s\<down> \<union> mut_ghost_honorary_root s\<down>; white r s\<down>;
handshake_phase_inv s\<down> \<rbrakk>
\<Longrightarrow> no_grey_refs (s\<down>(mutator m := s\<down> (mutator m)\<lparr>ghost_honorary_grey := {r}\<rparr>))"
apply (elim disjE)
apply (clarsimp simp: atS_simps gc.handshake_invL_def loc)
apply (frule handshake_phase_invD)
apply (clarsimp simp: hp_step_rel_def)
apply (drule spec[where x=m])
apply (clarsimp simp: conj_disj_distribR[symmetric])
apply (drule (2) no_grey_refs_not_rootD)
apply blast
apply (clarsimp simp: atS_simps gc.handshake_invL_def loc)
apply (frule handshake_phase_invD)
apply (clarsimp simp: hp_step_rel_def)
apply (drule spec[where x=m])
apply (clarsimp simp: conj_disj_distribR[symmetric])
apply (drule (2) no_grey_refs_not_rootD)
apply blast
apply (clarsimp simp: atS_simps gc.handshake_invL_def loc)
apply (frule handshake_phase_invD)
apply (clarsimp simp: hp_step_rel_def)
apply (drule spec[where x=m])
apply clarsimp
apply (drule (2) no_grey_refs_not_rootD)
apply blast
apply (frule (2) handshake_sweep_mark_endD)
apply (drule spec[where x=m])
apply clarsimp
apply (drule (2) no_grey_refs_not_rootD)
apply blast
done
lemma (in mut_m) gc_W_empty_invL[intro]:
notes gc.gc_W_empty_invL_def[inv]
shows
"\<lbrace> handshake_invL \<^bold>\<and> mark_object_invL \<^bold>\<and> tso_lock_invL
\<^bold>\<and> mut_get_roots.mark_object_invL m
\<^bold>\<and> mut_store_del.mark_object_invL m
\<^bold>\<and> mut_store_ins.mark_object_invL m
\<^bold>\<and> gc.handshake_invL \<^bold>\<and> gc.obj_fields_marked_invL
\<^bold>\<and> gc.gc_W_empty_invL
\<^bold>\<and> LSTP (handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> valid_W_inv) \<rbrace>
mutator m
\<lbrace> gc.gc_W_empty_invL \<rbrace>"
apply vcg_nihe
(* apply vcg_ni -- very slow *)
apply(tactic \<open>
let val ctxt = @{context} in
TRY (HEADGOAL (vcg_clarsimp_tac ctxt))
THEN
PARALLEL_ALLGOALS (
vcg_sem_tac ctxt
THEN' (TRY o SELECT_GOAL (Local_Defs.unfold_tac ctxt (Named_Theorems.get ctxt @{named_theorems inv})))
THEN' (TRY o REPEAT_ALL_NEW (Tactic.match_tac ctxt @{thms conjI})) (* expose the location predicates, do not split the consequents *)
THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Tactic.match_tac ctxt @{thms impI}))
(* Preserve the label sets in atS but normalise the label in at; turn s' into s *)
THEN_ALL_NEW full_simp_tac ctxt (* FIXME vcg_ni uses asm_full_simp_tac here *)
THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Tactic.ematch_tac ctxt @{thms conjE}))
(* The effect of vcg_pre: should be cheap *)
THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Tactic.ematch_tac ctxt @{thms thin_locs} THEN' REPEAT1 o assume_tac ctxt))
THEN_ALL_NEW asm_full_simp_tac (ss_only (@{thms loc_simps} @ Named_Theorems.get ctxt @{named_theorems loc}) ctxt)
THEN_ALL_NEW (TRY o REPEAT_ALL_NEW (Rule_Insts.thin_tac ctxt "True" []))
THEN_ALL_NEW clarsimp_tac ctxt)
end
\<close>)
(* hs_noop_done *)
apply (clarsimp simp: atS_un gc.handshake_invL_def)
(* hs_get_roots_done: gc_W_empty *)
apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def)
apply (rule conjI)
apply (clarsimp simp: WL_def)
apply clarsimp
apply (drule mp)
apply blast
apply (clarsimp simp: WL_def)
apply (rename_tac xa)
apply (rule_tac x=xa in exI)
apply clarsimp
(* hs_get_roots_done: no_grey_refs *)
apply (simp add: no_grey_refs_def)
apply (simp add: no_grey_refs_def)
(* hs_get_work_done: gc_W_empty *)
apply (clarsimp simp: mut_m.gc_W_empty_mut_inv_def)
apply (rule conjI)
apply (clarsimp simp: WL_def)
apply clarsimp
apply (drule mp)
apply blast
apply (clarsimp simp: WL_def)
apply (rename_tac xa)
apply (rule_tac x=xa in exI)
apply clarsimp
(* hs_get_work_done: no_grey_refs *)
apply (simp add: no_grey_refs_def)
apply (simp add: no_grey_refs_def)
done
(*>*)
subsection\<open>Sweep loop invariants\<close>
locset_definition "sweep_loop_locs = prefixed ''sweep_loop''"
inv_definition (in gc) sweep_loop_invL :: "('field, 'mut, 'ref) gc_pred" where
"sweep_loop_invL =
(at_gc ''sweep_loop_check'' ( (\<^bold>\<not>(NULL gc_mark) \<^bold>\<longrightarrow> (\<lambda>s. obj_at (\<lambda>obj. Some (obj_mark obj) = gc_mark s) (gc_tmp_ref s) s))
\<^bold>\<and> ( NULL gc_mark \<^bold>\<longrightarrow> valid_ref \<^bold>$ gc_tmp_ref \<^bold>\<longrightarrow> marked \<^bold>$ gc_tmp_ref ) )
\<^bold>\<and> at_gc ''sweep_loop_free'' ( \<^bold>\<not>(NULL gc_mark) \<^bold>\<and> the \<circ> gc_mark \<^bold>\<noteq> gc_fM \<^bold>\<and> (\<lambda>s. obj_at (\<lambda>obj. Some (obj_mark obj) = gc_mark s) (gc_tmp_ref s) s) )
\<^bold>\<and> at_gc ''sweep_loop_ref_done'' (valid_ref \<^bold>$ gc_tmp_ref \<^bold>\<longrightarrow> marked \<^bold>$ gc_tmp_ref)
\<^bold>\<and> atS_gc sweep_loop_locs (\<^bold>\<forall>r. \<^bold>\<not>(\<langle>r\<rangle> \<^bold>\<in> gc_refs) \<^bold>\<longrightarrow> valid_ref r \<^bold>\<longrightarrow> marked r)
\<^bold>\<and> atS_gc black_heap_locs (\<^bold>\<forall>r. valid_ref r \<^bold>\<longrightarrow> marked r)
\<^bold>\<and> atS_gc (prefixed ''sweep_loop_'' - { ''sweep_loop_choose_ref'' }) (gc_tmp_ref \<^bold>\<in> gc_refs))"
(*<*)
lemma (in gc) sweep_loop_invL_eq_imp:
"eq_imp (\<lambda>(_::unit) (s :: ('field, 'mut, 'ref) gc_pred_state). (AT s gc, s\<down> gc, sys_fM s\<down>, Option.map_option obj_mark \<circ> sys_heap s\<down>))
sweep_loop_invL"
apply (clarsimp simp: eq_imp_def inv)
apply (rename_tac s s')
apply (subgoal_tac "\<forall>r. valid_ref r s\<down> \<longleftrightarrow> valid_ref r s'\<down>")
apply (subgoal_tac "\<forall>P r. obj_at (\<lambda>obj. P (obj_mark obj)) r s\<down> \<longleftrightarrow> obj_at (\<lambda>obj. P (obj_mark obj)) r s'\<down>")
apply (frule_tac x="\<lambda>mark. Some mark = gc_mark s'\<down>" in spec)
apply (frule_tac x="\<lambda>mark. mark = sys_fM s'\<down>" in spec)
apply clarsimp
apply (clarsimp simp: fun_eq_iff split: obj_at_splits)
apply (rename_tac r)
apply ( (drule_tac x=r in spec)+, auto)[1]
apply (clarsimp simp: fun_eq_iff split: obj_at_splits)
apply (rename_tac r)
apply (drule_tac x=r in spec, auto)[1]
apply (metis map_option_eq_Some)+
done
lemmas gc_sweep_loop_invL_niE[nie] =
iffD1[OF gc.sweep_loop_invL_eq_imp[simplified eq_imp_simps, rule_format, unfolded conj_explode, rule_format], rotated -1]
lemma (in gc) sweep_loop_invL[intro]:
"\<lbrace> fM_fA_invL \<^bold>\<and> phase_invL \<^bold>\<and> sweep_loop_invL \<^bold>\<and> tso_lock_invL
\<^bold>\<and> LSTP (phase_rel_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> valid_W_inv) \<rbrace>
gc
\<lbrace> sweep_loop_invL \<rbrace>"
apply (vcg_jackhammer simp: no_grey_refs_def phase_rel_inv_def phase_rel_def)
apply (rename_tac s s' x)
apply (case_tac "x \<in> gc_refs s\<down>")
apply force
apply blast
apply (clarsimp split: obj_at_splits)
done
lemma sweep_loop_sweep_locs[iff]:
"sweep_loop_locs \<subseteq> sweep_locs"
by (auto simp: sweep_loop_locs_def sweep_locs_def intro: append_prefixD)
lemma sweep_locs_subseteq_fM_tso_empty_locs[iff]:
"sweep_locs \<subseteq> fM_tso_empty_locs"
by (auto simp: sweep_locs_def fM_tso_empty_locs_def)
lemma sweep_loop_locs_fM_eq_locs:
"sweep_loop_locs \<subseteq> fM_eq_locs"
by (auto simp: sweep_loop_locs_def fM_eq_locs_def sweep_locs_def)
lemma sweep_loop_locs_fA_eq_locs:
"sweep_loop_locs \<subseteq> fA_eq_locs"
apply (simp add: sweep_loop_locs_def fA_eq_locs_def sweep_locs_def)
apply (intro subset_insertI2)
apply (auto intro: append_prefixD)
done
lemma black_heap_locs_subseteq_fM_tso_empty_locs[iff]:
"black_heap_locs \<subseteq> fM_tso_empty_locs"
by (auto simp: black_heap_locs_def fM_tso_empty_locs_def)
lemma black_heap_locs_fM_eq_locs:
"black_heap_locs \<subseteq> fM_eq_locs"
by (simp add: black_heap_locs_def fM_eq_locs_def)
lemma black_heap_locs_fA_eq_locs:
"black_heap_locs \<subseteq> fA_eq_locs"
by (simp add: black_heap_locs_def fA_eq_locs_def sweep_locs_def)
lemma (in gc) fM_invL_tso_emptyD:
"\<lbrakk> atS gc ls s; fM_fA_invL s; ls \<subseteq> fM_tso_empty_locs \<rbrakk> \<Longrightarrow> tso_pending_fM gc s\<down> = []"
by (auto simp: fM_fA_invL_def dest: atS_mono)
lemma gc_sweep_loop_invL_locsE[rule_format]:
"(atS gc (sweep_locs \<union> black_heap_locs) s \<longrightarrow> False) \<Longrightarrow> gc.sweep_loop_invL s"
apply (simp add: gc.sweep_loop_invL_def atS_un)
apply (auto simp: loc atS_simps dest: atS_mono)
apply (clarsimp simp: atS_def)
apply (rename_tac x)
apply (drule_tac x=x in bspec)
apply (auto simp: sweep_locs_def intro: append_prefixD)
done
lemma (in sys) gc_sweep_loop_invL[intro]:
"\<lbrace> gc.fM_fA_invL \<^bold>\<and> gc.gc_W_empty_invL \<^bold>\<and> gc.handshake_invL \<^bold>\<and> gc.phase_invL \<^bold>\<and> gc.sweep_loop_invL
\<^bold>\<and> LSTP (mutators_phase_inv \<^bold>\<and> tso_writes_inv \<^bold>\<and> valid_W_inv) \<rbrace>
sys
\<lbrace> gc.sweep_loop_invL \<rbrace>"
apply vcg_nihe
apply vcg_ni
apply (clarsimp simp: do_write_action_def
split: mem_write_action.splits
elim: gc_sweep_loop_invL_niE) (* FIXME elimination rule using tso_writes_inv *)
(* mw_Mark *)
apply (rename_tac s s' p ws ref bool)
apply (rule gc_sweep_loop_invL_locsE)
apply (simp only: gc.gc_W_empty_invL_def no_grey_refs_locs_def cong del: atS_state_cong)
apply (clarsimp simp: atS_un)
apply (erule disjE)
apply clarsimp
apply (drule (1) no_grey_refs_no_pending_marks)
apply (clarsimp simp: filter_empty_conv)
apply (drule_tac x=p in spec)
apply fastforce
apply clarsimp
apply (drule (1) no_grey_refs_no_pending_marks)
apply (clarsimp simp: filter_empty_conv)
apply (drule_tac x=p in spec)
apply fastforce
(* mw_Mutate *)
apply (erule gc_sweep_loop_invL_niE, simp_all add: fun_eq_iff)[1] (* FIXME should be automatic *)
(* mw_fA *)
apply (erule gc_sweep_loop_invL_niE, simp_all add: fun_eq_iff)[1] (* FIXME should be automatic *)
(* mw_fM *)
apply (rename_tac s s' p ws bool)
apply (rule gc_sweep_loop_invL_locsE)
apply (case_tac p, clarsimp+)
apply (drule (1) gc.fM_invL_tso_emptyD)
apply simp_all
(* mv_Phase *)
apply (erule gc_sweep_loop_invL_niE, simp_all add: fun_eq_iff)[1] (* FIXME should be automatic *)
done (* FIXME weird: expect more aggressive use of gc_sweep_loop_invL_niE by clarsimp *)
lemma (in mut_m) gc_sweep_loop_invL[intro]:
"\<lbrace> gc.fM_fA_invL \<^bold>\<and> gc.handshake_invL \<^bold>\<and> gc.sweep_loop_invL
\<^bold>\<and> LSTP (mutators_phase_inv \<^bold>\<and> valid_refs_inv) \<rbrace>
mutator m
\<lbrace> gc.sweep_loop_invL \<rbrace>"
apply vcg_nihe
apply vcg_ni
apply (clarsimp simp: inv gc.sweep_loop_invL_def gc.fM_fA_invL_def)
apply (intro allI conjI impI)
(* four subgoals *)
apply ((erule (1) thin_locs)+)[1]
apply (force simp: loc)
apply ((erule (1) thin_locs)+)[1]
apply (force simp: loc)
apply (rename_tac s s' ra)
apply (drule mp, erule atS_mono[OF _ sweep_loop_locs_fA_eq_locs])
apply (drule mp, erule atS_mono[OF _ sweep_loop_locs_fM_eq_locs])
apply force
apply (rename_tac s s' ra)
apply (drule mp, erule atS_mono[OF _ black_heap_locs_fA_eq_locs])
apply (drule mp, erule atS_mono[OF _ black_heap_locs_fM_eq_locs])
apply force
done (* FIXME crappy split *)
lemma (in gc) sys_phase_inv[intro]:
"\<lbrace> fM_fA_invL \<^bold>\<and> gc_W_empty_invL \<^bold>\<and> handshake_invL \<^bold>\<and> obj_fields_marked_invL
\<^bold>\<and> phase_invL \<^bold>\<and> sweep_loop_invL
\<^bold>\<and> LSTP (phase_rel_inv \<^bold>\<and> sys_phase_inv \<^bold>\<and> valid_W_inv \<^bold>\<and> tso_writes_inv) \<rbrace>
gc
\<lbrace> LSTP sys_phase_inv \<rbrace>"
apply vcg_jackhammer
apply ( (erule black_heap_no_greys, simp)
| fastforce dest!: phase_rel_invD simp: phase_rel_def filter_empty_conv )+
done
lemma (in gc) no_black_refs_sweep_loop_free[simp]:
"no_black_refs s \<Longrightarrow> no_black_refs (s(sys := s sys\<lparr>heap := (sys_heap s)(gc_tmp_ref s := None)\<rparr>))"
by (simp add: no_black_refs_def)
lemma (in gc) no_black_refs_load_W[simp]:
"\<lbrakk> no_black_refs s; gc_W s = {} \<rbrakk>
\<Longrightarrow> no_black_refs (s(gc := s gc\<lparr>W := sys_W s\<rparr>, sys := s sys\<lparr>W := {}\<rparr>))"
by (simp add: no_black_refs_def)
lemma marked_deletions_sweep_loop_free[simp]:
"\<lbrakk> mut_m.marked_deletions m s; mut_m.reachable_snapshot_inv m s; no_grey_refs s; white r s \<rbrakk>
\<Longrightarrow> mut_m.marked_deletions m (s(sys := s sys\<lparr>heap := (sys_heap s)(r := None)\<rparr>))"
apply (clarsimp simp: mut_m.marked_deletions_def split: mem_write_action.splits)
apply (rename_tac ref field option)
apply (drule_tac x="mw_Mutate ref field option" in spec)
apply clarsimp
apply (clarsimp simp: obj_at_field_on_heap_def split: option.splits)
apply (rule conjI)
apply (clarsimp simp: mut_m.reachable_snapshot_inv_def)
apply (drule spec[where x=r], clarsimp simp: in_snapshot_def)
apply (drule mp, auto simp: mut_m.reachable_def mut_m.tso_write_refs_def split: mem_write_action.splits)[1] (* FIXME rule *)
apply (drule grey_protects_whiteD)
apply (clarsimp simp: no_grey_refs_def)
apply clarsimp
apply (rule conjI)
apply clarsimp
apply clarsimp
apply (rule conjI)
apply (clarsimp simp: mut_m.reachable_snapshot_inv_def)
apply (drule spec[where x=r], clarsimp simp: in_snapshot_def)
apply (drule mp, auto simp: mut_m.reachable_def mut_m.tso_write_refs_def split: mem_write_action.splits)[1] (* FIXME rule *)
apply (drule grey_protects_whiteD)
apply (clarsimp simp: no_grey_refs_def)
apply (clarsimp split: obj_at_splits)
done
lemma (in gc) mutator_phase_inv[intro]:
"\<lbrace> fM_fA_invL \<^bold>\<and> gc_W_empty_invL \<^bold>\<and> handshake_invL \<^bold>\<and> obj_fields_marked_invL \<^bold>\<and> sweep_loop_invL
\<^bold>\<and> gc_mark.mark_object_invL
\<^bold>\<and> LSTP (handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> tso_writes_inv \<^bold>\<and> valid_refs_inv \<^bold>\<and> valid_W_inv) \<rbrace>
gc
\<lbrace> LSTP (mut_m.mutator_phase_inv m) \<rbrace>"
apply vcg_jackhammer
apply (simp_all add: mutator_phase_inv_aux_case split: handshake_phase.splits)
(* sweep_loop_free *)
apply (intro allI conjI impI)
apply (drule mut_m.handshake_phase_invD[where m=m], clarsimp simp: hp_step_rel_def)
apply (rule mut_m.reachable_snapshot_inv_sweep_loop_free, simp_all)
(* ''mark_loop_get_work_load_W'' *)
apply clarsimp
apply (drule spec[where x=m])
apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) (* FIXME rule *)
(* mark_loop_blacken *)
apply (drule spec[where x=m])
apply clarsimp
apply (intro allI conjI impI)
apply clarsimp
apply (drule mut_m.handshake_phase_invD[where m=m], clarsimp simp: hp_step_rel_def)
apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def)
apply (rename_tac s s' x)
apply (drule_tac x=x in spec)
apply clarsimp
apply (rename_tac s s' x xa)
apply (drule_tac x=xa in spec)
apply clarsimp
apply (rule obj_fields_marked_inv_has_white_path_to_blacken, simp_all)[1]
(* mark_loop_mo_co_mark *)
apply clarsimp
apply (erule mut_m.reachable_snapshot_inv_mo_co_mark) (* FIXME hoist to the top level *)
apply blast
(* mark_loop_get_roots_load_W *)
apply clarsimp
apply (drule spec[where x=m])
apply (clarsimp simp: mut_m.reachable_snapshot_inv_def in_snapshot_def grey_protects_white_def) (* FIXME rule *)
done
lemma (in mut_m) sys_phase_inv[intro]:
"\<lbrace> handshake_invL
\<^bold>\<and> mark_object_invL
\<^bold>\<and> mut_get_roots.mark_object_invL m
\<^bold>\<and> mut_store_del.mark_object_invL m
\<^bold>\<and> mut_store_ins.mark_object_invL m
\<^bold>\<and> LSTP (fA_rel_inv \<^bold>\<and> fM_rel_inv \<^bold>\<and> handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> phase_rel_inv \<^bold>\<and> sys_phase_inv \<^bold>\<and> valid_refs_inv) \<rbrace>
mutator m
\<lbrace> LSTP sys_phase_inv \<rbrace>"
apply (vcg_jackhammer simp: fA_rel_inv_def fM_rel_inv_def)
apply (simp_all add: sys_phase_inv_aux_case heap_colours_colours
split: handshake_phase.splits if_splits)
(* alloc *)
subgoal by (clarsimp simp: fA_rel_def fM_rel_def no_black_refs_def
dest!: handshake_phase_invD phase_rel_invD
split: handshake_phase.splits)
(* store_ins_mo_co_mark *)
subgoal by (fastforce simp: fA_rel_def fM_rel_def hp_step_rel_def
dest!: handshake_phase_invD phase_rel_invD)
subgoal
apply (drule spec[where x=m])
apply (rule conjI)
apply (clarsimp simp: hp_step_rel_def phase_rel_def conj_disj_distribR[symmetric]
dest!: handshake_phase_invD phase_rel_invD)
apply (elim disjE, simp_all add: no_grey_refs_not_rootD)[1]
apply (clarsimp simp: hp_step_rel_def phase_rel_def
dest!: handshake_phase_invD phase_rel_invD)
apply (elim disjE, simp_all add: no_grey_refs_not_rootD)[1]
apply clarsimp
apply (elim disjE, simp_all add: no_grey_refs_not_rootD filter_empty_conv)[1]
apply fastforce
done
(* store_del_mo_co_unlock *)
subgoal by (fastforce simp: fA_rel_def fM_rel_def hp_step_rel_def
dest!: handshake_phase_invD phase_rel_invD)
subgoal
apply (drule spec[where x=m])
apply (rule conjI)
apply (clarsimp simp: hp_step_rel_def phase_rel_def conj_disj_distribR[symmetric] no_grey_refs_not_rootD
dest!: handshake_phase_invD phase_rel_invD)
apply (clarsimp simp: hp_step_rel_def phase_rel_def
dest!: handshake_phase_invD phase_rel_invD)
apply (elim disjE, simp_all add: no_grey_refs_not_rootD)[1]
apply clarsimp
apply (elim disjE, simp_all add: no_grey_refs_not_rootD filter_empty_conv)[1]
apply fastforce
done
(* hs_get_roots_done *)
subgoal by (clarsimp simp: hp_step_rel_def
dest!: handshake_phase_invD phase_rel_invD)
subgoal by (clarsimp simp: hp_step_rel_def
dest!: handshake_phase_invD phase_rel_invD)
subgoal by (clarsimp simp: hp_step_rel_def
dest!: handshake_phase_invD phase_rel_invD)
subgoal by (clarsimp simp: hp_step_rel_def
dest!: handshake_phase_invD phase_rel_invD)
subgoal
apply (clarsimp simp: hp_step_rel_def phase_rel_def filter_empty_conv
dest!: handshake_phase_invD phase_rel_invD)
apply auto
done
(* hs_get_roots_loop_mo_co_mark *)
subgoal by (clarsimp simp: hp_step_rel_def
dest!: handshake_phase_invD phase_rel_invD)
subgoal
apply (clarsimp simp: hp_step_rel_def phase_rel_def filter_empty_conv
dest!: handshake_phase_invD phase_rel_invD)
apply auto
done
(* hs_get_work_done *)
subgoal by (clarsimp simp: hp_step_rel_def
dest!: handshake_phase_invD phase_rel_invD)
subgoal by (clarsimp simp: hp_step_rel_def
dest!: handshake_phase_invD phase_rel_invD)
subgoal by (clarsimp simp: hp_step_rel_def
dest!: handshake_phase_invD phase_rel_invD)
subgoal by (clarsimp simp: hp_step_rel_def
dest!: handshake_phase_invD phase_rel_invD)
subgoal
apply (clarsimp simp: hp_step_rel_def phase_rel_def filter_empty_conv
dest!: handshake_phase_invD phase_rel_invD)
apply auto
done
done
lemma no_grey_refs_no_marks[simp]:
"\<lbrakk> no_grey_refs s; valid_W_inv s \<rbrakk> \<Longrightarrow> \<not>sys_mem_write_buffers p s = mw_Mark r fl # ws"
by (auto simp: no_grey_refs_def)
context sys
begin
lemma black_heap_dequeue_mark[iff]:
"\<lbrakk> sys_mem_write_buffers p s = mw_Mark r fl # ws; black_heap s; valid_W_inv s \<rbrakk>
\<Longrightarrow> black_heap (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (obj_mark_update (\<lambda>_. fl)) (sys_heap s r)), mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))"
by (auto simp: black_heap_def)
lemma black_heap_dequeue_ref[iff]:
"\<lbrakk> sys_mem_write_buffers p s = mw_Mutate r f r' # ws; black_heap s \<rbrakk>
\<Longrightarrow> black_heap (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (\<lambda>obj. obj\<lparr>obj_fields := (obj_fields obj)(f := r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))"
by (simp add: black_heap_def black_def)
lemma white_heap_dequeue_fM[iff]:
"black_heap s\<down>
\<Longrightarrow> white_heap (s\<down>(sys := s\<down> sys\<lparr>fM := \<not> sys_fM s\<down>, mem_write_buffers := (mem_write_buffers (s\<down> sys))(gc := ws)\<rparr>))"
by (auto simp: black_heap_def white_heap_def)
lemma black_heap_dequeue_fM[iff]:
"\<lbrakk> white_heap s\<down>; no_grey_refs s\<down> \<rbrakk>
\<Longrightarrow> black_heap (s\<down>(sys := s\<down> sys\<lparr>fM := \<not> sys_fM s\<down>, mem_write_buffers := (mem_write_buffers (s\<down> sys))(gc := ws)\<rparr>))"
by (auto simp: black_heap_def white_heap_def no_grey_refs_def)
lemma white_heap_dequeue_ref[iff]:
"\<lbrakk> sys_mem_write_buffers p s = mw_Mutate r f r' # ws; white_heap s \<rbrakk>
\<Longrightarrow> white_heap (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (\<lambda>obj. obj\<lparr>obj_fields := (obj_fields obj)(f := r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))"
by (simp add: white_heap_def)
lemma (in sys) sys_phase_inv[intro]:
"\<lbrace> LSTP (fA_rel_inv \<^bold>\<and> fM_rel_inv \<^bold>\<and> handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> phase_rel_inv \<^bold>\<and> sys_phase_inv \<^bold>\<and> tso_writes_inv \<^bold>\<and> valid_W_inv) \<rbrace>
sys
\<lbrace> LSTP sys_phase_inv \<rbrace>"
by (vcg_jackhammer simp: fA_rel_inv_def fM_rel_inv_def p_not_sys)
(clarsimp simp: do_write_action_def sys_phase_inv_aux_case
split: mem_write_action.splits handshake_phase.splits if_splits;
erule disjE; clarsimp simp: fA_rel_def fM_rel_def; fail)
end
lemma valid_W_inv_unlockE[iff]:
"\<lbrakk> sys_mem_lock s = Some p; sys_mem_write_buffers p s = [];
\<And>r. r \<in> ghost_honorary_grey (s p) \<Longrightarrow> marked r s;
valid_W_inv s
\<rbrakk> \<Longrightarrow> valid_W_inv (s(sys := mem_lock_update Map.empty (s sys)))"
unfolding valid_W_inv_def by clarsimp (metis emptyE empty_set)
lemma valid_W_inv_mark:
"\<lbrakk> sys_mem_lock s = Some p; white w s; valid_W_inv s \<rbrakk>
\<Longrightarrow> w \<in> ghost_honorary_grey (s p) \<or> (\<forall>q. w \<notin> WL q s)"
by (metis Un_iff WL_def marked_not_white option.inject valid_W_invD(1) valid_W_invD3(2))
lemma (in gc) valid_W_inv[intro]:
notes valid_W_invD2[dest!, simp]
notes valid_W_invD3[dest, simp]
shows
"\<lbrace> fM_invL \<^bold>\<and> gc_mark.mark_object_invL \<^bold>\<and> gc_W_empty_invL
\<^bold>\<and> obj_fields_marked_invL
\<^bold>\<and> sweep_loop_invL \<^bold>\<and> tso_lock_invL
\<^bold>\<and> LSTP (tso_writes_inv \<^bold>\<and> valid_W_inv) \<rbrace>
gc
\<lbrace> LSTP valid_W_inv \<rbrace>"
apply (vcg_jackhammer simp: fM_rel_def)
(* sweep loop free: what's with the case splitting? *)
subgoal for s s'
apply (subst valid_W_inv_def)
apply (intro allI conjI impI; clarsimp simp: p_not_sys split: if_splits)
apply blast
apply blast
apply blast
apply blast
apply (rename_tac p x)
apply (case_tac p; auto)
apply (rename_tac p)
apply (case_tac p; force simp: no_grey_refs_def)
done
(* mark_loop_get_work_load_W *)
subgoal by (subst valid_W_inv_def) (auto simp: all_conj_distrib split: process_name.splits)
(* mark_loop_blacken *)
subgoal by (subst valid_W_inv_def) (auto simp: all_conj_distrib)
(* mark_loop_mo_co_W *)
subgoal
apply (subst valid_W_inv_def)
apply clarsimp
apply blast
done
(* mark_loop_mo_co_unlock *)
subgoal
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
apply blast
apply blast
apply (auto iff: p_not_sys split: if_splits)[1]
apply blast
apply blast
apply blast
apply force
done
(* mark_loop_mo_co_mark *)
subgoal
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
apply blast
apply blast
apply blast
apply (frule (2) valid_W_inv_mark; auto)[1]
apply (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) (* FIXME want a cheaper contradiction between white and marked *)
apply (clarsimp dest!: valid_W_invD(1) split: obj_at_splits)
apply (clarsimp dest!: valid_W_invD(1) split: obj_at_splits)
apply blast
apply blast
apply blast
apply blast
apply blast
apply blast
apply force
done
(* mark_loop_mo_co_lock *)
subgoal
apply (subst valid_W_inv_def)
apply (auto simp: all_conj_distrib)
done
(* ''mark_loop_get_roots_load_W'' *)
subgoal
apply (subst valid_W_inv_def)
apply (auto simp: all_conj_distrib split: process_name.splits)
done
done
lemma (in mut_m) valid_W_inv[intro]:
notes valid_W_invD2[dest!, simp]
notes valid_W_invD3[dest, simp]
shows
"\<lbrace> handshake_invL \<^bold>\<and> mark_object_invL \<^bold>\<and> tso_lock_invL
\<^bold>\<and> mut_get_roots.mark_object_invL m
\<^bold>\<and> mut_store_del.mark_object_invL m
\<^bold>\<and> mut_store_ins.mark_object_invL m
\<^bold>\<and> LSTP (fM_rel_inv \<^bold>\<and> sys_phase_inv \<^bold>\<and> tso_writes_inv \<^bold>\<and> valid_refs_inv \<^bold>\<and> valid_W_inv) \<rbrace>
mutator m
\<lbrace> LSTP valid_W_inv \<rbrace>"
apply (vcg_jackhammer simp: fM_rel_inv_def fM_rel_def)
(* alloc *)
subgoal
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI; auto)
done
(* store ins mo co W *)
subgoal
apply (subst valid_W_inv_def)
apply clarsimp
apply blast
done
(* store ins mo co unlock *)
subgoal for s s' y
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal for p x by (case_tac "p = mutator m"; force split: if_splits)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by force
done
(* store ins mo co mark *)
subgoal
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by (blast dest: valid_W_inv_mark)
subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) (* FIXME want a cheaper contradiction between white and marked *)
subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits)
subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by force
done
(* store ins mo co lock *)
subgoal
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal by force
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by force
subgoal by blast
subgoal by force
done
(* store del mo co W *)
subgoal
apply (subst valid_W_inv_def)
apply clarsimp
apply blast
done
(* store del mo co unlock *)
subgoal for s s' y
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal for p x by (cases "p = mutator m") (auto split: if_splits)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by force
done
(* store del mo co mark *)
subgoal
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by (frule (2) valid_W_inv_mark; auto)
subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) (* FIXME want a cheaper contradiction between white and marked *)
subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits)
subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by force
done
(* store del mo co lock *)
subgoal
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal by force
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by force
subgoal by blast
subgoal by force
done
(* get roots done *)
subgoal
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by (auto split: if_splits obj_at_splits process_name.splits)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by (auto split: if_splits obj_at_splits process_name.splits)
subgoal by blast
subgoal by blast
done
(* hs get roots loop mo co W *)
subgoal
apply (subst valid_W_inv_def)
apply clarsimp
apply blast
done
(* hs get roots loop mo co unlock *)
subgoal for s s' y
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal for p x by (cases "p = mutator m") (auto split: if_splits)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by force
done
(* hs get roots loop mo co mark *)
subgoal
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by (frule (2) valid_W_inv_mark; auto)
subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits) (* FIXME want a cheaper contradiction between white and marked *)
subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits)
subgoal by (clarsimp dest!: valid_W_invD(1) split: obj_at_splits)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by force
done
(* hs get roots loop mo co lock *)
subgoal
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal by force
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by force
subgoal by blast
subgoal by force
done
(* hs get work done *)
subgoal
apply (subst valid_W_inv_def)
apply (clarsimp simp: all_conj_distrib)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by (auto split: if_splits obj_at_splits process_name.splits)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by force
subgoal by blast
subgoal by blast
done
done
lemma (in sys) valid_W_inv[intro]:
notes valid_W_invD2[dest!, simp]
notes valid_W_invD3[dest, simp]
notes valid_W_invD4[dest!, simp]
notes o_def[simp]
shows
"\<lbrace> LSTP (fM_rel_inv \<^bold>\<and> sys_phase_inv \<^bold>\<and> tso_writes_inv \<^bold>\<and> valid_refs_inv \<^bold>\<and> valid_W_inv) \<rbrace>
sys
\<lbrace> LSTP valid_W_inv \<rbrace>"
apply vcg_jackhammer
apply (subst valid_W_inv_def)
apply (clarsimp simp: do_write_action_def all_conj_distrib fM_rel_inv_def
split: mem_write_action.splits)
(* mw_Mark *)
subgoal
apply (intro allI conjI impI)
apply blast
apply blast
apply blast
apply blast
apply blast
apply force
apply blast
apply blast
apply blast
apply (force simp: filter_empty_conv)
apply force
apply blast
done
(* mw_Mutate, mw_fA *)
subgoal by (intro allI conjI impI; auto)
subgoal by (intro allI conjI impI; auto)
(* mw_fM *)
subgoal for s s' p ws bool
apply (clarsimp simp: fM_rel_def)
apply (rule disjE[OF iffD1[OF p_not_sys]], assumption)
prefer 2
apply clarsimp
apply (case_tac "sys_ghost_handshake_phase s\<down> = hp_Idle"; clarsimp)
apply (intro allI conjI impI)
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by blast
subgoal by force
subgoal by (force dest: no_grey_refs_no_pending_marks)
subgoal by blast
subgoal by force
subgoal by (fastforce dest: no_grey_refs_no_pending_marks)
subgoal by (fastforce dest: no_grey_refs_no_pending_marks)
done
(* mw_Phase *)
apply (intro allI conjI impI; auto)
done
lemma (in gc) strong_tricolour_inv[intro]:
"\<lbrace> fM_fA_invL \<^bold>\<and> gc_W_empty_invL \<^bold>\<and> gc_mark.mark_object_invL \<^bold>\<and> obj_fields_marked_invL \<^bold>\<and> sweep_loop_invL
\<^bold>\<and> LSTP (strong_tricolour_inv \<^bold>\<and> tso_writes_inv \<^bold>\<and> valid_W_inv) \<rbrace>
gc
\<lbrace> LSTP strong_tricolour_inv \<rbrace>"
apply (vcg_jackhammer simp: strong_tricolour_inv_def)
apply (fastforce elim!: obj_fields_marked_inv_blacken)
done
lemma (in mut_m) strong_tricolour[intro]:
"\<lbrace> mark_object_invL
\<^bold>\<and> mut_get_roots.mark_object_invL m
\<^bold>\<and> mut_store_del.mark_object_invL m
\<^bold>\<and> mut_store_ins.mark_object_invL m
\<^bold>\<and> LSTP (fA_rel_inv \<^bold>\<and> fM_rel_inv \<^bold>\<and> handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> strong_tricolour_inv \<^bold>\<and> sys_phase_inv \<^bold>\<and> valid_refs_inv) \<rbrace>
mutator m
\<lbrace> LSTP strong_tricolour_inv \<rbrace>"
apply (vcg_jackhammer simp: strong_tricolour_inv_def fA_rel_inv_def fM_rel_inv_def)
apply (drule handshake_phase_invD)
apply (clarsimp simp: sys_phase_inv_aux_case
split: handshake_phase.splits if_splits)
apply (blast dest: heap_colours_colours)
apply (blast dest: heap_colours_colours)
apply (clarsimp simp: fA_rel_def fM_rel_def no_black_refs_def) (* FIXME rule *)
apply (drule_tac x=m in spec)
apply (clarsimp simp: hp_step_rel_def)
apply (elim disjE, simp_all)
apply (auto dest: no_black_refsD)[1]
apply (auto dest: no_black_refsD)[1]
apply (clarsimp simp: fA_rel_def fM_rel_def)
apply (clarsimp split: obj_at_splits)
apply (drule spec[where x=m])
apply (clarsimp simp: hp_step_rel_def)
apply (elim disjE; force simp: fA_rel_def fM_rel_def split: obj_at_splits)
apply (drule spec[where x=m])
apply (clarsimp simp: hp_step_rel_def)
apply (elim disjE; force simp: fA_rel_def fM_rel_def split: obj_at_splits)
done
lemma (in sys) strong_tricolour_inv[intro]:
"\<lbrace> LSTP (fM_rel_inv \<^bold>\<and> handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> strong_tricolour_inv \<^bold>\<and> sys_phase_inv \<^bold>\<and> tso_writes_inv \<^bold>\<and> valid_W_inv) \<rbrace>
sys
\<lbrace> LSTP strong_tricolour_inv \<rbrace>"
+supply [[simproc del: defined_all]]
apply (vcg_jackhammer simp: strong_tricolour_inv_def p_not_sys)
apply (clarsimp simp: do_write_action_def fM_rel_inv_def
split: mem_write_action.splits)
apply (rename_tac ref field)
(* mw_Mark *)
subgoal for s s' p ws x xa ref field
apply (frule (1) valid_W_invD2)
apply clarsimp
apply (case_tac "x = ref", simp_all)
apply (clarsimp simp: grey_def WL_def)
done
(* mw_Mutate *)
apply (elim disjE, simp_all)
apply (clarsimp simp: points_to_mw_Mutate)
apply (elim disjE, simp_all)
apply (rename_tac s s' ws x xa ref field option m)
apply (drule_tac m=m in mut_m.handshake_phase_invD)
apply (frule_tac x=m in spec)
apply (clarsimp simp: mutator_phase_inv_aux_case hp_step_rel_def
split: handshake_phase.splits)
apply (elim disjE, simp_all split: if_splits)
apply (blast dest!: heap_colours_colours)
apply (blast dest!: heap_colours_colours)
apply (drule_tac x=m in spec)
apply (clarsimp simp: no_black_refs_def)
apply (drule_tac x=m in spec)
apply (clarsimp simp: no_black_refsD)
(* FIXME split less *)
apply (drule_tac x=m in spec)
apply (clarsimp simp: mut_m.marked_insertions_def)
apply (drule_tac x="mw_Mutate x field (Some xa)" in spec)
apply (clarsimp dest!: marked_not_white)
apply (drule_tac x=m in spec)
apply (clarsimp simp: mut_m.marked_insertions_def)
apply (drule_tac x="mw_Mutate x field (Some xa)" in spec)
apply (clarsimp dest!: marked_not_white)
apply (drule_tac x=m in spec)
apply (clarsimp simp: mut_m.marked_insertions_def)
apply (drule_tac x="mw_Mutate x field (Some xa)" in spec)
apply (clarsimp dest!: marked_not_white)
apply (drule_tac x=m in spec)
apply (clarsimp simp: mut_m.marked_insertions_def)
apply (drule_tac x="mw_Mutate x field (Some xa)" in spec)
apply (clarsimp dest!: marked_not_white)
apply (drule_tac x=m in spec)
apply (clarsimp simp: mut_m.marked_insertions_def)
apply (drule_tac x="mw_Mutate x field (Some xa)" in spec)
apply (clarsimp dest!: marked_not_white)
apply clarsimp
apply (elim disjE, simp_all split: if_splits)
apply (blast dest!: heap_colours_colours)
apply (blast dest!: heap_colours_colours)
apply (drule_tac x=m in spec)
apply (clarsimp simp: no_black_refs_def)
apply (drule_tac x=m in spec)
apply (clarsimp simp: mut_m.marked_insertions_def)
apply (drule_tac x="mw_Mutate x field (Some xa)" in spec)
apply (clarsimp dest!: marked_not_white)
apply (drule_tac x=m in spec)
apply (clarsimp simp: mut_m.marked_insertions_def)
apply (drule_tac x="mw_Mutate x field (Some xa)" in spec)
apply (clarsimp dest!: marked_not_white)
(* mw_fM *)
apply (rename_tac s s' p ws x xa bool)
apply (erule disjE)
apply (clarsimp simp: fM_rel_def black_heap_def split: if_splits)
apply ( (drule_tac x=x in spec)+ )[1]
apply (frule colours_distinct)
apply (clarsimp split: obj_at_splits)
apply (clarsimp simp: white_heap_def)
apply ( (drule_tac x=xa in spec)+ )[1]
apply (clarsimp split: obj_at_splits)
apply fastforce
done
text\<open>Remaining non-interference proofs.\<close>
lemma marked_insertionsD:
"\<lbrakk> sys_mem_write_buffers (mutator m) s = mw_Mutate r f (Some r') # ws; mut_m.marked_insertions m s \<rbrakk>
\<Longrightarrow> marked r' s"
by (clarsimp simp: mut_m.marked_insertions_def)
lemma mut_hs_get_roots_loop_locs_subseteq_hs_get_roots:
"mut_hs_get_roots_loop_locs \<subseteq> prefixed ''hs_get_roots_''"
by (auto simp: mut_hs_get_roots_loop_locs_def intro: append_prefixD)
lemma mut_m_handshake_invL_get_roots:
"\<lbrakk> mut_m.handshake_invL m s; atS (mutator m) mut_hs_get_roots_loop_locs s \<rbrakk>
\<Longrightarrow> sys_handshake_type s\<down> = ht_GetRoots \<and> sys_handshake_pending m s\<down>"
unfolding mut_m.handshake_invL_def
apply (elim conjE)
apply (drule mp, erule (1) atS_mono[OF _ mut_hs_get_roots_loop_locs_subseteq_hs_get_roots])
done
lemma subseteq_mut_mo_valid_ref_locs:
"prefixed ''store_del_mo'' \<union> {''lop_store_ins''} \<subseteq> mut_mo_valid_ref_locs"
by (auto simp: mut_mo_valid_ref_locs_def intro: append_prefixD)
lemma subseteq_mut_mo_valid_ref_locs2:
"prefixed ''store_ins'' \<subseteq> mut_mo_valid_ref_locs"
by (clarsimp simp: mut_mo_valid_ref_locs_def)
lemma mut_phase_inv:
"\<lbrakk> ghost_handshake_phase (s (mutator m)) = hp_Mark \<or> ghost_handshake_phase (s (mutator m)) = hp_IdleMarkSweep \<and> sys_phase s \<noteq> ph_Idle;
mut_m.mutator_phase_inv_aux m (ghost_handshake_phase (s (mutator m))) s \<rbrakk>
\<Longrightarrow> mut_m.marked_insertions m s \<and> mut_m.marked_deletions m s"
by (cases "ghost_handshake_phase (s (mutator m))", simp_all)
lemma (in sys) mut_mark_object_invL[intro]:
notes mut_m.mark_object_invL_def[inv]
notes do_write_action_fM[where m=m, simp]
notes do_write_action_prj_simps(1)[simp del] do_write_action_prj_simps(2)[simp del]
notes mut_m_get_roots_no_fM_write[where m=m, simp]
notes mut_m_get_roots_no_phase_write[where m=m, simp]
notes mut_m_ghost_handshake_phase_not_hp_Idle[where m=m, simp]
notes atS_simps[simp] filter_empty_conv[simp]
shows "\<lbrace> mut_m.handshake_invL m \<^bold>\<and> mut_m.mark_object_invL m
\<^bold>\<and> LSTP (fA_rel_inv \<^bold>\<and> fM_rel_inv \<^bold>\<and> handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> phase_rel_inv \<^bold>\<and> valid_refs_inv \<^bold>\<and> valid_W_inv \<^bold>\<and> tso_writes_inv) \<rbrace>
sys
\<lbrace> mut_m.mark_object_invL m \<rbrace>"
apply vcg_nihe
apply vcg_ni
subgoal by (clarsimp simp: do_write_action_def filter_empty_conv p_not_sys
dest!: valid_W_invD2
elim!: obj_at_weakenE
split: mem_write_action.splits)
subgoal by (clarsimp simp: do_write_action_def filter_empty_conv p_not_sys
dest!: valid_W_invD2
elim!: obj_at_weakenE
split: mem_write_action.splits)
subgoal by (clarsimp simp: do_write_action_def filter_empty_conv p_not_sys loc
dest!: valid_W_invD2
elim!: obj_at_weakenE
split: mem_write_action.splits)
subgoal
apply (drule phase_rel_invD)
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (clarsimp simp: p_not_sys)
apply (erule disjE)
apply (clarsimp simp: hp_step_rel_def)
apply (elim disjE, simp_all add: phase_rel_def)[1]
apply force
done
subgoal
apply (clarsimp simp: do_write_action_def filter_empty_conv p_not_sys loc fM_rel_inv_def
dest!: valid_W_invD2
elim!: obj_at_weakenE
split: mem_write_action.splits)
apply (rule conjI)
apply (clarsimp elim!: obj_at_weakenE)
apply clarsimp
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (erule disjE)
apply (auto simp: fM_rel_def hp_step_rel_def)[1]
apply force
apply (erule disjE)
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (drule phase_rel_invD)
apply (clarsimp simp: hp_step_rel_def)
apply (auto simp: phase_rel_def)[1]
apply force
done
subgoal for s s' p w ws
apply (drule mp, erule atS_mono[OF _ subseteq_mut_mo_valid_ref_locs])
apply ((thin_tac "atS p ls s \<longrightarrow> Q" for p ls s Q)+)[1]
apply ((thin_tac "at p ls s \<longrightarrow> Q" for p ls s Q)+)[1]
apply (clarsimp simp: do_write_action_def p_not_sys loc
split: mem_write_action.splits if_splits)
apply (drule (1) valid_W_invD2)
apply (erule obj_at_field_on_heap_weakenE)
apply (fastforce split: obj_at_splits)
apply (drule (1) valid_W_invD2)
apply (erule obj_at_field_on_heap_weakenE)
apply (fastforce split: obj_at_splits)
apply (drule spec[where x=m])
apply (frule (1) mut_phase_inv)
apply (rename_tac refa fielda option)
apply (frule_tac y="refa" in valid_refs_invD3, simp_all)[1]
apply (frule_tac y="tmp_ref (s\<down> (mutator m))" in valid_refs_invD(2), simp_all)[1]
apply (clarsimp split: obj_at_splits option.splits)
apply force
apply (frule (1) marked_insertionsD)
apply (auto split: obj_at_splits; fail)[1]
apply (erule disjE) (* super messy case *)
apply force
apply (rule conjI)
apply (erule obj_at_field_on_heap_imp_valid_ref)
apply (clarsimp split: option.splits)
apply (drule phase_rel_invD)
apply (frule mut_m.handshake_phase_invD[where m=m])
apply (rename_tac ma x2)
apply (drule_tac m=ma in mut_m.handshake_phase_invD)
apply (frule spec[where x=m])
apply (drule_tac x=ma in spec)
apply (clarsimp simp: hp_step_rel_def)
apply (elim disjE, auto simp: phase_rel_def dest: marked_insertionsD; fail)[1]
apply (erule disjE; clarsimp)
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (clarsimp simp: fM_rel_inv_def fM_rel_def hp_step_rel_def)
apply (metis (no_types, lifting) handshake_phase.distinct(5) handshake_phase.distinct(7))
apply (erule disjE; clarsimp)
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (drule phase_rel_invD)
apply (clarsimp simp: hp_step_rel_def phase_rel_def)
done
subgoal for s s' p w ws y
apply (clarsimp simp: do_write_action_def p_not_sys
split: mem_write_action.splits if_splits)
apply (auto split: obj_at_splits; fail)[1]
apply (erule disjE; clarsimp)
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (clarsimp simp: fM_rel_inv_def fM_rel_def hp_step_rel_def)
apply (metis (no_types, lifting) handshake_phase.distinct(5) handshake_phase.distinct(7))
apply (erule disjE; clarsimp)
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (drule phase_rel_invD)
apply (clarsimp simp: hp_step_rel_def phase_rel_def)
done
subgoal for s s' p w ws
apply (drule mp, erule atS_mono[OF _ subseteq_mut_mo_valid_ref_locs2])
apply ((thin_tac "atS p ls s \<longrightarrow> Q" for p ls s Q)+)[1]
apply ((thin_tac "at p ls s \<longrightarrow> Q" for p ls s Q)+)[1]
apply (subst do_write_action_fM[where m=m], simp_all)[1]
apply (elim disjE, simp_all)[1]
apply (clarsimp simp: do_write_action_def p_not_sys
split: mem_write_action.splits if_splits)
apply (erule obj_at_field_on_heap_weakenE, auto)[1]
apply (erule obj_at_field_on_heap_weakenE, auto split: obj_at_splits)[1]
apply (drule spec[where x=m])
apply (frule (1) mut_phase_inv)
apply (rename_tac refa fielda option)
apply (frule_tac y="refa" in valid_refs_invD3, simp_all)[1]
apply (frule_tac y="tmp_ref (s\<down> (mutator m))" in valid_refs_invD(2), simp_all)[1]
apply (clarsimp split: obj_at_splits option.splits)
apply auto[1]
apply (frule (1) marked_insertionsD)
apply (auto split: obj_at_splits)[1]
apply (erule disjE) (* super messy case *)
apply force
apply (rule conjI)
apply (erule obj_at_field_on_heap_imp_valid_ref)
apply (clarsimp split: option.splits)
apply (drule phase_rel_invD)
apply (frule mut_m.handshake_phase_invD[where m=m])
apply (rename_tac ma x2)
apply (drule_tac m=ma in mut_m.handshake_phase_invD)
apply (frule spec[where x=m])
apply (drule_tac x=ma in spec)
apply (clarsimp simp: hp_step_rel_def)
apply (elim disjE, auto simp: phase_rel_def dest: marked_insertionsD)[1]
apply (erule disjE)
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (drule phase_rel_invD)
apply (clarsimp simp: hp_step_rel_def phase_rel_def)
apply force
done
done
lemma (in gc) mut_mark_object_invL[intro]:
notes mut_m.mark_object_invL_def[inv]
notes atS_simps[simp]
shows "\<lbrace> fM_fA_invL \<^bold>\<and> gc_W_empty_invL \<^bold>\<and> handshake_invL \<^bold>\<and> sweep_loop_invL
\<^bold>\<and> mut_m.handshake_invL m
\<^bold>\<and> mut_m.mark_object_invL m
\<^bold>\<and> LSTP (fM_rel_inv \<^bold>\<and> handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> sys_phase_inv) \<rbrace>
gc
\<lbrace> mut_m.mark_object_invL m \<rbrace>"
apply vcg_nihe
apply vcg_ni
subgoal
apply (drule (1) mut_m_handshake_invL_get_roots)
apply clarsimp
done
subgoal by (simp add: mut_m.handshake_invL_def)
subgoal by (fastforce simp: fM_rel_inv_def fM_rel_def hp_step_rel_def split: obj_at_splits)
subgoal
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (drule spec[where x=m])
apply (clarsimp simp: valid_null_ref_def hp_step_rel_def conj_disj_distribR[symmetric] split: option.splits)
apply (drule (1) mut_m.reachable_blackD)
apply blast
apply (clarsimp split: obj_at_splits)
done
subgoal
apply (drule mp, erule atS_mono[OF _ subseteq_mut_mo_valid_ref_locs])
apply ((thin_tac "atS p ls s \<longrightarrow> Q" for p ls s Q)+)[1]
apply ((thin_tac "at p ls s \<longrightarrow> Q" for p ls s Q)+)[1]
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (drule spec[where x=m])
apply (clarsimp simp: valid_null_ref_def hp_step_rel_def conj_disj_distribR[symmetric] split: option.splits)
apply (drule (1) mut_m.reachable_blackD)
apply blast
apply (auto simp: obj_at_field_on_heap_def black_def split: obj_at_splits option.splits)[1]
done
subgoal by (clarsimp split: obj_at_splits)
subgoal
apply (drule mp, erule atS_mono[OF _ subseteq_mut_mo_valid_ref_locs2])
apply ((thin_tac "atS p ls s \<longrightarrow> Q" for p ls s Q)+)[1]
apply ((thin_tac "at p ls s \<longrightarrow> Q" for p ls s Q)+)[1]
apply (drule mut_m.handshake_phase_invD[where m=m])
apply (drule spec[where x=m])
apply (clarsimp simp: valid_null_ref_def hp_step_rel_def conj_disj_distribR[symmetric] split: option.splits)
apply (drule (1) mut_m.reachable_blackD)
apply blast
apply (auto simp: obj_at_field_on_heap_def black_def split: obj_at_splits option.splits)[1]
done
done
lemma (in gc) mut_store_old_mark_object_invL[intro]:
notes mut_m.mark_object_invL_def[inv]
shows
"\<lbrace> fM_fA_invL \<^bold>\<and> handshake_invL \<^bold>\<and> sweep_loop_invL \<^bold>\<and> gc_W_empty_invL
\<^bold>\<and> mut_m.mark_object_invL m
\<^bold>\<and> mut_store_del.mark_object_invL m
\<^bold>\<and> LSTP (handshake_phase_inv \<^bold>\<and> mut_m.mutator_phase_inv m) \<rbrace>
gc
\<lbrace> mut_store_del.mark_object_invL m \<rbrace>"
apply vcg_nihe
apply vcg_ni
apply ( (clarsimp dest!: mut_m.handshake_phase_invD[where m=m]
simp: hp_step_rel_def conj_disj_distribR[symmetric]
, drule_tac r="gc_tmp_ref s\<down>" in mut_m.no_grey_refs_not_rootD[where m=m]
, auto split: obj_at_splits if_splits)[1] )+
done
lemma (in gc) mut_store_ins_mark_object_invL[intro]:
notes mut_m.mark_object_invL_def[inv]
shows
"\<lbrace> fM_fA_invL \<^bold>\<and> handshake_invL \<^bold>\<and> sweep_loop_invL \<^bold>\<and> gc_W_empty_invL
\<^bold>\<and> mut_m.mark_object_invL m
\<^bold>\<and> mut_store_ins.mark_object_invL m
\<^bold>\<and> LSTP (handshake_phase_inv \<^bold>\<and> mut_m.mutator_phase_inv m) \<rbrace>
gc
\<lbrace> mut_store_ins.mark_object_invL m \<rbrace>"
apply vcg_nihe
apply vcg_ni
apply ( (clarsimp dest!: mut_m.handshake_phase_invD[where m=m]
simp: hp_step_rel_def conj_disj_distribR[symmetric]
, drule_tac r="gc_tmp_ref s\<down>" in mut_m.no_grey_refs_not_rootD[where m=m]
, auto split: obj_at_splits if_splits)[1] )+
done
lemma obj_fields_marked_locs_subseteq_hp_IdleMarkSweep_locs:
"obj_fields_marked_locs \<subseteq> hp_IdleMarkSweep_locs"
apply (clarsimp simp: obj_fields_marked_locs_def hp_IdleMarkSweep_locs_def mark_loop_locs_def)
apply (drule mp)
apply (auto intro: append_prefixD)
done
lemma obj_fields_marked_locs_subseteq_hs_in_sync_locs:
"obj_fields_marked_locs \<subseteq> hs_in_sync_locs"
by (auto simp: obj_fields_marked_locs_def hs_in_sync_locs_def hs_done_locs_def
dest: prefix_same_cases)
lemma handshake_obj_fields_markedD:
"\<lbrakk> atS gc obj_fields_marked_locs s; gc.handshake_invL s \<rbrakk> \<Longrightarrow> sys_ghost_handshake_phase s\<down> = hp_IdleMarkSweep \<and> All (ghost_handshake_in_sync (s\<down> sys))"
apply (simp add: gc.handshake_invL_def)
apply (elim conjE)
apply (drule mp, erule atS_mono[OF _ obj_fields_marked_locs_subseteq_hp_IdleMarkSweep_locs])
apply (drule mp, erule atS_mono[OF _ obj_fields_marked_locs_subseteq_hs_in_sync_locs])
apply clarsimp
done
lemma mark_loop_mo_mark_loop_field_done_subseteq_hp_IdleMarkSweep_locs:
"prefixed ''mark_loop_mo'' \<union> {''mark_loop_mark_field_done''} \<subseteq> hp_IdleMarkSweep_locs"
apply (clarsimp simp: hp_IdleMarkSweep_locs_def mark_loop_locs_def)
apply (drule mp)
apply (auto intro: append_prefixD)
done
lemma mark_loop_mo_mark_loop_field_done_subseteq_hs_in_sync_locs:
"prefixed ''mark_loop_mo'' \<union> {''mark_loop_mark_field_done''} \<subseteq> hs_in_sync_locs"
by (auto simp: hs_in_sync_locs_def hs_done_locs_def
dest: prefix_same_cases)
lemma mark_loop_mo_mark_loop_field_done_hp_phaseD:
"\<lbrakk> atS gc (prefixed ''mark_loop_mo'' \<union> {''mark_loop_mark_field_done''}) s; gc.handshake_invL s \<rbrakk> \<Longrightarrow> sys_ghost_handshake_phase s\<down> = hp_IdleMarkSweep \<and> All (ghost_handshake_in_sync (s\<down> sys))"
apply (simp add: gc.handshake_invL_def)
apply (elim conjE)
apply (drule mp, erule atS_mono[OF _ mark_loop_mo_mark_loop_field_done_subseteq_hp_IdleMarkSweep_locs])
apply (drule mp, erule atS_mono[OF _ mark_loop_mo_mark_loop_field_done_subseteq_hs_in_sync_locs])
apply clarsimp
done
(* an alternative is some kind of ghost_honorary_XXX for the GC while marking *)
lemma gc_marking_reaches_mw_Mutate:
assumes xys: "\<forall>y. (x reaches y) s \<longrightarrow> valid_ref y s"
assumes xy: "(x reaches y) (s(sys := s sys\<lparr>heap := (sys_heap s)(r := Option.map_option (\<lambda>obj. obj\<lparr>obj_fields := (obj_fields obj)(f := opt_r')\<rparr>) (sys_heap s r)),
mem_write_buffers := (mem_write_buffers (s sys))(p := ws)\<rparr>))"
assumes sb: "sys_mem_write_buffers (mutator m) s = mw_Mutate r f opt_r' # ws"
assumes vri: "valid_refs_inv s"
shows "valid_ref y s"
proof -
from xy xys
have "\<exists>z. z \<in> {x} \<union> mut_m.tso_write_refs m s \<and> (z reaches y) s \<and> valid_ref y s"
proof(induct rule: rtranclp.induct)
case (rtrancl_refl x) then show ?case by auto
next
case (rtrancl_into_rtrancl x y z) with sb vri show ?case
apply (clarsimp simp: points_to_mw_Mutate)
apply (elim disjE)
apply (force dest: rtranclp.intros(2))
apply (force dest: rtranclp.intros(2))
apply clarsimp
apply (elim disjE)
apply (rule exI[where x=z])
apply (clarsimp simp: mut_m.tso_write_refs_def)
apply (rule valid_refs_invD(3)[where m=m and x=z], auto simp: mut_m.tso_write_refs_def)[1]
apply (force dest: rtranclp.intros(2))
apply clarsimp
apply (elim disjE)
apply (rule exI[where x=z])
apply (clarsimp simp: mut_m.tso_write_refs_def)
apply (rule valid_refs_invD(3)[where m=m and x=z], auto simp: mut_m.tso_write_refs_def)[1]
apply (force dest: rtranclp.intros(2))
done
qed
then show ?thesis by blast
qed
lemma (in sys) gc_obj_fields_marked_invL[intro]:
notes gc.obj_fields_marked_invL_def[inv]
shows
"\<lbrace> gc.fM_fA_invL \<^bold>\<and> gc.handshake_invL \<^bold>\<and> gc.obj_fields_marked_invL
\<^bold>\<and> LSTP (fM_rel_inv \<^bold>\<and> handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> tso_writes_inv \<^bold>\<and> valid_refs_inv \<^bold>\<and> valid_W_inv) \<rbrace>
sys
\<lbrace> gc.obj_fields_marked_invL \<rbrace>"
apply vcg_nihe
apply (vcg_ni simp: p_not_sys fM_rel_inv_def)
apply (clarsimp simp: gc.obj_fields_marked_inv_def)
apply (frule (1) handshake_obj_fields_markedD)
apply (clarsimp simp: do_write_action_def
split: mem_write_action.splits)
(* mark *)
subgoal for s s' p ws x ref bool
apply (frule (1) valid_W_invD2)
apply (drule_tac x=x in spec)
apply clarsimp
apply (erule obj_at_field_on_heap_weakenE)
apply (clarsimp split: obj_at_splits)
done
(* ref *)
subgoal for s s' p ws x x23
apply (erule disjE; clarsimp)
apply (rename_tac m)
apply (drule_tac m=m in mut_m.handshake_phase_invD; clarsimp simp: hp_step_rel_def conj_disj_distribR[symmetric])
apply (drule_tac x=m in spec; clarsimp)
apply (drule_tac x=x in spec; clarsimp)
apply (auto split: option.splits)
done
(* fM *)
subgoal for s s' p ws x x4
apply (erule disjE)
apply (auto simp: fM_rel_def filter_empty_conv)[1]
apply clarsimp
done
subgoal for s s' p w ws
apply (clarsimp simp: gc.obj_fields_marked_inv_def)
apply (frule (1) mark_loop_mo_mark_loop_field_done_hp_phaseD)
apply (clarsimp simp: do_write_action_def
split: mem_write_action.splits)
(* mark *)
apply (frule (1) valid_W_invD2)
apply (erule obj_at_field_on_heap_weakenE)
apply (clarsimp split: obj_at_splits)
(* ref *)
apply (erule disjE, clarsimp+)[1]
apply (rename_tac option m)
apply (drule_tac m=m in mut_m.handshake_phase_invD, clarsimp simp: hp_step_rel_def conj_disj_distribR[symmetric])
apply (drule_tac x=m in spec, clarsimp)
apply (rule conjI)
apply (rule_tac x="gc_tmp_ref s\<down>" and m=m in valid_refs_invD(3), auto simp: mut_m.tso_write_refs_def)[1]
apply (clarsimp split: option.splits dest!: marked_insertionD)
(* fM *)
apply (erule disjE)
apply (auto simp: fM_rel_def filter_empty_conv)[1]
apply clarsimp
done
subgoal for s s' p w ws x y
apply (clarsimp simp: do_write_action_def
split: mem_write_action.splits)
(* mark *)
subgoal
apply (drule_tac x=x in spec)
apply (drule mp, erule predicate2D[OF rtranclp_mono[OF predicate2I], rotated])
apply clarsimp
apply assumption
done
(* ref *)
subgoal
apply (clarsimp simp: atS_un)
apply (erule disjE; clarsimp)
apply (erule gc_marking_reaches_mw_Mutate; blast)
done
done
(* mark loop mark field done *)
subgoal
apply (clarsimp simp: do_write_action_def
split: mem_write_action.splits)
(* mark *)
apply fast
(* fM *)
apply (clarsimp simp: gc.handshake_invL_def loc atS_simps)
apply (erule disjE)
apply (auto simp: fM_rel_def filter_empty_conv)[1]
apply clarsimp
done
done
lemma reachable_sweep_loop_free:
"mut_m.reachable m r (s(sys := s sys\<lparr>heap := (sys_heap s)(r' := None)\<rparr>))
\<Longrightarrow> mut_m.reachable m r s"
apply (clarsimp simp: mut_m.reachable_def)
apply (rename_tac x)
apply (rule_tac x=x in exI, simp)
apply (erule predicate2D[OF rtranclp_mono[OF predicate2I], rotated], clarsimp split: if_splits obj_at_splits)
done
lemma valid_refs_inv_sweep_loop_free[simp]:
assumes "valid_refs_inv s"
assumes ngr: "no_grey_refs s"
assumes rsi: "\<forall>m'. mut_m.reachable_snapshot_inv m' s"
assumes "white r' s"
shows "valid_refs_inv (s(sys := s sys\<lparr>heap := (sys_heap s)(r' := None)\<rparr>))"
using assms
apply (clarsimp simp: valid_refs_inv_def grey_reachable_def no_grey_refs_def dest!: reachable_sweep_loop_free)
apply (drule mut_m.reachable_blackD[OF ngr spec[OF rsi]])
apply (auto split: obj_at_splits)
done
lemma (in gc) valid_refs_inv[intro]:
"\<lbrace> fM_fA_invL \<^bold>\<and> handshake_invL \<^bold>\<and> gc_W_empty_invL \<^bold>\<and> gc_mark.mark_object_invL \<^bold>\<and> obj_fields_marked_invL \<^bold>\<and> phase_invL \<^bold>\<and> sweep_loop_invL
\<^bold>\<and> LSTP (handshake_phase_inv \<^bold>\<and> mutators_phase_inv \<^bold>\<and> sys_phase_inv \<^bold>\<and> valid_refs_inv \<^bold>\<and> valid_W_inv) \<rbrace>
gc
\<lbrace> LSTP valid_refs_inv \<rbrace>"
apply vcg_jackhammer
(* sweep loop free *)
apply (drule (1) handshake_in_syncD)
apply clarsimp
apply (auto simp: valid_refs_inv_def grey_reachable_def)
done
lemma (in sys) valid_refs_inv[intro]:
"\<lbrace> LSTP (valid_refs_inv \<^bold>\<and> tso_writes_inv) \<rbrace> sys \<lbrace> LSTP valid_refs_inv \<rbrace>"
apply vcg_jackhammer
apply (auto simp: do_write_action_def p_not_sys
split: mem_write_action.splits)
done
lemma (in mut_m) valid_refs_inv_discard_roots[simp]:
"\<lbrakk> valid_refs_inv s; roots' \<subseteq> mut_roots s \<rbrakk>
\<Longrightarrow> valid_refs_inv (s(mutator m := s (mutator m)\<lparr>roots := roots'\<rparr>))"
by (auto simp: valid_refs_inv_def mut_m.reachable_def split: if_splits)
lemma (in mut_m) valid_refs_inv_load[simp]:
"\<lbrakk> valid_refs_inv s; sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref r'; r \<in> mut_roots s \<rbrakk>
\<Longrightarrow> valid_refs_inv (s(mutator m := s (mutator m)\<lparr>roots := mut_roots s \<union> Option.set_option r'\<rparr>))"
by (auto simp: valid_refs_inv_def)
lemma (in mut_m) valid_refs_inv_alloc[simp]:
"\<lbrakk> valid_refs_inv s; sys_heap s r' = None \<rbrakk>
\<Longrightarrow> valid_refs_inv (s(mutator m := s (mutator m)\<lparr>roots := insert r' (mut_roots s)\<rparr>, sys := s sys\<lparr>heap := sys_heap s(r' \<mapsto> \<lparr>obj_mark = fl, obj_fields = Map.empty\<rparr>)\<rparr>))"
apply (clarsimp simp: valid_refs_inv_def)
apply (clarsimp simp: mut_m.reachable_def split: if_splits)
apply (auto elim: converse_rtranclpE split: obj_at_splits)
done
lemma (in mut_m) valid_refs_inv_store_ins[simp]:
"\<lbrakk> valid_refs_inv s; r \<in> mut_roots s; (\<exists>r'. opt_r' = Some r') \<longrightarrow> the opt_r' \<in> mut_roots s \<rbrakk>
\<Longrightarrow> valid_refs_inv (s(mutator m := s (mutator m)\<lparr> ghost_honorary_root := {} \<rparr>,
sys := s sys\<lparr> mem_write_buffers := (mem_write_buffers (s sys))(mutator m := sys_mem_write_buffers (mutator m) s @ [mw_Mutate r f opt_r']) \<rparr>))"
apply (subst valid_refs_inv_def)
apply (clarsimp simp: grey_reachable_def mut_m.reachable_def)
apply (rule conjI)
apply clarsimp
apply (rename_tac x xa)
apply (case_tac "xa = m")
apply clarsimp
apply fastforce
apply clarsimp
apply (fastforce dest: valid_refs_invD)
apply fastforce
done
lemma (in mut_m) valid_refs_inv_deref_del[simp]:
"\<lbrakk> valid_refs_inv s; sys_read (mutator m) (mr_Ref r f) (s sys) = mv_Ref opt_r'; r \<in> mut_roots s; mut_ghost_honorary_root s = {} \<rbrakk>
\<Longrightarrow> valid_refs_inv (s(mutator m := s (mutator m)\<lparr>ghost_honorary_root := Option.set_option opt_r', ref := opt_r'\<rparr>))"
by (clarsimp simp: valid_refs_inv_def)
lemma (in mut_m) valid_refs_inv[intro]:
"\<lbrace> mark_object_invL
\<^bold>\<and> mut_get_roots.mark_object_invL m
\<^bold>\<and> mut_store_del.mark_object_invL m
\<^bold>\<and> mut_store_ins.mark_object_invL m
\<^bold>\<and> LSTP valid_refs_inv \<rbrace>
mutator m
\<lbrace> LSTP valid_refs_inv \<rbrace>"
apply vcg_jackhammer
(* store ins mo co mark - FIXME some elim/dest rule really gets in the way here *)
subgoal
apply (subst valid_refs_inv_def)
apply clarsimp
apply (rule conjI)
apply clarsimp
apply (erule (1) valid_refs_invD)
apply (clarsimp simp: grey_reachable_def)
apply (erule disjE)
apply (erule (1) valid_refs_invD, simp)
apply clarsimp
apply (erule (1) valid_refs_invD, simp)
done
(* store del mo co mark *)
subgoal
apply (subst valid_refs_inv_def)
apply clarsimp
apply (rule conjI)
apply clarsimp
apply (erule (1) valid_refs_invD)
apply (clarsimp simp: grey_reachable_def)
apply (erule disjE)
apply (erule (1) valid_refs_invD, simp)
apply clarsimp
apply (erule (1) valid_refs_invD, simp)
done
(* get roots done *)
subgoal by (clarsimp simp: valid_refs_inv_def grey_reachable_def)
(* get roots loop mo co mark *)
subgoal by (auto simp: valid_refs_inv_def grey_reachable_def)
(* get work done *)
subgoal by (clarsimp simp: valid_refs_inv_def grey_reachable_def)
done
(*>*)
(*<*)
end
(*>*)
diff --git a/thys/Constructive_Cryptography/Examples/Secure_Channel/Message_Authentication_Code.thy b/thys/Constructive_Cryptography/Examples/Secure_Channel/Message_Authentication_Code.thy
--- a/thys/Constructive_Cryptography/Examples/Secure_Channel/Message_Authentication_Code.thy
+++ b/thys/Constructive_Cryptography/Examples/Secure_Channel/Message_Authentication_Code.thy
@@ -1,1269 +1,1271 @@
section \<open>Security of message authentication\<close>
theory Message_Authentication_Code imports
System_Construction
begin
definition rnd :: "security \<Rightarrow> bool list set" where
"rnd \<eta> \<equiv> nlists UNIV \<eta>"
definition mac :: "security \<Rightarrow> bool list \<Rightarrow> bool list \<Rightarrow> bool list spmf" where
"mac \<eta> r m \<equiv> return_spmf r"
definition vld :: "security \<Rightarrow> bool list set" where
"vld \<eta> \<equiv> nlists UNIV \<eta>"
fun valid_mac_query :: "security \<Rightarrow> (bool list \<times> bool list) insec_query \<Rightarrow> bool" where
"valid_mac_query \<eta> (ForwardOrEdit (Some (a, m))) \<longleftrightarrow> a \<in> vld \<eta> \<and> m \<in> vld \<eta>"
| "valid_mac_query \<eta> _ = True"
fun sim :: "('b list \<times> 'b list) option + unit \<Rightarrow> ('b list \<times> 'b list) insec_query
\<Rightarrow> (('b list \<times> 'b list) option \<times> (('b list \<times> 'b list) option + unit), auth_query , 'b list option) gpv" where
"sim (Inr ()) _ = Done (None, Inr())"
| "sim (Inl None) (Edit (a', m')) = do { _ \<leftarrow> Pause Drop Done; Done (None, Inr ())}"
| "sim (Inl (Some (a, m))) (Edit (a', m')) = (if a = a' \<and> m = m'
then do { _ \<leftarrow> Pause Forward Done; Done (None, Inl (Some (a, m)))}
else do { _ \<leftarrow> Pause Drop Done; Done (None, Inr ())})"
| "sim (Inl None) Forward = do {
Pause Forward Done;
Done (None, Inl None) }"
| "sim (Inl (Some _)) Forward = do {
Pause Forward Done;
Done (None, Inr ()) }"
| "sim (Inl None) Drop = do {
Pause Drop Done;
Done (None, Inl None) }"
| "sim (Inl (Some _)) Drop = do {
Pause Drop Done;
Done (None, Inr ()) }"
| "sim (Inl (Some (a, m))) Look = do {
lo \<leftarrow> Pause Look Done;
(case lo of
Some m \<Rightarrow> Done (Some (a, m), Inl (Some (a, m)))
| None \<Rightarrow> Done (None, Inl (Some (a, m))))}"
| "sim (Inl None) Look = do {
lo \<leftarrow> Pause Look Done;
(case lo of
Some m \<Rightarrow> do {
a \<leftarrow> lift_spmf (spmf_of_set (nlists UNIV (length m)));
Done (Some (a, m), Inl (Some (a, m)))}
| None \<Rightarrow> Done (None, Inl None))}"
context
fixes \<eta> :: "security"
begin
private definition rorc_channel_send :: "((bool \<times> unit) \<times> (bool list \<Rightarrow> bool list option) \<times> (bool list \<times> bool list) cstate, bool list, unit) oracle'" where
"rorc_channel_send s m \<equiv> (if fst (fst s)
then return_spmf ((), (True, ()), snd s)
else do {
(r, s) \<leftarrow> (rorc.rnd_oracle (rnd \<eta>))\<dagger> (snd s) m;
a \<leftarrow> mac \<eta> r m;
(_, s) \<leftarrow> \<dagger>channel.send_oracle s (a, m);
return_spmf ((), (True, ()), s)
})"
private definition rorc_channel_recv :: "((bool \<times> unit) \<times> (bool list \<Rightarrow> bool list option) \<times> (bool list \<times> bool list) cstate, unit, bool list option) oracle'" where
"rorc_channel_recv s q \<equiv> do {
(m, s) \<leftarrow> \<dagger>\<dagger>channel.recv_oracle s ();
(case m of
None \<Rightarrow> return_spmf (None, s)
| Some (a, m) \<Rightarrow> do {
(r, s) \<leftarrow> \<dagger>(rorc.rnd_oracle (rnd \<eta>))\<dagger> s m;
a' \<leftarrow> mac \<eta> r m;
return_spmf (if a' = a then Some m else None, s)})
}"
private definition rorc_channel_recv_f :: "((bool list \<Rightarrow> bool list option) \<times> (bool list \<times> bool list) cstate, unit, bool list option) oracle'" where
"rorc_channel_recv_f s q \<equiv> do {
(am, (as, ams)) \<leftarrow> \<dagger>channel.recv_oracle s ();
(case am of
None \<Rightarrow> return_spmf (None, (as, ams))
| Some (a, m) \<Rightarrow> (case as m of
None \<Rightarrow> do {
a'' :: bool list \<leftarrow> spmf_of_set (nlists UNIV \<eta> - {a});
a' \<leftarrow> spmf_of_set (nlists UNIV \<eta>);
(if a' = a
then return_spmf (None, as(m := Some a''), ams)
else return_spmf (None, as(m := Some a'), ams)) }
| Some a' \<Rightarrow> return_spmf (if a' = a then Some m else None, as, ams)))}"
private fun lazy_channel_send :: "(bool list cstate \<times> (bool list \<times> bool list) option \<times> (bool list \<Rightarrow> bool list option), bool list, unit) oracle'" where
"lazy_channel_send (Void, es) m = return_spmf ((), (Store m, es))"
| "lazy_channel_send s m = return_spmf ((), s)"
private fun lazy_channel_recv :: "(bool list cstate \<times> (bool list \<times> bool list) option \<times> (bool list \<Rightarrow> bool list option), unit, bool list option) oracle'" where
"lazy_channel_recv (Collect m, None, as) () = return_spmf (Some m, (Fail, None, as))"
| "lazy_channel_recv (ms, Some (a', m'), as) () = (case as m' of
None \<Rightarrow> do {
a \<leftarrow> spmf_of_set (rnd \<eta>);
return_spmf (if a = a' then Some m' else None, cstate.Fail, None, as (m' := Some a))}
| Some a \<Rightarrow> return_spmf (if a = a' then Some m' else None, Fail, None, as))"
| "lazy_channel_recv s () = return_spmf (None, s)"
private fun lazy_channel_insec :: "(bool list cstate \<times> (bool list \<times> bool list) option \<times> (bool list \<Rightarrow> bool list option),
(bool list \<times> bool list) insec_query, (bool list \<times> bool list) option) oracle'" where
"lazy_channel_insec (Void, _, as) (Edit (a', m')) = return_spmf (None, (Collect m', Some (a', m'), as))"
| "lazy_channel_insec (Store m, _, as) (Edit (a', m')) = return_spmf (None, (Collect m', Some (a', m'), as))"
| "lazy_channel_insec (Store m, es) Forward = return_spmf (None, (Collect m, es))"
| "lazy_channel_insec (Store m, es) Drop = return_spmf (None, (Fail, es))"
| "lazy_channel_insec (Store m, None, as) Look = (case as m of
None \<Rightarrow> do {
a \<leftarrow> spmf_of_set (rnd \<eta>);
return_spmf (Some (a, m), Store m, None, as (m := Some a))}
| Some a \<Rightarrow> return_spmf (Some (a, m), Store m, None, as))"
| "lazy_channel_insec s _ = return_spmf (None, s)"
private fun lazy_channel_recv_f :: "(bool list cstate \<times> (bool list \<times> bool list) option \<times> (bool list \<Rightarrow> bool list option), unit, bool list option) oracle'" where
"lazy_channel_recv_f (Collect m, None, as) () = return_spmf (Some m, (Fail, None, as))"
| "lazy_channel_recv_f (ms, Some (a', m'), as) () = (case as m' of
None \<Rightarrow> do {
a \<leftarrow> spmf_of_set (rnd \<eta>);
return_spmf (None, Fail, None, as (m' := Some a))}
| Some a \<Rightarrow> return_spmf (if a = a' then Some m' else None, Fail, None, as))"
| "lazy_channel_recv_f s () = return_spmf (None, s)"
private abbreviation callee_auth_channel where
"callee_auth_channel callee \<equiv> lift_state_oracle extend_state_oracle (attach_callee callee auth_channel.auth_oracle)"
private abbreviation
"valid_insecQ \<equiv> {(x :: (bool list \<times> bool list) insec_query). case x of
ForwardOrEdit (Some (a, m)) \<Rightarrow> length a = id' \<eta> \<and> length m = id' \<eta>
| _ \<Rightarrow> True}"
private inductive S :: "(bool list cstate \<times> (bool list \<times> bool list) option \<times> (bool list \<Rightarrow> bool list option)) spmf
\<Rightarrow> ((bool \<times> unit) \<times> (bool list \<Rightarrow> bool list option) \<times> (bool list \<times> bool list) cstate) spmf \<Rightarrow> bool" where
"S (return_spmf (Void, None, Map.empty))
(return_spmf ((False, ()), Map.empty, Void))"
| "S (return_spmf (Store m, None, Map.empty))
(map_spmf (\<lambda>a. ((True, ()), [m \<mapsto> a], Store (a, m))) (spmf_of_set (nlists UNIV \<eta>)))"
if "length m = id' \<eta>"
| "S (return_spmf (Collect m, None, Map.empty))
(map_spmf (\<lambda>a. ((True, ()), [m \<mapsto> a], Collect (a, m))) (spmf_of_set (nlists UNIV \<eta>)))"
if "length m = id' \<eta>"
| "S (return_spmf (Store m, None, [m \<mapsto> a]))
(return_spmf ((True, ()), [m \<mapsto> a], Store (a, m)))"
if "length m = id' \<eta>" and "length a = id' \<eta>"
| "S (return_spmf (Collect m, None, [m \<mapsto> a]))
(return_spmf ((True, ()), [m \<mapsto> a], Collect (a, m)))"
if "length m = id' \<eta>" and "length a = id' \<eta>"
| "S (return_spmf (Fail, None, Map.empty))
(map_spmf (\<lambda>a. ((True, ()), [m \<mapsto> a], Fail)) (spmf_of_set (nlists UNIV \<eta>)))"
if "length m = id' \<eta>"
| "S (return_spmf (Fail, None, [m \<mapsto> a]))
(return_spmf ((True, ()), [m \<mapsto> a], Fail))"
if "length m = id' \<eta>" and "length a = id' \<eta>"
| "S (return_spmf (Collect m', Some (a', m'), Map.empty))
(return_spmf ((False, ()), Map.empty, Collect (a', m')))"
if "length m' = id' \<eta>" and "length a' = id' \<eta>"
| "S (return_spmf (Collect m', Some (a', m'), [m \<mapsto> a]))
(return_spmf ((True, ()), [m \<mapsto> a], Collect (a', m')))"
if "length m = id' \<eta>" and "length a = id' \<eta>" and "length m' = id' \<eta>" and "length a' = id' \<eta>"
| "S (return_spmf (Collect m', Some (a', m'), Map.empty))
(map_spmf (\<lambda>x. ((True, ()), [m \<mapsto> x], Collect (a', m'))) (spmf_of_set (nlists UNIV \<eta>)))"
if "length m = id' \<eta>" and "length m' = id' \<eta>" and "length a' = id' \<eta>"
| "S (map_spmf (\<lambda>x. (Fail, None, as(m' \<mapsto> x))) spmf_s)
(map_spmf (\<lambda>x. ((False, ()), as(m' \<mapsto> x), Fail)) spmf_s)"
if "length m' = id' \<eta>" and "lossless_spmf spmf_s"
| "S (map_spmf (\<lambda>x. (Fail, None, as(m' \<mapsto> x))) spmf_s)
(map_spmf (\<lambda>x. ((True, ()), as(m' \<mapsto> x), Fail)) spmf_s)"
if "length m' = id' \<eta>" and "lossless_spmf spmf_s"
| "S (return_spmf (Fail, None, [m' \<mapsto> a']))
(map_spmf (\<lambda>x. ((True, ()), [m \<mapsto> x, m' \<mapsto> a'], Fail)) (spmf_of_set (nlists UNIV \<eta>)))"
if "length m = id' \<eta>" and "length m'= id' \<eta>" and "length a' = id' \<eta>"
| "S (map_spmf (\<lambda>x. (Fail, None, [m' \<mapsto> x])) (spmf_of_set (nlists UNIV \<eta> \<inter> {x. x \<noteq> a'})))
(map_spmf (\<lambda>x. ((True, ()), [m \<mapsto> fst x, m' \<mapsto> snd x], Fail)) (spmf_of_set (nlists UNIV \<eta> \<times> nlists UNIV \<eta> \<inter> {x. snd x \<noteq> a'})))"
if "length m = id' \<eta>" and "length m'= id' \<eta>"
| "S (map_spmf (\<lambda>x. (Fail, None, as(m' \<mapsto> x))) spmf_s)
(map_spmf (\<lambda>p. ((True, ()), as(m' \<mapsto> fst p, m \<mapsto> snd p), Fail)) (mk_lossless (pair_spmf spmf_s (spmf_of_set (nlists UNIV \<eta>)))))"
if "length m = id' \<eta>" and "length m'= id' \<eta>" and "lossless_spmf spmf_s"
private lemma trace_eq_lazy:
assumes "\<eta> > 0"
shows "(valid_insecQ <+> nlists UNIV (id' \<eta>) <+> UNIV) \<turnstile>\<^sub>R
RES (lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv) (Void, None, Map.empty)
\<approx>
RES (\<dagger>\<dagger>insec_channel.insec_oracle \<oplus>\<^sub>O rorc_channel_send \<oplus>\<^sub>O rorc_channel_recv) ((False, ()), Map.empty, Void)"
(is "?A \<turnstile>\<^sub>R RES (?L1 \<oplus>\<^sub>O ?L2 \<oplus>\<^sub>O ?L3) ?SL \<approx> RES (?R1 \<oplus>\<^sub>O ?R2 \<oplus>\<^sub>O ?R3) ?SR")
proof -
note [simp] =
spmf.map_comp o_def map_bind_spmf bind_map_spmf bind_spmf_const exec_gpv_bind
insec_channel.insec_oracle.simps channel.send_oracle.simps channel.recv_oracle.simps
rorc_channel_send_def rorc_channel_recv_def rorc.rnd_oracle.simps mac_def rnd_def
have aux1: "nlists (UNIV::bool set) \<eta> \<times> nlists (UNIV::bool set) \<eta> \<inter> {x. snd x \<noteq> a'} \<noteq> {}" (is ?aux1)
and aux2: "nlists (UNIV::bool set) \<eta> \<inter> {x. x \<noteq> a'} \<noteq> {}" (is ?aux2) for a'
proof -
have "\<exists>a. a \<in> nlists (UNIV::bool set) \<eta> \<and> a \<noteq> a'" for a'
proof (cases "a' \<in> nlists (UNIV::bool set) \<eta>")
case True
show ?thesis
proof (rule ccontr)
have "2 ^ \<eta> = card (nlists (UNIV :: bool set) \<eta>)" by (simp add: card_nlists)
also assume "\<nexists>a. a \<in> nlists UNIV \<eta> \<and> a \<noteq> a'"
then have "nlists UNIV \<eta> = {a'}" using True by blast
then have fct:"card (nlists (UNIV :: bool set) \<eta>) = card {a'}" by simp
also have " card {a'} = 1" by simp
finally have "\<eta> = 0" by simp
then show "False" using assms by blast
qed
next
case False
obtain a where obt:"a \<in> nlists (UNIV::bool set) \<eta>" using assms by auto
then have "a \<noteq> a'" using False by blast
then show ?thesis using obt by auto
qed
then obtain a where o1: "a \<in> nlists (UNIV::bool set) \<eta>" and o2: "a \<noteq> a'" by blast
obtain m where "m \<in> nlists (UNIV::bool set) \<eta>" by blast
with o1 o2 have "(m, a) \<in> {x. snd x \<noteq> a'}" and "(m, a) \<in> nlists UNIV \<eta> \<times> nlists UNIV \<eta>" by auto
then show ?aux1 by blast
from o1 o2 have "a \<in> {x. x \<noteq> a'}" and "a \<in> nlists UNIV \<eta>" by auto
then show ?aux2 by blast
qed
have *: "?A \<turnstile>\<^sub>C ?L1 \<oplus>\<^sub>O ?L2 \<oplus>\<^sub>O ?L3(?SL) \<approx> ?R1 \<oplus>\<^sub>O ?R2 \<oplus>\<^sub>O ?R3(?SR)"
proof(rule trace'_eqI_sim[where S=S], goal_cases Init_OK Output_OK State_OK)
case Init_OK
then show ?case by (simp add: S.simps)
next
case (Output_OK p q query)
show ?case
proof (cases query)
case (Inl adv_query)
with Output_OK show ?thesis
proof cases
case (14 m m' a')
then show ?thesis using Output_OK(2) Inl
by(cases adv_query)(simp; subst (1 2) weight_spmf_of_set_finite; auto simp add: assms aux1 aux2)+
qed (auto simp add: weight_mk_lossless lossless_spmf_def split: aquery.splits option.splits)
next
case Snd_Rcv: (Inr query')
show ?thesis
proof (cases query')
case (Inl snd_query)
with Output_OK Snd_Rcv show ?thesis
proof cases
case (11 m' _ as)
with Output_OK(2) Snd_Rcv Inl show ?thesis
by(cases "snd_query = m'"; cases "as snd_query")(simp_all)
next
case (14 m m' a')
with Output_OK(2) Snd_Rcv Inl show ?thesis
by(simp; subst (1 2) weight_spmf_of_set_finite; auto simp add: assms aux1 aux2)
qed (auto simp add: weight_mk_lossless lossless_spmf_def)
next
case (Inr rcv_query)
with Output_OK Snd_Rcv show ?thesis
proof cases
case (10 m m' a')
with Output_OK(2) Snd_Rcv Inr show ?thesis by(cases "m = m'")(simp_all)
next
case (14 m m' a')
with Output_OK(2) Snd_Rcv Inr show ?thesis
by(simp; subst (1 2) weight_spmf_of_set_finite; auto simp add: assms aux1 aux2)
qed (auto simp add: weight_mk_lossless lossless_spmf_def split:option.splits)
qed
qed
next
case (State_OK p q query state answer state')
show ?case
proof (cases query)
case (Inl adv_query)
show ?thesis
proof (cases adv_query)
case Look
with State_OK Inl show ?thesis
proof cases
case Store_State_Channel: (2 m)
have[simp]: " a \<in> nlists UNIV \<eta> \<Longrightarrow>
S (cond_spmf_fst (map_spmf (\<lambda>x. (Inl (Some (x, m)), Store m, None, [m \<mapsto> x]))
(spmf_of_set (nlists UNIV \<eta>))) (Inl (Some (a, m))))
(cond_spmf_fst (map_spmf (\<lambda>x. (Inl (Some (x, m)), (True, ()), [m \<mapsto> x], Store (x, m)))
(spmf_of_set (nlists UNIV \<eta>))) (Inl (Some (a, m))))" for a
proof(subst (1 2) cond_spmf_fst_map_Pair1, goal_cases)
case 4
then show ?case
by(subst (1 2 3) inv_into_f_f, simp_all add: inj_on_def)
(rule S.intros, simp_all add: Store_State_Channel in_nlists_UNIV id'_def)
qed (simp_all add: id'_def inj_on_def)
from Store_State_Channel show ?thesis using State_OK Inl Look
by(clarsimp simp add: map_spmf_conv_bind_spmf[symmetric])
qed (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S.intros)
next
case (ForwardOrEdit foe)
with State_OK Inl show ?thesis
proof (cases foe)
case (Some am')
obtain a' m' where "am' = (a', m')" by (cases am') simp
with State_OK Inl ForwardOrEdit Some show ?thesis
by cases (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S.intros elim:S.cases)
qed (erule S.cases, auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S.intros)
next
case Drop
with State_OK Inl show ?thesis
by cases (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S.intros)
qed
next
case Snd_Rcv: (Inr query')
show ?thesis
proof (cases query')
case (Inl snd_query)
with State_OK Snd_Rcv show ?thesis
proof cases
case 1
with State_OK Snd_Rcv Inl show ?thesis
by(clarsimp simp add: map_spmf_conv_bind_spmf[symmetric])
(rule S.intros, simp add: in_nlists_UNIV)
next
case (8 m' a')
with State_OK Snd_Rcv Inl show ?thesis
by(clarsimp simp add: map_spmf_conv_bind_spmf[symmetric])
(rule S.intros, simp add: in_nlists_UNIV)
next
case (11 m' spmf_s as)
with State_OK Snd_Rcv Inl show ?thesis
by(auto simp add: bind_bind_conv_pair_spmf split_def split: if_splits option.splits intro!: S.intros)
((auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: S.intros), simp only: id'_def in_nlists_UNIV)
qed (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S.intros)
next
case (Inr rcv_query)
with State_OK Snd_Rcv show ?thesis
proof(cases)
case (8 m' a')
then show ?thesis using State_OK(2-) Snd_Rcv Inr
by (auto simp add: map_spmf_conv_bind_spmf[symmetric] image_def
cond_spmf_fst_def vimage_def aux1 aux2 assms intro:S.intros)
next
case (9 m a m' a')
then show ?thesis using State_OK(2-) Snd_Rcv Inr
by (cases "m = m'") (auto simp add: map_spmf_conv_bind_spmf[symmetric] cond_spmf_fst_def
vimage_def aux1 aux2 assms intro:S.intros split!: if_splits)
next
case (10 m m' a')
show ?thesis
proof (cases "m = m'")
case True
with 10 show ?thesis using State_OK(2-) Snd_Rcv Inr
by (auto simp add: map_spmf_conv_bind_spmf[symmetric] cond_spmf_fst_def
vimage_def aux1 aux2 assms split!: if_split intro:S.intros)
next
case False
have[simp]: "a' \<in> nlists UNIV \<eta> \<Longrightarrow> nlists (UNIV :: bool set) \<eta> \<times> nlists UNIV \<eta> \<inter> {x. snd x = a'} = nlists UNIV \<eta> \<times> {a'}"
by auto
from 10 False show ?thesis using State_OK(2-) Snd_Rcv Inr
by(simp add: bind_bind_conv_pair_spmf)
((auto simp add: bind_bind_conv_pair_spmf split_def image_def map_spmf_conv_bind_spmf[symmetric]
cond_spmf_fst_def vimage_def cond_spmf_spmf_of_set pair_spmf_of_set )
, (auto simp add: pair_spmf_of_set[symmetric] spmf_of_set_singleton pair_spmf_return_spmf2
map_spmf_of_set_inj_on[symmetric] simp del: map_spmf_of_set_inj_on intro: S.intros))
qed
qed (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S.intros)
qed
qed
qed
from * show ?thesis by simp
qed
private lemma game_difference:
defines "\<I> \<equiv> \<I>_uniform (Set.Collect (valid_mac_query \<eta>)) (insert None (Some ` (nlists UNIV \<eta> \<times> nlists UNIV \<eta>))) \<oplus>\<^sub>\<I>
(\<I>_uniform (vld \<eta>) UNIV \<oplus>\<^sub>\<I> \<I>_uniform UNIV (insert None (Some ` vld \<eta>)))"
assumes bound: "interaction_bounded_by' (\<lambda>_. True) \<A> q"
and lossless: "plossless_gpv \<I> \<A>"
and WT: "\<I> \<turnstile>g \<A> \<surd>"
shows
"\<bar>spmf (connect \<A> (RES (lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv_f) (Void, None, Map.empty))) True -
spmf (connect \<A> (RES (lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv) (Void, None, Map.empty))) True\<bar>
\<le> q / real (2 ^ \<eta>)" (is "?LHS \<le> _")
proof -
define lazy_channel_insec' where
"lazy_channel_insec' = (\<dagger>lazy_channel_insec :: (bool \<times> bool list cstate \<times> (bool list \<times> bool list) option \<times> (bool list \<Rightarrow> bool list option),
(bool list \<times> bool list) insec_query, (bool list \<times> bool list) option) oracle')"
define lazy_channel_send' where
"lazy_channel_send' = (\<dagger>lazy_channel_send :: (bool \<times> bool list cstate \<times> (bool list \<times> bool list) option \<times> (bool list \<Rightarrow> bool list option),
bool list, unit) oracle')"
define lazy_channel_recv' where
"lazy_channel_recv' = (\<lambda> (s :: bool \<times> bool list cstate \<times> (bool list \<times> bool list) option \<times> (bool list \<Rightarrow> bool list option)) (q :: unit).
(case s of
(flg, Collect m, None, as) \<Rightarrow> return_spmf (Some m, (flg, Fail, None, as))
| (flg, ms, Some (a', m'), as) \<Rightarrow> (case as m' of
None \<Rightarrow> do {
a \<leftarrow> spmf_of_set (rnd \<eta>);
return_spmf (if a = a' then Some m' else None, flg \<or> a = a', Fail, None, as (m' := Some a))}
| Some a \<Rightarrow> return_spmf (if a = a' then Some m' else None, flg, Fail, None, as))
| _ \<Rightarrow> return_spmf (None, s)))"
define lazy_channel_recv_f' where
"lazy_channel_recv_f' = (\<lambda> (s :: bool \<times> bool list cstate \<times> (bool list \<times> bool list) option \<times> (bool list \<Rightarrow> bool list option)) (q :: unit).
(case s of
(flg, Collect m, None, as) \<Rightarrow> return_spmf (Some m, (flg, Fail, None, as))
| (flg, ms, Some (a', m'), as) \<Rightarrow> (case as m' of
None \<Rightarrow> do {
a \<leftarrow> spmf_of_set (rnd \<eta>);
return_spmf (None, flg \<or> a = a', Fail, None, as (m' := Some a))}
| Some a \<Rightarrow> return_spmf (if a = a' then Some m' else None, flg, Fail, None, as))
| _ \<Rightarrow> return_spmf (None, s)))"
define game where
"game = (\<lambda>(\<A> :: ((bool list \<times> bool list) insec_query + bool list + unit, (bool list \<times> bool list) option + unit + bool list option) distinguisher) recv_oracle. do {
(b :: bool, (flg, ams, es, as))\<leftarrow> (exec_gpv (lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O recv_oracle) \<A> (False, Void, None, Map.empty));
return_spmf (b, flg) })"
have fact1: "spmf (connect \<A> (RES (lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv_f) (Void, None, Map.empty))) True
= spmf (connect \<A> (RES (lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv_f') (False, Void, None, Map.empty))) True"
proof -
let ?orc_lft = "lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv_f"
let ?orc_rgt = "lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv_f'"
have[simp]: "rel_spmf (rel_prod (=) (\<lambda>x y. x = snd y))
(lazy_channel_insec s q) (lazy_channel_insec' (flg, s) q) " for s flg q
by (cases s) (simp add: lazy_channel_insec'_def spmf_rel_map rel_prod_sel split_def option.rel_refl pmf.rel_refl split:option.split)
have[simp]: "rel_spmf (rel_prod (=) (\<lambda>x y. x = snd y))
(lazy_channel_send s q) (lazy_channel_send' (flg, s) q)" for s flg q
proof -
obtain ams es as where "s = (ams, es, as)" by (cases s)
then show ?thesis by (cases ams) (auto simp add: spmf_rel_map map_spmf_conv_bind_spmf split_def lazy_channel_send'_def)
qed
have[simp]: "rel_spmf (rel_prod (=) (\<lambda>x y. x = snd y))
(lazy_channel_recv_f s q) (lazy_channel_recv_f' (flg, s) q)" for s flg q
proof -
obtain ams es as where *: "s = (ams, es, as)" by (cases s)
show ?thesis
proof (cases es)
case None
with * show ?thesis by (cases ams) (simp_all add:lazy_channel_recv_f'_def)
next
case (Some am)
obtain a m where "am = (a, m)" by (cases am)
with * show ?thesis by (cases ams) (simp_all add: lazy_channel_recv_f'_def rel_spmf_bind_reflI split:option.split)
qed
qed
have[simp]: "rel_spmf (rel_prod (=) (\<lambda>x y. x = snd y))
((lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv_f) (ams, es, as) query)
((lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv_f') (flg, ams, es, as) query)" for flg ams es as query
proof (cases query)
case (Inl adv_query)
then show ?thesis
by(clarsimp simp add: spmf_rel_map map_spmf_conv_bind_spmf[symmetric] apfst_def map_prod_def split_beta)
((rule rel_spmf_mono[where A="rel_prod (=) (\<lambda>x y. x = snd y)"]), auto)
next
case (Inr query')
note Snd_Rcv = this
then show ?thesis
by (cases query', auto simp add: spmf_rel_map map_spmf_conv_bind_spmf[symmetric] split_beta)
((rule rel_spmf_mono[where A="rel_prod (=) (\<lambda>x y. x = snd y)"]); auto)+
qed
have[simp]: "rel_spmf (\<lambda>x y. fst x = fst y)
(exec_gpv ?orc_lft \<A> (Void, None, Map.empty)) (exec_gpv ?orc_rgt \<A> (False, Void, None, Map.empty))"
by(rule rel_spmf_mono[where A="rel_prod (=) (\<lambda>x y. x = snd y)"])
(auto simp add: gpv.rel_eq split_def intro!: rel_funI
exec_gpv_parametric[where CALL="(=)", THEN rel_funD, THEN rel_funD, THEN rel_funD])
show ?thesis
unfolding map_spmf_conv_bind_spmf exec_gpv_resource_of_oracle connect_def spmf_conv_measure_spmf
by(rule measure_spmf_parametric[where A="(=)", THEN rel_funD, THEN rel_funD])
(auto simp add: rel_pred_def spmf_rel_map map_spmf_conv_bind_spmf[symmetric] gpv.rel_eq split_def intro!: rel_funI)
qed
have fact2: "\<bar>spmf (connect \<A> (RES (lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv_f') (False, Void, None, Map.empty))) True -
spmf (connect \<A> (RES (lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv') (False, Void, None, Map.empty))) True\<bar>
\<le> Sigma_Algebra.measure (measure_spmf (game \<A> lazy_channel_recv_f')) {x. snd x}" (is "\<bar>spmf ?L _ - spmf ?R _ \<bar> \<le> _ ")
proof -
let ?orc_lft = "lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv_f'"
let ?orc_rgt = "lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv'"
have[simp]: "callee_invariant_on lazy_channel_insec' fst (\<I>_uniform (Set.Collect (valid_mac_query \<eta>)) UNIV)"
proof (unfold_locales, goal_cases)
case (1 state query answer state')
then show ?case
by(cases state; cases "fst (snd state)")(auto simp add: spmf_rel_map map_spmf_conv_bind_spmf split_def lazy_channel_insec'_def)
qed (auto intro: WT_calleeI)
have[simp]: "callee_invariant_on lazy_channel_send' fst (\<I>_uniform (vld \<eta>) UNIV)"
proof (unfold_locales, goal_cases)
case (1 state query answer state')
then show ?case
by(cases state; cases "fst (snd state)")(auto simp add: spmf_rel_map map_spmf_conv_bind_spmf split_def lazy_channel_send'_def)
qed (auto intro: WT_calleeI)
have[simp]: "callee_invariant lazy_channel_recv' fst"
proof (unfold_locales, goal_cases)
case (1 state query answer state')
then show ?case
by(cases state; cases "fst (snd state)")(auto simp add: lazy_channel_recv'_def split:option.splits)
qed (auto split:option.splits)
have[simp]: "callee_invariant lazy_channel_recv_f' fst"
proof (unfold_locales, goal_cases)
case (1 state query answer state')
then show ?case
by(cases state; cases "fst (snd state)")(auto simp add: lazy_channel_recv_f'_def split:option.splits)
qed (auto split:option.splits)
have[simp]: "lossless_spmf (lazy_channel_insec s q)" for s q
by(cases "(s, q)" rule: lazy_channel_insec.cases)(auto simp add: rnd_def split!: option.split)
have[simp]: "lossless_spmf (lazy_channel_send' s q)" for s q
by(cases s; cases "fst (snd s)")(simp_all add: lazy_channel_send'_def)
have[simp]: "lossless_spmf (lazy_channel_recv' s ())" for s
by(auto simp add: lazy_channel_recv'_def rnd_def split: prod.split cstate.split option.split)
have[simp]: "lossless_spmf (lazy_channel_recv_f' s ())" for s
by(auto simp add: lazy_channel_recv_f'_def rnd_def split: prod.split cstate.split option.split)
have[simp]: "rel_spmf (\<lambda>(a, s1') (b, s2'). (fst s2' \<longrightarrow> fst s1') \<and> (\<not> fst s2' \<longrightarrow> \<not> fst s1' \<and> a = b \<and> s1' = s2'))
(?orc_lft (flg, ams, es, as) query) (?orc_rgt (flg, ams, es, as) query)" for flg ams es as query
proof (cases query)
case (Inl adv_query)
then show ?thesis by (auto simp add: spmf_rel_map map_spmf_conv_bind_spmf intro: rel_spmf_bind_reflI)
next
case (Inr query')
note Snd_Rcv = this
show ?thesis
proof (cases query')
case (Inl snd_query)
with Snd_Rcv show ?thesis
by (auto simp add: spmf_rel_map map_spmf_conv_bind_spmf intro: rel_spmf_bind_reflI)
next
case (Inr rcv_query)
with Snd_Rcv show ?thesis
proof (cases es)
case (Some am')
obtain a' m' where "am' = (a', m')" by (cases am')
with Snd_Rcv Some Inr show ?thesis
by (cases ams) (auto simp add: spmf_rel_map map_spmf_conv_bind_spmf
lazy_channel_recv'_def lazy_channel_recv_f'_def rel_spmf_bind_reflI split:option.splits)
qed (cases ams; auto simp add: lazy_channel_recv'_def lazy_channel_recv_f'_def split:option.splits)
qed
qed
let ?\<I> = "\<I>_uniform (Set.Collect (valid_mac_query \<eta>)) UNIV \<oplus>\<^sub>\<I> (\<I>_uniform (vld \<eta>) UNIV \<oplus>\<^sub>\<I> \<I>_full)"
let ?\<A> = "mk_lossless_gpv (responses_\<I> \<I>) True \<A>"
have "plossless_gpv ?\<I> ?\<A>" using lossless WT
by(rule plossless_gpv_mk_lossless_gpv)(simp add: \<I>_def)
moreover have "?\<I> \<turnstile>g ?\<A> \<surd>" using WT by(rule WT_gpv_mk_lossless_gpv)(simp add: \<I>_def)
ultimately have "rel_spmf (\<lambda>x y. fst (snd x) = fst (snd y) \<and> (\<not> fst (snd y) \<longrightarrow> (fst x \<longleftrightarrow> fst y)))
(exec_gpv ?orc_lft ?\<A> (False, Void, None, Map.empty)) (exec_gpv ?orc_rgt ?\<A> (False, Void, None, Map.empty))"
by(auto simp add: \<I>_def intro!: exec_gpv_oracle_bisim_bad_plossless[where X="(=)" and
X_bad="\<lambda> _ _. True" and ?bad1.0="fst" and ?bad2.0="fst" and \<I> = "?\<I>", THEN rel_spmf_mono])
(auto simp add: lazy_channel_insec'_def intro: WT_calleeI)
also let ?I = "(\<lambda>(_, s1, s2, s3). pred_cstate (\<lambda>x. length x = \<eta>) s1 \<and> pred_option (\<lambda>(x, y). length x = \<eta> \<and> length y = \<eta>) s2 \<and> ran s3 \<subseteq> nlists UNIV \<eta>)"
have "callee_invariant_on (lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv_f') ?I \<I>"
apply(unfold_locales)
subgoal for s x y s'
+ supply [[simproc del: defined_all]]
apply(clarsimp simp add: \<I>_def; elim PlusE; clarsimp simp add: lazy_channel_insec'_def lazy_channel_send'_def lazy_channel_recv_f'_def)
subgoal for _ _ _ _ x'
by(cases "(snd s, x')" rule: lazy_channel_insec.cases)
(auto simp add: vld_def in_nlists_UNIV rnd_def split: option.split_asm)
subgoal by(cases "(snd s, projl (projr x))" rule: lazy_channel_send.cases)(auto simp add: vld_def in_nlists_UNIV)
subgoal by(cases "(snd s, ())" rule: lazy_channel_recv_f.cases)(auto 4 3 split: option.split_asm if_split_asm cstate.split_asm simp add: in_nlists_UNIV vld_def ran_def rnd_def option.pred_set )
done
subgoal for s
apply(clarsimp simp add: \<I>_def; intro conjI WT_calleeI; clarsimp simp add: lazy_channel_insec'_def lazy_channel_send'_def lazy_channel_recv_f'_def)
subgoal for _ _ _ _ x
by(cases "(snd s, x)" rule: lazy_channel_insec.cases)
(auto simp add: vld_def in_nlists_UNIV rnd_def ran_def split: option.split_asm)
subgoal by(cases "(snd s, ())" rule: lazy_channel_recv_f.cases)(auto 4 3 split: option.split_asm if_split_asm cstate.split_asm simp add: in_nlists_UNIV vld_def ran_def rnd_def)
done
done
then have "exec_gpv ?orc_lft ?\<A> (False, Void, None, Map.empty) = exec_gpv ?orc_lft \<A> (False, Void, None, Map.empty)"
using WT by(rule callee_invariant_on.exec_gpv_mk_lossless_gpv) simp
also have "callee_invariant_on (lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv') ?I \<I>"
apply(unfold_locales)
subgoal for s x y s'
+ supply [[simproc del: defined_all]]
apply(clarsimp simp add: \<I>_def; elim PlusE; clarsimp simp add: lazy_channel_insec'_def lazy_channel_send'_def lazy_channel_recv'_def)
subgoal for _ _ _ _ x'
by(cases "(snd s, x')" rule: lazy_channel_insec.cases)
(auto simp add: vld_def in_nlists_UNIV rnd_def split: option.split_asm)
subgoal by(cases "(snd s, projl (projr x))" rule: lazy_channel_send.cases)(auto simp add: vld_def in_nlists_UNIV)
subgoal by(cases "(snd s, ())" rule: lazy_channel_recv.cases)(auto 4 3 split: option.split_asm if_split_asm cstate.split_asm simp add: in_nlists_UNIV vld_def ran_def rnd_def option.pred_set )
done
subgoal for s
apply(clarsimp simp add: \<I>_def; intro conjI WT_calleeI; clarsimp simp add: lazy_channel_insec'_def lazy_channel_send'_def lazy_channel_recv'_def)
subgoal for _ _ _ _ x
by(cases "(snd s, x)" rule: lazy_channel_insec.cases)
(auto simp add: vld_def in_nlists_UNIV rnd_def ran_def split: option.split_asm)
subgoal by(cases "(snd s, ())" rule: lazy_channel_recv.cases)(auto 4 3 split: option.split_asm if_split_asm cstate.split_asm simp add: in_nlists_UNIV vld_def ran_def rnd_def)
done
done
then have "exec_gpv ?orc_rgt ?\<A> (False, Void, None, Map.empty) = exec_gpv ?orc_rgt \<A> (False, Void, None, Map.empty)"
using WT by(rule callee_invariant_on.exec_gpv_mk_lossless_gpv) simp
finally have "\<bar>Sigma_Algebra.measure (measure_spmf (game \<A> lazy_channel_recv_f')) {x. fst x}
- Sigma_Algebra.measure (measure_spmf (game \<A> lazy_channel_recv')) {x. fst x}\<bar>
\<le> Sigma_Algebra.measure (measure_spmf (game \<A> lazy_channel_recv_f')) {x. snd x}"
unfolding game_def
by - (rule fundamental_lemma[where ?bad2.0="snd"]; auto simp add: spmf_rel_map map_spmf_conv_bind_spmf[symmetric] split_def)
moreover have "Sigma_Algebra.measure (measure_spmf (game \<A> lazy_channel_recv_f')) {x. fst x} = spmf ?L True"
unfolding game_def
by(auto simp add: map_spmf_conv_bind_spmf exec_gpv_resource_of_oracle connect_def spmf_conv_measure_spmf)
(auto simp add: rel_pred_def intro!: rel_spmf_bind_reflI measure_spmf_parametric[where A="\<lambda>l r. fst l \<longleftrightarrow> r", THEN rel_funD, THEN rel_funD])
moreover have "spmf ?R True = Sigma_Algebra.measure (measure_spmf (game \<A> lazy_channel_recv')) {x. fst x}"
unfolding game_def
by(auto simp add: map_spmf_conv_bind_spmf exec_gpv_resource_of_oracle connect_def spmf_conv_measure_spmf)
(auto simp add: rel_pred_def intro!: rel_spmf_bind_reflI measure_spmf_parametric[where A="\<lambda>l r. l \<longleftrightarrow> fst r", THEN rel_funD, THEN rel_funD])
ultimately show ?thesis by simp
qed
have fact3: "spmf (connect \<A> (RES (lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv') (False, Void, None, Map.empty))) True
= spmf (connect \<A> (RES (lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv) (Void, None, Map.empty))) True"
proof -
let ?orc_lft = "lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv'"
let ?orc_rgt = "lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv"
have[simp]: "rel_spmf (rel_prod (=) (\<lambda>x y. y = snd x))
(lazy_channel_insec' (flg, s) q) (lazy_channel_insec s q)" for s flg q
by (cases s) (simp add: lazy_channel_insec'_def spmf_rel_map rel_prod_sel split_def option.rel_refl pmf.rel_refl split:option.split)
have[simp]: "rel_spmf (rel_prod (=) (\<lambda>x y. y = snd x))
(lazy_channel_send' (flg, s) q) (lazy_channel_send s q)" for s flg q
by(cases "(s, q)" rule: lazy_channel_send.cases)(auto simp add: spmf_rel_map map_spmf_conv_bind_spmf split_def lazy_channel_send'_def)
have[simp]: "rel_spmf (rel_prod (=) (\<lambda>x y. y = snd x))
(lazy_channel_recv' (flg, s) q) (lazy_channel_recv s q)" for s flg q
by(cases "(s, q)" rule: lazy_channel_recv.cases)(auto simp add: lazy_channel_recv'_def rel_spmf_bind_reflI split:option.split cstate.split)
have[simp]: "rel_spmf (rel_prod (=) (\<lambda>x y. y = snd x))
((lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv') (flg, ams, es, as) query)
((lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv) (ams, es, as) query)" for flg ams es as query
proof (cases query)
case (Inl adv_query)
then show ?thesis
by(auto simp add: spmf_rel_map map_spmf_conv_bind_spmf[symmetric] apfst_def map_prod_def split_beta intro: rel_spmf_mono[where A="rel_prod (=) (\<lambda>x y. y = snd x)"])
next
case (Inr query')
note Snd_Rcv = this
then show ?thesis
by (cases query', auto simp add: spmf_rel_map map_spmf_conv_bind_spmf[symmetric] split_beta)
((rule rel_spmf_mono[where A="rel_prod (=) (\<lambda>x y. y = snd x)"]); auto)+
qed
have[simp]: "rel_spmf (\<lambda>x y. fst x = fst y)
(exec_gpv ?orc_lft \<A> (False, Void, None, Map.empty)) (exec_gpv ?orc_rgt \<A> (Void, None, Map.empty))"
by(rule rel_spmf_mono[where A="rel_prod (=) (\<lambda>x y. y = snd x)"])
(auto simp add: gpv.rel_eq split_def intro!: rel_funI
exec_gpv_parametric[where CALL="(=)", THEN rel_funD, THEN rel_funD, THEN rel_funD])
show ?thesis
unfolding map_spmf_conv_bind_spmf exec_gpv_resource_of_oracle connect_def spmf_conv_measure_spmf
by(rule measure_spmf_parametric[where A="(=)", THEN rel_funD, THEN rel_funD])
(auto simp add: rel_pred_def spmf_rel_map map_spmf_conv_bind_spmf[symmetric] gpv.rel_eq split_def intro!: rel_funI)
qed
have fact4: "Sigma_Algebra.measure (measure_spmf (game \<A> lazy_channel_recv_f')) {x. snd x} \<le> q / real (2 ^ \<eta>)"
proof -
let ?orc_sum = "lazy_channel_insec' \<oplus>\<^sub>O lazy_channel_send' \<oplus>\<^sub>O lazy_channel_recv_f'"
have "Sigma_Algebra.measure (measure_spmf (game \<A> lazy_channel_recv_f')) {x. snd x}
= spmf (map_spmf (fst \<circ> snd) (exec_gpv ?orc_sum \<A> (False, Void, None, Map.empty))) True"
unfolding game_def
by (simp add: split_def map_spmf_conv_bind_spmf[symmetric] measure_map_spmf)
(simp add: spmf_conv_measure_spmf measure_map_spmf vimage_def)
also have "\<dots> \<le> (1 / real (card (nlists (UNIV :: bool set) \<eta>))) * real q"
proof -
have *: "spmf (map_spmf (fst \<circ> snd) (?orc_sum (False, ams, es, as) query)) True
\<le> 1 / real (card (nlists (UNIV :: bool set) \<eta>))" for ams es as query
proof (cases query)
case (Inl adv_query)
then show ?thesis
by(cases "((ams, es, as), adv_query)" rule: lazy_channel_insec.cases)
(auto simp add: lazy_channel_insec'_def rnd_def map_spmf_conv_bind_spmf bind_spmf_const split: option.split)
next
case (Inr query')
note Snd_Rcv = this
then show ?thesis
proof (cases query')
case (Inr rcv_query)
with Snd_Rcv show ?thesis
by (cases ams, auto simp add: lazy_channel_recv_f'_def map_spmf_conv_bind_spmf split_def split:option.split)
(auto simp add: spmf_of_set map_spmf_conv_bind_spmf[symmetric] rnd_def spmf_map vimage_def spmf_conv_measure_spmf[symmetric] split: split_indicator)
qed (cases ams, simp_all add: lazy_channel_send'_def)
qed
show ?thesis by (rule oi_True.interaction_bounded_by_exec_gpv_bad[where bad="fst"]) (auto simp add: * assms)
qed
also have "... = 1 / real (2 ^ \<eta>) * real q" by (simp add: card_nlists)
finally show ?thesis by simp
qed
have "?LHS \<le> Sigma_Algebra.measure (measure_spmf (game \<A> lazy_channel_recv_f')) {x. snd x}"
using fact1 fact2 fact3 by arith
also note fact4
finally show ?thesis .
qed
private inductive S' :: "(((bool list \<times> bool list) option + unit) \<times> unit \<times> bool list cstate) spmf \<Rightarrow>
(bool list cstate \<times> (bool list \<times> bool list) option \<times> (bool list \<Rightarrow> bool list option)) spmf \<Rightarrow> bool" where
"S' (return_spmf (Inl None, (), Void))
(return_spmf (Void, None, Map.empty))"
| "S' (return_spmf (Inl None, (), Store m))
(return_spmf (Store m, None, Map.empty))"
if "length m = id' \<eta>"
| "S' (return_spmf (Inr (), (), Collect m))
(return_spmf (Collect m, None, Map.empty))"
if "length m = id' \<eta>"
| "S' (return_spmf (Inl (Some (a, m)), (), Store m))
(return_spmf (Store m, None, [m \<mapsto> a]))"
if "length m = id' \<eta>"
| "S' (return_spmf (Inr (), (), Collect m))
(return_spmf (Collect m, None, [m \<mapsto> a]))"
if "length m = id' \<eta>"
| "S' (return_spmf (Inr (), (), Fail))
(return_spmf (Fail, None, Map.empty))"
| "S' (return_spmf (Inr (), (), Fail))
(return_spmf (Fail, None, [m \<mapsto> x]))"
if "length m = id' \<eta>"
| "S' (return_spmf (Inr (), (), Void))
(return_spmf (Collect m', Some (a', m'), Map.empty))"
if "length m' = id' \<eta>" and "length a' = id' \<eta>"
| "S' (return_spmf (Inr (), (), Fail))
(return_spmf (Collect m', Some (a', m'), Map.empty))"
if "length m' = id' \<eta>" and "length a' = id' \<eta>"
| "S' (return_spmf (Inr (), (), Store m))
(return_spmf (Collect m', Some (a', m'), Map.empty))"
if "length m = id' \<eta>" and "length m' = id' \<eta>" and "length a' = id' \<eta>"
| "S' (return_spmf (Inl (Some (a', m')), (), Collect m'))
(return_spmf (Collect m', Some (a', m'), [m' \<mapsto> a']))"
if "length m' = id' \<eta>" and "length a' = id' \<eta>"
| "S' (return_spmf (Inl None, (), cstate.Collect m))
(return_spmf (cstate.Collect m, None, Map.empty))"
if "length m = id' \<eta>"
| "S' (return_spmf (Inl None, (), cstate.Fail))
(return_spmf (cstate.Fail, None, Map.empty))"
| "S' (return_spmf (Inr (), (), Fail))
(return_spmf (Collect m', Some (a', m'), [m \<mapsto> a]))"
if "length m = id' \<eta>" and "length m' = id' \<eta>" and "length a' = id' \<eta>" and "m \<noteq> m'"
| "S' (return_spmf (Inr (), (), Fail))
(return_spmf (Collect m', Some (a', m'), [m \<mapsto> a]))"
if "length m = id' \<eta>" and "length m' = id' \<eta>" and "length a' = id' \<eta>" and "a \<noteq> a'"
| "S' (return_spmf (Inl None, (), Collect m'))
(return_spmf (Collect m', Some (a', m'), [m' \<mapsto> a']))"
if "length m' = id' \<eta>" and "length a' = id' \<eta>"
| "S' (return_spmf (Inr (), (), Collect m'))
(return_spmf (Collect m', Some (a', m'), [m' \<mapsto> a']))"
if "length m' = id' \<eta>" and "length a' = id' \<eta>"
| "S' (return_spmf (Inr (), (), Void))
(map_spmf (\<lambda>a'. (Fail, None, [m' \<mapsto> a'])) (spmf_of_set (nlists UNIV \<eta>)))"
if "length m' = id' \<eta>"
| "S' (return_spmf (Inr (), (), Fail))
(map_spmf (\<lambda>a'. (Fail, None, [m' \<mapsto> a'])) (spmf_of_set (nlists UNIV \<eta>)))"
if "length m' = id' \<eta>"
| "S' (return_spmf (Inr (), (), Store m))
(map_spmf (\<lambda>a'. (Fail, None, [m' \<mapsto> a'])) (spmf_of_set (nlists UNIV \<eta>)))"
if "length m = id' \<eta>" and "length m' = id' \<eta>"
| "S' (return_spmf (Inr (), (), Fail))
(map_spmf (\<lambda>a'. (Fail, None, [m \<mapsto> a, m' \<mapsto> a'])) (spmf_of_set (nlists UNIV \<eta>)))"
if "length m = id' \<eta>" and "length m' = id' \<eta>" and "m \<noteq> m'"
| "S' (return_spmf (Inl (Some (a', m')), (), Fail))
(return_spmf (Fail, None, [m' \<mapsto> a']))"
if "length m' = id' \<eta>" and "length a' = id' \<eta>"
| "S' (return_spmf (Inl None, (), Fail))
(return_spmf (Fail, None, [m' \<mapsto> a']))"
if "length m' = id' \<eta>" and "length a' = id' \<eta>"
private lemma trace_eq_sim:
shows "(valid_insecQ <+> nlists UNIV (id' \<eta>) <+> UNIV) \<turnstile>\<^sub>R
RES (callee_auth_channel sim \<oplus>\<^sub>O \<dagger>\<dagger>channel.send_oracle \<oplus>\<^sub>O \<dagger>\<dagger>channel.recv_oracle) (Inl None, (), Void)
\<approx>
RES (lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv_f) (Void, None, Map.empty)"
(is "?A \<turnstile>\<^sub>R RES (?L1 \<oplus>\<^sub>O ?L2 \<oplus>\<^sub>O ?L3) ?SL \<approx> RES (?R1 \<oplus>\<^sub>O ?R2 \<oplus>\<^sub>O ?R3) ?SR")
proof -
note [simp] =
spmf.map_comp o_def map_bind_spmf bind_map_spmf bind_spmf_const exec_gpv_bind
auth_channel.auth_oracle.simps channel.send_oracle.simps channel.recv_oracle.simps
rorc_channel_send_def rorc_channel_recv_def rnd_def
have *: "?A \<turnstile>\<^sub>C ?L1 \<oplus>\<^sub>O ?L2 \<oplus>\<^sub>O ?L3(?SL) \<approx> ?R1 \<oplus>\<^sub>O ?R2 \<oplus>\<^sub>O ?R3(?SR)"
proof(rule trace'_eqI_sim[where S=S'], goal_cases Init_OK Output_OK State_OK)
case Init_OK
then show ?case by (rule S'.intros)
next
case (Output_OK p q query)
show ?case
proof (cases query)
case (Inl adv_query)
with Output_OK show ?thesis
proof (cases adv_query)
case (ForwardOrEdit foe)
with Output_OK Inl show ?thesis
proof (cases foe)
case (Some am')
obtain a' m' where "am' = (a', m')" by (cases am') simp
with Output_OK Inl ForwardOrEdit Some show ?thesis
by cases (simp_all add: id'_def)
qed (erule S'.cases, simp_all add: id'_def)
qed (erule S'.cases, simp_all add: id'_def)+
next
case (Inr query')
with Output_OK show ?thesis by (cases; cases query') (simp_all)
qed
next
case (State_OK p q query state answer state')
show ?case
proof (cases query)
case (Inl adv_query)
show ?thesis
proof (cases adv_query)
case Look
with State_OK Inl show ?thesis
proof cases
case (2 m)
have "S' (return_spmf (Inl (Some (x, m)), (), Store m)) (return_spmf (Store m, None, [m \<mapsto> x]))" for x
by (rule S'.intros) (simp only: 2)
with 2 show ?thesis using State_OK(2-) Inl Look
by clarsimp (simp add: cond_spmf_spmf_of_set spmf_of_set_singleton map_spmf_conv_bind_spmf[symmetric]
split_beta cond_spmf_fst_def image_def vimage_def id'_def)
qed (auto simp add: map_spmf_const[symmetric] map_spmf_conv_bind_spmf[symmetric] image_def intro: S'.intros; simp add: id'_def)+
next
case (ForwardOrEdit foe)
show ?thesis
proof (cases foe)
case None
with State_OK Inl ForwardOrEdit show ?thesis
by cases(auto simp add: map_spmf_const[symmetric] map_spmf_conv_bind_spmf[symmetric] image_def S'.intros)
next
case (Some am')
obtain a' m' where [simp]: "am' = (a', m')" by (cases am')
from State_OK Inl ForwardOrEdit Some show ?thesis
proof cases
case (4 m a)
then show ?thesis using State_OK(2-) Inl ForwardOrEdit Some
by (auto simp add: if_distrib_exec_gpv if_distrib_map_spmf split_def image_def if_distrib[symmetric] intro: S'.intros cong: if_cong)
qed (auto simp add: map_spmf_const[symmetric] map_spmf_conv_bind_spmf[symmetric] image_def intro:S'.intros)
qed
next
case Drop
with State_OK Inl show ?thesis
by cases (auto simp add: map_spmf_const[symmetric] map_spmf_conv_bind_spmf[symmetric] image_def intro:S'.intros; simp add: id'_def)+
qed
next
case Snd_Rcv: (Inr query')
with State_OK show ?thesis
by(cases; cases query')
(auto simp add: map_spmf_const[symmetric] map_spmf_conv_bind_spmf[symmetric] image_def;
(rule S'.intros; simp add: in_nlists_UNIV id'_def))+
qed
qed
from * show ?thesis by simp
qed
private lemma real_resource_wiring: "macode.res (rnd \<eta>) (mac \<eta>) =
RES (\<dagger>\<dagger>insec_channel.insec_oracle \<oplus>\<^sub>O rorc_channel_send \<oplus>\<^sub>O rorc_channel_recv) ((False, ()), Map.empty, Void)"
(is "?L = ?R") including lifting_syntax
proof -
have *: "(\<lambda>x y. map_spmf (map_prod id lprodr) ((A \<oplus>\<^sub>O B) (rprodl x) y))
= (\<lambda>x yl. map_spmf (\<lambda>p. ((), lprodr (snd p))) (A (rprodl x) yl)) \<oplus>\<^sub>O
(\<lambda>x yr. map_spmf (\<lambda>p. (fst p, lprodr (snd p))) (B (rprodl x) yr)) " for A B C
proof -
have aux: "map_prod id g \<circ> apfst h = apfst h \<circ> map_prod id g" for f g h by auto
show ?thesis
by (auto simp add: aux plus_oracle_def spmf.map_comp[where f="apfst _", symmetric]
spmf.map_comp[where f="map_prod id lprodr"] sum.case_distrib[where h="map_spmf _"] cong:sum.case_cong split!:sum.splits)
(subst plus_oracle_def[symmetric], simp add: map_prod_def split_def)
qed
have "?L = RES (\<dagger>\<dagger>insec_channel.insec_oracle \<oplus>\<^sub>O (rprodl ---> id ---> map_spmf (map_prod id lprodr))
(lift_state_oracle extend_state_oracle (attach_callee
(\<lambda>s m. if s
then Done ((), True)
else do {
r \<leftarrow> Pause (Inl m) Done;
a \<leftarrow> lift_spmf (mac \<eta> (projl r) m);
_ \<leftarrow> Pause (Inr (a, m)) Done;
( Done ((), True))}) ((rorc.rnd_oracle (rnd \<eta>))\<dagger> \<oplus>\<^sub>O \<dagger>channel.send_oracle)) \<oplus>\<^sub>O
\<dagger>\<dagger>(\<lambda>s q. do {
(amo, s') \<leftarrow> \<dagger>channel.recv_oracle s ();
case amo of
None \<Rightarrow> return_spmf (None, s')
| Some (a, m) \<Rightarrow> do {
(r, s'') \<leftarrow> (rorc.rnd_oracle (rnd \<eta>))\<dagger> s' m;
a'\<leftarrow> mac \<eta> r m;
return_spmf (if a' = a then Some m else None, s'')}}))) ((False, ()), Map.empty, Void)"
proof -
note[simp] = attach_CNV_RES attach_callee_parallel_intercept attach_stateless_callee
resource_of_oracle_rprodl Rel_def prod.rel_eq[symmetric] extend_state_oracle_plus_oracle[symmetric]
conv_callee_parallel[symmetric] conv_callee_parallel_id_left[where s="()", symmetric]
o_def bind_map_spmf exec_gpv_bind split_def option.case_distrib[where h="\<lambda>gpv. exec_gpv _ gpv _"]
show ?thesis
unfolding channel.res_def rorc.res_def macode.res_def macode.routing_def
macode.\<pi>E_def macode.enm_def macode.dem_def interface_wiring
by (subst lift_state_oracle_extend_state_oracle | auto cong del: option.case_cong_weak intro: extend_state_oracle_parametric)+
qed
also have "... = ?R"
unfolding rorc_channel_send_def rorc_channel_recv_def extend_state_oracle_def
by(simp add: * split_def o_def map_fun_def spmf.map_comp extend_state_oracle_def lift_state_oracle_def
exec_gpv_bind if_distrib[where f="\<lambda>gpv. exec_gpv _ gpv _"] cong: if_cong)
(simp add: split_def o_def rprodl_def Pair_fst_Unity bind_map_spmf map_spmf_bind_spmf
if_distrib[where f="map_spmf _"] option.case_distrib[where h="map_spmf _"] cong: if_cong option.case_cong)
finally show ?thesis .
qed
private lemma ideal_resource_wiring: "(CNV callee s) |\<^sub>= 1\<^sub>C \<rhd> channel.res auth_channel.auth_oracle =
RES (callee_auth_channel callee \<oplus>\<^sub>O \<dagger>\<dagger>channel.send_oracle \<oplus>\<^sub>O \<dagger>\<dagger>channel.recv_oracle ) (s, (), Void)" (is "?L1 \<rhd> _ = ?R")
proof -
have[simp]: "\<I>_full, \<I>_full \<oplus>\<^sub>\<I> (\<I>_full \<oplus>\<^sub>\<I> \<I>_full) \<turnstile>\<^sub>C ?L1 \<sim> ?L1" (is "_, ?I \<turnstile>\<^sub>C _ \<sim> _")
by(rule eq_\<I>_converter_mono)
(rule parallel_converter2_eq_\<I>_cong eq_\<I>_converter_reflI WT_converter_\<I>_full \<I>_full_le_plus_\<I> order_refl plus_\<I>_mono)+
have[simp]: "?I \<turnstile>c (auth_channel.auth_oracle \<oplus>\<^sub>O channel.send_oracle \<oplus>\<^sub>O channel.recv_oracle) s \<surd>" for s
by(rule WT_plus_oracleI WT_parallel_oracle WT_callee_full; (unfold split_paired_all)?)+
have[simp]: "?L1 \<rhd> RES (auth_channel.auth_oracle \<oplus>\<^sub>O channel.send_oracle \<oplus>\<^sub>O channel.recv_oracle) Void = ?R"
by(simp add: conv_callee_parallel_id_right[where s'="()", symmetric] attach_CNV_RES
attach_callee_parallel_intercept resource_of_oracle_rprodl extend_state_oracle_plus_oracle)
show ?thesis unfolding channel.res_def
by (subst eq_\<I>_attach[OF WT_resource_of_oracle, where \<I>' = "?I" and conv="?L1" and conv'="?L1"]) simp_all
qed
lemma all_together:
defines "\<I> \<equiv> \<I>_uniform (Set.Collect (valid_mac_query \<eta>)) (insert None (Some ` (nlists UNIV \<eta> \<times> nlists UNIV \<eta>))) \<oplus>\<^sub>\<I>
(\<I>_uniform (vld \<eta>) UNIV \<oplus>\<^sub>\<I> \<I>_uniform UNIV (insert None (Some ` vld \<eta>)))"
assumes "\<eta> > 0"
and "interaction_bounded_by' (\<lambda>_. True) (\<A> \<eta>) q"
and lossless: "plossless_gpv \<I> (\<A> \<eta>)"
and WT: "\<I> \<turnstile>g \<A> \<eta> \<surd>"
shows
"\<bar>spmf (connect (\<A> \<eta>) (CNV sim (Inl None) |\<^sub>= 1\<^sub>C \<rhd> channel.res auth_channel.auth_oracle)) True -
spmf (connect (\<A> \<eta>) (macode.res (rnd \<eta>) (mac \<eta>))) True\<bar> \<le> q / real (2 ^ \<eta>)"
proof -
have *: "\<forall>a b. ma = Edit (a, b) \<longrightarrow> length a = \<eta> \<and> length b = \<eta> \<Longrightarrow> valid_mac_query \<eta> ma" for ma a b
by(cases "(\<eta>, ma)" rule: valid_mac_query.cases)(auto simp add: vld_def in_nlists_UNIV)
have assm4_alt: "outs_gpv \<I> (\<A> \<eta>) \<subseteq> valid_insecQ <+> nlists UNIV (id' \<eta>) <+> UNIV"
using assms(5)[THEN WT_gpv_outs_gpv] unfolding id'_def
by(rule ord_le_eq_trans) (auto simp add: * split: aquery.split option.split simp add: in_nlists_UNIV vld_def \<I>_def)
have "callee_invariant_on (callee_auth_channel sim \<oplus>\<^sub>O \<dagger>\<dagger>channel.send_oracle \<oplus>\<^sub>O \<dagger>\<dagger>channel.recv_oracle)
(\<lambda>(s1, _, s3). (\<forall>x y. s1 = Inl (Some (x, y)) \<longrightarrow> length x = \<eta> \<and> length y = \<eta>) \<and> pred_cstate (\<lambda>x. length x = \<eta>) s3) \<I>"
apply unfold_locales
subgoal for s x y s'
apply(cases "(fst s, projl x)" rule: sim.cases; cases "snd (snd s)") (*A MESS, but neither fastforce nor auto works for all 36 subgoals*)
apply(fastforce simp: channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)[1]
apply(fastforce simp: channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)[1]
apply(fastforce simp: channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)[1]
apply(fastforce simp: channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)[1]
apply(fastforce simp: auth_channel.auth_oracle.simps channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)[1]
apply(fastforce simp: auth_channel.auth_oracle.simps channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)[1]
apply(fastforce simp: auth_channel.auth_oracle.simps channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)[1]
apply(fastforce simp: auth_channel.auth_oracle.simps channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)[1]
apply(auto split: if_split_asm simp add: exec_gpv_bind auth_channel.auth_oracle.simps channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)[1]
apply(auto split: if_split_asm simp add: auth_channel.auth_oracle.simps channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)[1]
apply(fastforce split: if_split_asm simp add: exec_gpv_bind auth_channel.auth_oracle.simps channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)+
done
subgoal for s
apply(rule WT_calleeI)
subgoal for x
by(cases "(fst s, projl x)" rule: sim.cases; cases "snd (snd s)")
(auto split: if_split_asm simp add: exec_gpv_bind auth_channel.auth_oracle.simps channel.send_oracle.simps channel.recv_oracle.simps vld_def in_nlists_UNIV \<I>_def)
done
done
then have WT1: "\<I> \<turnstile>res RES (callee_auth_channel sim \<oplus>\<^sub>O \<dagger>\<dagger>channel.send_oracle \<oplus>\<^sub>O \<dagger>\<dagger>channel.recv_oracle) (Inl None, (), Void) \<surd>"
by(rule callee_invariant_on.WT_resource_of_oracle) simp
have "callee_invariant_on (lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv_f)
(\<lambda>(s1, s2, s3). pred_cstate (\<lambda>x. length x = \<eta>) s1 \<and> pred_option (\<lambda>(x, y). length x = \<eta> \<and> length y = \<eta>) s2 \<and> ran s3 \<subseteq> nlists UNIV \<eta>)
\<I>"
apply unfold_locales
subgoal for s x y s'
- apply(clarsimp simp add: \<I>_def; elim PlusE; clarsimp)
+ using [[simproc del: defined_all]] apply(clarsimp simp add: \<I>_def; elim PlusE; clarsimp)
subgoal for _ _ _ x'
by(cases "(s, x')" rule: lazy_channel_insec.cases)
(auto simp add: vld_def in_nlists_UNIV rnd_def split: option.split_asm)
subgoal by(cases "(s, projl (projr x))" rule: lazy_channel_send.cases)(auto simp add: vld_def in_nlists_UNIV)
subgoal by(cases "(s, ())" rule: lazy_channel_recv_f.cases)(auto 4 3 split: option.split_asm if_split_asm simp add: in_nlists_UNIV vld_def ran_def rnd_def)
done
subgoal for s
apply(clarsimp simp add: \<I>_def; intro conjI WT_calleeI; clarsimp)
subgoal for _ _ _ x
by(cases "(s, x)" rule: lazy_channel_insec.cases)
(auto simp add: vld_def in_nlists_UNIV rnd_def ran_def split: option.split_asm)
subgoal by(cases "(s, ())" rule: lazy_channel_recv_f.cases)(auto 4 3 split: option.split_asm if_split_asm simp add: in_nlists_UNIV vld_def ran_def rnd_def)
done
done
then have WT2: "\<I> \<turnstile>res RES (lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv_f) (Void, None, Map.empty) \<surd>"
by(rule callee_invariant_on.WT_resource_of_oracle) simp
have "callee_invariant_on (lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv)
(\<lambda>(s1, s2, s3). pred_cstate (\<lambda>x. length x = \<eta>) s1 \<and> pred_option (\<lambda>(x, y). length x = \<eta> \<and> length y = \<eta>) s2 \<and> ran s3 \<subseteq> nlists UNIV \<eta>)
\<I>"
apply unfold_locales
subgoal for s x y s'
- apply(clarsimp simp add: \<I>_def; elim PlusE; clarsimp)
+ using [[simproc del: defined_all]] apply(clarsimp simp add: \<I>_def; elim PlusE; clarsimp)
subgoal for _ _ _ x'
by(cases "(s, x')" rule: lazy_channel_insec.cases)
(auto simp add: vld_def in_nlists_UNIV rnd_def split: option.split_asm)
subgoal by(cases "(s, projl (projr x))" rule: lazy_channel_send.cases)(auto simp add: vld_def in_nlists_UNIV)
subgoal by(cases "(s, ())" rule: lazy_channel_recv.cases)(auto 4 3 split: option.split_asm if_split_asm simp add: in_nlists_UNIV vld_def ran_def rnd_def option.pred_set)
done
subgoal for s
apply(clarsimp simp add: \<I>_def; intro conjI WT_calleeI; clarsimp)
subgoal for _ _ _ x
by(cases "(s, x)" rule: lazy_channel_insec.cases)
(auto simp add: vld_def in_nlists_UNIV rnd_def ran_def split: option.split_asm)
subgoal by(cases "(s, ())" rule: lazy_channel_recv_f.cases)(auto 4 3 split: option.split_asm if_split_asm simp add: in_nlists_UNIV vld_def ran_def rnd_def)
done
done
then have WT3: "\<I> \<turnstile>res RES (lazy_channel_insec \<oplus>\<^sub>O lazy_channel_send \<oplus>\<^sub>O lazy_channel_recv) (Void, None, Map.empty) \<surd>"
by(rule callee_invariant_on.WT_resource_of_oracle) simp
have "callee_invariant_on (\<dagger>\<dagger>insec_channel.insec_oracle \<oplus>\<^sub>O rorc_channel_send \<oplus>\<^sub>O rorc_channel_recv)
(\<lambda>(_, m, s). ran m \<subseteq> nlists UNIV \<eta> \<and> pred_cstate (\<lambda>(x, y). length x = \<eta> \<and> length y = \<eta>) s) \<I>"
apply(unfold_locales)
subgoal for s x y s'
- apply(clarsimp simp add: \<I>_def; elim PlusE; clarsimp)
+ using [[simproc del: defined_all]] apply(clarsimp simp add: \<I>_def; elim PlusE; clarsimp)
subgoal for _ _ _ x'
by(cases "(snd (snd s), x')" rule: insec_channel.insec_oracle.cases)
(auto simp add: vld_def in_nlists_UNIV rnd_def insec_channel.insec_oracle.simps split: option.split_asm)
subgoal
by(cases "snd (snd s)")
(auto 4 3 simp add: channel.send_oracle.simps rorc_channel_send_def rorc.rnd_oracle.simps in_nlists_UNIV rnd_def vld_def mac_def ran_def split: option.split_asm if_split_asm)
subgoal
by(cases "snd (snd s)")
(auto 4 4 simp add: rorc_channel_recv_def channel.recv_oracle.simps rorc.rnd_oracle.simps rnd_def mac_def ran_def iff: in_nlists_UNIV split: option.split_asm if_split_asm)
done
subgoal for s
apply(clarsimp simp add: \<I>_def; intro conjI WT_calleeI; clarsimp)
subgoal for _ _ _ x'
by(cases "(snd (snd s), x')" rule: insec_channel.insec_oracle.cases)
(auto simp add: vld_def in_nlists_UNIV rnd_def insec_channel.insec_oracle.simps split: option.split_asm)
subgoal
by(cases "snd (snd s)")
(auto simp add: rorc_channel_recv_def channel.recv_oracle.simps rorc.rnd_oracle.simps rnd_def mac_def vld_def ran_def iff: in_nlists_UNIV split: option.split_asm if_split_asm)
done
done
then have WT4: "\<I> \<turnstile>res RES (\<dagger>\<dagger>insec_channel.insec_oracle \<oplus>\<^sub>O rorc_channel_send \<oplus>\<^sub>O rorc_channel_recv) ((False, ()), Map.empty, Void) \<surd>"
by(rule callee_invariant_on.WT_resource_of_oracle) simp
note game_difference[OF assms(3), folded \<I>_def, OF assms(4,5)]
also note connect_cong_trace[OF trace_eq_sim WT assm4_alt WT1 WT2, symmetric]
also note connect_cong_trace[OF trace_eq_lazy, OF assms(2), OF WT assm4_alt WT3 WT4]
also note ideal_resource_wiring[of sim, of "Inl None", symmetric]
also note real_resource_wiring[symmetric]
finally show ?thesis by simp
qed
end
context begin
interpretation MAC: macode "rnd \<eta>" "mac \<eta>" for \<eta> .
interpretation A_CHAN: auth_channel .
lemma WT_enm:
"X \<noteq> {} \<Longrightarrow> \<I>_uniform (vld \<eta>) UNIV, \<I>_uniform (vld \<eta>) X \<oplus>\<^sub>\<I> \<I>_uniform (X \<times> vld \<eta>) UNIV \<turnstile>\<^sub>C MAC.enm \<eta> \<surd>"
unfolding MAC.enm_def
by(rule WT_converter_of_callee) (auto simp add: mac_def)
lemma WT_dem: "\<I>_uniform UNIV (insert None (Some ` vld \<eta>)), \<I>_full \<oplus>\<^sub>\<I> \<I>_uniform UNIV (insert None (Some ` (nlists UNIV \<eta> \<times> nlists UNIV \<eta>))) \<turnstile>\<^sub>C MAC.dem \<eta> \<surd>"
unfolding MAC.dem_def
by (rule WT_converter_of_callee) (auto simp add: vld_def stateless_callee_def mac_def split: option.split_asm)
lemma valid_insec_query_of [simp]: "valid_mac_query \<eta> (insec_query_of x)"
by(cases x) simp_all
lemma secure_mac:
defines "\<I>_real \<equiv> \<lambda>\<eta>. \<I>_uniform {x. valid_mac_query \<eta> x} (insert None (Some ` (nlists UNIV \<eta> \<times> nlists UNIV \<eta>)))"
and "\<I>_ideal \<equiv> \<lambda>\<eta>. \<I>_uniform UNIV (insert None (Some ` nlists UNIV \<eta>))"
and "\<I>_common \<equiv> \<lambda>\<eta>. \<I>_uniform (vld \<eta>) UNIV \<oplus>\<^sub>\<I> \<I>_uniform UNIV (insert None (Some ` vld \<eta>))"
shows
"constructive_security MAC.res (\<lambda>_. A_CHAN.res) (\<lambda>_. CNV sim (Inl None))
\<I>_real \<I>_ideal \<I>_common (\<lambda>_. enat q) True (\<lambda>_. insec_auth_wiring)"
proof
show WT_res [WT_intro]: "\<I>_real \<eta> \<oplus>\<^sub>\<I> \<I>_common \<eta> \<turnstile>res MAC.res \<eta> \<surd>" for \<eta>
proof -
let ?I = "pred_cstate (\<lambda>(x, y). length x = \<eta> \<and> length y = \<eta>)"
have *: "callee_invariant_on (MAC.RO.rnd_oracle \<eta> \<oplus>\<^sub>O MAC.RO.rnd_oracle \<eta>) (\<lambda>m. ran m \<subseteq> vld \<eta>) (\<I>_uniform (vld \<eta>) (vld \<eta>) \<oplus>\<^sub>\<I> \<I>_full)" for \<eta>
apply unfold_locales
subgoal for s x y s' by(cases x; clarsimp split: option.split_asm; auto simp add: rnd_def vld_def)
subgoal by(clarsimp intro!: WT_calleeI split: option.split_asm; auto simp add: rnd_def in_nlists_UNIV vld_def ran_def)
done
have [WT_intro]: "\<I>_uniform (vld \<eta>) (vld \<eta>) \<oplus>\<^sub>\<I> \<I>_full \<turnstile>res MAC.RO.res \<eta> \<surd>" for \<eta>
unfolding MAC.RO.res_def by(rule callee_invariant_on.WT_resource_of_oracle[OF *]) simp
have [simp]: "\<I>_real \<eta> \<turnstile>c MAC.INSEC.insec_oracle s \<surd>" if "?I s" for s
apply(rule WT_calleeI)
subgoal for x using that by(cases "(s, x)" rule: MAC.INSEC.insec_oracle.cases)(auto simp add: \<I>_real_def in_nlists_UNIV)
done
have [simp]: " \<I>_uniform UNIV (insert None (Some ` (nlists UNIV \<eta> \<times> nlists UNIV \<eta>))) \<turnstile>c A_CHAN.recv_oracle s \<surd>"
if "?I s" for s :: "(bool list \<times> bool list) cstate" using that
by(cases s)(auto intro!: WT_calleeI simp add: in_nlists_UNIV)
have *: "callee_invariant_on (MAC.INSEC.insec_oracle \<oplus>\<^sub>O A_CHAN.send_oracle \<oplus>\<^sub>O A_CHAN.recv_oracle) ?I
(\<I>_real \<eta> \<oplus>\<^sub>\<I> (\<I>_uniform (vld \<eta> \<times> vld \<eta>) UNIV \<oplus>\<^sub>\<I> \<I>_uniform UNIV (insert None (Some ` (nlists UNIV \<eta> \<times> nlists UNIV \<eta>)))))"
apply unfold_locales
subgoal for s x y s'
by(cases s; cases "(s, projl x)" rule: MAC.INSEC.insec_oracle.cases)(auto simp add: \<I>_real_def vld_def in_nlists_UNIV)
subgoal by(auto intro: WT_calleeI)
done
have [WT_intro]: "\<I>_real \<eta> \<oplus>\<^sub>\<I> (\<I>_uniform (vld \<eta> \<times> vld \<eta>) UNIV \<oplus>\<^sub>\<I> \<I>_uniform UNIV (insert None (Some ` (nlists UNIV \<eta> \<times> nlists UNIV \<eta>)))) \<turnstile>res MAC.INSEC.res \<surd>"
unfolding MAC.INSEC.res_def
by(rule callee_invariant_on.WT_resource_of_oracle[OF *]) simp
show ?thesis
unfolding \<I>_common_def MAC.res_def
by(rule WT_intro WT_enm[where X="vld \<eta>"] WT_dem | (simp add: vld_def; fail))+
qed
let ?I = "\<lambda>\<eta>. pred_cstate (\<lambda>x. length x = \<eta>)"
have WT_auth: "\<I>_uniform UNIV (insert None (Some ` nlists UNIV \<eta>)) \<turnstile>c A_CHAN.auth_oracle s \<surd>"
if "?I \<eta> s" for \<eta> s
apply(rule WT_calleeI)
subgoal for x using that by(cases "(s, x)" rule: A_CHAN.auth_oracle.cases; auto simp add: in_nlists_UNIV)
done
have WT_recv: "\<I>_uniform UNIV (insert None (Some ` vld \<eta>)) \<turnstile>c A_CHAN.recv_oracle s \<surd>"
if "?I \<eta> s" for \<eta> s using that
by(cases s)(auto intro!: WT_calleeI simp add: vld_def in_nlists_UNIV)
have *: "callee_invariant_on (A_CHAN.auth_oracle \<oplus>\<^sub>O A_CHAN.send_oracle \<oplus>\<^sub>O A_CHAN.recv_oracle)
(?I \<eta>) (\<I>_ideal \<eta> \<oplus>\<^sub>\<I> \<I>_common \<eta>)" for \<eta>
apply(unfold_locales)
subgoal for s x y s'
by(cases "(s, projl x)" rule: A_CHAN.auth_oracle.cases; cases "projr x")(auto simp add: \<I>_common_def vld_def in_nlists_UNIV)
subgoal for s using WT_auth WT_recv by(auto simp add: \<I>_common_def \<I>_ideal_def intro: WT_calleeI)
done
show WT_auth: "\<I>_ideal \<eta> \<oplus>\<^sub>\<I> \<I>_common \<eta> \<turnstile>res A_CHAN.res \<surd>" for \<eta>
unfolding A_CHAN.res_def by(rule callee_invariant_on.WT_resource_of_oracle[OF *]) simp
let ?I_sim = "\<lambda>\<eta> (s :: (bool list \<times> bool list) option + unit). \<forall>x y. s = Inl (Some (x, y)) \<longrightarrow> length x = \<eta> \<and> length y = \<eta>"
have "\<I>_real \<eta>, \<I>_ideal \<eta> \<turnstile>\<^sub>C CNV sim s \<surd>" if "?I_sim \<eta> s" for \<eta> s using that
apply(coinduction arbitrary: s)
subgoal for q r conv' s by(cases "(s, q)" rule: sim.cases)(auto simp add: \<I>_real_def \<I>_ideal_def vld_def in_nlists_UNIV)
done
then show [WT_intro]: "\<I>_real \<eta>, \<I>_ideal \<eta> \<turnstile>\<^sub>C CNV sim (Inl None) \<surd>" for \<eta> by simp
{ fix \<A> :: "security \<Rightarrow> ((bool list \<times> bool list) insec_query + bool list + unit, (bool list \<times> bool list) option + unit + bool list option) distinguisher"
assume WT: "\<I>_real \<eta> \<oplus>\<^sub>\<I> \<I>_common \<eta> \<turnstile>g \<A> \<eta> \<surd>" for \<eta>
assume bounded[simplified]: "interaction_bounded_by' (\<lambda>_. True) (\<A> \<eta>) q" for \<eta>
assume lossless[simplified]: "True \<Longrightarrow> plossless_gpv (\<I>_real \<eta> \<oplus>\<^sub>\<I> \<I>_common \<eta>) (\<A> \<eta>)" for \<eta>
show "negligible (\<lambda>\<eta>. advantage (\<A> \<eta>) (CNV sim (Inl None) |\<^sub>= 1\<^sub>C \<rhd> A_CHAN.res) (MAC.res \<eta>))"
proof -
have f1: "negligible (\<lambda>\<eta>. q * (1 / real (2 ^ \<eta>)))" (is "negligible ?ov")
by(rule negligible_poly_times[where n=0]) (simp add: negligible_inverse_powerI)+
have f2: "(\<lambda>\<eta>. spmf (connect (\<A> \<eta>) (CNV sim (Inl None) |\<^sub>= 1\<^sub>C \<rhd> A_CHAN.res)) True -
spmf (connect (\<A> \<eta>) (MAC.res \<eta>)) True) \<in> O(?ov)" (is "?L \<in> _")
by (auto simp add: bigo_def intro!: all_together[simplified] bounded eventually_at_top_linorderI[where c=1]
exI[where x=1] lossless[unfolded \<I>_real_def \<I>_common_def] WT[unfolded \<I>_real_def \<I>_common_def])
have "negligible ?L" using f1 f2 by (rule negligible_mono[of "?ov"])
then show ?thesis by (simp add: advantage_def)
qed
next
let ?cnv = "map_converter id id insec_query_of auth_response_of 1\<^sub>C"
show "\<exists>cnv. \<forall>\<D>. (\<forall>\<eta>. \<I>_ideal \<eta> \<oplus>\<^sub>\<I> \<I>_common \<eta> \<turnstile>g \<D> \<eta> \<surd>) \<longrightarrow>
(\<forall>\<eta>. interaction_bounded_by' (\<lambda>_. True) (\<D> \<eta>) q) \<longrightarrow>
(\<forall>\<eta>. True \<longrightarrow> plossless_gpv (\<I>_ideal \<eta> \<oplus>\<^sub>\<I> \<I>_common \<eta>) (\<D> \<eta>)) \<longrightarrow>
(\<forall>\<eta>. wiring (\<I>_ideal \<eta>) (\<I>_real \<eta>) (cnv \<eta>) (insec_query_of, map_option snd)) \<and>
negligible (\<lambda>\<eta>. advantage (\<D> \<eta>) A_CHAN.res (cnv \<eta> |\<^sub>= 1\<^sub>C \<rhd> MAC.res \<eta>))"
proof(intro exI[where x="\<lambda>_. ?cnv"] strip conjI)
fix \<D> :: "security \<Rightarrow> (auth_query + bool list + unit, bool list option + unit + bool list option) distinguisher"
assume WT_D [rule_format, WT_intro]: "\<forall>\<eta>. \<I>_ideal \<eta> \<oplus>\<^sub>\<I> \<I>_common \<eta> \<turnstile>g \<D> \<eta> \<surd>"
let ?A = "\<lambda>\<eta>. UNIV <+> nlists UNIV \<eta> <+> UNIV"
have WT1: "\<I>_ideal \<eta>, map_\<I> insec_query_of (map_option snd) (\<I>_real \<eta>) \<turnstile>\<^sub>C 1\<^sub>C \<surd>" for \<eta>
using WT_converter_id order_refl by(rule WT_converter_mono)(auto simp add: le_\<I>_def \<I>_ideal_def \<I>_real_def)
have WT0: "\<I>_ideal \<eta> \<oplus>\<^sub>\<I> \<I>_common \<eta> \<turnstile>res map_converter id id insec_query_of (map_option snd) 1\<^sub>C |\<^sub>= 1\<^sub>C \<rhd> MAC.res \<eta> \<surd>" for \<eta>
by(rule WT1 WT_intro | simp)+
have WT2: "\<I>_ideal \<eta>, map_\<I> insec_query_of (map_option snd) (\<I>_real \<eta>) \<turnstile>\<^sub>C 1\<^sub>C \<surd>" for \<eta>
using WT_converter_id order_refl
by(rule WT_converter_mono)(auto simp add: le_\<I>_def \<I>_ideal_def \<I>_real_def)
have WT_cnv: "\<I>_ideal \<eta>, \<I>_real \<eta> \<turnstile>\<^sub>C ?cnv \<surd>" for \<eta>
by(rule WT_converter_map_converter)(simp_all add: WT2)
from eq_\<I>_converter_reflI[OF this] this
show "wiring (\<I>_ideal \<eta>) (\<I>_real \<eta>) ?cnv insec_auth_wiring" for \<eta> ..
define res' :: "security \<Rightarrow> _ \<Rightarrow> auth_query + (bool list + bool list \<times> bool list) + bool list + unit \<Rightarrow> _"
where "res' \<eta> =
map_fun id (map_fun insec_query_of (map_spmf (map_prod (map_option snd) id))) \<dagger>MAC.INSEC.insec_oracle \<oplus>\<^sub>O
((MAC.RO.rnd_oracle \<eta>)\<dagger> \<oplus>\<^sub>O \<dagger>A_CHAN.send_oracle) \<oplus>\<^sub>O (MAC.RO.rnd_oracle \<eta>)\<dagger> \<oplus>\<^sub>O \<dagger>A_CHAN.recv_oracle"
for \<eta>
have push: "map_resource (map_sum f id) (map_sum g id) ((1\<^sub>C |\<^sub>= conv) \<rhd> res) =
(1\<^sub>C |\<^sub>= conv) \<rhd> map_resource (map_sum f id) (map_sum g id) res"
for f g conv res
proof -
have "map_resource (map_sum f id) (map_sum g id) ((1\<^sub>C |\<^sub>= conv) \<rhd> res) = map_converter f g id id 1\<^sub>C |\<^sub>= conv \<rhd> res"
by(simp add: attach_map_converter parallel_converter2_map1_out sum.map_id0)
also have "\<dots> = (1\<^sub>C |\<^sub>= conv) \<rhd> map_resource (map_sum f id) (map_sum g id) res"
by(subst map_converter_id_move_right)(simp add: parallel_converter2_map1_out sum.map_id0 attach_map_converter)
finally show ?thesis .
qed
have res': "map_resource (map_sum insec_query_of id) (map_sum (map_option snd) id) (MAC.res \<eta>) =
1\<^sub>C |\<^sub>= MAC.enm \<eta> |\<^sub>= MAC.dem \<eta> \<rhd> RES (res' \<eta>) (Map.empty, Void)" for \<eta>
unfolding MAC.res_def MAC.RO.res_def MAC.INSEC.res_def interface_wiring push
by(simp add: parallel_converter2_map1_out sum.map_id0 attach_map_converter map_resource_resource_of_oracle map_sum_plus_oracle prod.map_id0 option.map_id0 map_fun_id res'_def)
define res'' :: "security \<Rightarrow> (unit \<times> bool \<times> unit) \<times> (bool list \<Rightarrow> bool list option) \<times> _ cstate \<Rightarrow> auth_query + bool list + unit \<Rightarrow> _"
where "res'' \<eta> = map_fun rprodl (map_fun id (map_spmf (map_prod id lprodr)))
(lift_state_oracle extend_state_oracle \<dagger>(map_fun id (map_fun insec_query_of (map_spmf (map_prod (map_option snd) id))) \<dagger>MAC.INSEC.insec_oracle) \<oplus>\<^sub>O
\<dagger>(map_fun rprodl (map_fun id (map_spmf (map_prod id lprodr)))
(lift_state_oracle extend_state_oracle
(attach_callee
(\<lambda>bs m. if bs then Done ((), True) else Pause (Inl m) Done \<bind> (\<lambda>r. lift_spmf (mac \<eta> (projl r) m) \<bind> (\<lambda>a. Pause (Inr (a, m)) Done \<bind> (\<lambda>_. Done ((), True)))))
((MAC.RO.rnd_oracle \<eta>)\<dagger> \<oplus>\<^sub>O \<dagger>A_CHAN.send_oracle)) \<oplus>\<^sub>O
\<dagger>\<dagger>(\<lambda>s q. \<dagger>A_CHAN.recv_oracle s () \<bind>
(\<lambda>x. case x of (None, s') \<Rightarrow> return_spmf (None, s')
| (Some (x1, x2a), s') \<Rightarrow> (MAC.RO.rnd_oracle \<eta>)\<dagger> s' x2a \<bind> (\<lambda>(x, s'). mac \<eta> x x2a \<bind> (\<lambda>y. return_spmf (if y = x1 then Some x2a else None, s'))))))))"
for \<eta>
have "?cnv |\<^sub>= 1\<^sub>C \<rhd> MAC.res \<eta> = 1\<^sub>C |\<^sub>= MAC.enm \<eta> |\<^sub>= MAC.dem \<eta> \<rhd> RES (res' \<eta>) (Map.empty, Void)" for \<eta>
by(simp add: parallel_converter2_map1_out attach_map_converter sum.map_id0 res' attach_compose[symmetric] comp_converter_parallel2 comp_converter_id_left)
also have "\<dots> \<eta> = RES (res'' \<eta>) (((), False, ()), Map.empty, Void)" for \<eta>
unfolding MAC.enm_def MAC.dem_def conv_callee_parallel[symmetric] conv_callee_parallel_id_left[where s="()", symmetric] attach_CNV_RES
unfolding res'_def res''_def attach_callee_parallel_intercept attach_stateless_callee attach_callee_id_oracle prod.rel_eq[symmetric]
by(simp add: extend_state_oracle_plus_oracle[symmetric] rprodl_extend_state_oracle sum.case_distrib[where h="\<lambda>x. exec_gpv _ x _"]
option.case_distrib[where h="\<lambda>x. exec_gpv _ x _"] prod.case_distrib[where h="\<lambda>x. exec_gpv _ x _"] exec_gpv_bind bind_map_spmf o_def
cong: sum.case_cong option.case_cong)
also
define S :: "security \<Rightarrow> bool list cstate \<Rightarrow> (unit \<times> bool \<times> unit) \<times> (bool list \<Rightarrow> bool list option) \<times> (bool list \<times> bool list) cstate \<Rightarrow> bool"
where "S \<eta> l r = (case (l, r) of
(Void, ((_, False, _), m, Void)) \<Rightarrow> m = Map.empty
| (Store x, ((_, True, _), m, Store (y, z))) \<Rightarrow> length x = \<eta> \<and> length y = \<eta> \<and> length z = \<eta> \<and> m = [z \<mapsto> y] \<and> x = z
| (Collect x, ((_, True, _), m, Collect (y, z))) \<Rightarrow> length x = \<eta> \<and> length y = \<eta> \<and> length z = \<eta> \<and> m = [z \<mapsto> y] \<and> x = z
| (Fail, ((_, True, _), m, Fail)) \<Rightarrow> True
| _ \<Rightarrow> False)" for \<eta> l r
note [simp] = spmf_rel_map bind_map_spmf exec_gpv_bind
have "connect (\<D> \<eta>) (?cnv |\<^sub>= 1\<^sub>C \<rhd> MAC.res \<eta>) = connect (\<D> \<eta>) A_CHAN.res" for \<eta>
unfolding calculation using WT_D _ WT_auth
apply(rule connect_eq_resource_cong[symmetric])
unfolding A_CHAN.res_def
apply(rule eq_resource_on_resource_of_oracleI[where S="S \<eta>"])
apply(rule rel_funI)+
subgoal for s s' q q'
by(cases q; cases "projl q"; cases "projr q"; clarsimp simp add: eq_on_def S_def res''_def split: cstate.split_asm bool.split_asm; clarsimp simp add: rel_spmf_return_spmf1 rnd_def mac_def bind_UNION \<I>_common_def vld_def in_nlists_UNIV S_def)+
subgoal by(simp add: S_def)
done
then show adv: "negligible (\<lambda>\<eta>. advantage (\<D> \<eta>) A_CHAN.res (?cnv |\<^sub>= 1\<^sub>C \<rhd> MAC.res \<eta>))"
by(simp add: advantage_def)
qed
}
qed
end
end
diff --git a/thys/CoreC++/Execute.thy b/thys/CoreC++/Execute.thy
--- a/thys/CoreC++/Execute.thy
+++ b/thys/CoreC++/Execute.thy
@@ -1,1377 +1,1377 @@
(* Title: CoreC++
Author: Daniel Wasserrab, Stefan Berghofer
Maintainer: Daniel Wasserrab <wasserra at fmi.uni-passau.de>
*)
section \<open>Code generation for Semantics and Type System\<close>
theory Execute
imports BigStep WellType
"HOL-Library.AList_Mapping"
"HOL-Library.Code_Target_Numeral"
begin
subsection\<open>General redefinitions\<close>
inductive app :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool"
where
"app [] ys ys"
| "app xs ys zs \<Longrightarrow> app (x # xs) ys (x # zs)"
theorem app_eq1: "\<And>ys zs. zs = xs @ ys \<Longrightarrow> app xs ys zs"
apply (induct xs)
apply simp
apply (rule app.intros)
apply simp
apply (iprover intro: app.intros)
done
theorem app_eq2: "app xs ys zs \<Longrightarrow> zs = xs @ ys"
by (erule app.induct) simp_all
theorem app_eq: "app xs ys zs = (zs = xs @ ys)"
apply (rule iffI)
apply (erule app_eq2)
apply (erule app_eq1)
done
code_pred
(modes:
i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool, i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool,
o \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, o \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool as reverse_app)
app
.
declare rtranclp_rtrancl_eq[code del]
lemmas [code_pred_intro] = rtranclp.rtrancl_refl converse_rtranclp_into_rtranclp
code_pred
(modes:
(i => o => bool) => i => i => bool,
(i => o => bool) => i => o => bool)
rtranclp
by(erule converse_rtranclpE) blast+
definition Set_project :: "('a \<times> 'b) set => 'a => 'b set"
where "Set_project A a = {b. (a, b) \<in> A}"
lemma Set_project_set [code]:
"Set_project (set xs) a = set (List.map_filter (\<lambda>(a', b). if a = a' then Some b else None) xs)"
by(auto simp add: Set_project_def map_filter_def intro: rev_image_eqI split: if_split_asm)
text\<open>Redefine map Val vs\<close>
inductive map_val :: "expr list \<Rightarrow> val list \<Rightarrow> bool"
where
Nil: "map_val [] []"
| Cons: "map_val xs ys \<Longrightarrow> map_val (Val y # xs) (y # ys)"
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> o \<Rightarrow> bool)
map_val
.
inductive map_val2 :: "expr list \<Rightarrow> val list \<Rightarrow> expr list \<Rightarrow> bool"
where
Nil: "map_val2 [] [] []"
| Cons: "map_val2 xs ys zs \<Longrightarrow> map_val2 (Val y # xs) (y # ys) zs"
| Throw: "map_val2 (throw e # xs) [] (throw e # xs)"
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool)
map_val2
.
theorem map_val_conv: "(xs = map Val ys) = map_val xs ys"
(*<*)
proof -
have "\<And>ys. xs = map Val ys \<Longrightarrow> map_val xs ys"
apply (induct xs type:list)
apply (case_tac ys)
apply simp
apply (rule map_val.Nil)
apply simp
apply (case_tac ys)
apply simp
apply simp
- apply (erule conjE)
+
apply (rule map_val.Cons)
apply simp
done
moreover have "map_val xs ys \<Longrightarrow> xs = map Val ys"
by (erule map_val.induct) simp+
ultimately show ?thesis ..
qed
(*>*)
theorem map_val2_conv:
"(xs = map Val ys @ throw e # zs) = map_val2 xs ys (throw e # zs)"
(*<*)
proof -
have "\<And>ys. xs = map Val ys @ throw e # zs \<Longrightarrow> map_val2 xs ys (throw e # zs)"
apply (induct xs type:list)
apply (case_tac ys)
apply simp
apply simp
apply simp
apply (case_tac ys)
apply simp
apply (rule map_val2.Throw)
apply simp
apply (rule map_val2.Cons)
apply simp
done
moreover have "map_val2 xs ys (throw e # zs) \<Longrightarrow> xs = map Val ys @ throw e # zs"
by (erule map_val2.induct) simp+
ultimately show ?thesis ..
qed
(*>*)
subsection\<open>Code generation\<close>
lemma subclsRp_code [code_pred_intro]:
"\<lbrakk> class P C = \<lfloor>(Bs, rest)\<rfloor>; Predicate_Compile.contains (set Bs) (Repeats D) \<rbrakk> \<Longrightarrow> subclsRp P C D"
by(auto intro: subclsRp.intros simp add: contains_def)
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool)
subclsRp
by(erule subclsRp.cases)(fastforce simp add: Predicate_Compile.contains_def)
lemma subclsR_code [code_pred_inline]:
"P \<turnstile> C \<prec>\<^sub>R D \<longleftrightarrow> subclsRp P C D"
by(simp add: subclsR_def)
lemma subclsSp_code [code_pred_intro]:
"\<lbrakk> class P C = \<lfloor>(Bs, rest)\<rfloor>; Predicate_Compile.contains (set Bs) (Shares D) \<rbrakk> \<Longrightarrow> subclsSp P C D"
by(auto intro: subclsSp.intros simp add: Predicate_Compile.contains_def)
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool)
subclsSp
by(erule subclsSp.cases)(fastforce simp add: Predicate_Compile.contains_def)
declare SubobjsR_Base [code_pred_intro]
lemma SubobjsR_Rep_code [code_pred_intro]:
"\<lbrakk>subclsRp P C D; Subobjs\<^sub>R P D Cs\<rbrakk> \<Longrightarrow> Subobjs\<^sub>R P C (C # Cs)"
by(simp add: SubobjsR_Rep subclsR_def)
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool)
Subobjs\<^sub>R
by(erule Subobjs\<^sub>R.cases)(auto simp add: subclsR_code)
lemma subcls1p_code [code_pred_intro]:
"\<lbrakk>class P C = Some (Bs,rest); Predicate_Compile.contains (baseClasses Bs) D \<rbrakk> \<Longrightarrow> subcls1p P C D"
by(auto intro: subcls1p.intros simp add: Predicate_Compile.contains_def)
code_pred (modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool)
subcls1p
by(fastforce elim!: subcls1p.cases simp add: Predicate_Compile.contains_def)
declare Subobjs_Rep [code_pred_intro]
lemma Subobjs_Sh_code [code_pred_intro]:
"\<lbrakk> (subcls1p P)^** C C'; subclsSp P C' D; Subobjs\<^sub>R P D Cs\<rbrakk>
\<Longrightarrow> Subobjs P C Cs"
by(rule Subobjs_Sh)(simp_all add: rtrancl_def subcls1_def subclsS_def)
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool)
Subobjs
by(erule Subobjs.cases)(auto simp add: rtrancl_def subcls1_def subclsS_def)
definition widen_unique :: "prog \<Rightarrow> cname \<Rightarrow> cname \<Rightarrow> path \<Rightarrow> bool"
where "widen_unique P C D Cs \<longleftrightarrow> (\<forall>Cs'. Subobjs P C Cs' \<longrightarrow> last Cs' = D \<longrightarrow> Cs = Cs')"
code_pred [inductify, skip_proof] widen_unique .
lemma widen_subcls':
"\<lbrakk>Subobjs P C Cs'; last Cs' = D; widen_unique P C D Cs' \<rbrakk>
\<Longrightarrow> P \<turnstile> Class C \<le> Class D"
by(rule widen_subcls,auto simp:path_unique_def widen_unique_def)
declare
widen_refl [code_pred_intro]
widen_subcls' [code_pred_intro widen_subcls]
widen_null [code_pred_intro]
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool)
widen
by(erule widen.cases)(auto simp add: path_unique_def widen_unique_def)
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool)
leq_path1p
.
lemma leq_path_unfold: "P,C \<turnstile> Cs \<sqsubseteq> Ds \<longleftrightarrow> (leq_path1p P C)^** Cs Ds"
by(simp add: leq_path1_def rtrancl_def)
code_pred
(modes: i => i => i => o => bool, i => i => i => i => bool)
[inductify,skip_proof]
path_via
.
lemma path_unique_eq [code_pred_def]: "P \<turnstile> Path C to D unique \<longleftrightarrow>
(\<exists>Cs. Subobjs P C Cs \<and> last Cs = D \<and> (\<forall>Cs'. Subobjs P C Cs' \<longrightarrow> last Cs' = D \<longrightarrow> Cs = Cs'))"
by(auto simp add: path_unique_def)
code_pred
(modes: i => i => o => bool, i => i => i => bool)
[inductify, skip_proof]
path_unique .
text \<open>Redefine MethodDefs and FieldDecls\<close>
(* FIXME: These predicates should be defined inductively in the first place! *)
definition MethodDefs' :: "prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> path \<Rightarrow> method \<Rightarrow> bool" where
"MethodDefs' P C M Cs mthd \<equiv> (Cs, mthd) \<in> MethodDefs P C M"
lemma [code_pred_intro]:
"Subobjs P C Cs \<Longrightarrow> class P (last Cs) = \<lfloor>(Bs,fs,ms)\<rfloor> \<Longrightarrow> map_of ms M = \<lfloor>mthd\<rfloor> \<Longrightarrow>
MethodDefs' P C M Cs mthd"
by (simp add: MethodDefs_def MethodDefs'_def)
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool)
MethodDefs'
by(fastforce simp add: MethodDefs_def MethodDefs'_def)
definition FieldDecls' :: "prog \<Rightarrow> cname \<Rightarrow> vname \<Rightarrow> path \<Rightarrow> ty \<Rightarrow> bool" where
"FieldDecls' P C F Cs T \<equiv> (Cs, T) \<in> FieldDecls P C F"
lemma [code_pred_intro]:
"Subobjs P C Cs \<Longrightarrow> class P (last Cs) = \<lfloor>(Bs,fs,ms)\<rfloor> \<Longrightarrow> map_of fs F = \<lfloor>T\<rfloor> \<Longrightarrow>
FieldDecls' P C F Cs T"
by (simp add: FieldDecls_def FieldDecls'_def)
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool)
FieldDecls'
by(fastforce simp add: FieldDecls_def FieldDecls'_def)
definition MinimalMethodDefs' :: "prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> path \<Rightarrow> method \<Rightarrow> bool" where
"MinimalMethodDefs' P C M Cs mthd \<equiv> (Cs, mthd) \<in> MinimalMethodDefs P C M"
definition MinimalMethodDefs_unique :: "prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> path \<Rightarrow> bool"
where
"MinimalMethodDefs_unique P C M Cs \<longleftrightarrow>
(\<forall>Cs' mthd. MethodDefs' P C M Cs' mthd \<longrightarrow> (leq_path1p P C)^** Cs' Cs \<longrightarrow> Cs' = Cs)"
code_pred [inductify, skip_proof] MinimalMethodDefs_unique .
lemma [code_pred_intro]:
"MethodDefs' P C M Cs mthd \<Longrightarrow> MinimalMethodDefs_unique P C M Cs \<Longrightarrow>
MinimalMethodDefs' P C M Cs mthd"
by (fastforce simp add:MinimalMethodDefs_def MinimalMethodDefs'_def MethodDefs'_def MinimalMethodDefs_unique_def leq_path_unfold)
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool)
MinimalMethodDefs'
by(fastforce simp add:MinimalMethodDefs_def MinimalMethodDefs'_def MethodDefs'_def MinimalMethodDefs_unique_def leq_path_unfold)
definition LeastMethodDef_unique :: "prog \<Rightarrow> cname \<Rightarrow> mname \<Rightarrow> path \<Rightarrow> bool"
where
"LeastMethodDef_unique P C M Cs \<longleftrightarrow>
(\<forall>Cs' mthd'. MethodDefs' P C M Cs' mthd' \<longrightarrow> (leq_path1p P C)^** Cs Cs')"
code_pred [inductify, skip_proof] LeastMethodDef_unique .
lemma LeastMethodDef_unfold:
"P \<turnstile> C has least M = mthd via Cs \<longleftrightarrow>
MethodDefs' P C M Cs mthd \<and> LeastMethodDef_unique P C M Cs"
by(fastforce simp add: LeastMethodDef_def MethodDefs'_def leq_path_unfold LeastMethodDef_unique_def)
lemma LeastMethodDef_intro [code_pred_intro]:
"\<lbrakk> MethodDefs' P C M Cs mthd; LeastMethodDef_unique P C M Cs \<rbrakk>
\<Longrightarrow> P \<turnstile> C has least M = mthd via Cs"
by(simp add: LeastMethodDef_unfold LeastMethodDef_unique_def)
code_pred (modes: i => i => i => o => o => bool)
LeastMethodDef
by(simp add: LeastMethodDef_unfold LeastMethodDef_unique_def)
definition OverriderMethodDefs' :: "prog \<Rightarrow> subobj \<Rightarrow> mname \<Rightarrow> path \<Rightarrow> method \<Rightarrow> bool" where
"OverriderMethodDefs' P R M Cs mthd \<equiv> (Cs, mthd) \<in> OverriderMethodDefs P R M"
lemma Overrider1 [code_pred_intro]:
"P \<turnstile> (ldc R) has least M = mthd' via Cs' \<Longrightarrow>
MinimalMethodDefs' P (mdc R) M Cs mthd \<Longrightarrow>
last (snd R) = hd Cs' \<Longrightarrow> (leq_path1p P (mdc R))^** Cs (snd R @ tl Cs') \<Longrightarrow>
OverriderMethodDefs' P R M Cs mthd"
apply(simp add:OverriderMethodDefs_def OverriderMethodDefs'_def MinimalMethodDefs'_def appendPath_def leq_path_unfold)
apply(rule_tac x="Cs'" in exI)
apply clarsimp
apply(cases mthd')
apply blast
done
lemma Overrider2 [code_pred_intro]:
"P \<turnstile> (ldc R) has least M = mthd' via Cs' \<Longrightarrow>
MinimalMethodDefs' P (mdc R) M Cs mthd \<Longrightarrow>
last (snd R) \<noteq> hd Cs' \<Longrightarrow> (leq_path1p P (mdc R))^** Cs Cs' \<Longrightarrow>
OverriderMethodDefs' P R M Cs mthd"
by(auto simp add:OverriderMethodDefs_def OverriderMethodDefs'_def MinimalMethodDefs'_def appendPath_def leq_path_unfold simp del: split_paired_Ex)
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool)
OverriderMethodDefs'
apply(clarsimp simp add: OverriderMethodDefs'_def MinimalMethodDefs'_def MethodDefs'_def OverriderMethodDefs_def appendPath_def leq_path_unfold)
apply(case_tac "last xb = hd Cs'")
apply(simp)
apply(thin_tac "PROP _")
apply(simp add: leq_path1_def)
done
definition WTDynCast_ex :: "prog \<Rightarrow> cname \<Rightarrow> cname \<Rightarrow> bool"
where "WTDynCast_ex P D C \<longleftrightarrow> (\<exists>Cs. P \<turnstile> Path D to C via Cs)"
code_pred [inductify, skip_proof] WTDynCast_ex .
lemma WTDynCast_new:
"\<lbrakk>P,E \<turnstile> e :: Class D; is_class P C;
P \<turnstile> Path D to C unique \<or> \<not> WTDynCast_ex P D C\<rbrakk>
\<Longrightarrow> P,E \<turnstile> Cast C e :: Class C"
by(rule WTDynCast)(auto simp add: WTDynCast_ex_def)
definition WTStaticCast_sub :: "prog \<Rightarrow> cname \<Rightarrow> cname \<Rightarrow> bool"
where "WTStaticCast_sub P C D \<longleftrightarrow>
P \<turnstile> Path D to C unique \<or>
((subcls1p P)^** C D \<and> (\<forall>Cs. P \<turnstile> Path C to D via Cs \<longrightarrow> Subobjs\<^sub>R P C Cs))"
code_pred [inductify, skip_proof] WTStaticCast_sub .
lemma WTStaticCast_new:
"\<lbrakk>P,E \<turnstile> e :: Class D; is_class P C; WTStaticCast_sub P C D \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<lparr>C\<rparr>e :: Class C"
by (rule WTStaticCast)(auto simp add: WTStaticCast_sub_def subcls1_def rtrancl_def)
lemma WTBinOp1: "\<lbrakk> P,E \<turnstile> e\<^sub>1 :: T; P,E \<turnstile> e\<^sub>2 :: T\<rbrakk>
\<Longrightarrow> P,E \<turnstile> e\<^sub>1 \<guillemotleft>Eq\<guillemotright> e\<^sub>2 :: Boolean"
apply (rule WTBinOp)
apply assumption+
apply simp
done
lemma WTBinOp2: "\<lbrakk> P,E \<turnstile> e\<^sub>1 :: Integer; P,E \<turnstile> e\<^sub>2 :: Integer \<rbrakk>
\<Longrightarrow> P,E \<turnstile> e\<^sub>1 \<guillemotleft>Add\<guillemotright> e\<^sub>2 :: Integer"
apply (rule WTBinOp)
apply assumption+
apply simp
done
lemma LeastFieldDecl_unfold [code_pred_def]:
"P \<turnstile> C has least F:T via Cs \<longleftrightarrow>
FieldDecls' P C F Cs T \<and> (\<forall>Cs' T'. FieldDecls' P C F Cs' T' \<longrightarrow> (leq_path1p P C)^** Cs Cs')"
by(auto simp add: LeastFieldDecl_def FieldDecls'_def leq_path_unfold)
code_pred [inductify, skip_proof] LeastFieldDecl .
lemmas [code_pred_intro] = WT_WTs.WTNew
declare
WTDynCast_new[code_pred_intro WTDynCast_new]
WTStaticCast_new[code_pred_intro WTStaticCast_new]
lemmas [code_pred_intro] = WT_WTs.WTVal WT_WTs.WTVar
declare
WTBinOp1[code_pred_intro WTBinOp1]
WTBinOp2 [code_pred_intro WTBinOp2]
lemmas [code_pred_intro] =
WT_WTs.WTLAss WT_WTs.WTFAcc WT_WTs.WTFAss WT_WTs.WTCall WTStaticCall
WT_WTs.WTBlock WT_WTs.WTSeq WT_WTs.WTCond WT_WTs.WTWhile WT_WTs.WTThrow
lemmas [code_pred_intro] = WT_WTs.WTNil WT_WTs.WTCons
code_pred
(modes: WT: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool
and WTs: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool)
WT
proof -
case WT
from WT.prems show thesis
proof(cases (no_simp) rule: WT.cases)
case WTDynCast thus thesis
by(rule WT.WTDynCast_new[OF refl, unfolded WTDynCast_ex_def, simplified])
next
case WTStaticCast thus ?thesis
unfolding subcls1_def rtrancl_def mem_Collect_eq prod.case
by(rule WT.WTStaticCast_new[OF refl, unfolded WTStaticCast_sub_def])
next
case WTBinOp thus ?thesis
by(split bop.split_asm)(simp_all, (erule (4) WT.WTBinOp1[OF refl] WT.WTBinOp2[OF refl])+)
qed(assumption|erule (2) WT.that[OF refl])+
next
case WTs
from WTs.prems show thesis
by(cases (no_simp) rule: WTs.cases)(assumption|erule (2) WTs.that[OF refl])+
qed
lemma casts_to_code [code_pred_intro]:
"(case T of Class C \<Rightarrow> False | _ \<Rightarrow> True) \<Longrightarrow> P \<turnstile> T casts v to v"
"P \<turnstile> Class C casts Null to Null"
"\<lbrakk>Subobjs P (last Cs) Cs'; last Cs' = C;
last Cs = hd Cs'; Cs @ tl Cs' = Ds\<rbrakk>
\<Longrightarrow> P \<turnstile> Class C casts Ref(a,Cs) to Ref(a,Ds)"
"\<lbrakk>Subobjs P (last Cs) Cs'; last Cs' = C; last Cs \<noteq> hd Cs'\<rbrakk>
\<Longrightarrow> P \<turnstile> Class C casts Ref(a,Cs) to Ref(a,Cs')"
by(auto intro: casts_to.intros simp add: path_via_def appendPath_def)
code_pred (modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool)
casts_to
apply(erule casts_to.cases)
apply(fastforce split: ty.splits)
apply simp
apply(fastforce simp add: appendPath_def path_via_def split: if_split_asm)
done
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool)
Casts_to
.
lemma card_eq_1_iff_ex1: "x \<in> A \<Longrightarrow> card A = 1 \<longleftrightarrow> A = {x}"
apply(rule iffI)
apply(rule equalityI)
apply(rule subsetI)
apply(subgoal_tac "card {x, xa} \<le> card A")
apply(auto intro: ccontr)[1]
apply(rule card_mono)
apply simp_all
apply(metis Suc_n_not_n card_infinite)
done
lemma FinalOverriderMethodDef_unfold [code_pred_def]:
"P \<turnstile> R has overrider M = mthd via Cs \<longleftrightarrow>
OverriderMethodDefs' P R M Cs mthd \<and>
(\<forall>Cs' mthd'. OverriderMethodDefs' P R M Cs' mthd' \<longrightarrow> Cs = Cs' \<and> mthd = mthd')"
by(auto simp add: FinalOverriderMethodDef_def OverriderMethodDefs'_def card_eq_1_iff_ex1 simp del: One_nat_def)
code_pred
(modes: i => i => i => o => o => bool)
[inductify, skip_proof]
FinalOverriderMethodDef
.
code_pred
(modes: i => i => i => i => o => o => bool, i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> bool)
[inductify]
SelectMethodDef
.
text \<open>Isomorphic subo with mapping instead of a map\<close>
type_synonym subo' = "(path \<times> (vname, val) mapping)"
type_synonym obj' = "cname \<times> subo' set"
lift_definition init_class_fieldmap' :: "prog \<Rightarrow> cname \<Rightarrow> (vname, val) mapping" is "init_class_fieldmap" .
lemma init_class_fieldmap'_code [code]:
"init_class_fieldmap' P C =
Mapping (map (\<lambda>(F,T).(F,default_val T)) (fst(snd(the(class P C)))) )"
by transfer(simp add: init_class_fieldmap_def)
lift_definition init_obj' :: "prog \<Rightarrow> cname \<Rightarrow> subo' \<Rightarrow> bool" is init_obj .
lemma init_obj'_intros [code_pred_intro]:
"Subobjs P C Cs \<Longrightarrow> init_obj' P C (Cs, init_class_fieldmap' P (last Cs))"
by(transfer)(rule init_obj.intros)
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> bool as init_obj_pred)
init_obj'
by transfer(erule init_obj.cases, blast)
lemma init_obj_pred_conv: "set_of_pred (init_obj_pred P C) = Collect (init_obj' P C)"
by(auto elim: init_obj_predE intro: init_obj_predI)
lift_definition blank' :: "prog \<Rightarrow> cname \<Rightarrow> obj'" is "blank" .
lemma blank'_code [code]:
"blank' P C = (C, set_of_pred (init_obj_pred P C))"
unfolding init_obj_pred_conv by transfer(simp add: blank_def)
type_synonym heap' = "addr \<rightharpoonup> obj'"
abbreviation
cname_of' :: "heap' \<Rightarrow> addr \<Rightarrow> cname" where
"\<And>hp. cname_of' hp a == fst (the (hp a))"
lift_definition new_Addr' :: "heap' \<Rightarrow> addr option" is "new_Addr" .
lift_definition start_heap' :: "prog \<Rightarrow> heap'" is "start_heap" .
lemma start_heap'_code [code]:
"start_heap' P = Map.empty (addr_of_sys_xcpt NullPointer \<mapsto> blank' P NullPointer)
(addr_of_sys_xcpt ClassCast \<mapsto> blank' P ClassCast)
(addr_of_sys_xcpt OutOfMemory \<mapsto> blank' P OutOfMemory)"
by transfer(simp add: start_heap_def)
type_synonym
state' = "heap' \<times> locals"
lift_definition hp' :: "state' \<Rightarrow> heap'" is hp .
lemma hp'_code [code]: "hp' = fst"
by transfer simp
lift_definition lcl' :: "state' \<Rightarrow> locals" is lcl .
lemma lcl_code [code]: "lcl' = snd"
by transfer simp
lift_definition eval' :: "prog \<Rightarrow> env \<Rightarrow> expr \<Rightarrow> state' \<Rightarrow> expr \<Rightarrow> state' \<Rightarrow> bool"
("_,_ \<turnstile> ((1\<langle>_,/_\<rangle>) \<Rightarrow>''/ (1\<langle>_,/_\<rangle>))" [51,0,0,0,0] 81)
is eval .
lift_definition evals' :: "prog \<Rightarrow> env \<Rightarrow> expr list \<Rightarrow> state' \<Rightarrow> expr list \<Rightarrow> state' \<Rightarrow> bool"
("_,_ \<turnstile> ((1\<langle>_,/_\<rangle>) [\<Rightarrow>'']/ (1\<langle>_,/_\<rangle>))" [51,0,0,0,0] 81)
is evals .
lemma New':
"\<lbrakk> new_Addr' h = Some a; h' = h(a\<mapsto>(blank' P C)) \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>new C,(h,l)\<rangle> \<Rightarrow>' \<langle>ref (a,[C]),(h',l)\<rangle>"
by transfer(unfold blank_def, rule New)
lemma NewFail':
"new_Addr' h = None \<Longrightarrow>
P,E \<turnstile> \<langle>new C, (h,l)\<rangle> \<Rightarrow>' \<langle>THROW OutOfMemory,(h,l)\<rangle>"
by transfer(rule NewFail)
lemma StaticUpCast':
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a,Cs),s\<^sub>1\<rangle>; P \<turnstile> Path last Cs to C via Cs'; Ds = Cs@\<^sub>pCs' \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>\<lparr>C\<rparr>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a,Ds),s\<^sub>1\<rangle>"
by transfer(rule StaticUpCast)
lemma StaticDownCast'_new: (* requires reverse append *)
"\<lbrakk>P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a,Ds),s\<^sub>1\<rangle>; app Cs [C] Ds'; app Ds' Cs' Ds\<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>\<lparr>C\<rparr>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref(a,Cs@[C]),s\<^sub>1\<rangle>"
apply transfer
apply (rule StaticDownCast)
apply (simp add: app_eq)
done
lemma StaticCastNull':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>\<lparr>C\<rparr>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle>"
by transfer(rule StaticCastNull)
lemma StaticCastFail'_new: (* manual unfolding of subcls *)
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle>\<Rightarrow>' \<langle>ref (a,Cs),s\<^sub>1\<rangle>; \<not> (subcls1p P)^** (last Cs) C; C \<notin> set Cs\<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>\<lparr>C\<rparr>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>THROW ClassCast,s\<^sub>1\<rangle>"
apply transfer
by (fastforce intro:StaticCastFail simp add: rtrancl_def subcls1_def)
lemma StaticCastThrow':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>\<lparr>C\<rparr>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle>"
by transfer(rule StaticCastThrow)
lemma StaticUpDynCast':
"\<lbrakk>P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref(a,Cs),s\<^sub>1\<rangle>; P \<turnstile> Path last Cs to C unique;
P \<turnstile> Path last Cs to C via Cs'; Ds = Cs@\<^sub>pCs' \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref(a,Ds),s\<^sub>1\<rangle>"
by transfer(rule StaticUpDynCast)
lemma StaticDownDynCast'_new: (* requires reverse append *)
"\<lbrakk>P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a,Ds),s\<^sub>1\<rangle>; app Cs [C] Ds'; app Ds' Cs' Ds\<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref(a,Cs@[C]),s\<^sub>1\<rangle>"
apply transfer
apply (rule StaticDownDynCast)
apply (simp add: app_eq)
done
lemma DynCast':
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a,Cs),(h,l)\<rangle>; h a = Some(D,S);
P \<turnstile> Path D to C via Cs'; P \<turnstile> Path D to C unique \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a,Cs'),(h,l)\<rangle>"
by transfer(rule DynCast)
lemma DynCastNull':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle>"
by transfer(rule DynCastNull)
lemma DynCastFail':
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle>\<Rightarrow>' \<langle>ref (a,Cs),(h,l)\<rangle>; h a = Some(D,S); \<not> P \<turnstile> Path D to C unique;
\<not> P \<turnstile> Path last Cs to C unique; C \<notin> set Cs \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,(h,l)\<rangle>"
by transfer(rule DynCastFail)
lemma DynCastThrow':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>Cast C e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle>"
by transfer(rule DynCastThrow)
lemma Val':
"P,E \<turnstile> \<langle>Val v,s\<rangle> \<Rightarrow>' \<langle>Val v,s\<rangle>"
by transfer(rule Val)
lemma BinOp':
"\<lbrakk> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>Val v\<^sub>2,s\<^sub>2\<rangle>;
binop(bop,v\<^sub>1,v\<^sub>2) = Some v \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0\<rangle>\<Rightarrow>'\<langle>Val v,s\<^sub>2\<rangle>"
by transfer(rule BinOp)
lemma BinOpThrow1':
"P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2, s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e,s\<^sub>1\<rangle>"
by transfer(rule BinOpThrow1)
lemma BinOpThrow2':
"\<lbrakk> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val v\<^sub>1,s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>throw e,s\<^sub>2\<rangle> \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e,s\<^sub>2\<rangle>"
by transfer(rule BinOpThrow2)
lemma Var':
"l V = Some v \<Longrightarrow>
P,E \<turnstile> \<langle>Var V,(h,l)\<rangle> \<Rightarrow>' \<langle>Val v,(h,l)\<rangle>"
by transfer(rule Var)
lemma LAss':
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val v,(h,l)\<rangle>; E V = Some T;
P \<turnstile> T casts v to v'; l' = l(V\<mapsto>v') \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>V:=e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val v',(h,l')\<rangle>"
by (transfer) (erule (3) LAss)
lemma LAssThrow':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>V:=e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle>"
by transfer(rule LAssThrow)
lemma FAcc'_new: (* iteration over set *)
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a,Cs'),(h,l)\<rangle>; h a = Some(D,S);
Ds = Cs'@\<^sub>pCs; Predicate_Compile.contains (Set_project S Ds) fs; Mapping.lookup fs F = Some v \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>e\<bullet>F{Cs},s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val v,(h,l)\<rangle>"
unfolding Set_project_def mem_Collect_eq Predicate_Compile.contains_def
by transfer(rule FAcc)
lemma FAccNull':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>e\<bullet>F{Cs},s\<^sub>0\<rangle> \<Rightarrow>' \<langle>THROW NullPointer,s\<^sub>1\<rangle>"
by transfer(rule FAccNull)
lemma FAccThrow':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>e\<bullet>F{Cs},s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle>"
by transfer(rule FAccThrow)
lemma FAss'_new: (* iteration over set *)
"\<lbrakk> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a,Cs'),s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>Val v,(h\<^sub>2,l\<^sub>2)\<rangle>;
h\<^sub>2 a = Some(D,S); P \<turnstile> (last Cs') has least F:T via Cs; P \<turnstile> T casts v to v';
Ds = Cs'@\<^sub>pCs; Predicate_Compile.contains (Set_project S Ds) fs; fs' = Mapping.update F v' fs;
S' = S - {(Ds,fs)} \<union> {(Ds,fs')}; h\<^sub>2' = h\<^sub>2(a\<mapsto>(D,S'))\<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>1\<bullet>F{Cs}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val v',(h\<^sub>2',l\<^sub>2)\<rangle>"
unfolding Predicate_Compile.contains_def Set_project_def mem_Collect_eq
by transfer(rule FAss)
lemma FAssNull':
"\<lbrakk> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>Val v,s\<^sub>2\<rangle> \<rbrakk> \<Longrightarrow>
P,E \<turnstile> \<langle>e\<^sub>1\<bullet>F{Cs}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>THROW NullPointer,s\<^sub>2\<rangle>"
by transfer(rule FAssNull)
lemma FAssThrow1':
"P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>e\<^sub>1\<bullet>F{Cs}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle>"
by transfer(rule FAssThrow1)
lemma FAssThrow2':
"\<lbrakk> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val v,s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>2\<rangle> \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>1\<bullet>F{Cs}:=e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>2\<rangle>"
by transfer(rule FAssThrow2)
lemma CallObjThrow':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>Call e Copt M es,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle>"
by transfer(rule CallObjThrow)
lemma CallParamsThrow'_new: (* requires inverse map Val and append *)
"\<lbrakk> P,E \<turnstile> \<langle>e,s0\<rangle> \<Rightarrow>' \<langle>Val v,s1\<rangle>; P,E \<turnstile> \<langle>es,s1\<rangle> [\<Rightarrow>'] \<langle>evs,s2\<rangle>;
map_val2 evs vs (throw ex # es') \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>Call e Copt M es,s0\<rangle> \<Rightarrow>' \<langle>throw ex,s2\<rangle>"
apply transfer
apply(rule eval_evals.CallParamsThrow, assumption+)
apply(simp add: map_val2_conv[symmetric])
done
lemma Call'_new: (* requires inverse map Val *)
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a,Cs),s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>'] \<langle>evs,(h\<^sub>2,l\<^sub>2)\<rangle>;
map_val evs vs;
h\<^sub>2 a = Some(C,S); P \<turnstile> last Cs has least M = (Ts',T',pns',body') via Ds;
P \<turnstile> (C,Cs@\<^sub>pDs) selects M = (Ts,T,pns,body) via Cs'; length vs = length pns;
P \<turnstile> Ts Casts vs to vs'; l\<^sub>2' = [this\<mapsto>Ref (a,Cs'), pns[\<mapsto>]vs'];
new_body = (case T' of Class D \<Rightarrow> \<lparr>D\<rparr>body | _ \<Rightarrow> body);
P,E(this\<mapsto>Class(last Cs'), pns[\<mapsto>]Ts) \<turnstile> \<langle>new_body,(h\<^sub>2,l\<^sub>2')\<rangle> \<Rightarrow>' \<langle>e',(h\<^sub>3,l\<^sub>3)\<rangle> \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow>' \<langle>e',(h\<^sub>3,l\<^sub>2)\<rangle>"
apply transfer
apply(rule Call)
apply assumption+
apply(simp add: map_val_conv[symmetric])
apply assumption+
done
lemma StaticCall'_new: (* requires inverse map Val *)
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a,Cs),s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>'] \<langle>evs,(h\<^sub>2,l\<^sub>2)\<rangle>;
map_val evs vs;
P \<turnstile> Path (last Cs) to C unique; P \<turnstile> Path (last Cs) to C via Cs'';
P \<turnstile> C has least M = (Ts,T,pns,body) via Cs'; Ds = (Cs@\<^sub>pCs'')@\<^sub>pCs';
length vs = length pns; P \<turnstile> Ts Casts vs to vs';
l\<^sub>2' = [this\<mapsto>Ref (a,Ds), pns[\<mapsto>]vs'];
P,E(this\<mapsto>Class(last Ds), pns[\<mapsto>]Ts) \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2')\<rangle> \<Rightarrow>' \<langle>e',(h\<^sub>3,l\<^sub>3)\<rangle> \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>e\<bullet>(C::)M(ps),s\<^sub>0\<rangle> \<Rightarrow>' \<langle>e',(h\<^sub>3,l\<^sub>2)\<rangle>"
apply transfer
apply(rule StaticCall)
apply(assumption)+
apply(simp add: map_val_conv[symmetric])
apply assumption+
done
lemma CallNull'_new: (* requires inverse map Val *)
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>'] \<langle>evs,s\<^sub>2\<rangle>; map_val evs vs \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>Call e Copt M es,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>THROW NullPointer,s\<^sub>2\<rangle>"
apply transfer
apply(rule CallNull, assumption+)
apply(simp add: map_val_conv[symmetric])
done
lemma Block':
"\<lbrakk>P,E(V \<mapsto> T) \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0,l\<^sub>0(V:=None))\<rangle> \<Rightarrow>' \<langle>e\<^sub>1,(h\<^sub>1,l\<^sub>1)\<rangle> \<rbrakk> \<Longrightarrow>
P,E \<turnstile> \<langle>{V:T; e\<^sub>0},(h\<^sub>0,l\<^sub>0)\<rangle> \<Rightarrow>' \<langle>e\<^sub>1,(h\<^sub>1,l\<^sub>1(V:=l\<^sub>0 V))\<rangle>"
by transfer(rule Block)
lemma Seq':
"\<lbrakk> P,E \<turnstile> \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val v,s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>e\<^sub>2,s\<^sub>2\<rangle> \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>e\<^sub>2,s\<^sub>2\<rangle>"
by transfer(rule Seq)
lemma SeqThrow':
"P,E \<turnstile> \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>e\<^sub>0;;e\<^sub>1,s\<^sub>0\<rangle>\<Rightarrow>'\<langle>throw e,s\<^sub>1\<rangle>"
by transfer(rule SeqThrow)
lemma CondT':
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>true,s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>e',s\<^sub>2\<rangle> \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>e',s\<^sub>2\<rangle>"
by transfer(rule CondT)
lemma CondF':
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>false,s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>e',s\<^sub>2\<rangle> \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>e',s\<^sub>2\<rangle>"
by transfer(rule CondF)
lemma CondThrow':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>if (e) e\<^sub>1 else e\<^sub>2, s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle>"
by transfer(rule CondThrow)
lemma WhileF':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>false,s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>unit,s\<^sub>1\<rangle>"
by transfer(rule WhileF)
lemma WhileT':
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>true,s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>Val v\<^sub>1,s\<^sub>2\<rangle>;
P,E \<turnstile> \<langle>while (e) c,s\<^sub>2\<rangle> \<Rightarrow>' \<langle>e\<^sub>3,s\<^sub>3\<rangle> \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>e\<^sub>3,s\<^sub>3\<rangle>"
by transfer(rule WhileT)
lemma WhileCondThrow':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle> throw e',s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle>"
by transfer(rule WhileCondThrow)
lemma WhileBodyThrow':
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>true,s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>2\<rangle>\<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>while (e) c,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>2\<rangle>"
by transfer(rule WhileBodyThrow)
lemma Throw':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref r,s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Throw r,s\<^sub>1\<rangle>"
by transfer(rule eval_evals.Throw)
lemma ThrowNull':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>THROW NullPointer,s\<^sub>1\<rangle>"
by transfer(rule ThrowNull)
lemma ThrowThrow':
"P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>throw e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle>"
by transfer(rule ThrowThrow)
lemma Nil':
"P,E \<turnstile> \<langle>[],s\<rangle> [\<Rightarrow>'] \<langle>[],s\<rangle>"
by transfer(rule eval_evals.Nil)
lemma Cons':
"\<lbrakk> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val v,s\<^sub>1\<rangle>; P,E \<turnstile> \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>'] \<langle>es',s\<^sub>2\<rangle> \<rbrakk>
\<Longrightarrow> P,E \<turnstile> \<langle>e#es,s\<^sub>0\<rangle> [\<Rightarrow>'] \<langle>Val v # es',s\<^sub>2\<rangle>"
by transfer(rule eval_evals.Cons)
lemma ConsThrow':
"P,E \<turnstile> \<langle>e, s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e', s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>e#es, s\<^sub>0\<rangle> [\<Rightarrow>'] \<langle>throw e' # es, s\<^sub>1\<rangle>"
by transfer(rule ConsThrow)
text \<open>Axiomatic heap address model refinement\<close>
partial_function (option) lowest :: "(nat \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> nat option"
where
[code]: "lowest P n = (if P n then Some n else lowest P (Suc n))"
axiomatization
where
new_Addr'_code [code]: "new_Addr' h = lowest (Option.is_none \<circ> h) 0"
\<comment> \<open>admissible: a tightening of the specification of @{const new_Addr'}\<close>
lemma eval'_cases
[consumes 1,
case_names New NewFail StaticUpCast StaticDownCast StaticCastNull StaticCastFail
StaticCastThrow StaticUpDynCast StaticDownDynCast DynCast DynCastNull DynCastFail
DynCastThrow Val BinOp BinOpThrow1 BinOpThrow2 Var LAss LAssThrow FAcc FAccNull FAccThrow
FAss FAssNull FAssThrow1 FAssThrow2 CallObjThrow CallParamsThrow Call StaticCall CallNull
Block Seq SeqThrow CondT CondF CondThrow WhileF WhileT WhileCondThrow WhileBodyThrow
Throw ThrowNull ThrowThrow]:
assumes "P,x \<turnstile> \<langle>y,z\<rangle> \<Rightarrow>' \<langle>u,v\<rangle>"
and "\<And>h a h' C E l. x = E \<Longrightarrow> y = new C \<Longrightarrow> z = (h, l) \<Longrightarrow> u = ref (a, [C]) \<Longrightarrow>
v = (h', l) \<Longrightarrow> new_Addr' h = \<lfloor>a\<rfloor> \<Longrightarrow> h' = h(a \<mapsto> blank' P C) \<Longrightarrow> thesis"
and "\<And>h E C l. x = E \<Longrightarrow> y = new C \<Longrightarrow> z = (h, l) \<Longrightarrow>
u = Throw (addr_of_sys_xcpt OutOfMemory, [OutOfMemory]) \<Longrightarrow>
v = (h, l) \<Longrightarrow> new_Addr' h = None \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 a Cs s\<^sub>1 C Cs' Ds. x = E \<Longrightarrow> y = \<lparr>C\<rparr>e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = ref (a, Ds) \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a, Cs),s\<^sub>1\<rangle> \<Longrightarrow>
P \<turnstile> Path last Cs to C via Cs' \<Longrightarrow> Ds = Cs @\<^sub>p Cs' \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 a Cs C Cs' s\<^sub>1. x = E \<Longrightarrow> y = \<lparr>C\<rparr>e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = ref (a, Cs @ [C]) \<Longrightarrow>
v = s\<^sub>1 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a, Cs @ [C] @ Cs'),s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 s\<^sub>1 C. x = E \<Longrightarrow> y = \<lparr>C\<rparr>e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = null \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow>
P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 a Cs s\<^sub>1 C. x = E \<Longrightarrow> y = \<lparr>C\<rparr>e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = Throw (addr_of_sys_xcpt ClassCast, [ClassCast]) \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow>
P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a, Cs),s\<^sub>1\<rangle> \<Longrightarrow> (last Cs, C) \<notin> (subcls1 P)\<^sup>* \<Longrightarrow> C \<notin> set Cs \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 e' s\<^sub>1 C. x = E \<Longrightarrow> y = \<lparr>C\<rparr>e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = throw e' \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow>
P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 a Cs s\<^sub>1 C Cs' Ds. x = E \<Longrightarrow> y = Cast C e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = ref (a, Ds) \<Longrightarrow>
v = s\<^sub>1 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a, Cs),s\<^sub>1\<rangle> \<Longrightarrow> P \<turnstile> Path last Cs to C unique \<Longrightarrow>
P \<turnstile> Path last Cs to C via Cs' \<Longrightarrow> Ds = Cs @\<^sub>p Cs' \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 a Cs C Cs' s\<^sub>1. x = E \<Longrightarrow> y = Cast C e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = ref (a, Cs @ [C]) \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a, Cs @ [C] @ Cs'),s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 a Cs h l D S C Cs'. x = E \<Longrightarrow> y = Cast C e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = ref (a, Cs') \<Longrightarrow> v = (h, l) \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a, Cs),(h, l)\<rangle> \<Longrightarrow>
h a = \<lfloor>(D, S)\<rfloor> \<Longrightarrow> P \<turnstile> Path D to C via Cs' \<Longrightarrow> P \<turnstile> Path D to C unique \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 s\<^sub>1 C. x = E \<Longrightarrow> y = Cast C e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = null \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow>
P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 a Cs h l D S C. x = E \<Longrightarrow> y = Cast C e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = null \<Longrightarrow>
v = (h, l) \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a, Cs),(h, l)\<rangle> \<Longrightarrow> h a = \<lfloor>(D, S)\<rfloor> \<Longrightarrow>
\<not> P \<turnstile> Path D to C unique \<Longrightarrow> \<not> P \<turnstile> Path last Cs to C unique \<Longrightarrow> C \<notin> set Cs \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 e' s\<^sub>1 C. x = E \<Longrightarrow> y = Cast C e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = throw e' \<Longrightarrow> v = s\<^sub>1
\<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E va s. x = E \<Longrightarrow> y = Val va \<Longrightarrow> z = s \<Longrightarrow> u = Val va \<Longrightarrow> v = s \<Longrightarrow> thesis"
and "\<And>E e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 v\<^sub>2 s\<^sub>2 bop va. x = E \<Longrightarrow> y = e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = Val va \<Longrightarrow> v = s\<^sub>2 \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val v\<^sub>1,s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>Val v\<^sub>2,s\<^sub>2\<rangle> \<Longrightarrow> binop (bop, v\<^sub>1, v\<^sub>2) = \<lfloor>va\<rfloor> \<Longrightarrow> thesis"
and "\<And>E e\<^sub>1 s\<^sub>0 e s\<^sub>1 bop e\<^sub>2. x = E \<Longrightarrow> y = e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = throw e \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow>
P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e\<^sub>1 s\<^sub>0 v\<^sub>1 s\<^sub>1 e\<^sub>2 e s\<^sub>2 bop. x = E \<Longrightarrow> y = e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = throw e \<Longrightarrow>
v = s\<^sub>2 \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val v\<^sub>1,s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>throw e,s\<^sub>2\<rangle> \<Longrightarrow> thesis"
and "\<And>l V va E h. x = E \<Longrightarrow> y = Var V \<Longrightarrow> z = (h, l) \<Longrightarrow> u = Val va \<Longrightarrow> v = (h, l) \<Longrightarrow>
l V = \<lfloor>va\<rfloor> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 va h l V T v' l'. x = E \<Longrightarrow> y = V:=e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = Val v' \<Longrightarrow>
v = (h, l') \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val va,(h, l)\<rangle> \<Longrightarrow>
E V = \<lfloor>T\<rfloor> \<Longrightarrow> P \<turnstile> T casts va to v' \<Longrightarrow> l' = l(V \<mapsto> v') \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 e' s\<^sub>1 V. x = E \<Longrightarrow> y = V:=e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = throw e' \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow>
P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 a Cs' h l D S Ds Cs fs F va. x = E \<Longrightarrow> y = e\<bullet>F{Cs} \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = Val va \<Longrightarrow> v = (h, l) \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a, Cs'),(h, l)\<rangle> \<Longrightarrow>
h a = \<lfloor>(D, S)\<rfloor> \<Longrightarrow> Ds = Cs' @\<^sub>p Cs \<Longrightarrow> (Ds, fs) \<in> S \<Longrightarrow> Mapping.lookup fs F = \<lfloor>va\<rfloor> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 s\<^sub>1 F Cs. x = E \<Longrightarrow> y = e\<bullet>F{Cs} \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = Throw (addr_of_sys_xcpt NullPointer, [NullPointer]) \<Longrightarrow>
v = s\<^sub>1 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 e' s\<^sub>1 F Cs. x = E \<Longrightarrow> y = e\<bullet>F{Cs} \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = throw e' \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow>
P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e\<^sub>1 s\<^sub>0 a Cs' s\<^sub>1 e\<^sub>2 va h\<^sub>2 l\<^sub>2 D S F T Cs v' Ds fs fs' S' h\<^sub>2'.
x = E \<Longrightarrow> y = e\<^sub>1\<bullet>F{Cs} := e\<^sub>2 \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = Val v' \<Longrightarrow> v = (h\<^sub>2', l\<^sub>2) \<Longrightarrow>
P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a, Cs'),s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>Val va,(h\<^sub>2, l\<^sub>2)\<rangle> \<Longrightarrow>
h\<^sub>2 a = \<lfloor>(D, S)\<rfloor> \<Longrightarrow> P \<turnstile> last Cs' has least F:T via Cs \<Longrightarrow>
P \<turnstile> T casts va to v' \<Longrightarrow> Ds = Cs' @\<^sub>p Cs \<Longrightarrow> (Ds, fs) \<in> S \<Longrightarrow> fs' = Mapping.update F v' fs \<Longrightarrow>
S' = S - {(Ds, fs)} \<union> {(Ds, fs')} \<Longrightarrow> h\<^sub>2' = h\<^sub>2(a \<mapsto> (D, S')) \<Longrightarrow> thesis"
and "\<And>E e\<^sub>1 s\<^sub>0 s\<^sub>1 e\<^sub>2 va s\<^sub>2 F Cs. x = E \<Longrightarrow> y = e\<^sub>1\<bullet>F{Cs} := e\<^sub>2 \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = Throw (addr_of_sys_xcpt NullPointer, [NullPointer]) \<Longrightarrow>
v = s\<^sub>2 \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>Val va,s\<^sub>2\<rangle> \<Longrightarrow> thesis"
and "\<And>E e\<^sub>1 s\<^sub>0 e' s\<^sub>1 F Cs e\<^sub>2. x = E \<Longrightarrow> y = e\<^sub>1\<bullet>F{Cs} := e\<^sub>2 \<Longrightarrow>
z = s\<^sub>0 \<Longrightarrow> u = throw e' \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e\<^sub>1 s\<^sub>0 va s\<^sub>1 e\<^sub>2 e' s\<^sub>2 F Cs. x = E \<Longrightarrow> y = e\<^sub>1\<bullet>F{Cs} := e\<^sub>2 \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = throw e' \<Longrightarrow> v = s\<^sub>2 \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val va,s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>2\<rangle> \<Longrightarrow>
thesis"
and "\<And>E e s\<^sub>0 e' s\<^sub>1 Copt M es. x = E \<Longrightarrow> y = Call e Copt M es \<Longrightarrow>
z = s\<^sub>0 \<Longrightarrow> u = throw e' \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 va s\<^sub>1 es vs ex es' s\<^sub>2 Copt M. x = E \<Longrightarrow> y = Call e Copt M es \<Longrightarrow>
z = s\<^sub>0 \<Longrightarrow> u = throw ex \<Longrightarrow> v = s\<^sub>2 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val va,s\<^sub>1\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>'] \<langle>map Val vs @ throw ex # es',s\<^sub>2\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 a Cs s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C S M Ts' T' pns' body' Ds Ts T pns body Cs' vs' l\<^sub>2' new_body e'
h\<^sub>3 l\<^sub>3. x = E \<Longrightarrow> y = Call e None M ps \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = e' \<Longrightarrow> v = (h\<^sub>3, l\<^sub>2) \<Longrightarrow>
P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a, Cs),s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>'] \<langle>map Val vs,(h\<^sub>2, l\<^sub>2)\<rangle> \<Longrightarrow>
h\<^sub>2 a = \<lfloor>(C, S)\<rfloor> \<Longrightarrow> P \<turnstile> last Cs has least M = (Ts', T', pns', body') via Ds \<Longrightarrow>
P \<turnstile> (C,Cs @\<^sub>p Ds) selects M = (Ts, T, pns, body) via Cs' \<Longrightarrow> length vs = length pns \<Longrightarrow>
P \<turnstile> Ts Casts vs to vs' \<Longrightarrow> l\<^sub>2' = [this \<mapsto> Ref (a, Cs'), pns [\<mapsto>] vs'] \<Longrightarrow>
new_body = (case T' of Class D \<Rightarrow> \<lparr>D\<rparr>body | _ \<Rightarrow> body) \<Longrightarrow>
P,E(this \<mapsto> Class (last Cs'), pns [\<mapsto>] Ts) \<turnstile> \<langle>new_body,(h\<^sub>2, l\<^sub>2')\<rangle> \<Rightarrow>' \<langle>e',(h\<^sub>3, l\<^sub>3)\<rangle> \<Longrightarrow>
thesis"
and "\<And>E e s\<^sub>0 a Cs s\<^sub>1 ps vs h\<^sub>2 l\<^sub>2 C Cs'' M Ts T pns body Cs' Ds vs' l\<^sub>2' e' h\<^sub>3 l\<^sub>3.
x = E \<Longrightarrow> y = Call e \<lfloor>C\<rfloor> M ps \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = e' \<Longrightarrow> v = (h\<^sub>3, l\<^sub>2) \<Longrightarrow>
P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref (a, Cs),s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>'] \<langle>map Val vs,(h\<^sub>2, l\<^sub>2)\<rangle> \<Longrightarrow>
P \<turnstile> Path last Cs to C unique \<Longrightarrow> P \<turnstile> Path last Cs to C via Cs'' \<Longrightarrow>
P \<turnstile> C has least M = (Ts, T, pns, body) via Cs' \<Longrightarrow> Ds = (Cs @\<^sub>p Cs'') @\<^sub>p Cs' \<Longrightarrow>
length vs = length pns \<Longrightarrow> P \<turnstile> Ts Casts vs to vs' \<Longrightarrow>
l\<^sub>2' = [this \<mapsto> Ref (a, Ds), pns [\<mapsto>] vs'] \<Longrightarrow>
P,E(this \<mapsto> Class (last Ds), pns [\<mapsto>] Ts) \<turnstile> \<langle>body,(h\<^sub>2, l\<^sub>2')\<rangle> \<Rightarrow>' \<langle>e',(h\<^sub>3, l\<^sub>3)\<rangle> \<Longrightarrow>
thesis"
and "\<And>E e s\<^sub>0 s\<^sub>1 es vs s\<^sub>2 Copt M. x = E \<Longrightarrow> y = Call e Copt M es \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = Throw (addr_of_sys_xcpt NullPointer, [NullPointer]) \<Longrightarrow>
v = s\<^sub>2 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>'] \<langle>map Val vs,s\<^sub>2\<rangle> \<Longrightarrow> thesis"
and "\<And>E V T e\<^sub>0 h\<^sub>0 l\<^sub>0 e\<^sub>1 h\<^sub>1 l\<^sub>1.
x = E \<Longrightarrow> y = {V:T; e\<^sub>0} \<Longrightarrow> z = (h\<^sub>0, l\<^sub>0) \<Longrightarrow> u = e\<^sub>1 \<Longrightarrow>
v = (h\<^sub>1, l\<^sub>1(V := l\<^sub>0 V)) \<Longrightarrow> P,E(V \<mapsto> T) \<turnstile> \<langle>e\<^sub>0,(h\<^sub>0, l\<^sub>0(V := None))\<rangle> \<Rightarrow>' \<langle>e\<^sub>1,(h\<^sub>1, l\<^sub>1)\<rangle> \<Longrightarrow> thesis"
and "\<And>E e\<^sub>0 s\<^sub>0 va s\<^sub>1 e\<^sub>1 e\<^sub>2 s\<^sub>2. x = E \<Longrightarrow> y = e\<^sub>0;; e\<^sub>1 \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = e\<^sub>2 \<Longrightarrow>
v = s\<^sub>2 \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>Val va,s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>e\<^sub>2,s\<^sub>2\<rangle> \<Longrightarrow> thesis"
and "\<And>E e\<^sub>0 s\<^sub>0 e s\<^sub>1 e\<^sub>1. x = E \<Longrightarrow> y = e\<^sub>0;; e\<^sub>1 \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = throw e \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow>
P,E \<turnstile> \<langle>e\<^sub>0,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e,s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 s\<^sub>1 e\<^sub>1 e' s\<^sub>2 e\<^sub>2. x = E \<Longrightarrow> y = if (e) e\<^sub>1 else e\<^sub>2 \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = e' \<Longrightarrow>
v = s\<^sub>2 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>true,s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>1,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>e',s\<^sub>2\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 s\<^sub>1 e\<^sub>2 e' s\<^sub>2 e\<^sub>1. x = E \<Longrightarrow> y = if (e) e\<^sub>1 else e\<^sub>2 \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = e' \<Longrightarrow> v = s\<^sub>2 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>false,s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>e\<^sub>2,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>e',s\<^sub>2\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 e' s\<^sub>1 e\<^sub>1 e\<^sub>2. x = E \<Longrightarrow> y = if (e) e\<^sub>1 else e\<^sub>2 \<Longrightarrow>
z = s\<^sub>0 \<Longrightarrow> u = throw e' \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 s\<^sub>1 c. x = E \<Longrightarrow> y = while (e) c \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = unit \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow>
P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>false,s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 s\<^sub>1 c v\<^sub>1 s\<^sub>2 e\<^sub>3 s\<^sub>3. x = E \<Longrightarrow> y = while (e) c \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = e\<^sub>3 \<Longrightarrow>
v = s\<^sub>3 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>true,s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>Val v\<^sub>1,s\<^sub>2\<rangle> \<Longrightarrow>
P,E \<turnstile> \<langle>while (e) c,s\<^sub>2\<rangle> \<Rightarrow>' \<langle>e\<^sub>3,s\<^sub>3\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 e' s\<^sub>1 c. x = E \<Longrightarrow> y = while (e) c \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = throw e' \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow>
P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 s\<^sub>1 c e' s\<^sub>2. x = E \<Longrightarrow> y = while (e) c \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow> u = throw e' \<Longrightarrow>
v = s\<^sub>2 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>true,s\<^sub>1\<rangle> \<Longrightarrow> P,E \<turnstile> \<langle>c,s\<^sub>1\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>2\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 r s\<^sub>1. x = E \<Longrightarrow> y = throw e \<Longrightarrow>
z = s\<^sub>0 \<Longrightarrow> u = Throw r \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>ref r,s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 s\<^sub>1. x = E \<Longrightarrow> y = throw e \<Longrightarrow> z = s\<^sub>0 \<Longrightarrow>
u = Throw (addr_of_sys_xcpt NullPointer, [NullPointer]) \<Longrightarrow>
v = s\<^sub>1 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>null,s\<^sub>1\<rangle> \<Longrightarrow> thesis"
and "\<And>E e s\<^sub>0 e' s\<^sub>1. x = E \<Longrightarrow> y = throw e \<Longrightarrow>
z = s\<^sub>0 \<Longrightarrow> u = throw e' \<Longrightarrow> v = s\<^sub>1 \<Longrightarrow> P,E \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow>' \<langle>throw e',s\<^sub>1\<rangle> \<Longrightarrow> thesis"
shows thesis
using assms
by(transfer)(erule eval.cases, unfold blank_def, assumption+)
lemmas [code_pred_intro] = New' NewFail' StaticUpCast'
declare StaticDownCast'_new[code_pred_intro StaticDownCast']
lemmas [code_pred_intro] = StaticCastNull'
declare StaticCastFail'_new[code_pred_intro StaticCastFail']
lemmas [code_pred_intro] = StaticCastThrow' StaticUpDynCast'
declare
StaticDownDynCast'_new[code_pred_intro StaticDownDynCast']
DynCast'[code_pred_intro DynCast']
lemmas [code_pred_intro] = DynCastNull'
declare DynCastFail'[code_pred_intro DynCastFail']
lemmas [code_pred_intro] = DynCastThrow' Val' BinOp' BinOpThrow1'
declare BinOpThrow2'[code_pred_intro BinOpThrow2']
lemmas [code_pred_intro] = Var' LAss' LAssThrow'
declare FAcc'_new[code_pred_intro FAcc']
lemmas [code_pred_intro] = FAccNull' FAccThrow'
declare FAss'_new[code_pred_intro FAss']
lemmas [code_pred_intro] = FAssNull' FAssThrow1'
declare FAssThrow2'[code_pred_intro FAssThrow2']
lemmas [code_pred_intro] = CallObjThrow'
declare
CallParamsThrow'_new[code_pred_intro CallParamsThrow']
Call'_new[code_pred_intro Call']
StaticCall'_new[code_pred_intro StaticCall']
CallNull'_new[code_pred_intro CallNull']
lemmas [code_pred_intro] = Block' Seq'
declare SeqThrow'[code_pred_intro SeqThrow']
lemmas [code_pred_intro] = CondT'
declare
CondF'[code_pred_intro CondF']
CondThrow'[code_pred_intro CondThrow']
lemmas [code_pred_intro] = WhileF' WhileT'
declare
WhileCondThrow'[code_pred_intro WhileCondThrow']
WhileBodyThrow'[code_pred_intro WhileBodyThrow']
lemmas [code_pred_intro] = Throw'
declare ThrowNull'[code_pred_intro ThrowNull']
lemmas [code_pred_intro] = ThrowThrow'
lemmas [code_pred_intro] = Nil' Cons' ConsThrow'
code_pred
(modes: eval': i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool as big_step
and evals': i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool as big_steps)
eval'
proof -
case eval'
from eval'.prems show thesis
proof(cases (no_simp) rule: eval'_cases)
case (StaticDownCast E C e s\<^sub>0 a Cs Cs' s\<^sub>1)
moreover
have "app a [Cs] (a @ [Cs])" "app (a @ [Cs]) Cs' (a @ [Cs] @ Cs')"
by(simp_all add: app_eq)
ultimately show ?thesis by(rule eval'.StaticDownCast'[OF refl])
next
case StaticCastFail thus ?thesis
unfolding rtrancl_def subcls1_def mem_Collect_eq prod.case
by(rule eval'.StaticCastFail'[OF refl])
next
case (StaticDownDynCast E e s\<^sub>0 a Cs C Cs' s\<^sub>1)
moreover have "app Cs [C] (Cs @ [C])" "app (Cs @ [C]) Cs' (Cs @ [C] @ Cs')"
by(simp_all add: app_eq)
ultimately show thesis by(rule eval'.StaticDownDynCast'[OF refl])
next
case DynCast thus ?thesis by(rule eval'.DynCast'[OF refl])
next
case DynCastFail thus ?thesis by(rule eval'.DynCastFail'[OF refl])
next
case BinOpThrow2 thus ?thesis by(rule eval'.BinOpThrow2'[OF refl])
next
case FAcc thus ?thesis
by(rule eval'.FAcc'[OF refl, unfolded Predicate_Compile.contains_def Set_project_def mem_Collect_eq])
next
case FAss thus ?thesis
by(rule eval'.FAss'[OF refl, unfolded Predicate_Compile.contains_def Set_project_def mem_Collect_eq])
next
case FAssThrow2 thus ?thesis by(rule eval'.FAssThrow2'[OF refl])
next
case (CallParamsThrow E e s\<^sub>0 v s\<^sub>1 es vs ex es' s\<^sub>2 Copt M)
moreover have "map_val2 (map Val vs @ throw ex # es') vs (throw ex # es')"
by(simp add: map_val2_conv[symmetric])
ultimately show ?thesis by(rule eval'.CallParamsThrow'[OF refl])
next
case (Call E e s\<^sub>0 a Cs s\<^sub>1 ps vs)
moreover have "map_val (map Val vs) vs" by(simp add: map_val_conv[symmetric])
ultimately show ?thesis by-(rule eval'.Call'[OF refl])
next
case (StaticCall E e s\<^sub>0 a Cs s\<^sub>1 ps vs)
moreover have "map_val (map Val vs) vs" by(simp add: map_val_conv[symmetric])
ultimately show ?thesis by-(rule eval'.StaticCall'[OF refl])
next
case (CallNull E e s\<^sub>0 s\<^sub>1 es vs)
moreover have "map_val (map Val vs) vs" by(simp add: map_val_conv[symmetric])
ultimately show ?thesis by-(rule eval'.CallNull'[OF refl])
next
case SeqThrow thus ?thesis by(rule eval'.SeqThrow'[OF refl])
next
case CondF thus ?thesis by(rule eval'.CondF'[OF refl])
next
case CondThrow thus ?thesis by(rule eval'.CondThrow'[OF refl])
next
case WhileCondThrow thus ?thesis by(rule eval'.WhileCondThrow'[OF refl])
next
case WhileBodyThrow thus ?thesis by(rule eval'.WhileBodyThrow'[OF refl])
next
case ThrowNull thus ?thesis by(rule eval'.ThrowNull'[OF refl])
qed(assumption|erule (4) eval'.that[OF refl])+
next
case evals'
from evals'.prems evals'.that[OF refl]
show thesis by transfer(erule evals.cases)
qed
subsection \<open>Examples\<close>
declare [[values_timeout = 180]]
values [expected "{Val (Intg 5)}"]
"{fst (e', s') | e' s'.
[],Map.empty \<turnstile> \<langle>{''V'':Integer; ''V'' := Val(Intg 5);; Var ''V''},(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val (Intg 11)}"]
"{fst (e', s') | e' s'.
[],Map.empty \<turnstile> \<langle>(Val(Intg 5)) \<guillemotleft>Add\<guillemotright> (Val(Intg 6)),(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val (Intg 83)}"]
"{fst (e', s') | e' s'.
[],[''V''\<mapsto>Integer] \<turnstile> \<langle>(Var ''V'') \<guillemotleft>Add\<guillemotright> (Val(Intg 6)),
(Map.empty,[''V''\<mapsto>Intg 77])\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Some (Intg 6)}"]
"{lcl' (snd (e', s')) ''V'' | e' s'.
[],[''V''\<mapsto>Integer] \<turnstile> \<langle>''V'' := Val(Intg 6),(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Some (Intg 12)}"]
"{lcl' (snd (e', s')) ''mult'' | e' s'.
[],[''V''\<mapsto>Integer,''a''\<mapsto>Integer,''b''\<mapsto>Integer,''mult''\<mapsto>Integer]
\<turnstile> \<langle>(''a'' := Val(Intg 3));;(''b'' := Val(Intg 4));;(''mult'' := Val(Intg 0));;
(''V'' := Val(Intg 1));;
while (Var ''V'' \<guillemotleft>Eq\<guillemotright> Val(Intg 1))((''mult'' := Var ''mult'' \<guillemotleft>Add\<guillemotright> Var ''b'');;
(''a'' := Var ''a'' \<guillemotleft>Add\<guillemotright> Val(Intg (- 1)));;
(''V'' := (if(Var ''a'' \<guillemotleft>Eq\<guillemotright> Val(Intg 0)) Val(Intg 0) else Val(Intg 1)))),
(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val (Intg 30)}"]
"{fst (e', s') | e' s'.
[],[''a''\<mapsto>Integer, ''b''\<mapsto>Integer, ''c''\<mapsto> Integer, ''cond''\<mapsto>Boolean]
\<turnstile> \<langle>''a'' := Val(Intg 17);; ''b'' := Val(Intg 13);;
''c'' := Val(Intg 42);; ''cond'' := true;;
if (Var ''cond'') (Var ''a'' \<guillemotleft>Add\<guillemotright> Var ''b'') else (Var ''a'' \<guillemotleft>Add\<guillemotright> Var ''c''),
(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e',s'\<rangle>}"
text \<open>progOverrider examples\<close>
definition
classBottom :: "cdecl" where
"classBottom = (''Bottom'', [Repeats ''Left'', Repeats ''Right''],
[(''x'',Integer)],[])"
definition
classLeft :: "cdecl" where
"classLeft = (''Left'', [Repeats ''Top''],[],[(''f'', [Class ''Top'', Integer],Integer, [''V'',''W''],Var this \<bullet> ''x'' {[''Left'',''Top'']} \<guillemotleft>Add\<guillemotright> Val (Intg 5))])"
definition
classRight :: "cdecl" where
"classRight = (''Right'', [Shares ''Right2''],[],
[(''f'', [Class ''Top'', Integer], Integer,[''V'',''W''],Var this \<bullet> ''x'' {[''Right2'',''Top'']} \<guillemotleft>Add\<guillemotright> Val (Intg 7)),(''g'',[],Class ''Left'',[],new ''Left'')])"
definition
classRight2 :: "cdecl" where
"classRight2 = (''Right2'', [Repeats ''Top''],[],
[(''f'', [Class ''Top'', Integer], Integer,[''V'',''W''],Var this \<bullet> ''x'' {[''Right2'',''Top'']} \<guillemotleft>Add\<guillemotright> Val (Intg 9)),(''g'',[],Class ''Top'',[],new ''Top'')])"
definition
classTop :: "cdecl" where
"classTop = (''Top'', [], [(''x'',Integer)],[])"
definition
progOverrider :: "cdecl list" where
"progOverrider = [classBottom, classLeft, classRight, classRight2, classTop]"
values [expected "{Val(Ref(0,[''Bottom'',''Left'']))}"] \<comment> \<open>dynCastSide\<close>
"{fst (e', s') | e' s'.
progOverrider,[''V''\<mapsto>Class ''Right''] \<turnstile>
\<langle>''V'' := new ''Bottom'' ;; Cast ''Left'' (Var ''V''),(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val(Ref(0,[''Right'']))}"] \<comment> \<open>dynCastViaSh\<close>
"{fst (e', s') | e' s'.
progOverrider,[''V''\<mapsto>Class ''Right2''] \<turnstile>
\<langle>''V'' := new ''Right'' ;; Cast ''Right'' (Var ''V''),(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val (Intg 42)}"] \<comment> \<open>block\<close>
"{fst (e', s') | e' s'.
progOverrider,[''V''\<mapsto>Integer]
\<turnstile> \<langle>''V'' := Val(Intg 42) ;; {''V'':Class ''Left''; ''V'' := new ''Bottom''} ;; Var ''V'',
(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val (Intg 8)}"] \<comment> \<open>staticCall\<close>
"{fst (e', s') | e' s'.
progOverrider,[''V''\<mapsto>Class ''Right'',''W''\<mapsto>Class ''Bottom'']
\<turnstile> \<langle>''V'' := new ''Bottom'' ;; ''W'' := new ''Bottom'' ;;
((Cast ''Left'' (Var ''W''))\<bullet>''x''{[''Left'',''Top'']} := Val(Intg 3));;
(Var ''W''\<bullet>(''Left''::)''f''([Var ''V'',Val(Intg 2)])),(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val (Intg 12)}"] \<comment> \<open>call\<close>
"{fst (e', s') | e' s'.
progOverrider,[''V''\<mapsto>Class ''Right2'',''W''\<mapsto>Class ''Left'']
\<turnstile> \<langle>''V'' := new ''Right'' ;; ''W'' := new ''Left'' ;;
(Var ''V''\<bullet>''f''([Var ''W'',Val(Intg 42)])) \<guillemotleft>Add\<guillemotright> (Var ''W''\<bullet>''f''([Var ''V'',Val(Intg 13)])),
(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val(Intg 13)}"] \<comment> \<open>callOverrider\<close>
"{fst (e', s') | e' s'.
progOverrider,[''V''\<mapsto>Class ''Right2'',''W''\<mapsto>Class ''Left'']
\<turnstile> \<langle>''V'' := new ''Bottom'';; (Var ''V'' \<bullet> ''x'' {[''Right2'',''Top'']} := Val(Intg 6));;
''W'' := new ''Left'' ;; Var ''V''\<bullet>''f''([Var ''W'',Val(Intg 42)]),
(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val(Ref(1,[''Left'',''Top'']))}"] \<comment> \<open>callClass\<close>
"{fst (e', s') | e' s'.
progOverrider,[''V''\<mapsto>Class ''Right2'']
\<turnstile> \<langle>''V'' := new ''Right'' ;; Var ''V''\<bullet>''g''([]),(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val(Intg 42)}"] \<comment> \<open>fieldAss\<close>
"{fst (e', s') | e' s'.
progOverrider,[''V''\<mapsto>Class ''Right2'']
\<turnstile> \<langle>''V'' := new ''Right'' ;;
(Var ''V''\<bullet>''x''{[''Right2'',''Top'']} := (Val(Intg 42))) ;;
(Var ''V''\<bullet>''x''{[''Right2'',''Top'']}),(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
text \<open>typing rules\<close>
values [expected "{Class ''Bottom''}"] \<comment> \<open>typeNew\<close>
"{T. progOverrider,Map.empty \<turnstile> new ''Bottom'' :: T}"
values [expected "{Class ''Left''}"] \<comment> \<open>typeDynCast\<close>
"{T. progOverrider,Map.empty \<turnstile> Cast ''Left'' (new ''Bottom'') :: T}"
values [expected "{Class ''Left''}"] \<comment> \<open>typeStaticCast\<close>
"{T. progOverrider,Map.empty \<turnstile> \<lparr>''Left''\<rparr> (new ''Bottom'') :: T}"
values [expected "{Integer}"] \<comment> \<open>typeVal\<close>
"{T. [],Map.empty \<turnstile> Val(Intg 17) :: T}"
values [expected "{Integer}"] \<comment> \<open>typeVar\<close>
"{T. [],[''V'' \<mapsto> Integer] \<turnstile> Var ''V'' :: T}"
values [expected "{Boolean}"] \<comment> \<open>typeBinOp\<close>
"{T. [],Map.empty \<turnstile> (Val(Intg 5)) \<guillemotleft>Eq\<guillemotright> (Val(Intg 6)) :: T}"
values [expected "{Class ''Top''}"] \<comment> \<open>typeLAss\<close>
"{T. progOverrider,[''V'' \<mapsto> Class ''Top''] \<turnstile> ''V'' := (new ''Left'') :: T}"
values [expected "{Integer}"] \<comment> \<open>typeFAcc\<close>
"{T. progOverrider,Map.empty \<turnstile> (new ''Right'')\<bullet>''x''{[''Right2'',''Top'']} :: T}"
values [expected "{Integer}"] \<comment> \<open>typeFAss\<close>
"{T. progOverrider,Map.empty \<turnstile> (new ''Right'')\<bullet>''x''{[''Right2'',''Top'']} :: T}"
values [expected "{Integer}"] \<comment> \<open>typeStaticCall\<close>
"{T. progOverrider,[''V''\<mapsto>Class ''Left'']
\<turnstile> ''V'' := new ''Left'' ;; Var ''V''\<bullet>(''Left''::)''f''([new ''Top'', Val(Intg 13)]) :: T}"
values [expected "{Class ''Top''}"] \<comment> \<open>typeCall\<close>
"{T. progOverrider,[''V''\<mapsto>Class ''Right2'']
\<turnstile> ''V'' := new ''Right'' ;; Var ''V''\<bullet>''g''([]) :: T}"
values [expected "{Class ''Top''}"] \<comment> \<open>typeBlock\<close>
"{T. progOverrider,Map.empty \<turnstile> {''V'':Class ''Top''; ''V'' := new ''Left''} :: T}"
values [expected "{Integer}"] \<comment> \<open>typeCond\<close>
"{T. [],Map.empty \<turnstile> if (true) Val(Intg 6) else Val(Intg 9) :: T}"
values [expected "{Void}"] \<comment> \<open>typeWhile\<close>
"{T. [],Map.empty \<turnstile> while (false) Val(Intg 17) :: T}"
values [expected "{Void}"] \<comment> \<open>typeThrow\<close>
"{T. progOverrider,Map.empty \<turnstile> throw (new ''Bottom'') :: T}"
values [expected "{Integer}"] \<comment> \<open>typeBig\<close>
"{T. progOverrider,[''V''\<mapsto>Class ''Right2'',''W''\<mapsto>Class ''Left'']
\<turnstile> ''V'' := new ''Right'' ;; ''W'' := new ''Left'' ;;
(Var ''V''\<bullet>''f''([Var ''W'', Val(Intg 7)])) \<guillemotleft>Add\<guillemotright> (Var ''W''\<bullet>''f''([Var ''V'', Val(Intg 13)]))
:: T}"
text \<open>progDiamond examples\<close>
definition
classDiamondBottom :: "cdecl" where
"classDiamondBottom = (''Bottom'', [Repeats ''Left'', Repeats ''Right''],[(''x'',Integer)],
[(''g'', [],Integer, [],Var this \<bullet> ''x'' {[''Bottom'']} \<guillemotleft>Add\<guillemotright> Val (Intg 5))])"
definition
classDiamondLeft :: "cdecl" where
"classDiamondLeft = (''Left'', [Repeats ''TopRep'',Shares ''TopSh''],[],[])"
definition
classDiamondRight :: "cdecl" where
"classDiamondRight = (''Right'', [Repeats ''TopRep'',Shares ''TopSh''],[],
[(''f'', [Integer], Boolean,[''i''], Var ''i'' \<guillemotleft>Eq\<guillemotright> Val (Intg 7))])"
definition
classDiamondTopRep :: "cdecl" where
"classDiamondTopRep = (''TopRep'', [], [(''x'',Integer)],
[(''g'', [],Integer, [], Var this \<bullet> ''x'' {[''TopRep'']} \<guillemotleft>Add\<guillemotright> Val (Intg 10))])"
definition
classDiamondTopSh :: "cdecl" where
"classDiamondTopSh = (''TopSh'', [], [],
[(''f'', [Integer], Boolean,[''i''], Var ''i'' \<guillemotleft>Eq\<guillemotright> Val (Intg 3))])"
definition
progDiamond :: "cdecl list" where
"progDiamond = [classDiamondBottom, classDiamondLeft, classDiamondRight, classDiamondTopRep, classDiamondTopSh]"
values [expected "{Val(Ref(0,[''Bottom'',''Left'']))}"] \<comment> \<open>cast1\<close>
"{fst (e', s') | e' s'.
progDiamond,[''V''\<mapsto>Class ''Left''] \<turnstile> \<langle>''V'' := new ''Bottom'',
(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val(Ref(0,[''TopSh'']))}"] \<comment> \<open>cast2\<close>
"{fst (e', s') | e' s'.
progDiamond,[''V''\<mapsto>Class ''TopSh''] \<turnstile> \<langle>''V'' := new ''Bottom'',
(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{}"] \<comment> \<open>typeCast3 not typeable\<close>
"{T. progDiamond,[''V''\<mapsto>Class ''TopRep''] \<turnstile> ''V'' := new ''Bottom'' :: T}"
values [expected "{
Val(Ref(0,[''Bottom'', ''Left'', ''TopRep''])),
Val(Ref(0,[''Bottom'', ''Right'', ''TopRep'']))
}"] \<comment> \<open>cast3\<close>
"{fst (e', s') | e' s'.
progDiamond,[''V''\<mapsto>Class ''TopRep''] \<turnstile> \<langle>''V'' := new ''Bottom'',
(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
values [expected "{Val(Intg 17)}"] \<comment> \<open>fieldAss\<close>
"{fst (e', s') | e' s'.
progDiamond,[''V''\<mapsto>Class ''Bottom'']
\<turnstile> \<langle>''V'' := new ''Bottom'' ;;
((Var ''V'')\<bullet>''x''{[''Bottom'']} := (Val(Intg 17))) ;;
((Var ''V'')\<bullet>''x''{[''Bottom'']}),(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e',s'\<rangle>}"
values [expected "{Val Null}"] \<comment> \<open>dynCastNull\<close>
"{fst (e', s') | e' s'.
progDiamond,Map.empty \<turnstile> \<langle>Cast ''Right'' null,(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e',s'\<rangle>}"
values [expected "{Val (Ref(0, [''Right'']))}"] \<comment> \<open>dynCastViaSh\<close>
"{fst (e', s') | e' s'.
progDiamond,[''V''\<mapsto>Class ''TopSh'']
\<turnstile> \<langle>''V'' := new ''Right'' ;; Cast ''Right'' (Var ''V''),(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e',s'\<rangle>}"
values [expected "{Val Null}"] \<comment> \<open>dynCastFail\<close>
"{fst (e', s') | e' s'.
progDiamond,[''V''\<mapsto>Class ''TopRep'']
\<turnstile> \<langle>''V'' := new ''Right'' ;; Cast ''Bottom'' (Var ''V''),(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e',s'\<rangle>}"
values [expected "{Val (Ref(0, [''Bottom'', ''Left'']))}"] \<comment> \<open>dynCastSide\<close>
"{fst (e', s') | e' s'.
progDiamond,[''V''\<mapsto>Class ''Right'']
\<turnstile> \<langle>''V'' := new ''Bottom'' ;; Cast ''Left'' (Var ''V''),(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e',s'\<rangle>}"
text \<open>failing g++ example\<close>
definition
classD :: "cdecl" where
"classD = (''D'', [Shares ''A'', Shares ''B'', Repeats ''C''],[],[])"
definition
classC :: "cdecl" where
"classC = (''C'', [Shares ''A'', Shares ''B''],[],
[(''f'',[],Integer,[],Val(Intg 42))])"
definition
classB :: "cdecl" where
"classB = (''B'', [],[],
[(''f'',[],Integer,[],Val(Intg 17))])"
definition
classA :: "cdecl" where
"classA = (''A'', [],[],
[(''f'',[],Integer,[],Val(Intg 13))])"
definition
ProgFailing :: "cdecl list" where
"ProgFailing = [classA,classB,classC,classD]"
values [expected "{Val (Intg 42)}"] \<comment> \<open>callFailGplusplus\<close>
"{fst (e', s') | e' s'.
ProgFailing,Map.empty
\<turnstile> \<langle>{''V'':Class ''D''; ''V'' := new ''D'';; Var ''V''\<bullet>''f''([])},
(Map.empty,Map.empty)\<rangle> \<Rightarrow>' \<langle>e', s'\<rangle>}"
end
diff --git a/thys/Dijkstra_Shortest_Path/Dijkstra.thy b/thys/Dijkstra_Shortest_Path/Dijkstra.thy
--- a/thys/Dijkstra_Shortest_Path/Dijkstra.thy
+++ b/thys/Dijkstra_Shortest_Path/Dijkstra.thy
@@ -1,980 +1,981 @@
section \<open>Dijkstra's Algorithm\<close>
theory Dijkstra
imports
Graph
Dijkstra_Misc
Collections.Refine_Dflt_ICF
Weight
begin
text \<open>
This theory defines Dijkstra's algorithm. First, a correct result of
Dijkstra's algorithm w.r.t. a graph and a start vertex is specified.
Then, the refinement
framework is used to specify Dijkstra's Algorithm, prove it correct, and
finally refine it to datatypes that are closer to an implementation than
the original specification.
\<close>
subsection "Graph's for Dijkstra's Algorithm"
text \<open>A graph annotated with weights.\<close>
locale weighted_graph = valid_graph G
for G :: "('V,'W::weight) graph"
subsection "Specification of Correct Result"
context weighted_graph
begin
text \<open>
A result of Dijkstra's algorithm is correct, if it is a map from nodes
\<open>v\<close> to the shortest path from the start node \<open>v0\<close> to
\<open>v\<close>. Iff there is no such path, the node is not in the map.
\<close>
definition is_shortest_path_map :: "'V \<Rightarrow> ('V \<rightharpoonup> ('V,'W) path) \<Rightarrow> bool"
where
"is_shortest_path_map v0 res \<equiv> \<forall>v\<in>V. (case res v of
None \<Rightarrow> \<not>(\<exists>p. is_path v0 p v) |
Some p \<Rightarrow> is_path v0 p v
\<and> (\<forall>p'. is_path v0 p' v \<longrightarrow> path_weight p \<le> path_weight p')
)"
end
text \<open>
The following function returns the weight of an optional path,
where \<open>None\<close> is interpreted as infinity.
\<close>
fun path_weight' where
"path_weight' None = top" |
"path_weight' (Some p) = Num (path_weight p)"
subsection "Dijkstra's Algorithm"
text \<open>
The state in the main loop of the algorithm consists of a workset
\<open>wl\<close> of vertexes that still need to be explored, and a map
\<open>res\<close> that contains the current shortest path for each vertex.
\<close>
type_synonym ('V,'W) state = "('V set) \<times> ('V \<rightharpoonup> ('V,'W) path)"
text \<open>
The preconditions of Dijkstra's algorithm, i.e., that it operates on a
valid and finite graph, and that the start node is a node of the graph,
are summarized in a locale.
\<close>
locale Dijkstra = weighted_graph G
for G :: "('V,'W::weight) graph"+
fixes v0 :: 'V
assumes finite[simp,intro!]: "finite V" "finite E"
assumes v0_in_V[simp, intro!]: "v0\<in>V"
assumes nonneg_weights[simp, intro]: "(v,w,v')\<in>edges G \<Longrightarrow> 0\<le>w"
begin
text \<open>Paths have non-negative weights.\<close>
lemma path_nonneg_weight: "is_path v p v' \<Longrightarrow> 0 \<le> path_weight p"
by (induct rule: is_path.induct) auto
text \<open>Invariant of the main loop:
\begin{itemize}
\item The workset only contains nodes of the graph.
\item If the result set contains a path for a node, it is actually a path,
and uses only intermediate vertices outside the workset.
\item For all vertices outside the workset, the result map contains the
shortest path.
\item For all vertices in the workset, the result map contains the
shortest path among all paths that only use intermediate vertices outside
the workset.
\end{itemize}
\<close>
definition "dinvar \<sigma> \<equiv> let (wl,res)=\<sigma> in
wl \<subseteq> V \<and>
(\<forall>v\<in>V. \<forall>p. res v = Some p \<longrightarrow> is_path v0 p v \<and> int_vertices p \<subseteq> V-wl) \<and>
(\<forall>v\<in>V-wl. \<forall>p. is_path v0 p v
\<longrightarrow> path_weight' (res v) \<le> path_weight' (Some p)) \<and>
(\<forall>v\<in>wl. \<forall>p. is_path v0 p v \<and> int_vertices p \<subseteq> V-wl
\<longrightarrow> path_weight' (res v) \<le> path_weight' (Some p)
)
"
text \<open>Sanity check: The invariant is strong enough to imply correctness
of result.\<close>
lemma invar_imp_correct: "dinvar ({},res) \<Longrightarrow> is_shortest_path_map v0 res"
unfolding dinvar_def is_shortest_path_map_def
by (auto simp: infty_unbox split: option.split)
text \<open>
The initial workset contains all vertices. The initial result maps
\<open>v0\<close> to the empty path, and all other vertices to \<open>None\<close>.
\<close>
definition dinit :: "('V,'W) state nres" where
"dinit \<equiv> SPEC ( \<lambda>(wl,res) .
wl=V \<and> res v0 = Some [] \<and> (\<forall>v\<in>V-{v0}. res v = None))"
text \<open>
The initial state satisfies the invariant.
\<close>
lemma dinit_invar: "dinit \<le> SPEC dinvar"
unfolding dinit_def
apply (intro refine_vcg)
apply (force simp: dinvar_def split: option.split)
done
text \<open>
In each iteration, the main loop of the algorithm pops a minimal node from
the workset, and then updates the result map accordingly.
\<close>
text \<open>
Pop a minimal node from the workset. The node is minimal in the sense that
the length of the current path for that node is minimal.
\<close>
definition pop_min :: "('V,'W) state \<Rightarrow> ('V \<times> ('V,'W) state) nres" where
"pop_min \<sigma> \<equiv> do {
let (wl,res)=\<sigma>;
ASSERT (wl\<noteq>{});
v \<leftarrow> RES (least_map (path_weight' \<circ> res) wl);
RETURN (v,(wl-{v},res))
}"
text \<open>
Updating the result according to a node \<open>v\<close> is done by checking,
for each successor node, whether the path over \<open>v\<close> is shorter than
the path currently stored into the result map.
\<close>
inductive update_spec :: "'V \<Rightarrow> ('V,'W) state \<Rightarrow> ('V,'W) state \<Rightarrow> bool"
where
"\<lbrakk> \<forall>v'\<in>V.
res' v' \<in> least_map path_weight' (
{ res v' } \<union> { Some (p@[(v,w,v')]) | p w. res v = Some p \<and> (v,w,v')\<in>E }
)
\<rbrakk> \<Longrightarrow> update_spec v (wl,res) (wl,res')"
text \<open>
In order to ease the refinement proof, we will assert the following
precondition for updating.
\<close>
definition update_pre :: "'V \<Rightarrow> ('V,'W) state \<Rightarrow> bool" where
"update_pre v \<sigma> \<equiv> let (wl,res)=\<sigma> in v\<in>V
\<and> (\<forall>v'\<in>V-wl. v'\<noteq>v \<longrightarrow> (\<forall>p. is_path v0 p v'
\<longrightarrow> path_weight' (res v') \<le> path_weight' (Some p)))
\<and> (\<forall>v'\<in>V. \<forall>p. res v' = Some p \<longrightarrow> is_path v0 p v')"
definition update :: "'V \<Rightarrow> ('V,'W) state \<Rightarrow> ('V,'W) state nres" where
"update v \<sigma> \<equiv> do {ASSERT (update_pre v \<sigma>); SPEC (update_spec v \<sigma>)}"
text \<open>Finally, we define Dijkstra's algorithm:\<close>
definition dijkstra where
"dijkstra \<equiv> do {
\<sigma>0\<leftarrow>dinit;
(_,res) \<leftarrow> WHILE\<^sub>T\<^bsup>dinvar\<^esup> (\<lambda>(wl,_). wl\<noteq>{})
(\<lambda>\<sigma>.
do { (v,\<sigma>') \<leftarrow> pop_min \<sigma>; update v \<sigma>' }
)
\<sigma>0;
RETURN res }
"
text \<open>The following theorem states (total) correctness of Dijkstra's
algorithm.\<close>
theorem dijkstra_correct: "dijkstra \<le> SPEC (is_shortest_path_map v0)"
unfolding dijkstra_def
unfolding dinit_def
unfolding pop_min_def update_def [abs_def]
thm refine_vcg
apply (refine_rcg
WHILEIT_rule[where R="inv_image {(x,y). x<y} (card \<circ> fst)"]
refine_vcg
)
(* TODO/FIXME: Should we built in such massaging of the goal into
refine_rcg ?*)
+ supply [[simproc del: defined_all]]
apply (simp_all split: prod.split_asm)
apply (tactic \<open>
ALLGOALS ((REPEAT_DETERM o Hypsubst.bound_hyp_subst_tac @{context})
THEN' asm_full_simp_tac @{context}
)\<close>)
proof -
fix wl res v
assume INV: "dinvar (wl,res)"
and LM: "v\<in>least_map (path_weight' \<circ> res) wl"
hence "v\<in>V" unfolding dinvar_def by (auto dest: least_map_elemD)
moreover
from INV have " \<forall>v'\<in>V - (wl-{v}). v' \<noteq> v \<longrightarrow>
(\<forall>p. is_path v0 p v' \<longrightarrow> path_weight' (res v') \<le> Num (path_weight p))"
by (auto simp: dinvar_def)
moreover from INV have "\<forall>v'\<in>V. \<forall>p. res v'=Some p \<longrightarrow> is_path v0 p v'"
by (auto simp: dinvar_def)
ultimately show "update_pre v (wl-{v},res)" by (auto simp: update_pre_def)
next
fix res
assume "dinvar ({}, res)"
thus "is_shortest_path_map v0 res"
by (rule invar_imp_correct)
next
show "wf (inv_image {(x, y). x < y} (card \<circ> fst))"
by (blast intro: wf_less)
next
fix wl res v \<sigma>''
assume
LM: "v\<in>least_map (path_weight' \<circ> res) wl" and
UD: "update_spec v (wl-{v},res) \<sigma>''" and
INV: "dinvar (wl,res)"
from LM have "v\<in>wl" by (auto dest: least_map_elemD)
moreover from UD have "fst \<sigma>'' = wl-{v}" by (auto elim: update_spec.cases)
moreover from INV have "finite wl"
unfolding dinvar_def by (auto dest: finite_subset)
ultimately show "card (fst \<sigma>'') < card wl"
apply simp
by (metis card_gt_0_iff diff_Suc_less empty_iff)
next
fix a and res :: "'V \<rightharpoonup> ('V,'W) path"
assume "a = V \<and> res v0 = Some [] \<and> (\<forall>v\<in>V-{v0}. res v = None)"
thus "dinvar (V,res)"
by (force simp: dinvar_def split: option.split)
next
fix wl res
assume INV: "dinvar (wl,res)"
hence
WL_SUBSET: "wl \<subseteq> V" and
PATH_VALID: "\<forall>v\<in>V. \<forall>p. res v = Some p
\<longrightarrow> is_path v0 p v \<and> int_vertices p \<subseteq> V - wl" and
NWL_MIN: "\<forall>v\<in>V - wl. \<forall>p. is_path v0 p v
\<longrightarrow> path_weight' (res v) \<le> Num (path_weight p)" and
WL_MIN: "\<forall>v\<in>wl. \<forall>p. is_path v0 p v \<and> int_vertices p \<subseteq> V - wl
\<longrightarrow> path_weight' (res v) \<le> Num (path_weight p)"
unfolding dinvar_def by auto
fix v \<sigma>''
assume V_LEAST: "v\<in>least_map (path_weight' o res) wl"
and "update_spec v (wl-{v},res) \<sigma>''"
then obtain res' where
[simp]: "\<sigma>''=(wl-{v},res')"
and CONSIDERED_NEW_PATHS: "\<forall>v'\<in>V. res' v' \<in> least_map path_weight'
(insert (res v')
({ Some (p@[(v,w,v')]) | p w. res v = Some p \<and> (v,w,v')\<in>E }))"
by (auto elim!: update_spec.cases)
from V_LEAST have V_MEM: "v\<in>wl" by (blast intro: least_map_elemD)
show "dinvar \<sigma>''"
apply (unfold dinvar_def, simp)
apply (intro conjI)
proof -
from WL_SUBSET show "wl-{v} \<subseteq> V" by auto
show "\<forall>va\<in>V. \<forall>p. res' va = Some p
\<longrightarrow> is_path v0 p va \<and> int_vertices p \<subseteq> V - (wl - {v})"
proof (intro ballI conjI impI allI)
fix v' p
assume V'_MEM: "v'\<in>V" and [simp]: "res' v' = Some p"
txt \<open>The new paths that we have added are valid and only use
intermediate vertices outside the workset.
This proof works as follows: A path @{term "res' v'"} is either
the old path, or has been assembled as a path over node @{term v}.
In the former case the proposition follows straightforwardly from the
invariant for the old state. In the latter case we get, by the invariant
for the old state, that the path over node @{term v} is valid.
Then, we observe that appending an edge to a valid path yields a valid
path again. Also, adding @{term v} as intermediate node is legal, as we
just removed @{term v} from the workset.
\<close>
with CONSIDERED_NEW_PATHS have "res' v' \<in> (insert (res v')
({ Some (p@[(v,w,v')]) | p w. res v = Some p \<and> (v,w,v')\<in>E }))"
by (rule_tac least_map_elemD) blast
moreover {
assume [symmetric,simp]: "res' v' = res v'"
from V'_MEM PATH_VALID have
"is_path v0 p v'"
"int_vertices p \<subseteq> V - (wl-{v})"
by force+
} moreover {
fix pv w
assume "res' v' = Some (pv@[(v,w,v')])"
and [simp]: "res v = Some pv"
and EDGE: "(v,w,v')\<in>E"
hence [simp]: "p = pv@[(v,w,v')]" by simp
from bspec[OF PATH_VALID rev_subsetD[OF V_MEM WL_SUBSET]] have
PATHV: "is_path v0 pv v" and IVV: "int_vertices pv \<subseteq> V - wl" by auto
hence
"is_path v0 p v'"
"int_vertices p \<subseteq> V - (wl-{v})"
by (auto simp: EDGE V'_MEM)
}
ultimately show
"is_path v0 p v'"
"int_vertices p \<subseteq> V - (wl-{v})"
by blast+
qed
txt \<open>
We show that already the {\em original} result stores the minimal
path for all vertices not in the {\em new} workset.
For vertices also not in the original workset, this follows
straightforwardly from the invariant.
For the vertex \<open>v\<close>, that has been removed from the
workset, we split a path \<open>p'\<close> to \<open>v\<close> at the point
\<open>u\<close> where it first enters the original workset.
As we chose \<open>v\<close> to be the vertex in the workset with the
minimal weight, its weight is less than the current weight of
\<open>u\<close>. As the vertices of the prefix of \<open>p'\<close> up to
\<open>u\<close> are not in the workset, the current weight of
\<open>u\<close> is less than the weight of the prefix of \<open>p'\<close>, and thus less than the weight of \<open>p'\<close>.
Together, the current weight of \<open>v\<close> is less than the weight of
\<open>p'\<close>.\<close>
have RES_MIN: "\<forall>v\<in>V - (wl - {v}). \<forall>p. is_path v0 p v
\<longrightarrow> path_weight' (res v) \<le> Num (path_weight p)"
proof (intro ballI allI impI)
fix v' p'
assume NOT_IN_WL: "v' \<in> V - (wl - {v})"
and PATH: "is_path v0 p' v'"
hence [simp, intro!]: "v'\<in>V" by auto
show "path_weight' (res v') \<le> Num (path_weight p')"
proof (cases "v' = v")
assume NE[simp]: "v'\<noteq>v"
from bspec[OF NWL_MIN, of v'] NOT_IN_WL PATH show
"path_weight' (res v') \<le> Num (path_weight p')" by auto
next
assume EQ[simp]: "v'=v"
from path_split_set'[OF PATH, of wl] V_MEM obtain p1 p2 u where
[simp]: "p'=p1@p2"
and P1: "is_path v0 p1 u"
and P2: "is_path u p2 v'"
and P1V: "int_vertices p1 \<subseteq> -wl"
and [simp]: "u\<in>wl"
by auto
from least_map_leD[OF V_LEAST]
have "path_weight' (res v') \<le> path_weight' (res u)"by auto
also from bspec[OF WL_MIN, of u] P1 P1V int_vertices_subset[OF P1]
have "path_weight' (res u) \<le> Num (path_weight p1)" by auto
also have "\<dots> \<le> Num (path_weight p')"
using path_nonneg_weight[OF P2]
apply (auto simp: infty_unbox )
by (metis add_0_right add_left_mono)
finally show ?thesis .
qed
qed
txt \<open>With the previous statement, we easily show the
third part of the invariant, as the new paths are not longer than the
old ones.
\<close>
show "\<forall>v\<in>V - (wl - {v}). \<forall>p. is_path v0 p v
\<longrightarrow> path_weight' (res' v) \<le> Num (path_weight p)"
proof (intro allI ballI impI)
fix v' p
assume NOT_IN_WL: "v' \<in> V - (wl - {v})"
and PATH: "is_path v0 p v'"
hence [simp, intro!]: "v'\<in>V" by auto
from bspec[OF CONSIDERED_NEW_PATHS, of v']
have "path_weight' (res' v') \<le> path_weight' (res v')"
by (auto dest: least_map_leD)
also from bspec[OF RES_MIN NOT_IN_WL] PATH
have "path_weight' (res v') \<le> Num (path_weight p)" by blast
finally show "path_weight' (res' v') \<le> Num (path_weight p)" .
qed
txt \<open>
Finally, we have to show that for nodes on the worklist,
the stored paths are not longer than any path using only nodes not
on the worklist. Compared to the situation before the step, those
path may also use the node \<open>v\<close>.
\<close>
show "\<forall>va\<in>wl - {v}. \<forall>p.
is_path v0 p va \<and> int_vertices p \<subseteq> V - (wl - {v})
\<longrightarrow> path_weight' (res' va) \<le> Num (path_weight p)"
proof (intro allI impI ballI, elim conjE)
fix v' p
assume IWS: "v'\<in>wl - {v}"
and PATH: "is_path v0 p v'"
and VERTICES: "int_vertices p \<subseteq> V - (wl - {v})"
from IWS WL_SUBSET have [simp, intro!]: "v'\<in>V" by auto
{
txt \<open>
If the path is empty, the proposition follows easily from the
invariant for the original states, as no intermediate nodes are
used at all.
\<close>
assume [simp]: "p=[]"
from bspec[OF CONSIDERED_NEW_PATHS, of v'] have
"path_weight' (res' v') \<le> path_weight' (res v')"
using IWS WL_SUBSET by (auto dest: least_map_leD)
also have "int_vertices p \<subseteq> V-wl" by auto
with WL_MIN IWS PATH
have "path_weight' (res v') \<le> Num (path_weight p)"
by (auto simp del: path_weight_empty)
finally have "path_weight' (res' v') \<le> Num (path_weight p)" .
} moreover {
fix p1 u w
assume [simp]: "p = p1@[(u,w,v')]"
txt \<open>If the path is not empty, we pick the last but one vertex, and
call it @{term u}.\<close>
from PATH have PATH1: "is_path v0 p1 u" and EDGE: "(u,w,v')\<in>E" by auto
from VERTICES have NIV: "u\<in>V - (wl-{v})" by simp
hence U_MEM[simp]: "u\<in>V" by auto
txt \<open>From @{thm [source] RES_MIN}, we know that @{term "res u"} holds
the shortest path to @{term u}. Thus \<open>p\<close> is longer than the
path that is constructed by replacing the prefix of @{term p} by
{term "res u"}\<close>
from NIV RES_MIN PATH1
have G: "Num (path_weight p1) \<ge> path_weight' (res u)" by simp
then obtain pu where [simp]: "res u = Some pu"
by (cases "res u") (auto simp: infty_unbox)
from G have "Num (path_weight p) \<ge> path_weight' (res u) + Num w"
by (auto simp: infty_unbox add_right_mono)
also
have "path_weight' (res u) + Num w \<ge> path_weight' (res' v')"
txt \<open>
The remaining argument depends on wether @{term u}
equals @{term v}.
In the case @{term "u\<noteq>v"}, all vertices of @{term "res u"} are
outside the original workset. Thus, appending the edge
@{term "(u,w,v')"} to @{term "res u"} yields a path to @{term v}
over intermediate nodes only outside the workset. By the invariant
for the original state, @{term "res v'"} is shorter than this path.
As a step does not replace paths by longer ones, also
@{term "res' v'"} is shorter.
In the case @{term "u=v"}, the step has
considered the path to \<open>v'\<close> over \<open>v\<close>, and thus the
result path is not longer.
\<close>
proof (cases "u=v")
assume "u\<noteq>v"
with NIV have NIV': "u\<in>V-wl" by auto
from bspec[OF PATH_VALID U_MEM] NIV'
have "is_path v0 pu u" and VU: "int_vertices (pu@[(u,w,v')]) \<subseteq> V-wl"
by auto
with EDGE have PV': "is_path v0 (pu@[(u,w,v')]) v'" by auto
with bspec[OF WL_MIN, of v'] IWS VU have
"path_weight' (res v') \<le> Num (path_weight (pu@[(u,w,v')]))"
by blast
hence "path_weight' (res u) + Num w \<ge> path_weight' (res v')"
by (auto simp: infty_unbox)
also from CONSIDERED_NEW_PATHS have
"path_weight' (res v') \<ge> path_weight' (res' v')"
by (auto dest: least_map_leD)
finally (order_trans[rotated]) show ?thesis .
next
assume [symmetric,simp]: "u=v"
from CONSIDERED_NEW_PATHS EDGE have
"path_weight' (res' v') \<le> path_weight' (Some (pu@[(v,w,v')]))"
by (rule_tac least_map_leD) auto
thus ?thesis by (auto simp: infty_unbox)
qed
finally (order_trans[rotated]) have
"path_weight' (res' v') \<le> Num (path_weight p)" .
} ultimately show "path_weight' (res' v') \<le> Num (path_weight p)"
using PATH apply (cases p rule: rev_cases) by auto
qed
qed
qed
subsection \<open>Structural Refinement of Update\<close>
text \<open>
Now that we have proved correct the initial version of the algorithm, we start
refinement towards an efficient implementation.
\<close>
text \<open>
First, the update function is refined to iterate over each successor of the
selected node, and update the result on demand.
\<close>
definition uinvar
:: "'V \<Rightarrow> 'V set \<Rightarrow> _ \<Rightarrow> ('W\<times>'V) set \<Rightarrow> ('V,'W) state \<Rightarrow> bool" where
"uinvar v wl res it \<sigma> \<equiv> let (wl',res')=\<sigma> in wl'=wl
\<and> (\<forall>v'\<in>V.
res' v' \<in> least_map path_weight' (
{ res v' } \<union> { Some (p@[(v,w,v')]) | p w. res v = Some p
\<and> (w,v') \<in> succ G v - it }
))
\<and> (\<forall>v'\<in>V. \<forall>p. res' v' = Some p \<longrightarrow> is_path v0 p v')
\<and> res' v = res v
"
definition update' :: "'V \<Rightarrow> ('V,'W) state \<Rightarrow> ('V,'W) state nres" where
"update' v \<sigma> \<equiv> do {
ASSERT (update_pre v \<sigma>);
let (wl,res) = \<sigma>;
let wv = path_weight' (res v);
let pv = res v;
FOREACH\<^bsup>uinvar v wl res\<^esup> (succ G v) (\<lambda>(w',v') (wl,res).
if (wv + Num w' < path_weight' (res v')) then do {
ASSERT (v'\<in>wl \<and> pv\<noteq>None);
RETURN (wl,res(v' \<mapsto> the pv@[(v,w',v')]))
} else RETURN (wl,res)
) (wl,res)}"
lemma update'_refines:
assumes "v'=v" and "\<sigma>'=\<sigma>"
shows "update' v' \<sigma>' \<le> \<Down>Id (update v \<sigma>)"
apply (simp only: assms)
unfolding update'_def update_def
apply (refine_rcg refine_vcg)
(*apply (intro refine_vcg conjI)*)
apply (simp_all only: singleton_iff)
proof -
fix wl res
assume "update_pre v (wl,res)"
thus "uinvar v wl res (succ G v) (wl,res)"
by (simp add: uinvar_def update_pre_def)
next
fix wl res it wl' res' v' w'
assume PRE: "update_pre v (wl,res)"
assume INV: "uinvar v wl res it (wl',res')"
assume MEM: "(w',v')\<in>it"
assume IT_SS: "it\<subseteq> succ G v"
assume LESS: "path_weight' (res v) + Num w' < path_weight' (res' v')"
from PRE have [simp, intro!]: "v\<in>V" by (simp add: update_pre_def)
from MEM IT_SS have [simp,intro!]: "v'\<in>V" using succ_subset
by auto
from LESS obtain pv where [simp]: "res v = Some pv"
by (cases "res v") auto
thus "res v \<noteq> None" by simp
have [simp]: "wl'=wl" and [simp]: "res' v = res v"
using INV unfolding uinvar_def by auto
from MEM IT_SS have EDGE[simp]: "(v,w',v')\<in>E"
unfolding succ_def by auto
with INV have [simp]: "is_path v0 pv v"
unfolding uinvar_def by auto
have "0\<le>w'" by (rule nonneg_weights[OF EDGE])
hence [simp]: "v'\<noteq>v" using LESS
by auto
hence [simp]: "v\<noteq>v'" by blast
show [simp]: "v'\<in>wl'" proof (rule ccontr)
assume [simp]: "v'\<notin>wl'"
hence [simp]: "v'\<in>V-wl" and [simp]: "v'\<notin>wl" by auto
note LESS
also
from INV have "path_weight' (res' v') \<le> path_weight' (res v')"
unfolding uinvar_def by (auto dest: least_map_leD)
also
from PRE have PW: "\<And>p. is_path v0 p v' \<Longrightarrow>
path_weight' (res v') \<le> path_weight' (Some p)"
unfolding update_pre_def
by auto
have P: "is_path v0 (pv@[(v,w',v')]) v'" by simp
from PW[OF P] have
"path_weight' (res v') \<le> Num (path_weight (pv@[(v,w',v')]))"
by auto
finally show False by (simp add: infty_unbox)
qed
show "uinvar v wl res (it-{(w',v')}) (wl',res'(v'\<mapsto>the (res v)@[(v,w',v')]))"
proof -
have "(res'(v'\<mapsto>the (res v)@[(v,w',v')])) v = res' v" by simp
moreover {
fix v'' assume VMEM: "v''\<in>V"
have "(res'(v'\<mapsto>the (res v)@[(v,w',v')])) v'' \<in> least_map path_weight' (
{ res v'' } \<union> { Some (p@[(v,w,v'')]) | p w. res v = Some p
\<and> (w,v'') \<in> succ G v - (it - {(w',v')}) }
) \<and> (\<forall>p. (res'(v'\<mapsto>the (res v)@[(v,w',v')])) v'' = Some p
\<longrightarrow> is_path v0 p v'')"
proof (cases "v''=v'")
case [simp]: False
have "{ Some (p@[(v,w,v'')]) | p w. res v = Some p
\<and> (w,v'') \<in> succ G v - (it - {(w',v')}) } =
{ Some (p@[(v,w,v'')]) | p w. res v = Some p
\<and> (w,v'') \<in> succ G v - it }"
by auto
with INV VMEM show ?thesis unfolding uinvar_def
by simp
next
case [simp]: True
have EQ: "{ res v'' } \<union> { Some (p@[(v,w,v'')]) | p w. res v = Some p
\<and> (w,v'') \<in> succ G v - (it - {(w',v')}) } =
insert (Some (pv@[(v,w',v')])) (
{ res v'' } \<union> { Some (p@[(v,w,v'')]) | p w. res v = Some p
\<and> (w,v'') \<in> succ G v - it })"
using MEM IT_SS
by auto
show ?thesis
apply (subst EQ)
apply simp
apply (rule least_map_insert_min)
apply (rule ballI)
proof -
fix r'
assume A:
"r' \<in> insert (res v')
{Some (pv @ [(v, w, v')]) |w. (w, v') \<in> succ G v \<and> (w, v') \<notin> it}"
from LESS have
"path_weight' (Some (pv @ [(v, w', v')])) < path_weight' (res' v')"
by (auto simp: infty_unbox)
also from INV[unfolded uinvar_def] have
"res' v' \<in> least_map path_weight' (
insert (res v')
{Some (pv @ [(v, w, v')]) |w. (w, v') \<in> succ G v \<and> (w, v') \<notin> it}
)"
by auto
with A have "path_weight' (res' v') \<le> path_weight' r'"
by (auto dest: least_map_leD)
finally show
"path_weight' (Some (pv @ [(v, w', v')])) \<le> path_weight' r'"
by simp
qed
qed
}
ultimately show ?thesis
unfolding uinvar_def Let_def
by auto
qed
next
fix wl res it w' v' wl' res'
assume INV: "uinvar v wl res it (wl',res')"
and NLESS: "\<not> path_weight' (res v) + Num w' < path_weight' (res' v')"
and IN_IT: "(w',v')\<in>it"
and IT_SS: "it \<subseteq> succ G v"
from IN_IT IT_SS have [simp, intro!]: "(w',v')\<in>succ G v" by auto
hence [simp,intro!]: "v'\<in>V" using succ_subset
by auto
show "uinvar v wl res (it - {(w',v')}) (wl',res')"
proof (cases "res v")
case [simp]: None
from INV show ?thesis
unfolding uinvar_def by auto
next
case [simp]: (Some p)
{
fix v''
assume [simp, intro!]: "v''\<in>V"
have "res' v'' \<in> least_map path_weight' (
{ res v'' } \<union> { Some (p@[(v,w,v'')]) | p w. res v = Some p
\<and> (w,v'') \<in> succ G v - (it - {(w',v')}) }
)" (is "_ \<in> least_map path_weight' ?S")
proof (cases "v''=v'")
case False with INV show ?thesis
unfolding uinvar_def by auto
next
case [simp]: True
have EQ: "?S = insert (Some (p@[(v,w',v')])) (
{ res v' } \<union> { Some (p@[(v,w,v'')]) | p w. res v = Some p
\<and> (w,v'') \<in> succ G v - it }
)"
by auto
from NLESS have
"path_weight' (res' v') \<le> path_weight' (Some (p@[(v,w',v')]))"
by (auto simp: infty_unbox)
thus ?thesis
apply (subst EQ)
apply (rule least_map_insert_nmin)
using INV unfolding uinvar_def apply auto []
apply simp
done
qed
} with INV
show ?thesis
unfolding uinvar_def by auto
qed
next
fix wl res \<sigma>'
assume "uinvar v wl res {} \<sigma>'"
thus "update_spec v (wl,res) \<sigma>'"
unfolding uinvar_def
apply (cases \<sigma>')
apply (auto intro: update_spec.intros simp: succ_def)
done
next
show "finite (succ G v)" by simp
qed
text \<open>We integrate the new update function into the main algorithm:\<close>
definition dijkstra' where
"dijkstra' \<equiv> do {
\<sigma>0 \<leftarrow> dinit;
(_,res) \<leftarrow> WHILE\<^sub>T\<^bsup>dinvar\<^esup> (\<lambda>(wl,_). wl\<noteq>{})
(\<lambda>\<sigma>. do {(v,\<sigma>') \<leftarrow> pop_min \<sigma>; update' v \<sigma>'})
\<sigma>0;
RETURN res
}"
lemma dijkstra'_refines: "dijkstra' \<le> \<Down>Id dijkstra"
proof -
note [refine] = update'_refines
have [refine]: "\<And>\<sigma> \<sigma>'. \<sigma>=\<sigma>' \<Longrightarrow> pop_min \<sigma> \<le> \<Down>Id (pop_min \<sigma>')" by simp
show ?thesis
unfolding dijkstra_def dijkstra'_def
apply (refine_rcg)
apply simp_all
done
qed
end
subsection \<open>Refinement to Cached Weights\<close>
text \<open>
Next, we refine the data types of the workset and the result map.
The workset becomes a map from nodes to their current weights.
The result map stores, in addition to the shortest path, also the
weight of the shortest path. Moreover, we store the shortest paths
in reversed order, which makes appending new edges more effcient.
These refinements allow to implement the workset as a priority queue,
and save recomputation of the path weights in the inner loop of the
algorithm.
\<close>
type_synonym ('V,'W) mwl = "('V \<rightharpoonup> 'W infty)"
type_synonym ('V,'W) mres = "('V \<rightharpoonup> (('V,'W) path \<times> 'W))"
type_synonym ('V,'W) mstate = "('V,'W) mwl \<times> ('V,'W) mres"
text \<open>
Map a path with cached weight to one without cached weight.
\<close>
fun mpath' :: "(('V,'W) path \<times> 'W) option \<rightharpoonup> ('V,'W) path" where
"mpath' None = None" |
"mpath' (Some (p,w)) = Some p"
fun mpath_weight' :: "(('V,'W) path \<times> 'W) option \<Rightarrow> ('W::weight) infty" where
"mpath_weight' None = top" |
"mpath_weight' (Some (p,w)) = Num w"
context Dijkstra
begin
definition \<alpha>w::"('V,'W) mwl \<Rightarrow> 'V set" where "\<alpha>w \<equiv> dom"
definition \<alpha>r::"('V,'W) mres \<Rightarrow> 'V \<rightharpoonup> ('V,'W) path" where
"\<alpha>r \<equiv> \<lambda>res v. case res v of None \<Rightarrow> None | Some (p,w) \<Rightarrow> Some (rev p)"
definition \<alpha>s:: "('V,'W) mstate \<Rightarrow> ('V,'W) state" where
"\<alpha>s \<equiv> map_prod \<alpha>w \<alpha>r"
text \<open>Additional invariants for the new state. They guarantee that
the cached weights are consistent.\<close>
definition res_invarm :: "('V \<rightharpoonup> (('V,'W) path\<times>'W)) \<Rightarrow> bool" where
"res_invarm res \<equiv> (\<forall>v. case res v of
None \<Rightarrow> True |
Some (p,w) \<Rightarrow> w = path_weight (rev p))"
definition dinvarm :: "('V,'W) mstate \<Rightarrow> bool" where
"dinvarm \<sigma> \<equiv> let (wl,res) = \<sigma> in
(\<forall>v\<in>dom wl. the (wl v) = mpath_weight' (res v)) \<and> res_invarm res
"
lemma mpath_weight'_correct: "\<lbrakk>dinvarm (wl,res)\<rbrakk> \<Longrightarrow>
mpath_weight' (res v) = path_weight' (\<alpha>r res v)
"
unfolding dinvarm_def res_invarm_def \<alpha>r_def
by (auto split: option.split option.split_asm)
lemma mpath'_correct: "\<lbrakk>dinvarm (wl,res)\<rbrakk> \<Longrightarrow>
mpath' (res v) = map_option rev (\<alpha>r res v)"
unfolding dinvarm_def \<alpha>r_def
by (auto split: option.split option.split_asm)
lemma wl_weight_correct:
assumes INV: "dinvarm (wl,res)"
assumes WLV: "wl v = Some w"
shows "path_weight' (\<alpha>r res v) = w"
proof -
from INV WLV have "w = mpath_weight' (res v)"
unfolding dinvarm_def by force
also from mpath_weight'_correct[OF INV] have
"\<dots> = path_weight' (\<alpha>r res v)" .
finally show ?thesis by simp
qed
text \<open>The initial state is constructed using an iterator:\<close>
definition mdinit :: "('V,'W) mstate nres" where
"mdinit \<equiv> do {
wl \<leftarrow> FOREACH V (\<lambda>v wl. RETURN (wl(v\<mapsto>Infty))) Map.empty;
RETURN (wl(v0\<mapsto>Num 0),[v0 \<mapsto> ([],0)])
}"
lemma mdinit_refines: "mdinit \<le> \<Down>(build_rel \<alpha>s dinvarm) dinit"
unfolding mdinit_def dinit_def
apply (rule build_rel_SPEC)
apply (intro FOREACH_rule[where I="\<lambda>it wl. (\<forall>v\<in>V-it. wl v = Some Infty) \<and>
dom wl = V-it"]
refine_vcg)
apply (auto
simp: \<alpha>s_def \<alpha>w_def \<alpha>r_def dinvarm_def res_invarm_def infty_unbox
split: if_split_asm
)
done
text \<open>The new pop function:\<close>
definition
mpop_min :: "('V,'W) mstate \<Rightarrow> ('V \<times> 'W infty \<times> ('V,'W) mstate) nres"
where
"mpop_min \<sigma> \<equiv> do {
let (wl,res) = \<sigma>;
(v,w,wl')\<leftarrow>prio_pop_min wl;
RETURN (v,w,(wl',res))
}"
lemma mpop_min_refines:
"\<lbrakk> (\<sigma>,\<sigma>') \<in> build_rel \<alpha>s dinvarm \<rbrakk> \<Longrightarrow>
mpop_min \<sigma> \<le>
\<Down>(build_rel
(\<lambda>(v,w,\<sigma>). (v,\<alpha>s \<sigma>))
(\<lambda>(v,w,\<sigma>). dinvarm \<sigma> \<and> w = mpath_weight' (snd \<sigma> v)))
(pop_min \<sigma>')"
\<comment> \<open>The two algorithms are structurally different, so we use the
nofail/inres method to prove refinement.\<close>
unfolding mpop_min_def pop_min_def prio_pop_min_def
apply (rule pw_ref_I)
apply rule
apply (auto simp add: refine_pw_simps \<alpha>s_def \<alpha>w_def refine_rel_defs
split: prod.split prod.split_asm)
apply (auto simp: dinvarm_def) []
apply (auto simp: mpath_weight'_correct wl_weight_correct) []
apply (auto
simp: wl_weight_correct
intro!: least_map.intros
) []
apply (metis restrict_map_eq(2))
done
text \<open>The new update function:\<close>
definition "uinvarm v wl res it \<sigma> \<equiv>
uinvar v wl res it (\<alpha>s \<sigma>) \<and> dinvarm \<sigma>"
definition mupdate :: "'V \<Rightarrow> 'W infty \<Rightarrow> ('V,'W) mstate \<Rightarrow> ('V,'W) mstate nres"
where
"mupdate v wv \<sigma> \<equiv> do {
ASSERT (update_pre v (\<alpha>s \<sigma>) \<and> wv=mpath_weight' (snd \<sigma> v));
let (wl,res) = \<sigma>;
let pv = mpath' (res v);
FOREACH\<^bsup>uinvarm v (\<alpha>w wl) (\<alpha>r res)\<^esup> (succ G v) (\<lambda>(w',v') (wl,res).
if (wv + Num w' < mpath_weight' (res v')) then do {
ASSERT (v'\<in>dom wl \<and> pv \<noteq> None);
ASSERT (wv \<noteq> Infty);
RETURN (wl(v'\<mapsto>wv + Num w'),
res(v' \<mapsto> ((v,w',v')#the pv,val wv + w') ))
} else RETURN (wl,res)
) (wl,res)
}"
lemma mupdate_refines:
assumes SREF: "(\<sigma>,\<sigma>')\<in>build_rel \<alpha>s dinvarm"
assumes WV: "wv = mpath_weight' (snd \<sigma> v)"
assumes VV': "v'=v"
shows "mupdate v wv \<sigma> \<le> \<Down>(build_rel \<alpha>s dinvarm) (update' v' \<sigma>')"
proof (simp only: VV')
{
txt \<open>Show that IF-condition is a refinement:\<close>
fix wl res wl' res' it w' v'
assume "uinvarm v (\<alpha>w wl) (\<alpha>r res) it (wl',res')"
and "dinvarm (wl,res)"
hence "mpath_weight' (res v) + Num w' < mpath_weight' (res' v') \<longleftrightarrow>
path_weight' (\<alpha>r res v) + Num w' < path_weight' (\<alpha>r res' v')"
unfolding uinvarm_def
by (auto simp add: mpath_weight'_correct)
} note COND_refine=this
{
txt \<open>THEN-case:\<close>
fix wl res wl' res' it w' v'
assume UINV: "uinvarm v (\<alpha>w wl) (\<alpha>r res) it (wl',res')"
and DINV: "dinvarm (wl,res)"
and "mpath_weight' (res v) + Num w' < mpath_weight' (res' v')"
and "path_weight' (\<alpha>r res v) + Num w' < path_weight' (\<alpha>r res' v')"
and V'MEM: "v'\<in>\<alpha>w wl'"
and NN: "\<alpha>r res v \<noteq> None"
from NN obtain pv wv where
ARV: "\<alpha>r res v = Some (rev pv)" and
RV: "res v = Some (pv,wv)"
unfolding \<alpha>r_def by (auto split: option.split_asm)
with DINV have [simp]: "wv = path_weight (rev pv)"
unfolding dinvarm_def res_invarm_def by (auto split: option.split_asm)
note [simp] = ARV RV
from V'MEM NN have "v'\<in>dom wl'" (is "?G1")
and "mpath' (res v) \<noteq> None" (is "?G2")
unfolding \<alpha>w_def \<alpha>r_def by (auto split: option.split_asm)
hence "\<And>x. \<alpha>w wl' = \<alpha>w (wl'(v'\<mapsto>x))" by (auto simp: \<alpha>w_def)
moreover have "mpath' (res v) = map_option rev (\<alpha>r res v)" using DINV
by (simp add: mpath'_correct)
ultimately have
"\<alpha>w wl' = \<alpha>w (wl'(v' \<mapsto> mpath_weight' (res v) + Num w'))
\<and> (\<alpha>r res')(v' \<mapsto> the (\<alpha>r res v)@[(v, w', v')])
= \<alpha>r (res'(v' \<mapsto> ((v, w', v')#the (mpath' (res v)),
val (mpath_weight' (res v)) + w')))" (is ?G3)
by (auto simp add: \<alpha>r_def intro!: ext)
have
"(dinvarm (wl'(v'\<mapsto>mpath_weight' (res v) + Num w'),
res'(v' \<mapsto> ((v,w',v') # the (mpath' (res v)),
val (mpath_weight' (res v)) + w'
))))" (is ?G4)
using UINV unfolding uinvarm_def dinvarm_def res_invarm_def
by (auto simp: infty_unbox split: option.split option.split_asm)
note \<open>?G1\<close> \<open>?G2\<close> \<open>?G3\<close> \<open>?G4\<close>
} note THEN_refine=this
note [refine2] = inj_on_id
note [simp] = refine_rel_defs
show "mupdate v wv \<sigma> \<le> \<Down>(build_rel \<alpha>s dinvarm) (update' v \<sigma>')"
using SREF WV
unfolding mupdate_def update'_def
apply -
apply (refine_rcg)
apply simp_all [3]
apply (simp add: \<alpha>s_def uinvarm_def)
apply (simp_all add: \<alpha>s_def COND_refine THEN_refine(1-2)) [3]
apply (rule ccontr,simp)
using THEN_refine(3,4)
apply (auto simp: \<alpha>s_def) []
txt \<open>The ELSE-case is trivial:\<close>
apply simp
done
qed
text \<open>Finally, we assemble the refined algorithm:\<close>
definition mdijkstra where
"mdijkstra \<equiv> do {
\<sigma>0 \<leftarrow> mdinit;
(_,res) \<leftarrow> WHILE\<^sub>T\<^bsup>dinvarm\<^esup> (\<lambda>(wl,_). dom wl\<noteq>{})
(\<lambda>\<sigma>. do { (v,wv,\<sigma>') \<leftarrow> mpop_min \<sigma>; mupdate v wv \<sigma>' } )
\<sigma>0;
RETURN res
}"
lemma mdijkstra_refines: "mdijkstra \<le> \<Down>(build_rel \<alpha>r res_invarm) dijkstra'"
proof -
note [refine] = mdinit_refines mpop_min_refines mupdate_refines
show ?thesis
unfolding mdijkstra_def dijkstra'_def
apply (refine_rcg)
apply (simp_all split: prod.split
add: \<alpha>s_def \<alpha>w_def dinvarm_def refine_rel_defs)
done
qed
end
end
diff --git a/thys/Finger-Trees/FingerTree.thy b/thys/Finger-Trees/FingerTree.thy
--- a/thys/Finger-Trees/FingerTree.thy
+++ b/thys/Finger-Trees/FingerTree.thy
@@ -1,2651 +1,2653 @@
section "2-3 Finger Trees"
theory FingerTree
imports Main
begin
text \<open>
We implement and prove correct 2-3 finger trees as described by Ralf Hinze
and Ross Paterson\cite{HiPa06}.
\<close>
text \<open>
This theory is organized as follows:
Section~\ref{sec:datatype} contains the finger-tree datatype, its invariant
and its abstraction function to lists.
The Section~\ref{sec:operations} contains the operations
on finger trees and their correctness lemmas.
Section~\ref{sec:hide_invar} contains a finger tree datatype with implicit
invariant, and, finally, Section~\ref{sec:doc} contains a documentation
of the implemented operations.
\<close>
text_raw \<open>\paragraph{Technical Issues}\<close>
text \<open>
As Isabelle lacks proper support of namespaces, we
try to simulate namespaces by locales.
The problem is, that we define lots of internal functions that
should not be exposed to the user at all.
Moreover, we define some functions with names equal to names
from Isabelle's standard library. These names make perfect sense
in the context of FingerTrees, however, they shall not be exposed
to anyone using this theory indirectly, hiding the standard library
names there.
Our approach puts all functions and lemmas inside the locale
{\em FingerTree\_loc},
and then interprets this locale with the prefix {\em FingerTree}.
This makes all definitions visible outside the locale, with
qualified names. Inside the locale, however, one can use unqualified names.
\<close>
subsection "Datatype definition"
text_raw\<open>\label{sec:datatype}\<close>
locale FingerTreeStruc_loc
text \<open>
Nodes: Non empty 2-3 trees, with all elements stored within the leafs plus a
cached annotation
\<close>
datatype ('e,'a) Node = Tip 'e 'a |
Node2 'a "('e,'a) Node" "('e,'a) Node" |
Node3 'a "('e,'a) Node" "('e,'a) Node" "('e,'a) Node"
text \<open>Digit: one to four ordered Nodes\<close>
datatype ('e,'a) Digit = One "('e,'a) Node" |
Two "('e,'a) Node" "('e,'a) Node" |
Three "('e,'a) Node" "('e,'a) Node" "('e,'a) Node" |
Four "('e,'a) Node" "('e,'a) Node" "('e,'a) Node" "('e,'a) Node"
text \<open>FingerTreeStruc:
The empty tree, a single node or some nodes and a deeper tree\<close>
datatype ('e, 'a) FingerTreeStruc =
Empty |
Single "('e,'a) Node" |
Deep 'a "('e,'a) Digit" "('e,'a) FingerTreeStruc" "('e,'a) Digit"
subsubsection "Invariant"
context FingerTreeStruc_loc
begin
text_raw \<open>\paragraph{Auxiliary functions}\ \\\<close>
text \<open>Readout the cached annotation of a node\<close>
primrec gmn :: "('e,'a::monoid_add) Node \<Rightarrow> 'a" where
"gmn (Tip e a) = a" |
"gmn (Node2 a _ _) = a" |
"gmn (Node3 a _ _ _) = a"
text \<open>The annotation of a digit is computed on the fly\<close>
primrec gmd :: "('e,'a::monoid_add) Digit \<Rightarrow> 'a" where
"gmd (One a) = gmn a" |
"gmd (Two a b) = (gmn a) + (gmn b)"|
"gmd (Three a b c) = (gmn a) + (gmn b) + (gmn c)"|
"gmd (Four a b c d) = (gmn a) + (gmn b) + (gmn c) + (gmn d)"
text \<open>Readout the cached annotation of a finger tree\<close>
primrec gmft :: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> 'a" where
"gmft Empty = 0" |
"gmft (Single nd) = gmn nd" |
"gmft (Deep a _ _ _) = a"
text \<open>Depth and cached annotations have to be correct\<close>
fun is_leveln_node :: "nat \<Rightarrow> ('e,'a) Node \<Rightarrow> bool" where
"is_leveln_node 0 (Tip _ _) \<longleftrightarrow> True" |
"is_leveln_node (Suc n) (Node2 _ n1 n2) \<longleftrightarrow>
is_leveln_node n n1 \<and> is_leveln_node n n2" |
"is_leveln_node (Suc n) (Node3 _ n1 n2 n3) \<longleftrightarrow>
is_leveln_node n n1 \<and> is_leveln_node n n2 \<and> is_leveln_node n n3" |
"is_leveln_node _ _ \<longleftrightarrow> False"
primrec is_leveln_digit :: "nat \<Rightarrow> ('e,'a) Digit \<Rightarrow> bool" where
"is_leveln_digit n (One n1) \<longleftrightarrow> is_leveln_node n n1" |
"is_leveln_digit n (Two n1 n2) \<longleftrightarrow> is_leveln_node n n1 \<and>
is_leveln_node n n2" |
"is_leveln_digit n (Three n1 n2 n3) \<longleftrightarrow> is_leveln_node n n1 \<and>
is_leveln_node n n2 \<and> is_leveln_node n n3" |
"is_leveln_digit n (Four n1 n2 n3 n4) \<longleftrightarrow> is_leveln_node n n1 \<and>
is_leveln_node n n2 \<and> is_leveln_node n n3 \<and> is_leveln_node n n4"
primrec is_leveln_ftree :: "nat \<Rightarrow> ('e,'a) FingerTreeStruc \<Rightarrow> bool" where
"is_leveln_ftree n Empty \<longleftrightarrow> True" |
"is_leveln_ftree n (Single nd) \<longleftrightarrow> is_leveln_node n nd" |
"is_leveln_ftree n (Deep _ l t r) \<longleftrightarrow> is_leveln_digit n l \<and>
is_leveln_digit n r \<and> is_leveln_ftree (Suc n) t"
primrec is_measured_node :: "('e,'a::monoid_add) Node \<Rightarrow> bool" where
"is_measured_node (Tip _ _) \<longleftrightarrow> True" |
"is_measured_node (Node2 a n1 n2) \<longleftrightarrow> ((is_measured_node n1) \<and>
(is_measured_node n2)) \<and> (a = (gmn n1) + (gmn n2))" |
"is_measured_node (Node3 a n1 n2 n3) \<longleftrightarrow> ((is_measured_node n1) \<and>
(is_measured_node n2) \<and> (is_measured_node n3)) \<and>
(a = (gmn n1) + (gmn n2) + (gmn n3))"
primrec is_measured_digit :: "('e,'a::monoid_add) Digit \<Rightarrow> bool" where
"is_measured_digit (One a) = is_measured_node a" |
"is_measured_digit (Two a b) =
((is_measured_node a) \<and> (is_measured_node b))"|
"is_measured_digit (Three a b c) =
((is_measured_node a) \<and> (is_measured_node b) \<and> (is_measured_node c))"|
"is_measured_digit (Four a b c d) = ((is_measured_node a) \<and>
(is_measured_node b) \<and> (is_measured_node c) \<and> (is_measured_node d))"
primrec is_measured_ftree :: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> bool" where
"is_measured_ftree Empty \<longleftrightarrow> True" |
"is_measured_ftree (Single n1) \<longleftrightarrow> (is_measured_node n1)" |
"is_measured_ftree (Deep a l m r) \<longleftrightarrow> ((is_measured_digit l) \<and>
(is_measured_ftree m) \<and> (is_measured_digit r)) \<and>
(a = ((gmd l) + (gmft m) + (gmd r)))"
text "Structural invariant for finger trees"
definition "ft_invar t == is_leveln_ftree 0 t \<and> is_measured_ftree t"
subsubsection "Abstraction to Lists"
primrec nodeToList :: "('e,'a) Node \<Rightarrow> ('e \<times> 'a) list" where
"nodeToList (Tip e a) = [(e,a)]"|
"nodeToList (Node2 _ a b) = (nodeToList a) @ (nodeToList b)"|
"nodeToList (Node3 _ a b c)
= (nodeToList a) @ (nodeToList b) @ (nodeToList c)"
primrec digitToList :: "('e,'a) Digit \<Rightarrow> ('e \<times> 'a) list" where
"digitToList (One a) = nodeToList a"|
"digitToList (Two a b) = (nodeToList a) @ (nodeToList b)"|
"digitToList (Three a b c)
= (nodeToList a) @ (nodeToList b) @ (nodeToList c)"|
"digitToList (Four a b c d)
= (nodeToList a) @ (nodeToList b) @ (nodeToList c) @ (nodeToList d)"
text "List representation of a finger tree"
primrec toList :: "('e ,'a) FingerTreeStruc \<Rightarrow> ('e \<times> 'a) list" where
"toList Empty = []"|
"toList (Single a) = nodeToList a"|
"toList (Deep _ pr m sf) = (digitToList pr) @ (toList m) @ (digitToList sf)"
lemma nodeToList_empty: "nodeToList nd \<noteq> Nil"
by (induct nd) auto
lemma digitToList_empty: "digitToList d \<noteq> Nil"
by (cases d, auto simp add: nodeToList_empty)
text \<open>Auxiliary lemmas\<close>
lemma gmn_correct:
assumes "is_measured_node nd"
shows "gmn nd = sum_list (map snd (nodeToList nd))"
by (insert assms, induct nd) (auto simp add: add.assoc)
lemma gmd_correct:
assumes "is_measured_digit d"
shows "gmd d = sum_list (map snd (digitToList d))"
by (insert assms, cases d, auto simp add: gmn_correct add.assoc)
lemma gmft_correct: "is_measured_ftree t
\<Longrightarrow> (gmft t) = sum_list (map snd (toList t))"
by (induct t, auto simp add: ft_invar_def gmd_correct gmn_correct add.assoc)
lemma gmft_correct2: "ft_invar t \<Longrightarrow> (gmft t) = sum_list (map snd (toList t))"
by (simp only: ft_invar_def gmft_correct)
subsection \<open>Operations\<close>
text_raw\<open>\label{sec:operations}\<close>
subsubsection \<open>Empty tree\<close>
lemma Empty_correct[simp]:
"toList Empty = []"
"ft_invar Empty"
by (simp_all add: ft_invar_def)
text \<open>Exactly the empty finger tree represents the empty list\<close>
lemma toList_empty: "toList t = [] \<longleftrightarrow> t = Empty"
by (induct t, auto simp add: nodeToList_empty digitToList_empty)
subsubsection \<open>Annotation\<close>
text "Sum of annotations of all elements of a finger tree"
definition annot :: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> 'a"
where "annot t = gmft t"
lemma annot_correct:
"ft_invar t \<Longrightarrow> annot t = sum_list (map snd (toList t))"
using gmft_correct
unfolding annot_def
by (simp add: gmft_correct2)
subsubsection \<open>Appending\<close>
text \<open>Auxiliary functions to fill in the annotations\<close>
definition deep:: "('e,'a::monoid_add) Digit \<Rightarrow> ('e,'a) FingerTreeStruc
\<Rightarrow> ('e,'a) Digit \<Rightarrow> ('e, 'a) FingerTreeStruc" where
"deep pr m sf = Deep ((gmd pr) + (gmft m) + (gmd sf)) pr m sf"
definition node2 where
"node2 nd1 nd2 = Node2 ((gmn nd1)+(gmn nd2)) nd1 nd2"
definition node3 where
"node3 nd1 nd2 nd3 = Node3 ((gmn nd1)+(gmn nd2)+(gmn nd3)) nd1 nd2 nd3"
text "Append a node at the left end"
fun nlcons :: "('e,'a::monoid_add) Node \<Rightarrow> ('e,'a) FingerTreeStruc
\<Rightarrow> ('e,'a) FingerTreeStruc"
where
\<comment> \<open>Recursively we append a node, if the digit is full we push down a node3\<close>
"nlcons a Empty = Single a" |
"nlcons a (Single b) = deep (One a) Empty (One b)" |
"nlcons a (Deep _ (One b) m sf) = deep (Two a b) m sf" |
"nlcons a (Deep _ (Two b c) m sf) = deep (Three a b c) m sf" |
"nlcons a (Deep _ (Three b c d) m sf) = deep (Four a b c d) m sf" |
"nlcons a (Deep _ (Four b c d e) m sf)
= deep (Two a b) (nlcons (node3 c d e) m) sf"
text "Append a node at the right end"
fun nrcons :: "('e,'a::monoid_add) FingerTreeStruc
\<Rightarrow> ('e,'a) Node \<Rightarrow> ('e,'a) FingerTreeStruc" where
\<comment> \<open>Recursively we append a node, if the digit is full we push down a node3\<close>
"nrcons Empty a = Single a" |
"nrcons (Single b) a = deep (One b) Empty (One a)" |
"nrcons (Deep _ pr m (One b)) a = deep pr m (Two b a)"|
"nrcons (Deep _ pr m (Two b c)) a = deep pr m (Three b c a)" |
"nrcons (Deep _ pr m (Three b c d)) a = deep pr m (Four b c d a)" |
"nrcons (Deep _ pr m (Four b c d e)) a
= deep pr (nrcons m (node3 b c d)) (Two e a)"
lemma nlcons_invlevel: "\<lbrakk>is_leveln_ftree n t; is_leveln_node n nd\<rbrakk>
\<Longrightarrow> is_leveln_ftree n (nlcons nd t)"
by (induct t arbitrary: n nd rule: nlcons.induct)
(auto simp add: deep_def node3_def)
lemma nlcons_invmeas: "\<lbrakk>is_measured_ftree t; is_measured_node nd\<rbrakk>
\<Longrightarrow> is_measured_ftree (nlcons nd t)"
by (induct t arbitrary: nd rule: nlcons.induct)
(auto simp add: deep_def node3_def)
lemmas nlcons_inv = nlcons_invlevel nlcons_invmeas
lemma nlcons_list: "toList (nlcons a t) = (nodeToList a) @ (toList t)"
apply (induct t arbitrary: a rule: nlcons.induct)
apply (auto simp add: deep_def toList_def node3_def)
done
lemma nrcons_invlevel: "\<lbrakk>is_leveln_ftree n t; is_leveln_node n nd\<rbrakk>
\<Longrightarrow> is_leveln_ftree n (nrcons t nd)"
apply (induct t nd arbitrary: nd n rule:nrcons.induct)
apply(auto simp add: deep_def node3_def)
done
lemma nrcons_invmeas: "\<lbrakk>is_measured_ftree t; is_measured_node nd\<rbrakk>
\<Longrightarrow> is_measured_ftree (nrcons t nd)"
apply (induct t nd arbitrary: nd rule:nrcons.induct)
apply(auto simp add: deep_def node3_def)
done
lemmas nrcons_inv = nrcons_invlevel nrcons_invmeas
lemma nrcons_list: "toList (nrcons t a) = (toList t) @ (nodeToList a)"
apply (induct t a arbitrary: a rule: nrcons.induct)
apply (auto simp add: deep_def toList_def node3_def)
done
text "Append an element at the left end"
definition lcons :: "('e \<times> 'a::monoid_add)
\<Rightarrow> ('e,'a) FingerTreeStruc \<Rightarrow> ('e,'a) FingerTreeStruc" (infixr "\<lhd>" 65) where
"a \<lhd> t = nlcons (Tip (fst a) (snd a)) t"
lemma lcons_correct:
assumes "ft_invar t"
shows "ft_invar (a \<lhd> t)" and "toList (a \<lhd> t) = a # (toList t)"
using assms
unfolding ft_invar_def
by (simp_all add: lcons_def nlcons_list nlcons_invlevel nlcons_invmeas)
lemma lcons_inv:"ft_invar t \<Longrightarrow> ft_invar (a \<lhd> t)"
by (rule lcons_correct)
lemma lcons_list[simp]: "toList (a \<lhd> t) = a # (toList t)"
by (simp add: lcons_def nlcons_list)
text "Append an element at the right end"
definition rcons
:: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> ('e \<times> 'a) \<Rightarrow> ('e,'a) FingerTreeStruc"
(infixl "\<rhd>" 65) where
"t \<rhd> a = nrcons t (Tip (fst a) (snd a))"
lemma rcons_correct:
assumes "ft_invar t"
shows "ft_invar (t \<rhd> a)" and "toList (t \<rhd> a) = (toList t) @ [a]"
using assms
by (auto simp add: nrcons_inv ft_invar_def rcons_def nrcons_list)
lemma rcons_inv:"ft_invar t \<Longrightarrow> ft_invar (t \<rhd> a)"
by (rule rcons_correct)
lemma rcons_list[simp]: "toList (t \<rhd> a) = (toList t) @ [a]"
by(auto simp add: nrcons_list rcons_def)
subsubsection \<open>Convert list to tree\<close>
primrec toTree :: "('e \<times> 'a::monoid_add) list \<Rightarrow> ('e,'a) FingerTreeStruc" where
"toTree [] = Empty"|
"toTree (a#xs) = a \<lhd> (toTree xs)"
lemma toTree_correct[simp]:
"ft_invar (toTree l)"
"toList (toTree l) = l"
apply (induct l)
apply (simp add: ft_invar_def)
apply simp
apply (simp add: toTree_def lcons_list lcons_inv)
apply (simp add: toTree_def lcons_list lcons_inv)
done
text \<open>
Note that this lemma is a completeness statement of our implementation,
as it can be read as:
,,All lists of elements have a valid representation as a finger tree.''
\<close>
subsubsection \<open>Detaching leftmost/rightmost element\<close>
primrec digitToTree :: "('e,'a::monoid_add) Digit \<Rightarrow> ('e,'a) FingerTreeStruc"
where
"digitToTree (One a) = Single a"|
"digitToTree (Two a b) = deep (One a) Empty (One b)"|
"digitToTree (Three a b c) = deep (Two a b) Empty (One c)"|
"digitToTree (Four a b c d) = deep (Two a b) Empty (Two c d)"
primrec nodeToDigit :: "('e,'a) Node \<Rightarrow> ('e,'a) Digit" where
"nodeToDigit (Tip e a) = One (Tip e a)"|
"nodeToDigit (Node2 _ a b) = Two a b"|
"nodeToDigit (Node3 _ a b c) = Three a b c"
fun nlistToDigit :: "('e,'a) Node list \<Rightarrow> ('e,'a) Digit" where
"nlistToDigit [a] = One a" |
"nlistToDigit [a,b] = Two a b" |
"nlistToDigit [a,b,c] = Three a b c" |
"nlistToDigit [a,b,c,d] = Four a b c d"
primrec digitToNlist :: "('e,'a) Digit \<Rightarrow> ('e,'a) Node list" where
"digitToNlist (One a) = [a]" |
"digitToNlist (Two a b) = [a,b] " |
"digitToNlist (Three a b c) = [a,b,c]" |
"digitToNlist (Four a b c d) = [a,b,c,d]"
text \<open>Auxiliary function to unwrap a Node element\<close>
primrec n_unwrap:: "('e,'a) Node \<Rightarrow> ('e \<times> 'a)" where
"n_unwrap (Tip e a) = (e,a)"|
"n_unwrap (Node2 _ a b) = undefined"|
"n_unwrap (Node3 _ a b c) = undefined"
type_synonym ('e,'a) ViewnRes = "(('e,'a) Node \<times> ('e,'a) FingerTreeStruc) option"
lemma viewnres_cases:
fixes r :: "('e,'a) ViewnRes"
obtains (Nil) "r=None" |
(Cons) a t where "r=Some (a,t)"
by (cases r) auto
lemma viewnres_split:
"P (case_option f1 (case_prod f2) x) =
((x = None \<longrightarrow> P f1) \<and> (\<forall>a b. x = Some (a,b) \<longrightarrow> P (f2 a b)))"
by (auto split: option.split prod.split)
text \<open>Detach the leftmost node. Return @{const None} on empty finger tree.\<close>
fun viewLn :: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> ('e,'a) ViewnRes" where
"viewLn Empty = None"|
"viewLn (Single a) = Some (a, Empty)"|
"viewLn (Deep _ (Two a b) m sf) = Some (a, (deep (One b) m sf))"|
"viewLn (Deep _ (Three a b c) m sf) = Some (a, (deep (Two b c) m sf))"|
"viewLn (Deep _ (Four a b c d) m sf) = Some (a, (deep (Three b c d) m sf))"|
"viewLn (Deep _ (One a) m sf) =
(case viewLn m of
None \<Rightarrow> Some (a, (digitToTree sf)) |
Some (b, m2) \<Rightarrow> Some (a, (deep (nodeToDigit b) m2 sf)))"
text \<open>Detach the rightmost node. Return @{const None} on empty finger tree.\<close>
fun viewRn :: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> ('e,'a) ViewnRes" where
"viewRn Empty = None" |
"viewRn (Single a) = Some (a, Empty)" |
"viewRn (Deep _ pr m (Two a b)) = Some (b, (deep pr m (One a)))" |
"viewRn (Deep _ pr m (Three a b c)) = Some (c, (deep pr m (Two a b)))" |
"viewRn (Deep _ pr m (Four a b c d)) = Some (d, (deep pr m (Three a b c)))" |
"viewRn (Deep _ pr m (One a)) =
(case viewRn m of
None \<Rightarrow> Some (a, (digitToTree pr))|
Some (b, m2) \<Rightarrow> Some (a, (deep pr m2 (nodeToDigit b))))"
(* TODO: Head, last geht auch in O(1) !!! *)
lemma
digitToTree_inv: "is_leveln_digit n d \<Longrightarrow> is_leveln_ftree n (digitToTree d)"
"is_measured_digit d \<Longrightarrow> is_measured_ftree (digitToTree d)"
apply (cases d,auto simp add: deep_def)
apply (cases d,auto simp add: deep_def)
done
lemma digitToTree_list: "toList (digitToTree d) = digitToList d"
by (cases d) (auto simp add: deep_def)
lemma nodeToDigit_inv:
"is_leveln_node (Suc n) nd \<Longrightarrow> is_leveln_digit n (nodeToDigit nd) "
"is_measured_node nd \<Longrightarrow> is_measured_digit (nodeToDigit nd)"
by (cases nd, auto) (cases nd, auto)
lemma nodeToDigit_list: "digitToList (nodeToDigit nd) = nodeToList nd"
by (cases nd,auto)
lemma viewLn_empty: "t \<noteq> Empty \<longleftrightarrow> (viewLn t) \<noteq> None"
proof (cases t)
case Empty thus ?thesis by simp
next
case (Single Node) thus ?thesis by simp
next
case (Deep a l x r) thus ?thesis
apply(auto)
apply(case_tac l)
apply(auto)
apply(cases "viewLn x")
apply(auto)
done
qed
lemma viewLn_inv: "\<lbrakk>
is_measured_ftree t; is_leveln_ftree n t; viewLn t = Some (nd, s)
\<rbrakk> \<Longrightarrow> is_measured_ftree s \<and> is_measured_node nd \<and>
is_leveln_ftree n s \<and> is_leveln_node n nd"
apply(induct t arbitrary: n nd s rule: viewLn.induct)
apply(simp add: viewLn_empty)
apply(simp)
apply(auto simp add: deep_def)[1]
apply(auto simp add: deep_def)[1]
apply(auto simp add: deep_def)[1]
proof -
fix ux a m sf n nd s
assume av: "\<And>n nd s.
\<lbrakk>is_measured_ftree m; is_leveln_ftree n m; viewLn m = Some (nd, s)\<rbrakk>
\<Longrightarrow> is_measured_ftree s \<and>
is_measured_node nd \<and> is_leveln_ftree n s \<and> is_leveln_node n nd "
" is_measured_ftree (Deep ux (One a) m sf) "
"is_leveln_ftree n (Deep ux (One a) m sf)"
"viewLn (Deep ux (One a) m sf) = Some (nd, s)"
thus "is_measured_ftree s \<and>
is_measured_node nd \<and> is_leveln_ftree n s \<and> is_leveln_node n nd"
proof (cases "viewLn m" rule: viewnres_cases)
case Nil
with av(4) have v1: "nd = a" "s = digitToTree sf"
by auto
from v1 av(2,3) show "is_measured_ftree s \<and>
is_measured_node nd \<and> is_leveln_ftree n s \<and> is_leveln_node n nd"
apply(auto)
apply(auto simp add: digitToTree_inv)
done
next
case (Cons b m2)
with av(4) have v2: "nd = a" "s = (deep (nodeToDigit b) m2 sf)"
apply (auto simp add: deep_def)
done
note myiv = av(1)[of "Suc n" b m2]
from v2 av(2,3) have "is_measured_ftree m \<and> is_leveln_ftree (Suc n) m"
apply(simp)
done
hence bv: "is_measured_ftree m2 \<and>
is_measured_node b \<and> is_leveln_ftree (Suc n) m2 \<and> is_leveln_node (Suc n) b"
using myiv Cons
apply(simp)
done
with av(2,3) v2 show "is_measured_ftree s \<and>
is_measured_node nd \<and> is_leveln_ftree n s \<and> is_leveln_node n nd"
apply(auto simp add: deep_def nodeToDigit_inv)
done
qed
qed
lemma viewLn_list: " viewLn t = Some (nd, s)
\<Longrightarrow> toList t = (nodeToList nd) @ (toList s)"
+ supply [[simproc del: defined_all]]
apply(induct t arbitrary: nd s rule: viewLn.induct)
apply(simp)
apply(simp)
apply(simp)
apply(simp add: deep_def)
apply(auto simp add: toList_def)[1]
apply(simp)
apply(simp add: deep_def)
apply(auto simp add: toList_def)[1]
apply(simp)
apply(simp add: deep_def)
apply(auto simp add: toList_def)[1]
apply(simp)
subgoal premises prems for a m sf nd s
using prems
proof (cases "viewLn m" rule: viewnres_cases)
case Nil
hence av: "m = Empty" by (metis viewLn_empty)
from av prems
show "nodeToList a @ toList m @ digitToList sf = nodeToList nd @ toList s"
by (auto simp add: digitToTree_list)
next
case (Cons b m2)
with prems have bv: "nd = a" "s = (deep (nodeToDigit b) m2 sf)"
by (auto simp add: deep_def)
with Cons prems
show "nodeToList a @ toList m @ digitToList sf = nodeToList nd @ toList s"
apply(simp)
apply(simp add: deep_def)
apply(simp add: deep_def nodeToDigit_list)
done
qed
done
lemma viewRn_empty: "t \<noteq> Empty \<longleftrightarrow> (viewRn t) \<noteq> None"
proof (cases t)
case Empty thus ?thesis by simp
next
case (Single Node) thus ?thesis by simp
next
case (Deep a l x r) thus ?thesis
apply(auto)
apply(case_tac r)
apply(auto)
apply(cases "viewRn x")
apply(auto)
done
qed
lemma viewRn_inv: "\<lbrakk>
is_measured_ftree t; is_leveln_ftree n t; viewRn t = Some (nd, s)
\<rbrakk> \<Longrightarrow> is_measured_ftree s \<and> is_measured_node nd \<and>
is_leveln_ftree n s \<and> is_leveln_node n nd"
apply(induct t arbitrary: n nd s rule: viewRn.induct)
apply(simp add: viewRn_empty)
apply(simp)
apply(auto simp add: deep_def)[1]
apply(auto simp add: deep_def)[1]
apply(auto simp add: deep_def)[1]
proof -
fix ux a m "pr" n nd s
assume av: "\<And>n nd s.
\<lbrakk>is_measured_ftree m; is_leveln_ftree n m; viewRn m = Some (nd, s)\<rbrakk>
\<Longrightarrow> is_measured_ftree s \<and>
is_measured_node nd \<and> is_leveln_ftree n s \<and> is_leveln_node n nd "
" is_measured_ftree (Deep ux pr m (One a)) "
"is_leveln_ftree n (Deep ux pr m (One a))"
"viewRn (Deep ux pr m (One a)) = Some (nd, s)"
thus "is_measured_ftree s \<and>
is_measured_node nd \<and> is_leveln_ftree n s \<and> is_leveln_node n nd"
proof (cases "viewRn m" rule: viewnres_cases)
case Nil
with av(4) have v1: "nd = a" "s = digitToTree pr"
by auto
from v1 av(2,3) show "is_measured_ftree s \<and>
is_measured_node nd \<and> is_leveln_ftree n s \<and> is_leveln_node n nd"
apply(auto)
apply(auto simp add: digitToTree_inv)
done
next
case (Cons b m2)
with av(4) have v2: "nd = a" "s = (deep pr m2 (nodeToDigit b))"
apply (auto simp add: deep_def)
done
note myiv = av(1)[of "Suc n" b m2]
from v2 av(2,3) have "is_measured_ftree m \<and> is_leveln_ftree (Suc n) m"
apply(simp)
done
hence bv: "is_measured_ftree m2 \<and>
is_measured_node b \<and> is_leveln_ftree (Suc n) m2 \<and> is_leveln_node (Suc n) b"
using myiv Cons
apply(simp)
done
with av(2,3) v2 show "is_measured_ftree s \<and>
is_measured_node nd \<and> is_leveln_ftree n s \<and> is_leveln_node n nd"
apply(auto simp add: deep_def nodeToDigit_inv)
done
qed
qed
lemma viewRn_list: "viewRn t = Some (nd, s)
\<Longrightarrow> toList t = (toList s) @ (nodeToList nd)"
+ supply [[simproc del: defined_all]]
apply(induct t arbitrary: nd s rule: viewRn.induct)
apply(simp)
apply(simp)
apply(simp)
apply(simp add: deep_def)
apply(auto simp add: toList_def)[1]
apply(simp)
apply(simp add: deep_def)
apply(auto simp add: toList_def)[1]
apply(simp)
apply(simp add: deep_def)
apply(auto simp add: toList_def)[1]
apply(simp)
subgoal premises prems for pr m a nd s
proof (cases "viewRn m" rule: viewnres_cases)
case Nil
from Nil have av: "m = Empty" by (metis viewRn_empty)
from av prems
show "digitToList pr @ toList m @ nodeToList a = toList s @ nodeToList nd"
by (auto simp add: digitToTree_list)
next
case (Cons b m2)
with prems have bv: "nd = a" "s = (deep pr m2 (nodeToDigit b))"
apply(auto simp add: deep_def) done
with Cons prems
show "digitToList pr @ toList m @ nodeToList a = toList s @ nodeToList nd"
apply(simp)
apply(simp add: deep_def)
apply(simp add: deep_def nodeToDigit_list)
done
qed
done
type_synonym ('e,'a) viewres = "(('e \<times>'a) \<times> ('e,'a) FingerTreeStruc) option"
text \<open>Detach the leftmost element. Return @{const None} on empty finger tree.\<close>
definition viewL :: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> ('e,'a) viewres"
where
"viewL t = (case viewLn t of
None \<Rightarrow> None |
(Some (a, t2)) \<Rightarrow> Some ((n_unwrap a), t2))"
lemma viewL_correct:
assumes INV: "ft_invar t"
shows
"(t=Empty \<Longrightarrow> viewL t = None)"
"(t\<noteq>Empty \<Longrightarrow> (\<exists>a s. viewL t = Some (a, s) \<and> ft_invar s
\<and> toList t = a # toList s))"
proof -
assume "t=Empty" thus "viewL t = None" by (simp add: viewL_def)
next
assume NE: "t \<noteq> Empty"
from INV have INV': "is_leveln_ftree 0 t" "is_measured_ftree t"
by (simp_all add: ft_invar_def)
from NE have v1: "viewLn t \<noteq> None" by (auto simp add: viewLn_empty)
then obtain nd s where vn: "viewLn t = Some (nd, s)"
by (cases "viewLn t") (auto)
from this obtain a where v1: "viewL t = Some (a, s)"
by (auto simp add: viewL_def)
from INV' vn have
v2: "is_measured_ftree s \<and> is_leveln_ftree 0 s
\<and> is_leveln_node 0 nd \<and> is_measured_node nd"
"toList t = (nodeToList nd) @ (toList s)"
by (auto simp add: viewLn_inv[of t 0 nd s] viewLn_list[of t])
with v1 vn have v3: "nodeToList nd = [a]"
apply (auto simp add: viewL_def )
apply (induct nd)
apply (simp_all (no_asm_use))
done
with v1 v2
show "\<exists>a s. viewL t = Some (a, s) \<and> ft_invar s \<and> toList t = a # toList s"
by (auto simp add: ft_invar_def)
qed
lemma viewL_correct_empty[simp]: "viewL Empty = None"
by (simp add: viewL_def)
lemma viewL_correct_nonEmpty:
assumes "ft_invar t" "t \<noteq> Empty"
obtains a s where
"viewL t = Some (a, s)" "ft_invar s" "toList t = a # toList s"
using assms viewL_correct by blast
text \<open>Detach the rightmost element. Return @{const None} on empty finger tree.\<close>
definition viewR :: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> ('e,'a) viewres"
where
"viewR t = (case viewRn t of
None \<Rightarrow> None |
(Some (a, t2)) \<Rightarrow> Some ((n_unwrap a), t2))"
lemma viewR_correct:
assumes INV: "ft_invar t"
shows
"(t = Empty \<Longrightarrow> viewR t = None)"
"(t \<noteq> Empty \<Longrightarrow> (\<exists> a s. viewR t = Some (a, s) \<and> ft_invar s
\<and> toList t = toList s @ [a]))"
proof -
assume "t=Empty" thus "viewR t = None" by (simp add: viewR_def)
next
assume NE: "t \<noteq> Empty"
from INV have INV': "is_leveln_ftree 0 t" "is_measured_ftree t"
unfolding ft_invar_def by simp_all
from NE have v1: "viewRn t \<noteq> None" by (auto simp add: viewRn_empty)
then obtain nd s where vn: "viewRn t = Some (nd, s)"
by (cases "viewRn t") (auto)
from this obtain a where v1: "viewR t = Some (a, s)"
by (auto simp add: viewR_def)
from INV' vn have
v2: "is_measured_ftree s \<and> is_leveln_ftree 0 s
\<and> is_leveln_node 0 nd \<and> is_measured_node nd"
"toList t = (toList s) @ (nodeToList nd)"
by (auto simp add: viewRn_inv[of t 0 nd s] viewRn_list[of t])
with v1 vn have v3: "nodeToList nd = [a]"
apply (auto simp add: viewR_def )
apply (induct nd)
apply (simp_all (no_asm_use))
done
with v1 v2
show "\<exists>a s. viewR t = Some (a, s) \<and> ft_invar s \<and> toList t = toList s @ [a]"
unfolding ft_invar_def by auto
qed
lemma viewR_correct_empty[simp]: "viewR Empty = None"
unfolding viewR_def by simp
lemma viewR_correct_nonEmpty:
assumes "ft_invar t" "t \<noteq> Empty"
obtains a s where
"viewR t = Some (a, s)" "ft_invar s \<and> toList t = toList s @ [a]"
using assms viewR_correct by blast
text \<open>Finger trees viewed as a double-ended queue. The head and tail functions
here are only
defined for non-empty queues, while the view-functions were also defined for
empty finger trees.\<close>
text "Check for emptiness"
definition isEmpty :: "('e,'a) FingerTreeStruc \<Rightarrow> bool" where
[code del]: "isEmpty t = (t = Empty)"
lemma isEmpty_correct: "isEmpty t \<longleftrightarrow> toList t = []"
unfolding isEmpty_def by (simp add: toList_empty)
\<comment> \<open>Avoid comparison with @{text "(=)"}, and thus unnecessary equality-class
parameter on element types in generated code\<close>
lemma [code]: "isEmpty t = (case t of Empty \<Rightarrow> True | _ \<Rightarrow> False)"
apply (cases t)
apply (auto simp add: isEmpty_def)
done
text "Leftmost element"
definition head :: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> 'e \<times> 'a" where
"head t = (case viewL t of (Some (a, _)) \<Rightarrow> a)"
lemma head_correct:
assumes "ft_invar t" "t \<noteq> Empty"
shows "head t = hd (toList t)"
proof -
from assms viewL_correct
obtain a s where
v1:"viewL t = Some (a, s) \<and> ft_invar s \<and> toList t = a # toList s" by blast
hence v2: "head t = a" by (auto simp add: head_def)
from v1 have "hd (toList t) = a" by simp
with v2 show ?thesis by simp
qed
text "All but the leftmost element"
definition tail
:: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> ('e,'a) FingerTreeStruc"
where
"tail t = (case viewL t of (Some (_, m)) \<Rightarrow> m)"
lemma tail_correct:
assumes "ft_invar t" "t \<noteq> Empty"
shows "toList (tail t) = tl (toList t)" and "ft_invar (tail t)"
proof -
from assms viewL_correct
obtain a s where
v1:"viewL t = Some (a, s) \<and> ft_invar s \<and> toList t = a # toList s" by blast
hence v2: "tail t = s" by (auto simp add: tail_def)
from v1 have "tl (toList t) = toList s" by simp
with v1 v2 show
"toList (tail t) = tl (toList t)"
"ft_invar (tail t)"
by simp_all
qed
text "Rightmost element"
definition headR :: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> 'e \<times> 'a" where
"headR t = (case viewR t of (Some (a, _)) \<Rightarrow> a)"
lemma headR_correct:
assumes "ft_invar t" "t \<noteq> Empty"
shows "headR t = last (toList t)"
proof -
from assms viewR_correct
obtain a s where
v1:"viewR t = Some (a, s) \<and> ft_invar s \<and> toList t = toList s @ [a]" by blast
hence v2: "headR t = a" by (auto simp add: headR_def)
with v1 show ?thesis by auto
qed
text "All but the rightmost element"
definition tailR
:: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> ('e,'a) FingerTreeStruc"
where
"tailR t = (case viewR t of (Some (_, m)) \<Rightarrow> m)"
lemma tailR_correct:
assumes "ft_invar t" "t \<noteq> Empty"
shows "toList (tailR t) = butlast (toList t)" and "ft_invar (tailR t)"
proof -
from assms viewR_correct
obtain a s where
v1:"viewR t = Some (a, s) \<and> ft_invar s \<and> toList t = toList s @ [a]" by blast
hence v2: "tailR t = s" by (auto simp add: tailR_def)
with v1 show "toList (tailR t) = butlast (toList t)" and "ft_invar (tailR t)"
by auto
qed
subsubsection \<open>Concatenation\<close>
primrec lconsNlist :: "('e,'a::monoid_add) Node list
\<Rightarrow> ('e,'a) FingerTreeStruc \<Rightarrow> ('e,'a) FingerTreeStruc" where
"lconsNlist [] t = t" |
"lconsNlist (x#xs) t = nlcons x (lconsNlist xs t)"
primrec rconsNlist :: "('e,'a::monoid_add) FingerTreeStruc
\<Rightarrow> ('e,'a) Node list \<Rightarrow> ('e,'a) FingerTreeStruc" where
"rconsNlist t [] = t" |
"rconsNlist t (x#xs) = rconsNlist (nrcons t x) xs"
fun nodes :: "('e,'a::monoid_add) Node list \<Rightarrow> ('e,'a) Node list" where
"nodes [a, b] = [node2 a b]" |
"nodes [a, b, c] = [node3 a b c]" |
"nodes [a,b,c,d] = [node2 a b, node2 c d]" |
"nodes (a#b#c#xs) = (node3 a b c) # (nodes xs)"
text \<open>Recursively we concatenate two FingerTreeStrucs while we keep the
inner Nodes in a list\<close>
fun app3 :: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> ('e,'a) Node list
\<Rightarrow> ('e,'a) FingerTreeStruc \<Rightarrow> ('e,'a) FingerTreeStruc" where
"app3 Empty xs t = lconsNlist xs t" |
"app3 t xs Empty = rconsNlist t xs" |
"app3 (Single x) xs t = nlcons x (lconsNlist xs t)" |
"app3 t xs (Single x) = nrcons (rconsNlist t xs) x" |
"app3 (Deep _ pr1 m1 sf1) ts (Deep _ pr2 m2 sf2) =
deep pr1 (app3 m1
(nodes ((digitToNlist sf1) @ ts @ (digitToNlist pr2))) m2) sf2"
lemma lconsNlist_inv:
assumes "is_leveln_ftree n t"
and "is_measured_ftree t"
and "\<forall> x\<in>set xs. (is_leveln_node n x \<and> is_measured_node x)"
shows
"is_leveln_ftree n (lconsNlist xs t) \<and> is_measured_ftree (lconsNlist xs t)"
by (insert assms, induct xs, auto simp add: nlcons_invlevel nlcons_invmeas)
lemma rconsNlist_inv:
assumes "is_leveln_ftree n t"
and "is_measured_ftree t"
and "\<forall> x\<in>set xs. (is_leveln_node n x \<and> is_measured_node x)"
shows
"is_leveln_ftree n (rconsNlist t xs) \<and> is_measured_ftree (rconsNlist t xs)"
by (insert assms, induct xs arbitrary: t,
auto simp add: nrcons_invlevel nrcons_invmeas)
lemma nodes_inv:
assumes "\<forall> x \<in> set ts. is_leveln_node n x \<and> is_measured_node x"
and "length ts \<ge> 2"
shows "\<forall> x \<in> set (nodes ts). is_leveln_node (Suc n) x \<and> is_measured_node x"
proof (insert assms, induct ts rule: nodes.induct)
case (1 a b)
thus ?case by (simp add: node2_def)
next
case (2 a b c)
thus ?case by (simp add: node3_def)
next
case (3 a b c d)
thus ?case by (simp add: node2_def)
next
case (4 a b c v vb vc)
thus ?case by (simp add: node3_def)
next
show "\<lbrakk>\<forall>x\<in>set []. is_leveln_node n x \<and> is_measured_node x; 2 \<le> length []\<rbrakk>
\<Longrightarrow> \<forall>x\<in>set (nodes []). is_leveln_node (Suc n) x \<and> is_measured_node x"
by simp
next
show
"\<And>v. \<lbrakk>\<forall>x\<in>set [v]. is_leveln_node n x \<and> is_measured_node x; 2 \<le> length [v]\<rbrakk>
\<Longrightarrow> \<forall>x\<in>set (nodes [v]). is_leveln_node (Suc n) x \<and> is_measured_node x"
by simp
qed
lemma nodes_inv2:
assumes "is_leveln_digit n sf1"
and "is_measured_digit sf1"
and "is_leveln_digit n pr2"
and "is_measured_digit pr2"
and "\<forall> x \<in> set ts. is_leveln_node n x \<and> is_measured_node x"
shows
"\<forall>x\<in>set (nodes (digitToNlist sf1 @ ts @ digitToNlist pr2)).
is_leveln_node (Suc n) x \<and> is_measured_node x"
proof -
have v1:" \<forall>x\<in>set (digitToNlist sf1 @ ts @ digitToNlist pr2).
is_leveln_node n x \<and> is_measured_node x"
using assms
apply (simp add: digitToNlist_def)
apply (cases sf1)
apply (cases pr2)
apply simp_all
apply (cases pr2)
apply (simp_all)
apply (cases pr2)
apply (simp_all)
apply (cases pr2)
apply (simp_all)
done
have v2: "length (digitToNlist sf1 @ ts @ digitToNlist pr2) \<ge> 2"
apply (cases sf1)
apply (cases pr2)
apply simp_all
done
thus ?thesis
using v1 nodes_inv[of "digitToNlist sf1 @ ts @ digitToNlist pr2"]
by blast
qed
lemma app3_inv:
assumes "is_leveln_ftree n t1"
and "is_leveln_ftree n t2"
and "is_measured_ftree t1"
and "is_measured_ftree t2"
and "\<forall> x\<in>set xs. (is_leveln_node n x \<and> is_measured_node x)"
shows "is_leveln_ftree n (app3 t1 xs t2) \<and> is_measured_ftree (app3 t1 xs t2)"
proof (insert assms, induct t1 xs t2 arbitrary: n rule: app3.induct)
case (1 xs t n)
thus ?case using lconsNlist_inv by simp
next
case "2_1"
thus ?case by (simp add: rconsNlist_inv)
next
case "2_2"
thus ?case by (simp add: lconsNlist_inv rconsNlist_inv)
next
case "3_1"
thus ?case by (simp add: lconsNlist_inv nlcons_invlevel nlcons_invmeas )
next
case "3_2"
thus ?case
by (simp only: app3.simps)
(simp add: lconsNlist_inv nlcons_invlevel nlcons_invmeas)
next
case 4
thus ?case
by (simp only: app3.simps)
(simp add: rconsNlist_inv nrcons_invlevel nrcons_invmeas)
next
case (5 uu pr1 m1 sf1 ts uv pr2 m2 sf2 n)
thus ?case
proof -
have v1: "is_leveln_ftree (Suc n) m1"
and v2: "is_leveln_ftree (Suc n) m2"
using "5.prems" by (simp_all add: is_leveln_ftree_def)
have v3: "is_measured_ftree m1"
and v4: "is_measured_ftree m2"
using "5.prems" by (simp_all add: is_measured_ftree_def)
have v5: "is_leveln_digit n sf1"
"is_measured_digit sf1"
"is_leveln_digit n pr2"
"is_measured_digit pr2"
"\<forall>x\<in>set ts. is_leveln_node n x \<and> is_measured_node x"
using "5.prems"
by (simp_all add: is_leveln_ftree_def is_measured_ftree_def)
note v6 = nodes_inv2[OF v5]
note v7 = "5.hyps"[OF v1 v2 v3 v4 v6]
have v8: "is_leveln_digit n sf2"
"is_measured_digit sf2"
"is_leveln_digit n pr1"
"is_measured_digit pr1"
using "5.prems"
by (simp_all add: is_leveln_ftree_def is_measured_ftree_def)
show ?thesis using v7 v8
by (simp add: is_leveln_ftree_def is_measured_ftree_def deep_def)
qed
qed
primrec nlistToList:: "(('e, 'a) Node) list \<Rightarrow> ('e \<times> 'a) list" where
"nlistToList [] = []"|
"nlistToList (x#xs) = (nodeToList x) @ (nlistToList xs)"
lemma nodes_list: "length xs \<ge> 2 \<Longrightarrow> nlistToList (nodes xs) = nlistToList xs"
by (induct xs rule: nodes.induct) (auto simp add: node2_def node3_def)
lemma nlistToList_app:
"nlistToList (xs@ys) = (nlistToList xs) @ (nlistToList ys)"
by (induct xs arbitrary: ys, simp_all)
lemma nlistListLCons: "toList (lconsNlist xs t) = (nlistToList xs) @ (toList t)"
by (induct xs) (auto simp add: nlcons_list)
lemma nlistListRCons: "toList (rconsNlist t xs) = (toList t) @ (nlistToList xs)"
by (induct xs arbitrary: t) (auto simp add: nrcons_list)
lemma app3_list_lem1:
"nlistToList (nodes (digitToNlist sf1 @ ts @ digitToNlist pr2)) =
digitToList sf1 @ nlistToList ts @ digitToList pr2"
proof -
have len1: "length (digitToNlist sf1 @ ts @ digitToNlist pr2) \<ge> 2"
by (cases sf1,cases pr2,simp_all)
have "(nlistToList (digitToNlist sf1 @ ts @ digitToNlist pr2))
= (digitToList sf1 @ nlistToList ts @ digitToList pr2)"
apply (cases sf1, cases pr2)
apply (simp_all add: nlistToList_app)
apply (cases pr2, auto)
apply (cases pr2, auto)
apply (cases pr2, auto)
done
with nodes_list[OF len1] show ?thesis by simp
qed
lemma app3_list:
"toList (app3 t1 xs t2) = (toList t1) @ (nlistToList xs) @ (toList t2)"
apply (induct t1 xs t2 rule: app3.induct)
apply (simp_all add: nlistListLCons nlistListRCons nlcons_list nrcons_list)
apply (simp add: app3_list_lem1 deep_def)
done
definition app
:: "('e,'a::monoid_add) FingerTreeStruc \<Rightarrow> ('e,'a) FingerTreeStruc
\<Rightarrow> ('e,'a) FingerTreeStruc"
where "app t1 t2 = app3 t1 [] t2"
lemma app_correct:
assumes "ft_invar t1" "ft_invar t2"
shows "toList (app t1 t2) = (toList t1) @ (toList t2)"
and "ft_invar (app t1 t2)"
using assms
by (auto simp add: app3_inv app3_list ft_invar_def app_def)
lemma app_inv: "\<lbrakk>ft_invar t1;ft_invar t2\<rbrakk> \<Longrightarrow> ft_invar (app t1 t2)"
by (auto simp add: app3_inv ft_invar_def app_def)
lemma app_list[simp]: "toList (app t1 t2) = (toList t1) @ (toList t2)"
by (simp add: app3_list app_def)
subsubsection "Splitting"
type_synonym ('e,'a) SplitDigit =
"('e,'a) Node list \<times> ('e,'a) Node \<times> ('e,'a) Node list"
type_synonym ('e,'a) SplitTree =
"('e,'a) FingerTreeStruc \<times> ('e,'a) Node \<times> ('e,'a) FingerTreeStruc"
text \<open>Auxiliary functions to create a correct finger tree
even if the left or right digit is empty\<close>
fun deepL :: "('e,'a::monoid_add) Node list \<Rightarrow> ('e,'a) FingerTreeStruc
\<Rightarrow> ('e,'a) Digit \<Rightarrow> ('e,'a) FingerTreeStruc" where
"deepL [] m sf = (case (viewLn m) of None \<Rightarrow> digitToTree sf |
(Some (a, m2)) \<Rightarrow> deep (nodeToDigit a) m2 sf)" |
"deepL pr m sf = deep (nlistToDigit pr) m sf"
fun deepR :: "('e,'a::monoid_add) Digit \<Rightarrow> ('e,'a) FingerTreeStruc
\<Rightarrow> ('e,'a) Node list \<Rightarrow> ('e,'a) FingerTreeStruc" where
"deepR pr m [] = (case (viewRn m) of None \<Rightarrow> digitToTree pr |
(Some (a, m2)) \<Rightarrow> deep pr m2 (nodeToDigit a))" |
"deepR pr m sf = deep pr m (nlistToDigit sf)"
text \<open>Splitting a list of nodes\<close>
fun splitNlist :: "('a::monoid_add \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> ('e,'a) Node list
\<Rightarrow> ('e,'a) SplitDigit" where
"splitNlist p i [a] = ([],a,[])" |
"splitNlist p i (a#b) =
(let i2 = (i + gmn a) in
(if (p i2)
then ([],a,b)
else
(let (l,x,r) = (splitNlist p i2 b) in ((a#l),x,r))))"
text \<open>Splitting a digit by converting it into a list of nodes\<close>
definition splitDigit :: "('a::monoid_add \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> ('e,'a) Digit
\<Rightarrow> ('e,'a) SplitDigit" where
"splitDigit p i d = splitNlist p i (digitToNlist d)"
text \<open>Creating a finger tree from list of nodes\<close>
definition nlistToTree :: "('e,'a::monoid_add) Node list
\<Rightarrow> ('e,'a) FingerTreeStruc" where
"nlistToTree xs = lconsNlist xs Empty"
text \<open>Recursive splitting into a left and right tree and a center node\<close>
fun nsplitTree :: "('a::monoid_add \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> ('e,'a) FingerTreeStruc
\<Rightarrow> ('e,'a) SplitTree" where
"nsplitTree p i Empty = (Empty, Tip undefined undefined, Empty)"
\<comment> \<open>Making the function total\<close> |
"nsplitTree p i (Single ea) = (Empty,ea,Empty)" |
"nsplitTree p i (Deep _ pr m sf) =
(let
vpr = (i + gmd pr);
vm = (vpr + gmft m)
in
if (p vpr) then
(let (l,x,r) = (splitDigit p i pr) in
(nlistToTree l,x,deepL r m sf))
else (if (p vm) then
(let (ml,xs,mr) = (nsplitTree p vpr m);
(l,x,r) = (splitDigit p (vpr + gmft ml) (nodeToDigit xs)) in
(deepR pr ml l,x,deepL r mr sf))
else
(let (l,x,r) = (splitDigit p vm sf) in
(deepR pr m l,x,nlistToTree r))
))"
lemma nlistToTree_inv:
"\<forall> x \<in> set nl. is_measured_node x \<Longrightarrow> is_measured_ftree (nlistToTree nl)"
"\<forall> x \<in> set nl. is_leveln_node n x \<Longrightarrow> is_leveln_ftree n (nlistToTree nl)"
by (unfold nlistToTree_def, induct nl, auto simp add: nlcons_invmeas)
(induct nl, auto simp add: nlcons_invlevel)
lemma nlistToTree_list: "toList (nlistToTree nl) = nlistToList nl"
by (auto simp add: nlistToTree_def nlistListLCons)
lemma deepL_inv:
assumes "is_leveln_ftree (Suc n) m \<and> is_measured_ftree m"
and "is_leveln_digit n sf \<and> is_measured_digit sf"
and "\<forall> x \<in> set pr. (is_measured_node x \<and> is_leveln_node n x) \<and> length pr \<le> 4"
shows "is_leveln_ftree n (deepL pr m sf) \<and> is_measured_ftree (deepL pr m sf)"
apply (insert assms)
apply (induct "pr" m sf rule: deepL.induct)
apply (simp split: viewnres_split)
apply auto[1]
apply (simp_all add: digitToTree_inv deep_def)
proof -
fix m sf Node FingerTreeStruc
assume "is_leveln_ftree (Suc n) m" "is_measured_ftree m"
"is_leveln_digit n sf" "is_measured_digit sf"
"viewLn m = Some (Node, FingerTreeStruc)"
thus "is_leveln_digit n (nodeToDigit Node)
\<and> is_leveln_ftree (Suc n) FingerTreeStruc"
by (simp add: viewLn_inv[of m "Suc n" Node FingerTreeStruc] nodeToDigit_inv)
next
fix m sf Node FingerTreeStruc
assume assms1:
"is_leveln_ftree (Suc n) m" "is_measured_ftree m"
"is_leveln_digit n sf" "is_measured_digit sf"
"viewLn m = Some (Node, FingerTreeStruc)"
thus "is_measured_digit (nodeToDigit Node) \<and> is_measured_ftree FingerTreeStruc"
apply (auto simp only: viewLn_inv[of m "Suc n" Node FingerTreeStruc])
proof -
from assms1 have "is_measured_node Node \<and> is_leveln_node (Suc n) Node"
by (simp add: viewLn_inv[of m "Suc n" Node FingerTreeStruc])
thus "is_measured_digit (nodeToDigit Node)"
by (auto simp add: nodeToDigit_inv)
qed
next
fix v va
assume
"is_measured_node v \<and> is_leveln_node n (v:: ('a,'b) Node) \<and>
length (va::('a, 'b) Node list) \<le> 3 \<and>
(\<forall>x\<in>set va. is_measured_node x \<and> is_leveln_node n x \<and> length va \<le> 3)"
thus "is_leveln_digit n (nlistToDigit (v # va))
\<and> is_measured_digit (nlistToDigit (v # va))"
by(cases "v#va" rule: nlistToDigit.cases,simp_all)
qed
(*corollary deepL_inv':
assumes "is_leveln_ftree (Suc n) m" "is_measured_ftree m"
and "is_leveln_digit n sf" "is_measured_digit sf"
and "\<forall> x \<in> set pr. (is_measured_node x \<and> is_leveln_node n x)" "length pr \<le> 4"
shows "is_leveln_ftree n (deepL pr m sf)" "is_measured_ftree (deepL pr m sf)"
using assms deepL_inv by blast+
*)
lemma nlistToDigit_list:
assumes "1 \<le> length xs \<and> length xs \<le> 4"
shows "digitToList(nlistToDigit xs) = nlistToList xs"
by (insert assms, cases xs rule: nlistToDigit.cases,auto)
lemma deepL_list:
assumes "is_leveln_ftree (Suc n) m \<and> is_measured_ftree m"
and "is_leveln_digit n sf \<and> is_measured_digit sf"
and "\<forall> x \<in> set pr. (is_measured_node x \<and> is_leveln_node n x) \<and> length pr \<le> 4"
shows "toList (deepL pr m sf) = nlistToList pr @ toList m @ digitToList sf"
proof (insert assms, induct "pr" m sf rule: deepL.induct)
case (1 m sf)
thus ?case
proof (auto split: viewnres_split simp add: deep_def)
assume "viewLn m = None"
hence "m = Empty" by (metis viewLn_empty)
hence "toList m = []" by simp
thus "toList (digitToTree sf) = toList m @ digitToList sf"
by (simp add:digitToTree_list)
next
fix nd t
assume "viewLn m = Some (nd, t)"
"is_leveln_ftree (Suc n) m" "is_measured_ftree m"
hence "nodeToList nd @ toList t = toList m" by (metis viewLn_list)
thus "digitToList (nodeToDigit nd) @ toList t = toList m"
by (simp add: nodeToDigit_list)
qed
next
case (2 v va m sf)
thus ?case
apply (unfold deepL.simps)
apply (simp add: deep_def)
apply (simp add: nlistToDigit_list)
done
qed
lemma deepR_inv:
assumes "is_leveln_ftree (Suc n) m \<and> is_measured_ftree m"
and "is_leveln_digit n pr \<and> is_measured_digit pr"
and "\<forall> x \<in> set sf. (is_measured_node x \<and> is_leveln_node n x) \<and> length sf \<le> 4"
shows "is_leveln_ftree n (deepR pr m sf) \<and> is_measured_ftree (deepR pr m sf)"
apply (insert assms)
apply (induct "pr" m sf rule: deepR.induct)
apply (simp split: viewnres_split)
apply auto[1]
apply (simp_all add: digitToTree_inv deep_def)
proof -
fix m "pr" Node FingerTreeStruc
assume "is_leveln_ftree (Suc n) m" "is_measured_ftree m"
"is_leveln_digit n pr" "is_measured_digit pr"
"viewRn m = Some (Node, FingerTreeStruc)"
thus
"is_leveln_digit n (nodeToDigit Node)
\<and> is_leveln_ftree (Suc n) FingerTreeStruc"
by (simp add: viewRn_inv[of m "Suc n" Node FingerTreeStruc] nodeToDigit_inv)
next
fix m "pr" Node FingerTreeStruc
assume assms1:
"is_leveln_ftree (Suc n) m" "is_measured_ftree m"
"is_leveln_digit n pr" "is_measured_digit pr"
"viewRn m = Some (Node, FingerTreeStruc)"
thus "is_measured_ftree FingerTreeStruc \<and> is_measured_digit (nodeToDigit Node)"
apply (auto simp only: viewRn_inv[of m "Suc n" Node FingerTreeStruc])
proof -
from assms1 have "is_measured_node Node \<and> is_leveln_node (Suc n) Node"
by (simp add: viewRn_inv[of m "Suc n" Node FingerTreeStruc])
thus "is_measured_digit (nodeToDigit Node)"
by (auto simp add: nodeToDigit_inv)
qed
next
fix v va
assume
"is_measured_node v \<and> is_leveln_node n (v:: ('a,'b) Node) \<and>
length (va::('a, 'b) Node list) \<le> 3 \<and>
(\<forall>x\<in>set va. is_measured_node x \<and> is_leveln_node n x \<and> length va \<le> 3)"
thus "is_leveln_digit n (nlistToDigit (v # va)) \<and>
is_measured_digit (nlistToDigit (v # va))"
by(cases "v#va" rule: nlistToDigit.cases,simp_all)
qed
lemma deepR_list:
assumes "is_leveln_ftree (Suc n) m \<and> is_measured_ftree m"
and "is_leveln_digit n pr \<and> is_measured_digit pr"
and "\<forall> x \<in> set sf. (is_measured_node x \<and> is_leveln_node n x) \<and> length sf \<le> 4"
shows "toList (deepR pr m sf) = digitToList pr @ toList m @ nlistToList sf"
proof (insert assms, induct "pr" m sf rule: deepR.induct)
case (1 "pr" m)
thus ?case
proof (auto split: viewnres_split simp add: deep_def)
assume "viewRn m = None"
hence "m = Empty" by (metis viewRn_empty)
hence "toList m = []" by simp
thus "toList (digitToTree pr) = digitToList pr @ toList m"
by (simp add:digitToTree_list)
next
fix nd t
assume "viewRn m = Some (nd, t)" "is_leveln_ftree (Suc n) m"
"is_measured_ftree m"
hence "toList t @ nodeToList nd = toList m" by (metis viewRn_list)
thus "toList t @ digitToList (nodeToDigit nd) = toList m"
by (simp add: nodeToDigit_list)
qed
next
case (2 "pr" m v va)
thus ?case
apply (unfold deepR.simps)
apply (simp add: deep_def)
apply (simp add: nlistToDigit_list)
done
qed
primrec gmnl:: "('e, 'a::monoid_add) Node list \<Rightarrow> 'a" where
"gmnl [] = 0"|
"gmnl (x#xs) = gmn x + gmnl xs"
lemma gmnl_correct:
assumes "\<forall> x \<in> set xs. is_measured_node x"
shows "gmnl xs = sum_list (map snd (nlistToList xs))"
by (insert assms, induct xs) (auto simp add: add.assoc gmn_correct)
lemma splitNlist_correct:" \<lbrakk>
\<And>(a::'a) (b::'a). p a \<Longrightarrow> p (a + b);
\<not> p i;
p (i + gmnl (nl ::('e,'a::monoid_add) Node list));
splitNlist p i nl = (l, n, r)
\<rbrakk> \<Longrightarrow>
\<not> p (i + (gmnl l))
\<and>
p (i + (gmnl l) + (gmn n))
\<and>
nl = l @ n # r
"
proof (induct p i nl arbitrary: l n r rule: splitNlist.induct)
case 1 thus ?case by simp
next
case (2 p i a v va l n r) note IV = this
show ?case
proof (cases "p (i + (gmn a))")
case True with IV show ?thesis by simp
next
case False note IV2 = this IV thus ?thesis
proof -
obtain l1 n1 r1 where
v1[simp]: "splitNlist p (i + gmn a) (v # va) = (l1, n1, r1)"
by (cases "splitNlist p (i + gmn a) (v # va)", blast)
note miv = IV2(2)[of "i + gmn a" l1 n1 r1]
have v2:"p (i + gmn a + gmnl (v # va))"
using IV2(5) by (simp add: add.assoc)
note miv2 = miv[OF _ IV2(1) IV2(3) IV2(1) v2 v1]
have v3: "a # l1 = l" "n1 = n" "r1 = r" using IV2 v1 by auto
with miv2 have
v4: "\<not> p (i + gmn a + gmnl l1) \<and>
p (i + gmn a + gmnl l1 + gmn n1) \<and>
v # va = l1 @ n1 # r1"
by auto
with v2 v3 show ?thesis
by (auto simp add: add.assoc)
qed
qed
next
case 3 thus ?case by simp
qed
lemma digitToNlist_inv:
"is_measured_digit d \<Longrightarrow> (\<forall> x \<in> set (digitToNlist d). is_measured_node x)"
"is_leveln_digit n d \<Longrightarrow> (\<forall> x \<in> set (digitToNlist d). is_leveln_node n x)"
by (cases d, auto)(cases d, auto)
lemma gmnl_gmd:
"is_measured_digit d \<Longrightarrow> gmnl (digitToNlist d) = gmd d"
by (cases d, auto simp add: add.assoc)
lemma gmn_gmd:
"is_measured_node nd \<Longrightarrow> gmd (nodeToDigit nd) = gmn nd"
by (auto simp add: nodeToDigit_inv nodeToDigit_list gmn_correct gmd_correct)
lemma splitDigit_inv:
"\<lbrakk>
\<And>(a::'a) (b::'a). p a \<Longrightarrow> p (a + b);
\<not> p i;
is_measured_digit d;
is_leveln_digit n d;
p (i + gmd (d ::('e,'a::monoid_add) Digit));
splitDigit p i d = (l, nd, r)
\<rbrakk> \<Longrightarrow>
\<not> p (i + (gmnl l))
\<and>
p (i + (gmnl l) + (gmn nd))
\<and>
(\<forall> x \<in> set l. (is_measured_node x \<and> is_leveln_node n x))
\<and>
(\<forall> x \<in> set r. (is_measured_node x \<and> is_leveln_node n x))
\<and>
(is_measured_node nd \<and> is_leveln_node n nd )
"
proof -
fix p i d n l nd r
assume assms: "\<And>a b. p a \<Longrightarrow> p (a + b)" "\<not> p i" "is_measured_digit d"
"p (i + gmd d)" "splitDigit p i d = (l, nd, r)"
"is_leveln_digit n d"
from assms(3, 4) have v1: "p (i + gmnl (digitToNlist d))"
by (simp add: gmnl_gmd)
note snc = splitNlist_correct [of p i "digitToNlist d" l nd r]
from assms(5) have v2: "splitNlist p i (digitToNlist d) = (l, nd, r)"
by (simp add: splitDigit_def)
note snc1 = snc[OF assms(1) assms(2) v1 v2]
hence v3: "\<not> p (i + gmnl l) \<and> p (i + gmnl l + gmn nd) \<and>
digitToNlist d = l @ nd # r" by auto
from assms(3,6) have
v4:" \<forall> x \<in> set (digitToNlist d). is_measured_node x"
" \<forall> x \<in> set (digitToNlist d). is_leveln_node n x"
by(auto simp add: digitToNlist_inv)
with v3 have v5: "\<forall> x \<in> set l. (is_measured_node x \<and> is_leveln_node n x)"
"\<forall> x \<in> set r. (is_measured_node x \<and> is_leveln_node n x)"
"is_measured_node nd \<and> is_leveln_node n nd" by auto
with v3 v5 show
"\<not> p (i + gmnl l) \<and> p (i + gmnl l + gmn nd) \<and>
(\<forall>x\<in>set l. is_measured_node x \<and> is_leveln_node n x) \<and>
(\<forall>x\<in>set r. is_measured_node x \<and> is_leveln_node n x) \<and>
is_measured_node nd \<and> is_leveln_node n nd"
by auto
qed
lemma splitDigit_inv':
"\<lbrakk>
splitDigit p i d = (l, nd, r);
is_measured_digit d;
is_leveln_digit n d
\<rbrakk> \<Longrightarrow>
(\<forall> x \<in> set l. (is_measured_node x \<and> is_leveln_node n x))
\<and>
(\<forall> x \<in> set r. (is_measured_node x \<and> is_leveln_node n x))
\<and>
(is_measured_node nd \<and> is_leveln_node n nd )
"
apply (unfold splitDigit_def)
apply (cases d)
apply (auto split: if_split_asm simp add: Let_def)
done
lemma splitDigit_list: "splitDigit p i d = (l,n,r) \<Longrightarrow>
(digitToList d) = (nlistToList l) @ (nodeToList n) @ (nlistToList r)
\<and> length l \<le> 4 \<and> length r \<le> 4"
apply (unfold splitDigit_def)
apply (cases d)
apply (auto split: if_split_asm simp add: Let_def)
done
lemma gmnl_gmft: "\<forall> x \<in> set nl. is_measured_node x \<Longrightarrow>
gmft (nlistToTree nl) = gmnl nl"
by (auto simp add: gmnl_correct[of nl] nlistToTree_list[of nl]
nlistToTree_inv[of nl] gmft_correct[of "nlistToTree nl"])
lemma gmftR_gmnl:
assumes "is_leveln_ftree (Suc n) m \<and> is_measured_ftree m"
and "is_leveln_digit n pr \<and> is_measured_digit pr"
and "\<forall> x \<in> set sf. (is_measured_node x \<and> is_leveln_node n x) \<and> length sf \<le> 4"
shows "gmft (deepR pr m sf) = gmd pr + gmft m + gmnl sf"
proof-
from assms have
v1: "toList (deepR pr m sf) = digitToList pr @ toList m @ nlistToList sf"
by (auto simp add: deepR_list)
from assms have
v2: "is_measured_ftree (deepR pr m sf)"
by (auto simp add: deepR_inv)
with v1 have
v3: "gmft (deepR pr m sf) =
sum_list (map snd (digitToList pr @ toList m @ nlistToList sf))"
by (auto simp add: gmft_correct)
have
v4:"gmd pr + gmft m + gmnl sf =
sum_list (map snd (digitToList pr @ toList m @ nlistToList sf))"
by (auto simp add: gmd_correct gmft_correct gmnl_correct assms add.assoc)
with v3 show ?thesis by simp
qed
lemma nsplitTree_invpres: "\<lbrakk>
is_leveln_ftree n (s:: ('e,'a::monoid_add) FingerTreeStruc);
is_measured_ftree s;
\<not> p i;
p (i + (gmft s));
(nsplitTree p i s) = (l, nd, r)\<rbrakk>
\<Longrightarrow>
is_leveln_ftree n l
\<and>
is_measured_ftree l
\<and>
is_leveln_ftree n r
\<and>
is_measured_ftree r
\<and>
is_leveln_node n nd
\<and>
is_measured_node nd
"
proof (induct p i s arbitrary: n l nd r rule: nsplitTree.induct)
case 1
thus ?case by auto
next
case 2 thus ?case by auto
next
case (3 p i uu "pr" m sf n l nd r)
thus ?case
proof (cases "p (i + gmd pr)")
case True with 3 show ?thesis
proof -
obtain l1 x r1 where
l1xr1: "splitDigit p i pr = (l1,x,r1)"
by (cases "splitDigit p i pr", blast)
with True 3 have
v1: "l = nlistToTree l1" "nd = x" "r = deepL r1 m sf" by auto
from l1xr1 have
v2: "digitToList pr = nlistToList l1 @ nodeToList x @ nlistToList r1"
"length l1 \<le> 4" "length r1 \<le> 4"
by (auto simp add: splitDigit_list)
from 3(2,3) have
pr_m_sf_inv: "is_leveln_digit n pr \<and> is_measured_digit pr"
"is_leveln_ftree (Suc n) m \<and> is_measured_ftree m"
"is_leveln_digit n sf \<and> is_measured_digit sf" by simp_all
with 3(4,5) pr_m_sf_inv(1) True l1xr1
splitDigit_inv'[of p i "pr" l1 x r1 n] have
l1_x_r1_inv:
"\<forall> x \<in> set l1. (is_measured_node x \<and> is_leveln_node n x)"
"\<forall> x \<in> set r1. (is_measured_node x \<and> is_leveln_node n x)"
"is_measured_node x \<and> is_leveln_node n x"
by auto
from l1_x_r1_inv v1 v2(3) pr_m_sf_inv have
ziel3: "is_leveln_ftree n l \<and> is_measured_ftree l \<and>
is_leveln_ftree n r \<and> is_measured_ftree r \<and>
is_leveln_node n nd \<and> is_measured_node nd"
by (auto simp add: nlistToTree_inv deepL_inv)
thus ?thesis by simp
qed
next
case False note case1 = this with 3 show ?thesis
proof (cases "p (i + gmd pr + gmft m)")
case False with case1 3 show ?thesis
proof -
obtain l1 x r1 where
l1xr1: "splitDigit p (i + gmd pr + gmft m) sf = (l1,x,r1)"
by (cases "splitDigit p (i + gmd pr + gmft m) sf", blast)
with case1 False 3 have
v1: "l = deepR pr m l1" "nd = x" "r = nlistToTree r1" by auto
from l1xr1 have
v2: "digitToList sf = nlistToList l1 @ nodeToList x @ nlistToList r1"
"length l1 \<le> 4" "length r1 \<le> 4"
by (auto simp add: splitDigit_list)
from 3(2,3) have
pr_m_sf_inv: "is_leveln_digit n pr \<and> is_measured_digit pr"
"is_leveln_ftree (Suc n) m \<and> is_measured_ftree m"
"is_leveln_digit n sf \<and> is_measured_digit sf" by simp_all
from 3 have
v7: "p (i + gmd pr + gmft m + gmd sf)" by (auto simp add: add.assoc)
with pr_m_sf_inv 3(4) pr_m_sf_inv(3) case1 False l1xr1
splitDigit_inv'[of p "i + gmd pr + gmft m" sf l1 x r1 n]
have l1_x_r1_inv:
"\<forall> x \<in> set l1. (is_measured_node x \<and> is_leveln_node n x)"
"\<forall> x \<in> set r1. (is_measured_node x \<and> is_leveln_node n x)"
"is_measured_node x \<and> is_leveln_node n x"
by auto
from l1_x_r1_inv v1 v2(2) pr_m_sf_inv have
ziel3: "is_leveln_ftree n l \<and> is_measured_ftree l \<and>
is_leveln_ftree n r \<and> is_measured_ftree r \<and>
is_leveln_node n nd \<and> is_measured_node nd"
by (auto simp add: nlistToTree_inv deepR_inv)
from ziel3 show ?thesis by simp
qed
next
case True with case1 3 show ?thesis
proof -
obtain l1 x r1 where
l1_x_r1 :"nsplitTree p (i + gmd pr) m = (l1, x, r1)"
by (cases "nsplitTree p (i + gmd pr) m", blast)
from 3(2,3) have
pr_m_sf_inv: "is_leveln_digit n pr \<and> is_measured_digit pr"
"is_leveln_ftree (Suc n) m \<and> is_measured_ftree m"
"is_leveln_digit n sf \<and> is_measured_digit sf" by simp_all
with True case1
"3.hyps"[of "i + gmd pr" "i + gmd pr + gmft m" "Suc n" l1 x r1]
3(6) l1_x_r1
have l1_x_r1_inv:
"is_leveln_ftree (Suc n) l1 \<and> is_measured_ftree l1"
"is_leveln_ftree (Suc n) r1 \<and> is_measured_ftree r1"
"is_leveln_node (Suc n) x \<and> is_measured_node x"
by auto
obtain l2 x2 r2 where l2_x2_r2:
"splitDigit p (i + gmd pr + gmft l1) (nodeToDigit x) = (l2,x2,r2)"
by (cases "splitDigit p (i + gmd pr + gmft l1) (nodeToDigit x)",blast)
from l1_x_r1_inv have
ndx_inv: "is_leveln_digit n (nodeToDigit x) \<and>
is_measured_digit (nodeToDigit x)"
by (auto simp add: nodeToDigit_inv gmn_gmd)
note spdi = splitDigit_inv'[of p "i + gmd pr + gmft l1"
"nodeToDigit x" l2 x2 r2 n]
from ndx_inv l1_x_r1_inv(1) l2_x2_r2 3(4) have
l2_x2_r2_inv:
"\<forall>x\<in>set l2. is_measured_node x \<and> is_leveln_node n x"
"\<forall>x\<in>set r2. is_measured_node x \<and> is_leveln_node n x"
"is_measured_node x2 \<and> is_leveln_node n x2"
by (auto simp add: spdi)
note spdl = splitDigit_list[of p "i + gmd pr + gmft l1"
"nodeToDigit x" l2 x2 r2]
from l2_x2_r2 have
l2_x2_r2_list:
"digitToList (nodeToDigit x) =
nlistToList l2 @ nodeToList x2 @ nlistToList r2"
"length l2 \<le> 4 \<and> length r2 \<le> 4"
by (auto simp add: spdl)
from case1 True 3(6) l1_x_r1 l2_x2_r2 have
l_nd_r:
"l = deepR pr l1 l2"
"nd = x2"
"r = deepL r2 r1 sf"
by auto
note dr1 = deepR_inv[OF l1_x_r1_inv(1) pr_m_sf_inv(1)]
from dr1 l2_x2_r2_inv l2_x2_r2_list(2) l_nd_r have
l_inv: "is_leveln_ftree n l \<and> is_measured_ftree l"
by simp
note dl1 = deepL_inv[OF l1_x_r1_inv(2) pr_m_sf_inv(3)]
from dl1 l2_x2_r2_inv l2_x2_r2_list(2) l_nd_r have
r_inv: "is_leveln_ftree n r \<and> is_measured_ftree r"
by simp
from l2_x2_r2_inv l_nd_r have
nd_inv: "is_leveln_node n nd \<and> is_measured_node nd"
by simp
from l_inv r_inv nd_inv
show ?thesis by simp
qed
qed
qed
qed
lemma nsplitTree_correct: "\<lbrakk>
is_leveln_ftree n (s:: ('e,'a::monoid_add) FingerTreeStruc);
is_measured_ftree s;
\<And>(a::'a) (b::'a). p a \<Longrightarrow> p (a + b);
\<not> p i;
p (i + (gmft s));
(nsplitTree p i s) = (l, nd, r)\<rbrakk>
\<Longrightarrow> (toList s) = (toList l) @ (nodeToList nd) @ (toList r)
\<and>
\<not> p (i + (gmft l))
\<and>
p (i + (gmft l) + (gmn nd))
\<and>
is_leveln_ftree n l
\<and>
is_measured_ftree l
\<and>
is_leveln_ftree n r
\<and>
is_measured_ftree r
\<and>
is_leveln_node n nd
\<and>
is_measured_node nd
"
proof (induct p i s arbitrary: n l nd r rule: nsplitTree.induct)
case 1
thus ?case by auto
next
case 2 thus ?case by auto
next
case (3 p i uu "pr" m sf n l nd r)
thus ?case
proof (cases "p (i + gmd pr)")
case True with 3 show ?thesis
proof -
obtain l1 x r1 where
l1xr1: "splitDigit p i pr = (l1,x,r1)"
by (cases "splitDigit p i pr", blast)
with True 3(7) have
v1: "l = nlistToTree l1" "nd = x" "r = deepL r1 m sf" by auto
from l1xr1 have
v2: "digitToList pr = nlistToList l1 @ nodeToList x @ nlistToList r1"
"length l1 \<le> 4" "length r1 \<le> 4"
by (auto simp add: splitDigit_list)
from 3(2,3) have
pr_m_sf_inv: "is_leveln_digit n pr \<and> is_measured_digit pr"
"is_leveln_ftree (Suc n) m \<and> is_measured_ftree m"
"is_leveln_digit n sf \<and> is_measured_digit sf" by simp_all
with 3(4,5) pr_m_sf_inv(1) True l1xr1
splitDigit_inv[of p i "pr" n l1 x r1] have
l1_x_r1_inv:
"\<not> p (i + (gmnl l1))"
"p (i + (gmnl l1) + (gmn x))"
"\<forall> x \<in> set l1. (is_measured_node x \<and> is_leveln_node n x)"
"\<forall> x \<in> set r1. (is_measured_node x \<and> is_leveln_node n x)"
"is_measured_node x \<and> is_leveln_node n x"
by auto
from v2 v1 l1_x_r1_inv(4) pr_m_sf_inv have
ziel1: "toList (Deep uu pr m sf) = toList l @ nodeToList nd @ toList r"
by (auto simp add: nlistToTree_list deepL_list)
from l1_x_r1_inv(3) v1(1) have
v3: "gmft l = gmnl l1" by (simp add: gmnl_gmft)
with l1_x_r1_inv(1,2) v1 have
ziel2: " \<not> p (i + gmft l)"
"p (i + gmft l + gmn nd)"
by simp_all
from l1_x_r1_inv(3,4,5) v1 v2(3) pr_m_sf_inv have
ziel3: "is_leveln_ftree n l \<and> is_measured_ftree l \<and>
is_leveln_ftree n r \<and> is_measured_ftree r \<and>
is_leveln_node n nd \<and> is_measured_node nd"
by (auto simp add: nlistToTree_inv deepL_inv)
from ziel1 ziel2 ziel3 show ?thesis by simp
qed
next
case False note case1 = this with 3 show ?thesis
proof (cases "p (i + gmd pr + gmft m)")
case False with case1 3 show ?thesis
proof -
obtain l1 x r1 where
l1xr1: "splitDigit p (i + gmd pr + gmft m) sf = (l1,x,r1)"
by (cases "splitDigit p (i + gmd pr + gmft m) sf", blast)
with case1 False 3(7) have
v1: "l = deepR pr m l1" "nd = x" "r = nlistToTree r1" by auto
from l1xr1 have
v2: "digitToList sf = nlistToList l1 @ nodeToList x @ nlistToList r1"
"length l1 \<le> 4" "length r1 \<le> 4"
by (auto simp add: splitDigit_list)
from 3(2,3) have
pr_m_sf_inv: "is_leveln_digit n pr \<and> is_measured_digit pr"
"is_leveln_ftree (Suc n) m \<and> is_measured_ftree m"
"is_leveln_digit n sf \<and> is_measured_digit sf" by simp_all
from 3(3,6) have
v7: "p (i + gmd pr + gmft m + gmd sf)" by (auto simp add: add.assoc)
with pr_m_sf_inv 3(4) pr_m_sf_inv(3) case1 False l1xr1
splitDigit_inv[of p "i + gmd pr + gmft m" sf n l1 x r1]
have l1_x_r1_inv:
"\<not> p (i + gmd pr + gmft m + gmnl l1)"
"p (i + gmd pr + gmft m + gmnl l1 + gmn x)"
"\<forall> x \<in> set l1. (is_measured_node x \<and> is_leveln_node n x)"
"\<forall> x \<in> set r1. (is_measured_node x \<and> is_leveln_node n x)"
"is_measured_node x \<and> is_leveln_node n x"
by auto
from v2 v1 l1_x_r1_inv(3) pr_m_sf_inv have
ziel1: "toList (Deep uu pr m sf) = toList l @ nodeToList nd @ toList r"
by (auto simp add: nlistToTree_list deepR_list)
from l1_x_r1_inv(4) v1(3) have
v3: "gmft r = gmnl r1" by (simp add: gmnl_gmft)
with l1_x_r1_inv(1,2,3) pr_m_sf_inv v1 v2 have
ziel2: " \<not> p (i + gmft l)"
"p (i + gmft l + gmn nd)"
by (auto simp add: gmftR_gmnl add.assoc)
from l1_x_r1_inv(3,4,5) v1 v2(2) pr_m_sf_inv have
ziel3: "is_leveln_ftree n l \<and> is_measured_ftree l \<and>
is_leveln_ftree n r \<and> is_measured_ftree r \<and>
is_leveln_node n nd \<and> is_measured_node nd"
by (auto simp add: nlistToTree_inv deepR_inv)
from ziel1 ziel2 ziel3 show ?thesis by simp
qed
next
case True with case1 3 show ?thesis
proof -
obtain l1 x r1 where
l1_x_r1 :"nsplitTree p (i + gmd pr) m = (l1, x, r1)"
by (cases "nsplitTree p (i + gmd pr) m") blast
from 3(2,3) have
pr_m_sf_inv: "is_leveln_digit n pr \<and> is_measured_digit pr"
"is_leveln_ftree (Suc n) m \<and> is_measured_ftree m"
"is_leveln_digit n sf \<and> is_measured_digit sf" by simp_all
with True case1
"3.hyps"[of "i + gmd pr" "i + gmd pr + gmft m" "Suc n" l1 x r1]
3(4) l1_x_r1
have l1_x_r1_inv:
"\<not> p (i + gmd pr + gmft l1)"
"p (i + gmd pr + gmft l1 + gmn x)"
"is_leveln_ftree (Suc n) l1 \<and> is_measured_ftree l1"
"is_leveln_ftree (Suc n) r1 \<and> is_measured_ftree r1"
"is_leveln_node (Suc n) x \<and> is_measured_node x"
and l1_x_r1_list:
"toList m = toList l1 @ nodeToList x @ toList r1"
by auto
obtain l2 x2 r2 where l2_x2_r2:
"splitDigit p (i + gmd pr + gmft l1) (nodeToDigit x) = (l2,x2,r2)"
by (cases "splitDigit p (i + gmd pr + gmft l1) (nodeToDigit x)",blast)
from l1_x_r1_inv(2,5) have
ndx_inv: "is_leveln_digit n (nodeToDigit x) \<and>
is_measured_digit (nodeToDigit x)"
"p (i + gmd pr + gmft l1 + gmd (nodeToDigit x))"
by (auto simp add: nodeToDigit_inv gmn_gmd)
note spdi = splitDigit_inv[of p "i + gmd pr + gmft l1"
"nodeToDigit x" n l2 x2 r2]
from ndx_inv l1_x_r1_inv(1) l2_x2_r2 3(4) have
l2_x2_r2_inv:"\<not> p (i + gmd pr + gmft l1 + gmnl l2)"
"p (i + gmd pr + gmft l1 + gmnl l2 + gmn x2)"
"\<forall>x\<in>set l2. is_measured_node x \<and> is_leveln_node n x"
"\<forall>x\<in>set r2. is_measured_node x \<and> is_leveln_node n x"
"is_measured_node x2 \<and> is_leveln_node n x2"
by (auto simp add: spdi)
note spdl = splitDigit_list[of p "i + gmd pr + gmft l1"
"nodeToDigit x" l2 x2 r2]
from l2_x2_r2 have
l2_x2_r2_list:
"digitToList (nodeToDigit x) =
nlistToList l2 @ nodeToList x2 @ nlistToList r2"
"length l2 \<le> 4 \<and> length r2 \<le> 4"
by (auto simp add: spdl)
from case1 True 3(7) l1_x_r1 l2_x2_r2 have
l_nd_r:
"l = deepR pr l1 l2"
"nd = x2"
"r = deepL r2 r1 sf"
by auto
note dr1 = deepR_inv[OF l1_x_r1_inv(3) pr_m_sf_inv(1)]
from dr1 l2_x2_r2_inv(3) l2_x2_r2_list(2) l_nd_r have
l_inv: "is_leveln_ftree n l \<and> is_measured_ftree l"
by simp
note dl1 = deepL_inv[OF l1_x_r1_inv(4) pr_m_sf_inv(3)]
from dl1 l2_x2_r2_inv(4) l2_x2_r2_list(2) l_nd_r have
r_inv: "is_leveln_ftree n r \<and> is_measured_ftree r"
by simp
from l2_x2_r2_inv l_nd_r have
nd_inv: "is_leveln_node n nd \<and> is_measured_node nd"
by simp
from l_nd_r(1,2) l2_x2_r2_inv(1,2,3)
l1_x_r1_inv(3) l2_x2_r2_list(2) pr_m_sf_inv(1)
have split_point:
" \<not> p (i + gmft l)"
"p (i + gmft l + gmn nd)"
by (auto simp add: gmftR_gmnl add.assoc)
from l2_x2_r2_list have x_list:
"nodeToList x = nlistToList l2 @ nodeToList x2 @ nlistToList r2"
by (simp add: nodeToDigit_list)
from l1_x_r1_inv(3) pr_m_sf_inv(1)
l2_x2_r2_inv(3) l2_x2_r2_list(2) l_nd_r(1)
have l_list: "toList l = digitToList pr @ toList l1 @ nlistToList l2"
by (auto simp add: deepR_list)
from l1_x_r1_inv(4) pr_m_sf_inv(3) l2_x2_r2_inv(4)
l2_x2_r2_list(2) l_nd_r(3)
have r_list: "toList r = nlistToList r2 @ toList r1 @ digitToList sf"
by (auto simp add: deepL_list)
from x_list l1_x_r1_list l_list r_list l_nd_r
have "toList (Deep uu pr m sf) = toList l @ nodeToList nd @ toList r"
by auto
with split_point l_inv r_inv nd_inv
show ?thesis by simp
qed
qed
qed
qed
text \<open>
A predicate on the elements of a monoid is called {\em monotone},
iff, when it holds for some value $a$, it also holds for all values $a+b$:
\<close>
text \<open>Split a finger tree by a monotone predicate on the annotations, using
a given initial value. Intuitively, the elements are summed up from left to
right, and the split is done when the predicate first holds for the sum.
The predicate must not hold for the initial value of the summation, and must
hold for the sum of all elements.
\<close>
definition splitTree
:: "('a::monoid_add \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> ('e, 'a) FingerTreeStruc
\<Rightarrow> ('e, 'a) FingerTreeStruc \<times> ('e \<times> 'a) \<times> ('e, 'a) FingerTreeStruc"
where
"splitTree p i t = (let (l, x, r) = nsplitTree p i t in (l, (n_unwrap x), r))"
lemma splitTree_invpres:
assumes inv: "ft_invar (s:: ('e,'a::monoid_add) FingerTreeStruc)"
assumes init_ff: "\<not> p i"
assumes sum_tt: "p (i + annot s)"
assumes fmt: "(splitTree p i s) = (l, (e,a), r)"
shows "ft_invar l" and "ft_invar r"
proof -
obtain l1 nd r1 where
l1_nd_r1: "nsplitTree p i s = (l1, nd, r1)"
by (cases "nsplitTree p i s", blast)
with assms have
l0: "l = l1"
"(e,a) = n_unwrap nd"
"r = r1"
by (auto simp add: splitTree_def)
note nsp = nsplitTree_invpres[of 0 s p i l1 nd r1]
from assms have "p (i + gmft s)" by (simp add: ft_invar_def annot_def)
with assms l1_nd_r1 l0 have
v1:
"is_leveln_ftree 0 l \<and> is_measured_ftree l"
"is_leveln_ftree 0 r \<and> is_measured_ftree r"
"is_leveln_node 0 nd \<and> is_measured_node nd"
by (auto simp add: nsp ft_invar_def)
thus "ft_invar l" and "ft_invar r"
by (simp_all add: ft_invar_def annot_def)
qed
lemma splitTree_correct:
assumes inv: "ft_invar (s:: ('e,'a::monoid_add) FingerTreeStruc)"
assumes mono: "\<forall>a b. p a \<longrightarrow> p (a + b)"
assumes init_ff: "\<not> p i"
assumes sum_tt: "p (i + annot s)"
assumes fmt: "(splitTree p i s) = (l, (e,a), r)"
shows "(toList s) = (toList l) @ (e,a) # (toList r)"
and "\<not> p (i + annot l)"
and "p (i + annot l + a)"
and "ft_invar l" and "ft_invar r"
proof -
obtain l1 nd r1 where
l1_nd_r1: "nsplitTree p i s = (l1, nd, r1)"
by (cases "nsplitTree p i s", blast)
with assms have
l0: "l = l1"
"(e,a) = n_unwrap nd"
"r = r1"
by (auto simp add: splitTree_def)
note nsp = nsplitTree_correct[of 0 s p i l1 nd r1]
from assms have "p (i + gmft s)" by (simp add: ft_invar_def annot_def)
with assms l1_nd_r1 l0 have
v1:
"(toList s) = (toList l) @ (nodeToList nd) @ (toList r)"
"\<not> p (i + (gmft l))"
"p (i + (gmft l) + (gmn nd))"
"is_leveln_ftree 0 l \<and> is_measured_ftree l"
"is_leveln_ftree 0 r \<and> is_measured_ftree r"
"is_leveln_node 0 nd \<and> is_measured_node nd"
by (auto simp add: nsp ft_invar_def)
from v1(6) l0(2) have
ndea: "nd = Tip e a"
by (cases nd) auto
hence nd_list_inv: "nodeToList nd = [(e,a)]"
"gmn nd = a" by simp_all
with v1 show "(toList s) = (toList l) @ (e,a) # (toList r)"
and "\<not> p (i + annot l)"
and "p (i + annot l + a)"
and "ft_invar l" and "ft_invar r"
by (simp_all add: ft_invar_def annot_def)
qed
lemma splitTree_correctE:
assumes inv: "ft_invar (s:: ('e,'a::monoid_add) FingerTreeStruc)"
assumes mono: "\<forall>a b. p a \<longrightarrow> p (a + b)"
assumes init_ff: "\<not> p i"
assumes sum_tt: "p (i + annot s)"
obtains l e a r where
"(splitTree p i s) = (l, (e,a), r)" and
"(toList s) = (toList l) @ (e,a) # (toList r)" and
"\<not> p (i + annot l)" and
"p (i + annot l + a)" and
"ft_invar l" and "ft_invar r"
proof -
obtain l e a r where fmt: "(splitTree p i s) = (l, (e,a), r)"
by (cases "(splitTree p i s)") auto
from splitTree_correct[of s p, OF assms fmt] fmt
show ?thesis
by (blast intro: that)
qed
subsubsection \<open>Folding\<close>
fun foldl_node :: "('s \<Rightarrow> 'e \<times> 'a \<Rightarrow> 's) \<Rightarrow> 's \<Rightarrow> ('e,'a) Node \<Rightarrow> 's" where
"foldl_node f \<sigma> (Tip e a) = f \<sigma> (e,a)"|
"foldl_node f \<sigma> (Node2 _ a b) = foldl_node f (foldl_node f \<sigma> a) b"|
"foldl_node f \<sigma> (Node3 _ a b c) =
foldl_node f (foldl_node f (foldl_node f \<sigma> a) b) c"
primrec foldl_digit :: "('s \<Rightarrow> 'e \<times> 'a \<Rightarrow> 's) \<Rightarrow> 's \<Rightarrow> ('e,'a) Digit \<Rightarrow> 's" where
"foldl_digit f \<sigma> (One n1) = foldl_node f \<sigma> n1"|
"foldl_digit f \<sigma> (Two n1 n2) = foldl_node f (foldl_node f \<sigma> n1) n2"|
"foldl_digit f \<sigma> (Three n1 n2 n3) =
foldl_node f (foldl_node f (foldl_node f \<sigma> n1) n2) n3"|
"foldl_digit f \<sigma> (Four n1 n2 n3 n4) =
foldl_node f (foldl_node f (foldl_node f (foldl_node f \<sigma> n1) n2) n3) n4"
primrec foldr_node :: "('e \<times> 'a \<Rightarrow> 's \<Rightarrow> 's) \<Rightarrow> ('e,'a) Node \<Rightarrow> 's \<Rightarrow> 's" where
"foldr_node f (Tip e a) \<sigma> = f (e,a) \<sigma> "|
"foldr_node f (Node2 _ a b) \<sigma> = foldr_node f a (foldr_node f b \<sigma>)"|
"foldr_node f (Node3 _ a b c) \<sigma>
= foldr_node f a (foldr_node f b (foldr_node f c \<sigma>))"
primrec foldr_digit :: "('e \<times> 'a \<Rightarrow> 's \<Rightarrow> 's) \<Rightarrow> ('e,'a) Digit \<Rightarrow> 's \<Rightarrow> 's" where
"foldr_digit f (One n1) \<sigma> = foldr_node f n1 \<sigma>"|
"foldr_digit f (Two n1 n2) \<sigma> = foldr_node f n1 (foldr_node f n2 \<sigma>)"|
"foldr_digit f (Three n1 n2 n3) \<sigma> =
foldr_node f n1 (foldr_node f n2 (foldr_node f n3 \<sigma>))"|
"foldr_digit f (Four n1 n2 n3 n4) \<sigma> =
foldr_node f n1 (foldr_node f n2 (foldr_node f n3 (foldr_node f n4 \<sigma>)))"
lemma foldl_node_correct:
"foldl_node f \<sigma> nd = List.foldl f \<sigma> (nodeToList nd)"
by (induct nd arbitrary: "\<sigma>") (auto simp add: nodeToList_def)
lemma foldl_digit_correct:
"foldl_digit f \<sigma> d = List.foldl f \<sigma> (digitToList d)"
by (induct d arbitrary: "\<sigma>") (auto
simp add: digitToList_def foldl_node_correct)
lemma foldr_node_correct:
"foldr_node f nd \<sigma> = List.foldr f (nodeToList nd) \<sigma>"
by (induct nd arbitrary: "\<sigma>") (auto simp add: nodeToList_def)
lemma foldr_digit_correct:
"foldr_digit f d \<sigma> = List.foldr f (digitToList d) \<sigma>"
by (induct d arbitrary: "\<sigma>") (auto
simp add: digitToList_def foldr_node_correct)
text "Fold from left"
primrec foldl :: "('s \<Rightarrow> 'e \<times> 'a \<Rightarrow> 's) \<Rightarrow> 's \<Rightarrow> ('e,'a) FingerTreeStruc \<Rightarrow> 's"
where
"foldl f \<sigma> Empty = \<sigma>"|
"foldl f \<sigma> (Single nd) = foldl_node f \<sigma> nd"|
"foldl f \<sigma> (Deep _ d1 m d2) =
foldl_digit f (foldl f (foldl_digit f \<sigma> d1) m) d2"
lemma foldl_correct:
"foldl f \<sigma> t = List.foldl f \<sigma> (toList t)"
by (induct t arbitrary: "\<sigma>") (auto
simp add: toList_def foldl_node_correct foldl_digit_correct)
text "Fold from right"
primrec foldr :: "('e \<times> 'a \<Rightarrow> 's \<Rightarrow> 's) \<Rightarrow> ('e,'a) FingerTreeStruc \<Rightarrow> 's \<Rightarrow> 's"
where
"foldr f Empty \<sigma> = \<sigma>"|
"foldr f (Single nd) \<sigma> = foldr_node f nd \<sigma>"|
"foldr f (Deep _ d1 m d2) \<sigma>
= foldr_digit f d1 (foldr f m(foldr_digit f d2 \<sigma>))"
lemma foldr_correct:
"foldr f t \<sigma> = List.foldr f (toList t) \<sigma>"
by (induct t arbitrary: "\<sigma>") (auto
simp add: toList_def foldr_node_correct foldr_digit_correct)
subsubsection "Number of elements"
primrec count_node :: "('e, 'a) Node \<Rightarrow> nat" where
"count_node (Tip _ a) = 1" |
"count_node (Node2 _ a b) = count_node a + count_node b" |
"count_node (Node3 _ a b c) = count_node a + count_node b + count_node c"
primrec count_digit :: "('e,'a) Digit \<Rightarrow> nat" where
"count_digit (One a) = count_node a" |
"count_digit (Two a b) = count_node a + count_node b" |
"count_digit (Three a b c) = count_node a + count_node b + count_node c" |
"count_digit (Four a b c d)
= count_node a + count_node b + count_node c + count_node d"
lemma count_node_correct:
"count_node n = length (nodeToList n)"
by (induct n,auto simp add: nodeToList_def count_node_def)
lemma count_digit_correct:
"count_digit d = length (digitToList d)"
by (cases d, auto simp add: digitToList_def count_digit_def count_node_correct)
primrec count :: "('e,'a) FingerTreeStruc \<Rightarrow> nat" where
"count Empty = 0" |
"count (Single a) = count_node a" |
"count (Deep _ pr m sf) = count_digit pr + count m + count_digit sf"
lemma count_correct[simp]:
"count t = length (toList t)"
by (induct t,
auto simp add: toList_def count_def
count_digit_correct count_node_correct)
end
(* Expose finger tree functions as qualified names.
Generate code equations *)
interpretation FingerTreeStruc: FingerTreeStruc_loc .
(* Hide the concrete syntax *)
no_notation FingerTreeStruc.lcons (infixr "\<lhd>" 65)
no_notation FingerTreeStruc.rcons (infixl "\<rhd>" 65)
subsection "Hiding the invariant"
text_raw\<open>\label{sec:hide_invar}\<close>
text \<open>
In this section, we define the datatype of all FingerTrees that fulfill their
invariant, and define the operations to work on this datatype.
The advantage is, that the correctness lemmas do no longer contain
explicit invariant predicates, what makes them more handy to use.
\<close>
subsubsection "Datatype"
typedef (overloaded) ('e, 'a) FingerTree =
"{t :: ('e, 'a::monoid_add) FingerTreeStruc. FingerTreeStruc.ft_invar t}"
proof -
have "Empty \<in> ?FingerTree" by (simp)
then show ?thesis ..
qed
lemma Rep_FingerTree_invar[simp]: "FingerTreeStruc.ft_invar (Rep_FingerTree t)"
using Rep_FingerTree by simp
lemma [simp]:
"FingerTreeStruc.ft_invar t \<Longrightarrow> Rep_FingerTree (Abs_FingerTree t) = t"
using Abs_FingerTree_inverse by simp
lemma [simp, code abstype]: "Abs_FingerTree (Rep_FingerTree t) = t"
by (rule Rep_FingerTree_inverse)
typedef (overloaded) ('e,'a) viewres =
"{ r:: (('e \<times> 'a) \<times> ('e,'a::monoid_add) FingerTreeStruc) option .
case r of None \<Rightarrow> True | Some (a,t) \<Rightarrow> FingerTreeStruc.ft_invar t}"
apply (rule_tac x=None in exI)
apply auto
done
lemma [simp, code abstype]: "Abs_viewres (Rep_viewres x) = x"
by (rule Rep_viewres_inverse)
lemma Abs_viewres_inverse_None[simp]:
"Rep_viewres (Abs_viewres None) = None"
by (simp add: Abs_viewres_inverse)
lemma Abs_viewres_inverse_Some:
"FingerTreeStruc.ft_invar t \<Longrightarrow>
Rep_viewres (Abs_viewres (Some (a,t))) = Some (a,t)"
by (auto simp add: Abs_viewres_inverse)
definition [code]: "extract_viewres_isNone r == Rep_viewres r = None"
definition [code]: "extract_viewres_a r ==
case (Rep_viewres r) of Some (a,t) \<Rightarrow> a"
definition "extract_viewres_t r ==
case (Rep_viewres r) of None \<Rightarrow> Abs_FingerTree Empty
| Some (a,t) \<Rightarrow> Abs_FingerTree t"
lemma [code abstract]: "Rep_FingerTree (extract_viewres_t r) =
(case (Rep_viewres r) of None \<Rightarrow> Empty | Some (a,t) \<Rightarrow> t)"
apply (cases r)
apply (auto split: option.split option.split_asm
simp add: extract_viewres_t_def Abs_viewres_inverse_Some)
done
definition "extract_viewres r ==
if extract_viewres_isNone r then None
else Some (extract_viewres_a r, extract_viewres_t r)"
typedef (overloaded) ('e,'a) splitres =
"{ ((l,a,r):: (('e,'a) FingerTreeStruc \<times> ('e \<times> 'a) \<times> ('e,'a::monoid_add) FingerTreeStruc))
| l a r.
FingerTreeStruc.ft_invar l \<and> FingerTreeStruc.ft_invar r}"
apply (rule_tac x="(Empty,undefined,Empty)" in exI)
apply auto
done
lemma [simp, code abstype]: "Abs_splitres (Rep_splitres x) = x"
by (rule Rep_splitres_inverse)
lemma Abs_splitres_inverse:
"FingerTreeStruc.ft_invar r \<Longrightarrow> FingerTreeStruc.ft_invar s \<Longrightarrow>
Rep_splitres (Abs_splitres ((r,a,s))) = (r,a,s)"
by (auto simp add: Abs_splitres_inverse)
definition [code]: "extract_splitres_a r == case (Rep_splitres r) of (l,a,s) \<Rightarrow> a"
definition "extract_splitres_l r == case (Rep_splitres r) of (l,a,r) \<Rightarrow>
Abs_FingerTree l"
lemma [code abstract]: "Rep_FingerTree (extract_splitres_l r) = (case
(Rep_splitres r) of (l,a,r) \<Rightarrow> l)"
apply (cases r)
apply (auto split: option.split option.split_asm
simp add: extract_splitres_l_def Abs_splitres_inverse)
done
definition "extract_splitres_r r == case (Rep_splitres r) of (l,a,r) \<Rightarrow>
Abs_FingerTree r"
lemma [code abstract]: "Rep_FingerTree (extract_splitres_r r) = (case
(Rep_splitres r) of (l,a,r) \<Rightarrow> r)"
apply (cases r)
apply (auto split: option.split option.split_asm
simp add: extract_splitres_r_def Abs_splitres_inverse)
done
definition "extract_splitres r ==
(extract_splitres_l r,
extract_splitres_a r,
extract_splitres_r r)"
subsubsection "Definition of Operations"
locale FingerTree_loc
begin
definition [code]: "toList t == FingerTreeStruc.toList (Rep_FingerTree t)"
definition empty where "empty == Abs_FingerTree FingerTreeStruc.Empty"
lemma [code abstract]: "Rep_FingerTree empty = FingerTreeStruc.Empty"
by (simp add: empty_def)
lemma empty_rep: "t=empty \<longleftrightarrow> Rep_FingerTree t = Empty"
apply (auto simp add: empty_def)
apply (metis Rep_FingerTree_inverse)
done
definition [code]: "annot t == FingerTreeStruc.annot (Rep_FingerTree t)"
definition "toTree t == Abs_FingerTree (FingerTreeStruc.toTree t)"
lemma [code abstract]: "Rep_FingerTree (toTree t) = FingerTreeStruc.toTree t"
by (simp add: toTree_def)
definition "lcons a t ==
Abs_FingerTree (FingerTreeStruc.lcons a (Rep_FingerTree t))"
lemma [code abstract]:
"Rep_FingerTree (lcons a t) = (FingerTreeStruc.lcons a (Rep_FingerTree t))"
by (simp add: lcons_def FingerTreeStruc.lcons_correct)
definition "rcons t a ==
Abs_FingerTree (FingerTreeStruc.rcons (Rep_FingerTree t) a)"
lemma [code abstract]:
"Rep_FingerTree (rcons t a) = (FingerTreeStruc.rcons (Rep_FingerTree t) a)"
by (simp add: rcons_def FingerTreeStruc.rcons_correct)
definition "viewL_aux t ==
Abs_viewres (FingerTreeStruc.viewL (Rep_FingerTree t))"
definition "viewL t == extract_viewres (viewL_aux t)"
lemma [code abstract]:
"Rep_viewres (viewL_aux t) = (FingerTreeStruc.viewL (Rep_FingerTree t))"
apply (cases "(FingerTreeStruc.viewL (Rep_FingerTree t))")
apply (auto simp add: viewL_aux_def )
apply (cases "Rep_FingerTree t = Empty")
apply simp
apply (auto
elim!: FingerTreeStruc.viewL_correct_nonEmpty
[of "Rep_FingerTree t", simplified]
simp add: Abs_viewres_inverse_Some)
done
definition "viewR_aux t ==
Abs_viewres (FingerTreeStruc.viewR (Rep_FingerTree t))"
definition "viewR t == extract_viewres (viewR_aux t)"
lemma [code abstract]:
"Rep_viewres (viewR_aux t) = (FingerTreeStruc.viewR (Rep_FingerTree t))"
apply (cases "(FingerTreeStruc.viewR (Rep_FingerTree t))")
apply (auto simp add: viewR_aux_def )
apply (cases "Rep_FingerTree t = Empty")
apply simp
apply (auto
elim!: FingerTreeStruc.viewR_correct_nonEmpty
[of "Rep_FingerTree t", simplified]
simp add: Abs_viewres_inverse_Some)
done
definition [code]: "isEmpty t == FingerTreeStruc.isEmpty (Rep_FingerTree t)"
definition [code]: "head t = FingerTreeStruc.head (Rep_FingerTree t)"
definition "tail t \<equiv>
if t=empty then
empty
else
Abs_FingerTree (FingerTreeStruc.tail (Rep_FingerTree t))"
\<comment> \<open>Make function total, to allow abstraction\<close>
lemma [code abstract]: "Rep_FingerTree (tail t) =
(if (FingerTreeStruc.isEmpty (Rep_FingerTree t)) then Empty
else FingerTreeStruc.tail (Rep_FingerTree t))"
apply (simp add: tail_def FingerTreeStruc.tail_correct FingerTreeStruc.isEmpty_def empty_rep)
apply (auto simp add: empty_def)
done
definition [code]: "headR t = FingerTreeStruc.headR (Rep_FingerTree t)"
definition "tailR t \<equiv>
if t=empty then
empty
else
Abs_FingerTree (FingerTreeStruc.tailR (Rep_FingerTree t))"
lemma [code abstract]: "Rep_FingerTree (tailR t) =
(if (FingerTreeStruc.isEmpty (Rep_FingerTree t)) then Empty
else FingerTreeStruc.tailR (Rep_FingerTree t))"
apply (simp add: tailR_def FingerTreeStruc.tailR_correct FingerTreeStruc.isEmpty_def empty_rep)
apply (simp add: empty_def)
done
definition "app s t = Abs_FingerTree (
FingerTreeStruc.app (Rep_FingerTree s) (Rep_FingerTree t))"
lemma [code abstract]:
"Rep_FingerTree (app s t) =
FingerTreeStruc.app (Rep_FingerTree s) (Rep_FingerTree t)"
by (simp add: app_def FingerTreeStruc.app_correct)
definition "splitTree_aux p i t == if (\<not>p i \<and> p (i+annot t)) then
Abs_splitres (FingerTreeStruc.splitTree p i (Rep_FingerTree t))
else
Abs_splitres (Empty,undefined,Empty)"
definition "splitTree p i t == extract_splitres (splitTree_aux p i t)"
lemma [code abstract]:
"Rep_splitres (splitTree_aux p i t) = (if (\<not>p i \<and> p (i+annot t)) then
(FingerTreeStruc.splitTree p i (Rep_FingerTree t))
else
(Empty,undefined,Empty))"
using FingerTreeStruc.splitTree_invpres[of "Rep_FingerTree t" p i]
apply (auto simp add: splitTree_aux_def annot_def Abs_splitres_inverse)
apply (cases "FingerTreeStruc.splitTree p i (Rep_FingerTree t)")
apply (force simp add: Abs_FingerTree_inverse Abs_splitres_inverse)
done
definition foldl where
[code]: "foldl f \<sigma> t == FingerTreeStruc.foldl f \<sigma> (Rep_FingerTree t)"
definition foldr where
[code]: "foldr f t \<sigma> == FingerTreeStruc.foldr f (Rep_FingerTree t) \<sigma>"
definition count where
[code]: "count t == FingerTreeStruc.count (Rep_FingerTree t)"
subsubsection "Correctness statements"
lemma empty_correct: "toList t = [] \<longleftrightarrow> t=empty"
apply (unfold toList_def empty_rep)
apply (simp add: FingerTreeStruc.toList_empty)
done
lemma toList_of_empty[simp]: "toList empty = []"
apply (unfold toList_def empty_def)
apply (auto simp add: FingerTreeStruc.toList_empty)
done
lemma annot_correct: "annot t = sum_list (map snd (toList t))"
apply (unfold toList_def annot_def)
apply (simp add: FingerTreeStruc.annot_correct)
done
lemma toTree_correct: "toList (toTree l) = l"
apply (unfold toList_def toTree_def)
apply (simp add: FingerTreeStruc.toTree_correct)
done
lemma lcons_correct: "toList (lcons a t) = a#toList t"
apply (unfold toList_def lcons_def)
apply (simp add: FingerTreeStruc.lcons_correct)
done
lemma rcons_correct: "toList (rcons t a) = toList t@[a]"
apply (unfold toList_def rcons_def)
apply (simp add: FingerTreeStruc.rcons_correct)
done
lemma viewL_correct:
"t = empty \<Longrightarrow> viewL t = None"
"t \<noteq> empty \<Longrightarrow> \<exists>a s. viewL t = Some (a,s) \<and> toList t = a#toList s"
apply (unfold toList_def viewL_def viewL_aux_def
extract_viewres_def extract_viewres_isNone_def
extract_viewres_a_def
extract_viewres_t_def
empty_rep)
apply (simp add: FingerTreeStruc.viewL_correct)
apply (drule FingerTreeStruc.viewL_correct(2)[OF Rep_FingerTree_invar])
apply (auto simp add: Abs_viewres_inverse)
done
lemma viewL_empty[simp]: "viewL empty = None"
using viewL_correct by auto
lemma viewL_nonEmpty:
assumes "t\<noteq>empty"
obtains a s where "viewL t = Some (a,s)" "toList t = a#toList s"
using assms viewL_correct by blast
lemma viewR_correct:
"t = empty \<Longrightarrow> viewR t = None"
"t \<noteq> empty \<Longrightarrow> \<exists>a s. viewR t = Some (a,s) \<and> toList t = toList s@[a]"
apply (unfold toList_def viewR_def viewR_aux_def
extract_viewres_def extract_viewres_isNone_def
extract_viewres_a_def
extract_viewres_t_def
empty_rep)
apply (simp add: FingerTreeStruc.viewR_correct)
apply (drule FingerTreeStruc.viewR_correct(2)[OF Rep_FingerTree_invar])
apply (auto simp add: Abs_viewres_inverse)
done
lemma viewR_empty[simp]: "viewR empty = None"
using viewR_correct by auto
lemma viewR_nonEmpty:
assumes "t\<noteq>empty"
obtains a s where "viewR t = Some (a,s)" "toList t = toList s@[a]"
using assms viewR_correct by blast
lemma isEmpty_correct: "isEmpty t \<longleftrightarrow> t=empty"
apply (unfold toList_def isEmpty_def empty_rep)
apply (simp add: FingerTreeStruc.isEmpty_correct FingerTreeStruc.toList_empty)
done
lemma head_correct: "t\<noteq>empty \<Longrightarrow> head t = hd (toList t)"
apply (unfold toList_def head_def empty_rep)
apply (simp add: FingerTreeStruc.head_correct)
done
lemma tail_correct: "t\<noteq>empty \<Longrightarrow> toList (tail t) = tl (toList t)"
apply (unfold toList_def tail_def empty_rep)
apply (simp add: FingerTreeStruc.tail_correct)
done
lemma headR_correct: "t\<noteq>empty \<Longrightarrow> headR t = last (toList t)"
apply (unfold toList_def headR_def empty_rep)
apply (simp add: FingerTreeStruc.headR_correct)
done
lemma tailR_correct: "t\<noteq>empty \<Longrightarrow> toList (tailR t) = butlast (toList t)"
apply (unfold toList_def tailR_def empty_rep)
apply (simp add: FingerTreeStruc.tailR_correct)
done
lemma app_correct: "toList (app s t) = toList s @ toList t"
apply (unfold toList_def app_def)
apply (simp add: FingerTreeStruc.app_correct)
done
lemma splitTree_correct:
assumes mono: "\<forall>a b. p a \<longrightarrow> p (a + b)"
assumes init_ff: "\<not> p i"
assumes sum_tt: "p (i + annot s)"
assumes fmt: "(splitTree p i s) = (l, (e,a), r)"
shows "(toList s) = (toList l) @ (e,a) # (toList r)"
and "\<not> p (i + annot l)"
and "p (i + annot l + a)"
apply (rule
FingerTreeStruc.splitTree_correctE[
where p=p and s="Rep_FingerTree s",
OF _ mono init_ff sum_tt[unfolded annot_def],
simplified
])
using fmt
apply (unfold toList_def splitTree_aux_def splitTree_def annot_def
extract_splitres_def extract_splitres_l_def
extract_splitres_a_def extract_splitres_r_def) [1]
apply (auto split: if_split_asm prod.split_asm
simp add: init_ff sum_tt[unfolded annot_def] Abs_splitres_inverse) [1]
apply (rule
FingerTreeStruc.splitTree_correctE[
where p=p and s="Rep_FingerTree s",
OF _ mono init_ff sum_tt[unfolded annot_def],
simplified
])
using fmt
apply (unfold toList_def splitTree_aux_def splitTree_def annot_def
extract_splitres_def extract_splitres_l_def
extract_splitres_a_def extract_splitres_r_def) [1]
apply (auto split: if_split_asm prod.split_asm
simp add: init_ff sum_tt[unfolded annot_def] Abs_splitres_inverse) [1]
apply (rule
FingerTreeStruc.splitTree_correctE[
where p=p and s="Rep_FingerTree s",
OF _ mono init_ff sum_tt[unfolded annot_def],
simplified
])
using fmt
apply (unfold toList_def splitTree_aux_def splitTree_def annot_def
extract_splitres_def extract_splitres_l_def
extract_splitres_a_def extract_splitres_r_def) [1]
apply (auto split: if_split_asm prod.split_asm
simp add: init_ff sum_tt[unfolded annot_def] Abs_splitres_inverse) [1]
done
lemma splitTree_correctE:
assumes mono: "\<forall>a b. p a \<longrightarrow> p (a + b)"
assumes init_ff: "\<not> p i"
assumes sum_tt: "p (i + annot s)"
obtains l e a r where
"(splitTree p i s) = (l, (e,a), r)" and
"(toList s) = (toList l) @ (e,a) # (toList r)" and
"\<not> p (i + annot l)" and
"p (i + annot l + a)"
proof -
obtain l e a r where fmt: "(splitTree p i s) = (l, (e,a), r)"
by (cases "(splitTree p i s)") auto
from splitTree_correct[of p, OF assms fmt] fmt
show ?thesis
by (blast intro: that)
qed
lemma foldl_correct: "foldl f \<sigma> t = List.foldl f \<sigma> (toList t)"
apply (unfold toList_def foldl_def)
apply (simp add: FingerTreeStruc.foldl_correct)
done
lemma foldr_correct: "foldr f t \<sigma> = List.foldr f (toList t) \<sigma>"
apply (unfold toList_def foldr_def)
apply (simp add: FingerTreeStruc.foldr_correct)
done
lemma count_correct: "count t = length (toList t)"
apply (unfold toList_def count_def)
apply (simp add: FingerTreeStruc.count_correct)
done
end
interpretation FingerTree: FingerTree_loc .
text_raw\<open>\clearpage\<close>
subsection "Interface Documentation"
text_raw\<open>\label{sec:doc}\<close>
text \<open>
In this section, we list all supported operations on finger trees,
along with a short plaintext documentation and their correctness statements.
\<close>
(*#DOC
fun [no_spec] FingerTree.toList
Convert to list ($O(n)$)
fun FingerTree.empty
The empty finger tree ($O(1)$)
fun FingerTree.annot
Return sum of all annotations ($O(1)$)
fun FingerTree.toTree
Convert list to finger tree ($O(n\log(n))$)
fun FingerTree.lcons
Append element at the left end ($O(\log(n))$, $O(1)$ amortized)
fun FingerTree.rcons
Append element at the right end ($O(\log(n))$, $O(1)$ amortized)
fun FingerTree.viewL
Detach leftmost element ($O(\log(n))$, $O(1)$ amortized)
fun FingerTree.viewR
Detach rightmost element ($O(\log(n))$, $O(1)$ amortized)
fun FingerTree.isEmpty
Check whether tree is empty ($O(1)$)
fun FingerTree.head
Get leftmost element of non-empty tree ($O(\log(n))$)
fun FingerTree.tail
Get all but leftmost element of non-empty tree ($O(\log(n))$)
fun FingerTree.headR
Get rightmost element of non-empty tree ($O(\log(n))$)
fun FingerTree.tailR
Get all but rightmost element of non-empty tree ($O(\log(n))$)
fun FingerTree.app
Concatenate two finger trees ($O(\log(m+n))$)
fun [long_type] FingerTree.splitTree
Split tree by a monotone predicate. ($O(\log(n))$)
A predicate $p$ over the annotations is called monotone, iff, for all
annotations
$a,b$ with $p(a)$, we have already $p(a+b)$.
Splitting is done by specifying a monotone predicate $p$ that does not hold
for the initial value $i$ of the summation, but holds for $i$ plus the sum
of all annotations. The tree is then split at the position where $p$ starts to
hold for the sum of all elements up to that position.
fun [long_type] FingerTree.foldl
Fold with function from left
fun [long_type] FingerTree.foldr
Fold with function from right
fun FingerTree.count
Return the number of elements
*)
text \<open>
\underline{@{term_type "FingerTree.toList"}}\\
Convert to list ($O(n)$)\\
\underline{@{term_type "FingerTree.empty"}}\\
The empty finger tree ($O(1)$)\\
{\bf Spec} \<open>FingerTree.empty_correct\<close>:
@{thm [display] "FingerTree.empty_correct"}
\underline{@{term_type "FingerTree.annot"}}\\
Return sum of all annotations ($O(1)$)\\
{\bf Spec} \<open>FingerTree.annot_correct\<close>:
@{thm [display] "FingerTree.annot_correct"}
\underline{@{term_type "FingerTree.toTree"}}\\
Convert list to finger tree ($O(n\log(n))$)\\
{\bf Spec} \<open>FingerTree.toTree_correct\<close>:
@{thm [display] "FingerTree.toTree_correct"}
\underline{@{term_type "FingerTree.lcons"}}\\
Append element at the left end ($O(\log(n))$, $O(1)$ amortized)\\
{\bf Spec} \<open>FingerTree.lcons_correct\<close>:
@{thm [display] "FingerTree.lcons_correct"}
\underline{@{term_type "FingerTree.rcons"}}\\
Append element at the right end ($O(\log(n))$, $O(1)$ amortized)\\
{\bf Spec} \<open>FingerTree.rcons_correct\<close>:
@{thm [display] "FingerTree.rcons_correct"}
\underline{@{term_type "FingerTree.viewL"}}\\
Detach leftmost element ($O(\log(n))$, $O(1)$ amortized)\\
{\bf Spec} \<open>FingerTree.viewL_correct\<close>:
@{thm [display] "FingerTree.viewL_correct"}
\underline{@{term_type "FingerTree.viewR"}}\\
Detach rightmost element ($O(\log(n))$, $O(1)$ amortized)\\
{\bf Spec} \<open>FingerTree.viewR_correct\<close>:
@{thm [display] "FingerTree.viewR_correct"}
\underline{@{term_type "FingerTree.isEmpty"}}\\
Check whether tree is empty ($O(1)$)\\
{\bf Spec} \<open>FingerTree.isEmpty_correct\<close>:
@{thm [display] "FingerTree.isEmpty_correct"}
\underline{@{term_type "FingerTree.head"}}\\
Get leftmost element of non-empty tree ($O(\log(n))$)\\
{\bf Spec} \<open>FingerTree.head_correct\<close>:
@{thm [display] "FingerTree.head_correct"}
\underline{@{term_type "FingerTree.tail"}}\\
Get all but leftmost element of non-empty tree ($O(\log(n))$)\\
{\bf Spec} \<open>FingerTree.tail_correct\<close>:
@{thm [display] "FingerTree.tail_correct"}
\underline{@{term_type "FingerTree.headR"}}\\
Get rightmost element of non-empty tree ($O(\log(n))$)\\
{\bf Spec} \<open>FingerTree.headR_correct\<close>:
@{thm [display] "FingerTree.headR_correct"}
\underline{@{term_type "FingerTree.tailR"}}\\
Get all but rightmost element of non-empty tree ($O(\log(n))$)\\
{\bf Spec} \<open>FingerTree.tailR_correct\<close>:
@{thm [display] "FingerTree.tailR_correct"}
\underline{@{term_type "FingerTree.app"}}\\
Concatenate two finger trees ($O(\log(m+n))$)\\
{\bf Spec} \<open>FingerTree.app_correct\<close>:
@{thm [display] "FingerTree.app_correct"}
\underline{@{term "FingerTree.splitTree"}}
@{term_type [display] "FingerTree.splitTree"}
Split tree by a monotone predicate. ($O(\log(n))$)
A predicate $p$ over the annotations is called monotone, iff, for all
annotations
$a,b$ with $p(a)$, we have already $p(a+b)$.
Splitting is done by specifying a monotone predicate $p$ that does not hold
for the initial value $i$ of the summation, but holds for $i$ plus the sum
of all annotations. The tree is then split at the position where $p$ starts to
hold for the sum of all elements up to that position.\\
{\bf Spec} \<open>FingerTree.splitTree_correct\<close>:
@{thm [display] "FingerTree.splitTree_correct"}
\underline{@{term "FingerTree.foldl"}}
@{term_type [display] "FingerTree.foldl"}
Fold with function from left\\
{\bf Spec} \<open>FingerTree.foldl_correct\<close>:
@{thm [display] "FingerTree.foldl_correct"}
\underline{@{term "FingerTree.foldr"}}
@{term_type [display] "FingerTree.foldr"}
Fold with function from right\\
{\bf Spec} \<open>FingerTree.foldr_correct\<close>:
@{thm [display] "FingerTree.foldr_correct"}
\underline{@{term_type "FingerTree.count"}}\\
Return the number of elements\\
{\bf Spec} \<open>FingerTree.count_correct\<close>:
@{thm [display] "FingerTree.count_correct"}
\<close>
end
diff --git a/thys/Flyspeck-Tame/FaceDivisionProps.thy b/thys/Flyspeck-Tame/FaceDivisionProps.thy
--- a/thys/Flyspeck-Tame/FaceDivisionProps.thy
+++ b/thys/Flyspeck-Tame/FaceDivisionProps.thy
@@ -1,4955 +1,4966 @@
(* Author: Gertrud Bauer, Tobias Nipkow *)
section\<open>Properties of Face Division\<close>
theory FaceDivisionProps
imports Plane EnumeratorProps
begin
subsection\<open>Finality\<close>
(********************* makeFaceFinal ****************************)
lemma vertices_makeFaceFinal: "vertices(makeFaceFinal f g) = vertices g"
by(induct g)(simp add:vertices_graph_def makeFaceFinal_def)
lemma edges_makeFaceFinal: "\<E> (makeFaceFinal f g) = \<E> g"
proof -
{ fix fs
have "(\<Union>\<^bsub>f\<in>set (makeFaceFinalFaceList f fs)\<^esub> edges f) = (\<Union>\<^bsub>f\<in> set fs\<^esub> edges f)"
apply(unfold makeFaceFinalFaceList_def)
apply(induct f)
by(induct fs) simp_all }
thus ?thesis by(simp add:edges_graph_def makeFaceFinal_def)
qed
lemma in_set_repl_setFin:
"f \<in> set fs \<Longrightarrow> final f \<Longrightarrow> f \<in> set (replace f' [setFinal f'] fs)"
by (induct fs) auto
lemma in_set_repl: "f \<in> set fs \<Longrightarrow> f \<noteq> f' \<Longrightarrow> f \<in> set (replace f' fs' fs)"
by (induct fs) auto
lemma makeFaceFinals_preserve_finals:
"f \<in> set (finals g) \<Longrightarrow> f \<in> set (finals (makeFaceFinal f' g))"
by (induct g)
(simp add:makeFaceFinal_def finals_def makeFaceFinalFaceList_def
in_set_repl_setFin)
lemma len_faces_makeFaceFinal[simp]:
"|faces (makeFaceFinal f g)| = |faces g|"
by(simp add:makeFaceFinal_def makeFaceFinalFaceList_def)
lemma len_finals_makeFaceFinal:
"f \<in> \<F> g \<Longrightarrow> \<not> final f \<Longrightarrow> |finals (makeFaceFinal f g)| = |finals g| + 1"
by(simp add:makeFaceFinal_def finals_def makeFaceFinalFaceList_def
length_filter_replace1)
lemma len_nonFinals_makeFaceFinal:
"\<lbrakk> \<not> final f; f \<in> \<F> g\<rbrakk>
\<Longrightarrow> |nonFinals (makeFaceFinal f g)| = |nonFinals g| - 1"
by(simp add:makeFaceFinal_def nonFinals_def makeFaceFinalFaceList_def
length_filter_replace2)
lemma set_finals_makeFaceFinal[simp]: "distinct(faces g) \<Longrightarrow> f \<in> \<F> g \<Longrightarrow>
set(finals (makeFaceFinal f g)) = insert (setFinal f) (set(finals g))"
by(auto simp:finals_def makeFaceFinal_def makeFaceFinalFaceList_def
distinct_set_replace)
lemma splitFace_preserve_final:
"f \<in> set (finals g) \<Longrightarrow> \<not> final f' \<Longrightarrow>
f \<in> set (finals (snd (snd (splitFace g i j f' ns))))"
by (induct g) (auto simp add: splitFace_def finals_def split_def
intro: in_set_repl)
lemma splitFace_nonFinal_face:
"\<not> final (fst (snd (splitFace g i j f' ns)))"
by (simp add: splitFace_def split_def split_face_def)
lemma subdivFace'_preserve_finals:
"\<And>n i f' g. f \<in> set (finals g) \<Longrightarrow> \<not> final f' \<Longrightarrow>
f \<in> set (finals (subdivFace' g f' i n is))"
proof (induct "is")
case Nil then show ?case by(simp add:makeFaceFinals_preserve_finals)
next
case (Cons j "js") then show ?case
proof (cases j)
case None with Cons show ?thesis by simp
next
case (Some sj)
with Cons show ?thesis
by (auto simp: splitFace_preserve_final splitFace_nonFinal_face split_def)
qed
qed
lemma subdivFace_pres_finals:
"f \<in> set (finals g) \<Longrightarrow> \<not> final f' \<Longrightarrow>
f \<in> set (finals (subdivFace g f' is))"
by(simp add:subdivFace_def subdivFace'_preserve_finals)
declare Nat.diff_is_0_eq' [simp del]
(********************* section is_prefix ****************************)
subsection \<open>\<open>is_prefix\<close>\<close>
definition is_prefix :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"is_prefix ls vs \<equiv> (\<exists> bs. vs = ls @ bs)"
lemma is_prefix_add:
"is_prefix ls vs \<Longrightarrow> is_prefix (as @ ls) (as @ vs)" by (simp add: is_prefix_def)
lemma is_prefix_hd[simp]:
"is_prefix [l] vs = (l = hd vs \<and> vs \<noteq> [])"
apply (rule iffI) apply (auto simp: is_prefix_def)
apply (intro exI) apply (subgoal_tac "vs = hd vs # tl vs") apply assumption by auto
lemma is_prefix_f[simp]:
"is_prefix (a#as) (a#vs) = is_prefix as vs" by (auto simp: is_prefix_def)
lemma splitAt_is_prefix: "ram \<in> set vs \<Longrightarrow> is_prefix (fst (splitAt ram vs) @ [ram]) vs"
by (auto dest!: splitAt_ram simp: is_prefix_def)
(******************** section is_sublist *********************************)
subsection \<open>\<open>is_sublist\<close>\<close>
definition is_sublist :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"is_sublist ls vs \<equiv> (\<exists> as bs. vs = as @ ls @ bs)"
lemma is_prefix_sublist:
"is_prefix ls vs \<Longrightarrow> is_sublist ls vs" by (auto simp: is_prefix_def is_sublist_def)
lemma is_sublist_trans: "is_sublist as bs \<Longrightarrow> is_sublist bs cs \<Longrightarrow> is_sublist as cs"
apply (simp add: is_sublist_def) apply (elim exE)
apply (subgoal_tac "cs = (asaa @ asa) @ as @ (bsa @ bsaa)")
apply (intro exI) apply assumption by force
lemma is_sublist_add: "is_sublist as bs \<Longrightarrow> is_sublist as (xs @ bs @ ys)"
apply (simp add: is_sublist_def) apply (elim exE)
apply (subgoal_tac "xs @ bs @ ys = (xs @ asa) @ as @ (bsa @ ys)")
apply (intro exI) apply assumption by auto
lemma is_sublist_rec:
"is_sublist xs ys =
(if length xs > length ys then False else
if xs = take (length xs) ys then True else is_sublist xs (tl ys))"
proof (simp add:is_sublist_def, goal_cases)
case 1 show ?case
proof (standard, goal_cases)
case 1 show ?case
proof (standard, goal_cases)
case xs: 1
show ?case
proof (standard, goal_cases)
case 1 show ?case by auto
next
case 2 show ?case
proof (standard, goal_cases)
case 1
have "ys = take |xs| ys @ drop |xs| ys" by simp
also have "\<dots> = [] @ xs @ drop |xs| ys" by(simp add:xs[symmetric])
finally show ?case by blast
qed
qed
qed
next
case 2 show ?case
proof (standard, goal_cases)
case xs_neq: 1
show ?case
proof (standard, goal_cases)
case 1 show ?case by auto
next
case 2 show ?case
proof (standard, goal_cases)
case not_less: 1 show ?case
proof (standard, goal_cases)
case 1
then obtain as bs where ys: "ys = as @ xs @ bs" by blast
have "as \<noteq> []" using xs_neq ys by auto
then obtain a as' where "as = a # as'"
by (simp add:neq_Nil_conv) blast
hence "tl ys = as' @ xs @ bs" by(simp add:ys)
thus ?case by blast
next
case 2
then obtain as bs where ys: "tl ys = as @ xs @ bs" by blast
have "ys \<noteq> []" using xs_neq not_less by auto
then obtain y ys' where "ys = y # ys'"
by (simp add:neq_Nil_conv) blast
hence "ys = (y#as) @ xs @ bs" using ys by simp
thus ?case by blast
qed
qed
qed
qed
qed
qed
lemma not_sublist_len[simp]:
"|ys| < |xs| \<Longrightarrow> \<not> is_sublist xs ys"
by(simp add:is_sublist_rec)
lemma is_sublist_simp[simp]: "a \<noteq> v \<Longrightarrow> is_sublist (a#as) (v#vs) = is_sublist (a#as) vs"
proof
assume av: "a \<noteq> v" and subl: "is_sublist (a # as) (v # vs)"
then obtain rs ts where vvs: "v#vs = rs @ (a # as) @ ts" by (auto simp: is_sublist_def)
with av have "rs \<noteq> []" by auto
with vvs have "tl (v#vs) = tl rs @ a # as @ ts" by auto
then have "vs = tl rs @ a # as @ ts" by auto
then show "is_sublist (a # as) vs" by (auto simp: is_sublist_def)
next
assume av: "a \<noteq> v" and subl: "is_sublist (a # as) vs"
then show "is_sublist (a # as) (v # vs)" apply (auto simp: is_sublist_def) apply (intro exI)
apply (subgoal_tac "v # asa @ a # as @ bs = (v # asa) @ a # as @ bs") apply assumption by auto
qed
lemma is_sublist_id[simp]: "is_sublist vs vs" apply (auto simp: is_sublist_def) apply (intro exI)
apply (subgoal_tac "vs = [] @ vs @ []") by (assumption) auto
lemma is_sublist_in: "is_sublist (a#as) vs \<Longrightarrow> a \<in> set vs" by (auto simp: is_sublist_def)
lemma is_sublist_in1: "is_sublist [x,y] vs \<Longrightarrow> y \<in> set vs" by (auto simp: is_sublist_def)
lemma is_sublist_notlast[simp]: "distinct vs \<Longrightarrow> x = last vs \<Longrightarrow> \<not> is_sublist [x,y] vs"
proof
assume dvs: "distinct vs" and xl: "x = last vs" and subl:"is_sublist [x, y] vs"
then obtain rs ts where vs: "vs = rs @ x # y # ts" by (auto simp: is_sublist_def)
define as where "as = rs @ [x]"
define bs where "bs = y # ts"
then have bsne: "bs \<noteq> []" by auto
from as_def bs_def have vs2: "vs = as @ bs" using vs by auto
with as_def have xas: "x \<in> set as" by auto
from bsne vs2 have "last vs = last bs" by auto
with xl have "x = last bs" by auto
with bsne have "bs = (butlast bs) @ [x]" by auto
then have "x \<in> set bs" by (induct bs) auto
with xas vs2 dvs show False by auto
qed
lemma is_sublist_nth1: "is_sublist [x,y] ls \<Longrightarrow>
\<exists> i j. i < length ls \<and> j < length ls \<and> ls!i = x \<and> ls!j = y \<and> Suc i = j"
proof -
assume subl: "is_sublist [x,y] ls"
then obtain as bs where "ls = as @ x # y # bs" by (auto simp: is_sublist_def)
then have "(length as) < length ls \<and> (Suc (length as)) < length ls \<and> ls!(length as) = x
\<and> ls!(Suc (length as)) = y \<and> Suc (length as) = (Suc (length as))"
apply auto apply hypsubst_thin apply (induct as) by auto
then show ?thesis by auto
qed
lemma is_sublist_nth2: "\<exists> i j. i < length ls \<and> j < length ls \<and> ls!i = x \<and> ls!j = y \<and> Suc i = j \<Longrightarrow>
is_sublist [x,y] ls "
proof -
assume "\<exists> i j. i < length ls \<and> j < length ls \<and> ls!i = x \<and> ls!j = y \<and> Suc i = j"
then obtain i j where vors: "i < length ls \<and> j < length ls \<and> ls!i = x \<and> ls!j = y \<and> Suc i = j" by auto
then have "ls = take (Suc (Suc i)) ls @ drop (Suc (Suc i)) ls" by auto
with vors have "ls = take (Suc i) ls @ [ls! (Suc i)] @ drop (Suc (Suc i)) ls"
by (auto simp: take_Suc_conv_app_nth)
with vors have "ls = take i ls @ [ls!i] @ [ls! (Suc i)] @ drop (Suc (Suc i)) ls"
by (auto simp: take_Suc_conv_app_nth)
with vors show ?thesis by (auto simp: is_sublist_def)
qed
lemma is_sublist_tl: "is_sublist (a # as) vs \<Longrightarrow> is_sublist as vs" apply (simp add: is_sublist_def)
apply (elim exE) apply (intro exI)
apply (subgoal_tac "vs = (asa @ [a]) @ as @ bs") apply assumption by auto
lemma is_sublist_hd: "is_sublist (a # as) vs \<Longrightarrow> is_sublist [a] vs" apply (simp add: is_sublist_def) by auto
lemma is_sublist_hd_eq[simp]: "(is_sublist [a] vs) = (a \<in> set vs)" apply (rule_tac iffI)
apply (simp add: is_sublist_def) apply force
apply (simp add: is_sublist_def) apply (induct vs) apply force apply (case_tac "a = aa") apply force
apply (subgoal_tac "a \<in> set vs") apply simp apply (elim exE) apply (intro exI)
apply (subgoal_tac "aa # vs = (aa # as) @ a # bs") apply (assumption) by auto
lemma is_sublist_distinct_prefix:
"is_sublist (v#as) (v # vs) \<Longrightarrow> distinct (v # vs) \<Longrightarrow> is_prefix as vs"
proof -
assume d: "distinct (v # vs)" and subl: "is_sublist (v # as) (v # vs)"
from subl obtain rs ts where v_vs: "v # vs = rs @ (v # as) @ ts" by (simp add: is_sublist_def) auto
from d have v: "v \<notin> set vs" by auto
then have "\<not> is_sublist (v # as) vs" by (auto dest: is_sublist_hd)
with v_vs have "rs = []" apply (cases rs) by (auto simp: is_sublist_def)
with v_vs show "is_prefix as vs" by (auto simp: is_prefix_def)
qed
lemma is_sublist_distinct[intro]:
"is_sublist as vs \<Longrightarrow> distinct vs \<Longrightarrow> distinct as" by (auto simp: is_sublist_def)
lemma is_sublist_y_hd: "distinct vs \<Longrightarrow> y = hd vs \<Longrightarrow> \<not> is_sublist [x,y] vs"
proof
assume d: "distinct vs" and yh: "y = hd vs" and subl: "is_sublist [x, y] vs"
then obtain rs ts where vs: "vs = rs @ x # y # ts" by (auto simp: is_sublist_def)
define as where "as = rs @ [x]"
then have asne: "as \<noteq> []" by auto
define bs where "bs = y # ts"
then have bsne: "bs \<noteq> []" by auto
from as_def bs_def have vs2: "vs = as @ bs" using vs by auto
from bs_def have xbs: "y \<in> set bs" by auto
from vs2 asne have "hd vs = hd as" by simp
with yh have "y = hd as" by auto
with asne have "y \<in> set as" by (induct as) auto
with d xbs vs2 show False by auto
qed
lemma is_sublist_at1: "distinct (as @ bs) \<Longrightarrow> is_sublist [x,y] (as @ bs) \<Longrightarrow> x \<noteq> (last as) \<Longrightarrow>
is_sublist [x,y] as \<or> is_sublist [x,y] bs"
proof (cases "x \<in> set as")
assume d: "distinct (as @ bs)" and subl: "is_sublist [x, y] (as @ bs)" and xnl: "x \<noteq> last as"
define vs where "vs = as @ bs"
with d have dvs: "distinct vs" by auto
case True
with xnl subl have ind: "is_sublist (as@bs) vs \<Longrightarrow> is_sublist [x, y] as"
proof (induct as)
case Nil
then show ?case by force
next
case (Cons a as)
assume ih: "\<lbrakk>is_sublist (as@bs) vs; x \<noteq> last as; is_sublist [x,y] (as @ bs); x \<in> set as\<rbrakk> \<Longrightarrow>
is_sublist [x, y] as" and subl_aas_vs: "is_sublist ((a # as) @ bs) vs"
and xnl2: "x \<noteq> last (a # as)" and subl2: "is_sublist [x, y] ((a # as) @ bs)"
and x: "x \<in> set (a # as)"
then have rule1: "x \<noteq> a \<Longrightarrow> is_sublist [x,y] as" apply (cases "as = []") apply simp
apply (rule_tac ih) by (auto dest: is_sublist_tl)
from dvs subl_aas_vs have daas: "distinct (a # as @ bs)" apply (rule_tac is_sublist_distinct) by auto
from xnl2 have asne: "x = a \<Longrightarrow> as \<noteq> []" by auto
with subl2 daas have yhdas: "x = a \<Longrightarrow> y = hd as" apply simp apply (drule_tac is_sublist_distinct_prefix) by auto
with asne have "x = a \<Longrightarrow> as = y # tl as" by auto
with asne yhdas have "x = a \<Longrightarrow> is_prefix [x,y] (a # as)" by auto
then have rule2: "x = a \<Longrightarrow> is_sublist [x,y] (a # as)" by (simp add: is_prefix_sublist)
from rule1 rule2 show ?case by (cases "x = a") auto
qed
from vs_def d have "is_sublist [x, y] as" by (rule_tac ind) auto
then show ?thesis by auto
next
assume d: "distinct (as @ bs)" and subl: "is_sublist [x, y] (as @ bs)" and xnl: "x \<noteq> last as"
define ars where "ars = as"
case False
with ars_def have xars: "x \<notin> set ars" by auto
from subl have ind: "is_sublist as ars \<Longrightarrow> is_sublist [x, y] bs"
proof (induct as)
case Nil
then show ?case by auto
next
case (Cons a as)
assume ih: "\<lbrakk>is_sublist as ars; is_sublist [x, y] (as @ bs)\<rbrakk> \<Longrightarrow> is_sublist [x, y] bs"
and subl_aasbsvs: "is_sublist (a # as) ars" and subl2: "is_sublist [x, y] ((a # as) @ bs)"
from subl_aasbsvs ars_def False have "x \<noteq> a" by (auto simp:is_sublist_in)
with subl_aasbsvs subl2 show ?thesis apply (rule_tac ih) by (auto dest: is_sublist_tl)
qed
from ars_def have "is_sublist [x, y] bs" by (rule_tac ind) auto
then show ?thesis by auto
qed
lemma is_sublist_at4: "distinct (as @ bs) \<Longrightarrow> is_sublist [x,y] (as @ bs) \<Longrightarrow>
as \<noteq> [] \<Longrightarrow> x = last as \<Longrightarrow> y = hd bs"
proof -
assume d: "distinct (as @ bs)" and subl: "is_sublist [x,y] (as @ bs)"
and asne: "as \<noteq> []" and xl: "x = last as"
define vs where "vs = as @ bs"
with subl have "is_sublist [x,y] vs" by auto
then obtain rs ts where vs2: "vs = rs @ x # y # ts" by (auto simp: is_sublist_def)
from vs_def d have dvs:"distinct vs" by auto
from asne xl have as:"as = butlast as @ [x]" by auto
with vs_def have vs3: "vs = butlast as @ x # bs" by auto
from dvs vs2 vs3 have "rs = butlast as" apply (rule_tac dist_at1) by auto
then have "rs @ [x] = butlast as @ [x]" by auto
with as have "rs @ [x] = as" by auto
then have "as = rs @ [x]" by auto
with vs2 vs_def have "bs = y # ts" by auto
then show ?thesis by auto
qed
lemma is_sublist_at5: "distinct (as @ bs) \<Longrightarrow> is_sublist [x,y] (as @ bs) \<Longrightarrow>
is_sublist [x,y] as \<or> is_sublist [x,y] bs \<or> x = last as \<and> y = hd bs"
apply (case_tac "as = []") apply simp apply (cases "x = last as")
apply (subgoal_tac "y = hd bs") apply simp
apply (rule is_sublist_at4) apply assumption+ (*apply force+ *)
apply (drule_tac is_sublist_at1) by auto
lemma is_sublist_rev: "is_sublist [a,b] (rev zs) = is_sublist [b,a] zs"
apply (simp add: is_sublist_def)
apply (intro iffI) apply (elim exE) apply (intro exI)
apply (subgoal_tac "zs = (rev bs) @ b # a # rev as") apply assumption
apply (subgoal_tac "rev (rev zs) = rev (as @ a # b # bs)")
apply (thin_tac "rev zs = as @ a # b # bs") apply simp
apply simp
apply (elim exE) apply (intro exI) by force
lemma is_sublist_at5'[simp]:
"distinct as \<Longrightarrow> distinct bs \<Longrightarrow> set as \<inter> set bs = {} \<Longrightarrow> is_sublist [x,y] (as @ bs) \<Longrightarrow>
is_sublist [x,y] as \<or> is_sublist [x,y] bs \<or> x = last as \<and> y = hd bs"
apply (subgoal_tac "distinct (as @ bs)") apply (drule is_sublist_at5) by auto
lemma splitAt_is_sublist1R[simp]: "ram \<in> set vs \<Longrightarrow> is_sublist (fst (splitAt ram vs) @ [ram]) vs"
apply (auto dest!: splitAt_ram simp: is_sublist_def) apply (intro exI)
apply (subgoal_tac "vs = [] @ fst (splitAt ram vs) @ ram # snd (splitAt ram vs)") apply assumption by simp
lemma splitAt_is_sublist2R[simp]: "ram \<in> set vs \<Longrightarrow> is_sublist (ram # snd (splitAt ram vs)) vs"
apply (auto dest!: splitAt_ram splitAt_no_ram simp: is_sublist_def) apply (intro exI)
apply (subgoal_tac "vs = fst (splitAt ram vs) @ ram # snd (splitAt ram vs) @ []") apply assumption by auto
(*********************** section is_nextElem *****************************)
subsection \<open>\<open>is_nextElem\<close>\<close>
definition is_nextElem :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
"is_nextElem xs x y \<equiv> is_sublist [x,y] xs \<or> xs \<noteq> [] \<and> x = last xs \<and> y = hd xs"
lemma is_nextElem_a[intro]: "is_nextElem vs a b \<Longrightarrow> a \<in> set vs"
by (auto simp: is_nextElem_def is_sublist_def)
lemma is_nextElem_b[intro]: "is_nextElem vs a b \<Longrightarrow> b \<in> set vs"
by (auto simp: is_nextElem_def is_sublist_def)
lemma is_nextElem_last_hd[intro]: "distinct vs \<Longrightarrow> is_nextElem vs x y \<Longrightarrow>
x = last vs \<Longrightarrow> y = hd vs"
by (auto simp: is_nextElem_def)
lemma is_nextElem_last_ne[intro]: "distinct vs \<Longrightarrow> is_nextElem vs x y \<Longrightarrow>
x = last vs \<Longrightarrow> vs \<noteq> []"
by (auto simp: is_nextElem_def)
lemma is_nextElem_sublistI: "is_sublist [x,y] vs \<Longrightarrow> is_nextElem vs x y"
by (auto simp: is_nextElem_def)
lemma is_nextElem_nth1: "is_nextElem ls x y \<Longrightarrow> \<exists> i j. i < length ls
\<and> j < length ls \<and> ls!i = x \<and> ls!j = y \<and> (Suc i) mod (length ls) = j"
proof (cases "is_sublist [x,y] ls")
assume is_nextElem: "is_nextElem ls x y"
case True then show ?thesis apply (drule_tac is_sublist_nth1) by auto
next
assume is_nextElem: "is_nextElem ls x y"
case False with is_nextElem have hl: "ls \<noteq> [] \<and> last ls = x \<and> hd ls = y"
by (auto simp: is_nextElem_def)
then have j: "ls!0 = y" by (cases ls) auto
from hl have i: "ls!(length ls - 1) = x" by (cases ls rule: rev_exhaust) auto
from i j hl have "(length ls - 1) < length ls \<and> 0 < length ls \<and> ls!(length ls - 1) = x
\<and> ls!0 = y \<and> (Suc (length ls - 1)) mod (length ls) = 0" by auto
then show ?thesis apply (intro exI) .
qed
lemma is_nextElem_nth2: " \<exists> i j. i < length ls \<and> j < length ls \<and> ls!i = x \<and> ls!j = y
\<and> (Suc i) mod (length ls) = j \<Longrightarrow> is_nextElem ls x y"
proof -
assume "\<exists> i j. i < length ls \<and> j < length ls \<and> ls!i = x \<and> ls!j = y \<and> (Suc i) mod (length ls) = j"
then obtain i j where vors: "i < length ls \<and> j < length ls \<and> ls!i = x \<and> ls!j = y
\<and> (Suc i) mod (length ls) = j" by auto
then show ?thesis
proof (cases "Suc i = length ls")
case True with vors have "j = 0" by auto
(*ls ! i = last ls*)
with True vors show ?thesis apply (auto simp: is_nextElem_def)
apply (cases ls rule: rev_exhaust) apply auto apply (cases ls) by auto
next
case False with vors have "is_sublist [x,y] ls"
apply (rule_tac is_sublist_nth2) by auto
then show ?thesis by (simp add: is_nextElem_def)
qed
qed
lemma is_nextElem_rotate1_aux:
"is_nextElem (rotate m ls) x y \<Longrightarrow> is_nextElem ls x y"
proof -
assume is_nextElem: "is_nextElem (rotate m ls) x y"
define n where "n = m mod length ls"
then have rot_eq: "rotate m ls = rotate n ls"
by (auto intro: rotate_conv_mod)
with is_nextElem have "is_nextElem (rotate n ls) x y"
by simp
then obtain i j where vors:"i < length (rotate n ls) \<and> j < length (rotate n ls) \<and>
(rotate n ls)!i = x \<and> (rotate n ls)!j = y \<and>
(Suc i) mod (length (rotate n ls)) = j"
by (drule_tac is_nextElem_nth1) auto
then have lls: "0 < length ls"
by auto
define k where "k = (i+n) mod (length ls)"
with lls have sk: "k < length ls"
by simp
from k_def lls vors have "ls!k = (rotate n ls)!(i mod (length ls))"
by (simp add: nth_rotate)
with vors have lsk: "ls!k = x"
by simp
define l where "l = (j+n) mod (length ls)"
with lls have sl: "l < length ls"
by simp
from l_def lls vors have "ls!l = (rotate n ls)!(j mod (length ls))"
by (simp add: nth_rotate)
with vors have lsl: "ls!l = y"
by simp
from vors k_def l_def
have "(Suc i) mod length ls = j"
by simp
then have "(Suc i) mod length ls = j mod length ls"
by auto
then have "((Suc i) mod length ls + n mod (length ls)) mod length ls
= (j mod length ls + n mod (length ls)) mod length ls"
by simp
then have "((Suc i) + n) mod length ls = (j + n) mod length ls"
by (simp add: mod_simps)
with vors k_def l_def have "(Suc k) mod (length ls) = l"
by (simp add: mod_simps)
with sk lsk sl lsl
show ?thesis
by (auto intro: is_nextElem_nth2)
qed
lemma is_nextElem_rotate_eq[simp]: "is_nextElem (rotate m ls) x y = is_nextElem ls x y"
apply (auto dest: is_nextElem_rotate1_aux) apply (rule is_nextElem_rotate1_aux)
apply (subgoal_tac "is_nextElem (rotate (length ls - m mod length ls) (rotate m ls)) x y")
apply assumption by simp
lemma is_nextElem_congs_eq: "ls \<cong> ms \<Longrightarrow> is_nextElem ls x y = is_nextElem ms x y"
by (auto simp: congs_def)
lemma is_nextElem_rev[simp]: "is_nextElem (rev zs) a b = is_nextElem zs b a"
apply (simp add: is_nextElem_def is_sublist_rev)
apply (case_tac "zs = []") apply simp apply simp
apply (case_tac "a = hd zs") apply (case_tac "zs") apply simp apply simp apply simp
apply (case_tac "a = last (rev zs) \<and> b = last zs") apply simp
apply (case_tac "zs" rule: rev_exhaust) apply simp
apply (case_tac "ys") apply simp apply simp by force
lemma is_nextElem_circ:
"\<lbrakk> distinct xs; is_nextElem xs a b; is_nextElem xs b a \<rbrakk> \<Longrightarrow> |xs| \<le> 2"
apply(drule is_nextElem_nth1)
apply(drule is_nextElem_nth1)
apply (clarsimp)
apply(rename_tac i j)
apply(frule_tac i=j and j = "Suc i mod |xs|" in nth_eq_iff_index_eq)
apply assumption+
apply(frule_tac j=i and i = "Suc j mod |xs|" in nth_eq_iff_index_eq)
apply assumption+
apply(rule ccontr)
apply(simp add: distinct_conv_nth mod_Suc)
done
subsection\<open>\<open>nextElem, sublist, is_nextElem\<close>\<close>
lemma is_sublist_eq: "distinct vs \<Longrightarrow> c \<noteq> y \<Longrightarrow>
(nextElem vs c x = y) = is_sublist [x,y] vs"
proof -
assume d: "distinct vs" and c: "c \<noteq> y"
have r1: "nextElem vs c x = y \<Longrightarrow> is_sublist [x,y] vs"
proof -
assume fn: "nextElem vs c x = y"
with c show ?thesis by(drule_tac nextElem_cases)(auto simp: is_sublist_def)
qed
with d have r2: "is_sublist [x,y] vs \<Longrightarrow> nextElem vs c x = y"
apply (simp add: is_sublist_def) apply (elim exE) by auto
show ?thesis apply (intro iffI r1) by (auto intro: r2)
qed
lemma is_nextElem1: "distinct vs \<Longrightarrow> x \<in> set vs \<Longrightarrow> nextElem vs (hd vs) x = y \<Longrightarrow> is_nextElem vs x y"
proof -
assume d: "distinct vs" and x: "x \<in> set vs" and fn: "nextElem vs (hd vs) x = y"
from x have r0: "vs \<noteq> []" by auto
from d fn have r1: "x = last vs \<Longrightarrow> y = hd vs" by (auto)
from d fn have r3: "hd vs \<noteq> y \<Longrightarrow> (\<exists> a b. vs = a @ [x,y] @ b)" by (drule_tac nextElem_cases) auto
from x obtain n where xn:"x = vs!n" and nl: "n < length vs" by (auto simp: in_set_conv_nth)
define as where "as = take n vs"
define bs where "bs = drop (Suc n) vs"
from as_def bs_def xn nl have vs:"vs = as @ [x] @ bs" by (auto intro: id_take_nth_drop)
then have r2: "x \<noteq> last vs \<Longrightarrow> y \<noteq> hd vs"
proof -
assume notx: "x \<noteq> last vs"
from vs notx have "bs \<noteq> []" by auto
with vs have r2: "vs = as @ [x, hd bs] @ tl bs" by auto
with d have ineq: "hd bs \<noteq> hd vs" by (cases as) auto
from d fn r2 have "y = hd bs" by auto
with ineq show ?thesis by auto
qed
from r0 r1 r2 r3 show ?thesis apply (simp add:is_nextElem_def is_sublist_def)
apply (cases "x = last vs") by auto
qed
lemma is_nextElem2: "distinct vs \<Longrightarrow> x \<in> set vs \<Longrightarrow> is_nextElem vs x y \<Longrightarrow> nextElem vs (hd vs) x = y"
proof -
assume d: "distinct vs" and x: "x \<in> set vs" and is_nextElem: "is_nextElem vs x y"
then show ?thesis apply (simp add: is_nextElem_def) apply (cases "is_sublist [x,y] vs")
apply (cases "y = hd vs")
apply (simp add: is_sublist_def) apply (force dest: distinct_hd_not_cons)
apply (subgoal_tac "hd vs \<noteq> y") apply (simp add: is_sublist_eq) by auto
qed
lemma nextElem_is_nextElem:
"distinct xs \<Longrightarrow> x \<in> set xs \<Longrightarrow>
is_nextElem xs x y = (nextElem xs (hd xs) x = y)"
by (auto intro!: is_nextElem1 is_nextElem2)
lemma nextElem_congs_eq: "xs \<cong> ys \<Longrightarrow> distinct xs \<Longrightarrow> x \<in> set xs \<Longrightarrow>
nextElem xs (hd xs) x = nextElem ys (hd ys) x"
proof -
assume eq: "xs \<cong> ys" and dist: "distinct xs" and x: "x \<in> set xs"
define y where "y = nextElem xs (hd xs) x"
then have f1:"nextElem xs (hd xs) x = y" by auto
with dist x have "is_nextElem xs x y" by (auto intro: is_nextElem1)
with eq have "is_nextElem ys x y" by (simp add:is_nextElem_congs_eq)
with eq dist x have f2:"nextElem ys (hd ys) x = y"
by (auto simp: congs_distinct intro: is_nextElem2)
from f1 f2 show ?thesis by auto
qed
lemma is_sublist_is_nextElem: "distinct vs \<Longrightarrow> is_nextElem vs x y \<Longrightarrow> is_sublist as vs \<Longrightarrow> x \<in> set as \<Longrightarrow> x \<noteq> last as \<Longrightarrow> is_sublist [x,y] as"
proof -
assume d: "distinct vs" and is_nextElem: "is_nextElem vs x y" and subl: "is_sublist as vs" and xin: "x \<in> set as" and xnl: "x \<noteq> last as"
from xin have asne: "as \<noteq> []" by auto
with subl have vsne: "vs \<noteq> []" by (auto simp: is_sublist_def)
from subl obtain rs ts where vs: "vs = rs @ as @ ts" apply (simp add: is_sublist_def) apply (elim exE) by auto
with d xnl asne have "x \<noteq> last vs"
proof (cases "ts = []")
case True
with d xnl asne vs show ?thesis by force
next
define lastvs where "lastvs = last ts"
case False
with vs lastvs_def have vs2: "vs = rs @ as @ butlast ts @ [lastvs]" by auto
with d have "lastvs \<notin> set as" by auto
with xin have "lastvs \<noteq> x" by auto
with vs2 show ?thesis by auto
qed
with is_nextElem have subl_vs: "is_sublist [x,y] vs" by (auto simp: is_nextElem_def)
from d xin vs have "\<not> is_sublist [x] rs" by auto
then have nrs: "\<not> is_sublist [x,y] rs" by (auto dest: is_sublist_hd)
from d xin vs have "\<not> is_sublist [x] ts" by auto
then have nts: "\<not> is_sublist [x,y] ts" by (auto dest: is_sublist_hd)
from d xin vs have xnrs: "x \<notin> set rs" by auto
then have notrs: "\<not> is_sublist [x,y] rs" by (auto simp:is_sublist_in)
from xnrs have xnlrs: "rs \<noteq> [] \<Longrightarrow> x \<noteq> last rs" by (induct rs) auto
from d xin vs have xnts: "x \<notin> set ts" by auto
then have notts: "\<not> is_sublist [x,y] ts" by (auto simp:is_sublist_in)
from d vs subl_vs have "is_sublist [x,y] rs \<or> is_sublist [x,y] (as@ts)" apply (cases "rs = []") apply simp apply (rule_tac is_sublist_at1) by (auto intro!: xnlrs)
with notrs have "is_sublist [x,y] (as@ts)" by auto
with d vs xnl have "is_sublist [x,y] as \<or> is_sublist [x,y] ts" apply (rule_tac is_sublist_at1) by auto
with notts show "is_sublist [x,y] as" by auto
qed
subsection \<open>\<open>before\<close>\<close>
definition before :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
"before vs ram1 ram2 \<equiv> \<exists> a b c. vs = a @ ram1 # b @ ram2 # c"
lemma before_dist_fst_fst[simp]: "before vs ram1 ram2 \<Longrightarrow> distinct vs \<Longrightarrow> fst (splitAt ram2 (fst (splitAt ram1 vs))) = fst (splitAt ram1 (fst (splitAt ram2 vs)))"
apply (simp add: before_def) apply (elim exE)
apply (drule splitAt_dist_ram_all) by (auto dest!: pairD)
lemma before_dist_fst_snd[simp]: "before vs ram1 ram2 \<Longrightarrow> distinct vs \<Longrightarrow> fst (splitAt ram2 (snd (splitAt ram1 vs))) = snd (splitAt ram1 (fst (splitAt ram2 vs)))"
apply (simp add: before_def) apply (elim exE)
apply (drule_tac splitAt_dist_ram_all) by (auto dest!: pairD)
lemma before_dist_snd_fst[simp]: "before vs ram1 ram2 \<Longrightarrow> distinct vs \<Longrightarrow> snd (splitAt ram2 (fst (splitAt ram1 vs))) = snd (splitAt ram1 (snd (splitAt ram2 vs)))"
apply (simp add: before_def) apply (elim exE)
apply (drule_tac splitAt_dist_ram_all) by (auto dest!: pairD)
lemma before_dist_snd_snd[simp]: "before vs ram1 ram2 \<Longrightarrow> distinct vs \<Longrightarrow> snd (splitAt ram2 (snd (splitAt ram1 vs))) = fst (splitAt ram1 (snd (splitAt ram2 vs)))"
apply (simp add: before_def) apply (elim exE)
apply (drule_tac splitAt_dist_ram_all) by (auto dest!: pairD)
lemma before_dist_snd[simp]: "before vs ram1 ram2 \<Longrightarrow> distinct vs \<Longrightarrow> fst (splitAt ram1 (snd (splitAt ram2 vs))) = snd (splitAt ram2 vs)"
apply (simp add: before_def) apply (elim exE)
apply (drule_tac splitAt_dist_ram_all) by (auto dest!: pairD)
lemma before_dist_fst[simp]: "before vs ram1 ram2 \<Longrightarrow> distinct vs \<Longrightarrow> fst (splitAt ram1 (fst (splitAt ram2 vs))) = fst (splitAt ram1 vs)"
apply (simp add: before_def) apply (elim exE)
apply (drule_tac splitAt_dist_ram_all) by (auto dest!: pairD)
lemma before_or: "ram1 \<in> set vs \<Longrightarrow> ram2 \<in> set vs \<Longrightarrow> ram1 \<noteq> ram2 \<Longrightarrow> before vs ram1 ram2 \<or> before vs ram2 ram1"
proof -
assume r1: "ram1 \<in> set vs" and r2: "ram2 \<in> set vs" and r12: "ram1 \<noteq> ram2"
then show ?thesis
proof (cases "ram2 \<in> set (snd (splitAt ram1 vs))")
define a where "a = fst (splitAt ram1 vs)"
define b where "b = fst (splitAt ram2 (snd (splitAt ram1 vs)))"
define c where "c = snd (splitAt ram2 (snd (splitAt ram1 vs)))"
case True with r1 a_def b_def c_def have "vs = a @ [ram1] @ b @ [ram2] @ c"
by (auto dest!: splitAt_ram)
then show ?thesis apply (simp add: before_def) by auto
next
define ab where "ab = fst (splitAt ram1 vs)"
case False
with r1 r2 r12 ab_def have r2': "ram2 \<in> set ab" by (auto intro: splitAt_ram3)
define a where "a = fst (splitAt ram2 ab)"
define b where "b = snd (splitAt ram2 ab)"
define c where "c = snd (splitAt ram1 vs)"
from r1 ab_def c_def have "vs = ab @ [ram1] @ c" by (auto dest!: splitAt_ram)
with r2' a_def b_def have "vs = (a @ [ram2] @ b) @ [ram1] @ c" by (drule_tac splitAt_ram) simp
then show ?thesis apply (simp add: before_def) apply (rule disjI2) by auto
qed
qed
lemma before_r1:
"before vs r1 r2 \<Longrightarrow> r1 \<in> set vs" by (auto simp: before_def)
lemma before_r2:
"before vs r1 r2 \<Longrightarrow> r2 \<in> set vs" by (auto simp: before_def)
lemma before_dist_r2:
"distinct vs \<Longrightarrow> before vs r1 r2 \<Longrightarrow> r2 \<in> set (snd (splitAt r1 vs))"
proof -
assume d: "distinct vs" and b: "before vs r1 r2"
from d b have ex1: "\<exists>! s. (vs = (fst s) @ r1 # snd (s))" apply (drule_tac before_r1) apply (rule distinct_unique1) by auto
from d b ex1 show ?thesis apply (unfold before_def)
proof (elim exE ex1E)
fix a b c s
assume vs: "vs = a @ r1 # b @ r2 # c" and "\<forall>y. vs = fst y @ r1 # snd y \<longrightarrow> y = s"
then have "\<And> y. vs = fst y @ r1 # snd y \<longrightarrow> y = s" by (clarify, hypsubst_thin, auto)
then have single: "\<And> y. vs = fst y @ r1 # snd y \<Longrightarrow> y = s" by auto
define bc where "bc = b @ r2 # c"
with vs have vs2: "vs = a @ r1 # bc" by auto
from bc_def have r2: "r2 \<in> set bc" by auto
define t where "t = (a,bc)"
with vs2 have vs3: "vs = fst (t) @ r1 # snd (t)" by auto
with single have ts: "t = s" by (rule_tac single) auto
from b have "splitAt r1 vs = s" apply (drule_tac before_r1) apply (drule_tac splitAt_ram) by (rule single) auto
with ts have "t = splitAt r1 vs" by simp
with t_def have "bc = snd(splitAt r1 vs)" by simp
with r2 show ?thesis by simp
qed
qed
lemma before_dist_not_r2[intro]:
"distinct vs \<Longrightarrow> before vs r1 r2 \<Longrightarrow> r2 \<notin> set (fst (splitAt r1 vs))" apply (frule before_dist_r2) by (auto dest: splitAt_distinct_fst_snd)
lemma before_dist_r1:
"distinct vs \<Longrightarrow> before vs r1 r2 \<Longrightarrow> r1 \<in> set (fst (splitAt r2 vs))"
proof -
assume d: "distinct vs" and b: "before vs r1 r2"
from d b have ex1: "\<exists>! s. (vs = (fst s) @ r2 # snd (s))" apply (drule_tac before_r2) apply (rule distinct_unique1) by auto
from d b ex1 show ?thesis apply (unfold before_def)
proof (elim exE ex1E)
fix a b c s
assume vs: "vs = a @ r1 # b @ r2 # c" and "\<forall>y. vs = fst y @ r2 # snd y \<longrightarrow> y = s"
then have "\<And> y. vs = fst y @ r2 # snd y \<longrightarrow> y = s" by (clarify, hypsubst_thin, auto)
then have single: "\<And> y. vs = fst y @ r2 # snd y \<Longrightarrow> y = s" by auto
define ab where "ab = a @ r1 # b"
with vs have vs2: "vs = ab @ r2 # c" by auto
from ab_def have r1: "r1 \<in> set ab" by auto
define t where "t = (ab,c)"
with vs2 have vs3: "vs = fst (t) @ r2 # snd (t)" by auto
with single have ts: "t = s" by (rule_tac single) auto
from b have "splitAt r2 vs = s" apply (drule_tac before_r2) apply (drule_tac splitAt_ram) by (rule single) auto
with ts have "t = splitAt r2 vs" by simp
with t_def have "ab = fst(splitAt r2 vs)" by simp
with r1 show ?thesis by simp
qed
qed
lemma before_dist_not_r1[intro]:
"distinct vs \<Longrightarrow> before vs r1 r2 \<Longrightarrow> r1 \<notin> set (snd (splitAt r2 vs))" apply (frule before_dist_r1) by (auto dest: splitAt_distinct_fst_snd)
lemma before_snd:
"r2 \<in> set (snd (splitAt r1 vs)) \<Longrightarrow> before vs r1 r2"
proof -
assume r2: "r2 \<in> set (snd (splitAt r1 vs))"
from r2 have r1: "r1 \<in> set vs" apply (rule_tac ccontr) apply (drule splitAt_no_ram) by simp
define a where "a = fst (splitAt r1 vs)"
define bc where "bc = snd (splitAt r1 vs)"
define b where "b = fst (splitAt r2 bc)"
define c where "c = snd (splitAt r2 bc)"
from r1 a_def bc_def have vs: "vs = a @ [r1] @ bc" by (auto dest: splitAt_ram)
from r2 bc_def have r2: "r2 \<in> set bc" by simp
with b_def c_def have "bc = b @ [r2] @ c" by (auto dest: splitAt_ram)
with vs show ?thesis by (simp add: before_def) auto
qed
lemma before_fst:
"r2 \<in> set vs \<Longrightarrow> r1 \<in> set (fst (splitAt r2 vs)) \<Longrightarrow> before vs r1 r2"
proof -
assume r2: "r2 \<in> set vs" and r1: "r1 \<in> set (fst (splitAt r2 vs))"
define ab where "ab = fst (splitAt r2 vs)"
define c where "c = snd (splitAt r2 vs)"
define a where "a = fst (splitAt r1 ab)"
define b where "b = snd (splitAt r1 ab)"
from r2 ab_def c_def have vs: "vs = ab @ [r2] @ c" by (auto dest: splitAt_ram)
from r1 ab_def have r1: "r1 \<in> set ab" by simp
with a_def b_def have "ab = a @ [r1] @ b" by (auto dest: splitAt_ram)
with vs show ?thesis by (simp add: before_def) auto
qed
(* usefule simplifier rules *)
lemma before_dist_eq_fst:
"distinct vs \<Longrightarrow> r2 \<in> set vs \<Longrightarrow> r1 \<in> set (fst (splitAt r2 vs)) = before vs r1 r2"
by (auto intro: before_fst before_dist_r1)
lemma before_dist_eq_snd:
"distinct vs \<Longrightarrow> r2 \<in> set (snd (splitAt r1 vs)) = before vs r1 r2"
by (auto intro: before_snd before_dist_r2)
lemma before_dist_not1:
"distinct vs \<Longrightarrow> before vs ram1 ram2 \<Longrightarrow> \<not> before vs ram2 ram1"
proof
assume d: "distinct vs" and b1: "before vs ram2 ram1" and b2: "before vs ram1 ram2"
from b2 have r1: "ram1 \<in> set vs" by (drule_tac before_r1)
from d b1 have r2: "ram2 \<in> set (fst (splitAt ram1 vs))" by (rule before_dist_r1)
from d b2 have r2':"ram2 \<in> set (snd (splitAt ram1 vs))" by (rule before_dist_r2)
from d r1 r2 r2' show "False" by (drule_tac splitAt_distinct_fst_snd) auto
qed
lemma before_dist_not2:
"distinct vs \<Longrightarrow> ram1 \<in> set vs \<Longrightarrow> ram2 \<in> set vs \<Longrightarrow> ram1 \<noteq> ram2 \<Longrightarrow> \<not> (before vs ram1 ram2) \<Longrightarrow> before vs ram2 ram1"
proof -
assume "distinct vs" "ram1 \<in> set vs " "ram2 \<in> set vs" "ram1 \<noteq> ram2" "\<not> before vs ram1 ram2"
then show "before vs ram2 ram1" apply (frule_tac before_or) by auto
qed
lemma before_dist_eq:
"distinct vs \<Longrightarrow> ram1 \<in> set vs \<Longrightarrow> ram2 \<in> set vs \<Longrightarrow> ram1 \<noteq> ram2 \<Longrightarrow> ( \<not> (before vs ram1 ram2)) = before vs ram2 ram1"
by (auto intro: before_dist_not2 dest: before_dist_not1)
lemma before_vs:
"distinct vs \<Longrightarrow> before vs ram1 ram2 \<Longrightarrow> vs = fst (splitAt ram1 vs) @ ram1 # fst (splitAt ram2 (snd (splitAt ram1 vs))) @ ram2 # snd (splitAt ram2 vs)"
proof -
assume d: "distinct vs" and b: "before vs ram1 ram2"
define s where "s = snd (splitAt ram1 vs)"
from b have "ram1 \<in> set vs" by (auto simp: before_def)
with s_def have vs: "vs = fst (splitAt ram1 vs) @ [ram1] @ s" by (auto dest: splitAt_ram)
from d b s_def have "ram2 \<in> set s" by (auto intro: before_dist_r2)
then have snd: "s = fst (splitAt ram2 s) @ [ram2] @ snd (splitAt ram2 s)"
by (auto dest: splitAt_ram)
with vs have "vs = fst (splitAt ram1 vs) @ [ram1] @ fst (splitAt ram2 s) @ [ram2] @ snd (splitAt ram2 s)" by auto
with d b s_def show ?thesis by auto
qed
(************************ between lemmas *************************************)
subsection \<open>@{const between}\<close>
definition pre_between :: "'a list \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
"pre_between vs ram1 ram2 \<equiv>
distinct vs \<and> ram1 \<in> set vs \<and> ram2 \<in> set vs \<and> ram1 \<noteq> ram2"
declare pre_between_def [simp]
lemma pre_between_dist[intro]:
"pre_between vs ram1 ram2 \<Longrightarrow> distinct vs" by (auto simp: pre_between_def)
lemma pre_between_r1[intro]:
"pre_between vs ram1 ram2 \<Longrightarrow> ram1 \<in> set vs" by auto
lemma pre_between_r2[intro]:
"pre_between vs ram1 ram2 \<Longrightarrow> ram2 \<in> set vs" by auto
lemma pre_between_r12[intro]:
"pre_between vs ram1 ram2 \<Longrightarrow> ram1 \<noteq> ram2" by auto
lemma pre_between_symI:
"pre_between vs ram1 ram2 \<Longrightarrow> pre_between vs ram2 ram1" by auto
lemma pre_between_before[dest]:
"pre_between vs ram1 ram2 \<Longrightarrow> before vs ram1 ram2 \<or> before vs ram2 ram1" by (rule_tac before_or) auto
lemma pre_between_rotate1[intro]:
"pre_between vs ram1 ram2 \<Longrightarrow> pre_between (rotate1 vs) ram1 ram2" by auto
lemma pre_between_rotate[intro]:
"pre_between vs ram1 ram2 \<Longrightarrow> pre_between (rotate n vs) ram1 ram2" by auto
lemma(*<*) before_xor: (*>*)
"pre_between vs ram1 ram2 \<Longrightarrow> (\<not> before vs ram1 ram2) = before vs ram2 ram1"
by (simp add: before_dist_eq)
declare pre_between_def [simp del]
lemma between_simp1[simp]:
"before vs ram1 ram2 \<Longrightarrow> pre_between vs ram1 ram2 \<Longrightarrow>
between vs ram1 ram2 = fst (splitAt ram2 (snd (splitAt ram1 vs)))"
by (simp add: pre_between_def between_def split_def before_dist_eq_snd)
lemma between_simp2[simp]:
"before vs ram1 ram2 \<Longrightarrow> pre_between vs ram1 ram2 \<Longrightarrow>
between vs ram2 ram1 = snd (splitAt ram2 vs) @ fst (splitAt ram1 vs)"
proof -
assume b: "before vs ram1 ram2" and p: "pre_between vs ram1 ram2"
from p b have b2: "\<not> before vs ram2 ram1" apply (simp add: pre_between_def) by (auto dest: before_dist_not1)
with p have "ram2 \<notin> set (fst (splitAt ram1 vs))" by (simp add: pre_between_def before_dist_eq_fst)
then have "fst (splitAt ram1 vs) = fst (splitAt ram2 (fst (splitAt ram1 vs)))" by (auto dest: splitAt_no_ram)
then have "fst (splitAt ram2 (fst (splitAt ram1 vs))) = fst (splitAt ram1 vs)" by auto
with b2 b p show ?thesis apply (simp add: pre_between_def between_def split_def)
by (auto dest: before_dist_not_r1)
qed
lemma between_not_r1[intro]:
"distinct vs \<Longrightarrow> ram1 \<notin> set (between vs ram1 ram2)"
proof (cases "pre_between vs ram1 ram2")
assume d: "distinct vs"
case True then have p: "pre_between vs ram1 ram2" by auto
then show "ram1 \<notin> set (between vs ram1 ram2)"
proof (cases "before vs ram1 ram2")
case True with d p show ?thesis by (auto del: notI)
next
from p have p2: "pre_between vs ram2 ram1" by (auto intro: pre_between_symI)
case False with p have "before vs ram2 ram1" by auto
with d p2 show ?thesis by (auto del: notI)
qed
next
assume d:"distinct vs"
case False then have p: "\<not> pre_between vs ram1 ram2" by auto
then show ?thesis
proof (cases "ram1 = ram2")
case True with d have h1:"ram2 \<notin> set (snd (splitAt ram2 vs))" by (auto del: notI)
from True d have h2: "ram2 \<notin> set (fst (splitAt ram2 (fst (splitAt ram2 vs))))" by (auto del: notI)
with True d h1 show ?thesis by (auto simp: between_def split_def)
next
case False then have neq: "ram1 \<noteq> ram2" by auto
then show ?thesis
proof (cases "ram1 \<notin> set vs")
case True with d show ?thesis by (auto dest: splitAt_no_ram splitAt_in_fst simp: between_def split_def)
next
case False then have r1in: "ram1 \<in> set vs" by auto
then show ?thesis
proof (cases "ram2 \<notin> set vs")
from d have h1: "ram1 \<notin> set (fst (splitAt ram1 vs))" by (auto del: notI)
case True with d h1 show ?thesis
by (auto dest: splitAt_not1 splitAt_in_fst splitAt_ram
splitAt_no_ram simp: between_def split_def del: notI)
next
case False then have r2in: "ram2 \<in> set vs" by auto
with d neq r1in have "pre_between vs ram1 ram2"
by (auto simp: pre_between_def)
with p show ?thesis by auto
qed
qed
qed
qed
lemma between_not_r2[intro]:
"distinct vs \<Longrightarrow> ram2 \<notin> set (between vs ram1 ram2)"
proof (cases "pre_between vs ram1 ram2")
assume d: "distinct vs"
case True then have p: "pre_between vs ram1 ram2" by auto
then show "ram2 \<notin> set (between vs ram1 ram2)"
proof (cases "before vs ram1 ram2")
from d have "ram2 \<notin> set (fst (splitAt ram2 vs))" by (auto del: notI)
then have h1: "ram2 \<notin> set (snd (splitAt ram1 (fst (splitAt ram2 vs))))"
by (auto dest: splitAt_in_fst)
case True with d p h1 show ?thesis by (auto del: notI)
next
from p have p2: "pre_between vs ram2 ram1" by (auto intro: pre_between_symI)
case False with p have "before vs ram2 ram1" by auto
with d p2 show ?thesis by (auto del: notI)
qed
next
assume d:"distinct vs"
case False then have p: "\<not> pre_between vs ram1 ram2" by auto
then show ?thesis
proof (cases "ram1 = ram2")
case True with d have h1:"ram2 \<notin> set (snd (splitAt ram2 vs))" by (auto del: notI)
from True d have h2: "ram2 \<notin> set (fst (splitAt ram2 (fst (splitAt ram2 vs))))" by (auto del: notI)
with True d h1 show ?thesis by (auto simp: between_def split_def)
next
case False then have neq: "ram1 \<noteq> ram2" by auto
then show ?thesis
proof (cases "ram2 \<notin> set vs")
case True with d show ?thesis
by (auto dest: splitAt_no_ram splitAt_in_fst
splitAt_in_fst simp: between_def split_def)
next
case False then have r1in: "ram2 \<in> set vs" by auto
then show ?thesis
proof (cases "ram1 \<notin> set vs")
from d have h1: "ram1 \<notin> set (fst (splitAt ram1 vs))" by (auto del: notI)
case True with d h1 show ?thesis by (auto dest: splitAt_ram splitAt_no_ram simp: between_def split_def del: notI)
next
case False then have r2in: "ram1 \<in> set vs" by auto
with d neq r1in have "pre_between vs ram1 ram2" by (auto simp: pre_between_def)
with p show ?thesis by auto
qed
qed
qed
qed
lemma between_distinct[intro]:
"distinct vs \<Longrightarrow> distinct (between vs ram1 ram2)"
proof -
assume vs: "distinct vs"
define a where "a = fst (splitAt ram1 vs)"
define b where "b = snd (splitAt ram1 vs)"
from a_def b_def have ab: "(a,b) = splitAt ram1 vs" by auto
with vs have ab_disj:"set a \<inter> set b = {}" by (drule_tac splitAt_distinct_ab) auto
define c where "c = fst (splitAt ram2 a)"
define d where "d = snd (splitAt ram2 a)"
from c_def d_def have c_d: "(c,d) = splitAt ram2 a" by auto
with ab_disj have "set c \<inter> set b = {}" by (drule_tac splitAt_subset_ab) auto
with vs a_def b_def c_def show ?thesis
by (auto simp: between_def split_def splitAt_no_ram dest: splitAt_ram intro: splitAt_distinct_fst splitAt_distinct_snd)
qed
lemma between_distinct_r12:
"distinct vs \<Longrightarrow> ram1 \<noteq> ram2 \<Longrightarrow> distinct (ram1 # between vs ram1 ram2 @ [ram2])" by (auto del: notI)
lemma between_vs:
"before vs ram1 ram2 \<Longrightarrow> pre_between vs ram1 ram2 \<Longrightarrow>
vs = fst (splitAt ram1 vs) @ ram1 # (between vs ram1 ram2) @ ram2 # snd (splitAt ram2 vs)"
apply (simp) apply (frule pre_between_dist) apply (drule before_vs) by auto
lemma between_in:
"before vs ram1 ram2 \<Longrightarrow> pre_between vs ram1 ram2 \<Longrightarrow> x \<in> set vs \<Longrightarrow> x = ram1 \<or> x \<in> set (between vs ram1 ram2) \<or> x = ram2 \<or> x \<in> set (between vs ram2 ram1)"
proof -
assume b: "before vs ram1 ram2" and p: "pre_between vs ram1 ram2" and xin: "x \<in> set vs"
define a where "a = fst (splitAt ram1 vs)"
define b where "b = between vs ram1 ram2"
define c where "c = snd (splitAt ram2 vs)"
from p have "distinct vs" by auto
from p b a_def b_def c_def have "vs = a @ ram1 # b @ ram2 # c" apply (drule_tac between_vs) by auto
with xin have "x \<in> set (a @ ram1 # b @ ram2 # c)" by auto
then have "x \<in> set (a) \<or> x \<in> set (ram1 #b) \<or> x \<in> set (ram2 # c)" by auto
then have "x \<in> set (a) \<or> x = ram1 \<or> x \<in> set b \<or> x = ram2 \<or> x \<in> set c" by auto
then have "x \<in> set c \<or> x \<in> set (a) \<or> x = ram1 \<or> x \<in> set b \<or> x = ram2" by auto
then have "x \<in> set (c @ a) \<or> x = ram1 \<or> x \<in> set b \<or> x = ram2" by auto
with b p a_def b_def c_def show ?thesis by auto
qed
lemma
"before vs ram1 ram2 \<Longrightarrow> pre_between vs ram1 ram2 \<Longrightarrow>
hd vs \<noteq> ram1 \<Longrightarrow> (a,b) = splitAt (hd vs) (between vs ram2 ram1) \<Longrightarrow>
vs = [hd vs] @ b @ [ram1] @ (between vs ram1 ram2) @ [ram2] @ a"
proof -
assume b: "before vs ram1 ram2" and p: "pre_between vs ram1 ram2" and vs: "hd vs \<noteq> ram1" and ab: "(a,b) = splitAt (hd vs) (between vs ram2 ram1)"
from p have dist_b: "distinct (between vs ram2 ram1)" by (auto intro: between_distinct simp: pre_between_def)
with ab have "distinct a \<and> distinct b" by (auto intro: splitAt_distinct_a splitAt_distinct_b)
define r where "r = snd (splitAt ram1 vs)"
define btw where "btw = between vs ram2 ram1"
from p r_def have vs2: "vs = fst (splitAt ram1 vs) @ [ram1] @ r" by (auto dest: splitAt_ram simp: pre_between_def)
then have "fst (splitAt ram1 vs) = [] \<Longrightarrow> hd vs = ram1" by auto
with vs have neq: "fst (splitAt ram1 vs) \<noteq> []" by auto
with vs2 have vs_fst: "hd vs = hd (fst (splitAt ram1 vs))" by (induct ("fst (splitAt ram1 vs)")) auto
with neq have "hd vs \<in> set (fst (splitAt ram1 vs))" by auto
with b p have "hd vs \<in> set (between vs ram2 ram1)" by auto
with btw_def have help1: "btw = fst (splitAt (hd vs) btw) @ [hd vs] @ snd (splitAt (hd vs) btw)" by (auto dest: splitAt_ram)
from p b btw_def have "btw = snd (splitAt ram2 vs) @ fst (splitAt ram1 vs)" by auto
with neq have "btw = snd (splitAt ram2 vs) @ hd (fst (splitAt ram1 vs)) # tl (fst (splitAt ram1 vs))" by auto
with vs_fst have "btw = snd (splitAt ram2 vs) @ [hd vs] @ tl (fst (splitAt ram1 vs))" by auto
with help1 have eq: "snd (splitAt ram2 vs) @ [hd vs] @ tl (fst (splitAt ram1 vs)) = fst (splitAt (hd vs) btw) @ [hd vs] @ snd (splitAt (hd vs) btw)" by auto
from dist_b btw_def help1 have "distinct (fst (splitAt (hd vs) btw) @ [hd vs] @ snd (splitAt (hd vs) btw))" by auto
with eq have eq2: "snd (splitAt ram2 vs) = fst (splitAt (hd vs) btw) \<and> tl (fst (splitAt ram1 vs)) = snd (splitAt (hd vs) btw)" apply (rule_tac dist_at) by auto
with btw_def ab have a: "a = snd (splitAt ram2 vs)" by (auto dest: pairD)
from eq2 vs_fst have "hd (fst (splitAt ram1 vs)) # tl (fst (splitAt ram1 vs)) = hd vs # snd (splitAt (hd vs) btw)" by auto
with ab btw_def neq have hdb: "hd vs # b = fst (splitAt ram1 vs)" by (auto dest: pairD)
from b p have "vs = fst (splitAt ram1 vs) @ [ram1] @ fst (splitAt ram2 (snd (splitAt ram1 vs))) @ [ram2] @ snd (splitAt ram2 vs)" apply simp
apply (rule_tac before_vs) by (auto simp: pre_between_def)
with hdb have "vs = (hd vs # b) @ [ram1] @ fst (splitAt ram2 (snd (splitAt ram1 vs))) @ [ram2] @ snd (splitAt ram2 vs)" by auto
with a b p show ?thesis by (simp)
qed
lemma between_congs: "pre_between vs ram1 ram2 \<Longrightarrow> vs \<cong> vs' \<Longrightarrow> between vs ram1 ram2 = between vs' ram1 ram2"
proof -
have "\<And> us. pre_between us ram1 ram2 \<Longrightarrow> before us ram1 ram2 \<Longrightarrow> between us ram1 ram2 = between (rotate1 us) ram1 ram2"
proof -
fix us
assume vors: "pre_between us ram1 ram2" "before us ram1 ram2"
then have pb2: "pre_between (rotate1 us) ram1 ram2" by auto
with vors show "between us ram1 ram2 = between (rotate1 us) ram1 ram2"
proof (cases "us")
case Nil then show ?thesis by auto
next
case (Cons u' us')
with vors pb2 show ?thesis apply (auto simp: before_def)
apply (case_tac "a") apply auto
by (simp_all add: between_def split_def pre_between_def)
qed
qed
moreover have "\<And> us. pre_between us ram1 ram2 \<Longrightarrow> before us ram2 ram1 \<Longrightarrow> between us ram1 ram2 = between (rotate1 us) ram1 ram2"
proof -
fix us
assume vors: " pre_between us ram1 ram2" "before us ram2 ram1"
then have pb2: "pre_between (rotate1 us) ram1 ram2" by auto
with vors show "between us ram1 ram2 = between (rotate1 us) ram1 ram2"
proof (cases "us")
case Nil then show ?thesis by auto
next
case (Cons u' us')
with vors pb2 show ?thesis apply (auto simp: before_def)
apply (case_tac "a") apply auto
by (simp_all add: between_def split_def pre_between_def)
qed
qed
ultimately have "help": "\<And> us. pre_between us ram1 ram2 \<Longrightarrow> between us ram1 ram2 = between (rotate1 us) ram1 ram2"
apply (subgoal_tac "before us ram1 ram2 \<or> before us ram2 ram1") by auto
assume "vs \<cong> vs'" and pre_b: "pre_between vs ram1 ram2"
then obtain n where vs': "vs' = rotate n vs" by (auto simp: congs_def)
have "between vs ram1 ram2 = between (rotate n vs) ram1 ram2"
proof (induct n)
case 0 then show ?case by auto
next
case (Suc m) then show ?case apply simp
apply (subgoal_tac " between (rotate1 (rotate m vs)) ram1 ram2 = between (rotate m vs) ram1 ram2")
by (auto intro: "help" [symmetric] pre_b)
qed
with vs' show ?thesis by auto
qed
lemma between_inter_empty:
"pre_between vs ram1 ram2 \<Longrightarrow>
set (between vs ram1 ram2) \<inter> set (between vs ram2 ram1) = {}"
apply (case_tac "before vs ram1 ram2")
apply (simp add: pre_between_def)
apply (elim conjE)
apply (frule (1) before_vs)
apply (subgoal_tac "distinct (fst (splitAt ram1 vs) @
ram1 # fst (splitAt ram2 (snd (splitAt ram1 vs))) @ ram2 # snd (splitAt ram2 vs))")
apply (thin_tac "vs = fst (splitAt ram1 vs) @
ram1 # fst (splitAt ram2 (snd (splitAt ram1 vs))) @ ram2 # snd (splitAt ram2 vs)")
apply (frule (1) before_dist_fst_snd)
apply(simp)
apply blast
apply (simp only:)
apply (simp add: before_xor)
apply (subgoal_tac "pre_between vs ram2 ram1")
apply (simp add: pre_between_def)
apply (elim conjE)
apply (frule (1) before_vs)
apply (subgoal_tac "distinct (fst (splitAt ram2 vs) @
ram2 # fst (splitAt ram1 (snd (splitAt ram2 vs))) @ ram1 # snd (splitAt ram1 vs))")
apply (thin_tac "vs = fst (splitAt ram2 vs) @
ram2 # fst (splitAt ram1 (snd (splitAt ram2 vs))) @ ram1 # snd (splitAt ram1 vs)")
apply simp
apply blast
apply (simp only:)
by (rule pre_between_symI)
(*********************** between - is_nextElem *************************)
subsubsection \<open>\<open>between is_nextElem\<close>\<close>
lemma is_nextElem_or1: "pre_between vs ram1 ram2 \<Longrightarrow>
is_nextElem vs x y \<Longrightarrow> before vs ram1 ram2 \<Longrightarrow>
is_sublist [x,y] (ram1 # between vs ram1 ram2 @ [ram2])
\<or> is_sublist [x,y] (ram2 # between vs ram2 ram1 @ [ram1])"
proof -
assume p: "pre_between vs ram1 ram2" and is_nextElem: "is_nextElem vs x y" and b: "before vs ram1 ram2"
from p have r1: "ram1 \<in> set vs" by (auto simp: pre_between_def)
define bs where "bs = [ram1] @ (between vs ram1 ram2) @ [ram2]"
have rule1: "x \<in> set (ram1 # (between vs ram1 ram2)) \<Longrightarrow> is_sublist [x,y] bs"
proof -
assume xin:"x \<in> set (ram1 # (between vs ram1 ram2))"
with bs_def have xin2: "x \<in> set bs" by auto
define s where "s = snd (splitAt ram1 vs)"
from r1 s_def have sub1:"is_sublist (ram1 # s) vs" by (auto intro: splitAt_is_sublist2R)
from b p s_def have "ram2 \<in> set s" by (auto intro!: before_dist_r2 simp: pre_between_def)
then have "is_prefix (fst (splitAt ram2 s) @ [ram2]) s" by (auto intro!: splitAt_is_prefix)
then have "is_prefix ([ram1] @ ((fst (splitAt ram2 s)) @ [ram2])) ([ram1] @ s)" by (rule_tac is_prefix_add) auto
with sub1 have "is_sublist (ram1 # (fst (splitAt ram2 s)) @ [ram2]) vs" apply (rule_tac is_sublist_trans) apply (rule is_prefix_sublist)
by simp_all
with p b s_def bs_def have subl: "is_sublist bs vs" by (auto)
with p have db: "distinct bs" by (auto simp: pre_between_def)
with xin bs_def have xnlb:"x \<noteq> last bs" by auto
with p is_nextElem subl xin2 show "is_sublist [x,y] bs" apply (rule_tac is_sublist_is_nextElem) by (auto simp: pre_between_def)
qed
define bs2 where "bs2 = [ram2] @ (between vs ram2 ram1) @ [ram1]"
have rule2: "x \<in> set (ram2 # (between vs ram2 ram1)) \<Longrightarrow> is_sublist [x,y] bs2"
proof -
assume xin:"x \<in> set (ram2 # (between vs ram2 ram1))"
with bs2_def have xin2: "x \<in> set bs2" by auto
define cs1 where "cs1 = ram2 # snd (splitAt ram2 vs)"
then have cs1ne: "cs1 \<noteq> []" by auto
define cs2 where "cs2 = fst (splitAt ram1 vs)"
define cs2ram1 where "cs2ram1 = cs2 @ [ram1]"
from cs1_def cs2_def bs2_def p b have bs2: "bs2 = cs1 @ cs2 @ [ram1]" by (auto simp: pre_between_def)
have "x \<in> set cs1 \<Longrightarrow> x \<noteq> last cs1 \<Longrightarrow> is_sublist [x,y] cs1"
proof-
assume xin2: "x \<in> set cs1" and xnlcs1: "x \<noteq> last cs1"
from cs1_def p have "is_sublist cs1 vs" by (simp add: pre_between_def)
with p is_nextElem xnlcs1 xin2 show ?thesis apply (rule_tac is_sublist_is_nextElem) by (auto simp: pre_between_def)
qed
with bs2 have incs1nl: "x \<in> set cs1 \<Longrightarrow> x \<noteq> last cs1 \<Longrightarrow> is_sublist [x,y] bs2"
apply (auto simp: is_sublist_def) apply (intro exI)
apply (subgoal_tac "as @ x # y # bs @ cs2 @ [ram1] = as @ x # y # (bs @ cs2 @ [ram1])")
apply assumption by auto
have "x = last cs1 \<Longrightarrow> y = hd (cs2 @ [ram1])"
proof -
assume xl: "x = last cs1"
from p have "distinct vs" by auto
with p have "vs = fst (splitAt ram2 vs) @ ram2 # snd (splitAt ram2 vs)" by (auto intro: splitAt_ram)
with cs1_def have "last vs = last (fst (splitAt ram2 vs) @ cs1)" by auto
with cs1ne have "last vs = last cs1" by auto
with xl have "x = last vs" by auto
with p is_nextElem have yhd: "y = hd vs" by auto
from p have "vs = fst (splitAt ram1 vs) @ ram1 # snd (splitAt ram1 vs)" by (auto intro: splitAt_ram)
with yhd cs2ram1_def cs2_def have yhd2: "y = hd (cs2ram1 @ snd (splitAt ram1 vs))" by auto
from cs2ram1_def have "cs2ram1 \<noteq> []" by auto
then have "hd (cs2ram1 @ snd(splitAt ram1 vs)) = hd (cs2ram1)" by auto
with yhd2 cs2ram1_def show ?thesis by auto
qed
with bs2 cs1ne have "x = last cs1 \<Longrightarrow> is_sublist [x,y] bs2"
apply (auto simp: is_sublist_def) apply (intro exI)
apply (subgoal_tac "cs1 @ cs2 @ [ram1] = butlast cs1 @ last cs1 # hd (cs2 @ [ram1]) # tl (cs2 @ [ram1])")
apply assumption by auto
with incs1nl have incs1: "x \<in> set cs1 \<Longrightarrow> is_sublist [x,y] bs2" by auto
have "x \<in> set cs2 \<Longrightarrow> is_sublist [x,y] (cs2 @ [ram1])"
proof-
assume xin2': "x \<in> set cs2"
then have xin2: "x \<in> set (cs2 @ [ram1])" by auto
define fr2 where "fr2 = snd (splitAt ram1 vs)"
from p have "vs = fst (splitAt ram1 vs) @ ram1 # snd (splitAt ram1 vs)" by (auto intro: splitAt_ram)
with fr2_def cs2_def have "vs = cs2 @ [ram1] @ fr2" by simp
with p xin2' have "x \<noteq> ram1" by (auto simp: pre_between_def)
then have xnlcs2: "x \<noteq> last (cs2 @ [ram1])" by auto
from cs2_def p have "is_sublist (cs2 @ [ram1]) vs" by (simp add: pre_between_def)
with p is_nextElem xnlcs2 xin2 show ?thesis apply (rule_tac is_sublist_is_nextElem) by (auto simp: pre_between_def)
qed
with bs2 have "x \<in> set cs2 \<Longrightarrow> is_sublist [x,y] bs2"
apply (auto simp: is_sublist_def) apply (intro exI)
apply (subgoal_tac "cs1 @ as @ x # y # bs = (cs1 @ as) @ x # y # bs")
apply assumption by auto
with p b cs1_def cs2_def incs1 xin show ?thesis by auto
qed
from is_nextElem have "x \<in> set vs" by auto
with b p have "x = ram1 \<or> x \<in> set (between vs ram1 ram2) \<or> x = ram2 \<or> x \<in> set (between vs ram2 ram1)" by (rule_tac between_in) auto
then have "x \<in> set (ram1 # (between vs ram1 ram2)) \<or> x \<in> set (ram2 # (between vs ram2 ram1))" by auto
with rule1 rule2 bs_def bs2_def show ?thesis by auto
qed
lemma is_nextElem_or: "pre_between vs ram1 ram2 \<Longrightarrow> is_nextElem vs x y \<Longrightarrow>
is_sublist [x,y] (ram1 # between vs ram1 ram2 @ [ram2]) \<or> is_sublist [x,y] (ram2 # between vs ram2 ram1 @ [ram1])"
proof (cases "before vs ram1 ram2")
case True
assume "pre_between vs ram1 ram2" "is_nextElem vs x y"
with True show ?thesis by (rule_tac is_nextElem_or1)
next
assume p: "pre_between vs ram1 ram2" and is_nextElem: "is_nextElem vs x y"
from p have p': "pre_between vs ram2 ram1" by (auto intro: pre_between_symI)
case False with p have "before vs ram2 ram1" by auto
with p' is_nextElem show ?thesis apply (drule_tac is_nextElem_or1) apply assumption+ by auto
qed
lemma(*<*) between_eq2: (*>*)
"pre_between vs ram1 ram2 \<Longrightarrow>
before vs ram2 ram1 \<Longrightarrow>
\<exists>as bs cs. between vs ram1 ram2 = cs @ as \<and> vs = as @[ram2] @ bs @ [ram1] @ cs"
apply (subgoal_tac "pre_between vs ram2 ram1")
apply auto apply (intro exI conjI) apply simp apply (simp add: pre_between_def) apply auto
apply (frule_tac before_vs) apply auto by (auto simp: pre_between_def)
lemma is_sublist_same_len[simp]:
"length xs = length ys \<Longrightarrow> is_sublist xs ys = (xs = ys)"
apply(cases xs)
apply simp
apply(rename_tac a as)
apply(cases ys)
apply simp
apply(rename_tac b bs)
apply(case_tac "a = b")
apply(subst is_sublist_rec)
apply simp
apply simp
done
lemma is_nextElem_between_empty[simp]:
"distinct vs \<Longrightarrow> is_nextElem vs a b \<Longrightarrow> between vs a b = []"
apply (simp add: is_nextElem_def between_def split_def)
apply (cases "vs") apply simp+
apply (simp split: if_split_asm)
apply (case_tac "b = aa")
apply (simp add: is_sublist_def)
apply (erule disjE)
apply (elim exE)
apply (case_tac "as")
apply simp
apply simp
apply simp
apply (case_tac "list" rule: rev_exhaust)
apply simp
apply simp
apply simp
apply (subgoal_tac "aa # list = vs")
apply (thin_tac "vs = aa # list")
apply simp
apply (subgoal_tac "distinct vs")
apply (simp add: is_sublist_def)
apply (elim exE)
by auto
lemma is_nextElem_between_empty': "between vs a b = [] \<Longrightarrow> distinct vs \<Longrightarrow> a \<in> set vs \<Longrightarrow> b \<in> set vs \<Longrightarrow>
a \<noteq> b \<Longrightarrow> is_nextElem vs a b"
apply (simp add: is_nextElem_def between_def split_def split: if_split_asm)
apply (case_tac vs) apply simp
apply simp
apply (rule conjI)
apply (rule impI)
apply simp
apply (case_tac "aa = b")
apply simp
apply (rule impI)
apply (case_tac "list" rule: rev_exhaust)
apply simp
apply simp
apply (case_tac "a = y") apply simp
apply simp
apply (elim conjE)
apply (drule split_list_first)
apply (elim exE)
apply simp
apply (rule impI)
apply (subgoal_tac "b \<noteq> aa")
apply simp
apply (case_tac "a = aa")
apply simp
apply (case_tac "list") apply simp
apply simp
apply (case_tac "aaa = b") apply (simp add: is_sublist_def) apply force
apply simp
apply simp
apply (drule split_list_first)
apply (elim exE)
apply simp
apply (case_tac "zs") apply simp
apply simp
apply (case_tac "aaa = b")
apply simp
apply (simp add: is_sublist_def) apply force
apply simp
apply force
apply (case_tac vs) apply simp
apply simp
apply (rule conjI)
apply (rule impI) apply simp
apply (rule impI)
apply (case_tac "b = aa")
apply simp
apply (case_tac "list" rule: rev_exhaust) apply simp
apply simp
apply (case_tac "a = y") apply simp
apply simp
apply (drule split_list_first)
apply (elim exE)
apply simp
apply simp apply (case_tac "a = aa") by auto
lemma between_nextElem: "pre_between vs u v \<Longrightarrow>
between vs u (nextElem vs (hd vs) v) = between vs u v @ [v]"
apply(subgoal_tac "pre_between vs v u")
prefer 2 apply (clarsimp simp add:pre_between_def)
apply(case_tac "before vs u v")
apply(drule (1) between_eq2)
apply(clarsimp simp:pre_between_def hd_append split:list.split)
apply(simp add:between_def split_def)
apply(fastforce simp:neq_Nil_conv)
apply (simp only:before_xor)
apply(drule (1) between_eq2)
apply(clarsimp simp:pre_between_def hd_append split:list.split)
apply (simp add: append_eq_Cons_conv)
apply(fastforce simp:between_def split_def)
done
(******************** section split_face ********************************)
lemma nextVertices_in_face[simp]: "v \<in> \<V> f \<Longrightarrow> f\<^bsup>n\<^esup> \<bullet> v \<in> \<V> f"
proof -
assume v: "v \<in> \<V> f"
then have f: "vertices f \<noteq> []" by auto
show ?thesis apply (simp add: nextVertices_def)
apply (induct n) by (auto simp: f v)
qed
subsubsection \<open>\<open>is_nextElem edges\<close> equivalence\<close>
lemma is_nextElem_edges1: "distinct (vertices f) \<Longrightarrow> (a,b) \<in> edges (f::face) \<Longrightarrow> is_nextElem (vertices f) a b" apply (simp add: edges_face_def nextVertex_def) apply (rule is_nextElem1) by auto
lemma is_nextElem_edges2:
"distinct (vertices f) \<Longrightarrow> is_nextElem (vertices f) a b \<Longrightarrow>
(a,b) \<in> edges (f::face)"
apply (auto simp add: edges_face_def nextVertex_def)
apply (rule sym)
apply (rule is_nextElem2) by (auto intro: is_nextElem_a)
lemma is_nextElem_edges_eq[simp]:
"distinct (vertices (f::face)) \<Longrightarrow>
(a,b) \<in> edges f = is_nextElem (vertices f) a b"
by (auto intro: is_nextElem_edges1 is_nextElem_edges2)
(*********************** nextVertex *****************************)
subsubsection \<open>@{const nextVertex}\<close>
lemma nextElem_suc2: "distinct (vertices f) \<Longrightarrow> last (vertices f) = v \<Longrightarrow> v \<in> set (vertices f) \<Longrightarrow> f \<bullet> v = hd (vertices f)"
proof -
assume dist: "distinct (vertices f)" and last: "last (vertices f) = v" and v: "v \<in> set (vertices f)"
define ls where "ls = vertices f"
have ind: "\<And> c. distinct ls \<Longrightarrow> last ls = v \<Longrightarrow> v \<in> set ls \<Longrightarrow> nextElem ls c v = c"
proof (induct ls)
case Nil then show ?case by auto
next
case (Cons m ms)
then show ?case
proof (cases "m = v")
case True with Cons have "ms = []" apply (cases ms rule: rev_exhaust) by auto
then show ?thesis by auto
next
case False with Cons have a1: "v \<in> set ms" by auto
then have ms: "ms \<noteq> []" by auto
with False Cons ms have "nextElem ms c v = c" apply (rule_tac Cons) by simp_all
with False ms show ?thesis by simp
qed
qed
from dist ls_def last v have "nextElem ls (hd ls) v = hd ls" apply (rule_tac ind) by auto
with ls_def show ?thesis by (simp add: nextVertex_def)
qed
(*********************** split_face ****************************)
subsection \<open>@{const split_face}\<close>
definition pre_split_face :: "face \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool" where
"pre_split_face oldF ram1 ram2 newVertexList \<equiv>
distinct (vertices oldF) \<and> distinct (newVertexList)
\<and> \<V> oldF \<inter> set newVertexList = {}
\<and> ram1 \<in> \<V> oldF \<and> ram2 \<in> \<V> oldF \<and> ram1 \<noteq> ram2"
declare pre_split_face_def [simp]
lemma pre_split_face_p_between[intro]:
"pre_split_face oldF ram1 ram2 newVertexList \<Longrightarrow> pre_between (vertices oldF) ram1 ram2" by (simp add: pre_between_def)
lemma pre_split_face_symI:
"pre_split_face oldF ram1 ram2 newVertexList \<Longrightarrow> pre_split_face oldF ram2 ram1 newVertexList" by auto
lemma pre_split_face_rev[intro]:
"pre_split_face oldF ram1 ram2 newVertexList \<Longrightarrow> pre_split_face oldF ram1 ram2 (rev newVertexList)" by auto
lemma split_face_distinct1:
"(f12, f21) = split_face oldF ram1 ram2 newVertexList \<Longrightarrow> pre_split_face oldF ram1 ram2 newVertexList \<Longrightarrow>
distinct (vertices f12)"
proof -
assume split: "(f12, f21) = split_face oldF ram1 ram2 newVertexList" and p: "pre_split_face oldF ram1 ram2 newVertexList"
define old_vs where "old_vs = vertices oldF"
with p have d_old_vs: "distinct old_vs" by auto
from p have p2: "pre_between (vertices oldF) ram1 ram2" by auto
have rule1: "before old_vs ram1 ram2 \<Longrightarrow> distinct (vertices f12)"
proof -
assume b: "before old_vs ram1 ram2"
with split p have "f12 = Face (rev newVertexList @ ram1 # between (vertices oldF) ram1 ram2 @ [ram2]) Nonfinal" by (simp add: split_face_def)
then have h1:"vertices f12 = rev newVertexList @ ram1 # between (vertices oldF) ram1 ram2 @ [ram2]" by auto
from p have d1: "distinct(ram1 # between (vertices oldF) ram1 ram2 @ [ram2])" by (auto del: notI)
from b p p2 old_vs_def have d2: "set (ram1 # between (vertices oldF) ram1 ram2 @ [ram2]) \<inter> set newVertexList = {}"
by (auto dest!: splitAt_in_fst splitAt_in_snd)
with h1 d1 p show ?thesis by auto
qed
have rule2: "before old_vs ram2 ram1 \<Longrightarrow> distinct (vertices f12)"
proof -
assume b: "before old_vs ram2 ram1"
from p have p3: "pre_split_face oldF ram1 ram2 newVertexList"
by (auto intro: pre_split_face_symI)
then have p4: "pre_between (vertices oldF) ram2 ram1" by auto
with split p have
"f12 = Face (rev newVertexList @ ram1 # between (vertices oldF) ram1 ram2 @ [ram2]) Nonfinal"
by (simp add: split_face_def)
then have h1:"vertices f12 = rev newVertexList @ ram1 # between (vertices oldF) ram1 ram2 @ [ram2]"
by auto
from p3 have d1: "distinct(ram1 # between (vertices oldF) ram1 ram2 @ [ram2])"
by (auto del: notI)
from b p3 p4 old_vs_def
have d2: "set (ram1 # between (vertices oldF) ram1 ram2 @ [ram2]) \<inter> set newVertexList = {}"
by auto
with h1 d1 p show ?thesis by auto
qed
from p2 rule1 rule2 old_vs_def show ?thesis by auto
qed
lemma split_face_distinct1'[intro]:
"pre_split_face oldF ram1 ram2 newVertexList \<Longrightarrow>
distinct (vertices (fst(split_face oldF ram1 ram2 newVertexList)))"
apply (rule_tac split_face_distinct1)
by (auto simp del: pre_split_face_def simp: split_face_def)
lemma split_face_distinct2:
"(f12, f21) = split_face oldF ram1 ram2 newVertexList \<Longrightarrow>
pre_split_face oldF ram1 ram2 newVertexList \<Longrightarrow> distinct (vertices f21)"
proof -
assume split: "(f12, f21) = split_face oldF ram1 ram2 newVertexList"
and p: "pre_split_face oldF ram1 ram2 newVertexList"
define old_vs where "old_vs = vertices oldF"
with p have d_old_vs: "distinct old_vs" by auto
with p have p1: "pre_split_face oldF ram1 ram2 newVertexList" by auto
from p have p2: "pre_between (vertices oldF) ram1 ram2" by auto
have rule1: "before old_vs ram1 ram2 \<Longrightarrow> distinct (vertices f21)"
proof -
assume b: "before old_vs ram1 ram2"
with split p
have "f21 = Face (ram2 # between (vertices oldF) ram2 ram1 @ [ram1] @ newVertexList) Nonfinal"
by (simp add: split_face_def)
then have h1:"vertices f21 = ram2 # between (vertices oldF) ram2 ram1 @ [ram1] @ newVertexList"
by auto
from p have d1: "distinct(ram1 # between (vertices oldF) ram2 ram1 @ [ram2])"
by (auto del: notI)
from b p1 p2 old_vs_def
have d2: "set (ram2 # between (vertices oldF) ram2 ram1 @ [ram1]) \<inter> set newVertexList = {}"
by auto
with h1 d1 p1 show ?thesis by auto
qed
have rule2: "before old_vs ram2 ram1 \<Longrightarrow> distinct (vertices f21)"
proof -
assume b: "before old_vs ram2 ram1"
from p have p3: "pre_split_face oldF ram1 ram2 newVertexList"
by (auto intro: pre_split_face_symI)
then have p4: "pre_between (vertices oldF) ram2 ram1" by auto
with split p
have "f21 = Face (ram2 # between (vertices oldF) ram2 ram1 @ [ram1] @ newVertexList) Nonfinal"
by (simp add: split_face_def)
then have h1:"vertices f21 = ram2 # between (vertices oldF) ram2 ram1 @ [ram1] @ newVertexList"
by auto
from p3 have d1: "distinct(ram2 # between (vertices oldF) ram2 ram1 @ [ram1])"
by (auto del: notI)
from b p3 p4 old_vs_def
have d2: "set (ram2 # between (vertices oldF) ram2 ram1 @ [ram1]) \<inter> set newVertexList = {}"
by auto
with h1 d1 p1 show ?thesis by auto
qed
from p2 rule1 rule2 old_vs_def show ?thesis by auto
qed
lemma split_face_distinct2'[intro]:
"pre_split_face oldF ram1 ram2 newVertexList \<Longrightarrow> distinct (vertices (snd(split_face oldF ram1 ram2 newVertexList)))"
apply (rule_tac split_face_distinct2) by (auto simp del: pre_split_face_def simp: split_face_def)
declare pre_split_face_def [simp del]
lemma split_face_edges_or: "(f12, f21) = split_face oldF ram1 ram2 newVertexList \<Longrightarrow> pre_split_face oldF ram1 ram2 newVertexList \<Longrightarrow> (a, b) \<in> edges oldF \<Longrightarrow> (a,b) \<in> edges f12 \<or> (a,b) \<in> edges f21"
proof -
assume nf: "(f12, f21) = split_face oldF ram1 ram2 newVertexList" and p: "pre_split_face oldF ram1 ram2 newVertexList" and old:"(a, b) \<in> edges oldF"
from p have d1:"distinct (vertices oldF)" by auto
from nf p have d2: "distinct (vertices f12)" by (auto dest: pairD)
from nf p have d3: "distinct (vertices f21)" by (auto dest: pairD)
from p have p': "pre_between (vertices oldF) ram1 ram2" by auto
from old d1 have is_nextElem: "is_nextElem (vertices oldF) a b" by simp
with p have "is_sublist [a,b] (ram1 # (between (vertices oldF) ram1 ram2) @ [ram2]) \<or> is_sublist [a,b] (ram2 # (between (vertices oldF) ram2 ram1) @ [ram1])" apply (rule_tac is_nextElem_or) by auto
then have "is_nextElem (rev newVertexList @ ram1 # between (vertices oldF) ram1 ram2 @ [ram2]) a b
\<or> is_nextElem (ram2 # between (vertices oldF) ram2 ram1 @ ram1 # newVertexList) a b"
proof (cases "is_sublist [a,b] (ram1 # (between (vertices oldF) ram1 ram2) @ [ram2])")
case True then show ?thesis by (auto dest: is_sublist_add intro!: is_nextElem_sublistI)
next
case False
assume "is_sublist [a,b] (ram1 # (between (vertices oldF) ram1 ram2) @ [ram2])
\<or> is_sublist [a,b] (ram2 # (between (vertices oldF) ram2 ram1) @ [ram1])"
with False have "is_sublist [a,b] (ram2 # (between (vertices oldF) ram2 ram1) @ [ram1])" by auto
then show ?thesis apply (rule_tac disjI2) apply (rule_tac is_nextElem_sublistI)
apply (subgoal_tac "is_sublist [a, b] ([] @ (ram2 # between (vertices oldF) ram2 ram1 @ [ram1]) @ newVertexList)") apply force by (frule is_sublist_add)
qed
with nf d1 d2 d3 show ?thesis by (simp add: split_face_def)
qed
subsection \<open>\<open>verticesFrom\<close>\<close>
definition verticesFrom :: "face \<Rightarrow> vertex \<Rightarrow> vertex list" where
"verticesFrom f \<equiv> rotate_to (vertices f)"
lemmas verticesFrom_Def = verticesFrom_def rotate_to_def
lemma len_vFrom[simp]:
"v \<in> \<V> f \<Longrightarrow> |verticesFrom f v| = |vertices f|"
apply(drule split_list_first)
apply(clarsimp simp: verticesFrom_Def)
done
lemma verticesFrom_empty[simp]:
"v \<in> \<V> f \<Longrightarrow> (verticesFrom f v = []) = (vertices f = [])"
by(clarsimp simp: verticesFrom_Def)
lemma verticesFrom_congs:
"v \<in> \<V> f \<Longrightarrow> (vertices f) \<cong> (verticesFrom f v)"
by(simp add:verticesFrom_def cong_rotate_to)
lemma verticesFrom_eq_if_vertices_cong:
"\<lbrakk>distinct(vertices f); distinct(vertices f');
vertices f \<cong> vertices f'; x \<in> \<V> f \<rbrakk> \<Longrightarrow>
verticesFrom f x = verticesFrom f' x"
by(clarsimp simp:congs_def verticesFrom_Def splitAt_rotate_pair_conv)
lemma verticesFrom_in[intro]: "v \<in> \<V> f \<Longrightarrow> a \<in> \<V> f \<Longrightarrow> a \<in> set (verticesFrom f v)"
by (auto dest: verticesFrom_congs congs_pres_nodes)
lemma verticesFrom_in': "a \<in> set (verticesFrom f v) \<Longrightarrow> a \<noteq> v \<Longrightarrow> a \<in> \<V> f"
apply (cases "v \<in> \<V> f") apply (auto dest: verticesFrom_congs congs_pres_nodes)
by (simp add: verticesFrom_Def)
lemma set_verticesFrom:
"v \<in> \<V> f \<Longrightarrow> set (verticesFrom f v) = \<V> f"
apply(auto)
apply (auto simp add: verticesFrom_Def)
done
lemma verticesFrom_hd: "hd (verticesFrom f v) = v" by (simp add: verticesFrom_Def)
lemma verticesFrom_distinct[simp]: "distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow> distinct (verticesFrom f v)" apply (frule_tac verticesFrom_congs) by (auto simp: congs_distinct)
lemma verticesFrom_nextElem_eq:
"distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow> u \<in> \<V> f \<Longrightarrow>
nextElem (verticesFrom f v) (hd (verticesFrom f v)) u
= nextElem (vertices f) (hd (vertices f)) u" apply (subgoal_tac "(verticesFrom f v) \<cong> (vertices f)")
apply (rule nextElem_congs_eq) apply (auto simp: verticesFrom_congs congs_pres_nodes) apply (rule congs_sym)
by (simp add: verticesFrom_congs)
lemma nextElem_vFrom_suc1: "distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow> i < length (vertices f) \<Longrightarrow> last (verticesFrom f v) \<noteq> u \<Longrightarrow> (verticesFrom f v)!i = u \<Longrightarrow> f \<bullet> u = (verticesFrom f v)!(Suc i)"
proof -
assume dist: "distinct (vertices f)" and ith: "(verticesFrom f v)!i = u" and small_i: "i < length (vertices f)" and notlast: "last (verticesFrom f v) \<noteq> u" and v: "v \<in> \<V> f"
hence eq: "(vertices f) \<cong> (verticesFrom f v)" by (auto simp: verticesFrom_congs)
from ith eq small_i have "u \<in> set (verticesFrom f v)" by (auto simp: congs_length)
with eq have u: "u \<in> \<V> f" by (auto simp: congs_pres_nodes)
define ls where "ls = verticesFrom f v"
with dist ith have "ls!i = u" by auto
have ind: "\<And> c i. i < length ls \<Longrightarrow> distinct ls \<Longrightarrow> last ls \<noteq> u \<Longrightarrow> ls!i = u \<Longrightarrow> u \<in> set ls \<Longrightarrow>
nextElem ls c u = ls ! Suc i"
proof (induct ls)
case Nil then show ?case by auto
next
case (Cons m ms)
then show ?case
proof (cases "m = u")
case True with Cons have "u \<notin> set ms" by auto
with Cons True have i: "i = 0" apply (induct i) by auto
with Cons show ?thesis apply (simp split: if_split_asm) apply (cases ms) by simp_all
next
case False with Cons have a1: "u \<in> set ms" by auto
then have ms: "ms \<noteq> []" by auto
from False Cons have i: "0 < i" by auto
define i' where "i' = i - 1"
with i have i': "i = Suc i'" by auto
with False Cons i' ms have "nextElem ms c u = ms ! Suc i'" apply (rule_tac Cons) by simp_all
with False Cons i' ms show ?thesis by simp
qed
qed
from eq dist ith ls_def small_i notlast v
have "nextElem ls (hd ls) u = ls ! Suc i"
apply (rule_tac ind) by (auto simp: congs_length)
with dist u v ls_def show ?thesis by (simp add: nextVertex_def verticesFrom_nextElem_eq)
qed
lemma verticesFrom_nth: "distinct (vertices f) \<Longrightarrow> d < length (vertices f) \<Longrightarrow>
v \<in> \<V> f \<Longrightarrow> (verticesFrom f v)!d = f\<^bsup>d\<^esup> \<bullet> v"
proof (induct d)
case 0 then show ?case by (simp add: verticesFrom_Def nextVertices_def)
next
case (Suc n)
then have dist2: "distinct (verticesFrom f v)"
apply (frule_tac verticesFrom_congs) by (auto simp: congs_distinct)
from Suc have in2: "v \<in> set (verticesFrom f v)"
apply (frule_tac verticesFrom_congs) by (auto dest: congs_pres_nodes)
then have vFrom: "(verticesFrom f v) = butlast (verticesFrom f v) @ [last (verticesFrom f v)]"
apply (cases "(verticesFrom f v)" rule: rev_exhaust) by auto
from Suc show ?case
proof (cases "last (verticesFrom f v) = f\<^bsup>n\<^esup> \<bullet> v")
case True with Suc have "verticesFrom f v ! n = f\<^bsup>n\<^esup> \<bullet> v" by (rule_tac Suc) auto
with True have "last (verticesFrom f v) = verticesFrom f v ! n" by auto
with Suc dist2 in2 have "Suc n = length (verticesFrom f v)"
apply (rule_tac nth_last_Suc_n) by auto
with Suc show ?thesis by auto
next
case False with Suc show ?thesis apply (simp add: nextVertices_def) apply (rule sym)
apply (rule_tac nextElem_vFrom_suc1) by simp_all
qed
qed
lemma verticesFrom_length: "distinct (vertices f) \<Longrightarrow> v \<in> set (vertices f) \<Longrightarrow>
length (verticesFrom f v) = length (vertices f)"
by (auto intro: congs_length verticesFrom_congs)
lemma verticesFrom_between: "v' \<in> \<V> f \<Longrightarrow> pre_between (vertices f) u v \<Longrightarrow>
between (vertices f) u v = between (verticesFrom f v') u v"
by (auto intro!: between_congs verticesFrom_congs)
lemma verticesFrom_is_nextElem: "v \<in> \<V> f \<Longrightarrow>
is_nextElem (vertices f) a b = is_nextElem (verticesFrom f v) a b"
apply (rule is_nextElem_congs_eq) by (rule verticesFrom_congs)
lemma verticesFrom_is_nextElem_last: "v' \<in> \<V> f \<Longrightarrow> distinct (vertices f)
\<Longrightarrow> is_nextElem (verticesFrom f v') (last (verticesFrom f v')) v \<Longrightarrow> v = v'"
apply (subgoal_tac "distinct (verticesFrom f v')")
apply (subgoal_tac "last (verticesFrom f v') \<in> set (verticesFrom f v')")
apply (simp add: nextElem_is_nextElem)
apply (simp add: verticesFrom_hd)
apply (cases "(verticesFrom f v')" rule: rev_exhaust) apply (simp add: verticesFrom_Def)
by auto
lemma verticesFrom_is_nextElem_hd: "v' \<in> \<V> f \<Longrightarrow> distinct (vertices f)
\<Longrightarrow> is_nextElem (verticesFrom f v') u v' \<Longrightarrow> u = last (verticesFrom f v')"
apply (subgoal_tac "distinct (verticesFrom f v')")
apply (thin_tac "distinct (vertices f)") apply (auto simp: is_nextElem_def)
apply (drule is_sublist_y_hd) apply (simp add: verticesFrom_hd)
by auto
lemma verticesFrom_pres_nodes1: "v \<in> \<V> f \<Longrightarrow> \<V> f = set(verticesFrom f v)"
proof -
assume "v \<in> \<V> f"
then have "fst (splitAt v (vertices f)) @ [v] @ snd (splitAt v (vertices f)) = vertices f"
apply (drule_tac splitAt_ram) by simp
moreover have "set (fst (splitAt v (vertices f)) @ [v] @ snd (splitAt v (vertices f))) = set (verticesFrom f v)"
by (auto simp: verticesFrom_Def)
ultimately show ?thesis by simp
qed
lemma verticesFrom_pres_nodes: "v \<in> \<V> f \<Longrightarrow> w \<in> \<V> f \<Longrightarrow> w \<in> set (verticesFrom f v)"
by (auto dest: verticesFrom_pres_nodes1)
lemma before_verticesFrom: "distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow> w \<in> \<V> f \<Longrightarrow>
v \<noteq> w \<Longrightarrow> before (verticesFrom f v) v w"
proof -
assume v: "v \<in> \<V> f" and w: "w \<in> \<V> f" and neq: "v \<noteq> w"
have "hd (verticesFrom f v) = v" by (rule verticesFrom_hd)
with v have vf:"verticesFrom f v = v # tl (verticesFrom f v)"
apply (frule_tac verticesFrom_pres_nodes1)
apply (cases "verticesFrom f v") by auto
moreover with v w have "w \<in> set (verticesFrom f v)" by (auto simp: verticesFrom_pres_nodes)
ultimately have "w \<in> set (v # tl (verticesFrom f v))" by auto
with neq have "w \<in> set (tl (verticesFrom f v))" by auto
with vf have "verticesFrom f v =
[] @ v # fst (splitAt w (tl (verticesFrom f v))) @ w # snd (splitAt w (tl (verticesFrom f v)))"
by (auto dest: splitAt_ram)
then show ?thesis apply (unfold before_def) by (intro exI)
qed
lemma last_vFrom:
"\<lbrakk> distinct(vertices f); x \<in> \<V> f \<rbrakk> \<Longrightarrow> last(verticesFrom f x) = f\<^bsup>-1\<^esup> \<bullet> x"
apply(frule split_list)
apply(clarsimp simp:prevVertex_def verticesFrom_Def split:list.split)
done
lemma rotate_before_vFrom:
"\<lbrakk> distinct(vertices f); r \<in> \<V> f; r\<noteq>u \<rbrakk> \<Longrightarrow>
before (verticesFrom f r) u v \<Longrightarrow> before (verticesFrom f v) r u"
using [[linarith_neq_limit = 1]]
apply(frule split_list)
apply(clarsimp simp:verticesFrom_Def)
apply(rename_tac as bs)
apply(clarsimp simp:before_def)
apply(rename_tac xs ys zs)
apply(subst (asm) Cons_eq_append_conv)
apply clarsimp
apply(rename_tac bs')
apply(subst (asm) append_eq_append_conv2)
apply clarsimp
apply(rename_tac as')
apply(erule disjE)
defer
apply clarsimp
apply(rule_tac x = "v#zs" in exI)
apply(rule_tac x = "bs@as'" in exI)
apply simp
apply clarsimp
apply(subst (asm) append_eq_Cons_conv)
apply(erule disjE)
apply clarsimp
apply(rule_tac x = "v#zs" in exI)
apply simp apply blast
apply clarsimp
apply(rename_tac ys')
apply(subst (asm) append_eq_append_conv2)
apply clarsimp
apply(rename_tac vs')
apply(erule disjE)
apply clarsimp
apply(subst (asm) append_eq_Cons_conv)
apply(erule disjE)
apply clarsimp
apply(rule_tac x = "v#zs" in exI)
apply simp apply blast
apply clarsimp
apply(rule_tac x = "v#ys'@as" in exI)
apply simp apply blast
apply clarsimp
apply(rule_tac x = "v#zs" in exI)
apply simp apply blast
done
lemma before_between:
"\<lbrakk> before(verticesFrom f x) y z; distinct(vertices f); x \<in> \<V> f; x \<noteq> y \<rbrakk> \<Longrightarrow>
y \<in> set(between (vertices f) x z)"
apply(induct f)
apply(clarsimp simp:verticesFrom_Def before_def)
apply(frule split_list)
apply(clarsimp simp:Cons_eq_append_conv)
apply(subst (asm) append_eq_append_conv2)
apply clarsimp
apply(erule disjE)
apply(clarsimp simp:append_eq_Cons_conv)
prefer 2 apply(clarsimp simp add:between_def split_def)
apply(erule disjE)
apply (clarsimp simp:between_def split_def)
apply clarsimp
apply(subst (asm) append_eq_append_conv2)
apply clarsimp
apply(erule disjE)
prefer 2 apply(clarsimp simp add:between_def split_def)
apply(clarsimp simp:append_eq_Cons_conv)
apply(fastforce simp:between_def split_def)
done
lemma before_between2:
"\<lbrakk> before (verticesFrom f u) v w; distinct(vertices f); u \<in> \<V> f \<rbrakk>
\<Longrightarrow> u = v \<or> u \<in> set (between (vertices f) w v)"
apply(subgoal_tac "pre_between (vertices f) v w")
apply(subst verticesFrom_between)
apply assumption
apply(erule pre_between_symI)
apply(frule pre_between_r1)
apply(drule (1) verticesFrom_distinct)
using verticesFrom_hd[of f u]
apply(clarsimp simp add:before_def between_def split_def hd_append
split:if_split_asm)
apply(frule (1) verticesFrom_distinct)
apply(clarsimp simp:pre_between_def before_def simp del:verticesFrom_distinct)
apply(rule conjI)
apply(cases "v = u")
apply simp
apply(rule verticesFrom_in'[of v f u])apply simp apply assumption
apply(cases "w = u")
apply simp
apply(rule verticesFrom_in'[of w f u])apply simp apply assumption
done
(************************** splitFace ********************************)
subsection \<open>@{const splitFace}\<close>
definition pre_splitFace :: "graph \<Rightarrow> vertex \<Rightarrow> vertex \<Rightarrow> face \<Rightarrow> vertex list \<Rightarrow> bool" where
"pre_splitFace g ram1 ram2 oldF nvs \<equiv>
oldF \<in> \<F> g \<and> \<not> final oldF \<and> distinct (vertices oldF) \<and> distinct nvs
\<and> \<V> g \<inter> set nvs = {}
\<and> \<V> oldF \<inter> set nvs = {}
\<and> ram1 \<in> \<V> oldF \<and> ram2 \<in> \<V> oldF
\<and> ram1 \<noteq> ram2
\<and> (((ram1,ram2) \<notin> edges oldF \<and> (ram2,ram1) \<notin> edges oldF
\<and> (ram1, ram2) \<notin> edges g \<and> (ram2, ram1) \<notin> edges g) \<or> nvs \<noteq> [])"
declare pre_splitFace_def [simp]
lemma pre_splitFace_pre_split_face[simp]:
"pre_splitFace g ram1 ram2 oldF nvs \<Longrightarrow> pre_split_face oldF ram1 ram2 nvs"
by (simp add: pre_splitFace_def pre_split_face_def)
lemma pre_splitFace_oldF[simp]:
"pre_splitFace g ram1 ram2 oldF nvs \<Longrightarrow> oldF \<in> \<F> g"
apply (unfold pre_splitFace_def) by force
declare pre_splitFace_def [simp del]
lemma splitFace_split_face:
"oldF \<in> \<F> g \<Longrightarrow>
(f\<^sub>1, f\<^sub>2, newGraph) = splitFace g ram\<^sub>1 ram\<^sub>2 oldF newVs \<Longrightarrow>
(f\<^sub>1, f\<^sub>2) = split_face oldF ram\<^sub>1 ram\<^sub>2 newVs"
by (simp add: splitFace_def split_def)
(* split_face *)
lemma split_face_empty_ram2_ram1_in_f12:
"pre_split_face oldF ram1 ram2 [] \<Longrightarrow>
(f12, f21) = split_face oldF ram1 ram2 [] \<Longrightarrow> (ram2, ram1) \<in> edges f12"
proof -
assume split: "(f12, f21) = split_face oldF ram1 ram2 []"
"pre_split_face oldF ram1 ram2 []"
then have "ram2 \<in> \<V> f12" by (simp add: split_face_def)
moreover with split have "f12 \<bullet> ram2 = hd (vertices f12)"
apply (rule_tac nextElem_suc2)
apply (simp add: pre_split_face_def split_face_distinct1)
by (simp add: split_face_def)
with split have "ram1 = f12 \<bullet> ram2"
by (simp add: split_face_def)
ultimately show ?thesis by (simp add: edges_face_def)
qed
lemma split_face_empty_ram2_ram1_in_f12':
"pre_split_face oldF ram1 ram2 [] \<Longrightarrow>
(ram2, ram1) \<in> edges (fst (split_face oldF ram1 ram2 []))"
proof -
assume split: "pre_split_face oldF ram1 ram2 []"
define f12 where "f12 = fst (split_face oldF ram1 ram2 [])"
define f21 where "f21 = snd (split_face oldF ram1 ram2 [])"
then have "(f12, f21) = split_face oldF ram1 ram2 []" by (simp add: f12_def f21_def)
with split have "(ram2, ram1) \<in> edges f12"
by (rule split_face_empty_ram2_ram1_in_f12)
with f12_def show ?thesis by simp
qed
lemma splitFace_empty_ram2_ram1_in_f12:
"pre_splitFace g ram1 ram2 oldF [] \<Longrightarrow>
(f12, f21, newGraph) = splitFace g ram1 ram2 oldF [] \<Longrightarrow>
(ram2, ram1) \<in> edges f12"
proof -
assume pre: "pre_splitFace g ram1 ram2 oldF []"
then have oldF: "oldF \<in> \<F> g" by (unfold pre_splitFace_def) simp
assume sp: "(f12, f21, newGraph) = splitFace g ram1 ram2 oldF []"
with oldF have "(f12, f21) = split_face oldF ram1 ram2 []"
by (rule splitFace_split_face)
with pre sp show ?thesis
apply (unfold splitFace_def pre_splitFace_def)
apply (simp add: split_def)
apply (rule split_face_empty_ram2_ram1_in_f12')
apply (rule pre_splitFace_pre_split_face)
apply (rule pre)
done
qed
lemma splitFace_f12_new_vertices:
"(f12, f21, newGraph) = splitFace g ram1 ram2 oldF newVs \<Longrightarrow>
v \<in> set newVs \<Longrightarrow> v \<in> \<V> f12"
apply (unfold splitFace_def)
apply (simp add: split_def)
apply (unfold split_face_def Let_def)
by simp
lemma splitFace_add_vertices_direct[simp]:
"vertices (snd (snd (splitFace g ram1 ram2 oldF [countVertices g ..< countVertices g + n])))
= vertices g @ [countVertices g ..< countVertices g + n]"
apply (auto simp: splitFace_def split_def) apply (cases g)
apply auto apply (induct n) by auto
lemma splitFace_delete_oldF:
" (f12, f21, newGraph) = splitFace g ram1 ram2 oldF newVertexList \<Longrightarrow>
oldF \<noteq> f12 \<Longrightarrow> oldF \<noteq> f21 \<Longrightarrow> distinct (faces g) \<Longrightarrow>
oldF \<notin> \<F> newGraph"
by (simp add: splitFace_def split_def distinct_set_replace)
lemma splitFace_faces_1:
"(f12, f21, newGraph) = splitFace g ram1 ram2 oldF newVertexList \<Longrightarrow>
oldF \<in> \<F> g \<Longrightarrow>
set (faces newGraph) \<union> {oldF} = {f12, f21} \<union> set (faces g)"
(is "?oldF \<Longrightarrow> ?C \<Longrightarrow> ?A = ?B")
proof (intro equalityI subsetI)
fix x
assume "x \<in> ?A" and "?C" and "?oldF"
then show "x \<in> ?B" apply (simp add: splitFace_def split_def) by (auto dest: replace1)
next
fix x
assume "x \<in> ?B" and "?C" and "?oldF"
then show "x \<in> ?A" apply (simp add: splitFace_def split_def)
apply (cases "x = oldF") apply simp_all
apply (cases "x = f12") apply simp_all
apply (cases "x = f21") by (auto intro: replace3 replace4)
qed
lemma splitFace_distinct1[intro]:"pre_splitFace g ram1 ram2 oldF newVertexList \<Longrightarrow>
distinct (vertices (fst (snd (splitFace g ram1 ram2 oldF newVertexList))))"
apply (unfold splitFace_def split_def)
by (auto simp add: split_def)
lemma splitFace_distinct2[intro]:"pre_splitFace g ram1 ram2 oldF newVertexList \<Longrightarrow>
distinct (vertices (fst (splitFace g ram1 ram2 oldF newVertexList)))"
apply (unfold splitFace_def split_def)
by (auto simp add: split_def)
lemma splitFace_add_f21': "f' \<in> \<F> g' \<Longrightarrow> fst (snd (splitFace g' v a f' nvl))
\<in> \<F> (snd (snd (splitFace g' v a f' nvl)))"
apply (simp add: splitFace_def split_def) apply (rule disjI2)
apply (rule replace3) by simp_all
lemma split_face_help[simp]: "Suc 0 < |vertices (fst (split_face f' v a nvl))|"
by (simp add: split_face_def)
lemma split_face_help'[simp]: "Suc 0 < |vertices (snd (split_face f' v a nvl))|"
by (simp add: split_face_def)
lemma splitFace_split: "f \<in> \<F> (snd (snd (splitFace g v a f' nvl))) \<Longrightarrow>
f \<in> \<F> g
\<or> f = fst (splitFace g v a f' nvl)
\<or> f = (fst (snd (splitFace g v a f' nvl)))"
apply (simp add: splitFace_def split_def)
apply safe by (frule replace5) simp
lemma pre_FaceDiv_between1: "pre_splitFace g' ram1 ram2 f [] \<Longrightarrow>
\<not> between (vertices f) ram1 ram2 = []"
proof -
assume pre_f: "pre_splitFace g' ram1 ram2 f []"
then have pre_bet: "pre_between (vertices f) ram1 ram2" apply (unfold pre_splitFace_def)
by (simp add: pre_between_def)
then have pre_bet': "pre_between (verticesFrom f ram1) ram1 ram2"
by (simp add: pre_between_def set_verticesFrom)
from pre_f have dist': "distinct (verticesFrom f ram1)" apply (unfold pre_splitFace_def) by simp
from pre_f have ram2: "ram2 \<in> \<V> f" apply (unfold pre_splitFace_def) by simp
from pre_f have "\<not> is_nextElem (vertices f) ram1 ram2" apply (unfold pre_splitFace_def) by auto
with pre_f have "\<not> is_nextElem (verticesFrom f ram1) ram1 ram2" apply (unfold pre_splitFace_def)
by (simp add: verticesFrom_is_nextElem [symmetric])
moreover
from pre_f have "ram2 \<in> set (verticesFrom f ram1)" apply (unfold pre_splitFace_def) by auto
moreover
from pre_f have "ram2 \<noteq> ram1" apply (unfold pre_splitFace_def) by auto
ultimately have ram2_not: "ram2 \<noteq> hd (snd (splitAt ram1 (vertices f)) @ fst (splitAt ram1 (vertices f)))"
apply (simp add: is_nextElem_def verticesFrom_Def)
apply (cases "snd (splitAt ram1 (vertices f)) @ fst (splitAt ram1 (vertices f))")
apply simp apply simp
apply (simp add: is_sublist_def) by auto
from pre_f have before: "before (verticesFrom f ram1) ram1 ram2" apply (unfold pre_splitFace_def)
apply safe apply (rule before_verticesFrom) by auto
have "fst (splitAt ram2 (snd (splitAt ram1 (verticesFrom f ram1)))) = [] \<Longrightarrow> False"
proof -
assume "fst (splitAt ram2 (snd (splitAt ram1 (verticesFrom f ram1)))) = []"
moreover
from ram2 pre_f have "ram2 \<in> set (verticesFrom f ram1)"apply (unfold pre_splitFace_def)
by auto
with pre_f have "ram2 \<in> set (snd (splitAt ram1 (vertices f)) @ fst (splitAt ram1 (vertices f)))"
apply (simp add: verticesFrom_Def)
apply (unfold pre_splitFace_def) by auto
moreover
note dist'
ultimately have "ram2 = hd (snd (splitAt ram1 (vertices f)) @ fst (splitAt ram1 (vertices f)))"
apply (rule_tac ccontr)
apply (cases "(snd (splitAt ram1 (vertices f)) @ fst (splitAt ram1 (vertices f)))")
apply simp
apply simp
by (simp add: verticesFrom_Def del: distinct_append)
with ram2_not show ?thesis by auto
qed
with before pre_bet' have "between (verticesFrom f ram1) ram1 ram2 \<noteq> []" by auto
with pre_f pre_bet show ?thesis apply (unfold pre_splitFace_def) apply safe
by (simp add: verticesFrom_between)
qed
lemma pre_FaceDiv_between2: "pre_splitFace g' ram1 ram2 f [] \<Longrightarrow>
\<not> between (vertices f) ram2 ram1 = []"
proof -
assume pre_f: "pre_splitFace g' ram1 ram2 f []"
then have "pre_splitFace g' ram2 ram1 f []" apply (unfold pre_splitFace_def) by auto
then show ?thesis by (rule pre_FaceDiv_between1)
qed
(********************** Edges *******************)
definition Edges :: "vertex list \<Rightarrow> (vertex \<times> vertex) set" where
"Edges vs \<equiv> {(a,b). is_sublist [a,b] vs}"
lemma Edges_Nil[simp]: "Edges [] = {}"
by(simp add:Edges_def is_sublist_def)
lemma Edges_rev:
"Edges (rev (zs::vertex list)) = {(b,a). (a,b) \<in> Edges zs}"
by (auto simp add: Edges_def is_sublist_rev)
lemma in_Edges_rev[simp]:
"((a,b) : Edges (rev (zs::vertex list))) = ((b,a) \<in> Edges zs)"
by (simp add: Edges_rev)
lemma notinset_notinEdge1: "x \<notin> set xs \<Longrightarrow> (x,y) \<notin> Edges xs"
by(unfold Edges_def)(blast intro:is_sublist_in)
lemma notinset_notinEdge2: "y \<notin> set xs \<Longrightarrow> (x,y) \<notin> Edges xs"
by(unfold Edges_def)(blast intro:is_sublist_in1)
lemma in_Edges_in_set: "(x,y) : Edges vs \<Longrightarrow> x \<in> set vs \<and> y \<in> set vs"
by(unfold Edges_def)(blast intro:is_sublist_in is_sublist_in1)
lemma edges_conv_Edges:
"distinct(vertices(f::face)) \<Longrightarrow> \<E> f =
Edges (vertices f) \<union>
(if vertices f = [] then {} else {(last(vertices f), hd(vertices f))})"
by(induct f) (auto simp: Edges_def is_nextElem_def)
lemma Edges_Cons: "Edges(x#xs) =
(if xs = [] then {} else Edges xs \<union> {(x,hd xs)})"
apply(auto simp:Edges_def)
apply(rule ccontr)
apply(simp)
apply(erule thin_rl)
apply(erule contrapos_np)
apply(clarsimp simp:is_sublist_def Cons_eq_append_conv)
apply(rename_tac as bs)
apply(erule disjE)
apply simp
apply(clarsimp)
apply(rename_tac cs)
apply(rule_tac x = cs in exI)
apply(rule_tac x = bs in exI)
apply(rule HOL.refl)
apply(clarsimp simp:neq_Nil_conv)
apply(subst is_sublist_rec)
apply simp
apply(simp add:is_sublist_def)
apply(erule exE)+
apply(rename_tac as bs)
apply simp
apply(rule_tac x = "x#as" in exI)
apply(rule_tac x = bs in exI)
apply simp
done
lemma Edges_append: "Edges(xs @ ys) =
(if xs = [] then Edges ys else
if ys = [] then Edges xs else
Edges xs \<union> Edges ys \<union> {(last xs, hd ys)})"
apply(induct xs)
apply simp
apply (simp add:Edges_Cons)
apply blast
done
lemma Edges_rev_disj: "distinct xs \<Longrightarrow> Edges(rev xs) \<inter> Edges(xs) = {}"
apply(induct xs)
apply simp
apply(auto simp:Edges_Cons Edges_append last_rev
notinset_notinEdge1 notinset_notinEdge2)
done
lemma disj_sets_disj_Edges:
"set xs \<inter> set ys = {} \<Longrightarrow> Edges xs \<inter> Edges ys = {}"
by(unfold Edges_def)(blast intro:is_sublist_in is_sublist_in1)
lemma disj_sets_disj_Edges2:
"set ys \<inter> set xs = {} \<Longrightarrow> Edges xs \<inter> Edges ys = {}"
by(blast intro!:disj_sets_disj_Edges)
lemma finite_Edges[iff]: "finite(Edges xs)"
by(induct xs)(simp_all add:Edges_Cons)
lemma Edges_compl:
"\<lbrakk> distinct vs; x \<in> set vs; y \<in> set vs; x \<noteq> y \<rbrakk> \<Longrightarrow>
Edges(x # between vs x y @ [y]) \<inter> Edges(y # between vs y x @ [x]) = {}"
using [[linarith_neq_limit = 1]]
apply(subgoal_tac
"\<And>xs (ys::vertex list). xs \<noteq> [] \<Longrightarrow> set xs \<inter> set ys = {} \<Longrightarrow> hd xs \<notin> set ys")
prefer 2 apply(drule hd_in_set) apply(blast)
apply(frule split_list[of x])
apply clarsimp
apply(erule disjE)
apply(frule split_list[of y])
apply(clarsimp simp add:between_def split_def)
apply (clarsimp simp add:Edges_Cons Edges_append
notinset_notinEdge1 notinset_notinEdge2
disj_sets_disj_Edges disj_sets_disj_Edges2
Int_Un_distrib Int_Un_distrib2)
apply(fastforce)
apply(frule split_list[of y])
apply(clarsimp simp add:between_def split_def)
apply (clarsimp simp add:Edges_Cons Edges_append notinset_notinEdge1
notinset_notinEdge2 disj_sets_disj_Edges disj_sets_disj_Edges2
Int_Un_distrib Int_Un_distrib2)
apply fastforce
done
lemma Edges_disj:
"\<lbrakk> distinct vs; x \<in> set vs; z \<in> set vs; x \<noteq> y; y \<noteq> z;
y \<in> set(between vs x z) \<rbrakk> \<Longrightarrow>
Edges(x # between vs x y @ [y]) \<inter> Edges(y # between vs y z @ [z]) = {}"
apply(subgoal_tac
"\<And>xs (ys::vertex list). xs \<noteq> [] \<Longrightarrow> set xs \<inter> set ys = {} \<Longrightarrow> hd xs \<notin> set ys")
prefer 2 apply(drule hd_in_set) apply(blast)
apply(frule split_list[of x])
apply clarsimp
apply(erule disjE)
apply simp
apply(drule inbetween_inset)
apply(rule Edges_compl)
apply simp
apply simp
apply simp
apply simp
apply(erule disjE)
apply(frule split_list[of z])
apply(clarsimp simp add:between_def split_def)
apply(erule disjE)
apply(frule split_list[of y])
apply clarsimp
apply (clarsimp simp add:Edges_Cons Edges_append
notinset_notinEdge1 notinset_notinEdge2
disj_sets_disj_Edges disj_sets_disj_Edges2
Int_Un_distrib Int_Un_distrib2)
apply fastforce
apply(frule split_list[of y])
apply clarsimp
apply (clarsimp simp add:Edges_Cons Edges_append notinset_notinEdge1
notinset_notinEdge2 disj_sets_disj_Edges disj_sets_disj_Edges2
Int_Un_distrib Int_Un_distrib2)
apply fastforce
apply(frule split_list[of z])
apply(clarsimp simp add:between_def split_def)
apply(frule split_list[of y])
apply clarsimp
apply (clarsimp simp add:Edges_Cons Edges_append notinset_notinEdge1
notinset_notinEdge2 disj_sets_disj_Edges disj_sets_disj_Edges2
Int_Un_distrib Int_Un_distrib2)
apply fastforce
done
lemma edges_conv_Un_Edges:
"\<lbrakk> distinct(vertices(f::face)); x \<in> \<V> f; y \<in> \<V> f; x \<noteq> y \<rbrakk> \<Longrightarrow>
\<E> f = Edges(x # between (vertices f) x y @ [y]) \<union>
Edges(y # between (vertices f) y x @ [x])"
apply(simp add:edges_conv_Edges)
apply(rule conjI)
apply clarsimp
apply clarsimp
apply(frule split_list[of x])
apply clarsimp
apply(erule disjE)
apply(frule split_list[of y])
apply(clarsimp simp add:between_def split_def)
apply (clarsimp simp add:Edges_Cons Edges_append
notinset_notinEdge1 notinset_notinEdge2
disj_sets_disj_Edges disj_sets_disj_Edges2
Int_Un_distrib Int_Un_distrib2)
apply(fastforce)
apply(frule split_list[of y])
apply(clarsimp simp add:between_def split_def)
apply (clarsimp simp add:Edges_Cons Edges_append
notinset_notinEdge1 notinset_notinEdge2
disj_sets_disj_Edges disj_sets_disj_Edges2
Int_Un_distrib Int_Un_distrib2)
apply(fastforce)
done
lemma Edges_between_edges:
"\<lbrakk> (a,b) \<in> Edges (u # between (vertices(f::face)) u v @ [v]);
pre_split_face f u v vs \<rbrakk> \<Longrightarrow> (a,b) \<in> \<E> f"
apply(simp add:pre_split_face_def)
apply(induct f)
apply(simp add:edges_conv_Edges Edges_Cons)
apply clarify
apply(rename_tac list)
apply(case_tac "between list u v = []")
apply simp
apply(drule (4) is_nextElem_between_empty')
apply(simp add:Edges_def)
apply(subgoal_tac "pre_between list u v")
prefer 2 apply (simp add:pre_between_def)
apply(subgoal_tac "pre_between list v u")
prefer 2 apply (simp add:pre_between_def)
apply(case_tac "before list u v")
apply(drule (1) between_eq2)
apply clarsimp
apply(erule disjE)
apply (clarsimp simp:neq_Nil_conv)
apply(rule is_nextElem_sublistI)
apply(simp (no_asm) add:is_sublist_def)
apply blast
apply(rule is_nextElem_sublistI)
apply(clarsimp simp add:Edges_def is_sublist_def)
apply(rename_tac as bs cs xs ys)
apply(rule_tac x = "as @ u # xs" in exI)
apply(rule_tac x = "ys @ cs" in exI)
apply simp
apply (simp only:before_xor)
apply(drule (1) between_eq2)
apply clarsimp
apply(rename_tac as bs cs)
apply(erule disjE)
apply (clarsimp simp:neq_Nil_conv)
apply(case_tac cs)
apply clarsimp
apply(simp add:is_nextElem_def)
apply simp
apply(rule is_nextElem_sublistI)
apply(simp (no_asm) add:is_sublist_def)
apply(rule_tac x = "as @ v # bs" in exI)
apply simp
apply(rule_tac m1 = "|as|+1" in is_nextElem_rotate_eq[THEN iffD1])
apply simp
apply(simp add:rotate_drop_take)
apply(rule is_nextElem_sublistI)
apply(clarsimp simp add:Edges_def is_sublist_def)
apply(rename_tac xs ys)
apply(rule_tac x = "bs @ u # xs" in exI)
apply simp
done
(******************** split_face_edges ********************************)
(* split_face *)
lemma edges_split_face1: "pre_split_face f u v vs \<Longrightarrow>
\<E>(fst(split_face f u v vs)) =
Edges(v # rev vs @ [u]) \<union> Edges(u # between (vertices f) u v @ [v])"
apply(simp add: edges_conv_Edges split_face_distinct1')
apply(auto simp add:split_face_def Edges_Cons Edges_append)
done
lemma edges_split_face2: "pre_split_face f u v vs \<Longrightarrow>
\<E>(snd(split_face f u v vs)) =
Edges(u # vs @ [v]) \<union> Edges(v # between (vertices f) v u @ [u])"
apply(simp add: edges_conv_Edges split_face_distinct2')
apply(auto simp add:split_face_def Edges_Cons Edges_append)
done
lemma split_face_empty_ram1_ram2_in_f21:
"pre_split_face oldF ram1 ram2 [] \<Longrightarrow>
(f12, f21) = split_face oldF ram1 ram2 [] \<Longrightarrow> (ram1, ram2) \<in> edges f21"
proof -
assume split: "(f12, f21) = split_face oldF ram1 ram2 []"
"pre_split_face oldF ram1 ram2 []"
then have "ram1 \<in> \<V> f21" by (simp add: split_face_def)
moreover with split have "f21 \<bullet> ram1 = hd (vertices f21)"
apply (rule_tac nextElem_suc2)
apply (simp add: pre_split_face_def split_face_distinct2)
by (simp add: split_face_def)
with split have "ram2 = f21 \<bullet> ram1"
by (simp add: split_face_def)
ultimately show ?thesis by (simp add: edges_face_def)
qed
lemma split_face_empty_ram1_ram2_in_f21':
"pre_split_face oldF ram1 ram2 [] \<Longrightarrow>
(ram1, ram2) \<in> edges (snd (split_face oldF ram1 ram2 []))"
proof -
assume split: "pre_split_face oldF ram1 ram2 []"
define f12 where "f12 = fst (split_face oldF ram1 ram2 [])"
define f21 where "f21 = snd (split_face oldF ram1 ram2 [])"
then have "(f12, f21) = split_face oldF ram1 ram2 []" by (simp add: f12_def f21_def)
with split have "(ram1, ram2) \<in> edges f21"
by (rule split_face_empty_ram1_ram2_in_f21)
with f21_def show ?thesis by simp
qed
lemma splitFace_empty_ram1_ram2_in_f21:
"pre_splitFace g ram1 ram2 oldF [] \<Longrightarrow>
(f12, f21, newGraph) = splitFace g ram1 ram2 oldF [] \<Longrightarrow>
(ram1, ram2) \<in> edges f21"
proof -
assume pre: "pre_splitFace g ram1 ram2 oldF []"
then have oldF: "oldF \<in> \<F> g" by (unfold pre_splitFace_def) simp
assume sp: "(f12, f21, newGraph) = splitFace g ram1 ram2 oldF []"
with oldF have "(f12, f21) = split_face oldF ram1 ram2 []"
by (rule splitFace_split_face)
with pre sp show ?thesis
apply (unfold splitFace_def pre_splitFace_def)
apply (simp add: split_def)
apply (rule split_face_empty_ram1_ram2_in_f21')
apply (rule pre_splitFace_pre_split_face)
apply (rule pre)
done
qed
lemma splitFace_f21_new_vertices:
"(f12, f21, newGraph) = splitFace g ram1 ram2 oldF newVs \<Longrightarrow>
v \<in> set newVs \<Longrightarrow> v \<in> \<V> f21"
apply (unfold splitFace_def)
apply (simp add: split_def)
apply (unfold split_face_def)
by simp
lemma split_face_edges_f12:
assumes vors: "pre_split_face f ram1 ram2 vs"
"(f12, f21) = split_face f ram1 ram2 vs"
"vs \<noteq> []" "vs1 = between (vertices f) ram1 ram2" "vs1 \<noteq> []"
shows "edges f12 =
{(hd vs, ram1) , (ram1, hd vs1), (last vs1, ram2), (ram2, last vs)} \<union>
Edges(rev vs) \<union> Edges vs1" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto
from x vors show "x \<in> ?rhs"
apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f12)
apply (case_tac "c = ram2 \<and> d = last vs") apply simp apply simp apply (elim exE)
apply (case_tac "c = ram1") apply simp
apply (subgoal_tac "between (vertices f) ram1 ram2 @ [ram2] = d # bs")
apply (case_tac "between (vertices f) ram1 ram2") apply simp apply simp
apply (rule dist_at2) apply (rule dist_f12) apply (rule sym) apply simp apply simp
(* c \<noteq> ram1 *)
apply (case_tac "c \<in> set(rev vs)")
apply (subgoal_tac "distinct(rev vs)") apply (simp only: in_set_conv_decomp) apply (elim exE) apply simp
apply (case_tac "zs") apply simp
apply (subgoal_tac "ys = as") apply(drule last_rev) apply (simp)
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp
apply simp
apply (simp add:rev_swap)
apply (subgoal_tac "ys = as")
apply (clarsimp simp add: Edges_def is_sublist_def)
apply (rule conjI)
apply (subgoal_tac "\<exists>as bs. rev list @ [d, c] = as @ d # c # bs") apply simp apply (intro exI) apply simp
apply (subgoal_tac "\<exists>asa bs. rev list @ d # c # rev as = asa @ d # c # bs") apply simp apply (intro exI) apply simp
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp
apply (simp add: pre_split_face_def)
(* c \<notin> set vs *)
apply (case_tac "c \<in> set (between (vertices f) ram1 ram2)")
apply (subgoal_tac "distinct (between (vertices f) ram1 ram2)") apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp
apply (case_tac zs) apply simp apply (subgoal_tac "rev vs @ ram1 # ys = as") apply force
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp
apply simp
apply (subgoal_tac "rev vs @ ram1 # ys = as") apply (simp add: Edges_def is_sublist_def)
apply (subgoal_tac "(rev vs @ ram1 # ys) @ c # a # list @ [ram2] = as @ c # d # bs") apply simp
apply (rule conjI) apply (rule impI) apply (rule disjI2)+ apply (rule exI) apply force
apply (rule impI) apply (rule disjI2)+ apply force
apply force
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp
apply (thin_tac "rev vs @ ram1 # between (vertices f) ram1 ram2 @ [ram2] = as @ c # d # bs")
apply (subgoal_tac "distinct (vertices f12)") apply simp
apply (rule dist_f12)
(* c \<notin> set (between (vertices f) ram1 ram2) *)
apply (subgoal_tac "c = ram2") apply simp
apply (subgoal_tac "rev vs @ ram1 # between (vertices f) ram1 ram2 = as") apply force
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp
(* c \<noteq> ram2 *)
apply (subgoal_tac "c \<in> set (rev vs @ ram1 # between (vertices f) ram1 ram2 @ [ram2])")
apply (thin_tac "rev vs @ ram1 # between (vertices f) ram1 ram2 @ [ram2] = as @ c # d # bs") apply simp
by simp
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto
from x vors show "x \<in> ?lhs"
+ supply [[simproc del: defined_all]]
apply (simp add: dist_f12 is_nextElem_def is_sublist_def) apply (simp add: split_face_def)
apply (case_tac "c = ram2 \<and> d = last vs") apply simp
apply (rule disjCI) apply simp
apply (case_tac "c = hd vs \<and> d = ram1")
apply (case_tac "vs") apply simp
apply force
apply simp
apply (case_tac "c = ram1 \<and> d = hd (between (vertices f) ram1 ram2)")
apply (case_tac "between (vertices f) ram1 ram2") apply simp apply force
apply simp
apply (case_tac "c = last (between (vertices f) ram1 ram2) \<and> d = ram2")
apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp
apply simp
apply (intro exI) apply (subgoal_tac "rev vs @ ram1 # ys @ [y, ram2] = (rev vs @ ram1 # ys) @ y # ram2 # []")
apply assumption
apply simp
apply simp
apply (case_tac "(d,c) \<in> Edges vs") apply (simp add: Edges_def is_sublist_def)
apply (elim exE) apply simp apply (intro exI) apply simp
apply (simp add: Edges_def is_sublist_def)
apply (elim exE) apply simp apply (intro exI)
apply (subgoal_tac "rev vs @ ram1 # as @ c # d # bs @ [ram2] = (rev vs @ ram1 # as) @ c # d # bs @ [ram2]")
apply assumption
by simp
qed
lemma split_face_edges_f12_vs:
assumes vors: "pre_split_face f ram1 ram2 []"
"(f12, f21) = split_face f ram1 ram2 []"
"vs1 = between (vertices f) ram1 ram2" "vs1 \<noteq> []"
shows "edges f12 = {(ram2, ram1) , (ram1, hd vs1), (last vs1, ram2)} \<union>
Edges vs1" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto
from x vors show "x \<in> ?rhs"
apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f12)
apply (case_tac " c = ram2 \<and> d = ram1") apply simp
apply simp apply (elim exE)
apply (case_tac "c = ram1") apply simp
apply (subgoal_tac "as = []") apply simp
apply (case_tac "between (vertices f) ram1 ram2") apply simp
apply simp
apply (rule dist_at1) apply (rule dist_f12) apply force apply (rule sym) apply simp
(* a \<noteq> ram1 *)
apply (case_tac "c \<in> set (between (vertices f) ram1 ram2)")
apply (subgoal_tac "distinct (between (vertices f) ram1 ram2)") apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp
apply (case_tac zs) apply simp apply (subgoal_tac "ram1 # ys = as") apply force
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp
apply simp
apply (subgoal_tac "ram1 # ys = as") apply (simp add: Edges_def is_sublist_def)
apply (subgoal_tac "(ram1 # ys) @ c # a # list @ [ram2] = as @ c # d # bs") apply simp
apply (rule conjI) apply (rule impI) apply (rule disjI2)+ apply (rule exI) apply force
apply (rule impI) apply (rule disjI2)+ apply force
apply force
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp
apply (thin_tac "ram1 # between (vertices f) ram1 ram2 @ [ram2] = as @ c # d # bs")
apply (subgoal_tac "distinct (vertices f12)") apply simp
apply (rule dist_f12)
(* c \<notin> set (between (vertices f) ram1 ram2) *)
apply (subgoal_tac "c = ram2") apply simp
apply (subgoal_tac "ram1 # between (vertices f) ram1 ram2 = as") apply force
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp
(* c \<noteq> ram2 *)
apply (subgoal_tac "c \<in> set (ram1 # between (vertices f) ram1 ram2 @ [ram2])")
apply (thin_tac "ram1 # between (vertices f) ram1 ram2 @ [ram2] = as @ c # d # bs") apply simp
by simp
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto
from x vors show "x \<in> ?lhs"
+ supply [[simproc del: defined_all]]
apply (simp add: dist_f12 is_nextElem_def is_sublist_def) apply (simp add: split_face_def)
apply (case_tac "c = ram2 \<and> d = ram1") apply simp
apply (rule disjCI) apply simp
apply (case_tac "c = ram1 \<and> d = hd (between (vertices f) ram1 ram2)")
apply (case_tac "between (vertices f) ram1 ram2") apply simp
apply force
apply simp
apply (case_tac "c = last (between (vertices f) ram1 ram2) \<and> d = ram2")
apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp
apply simp
apply (intro exI) apply (subgoal_tac "ram1 # ys @ [y, ram2] = (ram1 # ys) @ y # ram2 # []")
apply assumption
apply simp
apply simp
apply (case_tac "(c, d) \<in> Edges vs") apply (simp add: Edges_def is_sublist_def)
apply (elim exE) apply simp apply (intro exI)
apply (subgoal_tac "ram1 # as @ c # d # bs @ [ram2] = (ram1 # as) @ c # d # (bs @ [ram2])") apply assumption
apply simp
apply (simp add: Edges_def is_sublist_def)
apply (elim exE) apply simp apply (intro exI)
apply (subgoal_tac "ram1 # as @ c # d # bs @ [ram2] = (ram1 # as) @ c # d # bs @ [ram2]")
apply assumption
by simp
qed
lemma split_face_edges_f12_bet:
assumes vors: "pre_split_face f ram1 ram2 vs"
"(f12, f21) = split_face f ram1 ram2 vs"
"vs \<noteq> []" "between (vertices f) ram1 ram2 = []"
shows "edges f12 = {(hd vs, ram1) , (ram1, ram2), (ram2, last vs)} \<union>
Edges(rev vs)" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto
from x vors show "x \<in> ?rhs"
apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f12)
apply (case_tac " c = ram2 \<and> d = last vs") apply simp
apply simp apply (elim exE)
apply (case_tac "c = ram1") apply simp
apply (subgoal_tac "rev vs = as") apply simp
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp
(* a \<noteq> ram1 *)
apply (case_tac "c \<in> set(rev vs)")
apply (subgoal_tac "distinct(rev vs)") apply (simp only: in_set_conv_decomp) apply (elim exE) apply simp
apply (case_tac "zs") apply simp apply (subgoal_tac "ys = as") apply (simp add:rev_swap)
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp apply simp
apply (subgoal_tac "ys = as") apply (simp add: Edges_def is_sublist_def rev_swap)
apply (rule conjI)
apply (subgoal_tac "\<exists>as bs. rev list @ [d, c] = as @ d # c # bs") apply simp apply (intro exI) apply simp
apply (subgoal_tac "\<exists>asa bs. rev list @ d # c # rev as = asa @ d # c # bs") apply simp apply (intro exI) apply simp
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp
apply (simp add: pre_split_face_def)
(* c \<notin> set vs *)
apply (subgoal_tac "c = ram2") apply simp
apply (subgoal_tac "rev vs @ [ram1] = as") apply force
apply (rule dist_at1) apply (rule dist_f12) apply (rule sym) apply simp apply simp
(* c \<noteq> ram2 *)
apply (subgoal_tac "c \<in> set (rev vs @ ram1 # [ram2])")
apply (thin_tac "rev vs @ ram1 # [ram2] = as @ c # d # bs") apply simp
by simp
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto
from x vors show "x \<in> ?lhs"
apply (simp add: dist_f12 is_nextElem_def is_sublist_def) apply (simp add: split_face_def)
apply (case_tac "c = hd vs \<and> d = ram1")
apply (case_tac "vs") apply simp
apply force
apply simp
apply (case_tac "c = ram1 \<and> d = ram2") apply force
apply simp
apply (case_tac "c = ram2 \<and> d = last vs")
apply (case_tac "vs" rule:rev_exhaust) apply simp
apply simp
apply (simp add: Edges_def is_sublist_def)
apply (elim exE) apply simp apply (rule conjI)
apply (rule impI) apply (rule disjI1) apply (intro exI)
apply (subgoal_tac "c # d # rev as @ [ram1, ram2] = [] @ c # d # rev as @ [ram1,ram2]") apply assumption apply simp
apply (rule impI) apply (rule disjI1) apply (intro exI) by simp
qed
lemma split_face_edges_f12_bet_vs:
assumes vors: "pre_split_face f ram1 ram2 []"
"(f12, f21) = split_face f ram1 ram2 []"
"between (vertices f) ram1 ram2 = []"
shows "edges f12 = {(ram2, ram1) , (ram1, ram2)}" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto
from x vors show "x \<in> ?rhs"
apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f12)
apply (case_tac " c = ram2 \<and> d = ram1") apply force
apply simp apply (elim exE)
apply (case_tac "c = ram1") apply simp
apply (case_tac "as") apply simp
apply (case_tac "list") apply simp apply simp
apply (case_tac "as") apply simp apply (case_tac "list") apply simp by simp
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f12: "distinct (vertices f12)" apply (rule_tac split_face_distinct1) by auto
from x vors show "x \<in> ?lhs"
apply (simp add: dist_f12 is_nextElem_def is_sublist_def) apply (simp add: split_face_def)
by auto
qed
lemma split_face_edges_f12_subset: "pre_split_face f ram1 ram2 vs \<Longrightarrow>
(f12, f21) = split_face f ram1 ram2 vs \<Longrightarrow> vs \<noteq> [] \<Longrightarrow>
{(hd vs, ram1), (ram2, last vs)} \<union> Edges(rev vs) \<subseteq> edges f12"
apply (case_tac "between (vertices f) ram1 ram2")
apply (frule split_face_edges_f12_bet) apply simp apply simp apply simp apply force
apply (frule split_face_edges_f12) apply simp+ by force
(*declare in_Edges_rev [simp del] rev_eq_Cons_iff [simp del]*)
lemma split_face_edges_f21:
assumes vors: "pre_split_face f ram1 ram2 vs"
"(f12, f21) = split_face f ram1 ram2 vs"
"vs \<noteq> []" "vs2 = between (vertices f) ram2 ram1" "vs2 \<noteq> []"
shows "edges f21 = {(last vs2, ram1) , (ram1, hd vs), (last vs, ram2), (ram2, hd vs2)} \<union>
Edges vs \<union> Edges vs2" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto
from x vors show "x \<in> ?rhs"
apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f21)
apply (case_tac " c = last vs \<and> d = ram2") apply simp
apply simp apply (elim exE)
apply (case_tac "c = ram1") apply simp
apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 = as")
apply clarsimp
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
(* a \<noteq> ram1 *)
apply (case_tac "c \<in> set vs")
apply (subgoal_tac "distinct vs")
apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp
apply (case_tac "zs") apply simp
apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # ys = as") apply force
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
apply simp
apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # ys = as")
apply (subgoal_tac "(ram2 # between (vertices f) ram2 ram1 @ ram1 # ys) @ c # a # list = as @ c # d # bs")
apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # ys @ c # a # list = as @ c # d # bs")
apply (simp add: Edges_def is_sublist_def)
apply(rule conjI)
apply (subgoal_tac "\<exists>as bs. ys @ [c, d] = as @ c # d # bs") apply simp apply force
apply force
apply force
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
apply (simp add: pre_split_face_def)
(* c \<notin> set (rev vs) *)
apply (case_tac "c \<in> set (between (vertices f) ram2 ram1)")
apply (subgoal_tac "distinct (between (vertices f) ram2 ram1)") apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp
apply (case_tac zs) apply simp apply (subgoal_tac "ram2 # ys = as") apply force
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
apply simp
apply (subgoal_tac "ram2 # ys = as") apply (simp add: Edges_def is_sublist_def)
apply (subgoal_tac "(ram2 # ys) @ c # a # list @ ram1 # vs = as @ c # d # bs")
apply (thin_tac "ram2 # ys @ c # a # list @ ram1 # vs = as @ c # d # bs")
apply (rule conjI) apply (rule impI) apply (rule disjI2)+ apply force
apply (rule impI) apply (rule disjI2)+ apply force
apply force
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
apply (subgoal_tac "distinct (vertices f21)")
apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # vs = as @ c # d # bs") apply simp
apply (rule dist_f21)
(* c \<notin> set (between (vertices f) ram2 ram1) *)
apply (subgoal_tac "c = ram2") apply simp
apply (subgoal_tac "[] = as") apply simp apply (case_tac "between (vertices f) ram2 ram1") apply simp apply simp
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
(* c \<noteq> ram2 *)
apply (subgoal_tac "c \<in> set (ram2 # between (vertices f) ram2 ram1 @ ram1 # vs)")
apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # vs = as @ c # d # bs") apply simp
by simp
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto
from x vors show "x \<in> ?lhs"
+ supply [[simproc del: defined_all]]
apply (simp add: dist_f21 is_nextElem_def is_sublist_def) apply (simp add: split_face_def)
apply (case_tac "c = ram2 \<and> d = hd (between (vertices f) ram2 ram1)") apply simp apply (rule disjI1)
apply (intro exI) apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # vs =
[] @ ram2 # hd (between (vertices f) ram2 ram1) # tl (between (vertices f) ram2 ram1) @ ram1 # vs") apply assumption apply simp
apply (case_tac "c = ram1 \<and> d = hd vs") apply (rule disjI1)
apply (case_tac "vs") apply simp
apply simp apply (intro exI)
apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # a # list =
(ram2 # between (vertices f) ram2 ram1) @ ram1 # a # list") apply assumption apply simp
apply (case_tac "c = last vs \<and> d = ram2")
apply (case_tac vs rule:rev_exhaust) apply simp
apply simp
apply simp
apply (case_tac "c = last (between (vertices f) ram2 ram1) \<and> d = ram1") apply (rule disjI1)
apply (case_tac "between (vertices f) ram2 ram1" rule: rev_exhaust) apply simp
apply (intro exI) apply simp
apply (subgoal_tac "ram2 # ys @ y # ram1 # vs = (ram2 # ys) @ y # ram1 # vs")
apply assumption
apply simp
apply simp
apply (case_tac "(c, d) \<in> Edges vs") apply (simp add: Edges_def is_sublist_def)
apply (elim exE) apply simp apply (rule conjI) apply (rule impI) apply (rule disjI1) apply (intro exI)
apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # as @ [c, d]
= (ram2 # between (vertices f) ram2 ram1 @ ram1 # as) @ c # d # []") apply assumption apply simp
apply (rule impI) apply (rule disjI1) apply (intro exI)
apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ ram1 # as @ c # d # bs
= (ram2 # between (vertices f) ram2 ram1 @ ram1 # as) @ c # d # bs") apply assumption apply simp
apply (simp add: Edges_def is_sublist_def)
apply (elim exE) apply simp apply (rule disjI1) apply (intro exI)
apply (subgoal_tac "ram2 # as @ c # d # bs @ ram1 # vs = (ram2 # as) @ c # d # (bs @ ram1 # vs)")
apply assumption by simp
qed
lemma split_face_edges_f21_vs:
assumes vors: "pre_split_face f ram1 ram2 []"
"(f12, f21) = split_face f ram1 ram2 []"
"vs2 = between (vertices f) ram2 ram1" "vs2 \<noteq> []"
shows "edges f21 = {(last vs2, ram1) , (ram1, ram2), (ram2, hd vs2)} \<union>
Edges vs2" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto
from x vors show "x \<in> ?rhs"
apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f21)
apply (case_tac " c = ram1 \<and> d = ram2") apply simp apply simp apply (elim exE)
apply (case_tac "c = ram1") apply simp
apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 = as")
apply (subgoal_tac "(ram2 # between (vertices f) ram2 ram1) @ [ram1] = as @ ram1 # d # bs")
apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ [ram1] = as @ ram1 # d # bs")
apply simp apply force
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
(* a \<noteq> ram1 *)
apply (case_tac "c \<in> set (between (vertices f) ram2 ram1)")
apply (subgoal_tac "distinct (between (vertices f) ram2 ram1)") apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp
apply (case_tac zs) apply simp apply (subgoal_tac "ram2 # ys = as") apply force
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp apply simp
apply (subgoal_tac "ram2 # ys = as") apply (simp add: Edges_def is_sublist_def)
apply (subgoal_tac "(ram2 # ys) @ c # a # list @ [ram1] = as @ c # d # bs")
apply (thin_tac "ram2 # ys @ c # a # list @ [ram1] = as @ c # d # bs")
apply (rule conjI) apply (rule impI) apply (rule disjI2)+ apply force
apply (rule impI) apply (rule disjI2)+ apply force apply force
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
apply (subgoal_tac "distinct (vertices f21)")
apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ [ram1] = as @ c # d # bs") apply simp
apply (rule dist_f21)
(* c \<notin> set (between (vertices f) ram2 ram1) *)
apply (subgoal_tac "c = ram2") apply simp
apply (subgoal_tac "[] = as") apply simp apply (case_tac "between (vertices f) ram2 ram1") apply simp apply simp
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
(* c \<noteq> ram2 *)
apply (subgoal_tac "c \<in> set (ram2 # between (vertices f) ram2 ram1 @ [ram1])")
apply (thin_tac "ram2 # between (vertices f) ram2 ram1 @ [ram1] = as @ c # d # bs") apply simp
by simp
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto
from x vors show "x \<in> ?lhs"
+ supply [[simproc del: defined_all]]
apply (simp add: dist_f21 is_nextElem_def is_sublist_def) apply (simp add: split_face_def)
apply (case_tac "c = ram2 \<and> d = hd (between (vertices f) ram2 ram1)") apply simp apply (rule disjI1)
apply (intro exI) apply (subgoal_tac "ram2 # between (vertices f) ram2 ram1 @ [ram1] =
[] @ ram2 # hd (between (vertices f) ram2 ram1) # tl (between (vertices f) ram2 ram1) @ [ram1]") apply assumption apply simp
apply (case_tac "c = ram1 \<and> d = ram2") apply (rule disjI2) apply simp apply simp
apply (case_tac "c = last (between (vertices f) ram2 ram1) \<and> d = ram1") apply (rule disjI1)
apply (case_tac "between (vertices f) ram2 ram1" rule: rev_exhaust) apply simp
apply (intro exI) apply simp
apply (subgoal_tac "ram2 # ys @ y # [ram1] = (ram2 # ys) @ y # [ram1]")
apply assumption apply simp apply simp
apply (simp add: Edges_def is_sublist_def)
apply (elim exE) apply simp apply (rule disjI1) apply (intro exI)
apply (subgoal_tac "ram2 # as @ c # d # bs @ [ram1] = (ram2 # as) @ c # d # (bs @ [ram1])")
apply assumption by simp
qed
lemma split_face_edges_f21_bet:
assumes vors: "pre_split_face f ram1 ram2 vs"
"(f12, f21) = split_face f ram1 ram2 vs"
"vs \<noteq> []" "between (vertices f) ram2 ram1 = []"
shows "edges f21 = {(ram1, hd vs), (last vs, ram2), (ram2, ram1)} \<union>
Edges vs" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto
from x vors show "x \<in> ?rhs"
apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f21)
apply (case_tac " c = last vs \<and> d = ram2") apply simp
apply simp apply (elim exE)
apply (case_tac "c = ram1") apply simp
apply (subgoal_tac "[ram2] = as") apply clarsimp
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
(* a \<noteq> ram1 *)
apply (case_tac "c \<in> set vs")
apply (subgoal_tac "distinct vs") apply (simp add: in_set_conv_decomp) apply (elim exE) apply simp
apply (case_tac "zs") apply simp
apply (subgoal_tac "ram2 # ram1 # ys = as") apply force
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
apply simp
apply (subgoal_tac "ram2 # ram1 # ys = as")
apply (subgoal_tac "(ram2 # ram1 # ys) @ c # a # list = as @ c # d # bs")
apply (thin_tac "ram2 # ram1 # ys @ c # a # list = as @ c # d # bs")
apply (simp add: Edges_def is_sublist_def) apply force
apply force
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
apply (simp add: pre_split_face_def)
(* c \<notin> set (rev vs) *)
apply (subgoal_tac "c = ram2") apply simp
apply (subgoal_tac "[] = as") apply simp
apply (rule dist_at1) apply (rule dist_f21) apply (rule sym) apply simp apply simp
(* c \<noteq> ram2 *)
apply (subgoal_tac "c \<in> set (ram2 # ram1 # vs)")
apply (thin_tac "ram2 # ram1 # vs = as @ c # d # bs") apply simp
by simp
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto
from x vors show "x \<in> ?lhs"
+ supply [[simproc del: defined_all]]
apply (simp add: dist_f21 is_nextElem_def is_sublist_def) apply (simp add: split_face_def)
apply (case_tac "c = ram2 \<and> d = ram1") apply simp apply (rule disjI1) apply force
apply (case_tac "c = ram1 \<and> d = hd vs") apply (rule disjI1)
apply (case_tac "vs") apply simp
apply simp apply (intro exI)
apply (subgoal_tac "ram2 # ram1 # a # list =
[ram2] @ ram1 # a # list") apply assumption apply simp
apply (case_tac "c = last vs \<and> d = ram2")
apply (case_tac vs rule: rev_exhaust) apply simp
apply simp
apply (simp add: Edges_def is_sublist_def)
apply (elim exE) apply simp apply (rule conjI) apply (rule impI) apply (rule disjI1) apply (intro exI)
apply (subgoal_tac "ram2 # ram1 # as @ [c, d]
= (ram2 # ram1 # as) @ c # d # []") apply assumption apply simp
apply (rule impI) apply (rule disjI1) apply (intro exI)
apply (subgoal_tac "ram2 # ram1 # as @ c # d # bs
= (ram2 # ram1 # as) @ c # d # bs") apply assumption by simp
qed
lemma split_face_edges_f21_bet_vs:
assumes vors: "pre_split_face f ram1 ram2 []"
"(f12, f21) = split_face f ram1 ram2 []"
"between (vertices f) ram2 ram1 = []"
shows "edges f21 = {(ram1, ram2), (ram2, ram1)}" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto
from x vors show "x \<in> ?rhs"
apply (simp add: split_face_def is_nextElem_def is_sublist_def dist_f21)
apply (case_tac " c = ram1 \<and> d = ram2") apply simp apply simp apply (elim exE)
apply (case_tac "as") apply simp apply (case_tac "list") by auto
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f21: "distinct (vertices f21)" apply (rule_tac split_face_distinct2) by auto
from x vors show "x \<in> ?lhs"
apply (simp add: dist_f21 is_nextElem_def is_sublist_def) apply (simp add: split_face_def)
by auto
qed
lemma split_face_edges_f21_subset: "pre_split_face f ram1 ram2 vs \<Longrightarrow>
(f12, f21) = split_face f ram1 ram2 vs \<Longrightarrow> vs \<noteq> [] \<Longrightarrow>
{(last vs, ram2), (ram1, hd vs)} \<union> Edges vs \<subseteq> edges f21"
apply (case_tac "between (vertices f) ram2 ram1")
apply (frule split_face_edges_f21_bet) apply simp apply simp apply simp apply force
apply (frule split_face_edges_f21) apply simp+ by force
lemma verticesFrom_ram1: "pre_split_face f ram1 ram2 vs \<Longrightarrow>
verticesFrom f ram1 = ram1 # between (vertices f) ram1 ram2 @ ram2 # between (vertices f) ram2 ram1"
apply (subgoal_tac "pre_between (vertices f) ram1 ram2")
apply (subgoal_tac "distinct (vertices f)")
apply (case_tac "before (vertices f) ram1 ram2")
apply (simp add: verticesFrom_Def)
apply (subgoal_tac "ram2 \<in> set (snd (splitAt ram1 (vertices f)))") apply (drule splitAt_ram)
apply (subgoal_tac "snd (splitAt ram2 (snd (splitAt ram1 (vertices f)))) = snd (splitAt ram2 (vertices f))")
apply simp apply (thin_tac "snd (splitAt ram1 (vertices f)) =
fst (splitAt ram2 (snd (splitAt ram1 (vertices f)))) @
ram2 # snd (splitAt ram2 (snd (splitAt ram1 (vertices f))))") apply simp
apply (rule before_dist_r2) apply simp apply simp
apply (subgoal_tac "before (vertices f) ram2 ram1")
apply (subgoal_tac "pre_between (vertices f) ram2 ram1")
apply (simp add: verticesFrom_Def)
apply (subgoal_tac "ram2 \<in> set (fst (splitAt ram1 (vertices f)))") apply (drule splitAt_ram)
apply (subgoal_tac "fst (splitAt ram2 (fst (splitAt ram1 (vertices f)))) = fst (splitAt ram2 (vertices f))")
apply simp apply (thin_tac "fst (splitAt ram1 (vertices f)) =
fst (splitAt ram2 (fst (splitAt ram1 (vertices f)))) @
ram2 # snd (splitAt ram2 (fst (splitAt ram1 (vertices f))))") apply simp
apply (rule before_dist_r1) apply simp apply simp apply (simp add: pre_between_def) apply force
apply (simp add: pre_split_face_def) by (rule pre_split_face_p_between)
lemma split_face_edges_f_vs1_vs2:
assumes vors: "pre_split_face f ram1 ram2 vs"
"between (vertices f) ram1 ram2 = []"
"between (vertices f) ram2 ram1 = []"
shows "edges f = {(ram2, ram1) , (ram1, ram2)}" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def)
from x vors show "x \<in> ?rhs" apply (simp add: dist_f)
apply (subgoal_tac "pre_between (vertices f) ram1 ram2")
apply (drule is_nextElem_or) apply assumption
apply (simp add: Edges_def)
apply (case_tac "is_sublist [c, d] [ram1, ram2]") apply (simp)
apply (simp) apply blast
by (rule pre_split_face_p_between)
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def)
from x vors show "x \<in> ?lhs" apply (simp add: dist_f)
apply (subgoal_tac "ram1 \<in> \<V> f") apply (simp add: verticesFrom_is_nextElem verticesFrom_ram1)
apply (simp add: is_nextElem_def) apply blast
by (simp add: pre_split_face_def)
qed
lemma split_face_edges_f_vs1:
assumes vors: "pre_split_face f ram1 ram2 vs"
"between (vertices f) ram1 ram2 = []"
"vs2 = between (vertices f) ram2 ram1" "vs2 \<noteq> []"
shows "edges f = {(last vs2, ram1) , (ram1, ram2), (ram2, hd vs2)} \<union>
Edges vs2" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def)
from vors have dist_vs2: "distinct (ram2 # vs2 @ [ram1])" apply (simp only:)
apply (rule between_distinct_r12) apply (rule dist_f) apply (rule not_sym) by (simp add: pre_split_face_def)
from x vors show "x \<in> ?rhs" apply (simp add: dist_f)
apply (subgoal_tac "pre_between (vertices f) ram1 ram2")
apply (drule is_nextElem_or) apply assumption
apply (simp add: Edges_def)
apply (case_tac "is_sublist [c, d] [ram1, ram2]")
apply simp
apply simp
apply(erule disjE) apply blast
apply (case_tac "c = ram2")
apply (case_tac "between (vertices f) ram2 ram1") apply simp
apply simp
apply (drule is_sublist_distinct_prefix)
apply (subgoal_tac "distinct (ram2 # vs2 @ [ram1])")
apply simp
apply (rule dist_vs2)
apply simp
apply (case_tac "c = ram1")
apply (subgoal_tac "\<not> is_sublist [c, d] (ram2 # vs2 @ [ram1])")
apply simp
apply (rule is_sublist_notlast)
apply (rule_tac dist_vs2)
apply simp
apply simp
apply (simp add: is_sublist_def)
apply (elim exE)
apply (case_tac "between (vertices f) ram2 ram1" rule: rev_exhaust) apply simp
apply simp
apply (case_tac "bs" rule: rev_exhaust) apply simp
apply simp
apply (rule disjI2)
apply (intro exI)
apply simp
apply (rule pre_split_face_p_between) by simp
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def)
from vors have dist_vs2: "distinct (ram2 # vs2 @ [ram1])" apply (simp only:)
apply (rule between_distinct_r12) apply (rule dist_f) apply (rule not_sym) by (simp add: pre_split_face_def)
- from x vors show "x \<in> ?lhs" apply (simp add: dist_f)
+ from x vors show "x \<in> ?lhs"
+ supply [[simproc del: defined_all]]
+ apply (simp add: dist_f)
apply (subgoal_tac "ram1 \<in> set (vertices f)") apply (simp add: verticesFrom_is_nextElem verticesFrom_ram1)
apply (simp add: is_nextElem_def)
apply (case_tac "c = last (between (vertices f) ram2 ram1) \<and> d = ram1") apply simp apply simp apply (rule disjI1)
apply (case_tac "c = ram1 \<and> d = ram2") apply (simp add: is_sublist_def) apply force apply simp
apply (case_tac "c = ram2 \<and> d = hd (between (vertices f) ram2 ram1)")
apply (case_tac "between (vertices f) ram2 ram1") apply simp apply (simp add: is_sublist_def) apply (intro exI)
apply (subgoal_tac "ram1 # ram2 # a # list =
[ram1] @ ram2 # a # (list)") apply assumption apply simp
apply simp
apply (subgoal_tac "is_sublist [c, d] ((ram1 #
[ram2]) @ between (vertices f) ram2 ram1 @ [])")
apply simp apply (rule is_sublist_add) apply (simp add: Edges_def)
by (simp add: pre_split_face_def)
qed
lemma split_face_edges_f_vs2:
assumes vors: "pre_split_face f ram1 ram2 vs"
"vs1 = between (vertices f) ram1 ram2" "vs1 \<noteq> []"
"between (vertices f) ram2 ram1 = []"
shows "edges f = {(ram2, ram1) , (ram1, hd vs1), (last vs1, ram2)} \<union>
Edges vs1" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def)
from vors have dist_vs1: "distinct (ram1 # vs1 @ [ram2])" apply (simp only:)
apply (rule between_distinct_r12) apply (rule dist_f) by (simp add: pre_split_face_def)
from x vors show "x \<in> ?rhs" apply (simp add: dist_f)
apply (subgoal_tac "pre_between (vertices f) ram1 ram2")
apply (drule is_nextElem_or) apply assumption
apply (simp add: Edges_def)
apply (case_tac "is_sublist [c, d] (ram1 # between (vertices f) ram1 ram2 @ [ram2])")
apply simp
apply (case_tac "c = ram1")
apply (case_tac "between (vertices f) ram1 ram2") apply simp
apply simp
apply (drule is_sublist_distinct_prefix)
apply (subgoal_tac "distinct (ram1 # vs1 @ [ram2])") apply simp
apply (rule dist_vs1)
apply simp
apply (case_tac "c = ram2")
apply (subgoal_tac "\<not> is_sublist [c, d] (ram1 # vs1 @ [ram2])") apply simp
apply (rule is_sublist_notlast) apply (rule_tac dist_vs1)
apply simp
apply simp
apply (simp add: is_sublist_def)
apply (elim exE)
apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp
apply simp
apply (case_tac "bs" rule: rev_exhaust) apply simp
apply simp
apply (rule disjI2)
apply (intro exI)
apply simp
apply simp
apply (rule pre_split_face_p_between) by simp
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def)
from vors have dist_vs1: "distinct (ram1 # vs1 @ [ram2])" apply (simp only:)
apply (rule between_distinct_r12) apply (rule dist_f) by (simp add: pre_split_face_def)
- from x vors show "x \<in> ?lhs" apply (simp add: dist_f)
+ from x vors show "x \<in> ?lhs"
+ supply [[simproc del: defined_all]]
+ apply (simp add: dist_f)
apply (subgoal_tac "ram1 \<in> \<V> f") apply (simp add: verticesFrom_is_nextElem verticesFrom_ram1)
apply (simp add: is_nextElem_def)
apply (case_tac "c = ram2 \<and> d = ram1") apply simp apply simp apply (rule disjI1)
apply (case_tac "c = ram1 \<and> d = hd (between (vertices f) ram1 ram2)")
apply (case_tac "between (vertices f) ram1 ram2") apply simp apply (force simp: is_sublist_def) apply simp
apply (case_tac "c = last (between (vertices f) ram1 ram2) \<and> d = ram2")
apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp apply (simp add: is_sublist_def)
apply (intro exI)
apply (subgoal_tac "ram1 # ys @ [y, ram2] =
(ram1 # ys) @ y # ram2 # []") apply assumption apply simp
apply simp
apply (simp add: Edges_def)
apply (subgoal_tac "is_sublist [c, d] ([ram1] @ between (vertices f) ram1 ram2 @ [ram2])")
apply simp apply (rule is_sublist_add) apply simp
by (simp add: pre_split_face_def)
qed
lemma split_face_edges_f:
assumes vors: "pre_split_face f ram1 ram2 vs"
"vs1 = between (vertices f) ram1 ram2" "vs1 \<noteq> []"
"vs2 = between (vertices f) ram2 ram1" "vs2 \<noteq> []"
shows "edges f = {(last vs2, ram1) , (ram1, hd vs1), (last vs1, ram2), (ram2, hd vs2)} \<union>
Edges vs1 \<union> Edges vs2" (is "?lhs = ?rhs")
proof (intro equalityI subsetI)
fix x
assume x: "x \<in> ?lhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def)
from vors have dist_vs1: "distinct (ram1 # vs1 @ [ram2])" apply (simp only:)
apply (rule between_distinct_r12) apply (rule dist_f) by (simp add: pre_split_face_def)
from vors have dist_vs2: "distinct (ram2 # vs2 @ [ram1])" apply (simp only:)
apply (rule between_distinct_r12) apply (rule dist_f) apply (rule not_sym) by (simp add: pre_split_face_def)
from x vors show "x \<in> ?rhs" apply (simp add: dist_f)
apply (subgoal_tac "pre_between (vertices f) ram1 ram2")
apply (drule is_nextElem_or) apply assumption apply (simp add: Edges_def)
apply (case_tac "is_sublist [c, d] (ram1 # between (vertices f) ram1 ram2 @ [ram2])") apply simp
apply (case_tac "c = ram1")
apply (case_tac "between (vertices f) ram1 ram2") apply simp apply simp
apply (drule is_sublist_distinct_prefix) apply (subgoal_tac "distinct (ram1 # vs1 @ [ram2])")
apply simp apply (rule dist_vs1) apply simp
apply (case_tac "c = ram2")
apply (subgoal_tac "\<not> is_sublist [c, d] (ram1 # vs1 @ [ram2])") apply simp
apply (rule is_sublist_notlast) apply (rule_tac dist_vs1) apply simp
apply simp apply (simp add: is_sublist_def) apply (elim exE)
apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp apply simp
apply (case_tac "bs" rule: rev_exhaust) apply simp apply simp
apply (rule disjI2) apply (rule disjI2) apply (rule disjI1) apply (intro exI) apply simp
apply simp
apply (case_tac "c = ram2")
apply (case_tac "between (vertices f) ram2 ram1") apply simp apply simp
apply (drule is_sublist_distinct_prefix) apply (subgoal_tac "distinct (ram2 # vs2 @ [ram1])")
apply simp apply (rule dist_vs2) apply simp
apply (case_tac "c = ram1")
apply (subgoal_tac "\<not> is_sublist [c, d] (ram2 # vs2 @ [ram1])") apply simp
apply (rule is_sublist_notlast) apply (rule_tac dist_vs2) apply simp
apply simp apply (simp add: is_sublist_def) apply (elim exE)
apply (case_tac "between (vertices f) ram2 ram1" rule: rev_exhaust) apply simp apply simp
apply (case_tac "bs" rule: rev_exhaust) apply simp apply simp
apply (rule disjI2) apply (rule disjI2) apply (rule disjI2) apply (intro exI) apply simp
apply (rule pre_split_face_p_between) by simp
next
fix x
assume x: "x \<in> ?rhs"
define c where "c = fst x"
define d where "d = snd x"
with c_def have [simp]: "x = (c,d)" by simp
from vors have dist_f: "distinct (vertices f)" by (simp add: pre_split_face_def)
from vors have dist_vs1: "distinct (ram1 # vs1 @ [ram2])" apply (simp only:)
apply (rule between_distinct_r12) apply (rule dist_f) by (simp add: pre_split_face_def)
from vors have dist_vs2: "distinct (ram2 # vs2 @ [ram1])" apply (simp only:)
apply (rule between_distinct_r12) apply (rule dist_f) apply (rule not_sym) by (simp add: pre_split_face_def)
- from x vors show "x \<in> ?lhs" apply (simp add: dist_f)
+ from x vors show "x \<in> ?lhs"
+ supply [[simproc del: defined_all]]
+ apply (simp add: dist_f)
apply (subgoal_tac "ram1 \<in> \<V> f") apply (simp add: verticesFrom_is_nextElem verticesFrom_ram1)
apply (simp add: is_nextElem_def)
apply (case_tac "c = last (between (vertices f) ram2 ram1) \<and> d = ram1") apply simp apply simp apply (rule disjI1)
apply (case_tac "c = ram1 \<and> d = hd (between (vertices f) ram1 ram2)")
apply (case_tac "between (vertices f) ram1 ram2") apply simp apply (force simp: is_sublist_def) apply simp
apply (case_tac "c = last (between (vertices f) ram1 ram2) \<and> d = ram2")
apply (case_tac "between (vertices f) ram1 ram2" rule: rev_exhaust) apply simp apply (simp add: is_sublist_def)
apply (intro exI)
apply (subgoal_tac "ram1 # ys @ y # ram2 # between (vertices f) ram2 ram1 =
(ram1 # ys) @ y # ram2 # (between (vertices f) ram2 ram1)") apply assumption apply simp apply simp
apply (case_tac "c = ram2 \<and> d = hd (between (vertices f) ram2 ram1)")
apply (case_tac "between (vertices f) ram2 ram1") apply simp apply (simp add: is_sublist_def) apply (intro exI)
apply (subgoal_tac "ram1 # between (vertices f) ram1 ram2 @ ram2 # a # list =
(ram1 # between (vertices f) ram1 ram2) @ ram2 # a # (list)") apply assumption apply simp apply simp
apply (case_tac "(c, d) \<in> Edges (between (vertices f) ram1 ram2)") apply (simp add: Edges_def)
apply (subgoal_tac "is_sublist [c, d] ([ram1] @ between (vertices f) ram1 ram2 @
(ram2 # between (vertices f) ram2 ram1))")
apply simp apply (rule is_sublist_add) apply simp
apply simp
apply (subgoal_tac "is_sublist [c, d] ((ram1 # between (vertices f) ram1 ram2 @
[ram2]) @ between (vertices f) ram2 ram1 @ [])")
apply simp apply (rule is_sublist_add) apply (simp add: Edges_def)
by (simp add: pre_split_face_def)
qed
lemma split_face_edges_f12_f21:
"pre_split_face f ram1 ram2 vs \<Longrightarrow> (f12, f21) = split_face f ram1 ram2 vs \<Longrightarrow>
vs \<noteq> []
\<Longrightarrow> edges f12 \<union> edges f21 = edges f \<union>
{(hd vs, ram1), (ram1, hd vs), (last vs, ram2), (ram2, last vs)} \<union>
Edges vs \<union>
Edges (rev vs)"
apply (case_tac "between (vertices f) ram1 ram2 = []")
apply (case_tac "between (vertices f) ram2 ram1 = []")
apply (simp add: split_face_edges_f12_bet split_face_edges_f21_bet split_face_edges_f_vs1_vs2)
apply force
apply (simp add: split_face_edges_f12_bet split_face_edges_f21 split_face_edges_f_vs1) apply force
apply (case_tac "between (vertices f) ram2 ram1 = []")
apply (simp add: split_face_edges_f21_bet split_face_edges_f12 split_face_edges_f_vs2) apply force
apply (simp add: split_face_edges_f21 split_face_edges_f12 split_face_edges_f) by force
lemma split_face_edges_f12_f21_vs:
"pre_split_face f ram1 ram2 [] \<Longrightarrow> (f12, f21) = split_face f ram1 ram2 []
\<Longrightarrow> edges f12 \<union> edges f21 = edges f \<union>
{(ram2, ram1), (ram1, ram2)}"
apply (case_tac "between (vertices f) ram1 ram2 = []")
apply (case_tac "between (vertices f) ram2 ram1 = []")
apply (simp add: split_face_edges_f12_bet_vs split_face_edges_f21_bet_vs split_face_edges_f_vs1_vs2)
apply force
apply (simp add: split_face_edges_f12_bet_vs split_face_edges_f21_vs split_face_edges_f_vs1) apply force
apply (case_tac "between (vertices f) ram2 ram1 = []")
apply (simp add: split_face_edges_f21_bet_vs split_face_edges_f12_vs split_face_edges_f_vs2) apply force
apply (simp add: split_face_edges_f21_vs split_face_edges_f12_vs split_face_edges_f) by force
lemma split_face_edges_f12_f21_sym:
"f \<in> \<F> g \<Longrightarrow>
pre_split_face f ram1 ram2 vs \<Longrightarrow> (f12, f21) = split_face f ram1 ram2 vs
\<Longrightarrow> ((a,b) \<in> edges f12 \<or> (a,b) \<in> edges f21) =
((a,b) \<in> edges f \<or>
(((b,a) \<in> edges f12 \<or> (b,a) \<in> edges f21) \<and>
((a,b) \<in> edges f12 \<or> (a,b) \<in> edges f21)))"
apply (subgoal_tac "((a,b) \<in> edges f12 \<union> edges f21) =
((a,b) \<in> edges f \<or> ((b,a) \<in> edges f12 \<union> edges f21) \<and> (a,b) \<in> edges f12 \<union> edges f21)") apply force
apply (case_tac "vs = []")
apply (subgoal_tac "pre_split_face f ram1 ram2 []")
apply (drule split_face_edges_f12_f21_vs) apply simp apply simp apply force apply simp
apply (drule split_face_edges_f12_f21) apply simp apply simp
apply simp by force
lemma splitFace_edges_g'_help: "pre_splitFace g ram1 ram2 f vs \<Longrightarrow>
(f12, f21, g') = splitFace g ram1 ram2 f vs \<Longrightarrow> vs \<noteq> [] \<Longrightarrow>
edges g' = edges g \<union> edges f \<union> Edges vs \<union> Edges(rev vs) \<union>
{(ram2, last vs), (hd vs, ram1), (ram1, hd vs), (last vs, ram2)}"
proof -
assume pre: "pre_splitFace g ram1 ram2 f vs"
and fdg: "(f12, f21, g') = splitFace g ram1 ram2 f vs"
and vs_neq: "vs \<noteq> []"
from pre fdg have split: "(f12, f21) = split_face f ram1 ram2 vs"
apply (unfold pre_splitFace_def) apply (elim conjE)
by (simp add: splitFace_split_face)
from fdg pre have "edges g' = (\<Union>\<^bsub>a\<in>set (replace f [f21] (faces g))\<^esub> edges a) \<union>
edges (f12)" by(auto simp: splitFace_def split_def edges_graph_def)
with pre vs_neq show ?thesis apply (simp add: UNION_eq) apply (rule equalityI) apply simp
apply (rule conjI) apply (rule subsetI) apply simp apply (erule bexE) apply (drule replace5)
apply (case_tac "xa \<in> \<F> g") apply simp
apply (subgoal_tac "x \<in> edges g") apply simp
apply (simp add: edges_graph_def) apply force
apply simp
apply (subgoal_tac "pre_split_face f ram1 ram2 vs")
apply (case_tac "between (vertices f) ram2 ram1 = []")
apply (frule split_face_edges_f21_bet) apply (rule split) apply simp apply simp
apply (case_tac "between (vertices f) ram1 ram2 = []")
apply (frule split_face_edges_f_vs1_vs2) apply simp apply simp apply simp apply force
apply (frule split_face_edges_f_vs2) apply simp apply simp apply simp apply force
apply (frule split_face_edges_f21) apply (rule split) apply simp apply simp apply simp
apply (case_tac "between (vertices f) ram1 ram2 = []")
apply (frule split_face_edges_f_vs1) apply simp apply simp apply simp apply simp apply force
apply (frule split_face_edges_f) apply simp apply simp apply simp apply simp apply force
apply simp
apply (subgoal_tac "pre_split_face f ram1 ram2 vs")
apply (case_tac "between (vertices f) ram1 ram2 = []")
apply (frule split_face_edges_f12_bet) apply (rule split) apply simp apply simp
apply (case_tac "between (vertices f) ram2 ram1 = []")
apply (frule split_face_edges_f_vs1_vs2) apply simp apply simp apply simp apply force
apply (frule split_face_edges_f_vs1) apply simp apply simp apply simp apply force
apply (frule split_face_edges_f12) apply (rule split) apply simp apply simp apply simp
apply (case_tac "between (vertices f) ram2 ram1 = []")
apply (frule split_face_edges_f_vs2) apply simp apply simp apply simp apply simp apply force
apply (frule split_face_edges_f) apply simp apply simp apply simp apply simp apply force
apply simp
apply simp
apply (subgoal_tac "pre_split_face f ram1 ram2 vs")
apply (subgoal_tac "(ram2, last vs) \<in> edges f12 \<and> (hd vs, ram1) \<in> edges f12")
apply (rule conjI) apply simp
apply (rule conjI) apply simp
apply (subgoal_tac "(ram1, hd vs) \<in> edges f21 \<and> (last vs, ram2) \<in> edges f21")
apply (rule conjI) apply (rule disjI1) apply (rule bexI) apply (elim conjE) apply simp
apply (rule replace3) apply(erule pre_splitFace_oldF) apply simp
apply (rule conjI) apply (rule disjI1) apply (rule bexI) apply (elim conjE) apply simp
apply (rule replace3) apply(erule pre_splitFace_oldF)
apply simp
apply (subgoal_tac "edges f \<subseteq> {y. \<exists>x\<in>set (replace f [f21] (faces g)). y \<in> edges x} \<union> edges f12")
apply (subgoal_tac "edges g \<subseteq> {y. \<exists>x\<in>set (replace f [f21] (faces g)). y \<in> edges x} \<union> edges f12")
apply (rule conjI) apply simp
apply (rule conjI) apply simp
apply (subgoal_tac "Edges(rev vs) \<subseteq> edges f12") apply (rule conjI) prefer 2 apply blast
apply (subgoal_tac "Edges vs \<subseteq> edges f21")
apply (subgoal_tac "Edges vs \<subseteq> {y. \<exists>x\<in>set (replace f [f21] (faces g)). y \<in> edges x}") apply blast
apply (rule subset_trans) apply assumption apply (rule subsetI) apply simp apply (rule bexI) apply simp
apply (rule replace3) apply(erule pre_splitFace_oldF) apply simp
(* jetzt alle 7 subgoals aufloesen *)
apply (frule split_face_edges_f21_subset) apply (rule split) apply simp apply simp
apply (frule split_face_edges_f12_subset) apply (rule split) apply simp apply simp
apply (simp add: edges_graph_def) apply (rule subsetI) apply simp apply (elim bexE)
apply (case_tac "xa = f") apply simp apply blast
apply (rule disjI1) apply (rule bexI) apply simp apply (rule replace4) apply simp apply force
apply (rule subsetI)
apply (subgoal_tac "\<exists> u v. x = (u,v)") apply (elim exE conjE)
apply (frule split_face_edges_or [OF split]) apply simp
apply (case_tac "(u, v) \<in> edges f12") apply simp apply simp
apply (rule bexI) apply (thin_tac "(u, v) \<in> edges f") apply assumption
apply (rule replace3) apply(erule pre_splitFace_oldF) apply simp apply simp
apply (frule split_face_edges_f21_subset) apply (rule split) apply simp apply simp
apply (frule split_face_edges_f12_subset) apply (rule split) apply simp apply simp
by simp
qed
lemma pre_splitFace_edges_f_in_g: "pre_splitFace g ram1 ram2 f vs \<Longrightarrow> edges f \<subseteq> edges g"
apply (simp add: edges_graph_def) by (force)
lemma pre_splitFace_edges_f_in_g2: "pre_splitFace g ram1 ram2 f vs \<Longrightarrow> x \<in> edges f \<Longrightarrow> x \<in> edges g"
apply (simp add: edges_graph_def) by (force)
lemma splitFace_edges_g': "pre_splitFace g ram1 ram2 f vs \<Longrightarrow>
(f12, f21, g') = splitFace g ram1 ram2 f vs \<Longrightarrow> vs \<noteq> [] \<Longrightarrow>
edges g' = edges g \<union> Edges vs \<union> Edges(rev vs) \<union>
{(ram2, last vs), (hd vs, ram1), (ram1, hd vs), (last vs, ram2)}"
apply (subgoal_tac "edges g \<union> edges f = edges g")
apply (frule splitFace_edges_g'_help) apply simp apply simp apply simp
apply (frule pre_splitFace_edges_f_in_g) by blast
lemma splitFace_edges_g'_vs: "pre_splitFace g ram1 ram2 f [] \<Longrightarrow>
(f12, f21, g') = splitFace g ram1 ram2 f [] \<Longrightarrow>
edges g' = edges g \<union> {(ram1, ram2), (ram2, ram1)}"
proof -
assume pre: "pre_splitFace g ram1 ram2 f []"
and fdg: "(f12, f21, g') = splitFace g ram1 ram2 f []"
from pre fdg have split: "(f12, f21) = split_face f ram1 ram2 []"
apply (unfold pre_splitFace_def) apply (elim conjE)
by (simp add: splitFace_split_face)
from fdg pre have "edges g' = (\<Union>\<^bsub>a\<in>set (replace f [f21] (faces g))\<^esub> edges a) \<union>
edges (f12)" by (auto simp: splitFace_def split_def edges_graph_def)
with pre show ?thesis apply (simp add: UNION_eq) apply (rule equalityI) apply simp
apply (rule conjI) apply (rule subsetI) apply simp apply (erule bexE) apply (drule replace5)
apply (case_tac "xa \<in> \<F> g") apply simp
apply (subgoal_tac "x \<in> edges g") apply simp
apply (simp add: edges_graph_def) apply force
apply simp
apply (subgoal_tac "pre_split_face f ram1 ram2 []")
apply (case_tac "between (vertices f) ram2 ram1 = []") apply (simp add: pre_FaceDiv_between2)
apply (frule split_face_edges_f21_vs) apply (rule split) apply simp apply simp apply simp
apply (case_tac "x = (ram1, ram2)") apply simp apply simp apply (rule disjI2)
apply (rule pre_splitFace_edges_f_in_g2) apply simp
apply (subgoal_tac "pre_split_face f ram1 ram2 []")
apply (frule split_face_edges_f) apply simp apply simp apply (rule pre_FaceDiv_between1) apply simp apply simp
apply simp apply force apply simp apply simp
apply (rule subsetI) apply simp
apply (subgoal_tac "pre_split_face f ram1 ram2 []")
apply (case_tac "between (vertices f) ram1 ram2 = []") apply (simp add: pre_FaceDiv_between1)
apply (frule split_face_edges_f12_vs) apply (rule split) apply simp apply simp apply simp
apply (case_tac "x = (ram2, ram1)") apply simp apply simp apply (rule disjI2)
apply (rule pre_splitFace_edges_f_in_g2) apply simp
apply (subgoal_tac "pre_split_face f ram1 ram2 []")
apply (frule split_face_edges_f) apply simp apply simp apply simp apply (rule pre_FaceDiv_between2) apply simp
apply simp apply force apply simp apply simp
apply simp
apply (subgoal_tac "pre_split_face f ram1 ram2 []")
apply (subgoal_tac "(ram1, ram2) \<in> edges f21")
apply (rule conjI) apply (rule disjI1) apply (rule bexI) apply simp apply (force)
apply (subgoal_tac "(ram2, ram1) \<in> edges f12")
apply (rule conjI) apply force
apply (rule subsetI) apply (simp add: edges_graph_def) apply (elim bexE)
apply (case_tac "xa = f") apply simp
apply (subgoal_tac "\<exists> u v. x = (u,v)") apply (elim exE conjE)
apply (subgoal_tac "pre_split_face f ram1 ram2 []")
apply (frule split_face_edges_or [OF split]) apply simp
apply (case_tac "(u, v) \<in> edges f12") apply simp apply simp apply force apply simp apply simp
apply (rule disjI1) apply (rule bexI) apply simp apply (rule replace4) apply simp apply force
apply (frule split_face_edges_f12_vs) apply simp apply (rule split) apply simp
apply (rule pre_FaceDiv_between1) apply simp apply simp
apply (frule split_face_edges_f21_vs) apply simp apply (rule split) apply simp
apply (rule pre_FaceDiv_between2) apply simp apply simp
by simp
qed
lemma splitFace_edges_incr:
"pre_splitFace g ram1 ram2 f vs \<Longrightarrow>
(f\<^sub>1, f\<^sub>2, g') = splitFace g ram1 ram2 f vs \<Longrightarrow>
edges g \<subseteq> edges g'"
apply(cases vs)
apply(simp add:splitFace_edges_g'_vs)
apply blast
apply(simp add:splitFace_edges_g')
apply blast
done
lemma snd_snd_splitFace_edges_incr:
"pre_splitFace g v\<^sub>1 v\<^sub>2 f vs \<Longrightarrow>
edges g \<subseteq> edges(snd(snd(splitFace g v\<^sub>1 v\<^sub>2 f vs)))"
apply(erule splitFace_edges_incr
[where f\<^sub>1 = "fst(splitFace g v\<^sub>1 v\<^sub>2 f vs)"
and f\<^sub>2 = "fst(snd(splitFace g v\<^sub>1 v\<^sub>2 f vs))"])
apply(auto)
done
(********************* Vorbereitung für subdivFace *****************)
(**** computes the list of ram vertices **********)
subsection \<open>\<open>removeNones\<close>\<close>
definition removeNones :: "'a option list \<Rightarrow> 'a list" where
"removeNones vOptionList \<equiv> [the x. x \<leftarrow> vOptionList, x \<noteq> None]"
(* "removeNones vOptionList \<equiv> map the [x \<in> vOptionList. x \<noteq> None]" *)
declare removeNones_def [simp]
lemma removeNones_inI[intro]: "Some a \<in> set ls \<Longrightarrow> a \<in> set (removeNones ls)" by (induct ls) auto
lemma removeNones_hd[simp]: "removeNones ( Some a # ls) = a # removeNones ls" by auto
lemma removeNones_last[simp]: "removeNones (ls @ [Some a]) = removeNones ls @ [a]" by auto
lemma removeNones_in[simp]: "removeNones (as @ Some a # bs) = removeNones as @ a # removeNones bs" by auto
lemma removeNones_none_hd[simp]: "removeNones ( None # ls) = removeNones ls" by auto
lemma removeNones_none_last[simp]: "removeNones (ls @ [None]) = removeNones ls" by auto
lemma removeNones_none_in[simp]: "removeNones (as @ None # bs) = removeNones (as @ bs)" by auto
lemma removeNones_empty[simp]: "removeNones [] = []" by auto
declare removeNones_def [simp del]
(************* natToVertexList ************)
subsection\<open>\<open>natToVertexList\<close>\<close>
(* Hilfskonstrukt natToVertexList *)
primrec natToVertexListRec ::
"nat \<Rightarrow> vertex \<Rightarrow> face \<Rightarrow> nat list \<Rightarrow> vertex option list"
where
"natToVertexListRec old v f [] = []" |
"natToVertexListRec old v f (i#is) =
(if i = old then None#natToVertexListRec i v f is
else Some (f\<^bsup>i\<^esup> \<bullet> v)
# natToVertexListRec i v f is)"
primrec natToVertexList ::
"vertex \<Rightarrow> face \<Rightarrow> nat list \<Rightarrow> vertex option list"
where
"natToVertexList v f [] = []" |
"natToVertexList v f (i#is) =
(if i = 0 then (Some v)#(natToVertexListRec i v f is) else [])"
subsection \<open>@{const indexToVertexList}\<close>
lemma nextVertex_inj:
"distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow>
i < length (vertices (f::face)) \<Longrightarrow> a < length (vertices f) \<Longrightarrow>
f\<^bsup>a\<^esup>\<bullet>v = f\<^bsup>i\<^esup>\<bullet>v \<Longrightarrow> i = a"
proof -
assume d: "distinct (vertices f)" and v: "v \<in> \<V> f" and i: "i < length (vertices (f::face))"
and a: "a < length (vertices f)" and eq:" f\<^bsup>a\<^esup>\<bullet>v = f\<^bsup>i\<^esup>\<bullet>v"
then have eq: "(verticesFrom f v)!a = (verticesFrom f v)!i " by (simp add: verticesFrom_nth)
define xs where "xs = verticesFrom f v"
with eq have eq: "xs!a = xs!i" by auto
from d v have z: "distinct (verticesFrom f v)" by auto
moreover
from xs_def a v d have "(verticesFrom f v) = take a xs @ xs ! a # drop (Suc a) xs"
by (auto intro: id_take_nth_drop simp: verticesFrom_length)
with eq have "(verticesFrom f v) = take a xs @ xs ! i # drop (Suc a) xs" by simp
moreover
from xs_def i v d have "(verticesFrom f v) = take i xs @ xs ! i # drop (Suc i) xs"
by (auto intro: id_take_nth_drop simp: verticesFrom_length)
ultimately have "take a xs = take i xs" by (rule dist_at1)
moreover
from v d have vertFrom[simp]: "length (vertices f) = length (verticesFrom f v)"
by (auto simp: verticesFrom_length)
from xs_def a i have "a < length xs" "i < length xs" by auto
moreover
have "\<And> a i. a < length xs \<Longrightarrow> i < length xs \<Longrightarrow> take a xs = take i xs \<Longrightarrow> a = i"
proof (induct xs)
case Nil then show ?case by auto
next
case (Cons x xs) then show ?case
apply (cases a) apply auto
apply (cases i) apply auto
apply (cases i) by auto
qed
ultimately show ?thesis by simp
qed
lemma a: "distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow> (\<forall>i \<in> set is. i < length (vertices f)) \<Longrightarrow>
(\<And>a. a < length (vertices f) \<Longrightarrow> hideDupsRec ((f \<bullet> ^^ a) v) [(f \<bullet> ^^ k) v. k \<leftarrow> is] = natToVertexListRec a v f is)"
proof (induct "is")
case Nil then show ?case by simp
next
case (Cons i "is") then show ?case
by (auto simp: nextVertices_def intro: nextVertex_inj)
qed
lemma indexToVertexList_natToVertexList_eq: "distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow>
(\<forall>i \<in> set is. i < length (vertices f)) \<Longrightarrow> is \<noteq> [] \<Longrightarrow>
hd is = 0 \<Longrightarrow> indexToVertexList f v is = natToVertexList v f is"
apply (cases "is") by (auto simp: a [where a = 0, simplified] indexToVertexList_def nextVertices_def)
(********************* Einschub Ende ***************************************)
(******** natToVertexListRec ************)
lemma nvlr_length: "\<And> old. (length (natToVertexListRec old v f ls)) = length ls"
apply (induct ls) by auto
lemma nvl_length[simp]: "hd e = 0 \<Longrightarrow> length (natToVertexList v f e) = length e"
apply (cases "e")
by (auto intro: nvlr_length)
lemma natToVertexListRec_length[simp]: "\<And> e f. length (natToVertexListRec e v f es) = length es"
by (induct es) auto
lemma natToVertexList_length[simp]: "incrIndexList es (length es) (length (vertices f)) \<Longrightarrow>
length (natToVertexList v f es) = length es" apply (case_tac es) by simp_all
lemma natToVertexList_nth_Suc: "incrIndexList es (length es) (length (vertices f)) \<Longrightarrow> Suc n < length es \<Longrightarrow>
(natToVertexList v f es)!(Suc n) = (if (es!n = es!(Suc n)) then None else Some (f\<^bsup>(es!Suc n)\<^esup> \<bullet> v))"
proof -
assume incr: "incrIndexList es (length es) (length (vertices f))" and n: "Suc n < length es"
have rec: "\<And> old n. Suc n < length es \<Longrightarrow>
(natToVertexListRec old v f es)!(Suc n) = (if (es!n = es!(Suc n)) then None else Some (f\<^bsup>(es!Suc n)\<^esup> \<bullet> v))"
proof (induct es)
case Nil then show ?case by auto
next
case (Cons e es)
note cons1 = this
then show ?case
proof (cases es)
case Nil with cons1 show ?thesis by simp
next
case (Cons e' es')
with cons1 show ?thesis
proof (cases n)
case 0 with Cons cons1 show ?thesis by simp
next
case (Suc m) with Cons cons1
have "\<And> old. natToVertexListRec old v f es ! Suc m = (if es ! m = es ! Suc m then None else Some (f\<^bsup>es ! Suc m\<^esup> \<bullet> v))"
by (rule_tac cons1) auto
then show ?thesis apply (cases "e = old") by (simp_all add: Suc)
qed
qed
qed
with n have "natToVertexListRec 0 v f es ! Suc n = (if es ! n = es ! Suc n then None else Some (f\<^bsup>es ! Suc n\<^esup> \<bullet> v))" by (rule_tac rec) auto
with incr show ?thesis by (cases es) auto
qed
lemma natToVertexList_nth_0: "incrIndexList es (length es) (length (vertices f)) \<Longrightarrow> 0 < length es \<Longrightarrow>
(natToVertexList v f es)!0 = Some (f\<^bsup>(es!0)\<^esup> \<bullet> v)"
apply (cases es)
apply (simp_all add: nextVertices_def)
by (subgoal_tac "a = 0") auto
lemma natToVertexList_hd[simp]:
"incrIndexList es (length es) (length (vertices f)) \<Longrightarrow> hd (natToVertexList v f es) = Some v"
apply (cases es) by (simp_all add: nextVertices_def)
lemma nth_last[intro]: "Suc i = length xs \<Longrightarrow> xs!i = last xs"
by (cases xs rule: rev_exhaust) auto
declare incrIndexList_help4 [simp del]
lemma natToVertexList_last[simp]:
"distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow> incrIndexList es (length es) (length (vertices f)) \<Longrightarrow> last (natToVertexList v f es) = Some (last (verticesFrom f v))"
proof -
assume vors: "distinct (vertices f)" "v \<in> \<V> f" and incr: "incrIndexList es (length es) (length (vertices f))"
define n' where "n' = length es - 2"
from incr have "1 < length es" by auto
with n'_def have n'l: "Suc (Suc n') = length es" by arith
from incr n'l have last_ntvl: "(natToVertexList v f es)!(Suc n') = last (natToVertexList v f es)" by auto
from n'l have last_es: "es!(Suc n') = last es" by auto
from n'l have "es!n' = last (butlast es)" apply (cases es rule: rev_exhaust) by (auto simp: nth_append)
with last_es incr have less: "es!n' < es!(Suc n')" by auto
from n'l have "Suc n' < length es" by arith
with incr less have "(natToVertexList v f es)!(Suc n') = (Some (f\<^bsup>(es!Suc n')\<^esup> \<bullet> v))" by (auto dest: natToVertexList_nth_Suc)
with incr last_ntvl last_es have rule1: "last (natToVertexList v f es) = Some (f\<^bsup>((length (vertices f)) - (Suc 0))\<^esup> \<bullet> v)" by auto
from incr have lvf: "1 < length (vertices f)" by auto
with vors have rule2: "verticesFrom f v ! ((length (vertices f)) - (Suc 0)) = f\<^bsup>((length (vertices f)) - (Suc 0))\<^esup> \<bullet> v" by (auto intro!: verticesFrom_nth)
from vors lvf have "verticesFrom f v ! ((length (vertices f)) - (Suc 0)) = last (verticesFrom f v)"
apply (rule_tac nth_last)
by (auto simp: verticesFrom_length)
with rule1 rule2 show ?thesis by auto
qed
lemma indexToVertexList_last[simp]:
"distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow> incrIndexList es (length es) (length (vertices f)) \<Longrightarrow> last (indexToVertexList f v es) = Some (last (verticesFrom f v))"
apply (subgoal_tac "indexToVertexList f v es = natToVertexList v f es") apply simp
apply (rule indexToVertexList_natToVertexList_eq) by auto
lemma nths_take: "\<And> n iset. \<forall> i \<in> iset. i < n \<Longrightarrow> nths (take n xs) iset = nths xs iset"
proof (induct xs)
case Nil then show ?case by simp
next
case (Cons x xs) then show ?case apply (simp add: nths_Cons) apply (cases n) apply simp apply (simp add: nths_Cons) apply (rule Cons) by auto
qed
lemma nths_reduceIndices: "\<And> iset. nths xs iset = nths xs {i. i < length xs \<and> i \<in> iset}"
proof (induct xs)
case Nil then show ?case by simp
next
case (Cons x xs) then
have "nths xs {j. Suc j \<in> iset} = nths xs {i. i < length xs \<and> i \<in> {j. Suc j \<in> iset}}" by (rule_tac Cons)
then show ?case by (simp add: nths_Cons)
qed
lemma natToVertexList_nths1: "distinct (vertices f) \<Longrightarrow>
v \<in> \<V> f \<Longrightarrow> vs = verticesFrom f v \<Longrightarrow>
incrIndexList es (length es) (length vs) \<Longrightarrow> n \<le> length es \<Longrightarrow>
nths (take (Suc (es!(n - 1))) vs) (set (take n es))
= removeNones (take n (natToVertexList v f es))"
proof (induct n)
case 0 then show ?case by simp
next
case (Suc n)
then have "nths (take (Suc (es ! (n - Suc 0))) (verticesFrom f v)) (set (take n es)) = removeNones (take n (natToVertexList v f es))"
"distinct (vertices f)" "v \<in> \<V> f" "vs = verticesFrom f v" "incrIndexList es (length es) (length (verticesFrom f v))" "Suc n \<le> length es" by auto (* does this improve the performance? *)
note suc1 = this
then have lvs: "length vs = length (vertices f)" by (auto intro: verticesFrom_length)
with suc1 have vsne: "vs \<noteq> []" by auto
with suc1 show ?case
proof (cases "natToVertexList v f es ! n")
case None then show ?thesis
proof (cases n)
case 0 with None suc1 lvs show ?thesis by (simp add: take_Suc_conv_app_nth natToVertexList_nth_0)
next
case (Suc n')
with None suc1 lvs have esn: "es!n = es!n'" by (simp add: natToVertexList_nth_Suc split: if_split_asm)
from Suc have n': "n - Suc 0 = n'" by auto
show ?thesis
proof (cases "Suc n = length es")
case True then
have small_n: "n < length es" by auto
from True have "take (Suc n) es = es" by auto
with small_n have "take n es @ [es!n] = es" by (simp add: take_Suc_conv_app_nth)
then have esn_simps: "take n es = butlast es \<and> es!n = last es" by (cases es rule: rev_exhaust) auto
from True Suc have n'l: "Suc n' = length (butlast es)" by auto
then have small_n': "n' < length (butlast es)" by auto
from Suc small_n have take_n': "take (Suc n') (butlast es @ [last es]) = take (Suc n') (butlast es)" by auto
from small_n have es_exh: "es = butlast es @ [last es]" by (cases es rule: rev_exhaust) auto
from n'l have "take (Suc n') (butlast es @ [last es]) = butlast es" by auto
with es_exh have "take (Suc n') es = butlast es" by auto
with small_n Suc have "take n' es @ [es!n'] = (butlast es)" by (simp add: take_Suc_conv_app_nth)
with small_n' have esn'_simps: "take n' es = butlast (butlast es) \<and> es!n' = last (butlast es)"
by (cases "butlast es" rule: rev_exhaust) auto
from suc1 have "last (butlast es) < last es" by auto
with esn esn_simps esn'_simps have False by auto (* have "last (butlast es) = last es" by auto *)
then show ?thesis by auto
next
case False with suc1 have le: "Suc n < length es" by auto
from suc1 le have "es = take (Suc n) es @ es!(Suc n) # drop (Suc (Suc n)) es" by (auto intro: id_take_nth_drop)
with suc1 have "increasing (take (Suc n) es @ es!(Suc n) # drop (Suc (Suc n)) es)" by auto
then have "\<forall> i \<in> (set (take (Suc n) es)). i \<le> es ! (Suc n)" by (auto intro: increasing2)
with suc1 have "\<forall> i \<in> (set (take n es)). i \<le> es ! (Suc n)" by (simp add: take_Suc_conv_app_nth)
then have seq: "nths (take (Suc (es ! Suc n)) (verticesFrom f v)) (set (take n es))
= nths (verticesFrom f v) (set (take n es))"
apply (rule_tac nths_take) by auto
from suc1 have "es = take n es @ es!n # drop (Suc n) es" by (auto intro: id_take_nth_drop)
with suc1 have "increasing (take n es @ es!n # drop (Suc n) es)" by auto
then have "\<forall> i \<in> (set (take n es)). i \<le> es ! n" by (auto intro: increasing2)
with suc1 esn have "\<forall> i \<in> (set (take n es)). i \<le> es ! n'" by (simp add: take_Suc_conv_app_nth)
with Suc have seq2: "nths (take (Suc (es ! n')) (verticesFrom f v)) (set (take n es))
= nths (verticesFrom f v) (set (take n es))"
apply (rule_tac nths_take) by auto
from Suc suc1 have "(insert (es ! n') (set (take n es))) = set (take n es)"
apply auto by (simp add: take_Suc_conv_app_nth)
with esn None suc1 seq seq2 n' show ?thesis by (simp add: take_Suc_conv_app_nth)
qed
qed
next
case (Some v') then show ?thesis
proof (cases n)
case 0
from suc1 lvs have "verticesFrom f v \<noteq> []" by auto
then have "verticesFrom f v = hd (verticesFrom f v) # tl (verticesFrom f v)" by auto
then have "verticesFrom f v = v # tl (verticesFrom f v)" by (simp add: verticesFrom_hd)
then obtain z where "verticesFrom f v = v # z" by auto
then have sub: "nths (verticesFrom f v) {0} = [v]" by (auto simp: nths_Cons)
from 0 suc1 have "es!0 = 0" by (cases es) auto
with 0 Some suc1 lvs sub vsne show ?thesis
by (simp add: take_Suc_conv_app_nth natToVertexList_nth_0 nextVertices_def take_Suc
nths_Cons verticesFrom_hd del:verticesFrom_empty)
next
case (Suc n')
with Some suc1 lvs have esn: "es!n \<noteq> es!n'" by (simp add: natToVertexList_nth_Suc split: if_split_asm)
from suc1 Suc have "Suc n' < length es" by auto
with suc1 lvs esn have "natToVertexList v f es !(Suc n') = Some (f\<^bsup>(es!(Suc n'))\<^esup> \<bullet> v)"
apply (simp add: natToVertexList_nth_Suc)
by (simp add: Suc)
with Suc have "natToVertexList v f es ! n = Some (f\<^bsup>(es!n)\<^esup> \<bullet> v)" by auto
with Some have v': "v' = f\<^bsup>(es!n)\<^esup> \<bullet> v" by simp
from Suc have n': "n - Suc 0 = n'" by auto
from suc1 Suc have "es = take (Suc n') es @ es!n # drop (Suc n) es" by (auto intro: id_take_nth_drop)
with suc1 have "increasing (take (Suc n') es @ es!n # drop (Suc n) es)" by auto
with suc1 Suc have "es!n' \<le> es!n" apply (auto intro!: increasing2)
by (auto simp: take_Suc_conv_app_nth)
with esn have smaller_n: "es!n' < es!n" by auto
from suc1 lvs have smaller: "(es!n) < length vs" by auto
from suc1 smaller lvs have "(verticesFrom f v)!(es!n) = f\<^bsup>(es!n)\<^esup> \<bullet> v" by (auto intro: verticesFrom_nth)
with v' have "(verticesFrom f v)!(es!n) = v'" by auto
then have sub1: "nths ([((verticesFrom f v)!(es!n))])
{j. j + (es!n) : (insert (es ! n) (set (take n es)))} = [v']" by auto
from suc1 smaller lvs have len: "length (take (es ! n) (verticesFrom f v)) = es!n" by auto
have "\<And>x. x \<in> (set (take n es)) \<Longrightarrow> x < (es ! n)"
proof -
fix x
assume x: "x \<in> set (take n es)"
from suc1 Suc have "es = take n' es @ es!n' # drop (Suc n') es" by (auto intro: id_take_nth_drop)
with suc1 have "increasing (take n' es @ es!n' # drop (Suc n') es)" by auto
then have "\<And> x. x \<in> set (take n' es) \<Longrightarrow> x \<le> es!n'" by (auto intro!: increasing2)
with x Suc suc1 have "x \<le> es!n'" by (auto simp: take_Suc_conv_app_nth)
with smaller_n show "x < es!n" by auto
qed
then have "{i. i < es ! n \<and> i \<in> set (take n es)} = (set (take n es))" by auto
then have elim_insert: "{i. i < es ! n \<and> i \<in> insert (es ! n) (set (take n es))} = (set (take n es))" by auto
have "nths (take (es ! n) (verticesFrom f v)) (insert (es ! n) (set (take n es))) =
nths (take (es ! n) (verticesFrom f v)) {i. i < length (take (es ! n) (verticesFrom f v))
\<and> i \<in> (insert (es ! n) (set (take n es)))}" by (rule nths_reduceIndices)
with len have "nths (take (es ! n) (verticesFrom f v)) (insert (es ! n) (set (take n es))) =
nths (take (es ! n) (verticesFrom f v)) {i. i < (es ! n) \<and> i \<in> (insert (es ! n) (set (take n es)))}"
by simp
with elim_insert have sub2: "nths (take (es ! n) (verticesFrom f v)) (insert (es ! n) (set (take n es))) =
nths (take (es ! n) (verticesFrom f v)) (set (take n es))" by simp
define m where "m = es!n - es!n'"
with smaller_n have mgz: "0 < m" by auto
with m_def have esn: "es!n = (es!n') + m" by auto
have helper: "\<And>x. x \<in> (set (take n es)) \<Longrightarrow> x \<le> (es ! n')"
proof -
fix x
assume x: "x \<in> set (take n es)"
from suc1 Suc have "es = take n' es @ es!n' # drop (Suc n') es" by (auto intro: id_take_nth_drop)
with suc1 have "increasing (take n' es @ es!n' # drop (Suc n') es)" by auto
then have "\<And> x. x \<in> set (take n' es) \<Longrightarrow> x \<le> es!n'" by (auto intro!: increasing2)
with x Suc suc1 show "x \<le> es!n'" by (auto simp: take_Suc_conv_app_nth)
qed
define m' where "m' = m - 1"
define Suc_es_n' where "Suc_es_n' = Suc (es!n')"
from smaller smaller_n have "Suc (es!n') < length vs" by auto
then have "min (length vs) (Suc (es ! n')) = Suc (es!n')" by arith
with Suc_es_n'_def have empty: "{j. j + length (take Suc_es_n' vs) \<in> set (take n es)} = {}"
apply auto apply (frule helper) by arith
from Suc_es_n'_def mgz esn m'_def have esn': "es!n = Suc_es_n' + m'" by auto
with smaller have "(take (Suc_es_n' + m') vs) = take (Suc_es_n') vs @ take m' (drop (Suc_es_n') vs)"
by (auto intro: take_add)
with esn' have "nths (take (es ! n) vs) (set (take n es))
= nths (take (Suc_es_n') vs @ take m' (drop (Suc_es_n') vs)) (set (take n es))" by auto
then have "nths (take (es ! n) vs) (set (take n es)) =
nths (take (Suc_es_n') vs) (set (take n es)) @
nths (take m' (drop (Suc_es_n') vs)) {j. j + length (take (Suc_es_n') vs) : (set (take n es))}"
by (simp add: nths_append)
with empty Suc_es_n'_def have "nths (take (es ! n) vs) (set (take n es)) =
nths (take (Suc (es!n')) vs) (set (take n es))" by simp
with suc1 sub2 have sub3: "nths (take (es ! n) (verticesFrom f v)) (insert (es ! n) (set (take n es))) =
nths (take (Suc (es!n')) (verticesFrom f v)) (set (take n es))" by simp
from smaller suc1 have "take (Suc (es ! n)) (verticesFrom f v)
= take (es ! n) (verticesFrom f v) @ [((verticesFrom f v)!(es!n))]"
by (auto simp: take_Suc_conv_app_nth)
with suc1 smaller have
"nths (take (Suc (es ! n)) (verticesFrom f v)) (insert (es ! n) (set (take n es))) =
nths (take (es ! n) (verticesFrom f v)) (insert (es ! n) (set (take n es)))
@ nths ([((verticesFrom f v)!(es!n))]) {j. j + (es!n) : (insert (es ! n) (set (take n es)))}"
by (auto simp: nths_append)
with sub1 sub3 have "nths (take (Suc (es ! n)) (verticesFrom f v)) (insert (es ! n) (set (take n es)))
= nths (take (Suc (es ! n')) (verticesFrom f v)) (set (take n es)) @ [v']" by auto
with Some suc1 lvs n' show ?thesis by (simp add: take_Suc_conv_app_nth)
qed
qed
qed
lemma natToVertexList_nths: "distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow>
incrIndexList es (length es) (length (vertices f)) \<Longrightarrow>
nths (verticesFrom f v) (set es) = removeNones (natToVertexList v f es)"
proof -
assume vors1: "distinct (vertices f)" "v \<in> \<V> f"
"incrIndexList es (length es) (length (vertices f))"
define vs where "vs = verticesFrom f v"
with vors1 have lvs: "length vs = length (vertices f)" by (auto intro: verticesFrom_length)
with vors1 vs_def have vors: "distinct (vertices f)" "v \<in> \<V> f"
"vs = verticesFrom f v" "incrIndexList es (length es) (length vs)" by auto
with lvs have vsne: "vs \<noteq> []" by auto
define n where "n = length es"
then have "es!(n - 1) = last es"
proof (cases n)
case 0 with n_def vors show ?thesis by (cases es) auto
next
case (Suc n')
with n_def have small_n': "n' < length es" by arith
from Suc n_def have "take (Suc n') es = es" by auto
with small_n' have "take n' es @ [es!n'] = es" by (simp add: take_Suc_conv_app_nth)
then have "es!n' = last es" by (cases es rule: rev_exhaust) auto
with Suc show ?thesis by auto
qed
with vors have "es!(n - 1) = (length vs) - 1" by auto
with vsne have "Suc (es! (n - 1)) = (length vs)" by auto
then have take_vs: "take (Suc (es!(n - 1))) vs = vs" by auto
from n_def vors have "n = length (natToVertexList v f es)" by auto
then have take_nTVL: "take n (natToVertexList v f es) = natToVertexList v f es" by auto
from n_def have take_es: "take n es = es" by auto
from n_def have "n \<le> length es" by auto
with vors have "nths (take (Suc (es!(n - 1))) vs) (set (take n es))
= removeNones (take n (natToVertexList v f es))" by (rule natToVertexList_nths1)
with take_vs take_nTVL take_es vs_def show ?thesis by simp
qed
lemma filter_Cons2:
"x \<notin> set ys \<Longrightarrow> [y\<leftarrow>ys. y = x \<or> P y] = [y\<leftarrow>ys. P y]"
by (induct ys) (auto)
lemma natToVertexList_removeNones:
"distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow>
incrIndexList es (length es) (length (vertices f)) \<Longrightarrow>
[x\<leftarrow>verticesFrom f v. x \<in> set (removeNones (natToVertexList v f es))]
= removeNones (natToVertexList v f es)"
proof -
assume vors: "distinct (vertices f)" "v \<in> \<V> f"
"incrIndexList es (length es) (length (vertices f))"
then have dist: "distinct (verticesFrom f v)" by auto
from vors have sub_eq: "nths (verticesFrom f v) (set es)
= removeNones (natToVertexList v f es)" by (rule natToVertexList_nths)
from dist have "[x \<leftarrow> verticesFrom f v.
x \<in> set (nths (verticesFrom f v) (set es))] = removeNones (natToVertexList v f es)"
apply (simp add: filter_in_nths)
by (simp add: sub_eq)
with sub_eq show ?thesis by simp
qed
(**************************** invalidVertexList ************)
definition is_duplicateEdge :: "graph \<Rightarrow> face \<Rightarrow> vertex \<Rightarrow> vertex \<Rightarrow> bool" where
"is_duplicateEdge g f a b \<equiv>
((a, b) \<in> edges g \<and> (a, b) \<notin> edges f \<and> (b, a) \<notin> edges f)
\<or> ((b, a) \<in> edges g \<and> (b, a) \<notin> edges f \<and> (a, b) \<notin> edges f)"
definition invalidVertexList :: "graph \<Rightarrow> face \<Rightarrow> vertex option list \<Rightarrow> bool" where
"invalidVertexList g f vs \<equiv>
\<exists>i < |vs|- 1.
case vs!i of None \<Rightarrow> False
| Some a \<Rightarrow> case vs!(i+1) of None \<Rightarrow> False
| Some b \<Rightarrow> is_duplicateEdge g f a b"
(**************************** subdivFace **********************)
subsection \<open>\<open>pre_subdivFace(')\<close>\<close>
definition pre_subdivFace_face :: "face \<Rightarrow> vertex \<Rightarrow> vertex option list \<Rightarrow> bool" where
"pre_subdivFace_face f v' vOptionList \<equiv>
[v \<leftarrow> verticesFrom f v'. v \<in> set (removeNones vOptionList)]
= (removeNones vOptionList)
\<and> \<not> final f \<and> distinct (vertices f)
\<and> hd (vOptionList) = Some v'
\<and> v' \<in> \<V> f
\<and> last (vOptionList) = Some (last (verticesFrom f v'))
\<and> hd (tl (vOptionList)) \<noteq> last (vOptionList)
\<and> 2 < | vOptionList |
\<and> vOptionList \<noteq> []
\<and> tl (vOptionList) \<noteq> []"
definition pre_subdivFace :: "graph \<Rightarrow> face \<Rightarrow> vertex \<Rightarrow> vertex option list \<Rightarrow> bool" where
"pre_subdivFace g f v' vOptionList \<equiv>
pre_subdivFace_face f v' vOptionList \<and> \<not> invalidVertexList g f vOptionList"
(* zu teilende Fläche, ursprüngliches v, erster Ram-Punkt, Anzahl der überlaufenen NOnes, rest der vol *)
definition pre_subdivFace' :: "graph \<Rightarrow> face \<Rightarrow> vertex \<Rightarrow> vertex \<Rightarrow> nat \<Rightarrow> vertex option list \<Rightarrow> bool" where
"pre_subdivFace' g f v' ram1 n vOptionList \<equiv>
\<not> final f \<and> v' \<in> \<V> f \<and> ram1 \<in> \<V> f
\<and> v' \<notin> set (removeNones vOptionList)
\<and> distinct (vertices f)
\<and> (
[v \<leftarrow> verticesFrom f v'. v \<in> set (removeNones vOptionList)]
= (removeNones vOptionList)
\<and> before (verticesFrom f v') ram1 (hd (removeNones vOptionList))
\<and> last (vOptionList) = Some (last (verticesFrom f v'))
\<and> vOptionList \<noteq> []
\<and> ((v' = ram1 \<and> (0 < n)) \<or> ((v' = ram1 \<and> (hd (vOptionList) \<noteq> Some (last (verticesFrom f v')))) \<or> (v' \<noteq> ram1)))
\<and> \<not> invalidVertexList g f vOptionList
\<and> (n = 0 \<and> hd (vOptionList) \<noteq> None \<longrightarrow> \<not> is_duplicateEdge g f ram1 (the (hd (vOptionList))))
\<or> (vOptionList = [] \<and> v' \<noteq> ram1)
)"
lemma pre_subdivFace_face_in_f[intro]: "pre_subdivFace_face f v ls \<Longrightarrow> Some a \<in> set ls \<Longrightarrow> a \<in> set (verticesFrom f v)"
apply (subgoal_tac "a \<in> set (removeNones ls)") apply (auto simp: pre_subdivFace_face_def)
apply (subgoal_tac "a \<in> set [v\<leftarrow>verticesFrom f v . v \<in> set (removeNones ls)]")
apply (thin_tac "[v\<leftarrow>verticesFrom f v . v \<in> set (removeNones ls)] = removeNones ls") by auto
lemma pre_subdivFace_in_f[intro]: "pre_subdivFace g f v ls \<Longrightarrow> Some a \<in> set ls \<Longrightarrow> a \<in> set (verticesFrom f v)"
by (auto simp: pre_subdivFace_def)
lemma pre_subdivFace_face_in_f'[intro]: "pre_subdivFace_face f v ls \<Longrightarrow> Some a \<in> set ls \<Longrightarrow> a \<in> \<V> f"
apply (cases "a = v") apply (force simp: pre_subdivFace_face_def)
apply (rule verticesFrom_in') apply (rule pre_subdivFace_face_in_f)
by auto
lemma filter_congs_shorten1: "distinct (verticesFrom f v) \<Longrightarrow> [v\<leftarrow>verticesFrom f v . v = a \<or> v \<in> set vs] = (a # vs)
\<Longrightarrow> [v\<leftarrow>verticesFrom f v . v \<in> set vs] = vs"
proof -
assume dist: "distinct (verticesFrom f v)" and eq: "[v\<leftarrow>verticesFrom f v . v = a \<or> v \<in> set vs] = (a # vs)"
have rule1: "\<And> vs a ys. distinct vs \<Longrightarrow> [v\<leftarrow>vs . v = a \<or> v \<in> set ys] = a # ys \<Longrightarrow> [v\<leftarrow>vs. v \<in> set ys] = ys"
proof -
fix vs a ys
assume dist: "distinct vs" and ays: "[v\<leftarrow>vs . v = a \<or> v \<in> set ys] = a # ys"
then have "distinct ([v\<leftarrow>vs . v = a \<or> v \<in> set ys])" by (rule_tac distinct_filter)
with ays have distys: "distinct (a # ys)" by simp
from dist distys ays show "[v\<leftarrow>vs. v \<in> set ys] = ys"
apply (induct vs) by (auto split: if_split_asm simp: filter_Cons2)
qed
from dist eq show ?thesis by (rule_tac rule1)
qed
lemma ovl_shorten: "distinct (verticesFrom f v) \<Longrightarrow>
[v\<leftarrow>verticesFrom f v . v \<in> set (removeNones (va # vol))] = (removeNones (va # vol))
\<Longrightarrow> [v\<leftarrow>verticesFrom f v . v \<in> set (removeNones (vol))] = (removeNones (vol))"
proof -
assume dist: "distinct (verticesFrom f v)"
and vors: "[v\<leftarrow>verticesFrom f v . v \<in> set (removeNones (va # vol))] = (removeNones (va # vol))"
then show ?thesis
proof (cases va)
case None with vors Cons show ?thesis by auto
next
case (Some a) with vors dist show ?thesis by (auto intro!: filter_congs_shorten1)
qed
qed
lemma pre_subdivFace_face_distinct: "pre_subdivFace_face f v vol \<Longrightarrow> distinct (removeNones vol)"
apply (auto dest!: verticesFrom_distinct simp: pre_subdivFace_face_def)
apply (subgoal_tac "distinct ([v\<leftarrow>verticesFrom f v . v \<in> set (removeNones vol)])") apply simp
apply (thin_tac "[v\<leftarrow>verticesFrom f v . v \<in> set (removeNones vol)] = removeNones vol") by auto
lemma invalidVertexList_shorten: "invalidVertexList g f vol \<Longrightarrow> invalidVertexList g f (v # vol)"
apply (simp add: invalidVertexList_def) apply auto apply (rule exI) apply safe
apply (subgoal_tac "(Suc i) < | vol |") apply assumption apply arith
apply auto apply (case_tac "vol!i") by auto
lemma pre_subdivFace_pre_subdivFace': "v \<in> \<V> f \<Longrightarrow> pre_subdivFace g f v (vo # vol) \<Longrightarrow>
pre_subdivFace' g f v v 0 (vol)"
proof -
assume vors: "v \<in> \<V> f" "pre_subdivFace g f v (vo # vol)"
then have vors': "v \<in> \<V> f" "pre_subdivFace_face f v (vo # vol)" "\<not> invalidVertexList g f (vo # vol)"
by (auto simp: pre_subdivFace_def)
then have r: "removeNones vol \<noteq> []" apply (cases "vol" rule: rev_exhaust) by (auto simp: pre_subdivFace_face_def)
then have "Some (hd (removeNones vol)) \<in> set vol" apply (induct vol) apply auto apply (case_tac a) by auto
then have "Some (hd (removeNones vol)) \<in> set (vo # vol)" by auto
with vors' have hd: "hd (removeNones vol) \<in> \<V> f" by (rule_tac pre_subdivFace_face_in_f')
from vors' have "Some v = vo" by (auto simp: pre_subdivFace_face_def)
with vors' have "v \<notin> set (tl (removeNones (vo # vol)))" apply (drule_tac pre_subdivFace_face_distinct) by auto
with vors' r have ne: "v \<noteq> hd (removeNones vol)" by (cases "removeNones vol") (auto simp: pre_subdivFace_face_def)
from vors' have dist: "distinct (removeNones (vo # vol))" apply (rule_tac pre_subdivFace_face_distinct) .
from vors' have invalid: "\<not> invalidVertexList g f vol" by (auto simp: invalidVertexList_shorten)
from ne hd vors' invalid dist show ?thesis apply (unfold pre_subdivFace'_def)
apply (simp add: pre_subdivFace'_def pre_subdivFace_face_def)
apply safe
apply (rule ovl_shorten)
apply (simp add: pre_subdivFace_face_def) apply assumption
apply (rule before_verticesFrom)
apply simp+
apply (simp add: invalidVertexList_def)
apply (erule allE)
apply (erule impE)
apply (subgoal_tac "0 < |vol|")
apply (thin_tac "Suc 0 < | vol |")
apply assumption
apply simp
apply (simp)
apply (case_tac "vol") apply simp by (simp add: is_duplicateEdge_def)
qed
lemma pre_subdivFace'_distinct: "pre_subdivFace' g f v' v n vol \<Longrightarrow> distinct (removeNones vol)"
apply (unfold pre_subdivFace'_def)
apply (cases vol) apply simp+
apply (elim conjE)
apply (drule_tac verticesFrom_distinct) apply assumption
apply (subgoal_tac "distinct [v\<leftarrow>verticesFrom f v' . v \<in> set (removeNones (a # list))]") apply force
apply (thin_tac "[v\<leftarrow>verticesFrom f v' . v \<in> set (removeNones (a # list))] = removeNones (a # list)")
by auto
lemma natToVertexList_pre_subdivFace_face:
"\<not> final f \<Longrightarrow> distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow> 2 < |es| \<Longrightarrow>
incrIndexList es (length es) (length (vertices f)) \<Longrightarrow>
pre_subdivFace_face f v (natToVertexList v f es)"
proof -
assume vors: "\<not> final f" "distinct (vertices f)" "v \<in> \<V> f" "2 < |es|"
"incrIndexList es (length es) (length (vertices f))"
then have lastOvl: "last (natToVertexList v f es) = Some (last (verticesFrom f v))" by auto
from vors have nvl_l: "2 < | natToVertexList v f es |"
by auto
from vors have "distinct [x\<leftarrow>verticesFrom f v . x \<in> set (removeNones (natToVertexList v f es))]" by auto
with vors have "distinct (removeNones (natToVertexList v f es))" by (simp add: natToVertexList_removeNones)
with nvl_l lastOvl have hd_last: "hd (tl (natToVertexList v f es)) \<noteq> last (natToVertexList v f es)" apply auto
apply (cases "natToVertexList v f es") apply simp
apply (case_tac "list" rule: rev_exhaust) apply simp
apply (case_tac "ys") apply simp
apply (case_tac "a") apply simp by simp
from vors lastOvl hd_last nvl_l show ?thesis
apply (auto intro: natToVertexList_removeNones simp: pre_subdivFace_face_def)
apply (cases es) apply auto
apply (cases es) apply auto
apply (subgoal_tac "0 < length list") apply (case_tac list) by (auto split: if_split_asm)
qed
lemma indexToVertexList_pre_subdivFace_face:
"\<not> final f \<Longrightarrow> distinct (vertices f) \<Longrightarrow> v \<in> \<V> f \<Longrightarrow> 2 < |es| \<Longrightarrow>
incrIndexList es (length es) (length (vertices f)) \<Longrightarrow>
pre_subdivFace_face f v (indexToVertexList f v es)"
apply (subgoal_tac "indexToVertexList f v es = natToVertexList v f es") apply simp
apply (rule natToVertexList_pre_subdivFace_face) apply assumption+
apply (rule indexToVertexList_natToVertexList_eq) by auto
lemma subdivFace_subdivFace'_eq: "pre_subdivFace g f v vol \<Longrightarrow> subdivFace g f vol = subdivFace' g f v 0 (tl vol)"
by (simp_all add: subdivFace_def pre_subdivFace_def pre_subdivFace_face_def)
lemma pre_subdivFace'_None:
"pre_subdivFace' g f v' v n (None # vol) \<Longrightarrow>
pre_subdivFace' g f v' v (Suc n) vol"
by(auto simp: pre_subdivFace'_def dest:invalidVertexList_shorten
split:if_split_asm)
declare verticesFrom_between [simp del]
(* zu zeigen:
1. vol \<noteq> [] \<Longrightarrow> [v\<in>verticesFrom f21 v' . v \<in> set (removeNones vol)] = removeNones vol
2. vol \<noteq> [] \<Longrightarrow> before (verticesFrom f21 v') u (hd (removeNones vol))
3. vol \<noteq> [] \<Longrightarrow> distinct (vertices f21)
4. vol \<noteq> [] \<Longrightarrow> last vol = Some (last (verticesFrom f21 v'))
*)
lemma verticesFrom_split: "v # tl (verticesFrom f v) = verticesFrom f v" by (auto simp: verticesFrom_Def)
lemma verticesFrom_v: "distinct (vertices f) \<Longrightarrow> vertices f = a @ v # b \<Longrightarrow> verticesFrom f v = v # b @ a"
by (simp add: verticesFrom_Def)
lemma splitAt_fst[simp]: "distinct xs \<Longrightarrow> xs = a @ v # b \<Longrightarrow> fst (splitAt v xs) = a"
by auto
lemma splitAt_snd[simp]: "distinct xs \<Longrightarrow> xs = a @ v # b \<Longrightarrow> snd (splitAt v xs) = b"
by auto
lemma verticesFrom_splitAt_v_fst[simp]:
"distinct (verticesFrom f v) \<Longrightarrow> fst (splitAt v (verticesFrom f v)) = []"
by (simp add: verticesFrom_Def)
lemma verticesFrom_splitAt_v_snd[simp]:
"distinct (verticesFrom f v) \<Longrightarrow> snd (splitAt v (verticesFrom f v)) = tl (verticesFrom f v)"
by (simp add: verticesFrom_Def)
lemma filter_distinct_at:
"distinct xs \<Longrightarrow> xs = (as @ u # bs) \<Longrightarrow> [v\<leftarrow>xs. v = u \<or> P v] = u # us \<Longrightarrow>
[v\<leftarrow>bs. P v] = us \<and> [v\<leftarrow>as. P v] = []"
apply (subgoal_tac "filter P as @ u # filter P bs = [] @ u # us")
apply (drule local_help') by (auto simp: filter_Cons2)
lemma filter_distinct_at3: "distinct xs \<Longrightarrow> xs = (as @ u # bs) \<Longrightarrow>
[v\<leftarrow>xs. v = u \<or> P v] = u # us \<Longrightarrow> \<forall> z \<in> set zs. z \<in> set as \<or> \<not> ( P z) \<Longrightarrow>
[v\<leftarrow>zs@bs. P v] = us"
apply (drule filter_distinct_at) apply assumption+ apply simp
by (induct zs) auto
lemma filter_distinct_at4: "distinct xs \<Longrightarrow> xs = (as @ u # bs)
\<Longrightarrow> [v\<leftarrow>xs. v = u \<or> v \<in> set us] = u # us
\<Longrightarrow> set zs \<inter> set us \<subseteq> {u} \<union> set as
\<Longrightarrow> [v \<leftarrow> zs@bs. v \<in> set us] = us"
proof -
assume vors: "distinct xs" "xs = (as @ u # bs)"
"[v\<leftarrow>xs. v = u \<or> v \<in> set us] = u # us"
"set zs \<inter> set us \<subseteq> {u} \<union> set as"
then have "distinct ([v\<leftarrow>xs. v = u \<or> v \<in> set us])" apply (rule_tac distinct_filter) by simp
with vors have dist: "distinct (u # us)" by auto
with vors show ?thesis
apply (rule_tac filter_distinct_at3) by assumption+ auto
qed
lemma filter_distinct_at5: "distinct xs \<Longrightarrow> xs = (as @ u # bs)
\<Longrightarrow> [v\<leftarrow>xs. v = u \<or> v \<in> set us] = u # us
\<Longrightarrow> set zs \<inter> set xs \<subseteq> {u} \<union> set as
\<Longrightarrow> [v \<leftarrow> zs@bs. v \<in> set us] = us"
proof -
assume vors: "distinct xs" "xs = (as @ u # bs)"
"[v\<leftarrow>xs. v = u \<or> v \<in> set us] = u # us"
"set zs \<inter> set xs \<subseteq> {u} \<union> set as"
have "set ([v\<leftarrow>xs. v = u \<or> v \<in> set us]) \<subseteq> set xs" by auto
with vors have "set (u # us) \<subseteq> set xs" by simp
then have "set us \<subseteq> set xs" by simp
with vors have "set zs \<inter> set us \<subseteq> set zs \<inter> insert u (set as \<union> set bs)" by auto
with vors show ?thesis apply (rule_tac filter_distinct_at4) apply assumption+ by auto
qed
lemma filter_distinct_at6: "distinct xs \<Longrightarrow> xs = (as @ u # bs)
\<Longrightarrow> [v\<leftarrow>xs. v = u \<or> v \<in> set us] = u # us
\<Longrightarrow> set zs \<inter> set xs \<subseteq> {u} \<union> set as
\<Longrightarrow> [v \<leftarrow> zs@bs. v \<in> set us] = us \<and> [v \<leftarrow> bs. v \<in> set us] = us"
proof -
assume vors: "distinct xs" "xs = (as @ u # bs)"
"[v \<leftarrow> xs. v = u \<or> v \<in> set us] = u # us"
"set zs \<inter> set xs \<subseteq> {u} \<union> set as"
then show ?thesis apply (rule_tac conjI) apply (rule_tac filter_distinct_at5) apply assumption+
apply (drule filter_distinct_at) apply assumption+ by auto
qed
lemma filter_distinct_at_special:
"distinct xs \<Longrightarrow> xs = (as @ u # bs)
\<Longrightarrow> [v\<leftarrow>xs. v = u \<or> v \<in> set us] = u # us
\<Longrightarrow> set zs \<inter> set xs \<subseteq> {u} \<union> set as
\<Longrightarrow> us = hd_us # tl_us
\<Longrightarrow> [v \<leftarrow> zs@bs. v \<in> set us] = us \<and> hd_us \<in> set bs"
proof -
assume vors: "distinct xs" "xs = (as @ u # bs)"
"[v\<leftarrow>xs. v = u \<or> v \<in> set us] = u # us"
"set zs \<inter> set xs \<subseteq> {u} \<union> set as"
"us = hd_us # tl_us"
then have "[v \<leftarrow> zs@bs. v \<in> set us] = us \<and> [v\<leftarrow>bs. v \<in> set us] = us"
by (rule_tac filter_distinct_at6)
with vors show ?thesis apply (rule_tac conjI) apply safe apply simp
apply (subgoal_tac "set (hd_us # tl_us) \<subseteq> set bs") apply simp
apply (subgoal_tac "set [v\<leftarrow>bs . v = hd_us \<or> v \<in> set tl_us] \<subseteq> set bs") apply simp
by (rule_tac filter_is_subset)
qed
(* später ggf. pre_splitFace eliminieren *)
(* nein, Elimination nicht sinnvoll *)
lemma pre_subdivFace'_Some1':
assumes pre_add: "pre_subdivFace' g f v' v n ((Some u) # vol)"
and pre_fdg: "pre_splitFace g v u f ws"
and fdg: "f21 = fst (snd (splitFace g v u f ws))"
and g': "g' = snd (snd (splitFace g v u f ws))"
shows "pre_subdivFace' g' f21 v' u 0 vol"
proof (cases "vol = []")
case True then show ?thesis using pre_add fdg pre_fdg
apply(unfold pre_subdivFace'_def pre_splitFace_def)
apply (simp add: splitFace_def split_face_def split_def del:distinct.simps)
apply (rule conjI)
apply(clarsimp)
apply(rule before_between)
apply(erule (5) rotate_before_vFrom)
apply(erule not_sym)
apply (clarsimp simp:between_distinct between_not_r1 between_not_r2)
apply(blast dest:inbetween_inset)
done
next
case False
with pre_add
have "removeNones vol \<noteq> []" apply (cases "vol" rule: rev_exhaust) by (auto simp: pre_subdivFace'_def)
then have removeNones_split: "removeNones vol = hd (removeNones vol) # tl (removeNones vol)" by auto
from pre_add have dist: "distinct (removeNones ((Some u) # vol))" by (rule_tac pre_subdivFace'_distinct)
from pre_add have v': "v' \<in> \<V> f" by (auto simp: pre_subdivFace'_def)
hence "(vertices f) \<cong> (verticesFrom f v')" by (rule verticesFrom_congs)
hence set_eq: "set (verticesFrom f v') = \<V> f"
apply (rule_tac sym) by (rule congs_pres_nodes)
from pre_fdg fdg have dist_f21: "distinct (vertices f21)" by auto
from pre_add have pre_bet': "pre_between (verticesFrom f v') u v"
apply (simp add: pre_between_def pre_subdivFace'_def)
apply (elim conjE) apply (thin_tac "n = 0 \<longrightarrow> \<not> is_duplicateEdge g f v u")
apply (thin_tac "v' = v \<and> 0 < n \<or> v' = v \<and> u \<noteq> last (verticesFrom f v') \<or> v' \<noteq> v")
apply (auto simp add: before_def)
apply (subgoal_tac "distinct (verticesFrom f v')") apply simp
apply (rule_tac verticesFrom_distinct) by auto
with pre_add have pre_bet: "pre_between (vertices f) u v"
apply (subgoal_tac "(vertices f) \<cong> (verticesFrom f v')")
apply (simp add: pre_between_def pre_subdivFace'_def)
by (auto dest: congs_pres_nodes intro: verticesFrom_congs simp: pre_subdivFace'_def)
from pre_bet pre_add have bet_eq[simp]: "between (vertices f) u v = between (verticesFrom f v') u v"
by (auto intro: verticesFrom_between simp: pre_subdivFace'_def)
from fdg have f21_split_face: "f21 = snd (split_face f v u ws)"
by (simp add: splitFace_def split_def)
then have f21: "f21 = Face (u # between (vertices f) u v @ v # ws) Nonfinal"
by (simp add: split_face_def)
with pre_add pre_bet'
have vert_f21: "vertices f21
= u # snd (splitAt u (verticesFrom f v')) @ fst (splitAt v (verticesFrom f v')) @ v # ws"
apply (drule_tac pre_between_symI)
by (auto simp: pre_subdivFace'_def between_simp2 intro: pre_between_symI)
moreover
from pre_add have "v \<in> set (verticesFrom f v')" by (auto simp: pre_subdivFace'_def before_def)
then have "verticesFrom f v' =
fst (splitAt v (verticesFrom f v')) @ v # snd (splitAt v (verticesFrom f v'))"
by (auto dest: splitAt_ram)
then have m: "v' # tl (verticesFrom f v')
= fst (splitAt v (verticesFrom f v')) @ v # snd (splitAt v (verticesFrom f v'))"
by (simp add: verticesFrom_split)
then have vv': "v \<noteq> v' \<Longrightarrow> fst (splitAt v (verticesFrom f v'))
= v' # tl (fst (splitAt v (verticesFrom f v')))"
by (cases "fst (splitAt v (verticesFrom f v'))") auto
ultimately have "v \<noteq> v' \<Longrightarrow> vertices f21
= u # snd (splitAt u (verticesFrom f v')) @ v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws"
by auto
moreover
with f21 have rule2: "v' \<in> \<V> f21" by auto
with dist_f21 have dist_f21_v': "distinct (verticesFrom f21 v')" by auto
ultimately have m1: "v \<noteq> v' \<Longrightarrow> verticesFrom f21 v'
= v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws @ u # snd (splitAt u (verticesFrom f v'))"
apply auto
apply (subgoal_tac "snd (splitAt v' (vertices f21)) = tl (fst (splitAt v (verticesFrom f v'))) @ v # ws")
apply (subgoal_tac "fst (splitAt v' (vertices f21)) = u # snd (splitAt u (verticesFrom f v'))")
apply (subgoal_tac "verticesFrom f21 v' = v' # snd (splitAt v' (vertices f21)) @ fst (splitAt v' (vertices f21))")
apply simp
apply (intro verticesFrom_v dist_f21) apply force
apply (subgoal_tac "distinct (vertices f21)") apply simp
apply (rule_tac dist_f21)
apply (subgoal_tac "distinct (vertices f21)") apply simp
by (rule_tac dist_f21)
from pre_add have dist_vf_v': "distinct (verticesFrom f v')" by (simp add: pre_subdivFace'_def)
with vert_f21 have m2: "v = v' \<Longrightarrow> verticesFrom f21 v' = v' # ws @ u # snd (splitAt u (verticesFrom f v'))"
apply auto apply (intro verticesFrom_v dist_f21) by simp
from pre_add have u: "u \<in> set (verticesFrom f v')" by (fastforce simp: pre_subdivFace'_def before_def)
then have split_u: "verticesFrom f v'
= fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v'))"
by (auto dest!: splitAt_ram)
then have rule1': "[v \<leftarrow> snd (splitAt u (verticesFrom f v')) . v \<in> set (removeNones vol)] = removeNones vol"
proof -
from split_u have "v' # tl (verticesFrom f v')
= fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v'))"
by (simp add: verticesFrom_split)
have "help": "set [] \<inter> set (verticesFrom f v') \<subseteq> {u} \<union> set (fst (splitAt u (verticesFrom f v')))" by auto
from split_u dist_vf_v' pre_add
have "[v \<leftarrow> [] @ snd (splitAt u (verticesFrom f v')) . v \<in> set (removeNones vol)] = removeNones vol"
apply (rule_tac filter_distinct_at5) apply assumption+
apply (simp add: pre_subdivFace'_def) by (rule "help")
then show ?thesis by auto
qed
then have inSnd_u: "\<And> x. x \<in> set (removeNones vol) \<Longrightarrow> x \<in> set (snd (splitAt u (verticesFrom f v')))"
apply (subgoal_tac "x \<in> set [v \<leftarrow> snd (splitAt u (verticesFrom f v')) . v \<in> set (removeNones vol)] \<Longrightarrow>
x \<in> set (snd (splitAt u (verticesFrom f v')))")
apply force apply (thin_tac "[v \<leftarrow> snd (splitAt u (verticesFrom f v')) . v \<in> set (removeNones vol)] = removeNones vol")
by simp
from split_u dist_vf_v' have notinFst_u: "\<And> x. x \<in> set (removeNones vol) \<Longrightarrow>
x \<notin> set ((fst (splitAt u (verticesFrom f v'))) @ [u])" apply (drule_tac inSnd_u)
apply (subgoal_tac "distinct ( fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v')))")
apply (thin_tac "verticesFrom f v'
= fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v'))")
apply simp apply safe
apply (subgoal_tac "x \<in> set (fst (splitAt u (verticesFrom f v'))) \<inter> set (snd (splitAt u (verticesFrom f v')))")
apply simp
apply (thin_tac "set (fst (splitAt u (verticesFrom f v'))) \<inter> set (snd (splitAt u (verticesFrom f v'))) = {}")
apply simp
by (simp only:)
from rule2 v' have "\<And> a b. is_nextElem (vertices f) a b \<and> a \<in> set (removeNones vol) \<and> b \<in> set (removeNones vol) \<Longrightarrow>
is_nextElem (vertices f21) a b"
proof -
fix a b
assume vors: "is_nextElem (vertices f) a b \<and> a \<in> set (removeNones vol) \<and> b \<in> set (removeNones vol)"
define vor_u where "vor_u = fst (splitAt u (verticesFrom f v'))"
define nach_u where "nach_u = snd (splitAt u (verticesFrom f v'))"
from vors v' have "is_nextElem (verticesFrom f v') a b" by (simp add: verticesFrom_is_nextElem)
moreover
from vors inSnd_u nach_u_def have "a \<in> set (nach_u)" by auto
moreover
from vors inSnd_u nach_u_def have "b \<in> set (nach_u)" by auto
moreover
from split_u vor_u_def nach_u_def have "verticesFrom f v' = vor_u @ u # nach_u" by auto
moreover
note dist_vf_v'
ultimately have "is_sublist [a,b] (nach_u)" apply (simp add: is_nextElem_def split:if_split_asm)
apply (subgoal_tac "b \<noteq> hd (vor_u @ u # nach_u)")
apply simp
apply (subgoal_tac "distinct (vor_u @ (u # nach_u))")
apply (drule is_sublist_at5)
apply simp
apply simp
apply (erule disjE)
apply (drule is_sublist_in1)+
apply (subgoal_tac "b \<in> set vor_u \<inter> set nach_u") apply simp
apply (thin_tac "set vor_u \<inter> set nach_u = {}")
apply simp
apply (erule disjE)
apply (subgoal_tac "distinct ([u] @ nach_u)")
apply (drule is_sublist_at5)
apply simp
apply simp
apply (erule disjE)
apply simp
apply simp
apply simp
apply (subgoal_tac "distinct (vor_u @ (u # nach_u))")
apply (drule is_sublist_at5) apply simp
apply (erule disjE)
apply (drule is_sublist_in1)+
apply simp
apply (erule disjE)
apply (drule is_sublist_in1)+ apply simp
apply simp
apply simp
apply simp
apply (cases "vor_u") by auto
with nach_u_def have "is_sublist [a,b] (snd (splitAt u (verticesFrom f v')))" by auto
then have "is_sublist [a,b] (verticesFrom f21 v')"
apply (cases "v = v'") apply (simp_all add: m1 m2)
apply (subgoal_tac "is_sublist [a, b] ((v' # ws @ [u]) @ snd (splitAt u (verticesFrom f v')) @ [])")
apply simp apply (rule is_sublist_add) apply simp
apply (subgoal_tac "is_sublist [a, b]
((v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws @ [u]) @ (snd (splitAt u (verticesFrom f v'))) @ [])")
apply simp apply (rule is_sublist_add) by simp
with rule2 show "is_nextElem (vertices f) a b \<and> a \<in> set (removeNones vol) \<and> b \<in> set (removeNones vol) \<Longrightarrow>
is_nextElem (vertices f21) a b" apply (simp add: verticesFrom_is_nextElem) by (auto simp: is_nextElem_def)
qed
with pre_add dist_f21 have rule5':
"\<And> a b. (a,b) \<in> edges f \<and> a \<in> set (removeNones vol) \<and> b \<in> set (removeNones vol) \<Longrightarrow> (a, b) \<in> edges f21"
by (simp add: is_nextElem_edges_eq pre_subdivFace'_def)
have rule1: "[v\<leftarrow>verticesFrom f21 v' . v \<in> set (removeNones vol)]
= removeNones vol \<and> hd (removeNones vol) \<in> set (snd (splitAt u (verticesFrom f v')))"
proof (cases "v = v'")
case True
from split_u have "v' # tl (verticesFrom f v')
= fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v'))"
by (simp add: verticesFrom_split)
then have "u \<noteq> v' \<Longrightarrow> fst (splitAt u (verticesFrom f v'))
= v' # tl (fst (splitAt u (verticesFrom f v')))" by (cases "fst (splitAt u (verticesFrom f v'))") auto
moreover
have "v' \<in> set (v' # tl (fst (splitAt u (verticesFrom f v'))))" by simp
ultimately have "u \<noteq> v' \<Longrightarrow> v' \<in> set (fst (splitAt u (verticesFrom f v')))" by simp
moreover
from pre_fdg have "set (v' # ws @ [u]) \<inter> set (verticesFrom f v') \<subseteq> {v', u}"
apply (simp add: set_eq)
by (unfold pre_splitFace_def) auto
ultimately have "help": "set (v' # ws @ [u]) \<inter> set (verticesFrom f v')
\<subseteq> {u} \<union> set (fst (splitAt u (verticesFrom f v')))" apply (rule_tac subset_trans)
apply assumption apply (cases "u = v'") by simp_all
from split_u dist_vf_v' pre_add pre_fdg removeNones_split have
"[v \<leftarrow> (v' # ws @ [u]) @ snd (splitAt u (verticesFrom f v')) . v \<in> set (removeNones vol)]
= removeNones vol \<and> hd (removeNones vol) \<in> set (snd (splitAt u (verticesFrom f v')))"
apply (rule_tac filter_distinct_at_special) apply assumption+
apply (simp add: pre_subdivFace'_def) apply (rule "help") .
with True m2 show ?thesis by auto
next
case False
with m1 dist_f21_v' have ne_uv': "u \<noteq> v'" by auto
define fst_u where "fst_u = fst (splitAt u (verticesFrom f v'))"
define fst_v where "fst_v = fst (splitAt v (verticesFrom f v'))"
from pre_add u dist_vf_v' have "v \<in> set (fst (splitAt u (verticesFrom f v')))"
apply (rule_tac before_dist_r1) by (auto simp: pre_subdivFace'_def)
with fst_u_def have "fst_u = fst (splitAt v (fst (splitAt u (verticesFrom f v'))))
@ v # snd (splitAt v (fst (splitAt u (verticesFrom f v'))))"
by (auto dest: splitAt_ram)
with pre_add fst_v_def pre_bet' have fst_u':"fst_u
= fst_v @ v # snd (splitAt v (fst (splitAt u (verticesFrom f v'))))" by (simp add: pre_subdivFace'_def)
from pre_fdg have "set (v # ws @ [u]) \<inter> set (verticesFrom f v') \<subseteq> {v, u}" apply (simp add: set_eq)
by (unfold pre_splitFace_def) auto
with fst_u' have "set (v # ws @ [u]) \<inter> set (verticesFrom f v') \<subseteq> {u} \<union> set fst_u" by auto
moreover
from fst_u' have "set fst_v \<subseteq> set fst_u" by auto
ultimately
have "(set (v # ws @ [u]) \<union> set fst_v) \<inter> set (verticesFrom f v') \<subseteq> {u} \<union> set fst_u" by auto
with fst_u_def fst_v_def
have "set (fst (splitAt v (verticesFrom f v')) @ v # ws @ [u]) \<inter> set (verticesFrom f v')
\<subseteq> {u} \<union> set (fst (splitAt u (verticesFrom f v')))" by auto
moreover
with False vv' have "v' # tl (fst (splitAt v (verticesFrom f v')))
= fst (splitAt v (verticesFrom f v'))" by auto
ultimately have "set ((v' # tl (fst (splitAt v (verticesFrom f v')))) @ v # ws @ [u]) \<inter> set (verticesFrom f v')
\<subseteq> {u} \<union> set (fst (splitAt u (verticesFrom f v')))"
by (simp only:)
then have "help": "set (v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws @ [u]) \<inter> set (verticesFrom f v')
\<subseteq> {u} \<union> set (fst (splitAt u (verticesFrom f v')))" by auto
from split_u dist_vf_v' pre_add pre_fdg removeNones_split have
"[v \<leftarrow> (v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws @ [u])
@ snd (splitAt u (verticesFrom f v')) . v \<in> set (removeNones vol)]
= removeNones vol \<and> hd (removeNones vol) \<in> set (snd (splitAt u (verticesFrom f v')))"
apply (rule_tac filter_distinct_at_special) apply assumption+
apply (simp add: pre_subdivFace'_def) apply (rule "help") .
with False m1 show ?thesis by auto
qed
from rule1 have "(hd (removeNones vol)) \<in> set (snd (splitAt u (verticesFrom f v')))" by auto
with m1 m2 dist_f21_v' have rule3: "before (verticesFrom f21 v') u (hd (removeNones vol))"
proof -
assume hd_ram: "(hd (removeNones vol)) \<in> set (snd (splitAt u (verticesFrom f v')))"
from m1 m2 dist_f21_v' have "distinct (snd (splitAt u (verticesFrom f v')))" apply (cases "v = v'")
by auto
moreover
define z1 where "z1 = fst (splitAt (hd (removeNones vol)) (snd (splitAt u (verticesFrom f v'))))"
define z2 where "z2 = snd (splitAt (hd (removeNones vol)) (snd (splitAt u (verticesFrom f v'))))"
note z1_def z2_def hd_ram
ultimately have "snd (splitAt u (verticesFrom f v')) = z1 @ (hd (removeNones vol)) # z2"
by (auto intro: splitAt_ram)
with m1 m2 show ?thesis apply (cases "v = v'") apply (auto simp: before_def)
apply (intro exI )
apply (subgoal_tac "v' # ws @ u # z1 @ hd (removeNones vol) # z2 = (v' # ws) @ u # z1 @ hd (removeNones vol) # z2")
apply assumption apply simp
apply (intro exI )
apply (subgoal_tac "v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws @ u # z1 @ hd (removeNones vol) # z2 =
(v' # tl (fst (splitAt v (verticesFrom f v'))) @ v # ws) @ u # z1 @ hd (removeNones vol) # z2")
apply assumption by simp
qed
from rule1 have ne:"(hd (removeNones vol)) \<in> set (snd (splitAt u (verticesFrom f v')))" by auto
with m1 m2 have "last (verticesFrom f21 v') = last (snd (splitAt u (verticesFrom f v')))"
apply (cases "snd (splitAt u (verticesFrom f v'))" rule: rev_exhaust) apply simp_all
apply (cases "v = v'") by simp_all
moreover
from ne have "last (fst (splitAt u (verticesFrom f v')) @ u # snd (splitAt u (verticesFrom f v')))
= last (snd (splitAt u (verticesFrom f v')))" by auto
moreover
note split_u
ultimately have rule4: "last (verticesFrom f v') = last (verticesFrom f21 v')" by simp
have l: "\<And> a b f v. v \<in> set (vertices f) \<Longrightarrow> is_nextElem (vertices f) a b = is_nextElem (verticesFrom f v) a b "
apply (rule is_nextElem_congs_eq) by (rule verticesFrom_congs)
define f12 where "f12 = fst (split_face f v u ws)"
then have f12_fdg: "f12 = fst (splitFace g v u f ws)"
by (simp add: splitFace_def split_def)
from pre_bet pre_add have bet_eq2[simp]: "between (vertices f) v u = between (verticesFrom f v') v u"
apply (drule_tac pre_between_symI)
by (auto intro: verticesFrom_between simp: pre_subdivFace'_def)
from f12_fdg have f12_split_face: "f12 = fst (split_face f v u ws)"
by (simp add: splitFace_def split_def)
then have f12: "f12 = Face (rev ws @ v # between (verticesFrom f v') v u @ [u]) Nonfinal"
by (simp add: split_face_def)
then have "vertices f12 = rev ws @ v # between (verticesFrom f v') v u @ [u]" by simp
with pre_add pre_bet' have vert_f12: "vertices f12
= rev ws @ v # snd (splitAt v (fst (splitAt u (verticesFrom f v')))) @ [u]"
apply (subgoal_tac "between (verticesFrom f v') v u = fst (splitAt u (snd (splitAt v (verticesFrom f v'))))")
apply (simp add: pre_subdivFace'_def)
apply (rule between_simp1)
apply (simp add: pre_subdivFace'_def)
apply (rule pre_between_symI) .
with dist_f21_v' have removeNones_vol_not_f12: "\<And> x. x \<in> set (removeNones vol) \<Longrightarrow> x \<notin> set (vertices f12)"
apply (frule_tac notinFst_u) apply (drule inSnd_u) apply simp
apply (case_tac "v = v'") apply (simp add: m1 m2)
apply (rule conjI) apply force
apply (rule conjI) apply (rule ccontr) apply simp
apply (subgoal_tac "x \<in> set ws \<inter> set (snd (splitAt u (verticesFrom f v')))")
apply simp apply (elim conjE)
apply (thin_tac "set ws \<inter> set (snd (splitAt u (verticesFrom f v'))) = {}")
apply simp
apply force
apply (simp add: m1 m2)
apply (rule conjI) apply force
apply (rule conjI) apply (rule ccontr) apply simp
apply (subgoal_tac "x \<in> set ws \<inter> set (snd (splitAt u (verticesFrom f v')))")
apply simp apply (elim conjE)
apply (thin_tac "set ws \<inter> set (snd (splitAt u (verticesFrom f v'))) = {}") apply simp
by force
from pre_fdg f12_split_face have dist_f12: "distinct (vertices f12)" by (auto intro: split_face_distinct1')
then have removeNones_vol_edges_not_f12: "\<And> x y. x \<in> set (removeNones vol) \<Longrightarrow> (x,y) \<notin> edges f12" (* ? *)
apply (drule_tac removeNones_vol_not_f12) by auto
from dist_f12 have removeNones_vol_edges_not_f12': "\<And> x y. y \<in> set (removeNones vol) \<Longrightarrow> (x,y) \<notin> edges f12"
apply (drule_tac removeNones_vol_not_f12) by auto
from f12_fdg pre_fdg g' fdg have face_set_eq: "\<F> g' \<union> {f} = {f12, f21} \<union> \<F> g"
apply (rule_tac splitFace_faces_1)
by (simp_all)
have rule5'': "\<And> a b. (a,b) \<in> edges g' \<and> (a,b) \<notin> edges g
\<and> a \<in> set (removeNones vol) \<and> b \<in> set (removeNones vol) \<Longrightarrow> (a, b) \<in> edges f21" (* ? *)
apply (simp add: edges_graph_def) apply safe
apply (case_tac "x = f") apply simp apply (rule rule5') apply safe
apply (subgoal_tac "x \<in> \<F> g' \<union> {f}") apply (thin_tac "x \<noteq> f")
apply (thin_tac "x \<in> set (faces g')") apply (simp only: add: face_set_eq)
apply safe apply (drule removeNones_vol_edges_not_f12) by auto
have rule5''': "\<And> a b. (a,b) \<in> edges g' \<and> (a,b) \<notin> edges g
\<and> a \<in> set (removeNones vol) \<and> b \<in> set (removeNones vol) \<Longrightarrow> (a, b) \<in> edges f21" (* ? *)
apply (simp add: edges_graph_def) apply safe
apply (case_tac "x = f") apply simp apply (rule rule5') apply safe
apply (subgoal_tac "x \<in> \<F> g' \<union> {f}") apply (thin_tac "x \<noteq> f")
apply (thin_tac "x \<in> \<F> g'") apply (simp only: add: face_set_eq)
apply safe apply (drule removeNones_vol_edges_not_f12) by auto
from pre_fdg fdg f12_fdg g' have edges_g'1: "ws \<noteq> [] \<Longrightarrow> edges g' = edges g \<union> Edges ws \<union> Edges(rev ws) \<union>
{(u, last ws), (hd ws, v), (v, hd ws), (last ws, u)}"
apply (rule_tac splitFace_edges_g') apply simp
apply (subgoal_tac "(f12, f21, g') = splitFace g v u f ws") apply assumption by auto
from pre_fdg fdg f12_fdg g' have edges_g'2: "ws = [] \<Longrightarrow> edges g' = edges g \<union>
{(v, u), (u, v)}"
apply (rule_tac splitFace_edges_g'_vs) apply simp
apply (subgoal_tac "(f12, f21, g') = splitFace g v u f []") apply assumption by auto
from f12_split_face f21_split_face have split: "(f12,f21) = split_face f v u ws" by simp
from pre_add have "\<not> invalidVertexList g f vol"
by (auto simp: pre_subdivFace'_def dest: invalidVertexList_shorten)
then have rule5: "\<not> invalidVertexList g' f21 vol"
apply (simp add: invalidVertexList_def)
apply (intro allI impI)
apply (case_tac "vol!i") apply simp+
apply (case_tac "vol!Suc i") apply simp+
apply (subgoal_tac "\<not> is_duplicateEdge g f a aa")
apply (thin_tac "\<forall>i<|vol| - Suc 0. \<not> (case vol ! i of None \<Rightarrow> False
| Some a \<Rightarrow> case_option False (is_duplicateEdge g f a) (vol ! (i+1)))")
apply (simp add: is_duplicateEdge_def)
apply (subgoal_tac "a \<in> set (removeNones vol) \<and> aa \<in> set (removeNones vol)")
apply (rule conjI)
apply (rule impI)
apply (case_tac "(a, aa) \<in> edges f")
apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (frule split_face_edges_or [OF split]) apply simp
apply (simp add: removeNones_vol_edges_not_f12)
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
apply (case_tac "(aa, a) \<in> edges f")
apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (frule split_face_edges_or [OF split]) apply simp
apply (simp add: removeNones_vol_edges_not_f12)
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
apply simp
apply (case_tac "ws = []") apply (frule edges_g'2) apply simp
apply (subgoal_tac "pre_split_face f v u []")
apply (subgoal_tac "(f12, f21) = split_face f v u ws")
apply (case_tac "between (vertices f) u v = []")
apply (frule split_face_edges_f21_bet_vs) apply simp apply simp
apply simp
apply (frule split_face_edges_f21_vs) apply simp apply simp apply simp
apply (case_tac "a = v \<and> aa = u") apply simp apply simp
apply (rule split)
apply (subgoal_tac "pre_split_face f v u ws") apply simp
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
apply (frule edges_g'1) apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (subgoal_tac "(f12, f21) = split_face f v u ws")
apply (case_tac "between (vertices f) u v = []")
apply (frule split_face_edges_f21_bet) apply simp apply simp apply simp
apply simp
apply (case_tac "a = u \<and> aa = last ws") apply simp apply simp
apply (case_tac "a = hd ws \<and> aa = v") apply simp apply simp
apply (case_tac "a = v \<and> aa = hd ws") apply simp apply simp
apply (case_tac "a = last ws \<and> aa = u") apply simp apply simp
apply (case_tac "(a, aa) \<in> Edges ws") apply simp
apply simp
apply (frule split_face_edges_f21) apply simp apply simp apply simp apply simp
apply (force)
apply (rule split)
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
apply (rule impI)
apply (case_tac "(aa,a) \<in> edges f") apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (frule split_face_edges_or [OF split]) apply simp
apply (simp add: removeNones_vol_edges_not_f12)
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
apply (case_tac "(a,aa) \<in> edges f") apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (frule split_face_edges_or [OF split]) apply simp
apply (simp add: removeNones_vol_edges_not_f12)
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
apply simp
apply (case_tac "ws = []") apply (frule edges_g'2) apply simp
apply (subgoal_tac "pre_split_face f v u []")
apply (subgoal_tac "(f12, f21) = split_face f v u ws")
apply (case_tac "between (vertices f) u v = []")
apply (frule split_face_edges_f21_bet_vs) apply simp apply simp
apply simp
apply (frule split_face_edges_f21_vs) apply simp apply simp apply simp
apply force
apply (rule split)
apply (subgoal_tac "pre_split_face f v u ws") apply simp
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
apply (frule edges_g'1) apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (subgoal_tac "(f12, f21) = split_face f v u ws")
apply (case_tac "between (vertices f) u v = []")
apply (frule split_face_edges_f21_bet) apply simp apply simp apply simp
apply (force)
apply (frule split_face_edges_f21) apply simp apply simp apply simp apply simp
apply (force)
apply (rule split)
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
apply (rule conjI)
apply (subgoal_tac "Some a \<in> set vol") apply (induct vol) apply simp apply force
apply (subgoal_tac "vol ! i \<in> set vol") apply simp
apply (rule nth_mem) apply arith
apply (subgoal_tac "Some aa \<in> set vol") apply (induct vol) apply simp apply force
apply (subgoal_tac "vol ! (Suc i) \<in> set vol") apply simp apply (rule nth_mem) apply arith
by auto
from pre_fdg dist_f21 v' have dists: "distinct (vertices f)" "distinct (vertices f12)"
"distinct (vertices f21)" "v' \<in> \<V> f"
apply auto defer
apply (drule splitFace_distinct2) apply (simp add: f12_fdg)
apply (unfold pre_splitFace_def) by simp
with pre_fdg have edges_or: "\<And> a b. (a,b) \<in> edges f \<Longrightarrow> (a,b) \<in> edges f12 \<or> (a,b) \<in> edges f21"
apply (rule_tac split_face_edges_or) apply (simp add: f12_split_face f21_split_face)
by simp+
from pre_fdg have dist_f: "distinct (vertices f)" apply (unfold pre_splitFace_def) by simp
(* lemma *)
from g' have edges_g': "edges g'
= (UN h:set(replace f [snd (split_face f v u ws)] (faces g)). edges h)
\<union> edges (fst (split_face f v u ws))"
by (auto simp add: splitFace_def split_def edges_graph_def)
(* lemma *)
from pre_fdg edges_g' have edges_g'_or:
"\<And> a b. (a,b) \<in> edges g' \<Longrightarrow>
(a,b) \<in> edges g \<or> (a,b) \<in> edges f12 \<or> (a,b) \<in> edges f21"
apply simp apply (case_tac "(a, b) \<in> edges (fst (split_face f v u ws))")
apply (simp add:f12_split_face) apply simp
apply (elim bexE) apply (simp add: f12_split_face) apply (case_tac "x \<in> \<F> g")
apply (induct g) apply (simp add: edges_graph_def) apply (rule disjI1)
apply (rule bexI) apply simp apply simp
apply (drule replace1) apply simp by (simp add: f21_split_face)
have rule6: "0 < |vol| \<Longrightarrow> \<not> invalidVertexList g f (Some u # vol) \<Longrightarrow>
(\<exists>y. hd vol = Some y) \<longrightarrow> \<not> is_duplicateEdge g' f21 u (the (hd vol))"
apply (rule impI)
apply (erule exE) apply simp apply (case_tac vol) apply simp+
apply (simp add: invalidVertexList_def) apply (erule allE) apply (erule impE) apply force
apply (simp)
apply (subgoal_tac "y \<notin> \<V> f12") defer apply (rule removeNones_vol_not_f12) apply simp
apply (simp add: is_duplicateEdge_def)
apply (subgoal_tac "y \<in> set (removeNones vol)")
apply (rule conjI)
apply (rule impI)
apply (case_tac "(u, y) \<in> edges f") apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (frule split_face_edges_or [OF split]) apply simp
apply (simp add: removeNones_vol_edges_not_f12')
apply (rule pre_splitFace_pre_split_face) apply simp apply (rule pre_fdg)
apply (case_tac "(y, u) \<in> edges f") apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (frule split_face_edges_or [OF split]) apply simp
apply (simp add: removeNones_vol_edges_not_f12)
apply (rule pre_splitFace_pre_split_face) apply simp apply (rule pre_fdg)
apply simp
apply (case_tac "ws = []") apply (frule edges_g'2) apply simp
apply (subgoal_tac "pre_split_face f v u []")
apply (subgoal_tac "(f12, f21) = split_face f v u ws")
apply (case_tac "between (vertices f) u v = []")
apply (frule split_face_edges_f21_bet_vs) apply simp apply simp apply simp
apply (frule split_face_edges_f21_vs) apply simp apply simp apply simp
apply force
apply (rule split)
apply (subgoal_tac "pre_split_face f v u ws") apply simp
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
apply (frule edges_g'1) apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (subgoal_tac "(f12, f21) = split_face f v u ws")
apply (case_tac "between (vertices f) u v = []")
apply (frule split_face_edges_f21_bet) apply simp apply simp apply simp
apply (force)
apply (frule split_face_edges_f21) apply simp apply simp apply simp apply simp
apply (force)
apply (rule split)
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
apply (rule impI)
apply (case_tac "(u, y) \<in> edges f") apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (frule split_face_edges_or [OF split]) apply simp apply (simp add: removeNones_vol_edges_not_f12')
apply (rule pre_splitFace_pre_split_face) apply simp apply (rule pre_fdg)
apply (case_tac "(y, u) \<in> edges f") apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (frule split_face_edges_or [OF split]) apply simp apply (simp add: removeNones_vol_edges_not_f12)
apply (rule pre_splitFace_pre_split_face) apply simp apply (rule pre_fdg)
apply simp
apply (case_tac "ws = []") apply (frule edges_g'2) apply simp
apply (subgoal_tac "pre_split_face f v u []")
apply (subgoal_tac "(f12, f21) = split_face f v u ws")
apply (case_tac "between (vertices f) u v = []")
apply (frule split_face_edges_f21_bet_vs) apply simp apply simp
apply simp
apply (frule split_face_edges_f21_vs) apply simp apply simp apply simp
apply force
apply (rule split)
apply (subgoal_tac "pre_split_face f v u ws") apply simp
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
apply (frule edges_g'1) apply simp
apply (subgoal_tac "pre_split_face f v u ws")
apply (subgoal_tac "(f12, f21) = split_face f v u ws")
apply (case_tac "between (vertices f) u v = []")
apply (frule split_face_edges_f21_bet) apply simp apply simp apply simp
apply (force)
apply (frule split_face_edges_f21) apply simp apply simp apply simp apply simp
apply (force)
apply (rule split)
apply (rule pre_splitFace_pre_split_face) apply (rule pre_fdg)
by simp
have u21: "u \<in> \<V> f21" by(simp add:f21)
from fdg have "\<not> final f21"
by(simp add:splitFace_def split_face_def split_def)
with pre_add rule1 rule2 rule3 rule4 rule5 rule6 dist_f21 False dist u21
show ?thesis by (simp_all add: pre_subdivFace'_def l)
qed
lemma before_filter: "\<And> ys. filter P xs = ys \<Longrightarrow> distinct xs \<Longrightarrow> before ys u v \<Longrightarrow> before xs u v"
supply subst_all [simp del]
apply (subgoal_tac "P u")
apply (subgoal_tac "P v")
apply (subgoal_tac "pre_between xs u v")
apply (rule ccontr) apply (simp add: before_xor)
apply (subgoal_tac "before ys v u")
apply (subgoal_tac "\<not> before ys v u")
apply simp
apply (rule before_dist_not1) apply force apply simp
apply (simp add: before_def) apply (elim exE) apply simp
apply (subgoal_tac "a @ u # b @ v # c = filter P aa @ v # filter P ba @ u # filter P ca")
apply (intro exI) apply assumption
apply simp
apply (subgoal_tac "u \<in> set ys \<and> v \<in> set ys \<and> u \<noteq> v") apply (simp add: pre_between_def) apply force
apply (subgoal_tac "distinct ys")
apply (simp add: before_def) apply (elim exE) apply simp
apply force
apply (subgoal_tac "v \<in> set (filter P xs)") apply force
apply (simp add: before_def) apply (elim exE) apply simp
apply (subgoal_tac "u \<in> set (filter P xs)") apply force
apply (simp add: before_def) apply (elim exE) by simp
lemma pre_subdivFace'_Some2: "pre_subdivFace' g f v' v 0 ((Some u) # vol) \<Longrightarrow> pre_subdivFace' g f v' u 0 vol"
apply (cases "vol = []")
apply (simp add: pre_subdivFace'_def)
apply (cases "u = v'") apply simp
apply(rule verticesFrom_in')
apply(rule last_in_set)
apply(simp add:verticesFrom_Def)
apply clarsimp
apply (simp add: pre_subdivFace'_def)
apply (elim conjE)
apply (thin_tac "v' = v \<and> u \<noteq> last (verticesFrom f v') \<or> v' \<noteq> v")
apply auto
apply(rule verticesFrom_in'[where v = v'])
apply(clarsimp simp:before_def)
apply simp
apply (rule ovl_shorten) apply simp
apply (subgoal_tac "[v \<leftarrow> verticesFrom f v' . v \<in> set (removeNones ((Some u) # vol))] = removeNones ((Some u) # vol)")
apply assumption
apply simp
apply (rule before_filter)
apply assumption
apply simp
apply (simp add: before_def)
apply (intro exI)
apply (subgoal_tac "u # removeNones vol = [] @ u # [] @ hd (removeNones vol) # tl (removeNones vol)") apply assumption
apply simp
apply (subgoal_tac "removeNones vol \<noteq> []") apply simp
apply (cases vol rule: rev_exhaust) apply simp_all
apply (simp add: invalidVertexList_shorten)
apply (simp add: is_duplicateEdge_def)
apply (case_tac "vol") apply simp
apply simp
apply (simp add: invalidVertexList_def)
apply (elim allE)
apply (rotate_tac -1)
apply (erule impE)
apply (subgoal_tac "0 < Suc |list|")
apply assumption
apply simp
apply simp
by (simp add: is_duplicateEdge_def)
lemma pre_subdivFace'_preFaceDiv: "pre_subdivFace' g f v' v n ((Some u) # vol)
\<Longrightarrow> f \<in> \<F> g \<Longrightarrow> (f \<bullet> v = u \<longrightarrow> n \<noteq> 0) \<Longrightarrow> \<V> f \<subseteq> \<V> g
\<Longrightarrow> pre_splitFace g v u f [countVertices g ..< countVertices g + n]"
proof -
assume pre_add: "pre_subdivFace' g f v' v n ((Some u) # vol)" and f: "f \<in> \<F> g"
and nextVert: "(f \<bullet> v = u \<longrightarrow> n \<noteq> 0)" and subset: "\<V> f \<subseteq> \<V> g"
have "distinct [countVertices g ..< countVertices g + n]" by (induct n) auto
moreover
have "\<V> g \<inter> set [countVertices g ..< countVertices g + n] = {}"
apply (cases g) by auto
with subset have "\<V> f \<inter> set [countVertices g ..< countVertices g + n] = {}" by auto
moreover
from pre_add have "\<V> f = set (verticesFrom f v')" apply (intro congs_pres_nodes verticesFrom_congs)
by (simp add: pre_subdivFace'_def)
with pre_add have "help": "v \<in> \<V> f \<and> u \<in> \<V> f \<and> v \<noteq> u"
apply (simp add: pre_subdivFace'_def before_def)
apply (elim conjE exE)
apply (subgoal_tac "distinct (verticesFrom f v')") apply force
apply (rule verticesFrom_distinct) by simp_all
moreover
from "help" pre_add nextVert have help1: "is_nextElem (vertices f) v u \<Longrightarrow> 0 < n" apply auto
apply (simp add: nextVertex_def)
by (simp add: nextElem_is_nextElem pre_subdivFace'_def)
moreover
have help2: "before (verticesFrom f v') v u \<Longrightarrow> distinct (verticesFrom f v') \<Longrightarrow> v \<noteq> v' \<Longrightarrow> \<not> is_nextElem (verticesFrom f v') u v"
apply (simp add: before_def is_nextElem_def verticesFrom_hd is_sublist_def) apply safe
apply (frule dist_at)
apply simp
apply (thin_tac "verticesFrom f v' = a @ v # b @ u # c")
apply (subgoal_tac "verticesFrom f v' = (as @ [u]) @ v # bs") apply assumption
apply simp apply (subgoal_tac "distinct (a @ v # b @ u # c)") apply force by simp
note pre_add f
moreover(*
have "\<And> m. {k. k < m} \<inter> {k. m \<le> k \<and> k < (m + n)} = {}" by auto
moreover*)
from pre_add f help2 help1 "help" have "[countVertices g..<countVertices g + n] = [] \<Longrightarrow> (v, u) \<notin> edges f \<and> (u, v) \<notin> edges f"
apply (cases "0 < n") apply (induct g) apply simp+
apply (simp add: pre_subdivFace'_def)
apply (rule conjI) apply force
apply (simp split: if_split_asm)
apply (rule ccontr) apply simp
apply (subgoal_tac "v = v'") apply simp apply (elim conjE) apply (simp only:)
apply (rule verticesFrom_is_nextElem_last) apply force apply force
apply (simp add: verticesFrom_is_nextElem [symmetric])
apply (cases "v = v'") apply simp
apply (subgoal_tac "v' \<in> \<V> f")
apply (thin_tac "u \<in> \<V> f")
apply (simp add: verticesFrom_is_nextElem)
apply (rule ccontr) apply simp
apply (subgoal_tac "v' \<in> \<V> f")
apply (drule verticesFrom_is_nextElem_hd) apply simp+
apply (elim conjE) apply (drule help2)
apply simp apply simp
apply (subgoal_tac "is_nextElem (vertices f) u v = is_nextElem (verticesFrom f v') u v")
apply simp
apply (rule verticesFrom_is_nextElem) by simp
ultimately
show ?thesis
apply (simp add: pre_subdivFace'_def)
apply (unfold pre_splitFace_def)
apply simp
apply (cases "0 < n") apply (induct g) apply (simp add: ivl_disj_int)
apply (auto simp: invalidVertexList_def is_duplicateEdge_def)
done
qed
lemma pre_subdivFace'_Some1:
"pre_subdivFace' g f v' v n ((Some u) # vol)
\<Longrightarrow> f \<in> \<F> g \<Longrightarrow> (f \<bullet> v = u \<longrightarrow> n \<noteq> 0) \<Longrightarrow> \<V> f \<subseteq> \<V> g
\<Longrightarrow> f21 = fst (snd (splitFace g v u f [countVertices g ..< countVertices g + n]))
\<Longrightarrow> g' = snd (snd (splitFace g v u f [countVertices g ..< countVertices g + n]))
\<Longrightarrow> pre_subdivFace' g' f21 v' u 0 vol"
apply (subgoal_tac "pre_splitFace g v u f [countVertices g ..< countVertices g + n]")
apply (rule pre_subdivFace'_Some1') apply assumption+
apply (simp)
apply (rule pre_subdivFace'_preFaceDiv)
by auto
end
diff --git a/thys/Gabow_SCC/Find_Path_Impl.thy b/thys/Gabow_SCC/Find_Path_Impl.thy
--- a/thys/Gabow_SCC/Find_Path_Impl.thy
+++ b/thys/Gabow_SCC/Find_Path_Impl.thy
@@ -1,415 +1,414 @@
section \<open>Implementation of Safety Property Model Checker \label{sec:find_path_impl}\<close>
theory Find_Path_Impl
imports
Find_Path
CAVA_Automata.Digraph_Impl
begin
section \<open>Workset Algorithm\<close>
text \<open>A simple implementation is by a workset algorithm.\<close>
definition "ws_update E u p V ws \<equiv> RETURN (
V \<union> E``{u},
ws ++ (\<lambda>v. if (u,v)\<in>E \<and> v\<notin>V then Some (u#p) else None))"
definition "s_init U0 \<equiv> RETURN (None,U0,\<lambda>u. if u\<in>U0 then Some [] else None)"
definition "wset_find_path E U0 P \<equiv> do {
ASSERT (finite U0);
s0 \<leftarrow> s_init U0;
(res,_,_) \<leftarrow> WHILET
(\<lambda>(res,V,ws). res=None \<and> ws\<noteq>Map.empty)
(\<lambda>(res,V,ws). do {
ASSERT (ws\<noteq>Map.empty);
(u,p) \<leftarrow> SPEC (\<lambda>(u,p). ws u = Some p);
let ws=ws |` (-{u});
if P u then
RETURN (Some (rev p,u),V,ws)
else do {
ASSERT (finite (E``{u}));
ASSERT (dom ws \<subseteq> V);
(V,ws) \<leftarrow> ws_update E u p V ws;
RETURN (None,V,ws)
}
}) s0;
RETURN res
}"
lemma wset_find_path_correct:
fixes E :: "('v\<times>'v) set"
shows "wset_find_path E U0 P \<le> find_path E U0 P"
proof -
define inv where "inv = (\<lambda>(res,V,ws). case res of
None \<Rightarrow>
dom ws\<subseteq>V
\<and> finite (dom ws) \<comment> \<open>Derived\<close>
\<and> V\<subseteq>E\<^sup>*``U0
\<and> E``(V-dom ws) \<subseteq> V
\<and> (\<forall>v\<in>V-dom ws. \<not> P v)
\<and> U0 \<subseteq> V
\<and> (\<forall>v p. ws v = Some p
\<longrightarrow> ((\<forall>v\<in>set p. \<not>P v) \<and> (\<exists>u0\<in>U0. path E u0 (rev p) v)))
| Some (p,v) \<Rightarrow> (\<exists>u0\<in>U0. path E u0 p v \<and> P v \<and> (\<forall>v\<in>set p. \<not>P v)))"
define var where "var = inv_image
(brk_rel (finite_psupset (E\<^sup>*``U0) <*lex*> measure (card o dom)))
(\<lambda>(res::('v list \<times> 'v) option,V::'v set,ws::'v\<rightharpoonup>'v list).
(res\<noteq>None,V,ws))"
(*have [simp, intro!]: "wf var"
unfolding var_def
by (auto intro: FIN)*)
have [simp]: "\<And>u p V. dom (\<lambda>v. if (u, v) \<in> E \<and> v \<notin> V then Some (u # p)
else None) = E``{u} - V"
by (auto split: if_split_asm)
{
fix V ws u p
assume INV: "inv (None,V,ws)"
assume WSU: "ws u = Some p"
from INV WSU have
[simp]: "V \<subseteq> E\<^sup>*``U0"
and [simp]: "u \<in> V"
and UREACH: "\<exists>u0\<in>U0. (u0,u)\<in>E\<^sup>*"
and [simp]: "finite (dom ws)"
unfolding inv_def
apply simp_all
apply auto []
apply clarsimp
apply blast
done
have "(V \<union> E `` {u}, V) \<in> finite_psupset (E\<^sup>* `` U0) \<or>
V \<union> E `` {u} = V \<and>
card (E `` {u} - V \<union> (dom ws - {u})) < card (dom ws)"
proof (subst disj_commute, intro disjCI conjI)
assume "(V \<union> E `` {u}, V) \<notin> finite_psupset (E\<^sup>* `` U0)"
thus "V \<union> E `` {u} = V" using UREACH
by (auto simp: finite_psupset_def intro: rev_ImageI)
hence [simp]: "E``{u} - V = {}" by force
show "card (E `` {u} - V \<union> (dom ws - {u})) < card (dom ws)"
using WSU
by (auto intro: card_Diff1_less)
qed
} note wf_aux=this
{
fix V ws u p
assume FIN: "finite (E\<^sup>*``U0)"
assume "inv (None,V,ws)" "ws u = Some p"
then obtain u0 where "u0\<in>U0" "(u0,u)\<in>E\<^sup>*" unfolding inv_def
by clarsimp blast
hence "E``{u} \<subseteq> E\<^sup>*``U0" by (auto intro: rev_ImageI)
hence "finite (E``{u})" using FIN(1) by (rule finite_subset)
} note succs_finite=this
{
fix V ws u p
assume FIN: "finite (E\<^sup>*``U0)"
assume INV: "inv (None,V,ws)"
assume WSU: "ws u = Some p"
assume NVD: "\<not> P u"
have "inv (None, V \<union> E `` {u},
ws |` (- {u}) ++
(\<lambda>v. if (u, v) \<in> E \<and> v \<notin> V then Some (u # p)
else None))"
unfolding inv_def
apply (simp, intro conjI)
using INV WSU apply (auto simp: inv_def) []
using INV WSU apply (auto simp: inv_def) []
using INV WSU apply (auto simp: succs_finite FIN) []
using INV apply (auto simp: inv_def) []
using INV apply (auto simp: inv_def) []
using INV WSU apply (auto
simp: inv_def
intro: rtrancl_image_advance
) []
using INV WSU apply (auto simp: inv_def) []
using INV NVD apply (auto simp: inv_def) []
using INV NVD apply (auto simp: inv_def) []
using INV WSU NVD apply (fastforce
simp: inv_def restrict_map_def
intro!: path_conc path1
split: if_split_asm
) []
done
} note ip_aux=this
show ?thesis
unfolding wset_find_path_def find_path_def ws_update_def s_init_def
apply (refine_rcg refine_vcg le_ASSERTI
WHILET_rule[where
R = var and I = inv]
)
using [[goals_limit = 1]]
apply (auto simp: var_def) []
apply (auto
simp: inv_def dom_def
split: if_split_asm) []
apply simp
apply (auto simp: inv_def) []
apply (auto simp: var_def brk_rel_def) []
apply (simp add: succs_finite)
apply (auto simp: inv_def) []
apply clarsimp
apply (simp add: ip_aux)
apply clarsimp
apply (simp add: var_def brk_rel_def wf_aux) []
apply (fastforce
simp: inv_def
split: option.splits
intro: rev_ImageI
dest: Image_closed_trancl) []
done
qed
text \<open>We refine the algorithm to use a foreach-loop\<close>
definition "ws_update_foreach E u p V ws \<equiv>
FOREACH (LIST_SET_REV_TAG (E``{u})) (\<lambda>v (V,ws).
if v\<in>V then
RETURN (V,ws)
else do {
ASSERT (v\<notin>dom ws);
RETURN (insert v V,ws( v \<mapsto> u#p))
}
) (V,ws)"
lemma ws_update_foreach_refine[refine]:
assumes FIN: "finite (E``{u})"
assumes WSS: "dom ws \<subseteq> V"
assumes ID: "(E',E)\<in>Id" "(u',u)\<in>Id" "(p',p)\<in>Id" "(V',V)\<in>Id" "(ws',ws)\<in>Id"
shows "ws_update_foreach E' u' p' V' ws' \<le> \<Down>Id (ws_update E u p V ws)"
unfolding ID[simplified]
unfolding ws_update_foreach_def ws_update_def LIST_SET_REV_TAG_def
apply (refine_rcg refine_vcg FIN
FOREACH_rule[where I="\<lambda>it (V',ws').
V'=V \<union> (E``{u}-it)
\<and> dom ws' \<subseteq> V'
\<and> ws' = ws ++ (\<lambda>v. if (u,v)\<in>E \<and> v\<notin>it \<and> v\<notin>V then Some (u#p) else None)"]
)
using WSS
apply (auto
simp: Map.map_add_def
split: option.splits if_split_asm
intro!: ext[where 'a='a and 'b="'b list option"])
- apply fastforce+
done
definition "s_init_foreach U0 \<equiv> do {
(U0,ws) \<leftarrow> FOREACH U0 (\<lambda>x (U0,ws).
RETURN (insert x U0,ws(x\<mapsto>[]))) ({},Map.empty);
RETURN (None,U0,ws)
}"
lemma s_init_foreach_refine[refine]:
assumes FIN: "finite U0"
assumes ID: "(U0',U0)\<in>Id"
shows "s_init_foreach U0' \<le>\<Down>Id (s_init U0)"
unfolding s_init_foreach_def s_init_def ID[simplified]
apply (refine_rcg refine_vcg
FOREACH_rule[where
I = "\<lambda>it (U,ws).
U = U0-it
\<and> ws = (\<lambda>x. if x\<in>U0-it then Some [] else None)"]
)
apply (auto
simp: FIN
intro!: ext
)
done
definition "wset_find_path' E U0 P \<equiv> do {
ASSERT (finite U0);
s0\<leftarrow>s_init_foreach U0;
(res,_,_) \<leftarrow> WHILET
(\<lambda>(res,V,ws). res=None \<and> ws\<noteq>Map.empty)
(\<lambda>(res,V,ws). do {
ASSERT (ws\<noteq>Map.empty);
((u,p),ws) \<leftarrow> op_map_pick_remove ws;
if P u then
RETURN (Some (rev p,u),V,ws)
else do {
(V,ws) \<leftarrow> ws_update_foreach E u p V ws;
RETURN (None,V,ws)
}
})
s0;
RETURN res
}"
lemma wset_find_path'_refine:
"wset_find_path' E U0 P \<le> \<Down>Id (wset_find_path E U0 P)"
unfolding wset_find_path'_def wset_find_path_def
unfolding op_map_pick_remove_alt
apply (refine_rcg IdI)
apply assumption
apply simp_all
done
section \<open>Refinement to efficient data structures\<close>
schematic_goal wset_find_path'_refine_aux:
fixes U0::"'a set" and P::"'a \<Rightarrow> bool" and E::"('a\<times>'a) set"
and Pimpl :: "'ai \<Rightarrow> bool"
and node_rel :: "('ai \<times> 'a) set"
and node_eq_impl :: "'ai \<Rightarrow> 'ai \<Rightarrow> bool"
and node_hash_impl
and node_def_hash_size
assumes [autoref_rules]:
"(succi,E)\<in>\<langle>node_rel\<rangle>slg_rel"
"(Pimpl,P)\<in>node_rel \<rightarrow> bool_rel"
"(node_eq_impl, (=)) \<in> node_rel \<rightarrow> node_rel \<rightarrow> bool_rel"
"(U0',U0)\<in>\<langle>node_rel\<rangle>list_set_rel"
assumes [autoref_ga_rules]:
"is_bounded_hashcode node_rel node_eq_impl node_hash_impl"
"is_valid_def_hm_size TYPE('ai) node_def_hash_size"
notes [autoref_tyrel] =
TYRELI[where
R="\<langle>node_rel,\<langle>node_rel\<rangle>list_rel\<rangle>list_map_rel"]
TYRELI[where R="\<langle>node_rel\<rangle>map2set_rel (ahm_rel node_hash_impl)"]
(*notes [autoref_rules] =
IdI[of P, unfolded fun_rel_id_simp[symmetric]]*)
shows "(?c::?'c,wset_find_path' E U0 P) \<in> ?R"
unfolding wset_find_path'_def ws_update_foreach_def s_init_foreach_def
using [[autoref_trace_failed_id]]
using [[autoref_trace_intf_unif]]
using [[autoref_trace_pat]]
apply (autoref (keep_goal))
done
find_theorems list_map_update
concrete_definition wset_find_path_impl for node_eq_impl succi U0' Pimpl
uses wset_find_path'_refine_aux
section \<open>Autoref Setup\<close>
context begin interpretation autoref_syn .
lemma [autoref_itype]:
"find_path ::\<^sub>i \<langle>I\<rangle>\<^sub>ii_slg \<rightarrow>\<^sub>i \<langle>I\<rangle>\<^sub>ii_set \<rightarrow>\<^sub>i (I\<rightarrow>\<^sub>ii_bool)
\<rightarrow>\<^sub>i \<langle>\<langle>\<langle>\<langle>I\<rangle>\<^sub>ii_list, I\<rangle>\<^sub>ii_prod\<rangle>\<^sub>ii_option\<rangle>\<^sub>ii_nres" by simp
lemma wset_find_path_autoref[autoref_rules]:
fixes node_rel :: "('ai \<times> 'a) set"
assumes eq: "GEN_OP node_eq_impl (=) (node_rel\<rightarrow>node_rel\<rightarrow>bool_rel)"
assumes hash: "SIDE_GEN_ALGO (is_bounded_hashcode node_rel node_eq_impl node_hash_impl)"
assumes hash_dsz: "SIDE_GEN_ALGO (is_valid_def_hm_size TYPE('ai) node_def_hash_size)"
shows "(
wset_find_path_impl node_hash_impl node_def_hash_size node_eq_impl,
find_path)
\<in> \<langle>node_rel\<rangle>slg_rel \<rightarrow> \<langle>node_rel\<rangle>list_set_rel \<rightarrow> (node_rel\<rightarrow>bool_rel)
\<rightarrow> \<langle>\<langle>\<langle>node_rel\<rangle>list_rel\<times>\<^sub>rnode_rel\<rangle>option_rel\<rangle>nres_rel"
proof -
note EQI = GEN_OP_D[OF eq]
note HASHI = SIDE_GEN_ALGO_D[OF hash]
note DSZI = SIDE_GEN_ALGO_D[OF hash_dsz]
note wset_find_path_impl.refine[THEN nres_relD, OF _ _ EQI _ HASHI DSZI]
also note wset_find_path'_refine
also note wset_find_path_correct
finally show ?thesis
by (fastforce intro!: nres_relI)
qed
end
schematic_goal wset_find_path_transfer_aux:
"RETURN ?c \<le> wset_find_path_impl hashi dszi eqi E U0 P"
unfolding wset_find_path_impl_def
by (refine_transfer (post))
concrete_definition wset_find_path_code
for E ?U0.0 P uses wset_find_path_transfer_aux
lemmas [refine_transfer] = wset_find_path_code.refine
export_code wset_find_path_code checking SML
section \<open>Nontrivial paths\<close>
definition "find_path1_gen E u0 P \<equiv> do {
res \<leftarrow> find_path E (E``{u0}) P;
case res of None \<Rightarrow> RETURN None
| Some (p,v) \<Rightarrow> RETURN (Some (u0#p,v))
}"
lemma find_path1_gen_correct: "find_path1_gen E u0 P \<le> find_path1 E u0 P"
unfolding find_path1_gen_def find_path_def find_path1_def
apply (refine_rcg refine_vcg le_ASSERTI)
apply (auto
intro: path_prepend
dest: tranclD
elim: finite_subset[rotated]
)
done
schematic_goal find_path1_impl_aux:
fixes node_rel :: "('ai \<times> 'a) set"
assumes [autoref_rules]: "(node_eq_impl, (=)) \<in> node_rel \<rightarrow> node_rel \<rightarrow> bool_rel"
assumes [autoref_ga_rules]:
"is_bounded_hashcode node_rel node_eq_impl node_hash_impl"
"is_valid_def_hm_size TYPE('ai) node_def_hash_size"
shows "(?c,find_path1_gen::(_\<times>_) set \<Rightarrow> _) \<in> \<langle>node_rel\<rangle>slg_rel \<rightarrow> node_rel \<rightarrow> (node_rel \<rightarrow> bool_rel) \<rightarrow> \<langle>\<langle>\<langle>node_rel\<rangle>list_rel \<times>\<^sub>r node_rel\<rangle>option_rel\<rangle>nres_rel"
unfolding find_path1_gen_def[abs_def]
apply (autoref (trace,keep_goal))
done
lemma [autoref_itype]:
"find_path1 ::\<^sub>i \<langle>I\<rangle>\<^sub>ii_slg \<rightarrow>\<^sub>i I \<rightarrow>\<^sub>i (I\<rightarrow>\<^sub>ii_bool)
\<rightarrow>\<^sub>i \<langle>\<langle>\<langle>\<langle>I\<rangle>\<^sub>ii_list, I\<rangle>\<^sub>ii_prod\<rangle>\<^sub>ii_option\<rangle>\<^sub>ii_nres" by simp
concrete_definition find_path1_impl uses find_path1_impl_aux
lemma find_path1_autoref[autoref_rules]:
fixes node_rel :: "('ai \<times> 'a) set"
assumes eq: "GEN_OP node_eq_impl (=) (node_rel\<rightarrow>node_rel\<rightarrow>bool_rel)"
assumes hash: "SIDE_GEN_ALGO (is_bounded_hashcode node_rel node_eq_impl node_hash_impl)"
assumes hash_dsz: "SIDE_GEN_ALGO (is_valid_def_hm_size TYPE('ai) node_def_hash_size)"
shows "(find_path1_impl node_eq_impl node_hash_impl node_def_hash_size,find_path1)
\<in> \<langle>node_rel\<rangle>slg_rel \<rightarrow>node_rel \<rightarrow> (node_rel \<rightarrow> bool_rel) \<rightarrow>
\<langle>\<langle>\<langle>node_rel\<rangle>list_rel \<times>\<^sub>r node_rel\<rangle>Relators.option_rel\<rangle>nres_rel"
proof -
note EQI = GEN_OP_D[OF eq]
note HASHI = SIDE_GEN_ALGO_D[OF hash]
note DSZI = SIDE_GEN_ALGO_D[OF hash_dsz]
note R = find_path1_impl.refine[param_fo, THEN nres_relD, OF EQI HASHI DSZI]
note R
also note find_path1_gen_correct
finally show ?thesis by (blast intro: nres_relI)
qed
schematic_goal find_path1_transfer_aux:
"RETURN ?c \<le> find_path1_impl eqi hashi dszi E u P"
unfolding find_path1_impl_def
by refine_transfer
concrete_definition find_path1_code for E u P uses find_path1_transfer_aux
lemmas [refine_transfer] = find_path1_code.refine
end
diff --git a/thys/Generalized_Counting_Sort/Algorithm.thy b/thys/Generalized_Counting_Sort/Algorithm.thy
--- a/thys/Generalized_Counting_Sort/Algorithm.thy
+++ b/thys/Generalized_Counting_Sort/Algorithm.thy
@@ -1,2159 +1,2163 @@
(* Title: An Efficient Generalization of Counting Sort for Large, possibly Infinite Key Ranges
Author: Pasquale Noce
Software Engineer at HID Global, Italy
pasquale dot noce dot lavoro at gmail dot com
pasquale dot noce at hidglobal dot com
*)
section "Algorithm's description, analysis, and formalization"
theory Algorithm
imports Main
begin
text \<open>
\null
\emph{This paper is dedicated to Gaia, my sweet niece, whose arrival has blessed me and my family
with joy and tenderness.}
\null
\emph{Moreover, I would like to thank my colleague Iacopo Rippa, who patiently listened to the ideas
underlying sections \ref{SEC1} and \ref{SEC3}, and helped me expand those ideas into complete proofs
by providing me with valuable hints and test data.}
\<close>
subsection "Introduction"
subsubsection "Counting sort"
text \<open>
\emph{Counting sort} is a well-known algorithm that sorts a collection of objects of any kind, as
long as each such object is associated with a signed integer key, according to their respective keys
(cf. \cite{R1}, \cite{R2}). If $xs$ is the input array containing $n$ objects to be sorted, $out$ is
the output, sorted array, and $key$ is the function mapping objects to keys, counting sort works as
follows (assuming arrays to be zero-based):
\begin{enumerate}
\item
Search the minimum key $mi$ and the maximum key $ma$ occurring within $xs$ (which can be done via a
single loop over $xs$).
\item
Allocate an array $ns$ of $ma - mi + 2$ unsigned integers and initialize all its elements to 0.
\item
For each $i$ from 0 to $n-1$, increase $ns[key(xs[i]) - mi + 1]$ by 1.
\item
For each $i$ from 2 to $ma - mi$, increase $ns[i]$ by $ns[i - 1]$.
\item
For each $i$ from 0 to $n-1$, set $out[ns[key(xs[i]) - mi]]$ to $xs[i]$ and increase
$ns[key(xs[i]) - mi]$ by 1.
\end{enumerate}
Steps 1 and 2 take $O(n)$ and $O(ma - mi)$ time, respectively. Step 3 counts how many times each
possible key occurs within $xs$, and takes $O(n)$ time. Step 4 computes the offset within $out$ of
the first object in $xs$, if any, having each possible key, and takes $O(ma - mi)$ time. Finally,
step 5 fills $out$, taking $O(n)$ time. Thus, the overall running time is $O(n) + O(ma - mi)$, and
the same is obviously true of memory space.
If the range of all the keys possibly occurring within $xs$, henceforth briefly referred to as the
\emph{key range}, is known in advance, the first two steps can be skipped by using the minimum and
maximum keys in the key range as $mi$, $ma$ and pre-allocating (possibly statically, rather than
dynamically, in real-world implementations) an array $ns$ of size $ma - mi + 2$. However, this does
not affect the asymptotic running time and memory space required by the algorithm, since both keep
being $O(n) + O(ma - mi)$ independently of the distribution of the keys actually occurring in $xs$
within the key range.
As a result, counting sort is suitable for direct use, viz. not just as a subroutine of another
sorting algorithm such as radix sort, only if the key range is not significantly larger than $n$.
Indeed, if 100 objects with 100,000 possible keys have to be sorted, accomplishing this task by
allocating, and iterating over, an array of 100,000 unsigned integers to count keys' occurrences
would be quite impractical! Whence the question that this paper will try to answer: how can counting
sort be generalized for direct use in case of a large key range?
Solving this problem clearly requires to renounce having one counter per key, rather using a bounded
number of counters, independent of the key range's cardinality, and partitioning the key range into
some number of intervals compatible with the upper bound on the counters' number. The resulting key
intervals will then form as many \emph{buckets}, and what will have to be counted is the number of
the objects contained in each bucket.
Counting objects per bucket, rather than per single key, has the following major consequences, the
former good, the latter bad:
\begin{itemize}
\item
\emph{Keys are no longer constrained to be integers, but may rather be elements of any linear order,
even of infinite cardinality.}
\\In fact, in counting sort keys must be integers -- or anything else in one-to-one correspondence
with some subset of the integers, such as alphabet letters -- since this ensures that the key range
contains finitely many keys, so that finitely many counters are needed. Thus, the introduction of an
upper bound for the number of counters makes this constraint vanish. As a result, keys of any kind
are now allowed and the key range can even be infinite (mathematically, since any representation of
the key range on a computer will always be finite). Notably, rational or real numbers may be used as
keys, too.
\\This observation considerably extends the scope of application of the special case where function
$key$ matches the identity function. In counting sort, this option is viable only if the objects to
be sorted are themselves integers, whereas in the generalized algorithm it is viable whenever they
are elements of any linear order, which also happens if they are rational or real numbers.
\item
\emph{Recursion needs to be introduced, since any bucket containing more than one object is in turn
required to be sorted.}
\\In fact, nothing prevents multiple objects from falling in the same bucket, and while this happens
sorting is not accomplished. Therefore, the generalized algorithm must provide for recursive rounds,
where each round splits any bucket containing multiple objects into finer-grained buckets containing
fewer objects. Recursion will then go on until every bucket contains at most one object, viz. until
there remains no counter larger than one.
\end{itemize}
Of course, the fewer recursive rounds are required to complete sorting, the more the algorithm will
be efficient, whence the following, fundamental question: how to minimize the number of the rounds?
That is to say, how to maximize the probability that, as a result of the execution of a round, there
be at most one object in each bucket, so that no more rounds be required?
The intuitive answer is: first, by making the buckets equiprobable -- or at least, by making their
probabilities as much uniform as possible --, and second, by increasing the number of the buckets as
much as possible. Providing pen-and-paper proofs of both of these statements, and showing how they
can be enforced, is the purpose of the following sections.
\<close>
subsubsection "Buckets' probability -- Proof"
text \<open>
\label{SEC1}
Suppose that $k$ objects be split randomly among $n$ equiprobable buckets, where $k \leq n$. This
operation is equivalent to selecting at random a sequence of $k$ buckets, possibly with repetitions,
so that the first object be placed into the first bucket of the sequence, the second object into the
second bucket, and so on. Thus, the probability $P$ that each bucket will contain at most one object
-- which will be called \emph{event E} in what follows -- is equal to the probability of selecting a
sequence without repetitions among all the possible sequences of $k$ buckets formed with the $n$
given ones.
Since buckets are assumed to be equiprobable, so are all such sequences. Hence, $P$ is equal to the
ratio of the number of the sequences without repetitions to the number of all sequences, namely:
\begin{equation}
\label{EQ1}
P=\frac{n!}{(n-k)!n^k}
\end{equation}
In the special case where $k = n$, this equation takes the following, simpler form:
\begin{equation}
\label{EQ2}
P=\frac{n!}{n^n}
\end{equation}
Now, suppose that the $n$ buckets be no longer equiprobable, viz. that they no longer have the same,
uniform probability $1/n$, rather having arbitrary, nonuniform probabilities $p_1$, ..., $p_n$. The
equation for probability $P$ applying to this case can be obtained through an iterative procedure,
as follows.
Let $i$ be an index in the range 1 to $n$ such that $p_i$ is larger than $1/n$. After swapping index
$i$ for 1, let $x_1$ be the increment in probability $p_1$ with respect to $1/n$, so that $p_1 =
a_0/n + x_1$ with $a_0 = 1$ and $0 < x_1 \leq a_0(n-1)/n$ (as $p_1 = 1$ for $x_1 = a_0(n-1)/n$).
Then, let $P_1$ be the probability of event $E$ in case the first bucket has probability $p_1$ and
all the other $n-1$ buckets have the same, uniform probability $q_1 = a_0/n - x_1/(n-1)$.
If $k < n$, event $E$ occurs just in case either all $k$ objects fall in as many distinct buckets
with probability $q_1$, or $k - 1$ objects do so whereas the remaining object falls in the bucket
with probability $p_1$. As these events, say $E_A$ and $E_B$, are incompatible, $P_1$ matches the
sum of their respective probabilities.
Since all the possible choices of $k$ distinct buckets are mutually incompatible, while those of the
buckets containing any two distinct objects are mutually independent, the probability of event $E_A$
is equal to the product of the following factors:
\begin{itemize}
\item
The number of the sequences without repetitions of $k$ buckets formed with the $n-1$ ones with
probability $q_1$, i.e. $(n-1)!/(n-1-k)! = (n-k)(n-1)!/(n-k)!$.
\item
The probability of any such sequence, i.e. $q_1^k$.
\end{itemize}
By virtue of similar considerations, the probability of event $E_B$ turns out to match the product
of the following factors:
\begin{itemize}
\item
The number of the sequences without repetitions of $k - 1$ buckets formed with the $n-1$ ones with
probability $q_1$, i.e. $(n-1)!/(n-\bcancel{1}-k+\bcancel{1})! = (n-1)!/(n-k)!$.
\item
The probability of any such sequence, i.e. $q_1^{k-1}$.
\item
The number of the possible choices of the object falling in the first bucket, i.e. $k$.
\item
The probability of the first bucket, i.e. $p_1$.
\end{itemize}
Therefore, $P_1$ is provided by the following equation:
\begin{equation}
\label{EQ3}
\begin{split}
P_1&=\frac{(n-k)(n-1)!}{(n-k)!}\left(\frac{a_0}{n} - \frac{x_1}{n-1}\right)^k\\
&\quad+k\frac{(n-1)!}{(n-k)!}\left(\frac{a_0}{n} - \frac{x_1}{n-1}\right)^{k-1}
\left(\frac{a_0}{n} + x_1\right)
\end{split}
\end{equation}
The correctness of this equation is confirmed by the fact that its right-hand side matches that of
equation \eqref{EQ1} for $x_1 = 0$, since $P_1$ must degenerate to $P$ in this case. In fact, being
$a_0 = 1$, it results:
\begin{align*}
&\frac{(n-k)(n-1)!}{(n-k)!}\left(\frac{a_0}{n} - 0\right)^k
+k\frac{(n-1)!}{(n-k)!}\left(\frac{a_0}{n} - 0\right)^{k-1}\left(\frac{a_0}{n} + 0\right)\\
&\quad=(n-\bcancel{k}+\bcancel{k})\frac{(n-1)!}{(n-k)!}\left(\frac{a_0}{n}\right)^k\\
&\quad=\frac{n!}{(n-k)!n^k}
\end{align*}
If $k = n$, event $E_A$ is impossible, as there is no way to accommodate $n$ objects within $n-1$
buckets without repetitions. Thus, $P_1$ is given by the following equation, derived by deleting the
first addend and replacing $k$ with $n$ in the right-hand side of equation \eqref{EQ3}:
\begin{equation}
\label{EQ4}
P_1=n!\left(\frac{a_0}{n} - \frac{x_1}{n-1}\right)^{n-1}\left(\frac{a_0}{n} + x_1\right)
\end{equation}
Likewise, the right-hand side of this equation matches that of equation \eqref{EQ2} for $x_1 = 0$,
which confirms its correctness.
The conclusions reached so far can be given a concise form, suitable for generalization, through the
following definitions, where $i$ and $j$ are any two natural numbers such that $0 < k-j \leq n-i$
and $a_i$ is some assigned real number:
\begin{align*}
A_{i,j}&\equiv\frac{(n-i)!}{(n-i-k+j)!}\left(\frac{a_i}{n-i}\right)^{k-j}\\
F_{i,j}&\equiv(k-j+1)A_{i,j}p_i\\
G_{i,j}&\equiv A_{i,j-1}+F_{i,j}
\end{align*}
Then, denoting the value of $P$ in the uniform probability case with $P_0$, and $a_0(n-1)/n - x_1$
with $a_1$, so that $q_1 = a_1/(n-1)$, equations \eqref{EQ1}, \eqref{EQ3}, and \eqref{EQ4} can be
rewritten as follows:
\begin{align}
\label{EQ5}
P_0&=A_{0,0}\\
\label{EQ6}
P_1&=
\begin{cases}
G_{1,1}=A_{1,0}+kA_{1,1}p_1&\quad\text{if $k < n$,}\\
F_{1,1}=kA_{1,1}p_1&\quad\text{if $k = n$.}
\end{cases}
\end{align}
Even more than for their conciseness, these equations are significant insofar as they show that the
right-hand side of equation \eqref{EQ6} can be obtained from the one of equation \eqref{EQ5} by
replacing $A_{0,0}$ with either $G_{1,1}$ or $F_{1,1}$, depending on whether $k < n$ or $k = n$.
If $p_i$ matches $q_1$ for any $i$ in the range 2 to $n$, $P = P_1$, thus $P$ is given by equation
\eqref{EQ6}. Otherwise, the procedure that has led to equation \eqref{EQ6} can be applied again. For
some index $i$ in the range 2 to $n$ such that $p_i$ is larger than $q_1$, swap $i$ for 2, and let
$x_2 = p_2 - a_1/(n-1)$, $a_2 = a_1(n-2)/(n-1) - x_2$, with $0 < x_2 \leq a_1(n-2)/(n-1)$. Moreover,
let $P_2$ be the probability of event $E$ if the first two buckets have probabilities $p_1$, $p_2$
and the other $n-2$ buckets have the same probability $q_2 = a_2/(n-2)$.
Then, reasoning as before, it turns out that the equation for $P_2$ can be obtained from equation
\eqref{EQ6} by replacing:
\begin{itemize}
\item
$A_{1,0}$ with $G_{2,1}$ or $F_{2,1}$, depending on whether $k < n-1$ or $k = n-1$, and
\item
$A_{1,1}$ with $G_{2,2}$ or $F_{2,2}$, depending on whether $k-1 < n-1$, i.e. $k < n$, or $k-1 =
n-1$, i.e. $k = n$.
\end{itemize}
As a result, $P_2$ is provided by the following equation:
\begin{equation}
\label{EQ7}
P_2=
\begin{cases}
G_{2,1}+kG_{2,2}p_1=A_{2,0}+kA_{2,1}p_2+k[A_{2,1}+(k-1)A_{2,2}p_2]p_1\\
\quad\text{if $k < n-1$,}\\
F_{2,1}+kG_{2,2}p_1=kA_{2,1}p_2+k[A_{2,1}+(k-1)A_{2,2}p_2]p_1\\
\quad\text{if $k = n-1$,}\\
kF_{2,2}p_1=k(k-1)A_{2,2}p_2p_1\\
\quad\text{if $k = n$.}
\end{cases}
\end{equation}
Since the iterative procedure used to derive equations \eqref{EQ6} and \eqref{EQ7} can be further
applied as many times as required, it follows that for any nonuniform probability distribution
$p_1$, ..., $p_n$, the equation for $P$ can be obtained from equation \eqref{EQ5} with $n-1$ steps
at most, where each step consists of replacing terms of the form $A_{i,j}$ with terms of either form
$G_{i+1,j+1}$ or $F_{i+1,j+1}$, depending on whether $k-j < n-i$ or $k-j = n-i$.
Let us re-use letters $n$, $k$ in lieu of $n-i$ and $k-j$, and use letters $a$, $x$ as aliases for
$a_i$ and $x_{i+1}$. Then, any aforesaid replacement is equivalent to the insertion of either of the
following expressions, regarded as images of as many functions $G$, $F$ of real variable $x$:
\begin{align}
\nonumber
G(x)&=\frac{(n-k)(n-1)!}{(n-k)!}\left(\frac{a}{n} - \frac{x}{n-1}\right)^k\\
\label{EQ8}
&\quad+k\frac{(n-1)!}{(n-k)!}\left(\frac{a}{n} - \frac{x}{n-1}\right)^{k-1}
\left(\frac{a}{n} + x\right)
&& \text{for $k < n$,}\\
\label{EQ9}
F(x)&=n!\left(\frac{a}{n} - \frac{x}{n-1}\right)^{n-1}\left(\frac{a}{n} + x\right)
&& \text{for $k = n$}
\end{align}
in place of the following expression:
\begin{equation}
\label{EQ10}
\frac{n!}{(n-k)!}\left(\frac{a}{n}\right)^k=
\begin{cases}
G(0)&\quad\text{if $k < n$,}\\
F(0)&\quad\text{if $k = n$.}
\end{cases}
\end{equation}
Equation \eqref{EQ10} can be obtained from equations \eqref{EQ8} and \eqref{EQ9} in the same way as
equations \eqref{EQ3} and \eqref{EQ4} have previously been shown to match equations \eqref{EQ1} and
\eqref{EQ2} for $x_1 = 0$.
Since every such replacement takes place within a sum of nonnegative terms, $P$ can be proven to be
increasingly less than $P_0$ for increasingly nonuniform probability distributions -- which implies
that the probability of event $E$ is maximum in case of equiprobable buckets -- by proving that
functions $G$ and $F$ are strictly decreasing in $[0, b]$, where $b = a(n-1)/n$.
The slopes of the segments joining points $(0, G(0))$, $(b, G(b))$ and $(0, F(0))$, $(b, F(b))$ are:
\begin{align*}
\frac{G(b)-G(0)}{b-0}&=\frac{0-\dfrac{n!}{(n-k)!}\left(\dfrac{a}{n}\right)^k}{b}<0\text{,}\\
\frac{F(b)-F(0)}{b-0}&=\frac{0-n!\left(\dfrac{a}{n}\right)^n}{b}<0\text{.}
\end{align*}
Therefore, by Lagrange's mean value theorem, there exist $c, d \in (0, b)$ such that $G'(c) < 0$ and
$F'(d) < 0$. On the other hand, it is:
\begin{align*}
G'(x)&=-k\frac{(n-1)!}{(n-k)!}\frac{n-k}{n-1}\left(\frac{a}{n}-\frac{x}{n-1}\right)^{k-1}\\
&\quad-k\frac{(n-1)!}{(n-k)!}\frac{k-1}{n-1}\left(\frac{a}{n}-\frac{x}{n-1}\right)^{k-2}
\left(\frac{a}{n}+x\right)\\
&\quad+k\frac{(n-1)!}{(n-k)!}\left(\frac{a}{n}-\frac{x}{n-1}\right)^{k-1}\text{,}\\
F'(x)&=-n!\left(\frac{a}{n}-\frac{x}{n-1}\right)^{n-2}
\left(\frac{a}{n}+x\right)
+n!\left(\frac{a}{n}-\frac{x}{n-1}\right)^{n-1}\text{.}
\end{align*}
Thus, solving equations $G'(x) = 0$ and $F'(x) = 0$ for $x \neq b$, viz. for $a/n - x/(n-1) \neq 0$,
it results:
\begin{align*}
&G'(x)=0\\
&\quad\Rightarrow\bcancel{k\frac{(n-1)!}{(n-k)!}\left(\frac{a}{n}-\frac{x}{n-1}\right)^{k-2}}
\biggl[\frac{k-n}{n-1}\left(\frac{a}{n}-\frac{x}{n-1}\right)\\
&\quad\quad\quad+\frac{1-k}{n-1}\left(\frac{a}{n}+x\right)+\frac{a}{n}-\frac{x}{n-1}\biggr]=0\\
&\quad\Rightarrow\frac{1}{\bcancel{n(n-1)^2}}\bigl\{(k-n)[(n-1)a-nx]+(1-k)[(n-1)a+n(n-1)x]\\
&\quad\quad\quad+(n-1)^2a-n(n-1)x\bigr\}=0\\
&\quad\Rightarrow\bcancel{k(n-1)a}-knx-n(n-1)a+n^2x+(n-1)a+\bcancel{n(n-1)x}\\
&\quad\quad\quad-\bcancel{k(n-1)a}-kn(n-1)x+(n-1)^2a-\bcancel{n(n-1)x}=0\\
&\quad\Rightarrow-\bcancel{knx}-\bcancel{n^2a}+\bcancel{na}+n^2x+\bcancel{na}-\bcancel{a}
-kn^2x+\bcancel{knx}+\bcancel{n^2a}-\bcancel{2na}+\bcancel{a}=0\\
&\quad\Rightarrow\bcancel{n^2(1-k)}x=0\\
&\quad\Rightarrow x=0\text{,}
\end{align*}
\begin{align*}
&F'(x)=0\\
&\quad\Rightarrow\bcancel{n!\left(\frac{a}{n}-\frac{x}{n-1}\right)^{n-2}}
\left(-\bcancel{\frac{a}{n}}-x+\bcancel{\frac{a}{n}}-\frac{x}{n-1}\right)=0\\
&\quad\Rightarrow\bcancel{\frac{n}{1-n}}x=0\\
&\quad\Rightarrow x=0\text{.}
\end{align*}
Hence, there is no $x \in (0, b)$ such that $G'(x) = 0$ or $F'(x) = 0$. Moreover, if there existed
$y, z \in (0, b)$ such that $G'(y) > 0$ or $F'(z) > 0$, by Bolzano's theorem there would also exist
$u, v$ in the open intervals with endpoints $c, y$ and $d, z$, both included in $(0, b)$, such that
$G'(u) = 0$ or $F'(v) = 0$, which is not the case. Therefore, $G'(x)$ and $F'(x)$ are negative for
any $x \in (0, b)$, so that functions $G$ and $F$ are strictly decreasing in $[0, b]$, Q.E.D..
\<close>
subsubsection "Buckets' probability -- Implementation"
text \<open>
\label{SEC2}
Given $n > 1$ buckets, numbered with indices 0 to $n - 1$, and a finite set $A$ of objects having
minimum key $mi$ and maximum key $ma$, let $E(k)$, $I(mi,ma)$ be the following events, defined as
subsets of the whole range $R$ of function $key$, with $k$ varying over $R$:
\begin{align*}
E(k)&\equiv\{k'\in R.\;k'\leq k\}\\
I(mi,ma)&\equiv\{k'\in R.\;mi\leq k'\leq ma\}
\end{align*}
Furthermore, define functions $r$, $f$ as follows:
\begin{align*}
r(k,n,mi,ma)&\equiv (n-1)\cdot P(E(k)\;|\;I(mi,ma))\\
f(k,n,mi,ma)&\equiv floor(r(k,n,mi,ma))
\end{align*}
where $P(E(k)\;|\;I(mi,ma))$ denotes the conditional probability of event $E(k)$, viz. for a key not
to be larger than $k$, given event $I(mi,ma)$, viz. if the key is comprised between $mi$ and $ma$.
Then, the buckets' probabilities can be made as much uniform as possible by placing each object $x
\in A$ into the bucket whose index matches the following value:
\begin{equation*}
index(key,x,n,mi,ma)\equiv f(key(x),n,mi,ma)
\end{equation*}
For example, given $n = 5$ buckets, suppose that the image of set $A$ under function $key$ consists
of keys $k_1 = mi$, $k_2$, ..., $k_9 = ma$, where the conditional probabilities for a key comprised
between $k_1$ and $k_9$ to match each of these keys have the following values:
\begin{align*}
P_1&=0.05,\\
P_2&=0.05,\\
P_3&=0.15,\\
P_4&=0.075,\\
P_5&=0.2,\\
P_6&=0.025,\\
P_7&=0.1,\\
P_8&=0.25,\\
P_9&=0.1
\end{align*}
Evidently, there is no way of partitioning set $\{k_1, ..., k_9\}$ into five equiprobable subsets
comprised of contiguous keys. However, it results:
\begin{equation*}
floor\left(4\cdot\sum_{i=1}^{n}P_i\right)=
\begin{cases}
0&\quad\text{for $n = 1, 2$,}\\
1&\quad\text{for $n = 3, 4$,}\\
2&\quad\text{for $n = 5, 6, 7$,}\\
3&\quad\text{for $n = 8$,}\\
4&\quad\text{for $n = 9$.}
\end{cases}
\end{equation*}
Hence, in spite of the highly nonuniform distribution of the keys' probabilities -- key $k_8$'s
probability is 10 times that of key $k_6$ --, function $index$ manages all the same to split the
objects in $A$ so as to make the buckets' probabilities more uniform -- with the maximum one being
about 3 times the minimum one --, as follows:
\begin{itemize}
\item
Bucket 0 has probability 0.1, as it collects the objects with keys $k_1$, $k_2$.
\item
Bucket 1 has probability 0.225, as it collects the objects with keys $k_3$, $k_4$.
\item
Bucket 2 has probability 0.325, as it collects the objects with keys $k_5$, $k_6$, $k_7$.
\item
Bucket 3 has probability 0.25, as it collects the objects with key $k_8$.
\item
Bucket 4 has probability 0.1, as it collects the objects with key $k_9$.
\end{itemize}
Remarkably, function $index$ makes the buckets' probabilities exactly or almost uniform -- meaning
that the maximum one is at most twice the minimum nonzero one, possibly except for the last bucket
alone -- in the following most common, even though special, cases:
\begin{enumerate}
\item
$I(mi,ma)$ is a finite set of equiprobable keys.
\item
$I(mi,ma)$ is a closed interval of real numbers, i.e. $I(mi,ma) = [mi,ma] \subset \mathbb{R}$, with
$P(\{mi\}) = 0$, and function $r$ is continuous for $k \in [mi,ma]$.
\end{enumerate}
In case 1, let $m$ be the cardinality of $I(mi,ma)$. It is $m > 0$ since $mi \in I(mi,ma)$, so that
each key in $I(mi,ma)$ has probability $1/m$.
If $m \leq n - 1$, then $(n - 1)/m \geq 1$, thus $f$ is nonzero and strictly increasing for $k \in
I(mi,ma)$. Thus, in this subcase function $index$ fills exactly $m$ buckets, one for each single key
in $I(mi,ma)$, whereas the remaining $n - m$ buckets, particularly the first one, are left unused.
Therefore, every used bucket has probability $1/m$.
If $m > n - 1$ and $m$ is divisible by $n - 1$, let $q > 1$ be the quotient of the division, so that
$m = q(n - 1)$. Dividing both sides of this equation by $m(n - 1)$, it turns out that $1/(n - 1) =
q/m$, and then $1/(n - 1) - 1/m = (q - 1)/m$. Hence, $f$ matches zero for the first $q - 1$ keys in
$I(mi,ma)$, increases by one for each of the $n - 2$ subsequent groups of $q$ contiguous keys, and
reaches value $n - 1$ in correspondence with the last key. Indeed, $q - 1 + q(n - 2) + 1 = q +
q(n - 2) = q(n - 1) = m$.
Consequently, in this subcase function $index$ places the objects mapped to the first $q - 1$ keys
into the first bucket -- which then has probability $(q - 1)/m$ --, the objects mapped to the $i$-th
subsequent group of $q$ keys, where $1 \leq i \leq n - 2$, into the bucket with index $i$ -- which
then has probability $q/m$ -- and the objects mapped to the last key into the last bucket -- which
then has probability $1/m$ --. Since $2(q - 1)/m = 2q/m - 2/m \geq 2q/m - q/m = q/m$, the maximum
probability is at most twice the minimum one, excluding the last bucket if $q > 2$.
If $m > n - 1$ and $m$ is not divisible by $n - 1$, let $q$, $r$ be the quotient and the remainder
of the division, where $q > 0$ and $n - 1 > r > 0$. For any $i > 0$, it is:
\begin{align}
\nonumber
&m=q(n-1)+r\\
\nonumber
&\quad\Rightarrow\frac{\bcancel{m}}{\bcancel{m}(n-1)}=
\frac{q\bcancel{(n-1)}}{m\bcancel{(n-1)}}+\frac{r}{m(n-1)}\\
\nonumber
&\quad\Rightarrow\frac{i}{n-1}=\frac{iq}{m}+i\frac{r}{m(n-1)}\\
\label{EQ11}
&\quad\Rightarrow\frac{iq}{m}=\frac{i}{n-1}-\left(i\frac{r}{n-1}\right)\frac{1}{m}\\
\label{EQ12}
&\quad\Rightarrow\frac{iq+1}{m}=\frac{i}{n-1}+\left(1-i\frac{r}{n-1}\right)\frac{1}{m}
\end{align}
Both equations \eqref{EQ11} and \eqref{EQ12} have something significant to say for $i = 1$.
Equation \eqref{EQ11} takes the following form:
\begin{equation*}
\frac{q}{m}=\frac{1}{n-1}-\left(\frac{r}{n-1}\right)\frac{1}{m}
\end{equation*}
where $r/(n - 1) > 0$, so that $q/m < 1/(n - 1)$. This implies that, if $k$ is the first key in
$I(mi,ma)$ for which $f$ matches any given value, the subsequent $q - 1$ keys are never sufficient
to increase $f$ by one. Thus, function $index$ fills every bucket but the last one -- which collects
the objects mapped to the last key only -- with the objects mapped to $1 + q - 1 = q$ keys at least.
For its part, equation \eqref{EQ12} takes the following form:
\begin{equation*}
\frac{q+1}{m}=\frac{1}{n-1}+\left(1-\frac{r}{n-1}\right)\frac{1}{m}
\end{equation*}
where $1 - r/(n - 1) > 0$, so that $(q + 1)/m > 1/(n - 1)$. Therefore, the $q$ keys following any
aforesaid key $k$ are always sufficient to increase $f$ by one. Hence, function $index$ fills every
bucket with the objects mapped to $1 + q = q + 1$ keys at most. A further consequence is that $f$
changes from zero to one for $k$ matching the $(q + 1)$-th key in $I(mi,ma)$, which entails that the
first bucket collects the objects mapped to exactly the first $q$ keys.
Which is the first $i_1$, if any, such that the bucket with index $i_1$ collects the objects mapped
to $q + 1$, rather than $q$, keys? Such bucket, if any, is preceded by $i_1$ buckets (as indices are
zero-based), whose total probability is $i_1q/m$ (as each of those buckets accommodates a group of
$q$ keys). So, $i_1$ is the least index, if any, such that $0 < i_1 < n - 1$ and $[(i_1 + 1)q + 1]/m
< (i_1 + 1)/(n - 1)$. Rewriting the latter inequality using equation \eqref{EQ12}, it results:
\begin{align*}
&\bcancel{\frac{i_1+1}{n-1}}+\left[1-(i_1+1)\frac{r}{n-1}\right]\frac{1}{m}<
\bcancel{\frac{i_1+1}{n-1}}\\
&\quad\Rightarrow\left[1-(i_1+1)\frac{r}{n-1}\right]\bcancel{\frac{1}{m}}<0\\
&\quad\Rightarrow(i_1+1)\frac{r}{n-1}>1\\
&\quad\Rightarrow i_1>\frac{n-1}{r}-1
\end{align*}
where $(n - 1)/r - 1 > 0$ since $r < n - 1$. Hence, index $i_1$ there exists just in case:
\begin{align*}
&\frac{n-1}{r}-1<n-2\\
&\quad\Rightarrow\frac{\bcancel{n-1}}{r}<\bcancel{n-1}\\
&\quad\Rightarrow r>1
\end{align*}
Likewise, let $i_2$ be the next index, if any, such that the bucket with index $i_2$ accommodates a
group of $q + 1$ keys. Such bucket, if any, is preceded by $i_2 - 1$ buckets accommodating $q$ keys
and one bucket accommodating $q + 1$ keys, whose total probability is $(i_2q + 1)/m$. Thus, $i_2$ is
the least index, if any, such that $i_1 < i_2 < n - 1$ and $[(i_2 + 1)q + 2]/m < (i_2 + 1)/(n - 1)$.
Adding term $1/m$ to both sides of equation \eqref{EQ12}, the latter inequality can be rewritten as
follows:
\begin{align*}
&\bcancel{\frac{i_2+1}{n-1}}+\left[2-(i_2+1)\frac{r}{n-1}\right]\frac{1}{m}<
\bcancel{\frac{i_2+1}{n-1}}\\
&\quad\Rightarrow\left[2-(i_2+1)\frac{r}{n-1}\right]\bcancel{\frac{1}{m}}<0\\
&\quad\Rightarrow(i_2+1)\frac{r}{n-1}>2\\
&\quad\Rightarrow i_2>\frac{2(n-1)}{r}-1
\end{align*}
where $2(n - 1)/r - 1 > [(n - 1)/r - 1] + 1 \geq i_1$. Hence, index $i_2$ there exists just in case:
\begin{align*}
&\frac{2(n-1)}{r}-1<n-2\\
&\quad\Rightarrow\frac{2\bcancel{(n-1)}}{r}<\bcancel{n-1}\\
&\quad\Rightarrow r>2
\end{align*}
To sum up, in this subcase function $index$ turns out to work as follows:
\begin{itemize}
\item
The $r - 1$ buckets whose indices $i_j$ match the least solutions of inequalities $i_j > j(n - 1)/r
- 1$, for $1 \leq j \leq r - 1$, accommodate a group of $q + 1$ contiguous keys each, so that each
one has probability $(q + 1)/m$.
\item
The other $n - 1 - (r - 1) = n - r$ buckets excluding the last one, particularly the first bucket,
accommodate a group of $q$ contiguous keys each, so that each one has probability $q/m$.
\item
The last bucket accommodates the last key alone, so that its probability is $1/m$.
\end{itemize}
Indeed, $(q + 1)(r - 1) + q(n - r) + 1 = \bcancel{qr} - q + r - \bcancel{1} + qn - \bcancel{qr} +
\bcancel{1} = q(n - 1) + r = m$. Furthermore, being $2q/m \geq (q + 1)/m$, the maximum value among
buckets' probabilities is at most twice the minimum one, excluding the last bucket if $q > 2$.
Two further observations can be made concerning case 1. First, if $m > n - 1$, then the larger $q$
gets, the more efficient it becomes to use the buckets' number $n$ itself instead of $n - 1$ within
function $r$, placing the objects with index $n$, viz. mapped to the last key, into the bucket with
index $n - 1$. In fact, this ensures that all the buckets have almost uniform probabilities rather
than leaving a bucket, the last one, with a small, or even negligible, probability.
Second, if keys are integers and $I(mi,ma)$ includes all the integers comprised between $mi$ and
$ma$, it is $m = ma - mi + 1$, whereas the cardinality of set $E(k) \cap I(mi,ma)$ is $k - mi + 1$
for any $k \in I(mi,ma)$. Therefore, it results:
\begin{equation*}
r(k,n,mi,ma)=(n-1)\frac{k - mi + 1}{ma - mi + 1},
\end{equation*}
so that function $r$ resembles the approximate rank function $R$ described in \cite{R3}.
In case 2, let $Z$ be the set of the integers $i$ such that $0 \leq i \leq n-1$. As $r(k,n,mi,ma)$
matches 0 for $k = mi$ and $n-1$ for $k = ma$, by the intermediate value theorem, for each $i \in Z$
there exists a least $k_i \in [mi,ma]$ such that $r(k_i,n,mi,ma) = i$, where $k_0 = mi$. Then, let
$B_i = [k_i, k_{i+1})$ for each $i \in Z - \{n-1\}$ and $B_{n-1} = [k_{n-1}, ma]$.
For any $i \in Z - \{n-1\}$, $k \in B_i$, it is $r(k,n,mi,ma) \neq i+1$, since otherwise there would
exist some $k < k_{i+1}$ in $[mi,ma]$ such that $r(k,n,mi,ma) = i+1$. On the other hand, being $k <
k_{i+1}$, it is $r(k,n,mi,ma) \leq i+1$, since function $r$ is increasing with respect to variable
$k$. Hence, it turns out that $r(k,n,mi,ma) < i+1$. Moreover, the monotonicity of $r$ also implies
that $r(k,n,mi,ma) \geq i$. Therefore, it is $f(k,n,mi,ma) = i$, so that for any $i \in Z$, function
$index$ fills the bucket with index $i$ with the objects mapped to the keys in $B_i$.
Consequently, for each $i \in Z - \{n-1\}$, the probability of the bucket with index $i$ is:
\begin{align*}
&P(B_i\;|\;I(mi,ma))\\
&\quad=\frac{P(B_i\cap I(mi,ma))}{P(I(mi,ma))}\\
&\quad=\frac{P((k_i,k_{i+1}]\cap I(mi,ma))}{P(I(mi,ma))}\\
&\quad=\frac{P(E(k_{i+1})\cap I(mi,ma))-P(E(k_i)\cap I(mi,ma))}{P(I(mi,ma))}\\
&\quad=\frac{P(E(k_{i+1})\cap I(mi,ma))}{P(I(mi,ma))}-\frac{P(E(k_i)\cap I(mi,ma))}{P(I(mi,ma))}\\
&\quad=P(E(k_{i+1})\;|\;I(mi,ma))-P(E(k_i)\;|\;I(mi,ma))\\
&\quad=\frac{(n-1)\cdot P(E(k_{i+1})\;|\;I(mi,ma))-(n-1)\cdot P(E(k_i)\;|\;I(mi,ma))}{n-1}\\
&\quad=\frac{r(k_{i+1},n,mi,ma)-r(k_i,n,mi,ma)}{n-1}\\
&\quad=\frac{\bcancel{i}+1-\bcancel{i}}{n-1}\\
&\quad=\frac{1}{n-1}
\end{align*}
Observe that the computation uses:
\begin{itemize}
\item
The definition of conditional probability.
\item
The fact that events $B_i$ and $(k_i, k_{i+1}]$ differ by singletons $\{k_i\}$ and $\{k_{i+1}\}$,
whose probability is zero.
\\Indeed, it is $P(\{k_0\}) = P(\{mi\}) = 0$ by hypothesis, whereas for any $k \in (mi,ma]$, it is
$P(\{k\}) = 0$ due to the continuity of function $r$, and then of function $P(E(k) \cap I(mi,ma))$,
in point $k$. In fact, for any $k' \in (mi,k)$ it is $E(k') \cap I(mi,ma) = [mi,k'] \subset [mi,k)$,
so that $P(E(k') \cap I(mi,ma)) \leq P([mi,k))$. However, it is also $E(k) \cap I(mi,ma) = [mi,k] =
[mi,k) \cup \{k\}$, so that $P(E(k) \cap I(mi,ma)) = P([mi,k)) + P(\{k\})$. Thus, if $P(\{k\}) > 0$,
then $P(E(k) \cap I(mi,ma)) > P([mi,k))$, in contradiction with the assumption that:
\begin{equation*}
\lim_{k'\to k^-}P(E(k') \cap I(mi,ma))=P(E(k) \cap I(mi,ma))
\end{equation*}
\item
The fact that event $E(k_{i+1}) \cap I(mi,ma)$ is equal to the union of the disjoint events $E(k_i)
\cap I(mi,ma)$ and $(k_i, k_{i+1}] \cap I(mi,ma)$, so that the probability of the former event is
equal to the sum of the probabilities of the latter ones.
\end{itemize}
As a result, all the buckets but the last one are equiprobable, whereas the last one has probability
zero. Thus, in this case it is again more efficient to replace $n - 1$ with $n$ within function $r$,
assigning the objects with index $n$, viz. mapped to the keys falling in $B_n$, to the bucket with
index $n - 1$, which ensures that all the buckets have uniform probabilities.
If function $r$ is linear for $k \in [mi,ma]$, viz. if interval $[mi,ma]$ is endowed with a constant
probability density, then the function's graph (with factor $n - 1$ replaced by $n$) is the straight
line passing through points $(mi,0)$ and $(ma,n)$. Therefore, it results:
\begin{equation*}
r(k,n,mi,ma)=n\frac{k - mi}{ma - mi},
\end{equation*}
so that function $r$ matches the approximate rank function $R$ described in \cite{R3}.
\<close>
subsubsection "Buckets' number -- Proof"
text \<open>
\label{SEC3}
Given $n$ equiprobable buckets and $k$ objects to be partitioned randomly among such buckets, where
$1 < k \leq n$, the probability $P_{n,k}$ that each bucket will contain at most one object is given
by equation \eqref{EQ1}, namely:
\begin{equation*}
P_{n,k}=\frac{n!}{(n-k)!n^k}
\end{equation*}
Thus, it is:
\begin{align*}
&P_{n+1,k}-P_{n,k}\\
&\quad=\frac{(n+1)!}{(n-k+1)(n-k)!(n+1)^k}-\frac{n!}{(n-k)!n^k}\\
&\quad=\frac{(n+1)!n^k-n!(n-k+1)(n+1)^k}{(n-k+1)(n-k)!n^k(n+1)^k}
\end{align*}
Using the binomial theorem and Pascal's rule, the numerator of the fraction in the right-hand side
of this equation can be expressed as follows:
\begin{align*}
&(n+1)!n^k-(n-k+1)n!(n+1)^k\\
&\quad=n!(n+1)n^k+n!k(n+1)^k-n!n(n+1)^k-n!(n+1)^k\\
&\quad=n!n^{k+1}+n!n^k\\
&\quad\quad+n!k\left[n^k+\binom{k}{1}n^{k-1}+\binom{k}{2}n^{k-2}+...
+\binom{k}{k-1}n+\binom{k}{k}\right]\\
&\quad\quad-n!n\left[n^k+\binom{k}{1}n^{k-1}+\binom{k}{2}n^{k-2}+...
+\binom{k}{k-1}n+\binom{k}{k}\right]\\
&\quad\quad-n!\left[n^k+\binom{k}{1}n^{k-1}+\binom{k}{2}n^{k-2}+...
+\binom{k}{k-1}n+\binom{k}{k}\right]\\
&\quad=\bcancel{n!n^{k+1}}+\bcancel{n!n^k}+\bcancel{n!kn^k}\\
&\quad\quad+n!k\binom{k}{1}n^{k-1}+n!k\binom{k}{2}n^{k-2}+...
+n!k\binom{k}{k-1}n+n!k\\
&\quad\quad-\bcancel{n!n^{k+1}}-\bcancel{n!kn^k}-n!\binom{k}{2}n^{k-1}-...
-n!\binom{k}{k-1}n^2-n!\binom{k}{k}n\\
&\quad\quad-\bcancel{n!n^k}-n!\binom{k}{1}n^{k-1}-n!\binom{k}{2}n^{k-2}-...
-n!\binom{k}{k-1}n-n!\\
&\quad=n!(k-1)+n!n^{k-1}\left[k\binom{k}{1}-\binom{k}{1}-\binom{k}{2}\right]+...\\
&\quad\quad+n!n\left[k\binom{k}{k-1}-\binom{k}{k-1}-\binom{k}{k}\right]\\
&\quad=n!(k-1)+n!\cdot\sum_{i=1}^{k-1}
n^{k-i}\left\{k\binom{k}{i}-\left[\binom{k}{i}+\binom{k}{i+1}\right]\right\}\\
&\quad=n!(k-1)+n!\cdot\sum_{i=1}^{k-1}
n^{k-i}\left[k\binom{k}{i}-\binom{k+1}{i+1}\right]\\
&\quad=n!(k-1)+n!\cdot\sum_{i=1}^{k-1}
n^{k-i}\left[\frac{kk!}{i!(k-i)!}-\frac{(k+1)!}{(i+1)i!(k-i)!}\right]\\
&\quad=n!(k-1)+n!\cdot\sum_{i=1}^{k-1}
n^{k-i}k!\frac{(i+1)k-(k+1)}{(i+1)i!(k-i)!}\\
&\quad=n!(k-1)+n!\cdot\sum_{i=1}^{k-1}
n^{k-i}k!\frac{ik-1}{(i+1)i!(k-i)!}>0
\end{align*}
Therefore, for any fixed $k > 1$, sequence $(P_{n,k})_{n \geq k}$ is strictly increasing, viz. the
larger $n$ is, such is also the probability that each of the $n$ equiprobable buckets will contain
at most one of the $k$ given objects.
Moreover, it is $P_{n,k} = [n(n-1)(n-2)(n-3)...(n-k+1)]/n^k < n^k/n^k = 1$, as the product enclosed
within the square brackets comprises $k$ factors, one equal to $n$ and the other ones less than $n$.
On the other hand, it turns out that:
\begin{align*}
&n(n-1)(n-2)(n-3)...(n-k+1)\\
&\quad=n^2[(n-2)(n-3)...(n-k+1)]-n[(n-2)(n-3)...(n-k+1)]\\
&\quad\geq n^2[(n-2)(n-3)...(n-k+1)]-n\cdot n^{k-2}\\
&\quad=n[n(n-2)(n-3)...(n-k+1)]-n^{k-1}\\
&\quad=n\cdot n^2[(n-3)...(n-k+1)]-n\cdot 2n[(n-3)...(n-k+1)]-n^{k-1}\\
&\quad\geq n^3[(n-3)...(n-k+1)]-2n^2\cdot n^{k-3}-n^{k-1}\\
&\quad=n^2[n(n-3)...(n-k+1)]-(1+2)n^{k-1}\;...
\end{align*}
Hence, applying the same line of reasoning until the product within the square brackets disappears,
it results:
\begin{align*}
&n(n-1)(n-2)(n-3)...(n-k+1)\\
&\quad\geq n^k-[1+2+...+(k-1)]n^{k-1}\\
&\quad=n^k-\frac{k(k-1)}{2}n^{k-1},
\end{align*}
so that:
\begin{equation*}
P_{n,k}=\frac{n(n-1)(n-2)(n-3)...(n-k+1)}{n^k}\geq 1-\frac{k(k-1)}{2n}
\end{equation*}
Therefore, for any fixed $k > 1$, the terms of sequence $(P_{n,k})_{n \geq k}$ are comprised between
the corresponding terms of sequence $(1 - k(k-1)/2n)_{n \geq k}$ and constant sequence $(1)_{n \geq
k}$. Since both of these sequences converge to 1, by the squeeze theorem it is:
\begin{equation*}
\lim_{n\to\infty}P_{n,k}=1,
\end{equation*}
viz. the larger $n$ is, the closer to 1 is the probability that each of the $n$ equiprobable buckets
will contain at most one of the $k$ given objects.
As a result, the probability of placing at most one object into each bucket in any algorithm's round
is maximized by increasing the number of the buckets as much as possible, Q.E.D..
\<close>
subsubsection "Buckets' number -- Implementation"
text \<open>
\label{SEC4}
Let $n$ be the number of the objects to be sorted, and $p$ the upper bound on the counters' number
-- and then on the buckets' number as well, since there must be exactly one counter per bucket --.
This means that before the round begins, the objects to be split are located in $m$ buckets $B_1$,
..., $B_m$, where $0 < m \leq p$, respectively containing $n_1$, ..., $n_m$ objects, where $n_i > 0$
for each $i$ from 1 to $m$ and $n_1 + ... + n_m = n$.
Moreover, let $c$ be the number of the objects known to be the sole elements of their buckets, viz.
to require no partition into finer-grained buckets, at the beginning of a given algorithm's round.
Then, the number of the objects requiring to be split into finer-grained buckets in that round is
$n - c$, whereas the number of the available buckets is $p - c$, since $c$ counters must be left to
store as many 1s, one for each singleton bucket.
How to compute $c$? At first glance, the answer seems trivial: by counting, among the counters input
(either by the algorithm's caller or by the previous round) to the round under consideration, those
that match 1. However, this value would not take into account the fact that, for each non-singleton
bucket, the algorithm must find the leftmost occurrence of the minimum key, as well as the rightmost
occurrence of the maximum key, and place the corresponding objects into two new singleton buckets,
which shall be the first and the last finer-grained bucket, respectively.
The most fundamental reason for this is that, as a result of the partition of such a bucket, nothing
prevents all its objects from falling in the same finer-grained bucket -- particularly, this happens
whenever all its objects have the same key --, in which case the algorithm does not terminate unless
some object is removed from the bucket prior to the partition, so as to reduce its size. Just as
clearly, the algorithm must know where to place the finer-grained buckets containing the removed
objects with respect to the finer-grained buckets resulting from the partition. This is exactly what
is ensured by removing objects with minimum or maximum keys, whereas selecting the leftmost or the
rightmost ones, respectively, preserves the algorithm's stability.
Actually, the algorithm's termination requires the removal of at least one object per non-singleton
bucket, so the removal of one object only, either with minimum or maximum key, would be sufficient.
Nonetheless, the leftmost minimum and the rightmost maximum can be searched via a single loop, and
finding both of them enables to pass them as inputs to the function $index$ described in section
\ref{SEC2}, or to whatever other function used to split buckets into finer-grained ones. Moreover,
non-singleton buckets whose objects all have the same key can be detected as those whose minimum and
maximum keys are equal. This allows to optimize the algorithm by preventing it from unnecessarily
applying multiple recursive rounds to any such bucket; it shall rather be left as is, just replacing
its counter with as many 1s as its size to indicate that it is already sorted.
Therefore, as the round begins, the objects already known to be placed in singleton buckets are one
for each bucket whose counter matches 1, and two for each bucket whose counter is larger than 1. As
a result, $c$ shall be computed as follows. First, initialize $c$ to zero. Then, for each $i$ from 1
to $m$, increase $c$ by one if $n_i = 1$, by two otherwise.
Conversely, for any such $i$, the number $N_i$ of the objects contained in bucket $B_i$ having to be
partitioned into finer-grained buckets is 0 if $n_i = 1$, $n_i - 2$ otherwise, so that $N_1 + ... +
N_m = n - c$. According to the findings of section \ref{SEC3}, the number $N'_i$ of the resulting
finer-grained buckets should be maximized, and the most efficient way to do this is to render $N'_i$
proportional to $N_i$, since otherwise, viz. if some buckets were preferred to some other ones, the
unprivileged buckets would form as many bottlenecks.
This can be accomplished by means of the following procedure. First, initialize integers $R$ and $U$
to 0. Then, for each $i$ from 1 to $m$, check whether $N_i \leq 1$:
\begin{itemize}
\item
If so, set $N'_i$ to $N_i$.
\\In fact, no finer-grained bucket is necessary if there are no objects to be split, while a single
finer-grained bucket is sufficient for a lonely object.
\item
Otherwise, perform the integer division of $N_i \cdot (p - c) + R$ by $n - c$, and set integer $Q$
to the resulting quotient and $R$ to the resulting remainder. Then, if the minimum and maximum keys
occurring in bucket $B_i$ are equal, increase $U$ by $Q - N_i$, otherwise set $N'_i$ to $U + Q$ and
reset $U$ to 0.
\\In fact, as observed above, if its minimum and maximum keys are equal, bucket $B_i$ can be split
into $n_i = N_i + 2$ singleton buckets. Hence, the difference $Q - N_i$ between the number of the
available finer-grained buckets, i.e. $Q + 2$ (where 2 is the number of the buckets containing the
leftmost minimum and the rightmost maximum), and the number of those being used, i.e. $N_i + 2$, can
be added to the total number $U$ of the available finer-grained buckets still unused in the current
round. Such buckets can then be utilized as soon as a bucket $B_j$ with $N_j > 1$ whose minimum and
maximum keys do not match is encountered next, in addition to those already reserved for $B_j$.
\end{itemize}
Of course, for any $i$ from 1 to $m$ such that $N_i > 1$, it is $N_i \leq Q$ -- viz. the number of
the objects in $B_i$ to be split is not larger than that of the finer-grained buckets where they are
to be placed even if $U = 0$, so that the probability of placing at most one object into each bucket
is nonzero -- just in case $n - c \leq p - c$, i.e. $n \leq p$. Indeed, it will be formally proven
that for $n \leq p$, this procedure is successful in maximizing the buckets' number in each round
never exceeding the upper bound $p$.
\<close>
subsubsection "Generalized counting sort (GCsort)"
text \<open>
The conclusions of the efficiency analysis performed so far, put together, result in the following
\emph{generalized counting sort (GCsort)} algorithm.
Let $xs$ be the input array containing $n$ objects to be sorted, and $ns$ an array of $p$ integers,
where $0 < p$ and $n \leq p$. Moreover, let $xs'$ and $ns'$ be two further arrays of the same type
and size of $xs$ and $ns$, respectively, and let $i$, $i'$, and $j$ be as many integers.
Then, GCsort works as follows (assuming arrays to be zero-based):
\begin{enumerate}
\item
Initialize the first element of $ns$ to $n$ and any other element to 0.
\item
Check whether $ns$ contains any element larger than 1.
\\If not, terminate the algorithm and output $xs$ as the resulting sorted array.
\item
Initialize $i$, $i'$, and $j$ to 0.
\item
Check whether $ns[i] = 1$ or $ns[i] > 1$:
\begin{enumerate}
\item
In the former case, set $xs'[j]$ to $xs[j]$ and $ns'[i']$ to 1.
\\Then, increase $i'$ and $j$ by 1.
\item
In the latter case, partition the bucket comprised of objects $xs[j]$ to $xs[j+ns[i]-1]$ into
finer-grained buckets according to section \ref{SEC4}, storing the resulting $n'$ buckets in
$xs'[j]$ to $xs'[j+ns[i]-1]$ and their sizes in $ns'[i']$ to $ns'[i'+n'-1]$.
\\Then, increase $i'$ by $n'$ and $j$ by $ns[i]$.
\end{enumerate}
\item
Increase $i$ by 1, and then check whether $i < p$.
\\If so, go back to step 4.
\\Otherwise, perform the following operations:
\begin{enumerate}
\item
If $i' < p$, set integers $ns'[i']$ to $ns'[p-1]$ to 0.
\item
Swap addresses $xs$ and $xs'$, as well as addresses $ns$ and $ns'$.
\item
Go back to step 2.
\end{enumerate}
\end{enumerate}
Since the algorithm is tail-recursive, the memory space required for its execution matches the one
required for a single recursive round, which is $O(n) + O(p)$.
The best-case running time can be computed easily. The running time taken by step 1 is equal to $p$.
Moreover, the partition of a bucket into finer-grained ones is performed by determining their sizes,
computing the cumulative sum of such sizes, and rearranging the bucket's objects according to the
resulting offsets. All these operations only involve sequential, non-nested loops, which iterate
through either the objects or the finer-grained counters pertaining to the processed bucket alone.
Hence, the running time taken by a single recursive round is $O(n) + O(p)$, so that in the best case
where at most one round is executed after step 1, the running time taken by the algorithm is $O(n) +
O(p)$, too.
The asymptotic worst-case running time can be computed as follows. Let $t_{n,p}$ be the worst-case
running time taken by a single round. As $t_{n,p}$ is $O(n) + O(p)$, there exist three real numbers
$a > 0$, $b > 0$, and $c$ such that $t_{n,p} \leq an + bp + c$. Moreover, let $U_{n,p}$ be the set
of the $p$-tuples of natural numbers such that the sum of their elements matches $n$, and $max(u)$
the maximum element of a given $u \in U_{n,p}$. Finally, let $T_{n,p,u}$ be the worst-case running
time taken by the algorithm if it starts from step 2, viz. skipping step 1, using as initial content
of array $ns$ the $p$-tuple $u \in U_{n,p}$.
Then, it can be proven by induction on $max(u)$ that:
\begin{equation}
\label{EQ13}
T_{n,p,u}\leq
\begin{cases}
a\dfrac{max(u)}{2}n+\left[b\dfrac{max(u)}{2}+1\right]p+c\dfrac{max(u)}{2}\\
\quad\text{if $max(u)$ is even,}\\
\\
a\dfrac{max(u)-1}{2}n+\left[b\dfrac{max(u)-1}{2}+1\right]p+c\dfrac{max(u)-1}{2}\\
\quad\text{if $max(u)$ is odd}
\end{cases}
\end{equation}
In fact, if $max(u) = 0$, the initial $p$-tuple $u$ matches the all-zero one. Hence, the algorithm
executes step 2 just once and then immediately terminates. Therefore, the running time is $p$, which
matches the right-hand side of inequality \eqref{EQ13} for $max(u) = 0$.
If $max(u) = 1$, $u$ contains no element larger than 1. Thus, again, the algorithm terminates just
after the first execution of step 2. As a result, the running time is still $p$, which matches the
right-hand side of inequality \eqref{EQ13} for $max(u) = 1$.
If $max(u) = 2$, $u$ contains some element larger than 1, so that one round is executed, taking time
$t_{n,p}$ in the worst case. Once this round is over, array $ns$ will contain a $p$-tuple $u' \in
U_{n,p}$ such that $max(u') = 1$. Hence, step 2 is executed again, taking time $p$, and then the
algorithm terminates. As a result, it is:
\begin{equation*}
T_{n,p,u}\leq an+bp+c+p=an+(b+1)p+c,
\end{equation*}
which matches the right-hand side of inequality \eqref{EQ13} for $max(u) = 2$.
Finally, if $max(u) > 2$, $u$ has some element larger than 1, so one round is executed, taking time
$t_{n,p}$ in the worst case. Once this round is over, array $ns$ will contain a $p$-tuple $u' \in
U_{n,p}$ such that $max(u') \leq max(u) - 2$, because of the removal of the leftmost minimum and the
rightmost maximum from any non-singleton bucket. By the induction hypothesis, the worst-case time
$T_{n,p,u'}$ taken by the algorithm from this point onward complies with inequality \eqref{EQ13},
whose right-hand side is maximum if such is $max(u')$, viz. if $max(u') = max(u) - 2$.
As a result, if $max(u)$ is even, it is:
\begin{align*}
&T_{n,p,u}\leq an+bp+c\\
&\quad\quad+a\frac{max(u)-2}{2}n+\left[b\frac{max(u)-2}{2}+1\right]p+c\frac{max(u)-2}{2}\\
&\quad=a\frac{max(u)}{2}n+\left[b\frac{max(u)}{2}+1\right]p+c\frac{max(u)}{2},
\end{align*}
which matches the right-hand side of inequality \eqref{EQ13} for an even $max(u)$.
Similarly, if $max(u)$ is odd, it is:
\begin{align*}
&T_{n,p,u}\leq an+bp+c\\
&\quad\quad+a\frac{max(u)-3}{2}n+\left[b\frac{max(u)-3}{2}+1\right]p+c\frac{max(u)-3}{2}\\
&\quad=a\frac{max(u)-1}{2}n+\left[b\frac{max(u)-1}{2}+1\right]p+c\frac{max(u)-1}{2},
\end{align*}
which matches the right-hand side of inequality \eqref{EQ13} for an odd $max(u)$.
With this, the proof of inequality \eqref{EQ13} is complete. Now, let $T_{n,p}$ be the worst-case
time taken by the algorithm executed in full. Step 1 is executed first, taking time $p$. Then, array
$ns$ contains a $p$-tuple $u$ such that $max(u) = n$, and by definition, the worst-case time taken
by the algorithm from this point onward is $T_{n,p,u}$. Therefore, applying inequality \eqref{EQ13},
it turns out that:
\begin{align*}
&T_{n,p}=p+T_{n,p,u}\\
&\quad\leq
\begin{cases}
a\dfrac{n^2}{2}+\left(b\dfrac{n}{2}+2\right)p+c\dfrac{n}{2}
&\quad\text{if $n$ is even,}\\
\\
a\dfrac{n^2-n}{2}+\left(b\dfrac{n-1}{2}+2\right)p+c\dfrac{n-1}{2}
&\quad\text{if $n$ is odd}
\end{cases}
\end{align*}
As a result, the asymptotic worst-case running time taken by the algorithm is $O(n^2) + O(np)$.
\<close>
subsection "Formal definitions"
text \<open>
Here below, a formal definition of GCsort is provided, which will later enable to formally prove the
correctness of the algorithm. Henceforth, the main points of the formal definitions and proofs are
commented. For further information, see Isabelle documentation, particularly \cite{R5}, \cite{R6},
\cite{R7}, and \cite{R8}.
The following formalization of GCsort does not define any specific function @{text index} to be used
to split buckets into finer-grained ones. It rather defines only the type @{text index_sign} of such
functions, matching the signature of the function $index$ described in section \ref{SEC2}, along
with three predicates that whatever chosen @{text index} function is required to satisfy for GCsort
to work correctly:
\begin{itemize}
\item
Predicate @{text index_less} requires function @{text index} to map any object within a given bucket
to an index less than the number of the associated finer-grained buckets (\emph{less than} instead
of \emph{not larger than} since type @{typ "'a list"} is zero-based).
\item
Predicate @{text index_mono} requires function @{text index} to be monotonic with respect to the
keys of the objects within a given bucket.
\item
Predicate @{text index_same} requires function @{text index} to map any two distinct objects within
a given bucket that have the same key to the same index (premise \emph{distinct} is added to enable
this predicate to be used by the simplifier).
\end{itemize}
\null
\<close>
type_synonym ('a, 'b) index_sign = "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> nat"
definition index_less :: "('a, 'b::linorder) index_sign \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where
"index_less index key \<equiv>
\<forall>x n mi ma. key x \<in> {mi..ma} \<longrightarrow> 0 < n \<longrightarrow>
index key x n mi ma < n"
definition index_mono :: "('a, 'b::linorder) index_sign \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where
"index_mono index key \<equiv>
\<forall>x y n mi ma. {key x, key y} \<subseteq> {mi..ma} \<longrightarrow> key x \<le> key y \<longrightarrow>
index key x n mi ma \<le> index key y n mi ma"
definition index_same :: "('a, 'b::linorder) index_sign \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
where
"index_same index key \<equiv>
\<forall>x y n mi ma. key x \<in> {mi..ma} \<longrightarrow> x \<noteq> y \<longrightarrow> key x = key y \<longrightarrow>
index key x n mi ma = index key y n mi ma"
text \<open>
\null
Functions @{text bn_count} and @{text bn_comp} count, respectively, the objects known to be placed
in singleton buckets in a given round, and the finer-grained buckets available to partition a given
non-singleton bucket, according to section \ref{SEC4}.
\null
\<close>
fun bn_count :: "nat list \<Rightarrow> nat" where
"bn_count [] = 0" |
"bn_count (Suc (Suc (Suc (Suc n))) # ns) = Suc (Suc (bn_count ns))" |
"bn_count (n # ns) = n + bn_count ns"
fun bn_comp :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat" where
"bn_comp (Suc (Suc n)) p q r =
((Suc (Suc n) * p + r) div q, (Suc (Suc n) * p + r) mod q)" |
"bn_comp n p q r = (n, r)"
fun bn_valid :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
"bn_valid (Suc (Suc n)) p q = (q \<in> {0<..p})" |
"bn_valid n p q = True"
text \<open>
\null
Functions @{text mini} and @{text maxi} return the indices of the leftmost minimum and the rightmost
maximum within a given non-singleton bucket.
\null
\<close>
primrec (nonexhaustive) mini :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b::linorder) \<Rightarrow> nat" where
"mini (x # xs) key =
(let m = mini xs key in if xs = [] \<or> key x \<le> key (xs ! m) then 0 else Suc m)"
primrec (nonexhaustive) maxi :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b::linorder) \<Rightarrow> nat" where
"maxi (x # xs) key =
(let m = maxi xs key in if xs = [] \<or> key (xs ! m) < key x then 0 else Suc m)"
text \<open>
\null
Function @{text enum} counts the objects contained in each finer-grained bucket reserved for the
partition of a given non-singleton bucket.
\null
\<close>
primrec enum :: "'a list \<Rightarrow> ('a, 'b) index_sign \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow>
nat \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> nat list" where
"enum [] index key n mi ma = replicate n 0" |
"enum (x # xs) index key n mi ma =
(let i = index key x n mi ma;
ns = enum xs index key n mi ma
in ns[i := Suc (ns ! i)])"
text \<open>
\null
Function @{text offs} computes the cumulative sum of the resulting finer-grained buckets' sizes so
as to generate the associated offsets' list.
\null
\<close>
primrec offs :: "nat list \<Rightarrow> nat \<Rightarrow> nat list" where
"offs [] i = []" |
"offs (n # ns) i = i # offs ns (i + n)"
text \<open>
\null
Function @{text fill} fills the finer-grained buckets with their respective objects.
\null
\<close>
primrec fill :: "'a list \<Rightarrow> nat list \<Rightarrow> ('a, 'b) index_sign \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow>
nat \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> 'a option list" where
"fill [] ns index key n mi ma = replicate n None" |
"fill (x # xs) ns index key n mi ma =
(let i = index key x (length ns) mi ma;
ys = fill xs (ns[i := Suc (ns ! i)]) index key n mi ma
in ys[ns ! i := Some x])"
text \<open>
\null
Then, function @{text round} formalizes a single GCsort's recursive round.
\null
\<close>
definition round_suc_suc :: "('a, 'b::linorder) index_sign \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow>
'a list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat list \<times> 'a list" where
"round_suc_suc index key ws n n' u \<equiv>
let nmi = mini ws key; nma = maxi ws key;
xmi = ws ! nmi; xma = ws ! nma; mi = key xmi; ma = key xma
in if mi = ma
then (u + n' - n, replicate (Suc (Suc n)) (Suc 0), ws)
else
let k = case n of Suc (Suc i) \<Rightarrow> u + n' | _ \<Rightarrow> n;
zs = nths ws (- {nmi, nma}); ms = enum zs index key k mi ma
in (u + n' - k, Suc 0 # ms @ [Suc 0],
xmi # map the (fill zs (offs ms 0) index key n mi ma) @ [xma])"
fun round :: "('a, 'b::linorder) index_sign \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>
nat \<times> nat list \<times> 'a list \<Rightarrow> nat \<times> nat list \<times> 'a list" where
"round index key p q r (u, [], xs) = (u, [], xs)" |
"round index key p q r (u, 0 # ns, xs) = round index key p q r (u, ns, xs)" |
"round index key p q r (u, Suc 0 # ns, xs) =
(let (u', ns', xs') = round index key p q r (u, ns, tl xs)
in (u', Suc 0 # ns', hd xs # xs'))" |
"round index key p q r (u, Suc (Suc n) # ns, xs) =
(let ws = take (Suc (Suc n)) xs; (n', r') = bn_comp n p q r;
(v, ms', ws') = round_suc_suc index key ws n n' u;
(u', ns', xs') = round index key p q r' (v, ns, drop (Suc (Suc n)) xs)
in (u', ms' @ ns', ws' @ xs'))"
text \<open>
\null
Finally, function @{text gcsort_aux} formalizes GCsort. Since the algorithm is tail-recursive, this
function complies with the requirements for an auxiliary tail-recursive function applying to step 1
of the proof method described in \cite{R4} -- henceforth briefly referred to as the \emph{proof
method} --. This feature will later enable to formally prove the algorithm's correctness properties
by means of such method.
\null
\<close>
abbreviation gcsort_round :: "('a, 'b::linorder) index_sign \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow>
nat \<Rightarrow> nat list \<Rightarrow> 'a list \<Rightarrow> nat \<times> nat list \<times> 'a list" where
"gcsort_round index key p ns xs \<equiv>
round index key (p - bn_count ns) (length xs - bn_count ns) 0 (0, ns, xs)"
function gcsort_aux :: "('a, 'b::linorder) index_sign \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> nat \<Rightarrow>
nat \<times> nat list \<times> 'a list \<Rightarrow> nat \<times> nat list \<times> 'a list" where
"gcsort_aux index key p (u, ns, xs) = (if find (\<lambda>n. Suc 0 < n) ns = None
then (u, ns, xs)
else gcsort_aux index key p (gcsort_round index key p ns xs))"
by auto
text \<open>
\null
First of all, even before accomplishing step 2 of the proof method, it is necessary to prove that
function @{const gcsort_aux} always terminates by showing that the maximum bucket's size decreases
in each recursive round.
\null
\<close>
lemma add_zeros:
"foldl (+) (m :: nat) (replicate n 0) = m"
by (induction n, simp_all)
lemma add_suc:
"foldl (+) (Suc m) ns = Suc (foldl (+) m ns)"
by (induction ns arbitrary: m, simp_all)
lemma add_update:
"i < length ns \<Longrightarrow> foldl (+) m (ns[i := Suc (ns ! i)]) = Suc (foldl (+) m ns)"
by (induction ns arbitrary: i m, simp_all add: add_suc split: nat.split)
lemma add_le:
"(m :: nat) \<le> foldl (+) m ns"
by (induction ns arbitrary: m, simp_all, rule order_trans, rule le_add1)
lemma add_mono:
"(m :: nat) \<le> n \<Longrightarrow> foldl (+) m ns \<le> foldl (+) n ns"
by (induction ns arbitrary: m n, simp_all)
lemma add_max [rule_format]:
"ns \<noteq> [] \<longrightarrow> Max (set ns) \<le> foldl (+) (0 :: nat) ns"
by (induction ns, simp_all add: add_le, erule impCE, simp, rule ballI, drule bspec,
assumption, rule order_trans, assumption, rule add_mono, simp)
lemma enum_length:
"length (enum xs index key n mi ma) = n"
by (induction xs, simp_all add: Let_def)
lemma enum_add_le:
"foldl (+) 0 (enum xs index key n mi ma) \<le> length xs"
proof (induction xs, simp_all add: Let_def add_zeros)
fix x xs
let ?i = "index key x n mi ma"
assume "foldl (+) 0 (enum xs index key n mi ma) \<le> length xs"
(is "foldl _ _ ?ns \<le> _")
thus "foldl (+) 0 (?ns[?i := Suc (?ns ! ?i)]) \<le> Suc (length xs)"
by (cases "?i < length ?ns", simp_all add: add_update)
qed
lemma enum_max_le:
"0 < n \<Longrightarrow> Max (set (enum xs index key n mi ma)) \<le> length xs"
(is "_ \<Longrightarrow> Max (set ?ns) \<le> _")
by (insert add_max [of ?ns], insert enum_add_le [of xs index key n mi ma],
simp only: length_greater_0_conv [symmetric] enum_length, simp)
lemma mini_less:
"0 < length xs \<Longrightarrow> mini xs key < length xs"
by (induction xs, simp_all add: Let_def)
lemma maxi_less:
"0 < length xs \<Longrightarrow> maxi xs key < length xs"
by (induction xs, simp_all add: Let_def)
lemma mini_lb:
"x \<in> set xs \<Longrightarrow> key (xs ! mini xs key) \<le> key x"
by (induction xs, simp_all add: Let_def, auto)
lemma maxi_ub:
"x \<in> set xs \<Longrightarrow> key x \<le> key (xs ! maxi xs key)"
by (induction xs, simp_all add: Let_def, auto)
lemma mini_maxi_neq [rule_format]:
"Suc 0 < length xs \<longrightarrow> mini xs key \<noteq> maxi xs key"
proof (induction xs, simp_all add: Let_def, rule conjI, (rule impI)+,
(rule_tac [2] impI)+, rule_tac [2] notI, simp_all, rule ccontr)
fix x xs
assume "key (xs ! maxi xs key) < key x" and "key x \<le> key (xs ! mini xs key)"
hence "key (xs ! maxi xs key) < key (xs ! mini xs key)" by simp
moreover assume "xs \<noteq> []"
hence "0 < length xs" by simp
hence "mini xs key < length xs"
by (rule mini_less)
hence "xs ! mini xs key \<in> set xs" by simp
hence "key (xs ! mini xs key) \<le> key (xs ! maxi xs key)"
by (rule maxi_ub)
ultimately show False by simp
qed
lemma mini_maxi_nths:
"length (nths xs (- {mini xs key, maxi xs key})) =
(case length xs of 0 \<Rightarrow> 0 | Suc 0 \<Rightarrow> 0 | Suc (Suc n) \<Rightarrow> n)"
proof (simp add: length_nths split: nat.split, rule allI, rule conjI, rule_tac [2] allI,
(rule_tac [!] impI)+, simp add: length_Suc_conv, erule exE, simp, blast)
fix n
assume A: "length xs = Suc (Suc n)"
hence B: "Suc 0 < length xs" by simp
hence C: "0 < length xs" by arith
have "{i. i < Suc (Suc n) \<and> i \<noteq> mini xs key \<and> i \<noteq> maxi xs key} =
{..<Suc (Suc n)} - {mini xs key} - {maxi xs key}"
by blast
thus "card {i. i < Suc (Suc n) \<and> i \<noteq> mini xs key \<and> i \<noteq> maxi xs key} = n"
by (simp add: card_Diff_singleton_if, insert mini_maxi_neq [OF B, of key],
simp add: mini_less [OF C] maxi_less [OF C] A [symmetric])
qed
lemma mini_maxi_nths_le:
"length xs \<le> Suc (Suc n) \<Longrightarrow> length (nths xs (- {mini xs key, maxi xs key})) \<le> n"
by (simp add: mini_maxi_nths split: nat.split)
lemma round_nil:
"(fst (snd (round index key p q r t)) \<noteq> []) = (\<exists>n \<in> set (fst (snd t)). 0 < n)"
by (induction index key p q r t rule: round.induct,
simp_all add: round_suc_suc_def Let_def split: prod.split)
lemma round_max_eq [rule_format]:
"fst (snd t) \<noteq> [] \<longrightarrow> Max (set (fst (snd t))) = Suc 0 \<longrightarrow>
Max (set (fst (snd (round index key p q r t)))) = Suc 0"
proof (induction index key p q r t rule: round.induct, simp_all add: Let_def split:
prod.split del: all_simps, rule impI, (rule_tac [2] allI)+, (rule_tac [2] impI)+,
(rule_tac [3] allI)+, (rule_tac [3] impI)+, rule_tac [3] FalseE)
fix index p q r u ns xs and key :: "'a \<Rightarrow> 'b"
let ?t = "round index key p q r (u, ns, xs)"
assume "ns \<noteq> [] \<longrightarrow> Max (set ns) = Suc 0 \<longrightarrow>
Max (set (fst (snd ?t))) = Suc 0"
moreover assume A: "Max (insert 0 (set ns)) = Suc 0"
hence "ns \<noteq> []"
by (cases ns, simp_all)
moreover from this have "Max (set ns) = Suc 0"
using A by simp
ultimately show "Max (set (fst (snd ?t))) = Suc 0"
by simp
next
fix index p q r u ns xs u' ns' xs' and key :: "'a \<Rightarrow> 'b"
let ?t = "round index key p q r (u, ns, tl xs)"
assume A: "?t = (u', ns', xs')" and
"ns \<noteq> [] \<longrightarrow> Max (set ns) = Suc 0 \<longrightarrow> Max (set (fst (snd ?t))) = Suc 0"
hence B: "ns \<noteq> [] \<longrightarrow> Max (set ns) = Suc 0 \<longrightarrow> Max (set ns') = Suc 0"
by simp
assume C: "Max (insert (Suc 0) (set ns)) = Suc 0"
show "Max (insert (Suc 0) (set ns')) = Suc 0"
proof (cases "ns' = []", simp)
assume D: "ns' \<noteq> []"
hence "fst (snd ?t) \<noteq> []"
using A by simp
hence "\<exists>n \<in> set ns. 0 < n"
by (simp add: round_nil)
then obtain n where E: "n \<in> set ns" and F: "0 < n" ..
hence G: "ns \<noteq> []"
by (cases ns, simp_all)
moreover have "n \<le> Max (set ns)"
using E by (rule_tac Max_ge, simp_all)
hence "0 < Max (set ns)"
using F by simp
hence "Max (set ns) = Suc 0"
using C and G by simp
ultimately have "Max (set ns') = Suc 0"
using B by simp
thus ?thesis
using D by simp
qed
next
fix n ns
assume "Max (insert (Suc (Suc n)) (set ns)) = Suc 0"
thus False
by (cases ns, simp_all)
qed
lemma round_max_less [rule_format]:
"fst (snd t) \<noteq> [] \<longrightarrow> Suc 0 < Max (set (fst (snd t))) \<longrightarrow>
Max (set (fst (snd (round index key p q r t)))) < Max (set (fst (snd t)))"
proof (induction index key p q r t rule: round.induct, simp_all add: Let_def split:
prod.split del: all_simps, rule impI, (rule_tac [2] allI)+, (rule_tac [2] impI)+,
(rule_tac [3] allI)+, (rule_tac [3] impI)+, rule_tac [2] ballI)
fix index p q r u ns xs and key :: "'a \<Rightarrow> 'b"
let ?t = "round index key p q r (u, ns, xs)"
assume "ns \<noteq> [] \<longrightarrow> Suc 0 < Max (set ns) \<longrightarrow>
Max (set (fst (snd ?t))) < Max (set ns)"
moreover assume A: "Suc 0 < Max (insert 0 (set ns))"
hence "ns \<noteq> []"
by (cases ns, simp_all)
moreover from this have "Suc 0 < Max (set ns)"
using A by simp
ultimately show "Max (set (fst (snd ?t))) < Max (insert 0 (set ns))"
by simp
next
fix index p q r u ns xs u' ns' xs' i and key :: "'a \<Rightarrow> 'b"
let ?t = "round index key p q r (u, ns, tl xs)"
assume
"?t = (u', ns', xs')" and
"ns \<noteq> [] \<longrightarrow> Suc 0 < Max (set ns) \<longrightarrow>
Max (set (fst (snd ?t))) < Max (set ns)"
hence "ns \<noteq> [] \<longrightarrow> Suc 0 < Max (set ns) \<longrightarrow>
Max (set ns') < Max (set ns)"
by simp
moreover assume A: "Suc 0 < Max (insert (Suc 0) (set ns))"
hence B: "ns \<noteq> []"
by (cases ns, simp_all)
moreover from this have "Suc 0 < Max (set ns)"
using A by simp
ultimately have "Max (set ns') < Max (set ns)" by simp
moreover assume "i \<in> set ns'"
hence "i \<le> Max (set ns')"
by (rule_tac Max_ge, simp)
ultimately show "i < Max (insert (Suc 0) (set ns))"
using B by simp
next
fix index p q r u n ns n' r' v ms' ws' u' ns' xs'
and key :: "'a \<Rightarrow> 'b" and xs :: "'a list"
let ?ws = "take (Suc (Suc n)) xs"
let ?ys = "drop (Suc (Suc n)) xs"
let ?r = "\<lambda>n'. round_suc_suc index key ?ws n n' u"
let ?t = "\<lambda>r' v. round index key p q r' (v, ns, ?ys)"
assume
A: "?r n' = (v, ms', ws')" and
B: "?t r' v = (u', ns', xs')"
moreover assume "\<And>ws a b c d e f g h.
ws = ?ws \<Longrightarrow> a = bn_comp n p q r \<Longrightarrow> (b, c) = bn_comp n p q r \<Longrightarrow>
d = ?r b \<Longrightarrow> (e, f) = ?r b \<Longrightarrow> (g, h) = f \<Longrightarrow>
ns \<noteq> [] \<longrightarrow> Suc 0 < Max (set ns) \<longrightarrow>
Max (set (fst (snd (?t c e)))) < Max (set ns)" and
"bn_comp n p q r = (n', r')"
ultimately have C: "ns \<noteq> [] \<longrightarrow> Suc 0 < Max (set ns) \<longrightarrow>
Max (set ns') < Max (set ns)"
by simp
from A [symmetric] show "Max (set ms' \<union> set ns') <
Max (insert (Suc (Suc n)) (set ns))"
proof (simp add: round_suc_suc_def Let_def, subst Max_less_iff, simp_all,
rule_tac impI, simp add: Let_def split: if_split_asm, rule_tac ballI,
erule_tac UnE, simp add: Let_def split: if_split_asm, (erule_tac [1-2] conjE)+)
fix i
assume "i = Suc 0 \<or> i = Suc 0 \<and> 0 < n"
hence "i = Suc 0" by blast
hence "i < Suc (Suc n)" by simp
also have "\<dots> \<le> Max (insert (Suc (Suc n)) (set ns))"
by (rule Max_ge, simp_all)
finally show "i < Max (insert (Suc (Suc n)) (set ns))" .
next
fix i
let ?nmi = "mini ?ws key"
let ?nma = "maxi ?ws key"
let ?xmi = "?ws ! ?nmi"
let ?xma = "?ws ! ?nma"
let ?mi = "key ?xmi"
let ?ma = "key ?xma"
let ?k = "case n of 0 \<Rightarrow> n | Suc 0 \<Rightarrow> n | Suc (Suc i) \<Rightarrow> u + n'"
let ?zs = "nths ?ws (- {?nmi, ?nma})"
let ?ms = "enum ?zs index key ?k ?mi ?ma"
assume "i = Suc 0 \<or> i \<in> set ?ms"
moreover {
assume "i = Suc 0"
hence "i < Suc (Suc n)" by simp
}
moreover {
assume D: "i \<in> set ?ms"
hence "i \<le> Max (set ?ms)"
by (rule_tac Max_ge, simp)
moreover have "0 < length ?ms"
using D by (rule length_pos_if_in_set)
hence "0 < ?k"
by (simp add: enum_length)
hence "Max (set ?ms) \<le> length ?zs"
by (rule enum_max_le)
ultimately have "i \<le> length ?zs" by simp
moreover have "length ?zs \<le> n"
by (rule mini_maxi_nths_le, simp)
ultimately have "i < Suc (Suc n)" by simp
}
ultimately have "i < Suc (Suc n)" ..
also have "\<dots> \<le> Max (insert (Suc (Suc n)) (set ns))"
by (rule Max_ge, simp_all)
finally show "i < Max (insert (Suc (Suc n)) (set ns))" .
next
fix i
assume D: "i \<in> set ns'"
hence "0 < length ns'"
by (rule length_pos_if_in_set)
hence "fst (snd (?t r' v)) \<noteq> []"
using B by simp
hence E: "\<exists>n \<in> set ns. 0 < n"
by (simp add: round_nil)
hence F: "ns \<noteq> []"
by (cases ns, simp_all)
show "i < Max (insert (Suc (Suc n)) (set ns))"
proof (cases "Suc 0 < Max (set ns)")
case True
hence "Max (set ns') < Max (set ns)"
using C and F by simp
moreover have "i \<le> Max (set ns')"
using D by (rule_tac Max_ge, simp)
ultimately show ?thesis
using F by simp
next
case False
moreover from E obtain j where G: "j \<in> set ns" and H: "0 < j" ..
have "j \<le> Max (set ns)"
using G by (rule_tac Max_ge, simp)
hence "0 < Max (set ns)"
using H by simp
ultimately have "Max (set ns) = Suc 0" by simp
hence "Max (set (fst (snd (?t r' v)))) = Suc 0"
using F by (rule_tac round_max_eq, simp_all)
hence "Max (set ns') = Suc 0"
using B by simp
moreover have "i \<le> Max (set ns')"
using D by (rule_tac Max_ge, simp)
ultimately have "i < Suc (Suc n)" by simp
also have "\<dots> \<le> Max (insert (Suc (Suc n)) (set ns))"
by (rule Max_ge, simp_all)
finally show ?thesis .
qed
qed
qed
termination gcsort_aux
proof (relation "measure (\<lambda>(index, key, p, t). Max (set (fst (snd t))))",
simp_all add: find_None_iff, erule exE, erule conjE)
fix index p ns xs i and key :: "'a \<Rightarrow> 'b"
let ?t = "gcsort_round index key p ns xs"
assume A: "Suc 0 < i" and B: "i \<in> set ns"
have C: "0 < length ns"
using B by (rule length_pos_if_in_set)
moreover have "\<exists>i \<in> set ns. Suc 0 < i"
using A and B ..
hence "Suc 0 < Max (set ns)"
using C by (subst Max_gr_iff, simp_all)
ultimately have "Max (set (fst (snd ?t))) < Max (set (fst (snd (0, ns, xs))))"
by (insert round_max_less [of "(0, ns, xs)"], simp)
thus "Max (set (fst (snd ?t))) < Max (set ns)" by simp
qed
text \<open>
\null
Now steps 2, 3, and 4 of the proof method, which are independent of the properties to be proven, can
be accomplished. Particularly, function @{text gcsort} constitutes the complete formal definition of
GCsort, as it puts the algorithm's inputs and outputs into their expected form.
Observe that the conditional expression contained in the definition of function @{const gcsort_aux}
need not be reflected in the definition of inductive set @{text gcsort_set} as just one alternative
gives rise to a recursive call, viz. as its only purpose is to ensure the function's termination.
\null
\<close>
definition gcsort_in :: "'a list \<Rightarrow> nat \<times> nat list \<times> 'a list" where
"gcsort_in xs \<equiv> (0, [length xs], xs)"
definition gcsort_out :: "nat \<times> nat list \<times> 'a list \<Rightarrow> 'a list" where
"gcsort_out \<equiv> snd \<circ> snd"
definition gcsort :: "('a, 'b::linorder) index_sign \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> nat \<Rightarrow>
'a list \<Rightarrow> 'a list" where
"gcsort index key p xs \<equiv> gcsort_out (gcsort_aux index key p (gcsort_in xs))"
inductive_set gcsort_set :: "('a, 'b::linorder) index_sign \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> nat \<Rightarrow>
nat \<times> nat list \<times> 'a list \<Rightarrow> (nat \<times> nat list \<times> 'a list) set"
for index key p t where
R0: "t \<in> gcsort_set index key p t" |
R1: "(u, ns, xs) \<in> gcsort_set index key p t \<Longrightarrow>
gcsort_round index key p ns xs \<in> gcsort_set index key p t"
lemma gcsort_subset:
assumes A: "t' \<in> gcsort_set index key p t"
shows "gcsort_set index key p t' \<subseteq> gcsort_set index key p t"
by (rule subsetI, erule gcsort_set.induct, rule A, rule R1)
lemma gcsort_aux_set:
"gcsort_aux index key p t \<in> gcsort_set index key p t"
proof (induction index key p t rule: gcsort_aux.induct, simp, rule conjI,
rule_tac [!] impI, rule R0, simp)
fix index p u ns xs and key :: "'a \<Rightarrow> 'b"
let ?t = "gcsort_round index key p ns xs"
assume "gcsort_aux index key p ?t \<in> gcsort_set index key p ?t"
moreover have "(u, ns, xs) \<in> gcsort_set index key p (u, ns, xs)"
by (rule R0)
hence "?t \<in> gcsort_set index key p (u, ns, xs)"
by (rule R1)
hence "gcsort_set index key p ?t \<subseteq> gcsort_set index key p (u, ns, xs)"
by (rule gcsort_subset)
ultimately show "gcsort_aux index key p ?t
\<in> gcsort_set index key p (u, ns, xs)" ..
qed
subsection "Proof of a preliminary invariant"
text \<open>
This section is dedicated to the proof of the invariance of predicate @{text add_inv}, defined here
below, over inductive set @{const gcsort_set}. This invariant will later be used to prove GCsort's
correctness properties.
Another predicate, @{text bn_inv}, is also defined, using predicate @{const bn_valid} defined above.
\null
\<close>
fun bn_inv :: "nat \<Rightarrow> nat \<Rightarrow> nat \<times> nat list \<times> 'a list \<Rightarrow> bool" where
"bn_inv p q (u, ns, xs) =
(\<forall>n \<in> set ns. case n of Suc (Suc m) \<Rightarrow> bn_valid m p q | _ \<Rightarrow> True)"
fun add_inv :: "nat \<Rightarrow> nat \<times> nat list \<times> 'a list \<Rightarrow> bool" where
"add_inv n (u, ns, xs) =
(foldl (+) 0 ns = n \<and> length xs = n)"
lemma gcsort_add_input:
"add_inv (length xs) (0, [length xs], xs)"
by simp
lemma add_base:
"foldl (+) (k + m) ns = foldl (+) m ns + (k :: nat)"
by (induction ns arbitrary: m, simp_all, subst add.assoc, simp)
lemma add_base_zero:
"foldl (+) k ns = foldl (+) 0 ns + (k :: nat)"
by (insert add_base [of k 0 ns], simp)
lemma bn_count_le:
"bn_count ns \<le> foldl (+) 0 ns"
by (induction ns rule: bn_count.induct, simp_all add: add_suc, subst add_base_zero,
simp)
text \<open>
\null
Here below is the proof of the main property of predicate @{const bn_inv}, which states that if the
objects' number is not larger than the counters' upper bound, then, as long as there are buckets to
be split, the arguments $p$ and $q$ passed by function @{const round} to function @{const bn_comp}
are such that $0 < q \leq p$.
\null
\<close>
lemma bn_inv_intro [rule_format]:
"foldl (+) 0 ns \<le> p \<longrightarrow>
bn_inv (p - bn_count ns) (foldl (+) 0 ns - bn_count ns) (u, ns, xs)"
proof (induction ns, simp_all, (rule impI)+, subst (asm) (3) add_base_zero,
subst (1 2) add_base_zero, simp)
fix n ns
assume
A: "\<forall>x \<in> set ns. case x of 0 \<Rightarrow> True | Suc 0 \<Rightarrow> True | Suc (Suc m) \<Rightarrow>
bn_valid m (p - bn_count ns) (foldl (+) 0 ns - bn_count ns)" and
B: "foldl (+) 0 ns + n \<le> p"
show
"(case n of 0 \<Rightarrow> True | Suc 0 \<Rightarrow> True | Suc (Suc m) \<Rightarrow>
bn_valid m (p - bn_count (n # ns))
(foldl (+) 0 ns + n - bn_count (n # ns))) \<and>
(\<forall>x \<in> set ns. case x of 0 \<Rightarrow> True | Suc 0 \<Rightarrow> True | Suc (Suc m) \<Rightarrow>
bn_valid m (p - bn_count (n # ns))
(foldl (+) 0 ns + n - bn_count (n # ns)))"
+ using [[simproc del: defined_all]]
proof (rule conjI, rule_tac [2] ballI, simp_all split: nat.split, rule_tac [!] allI,
rule_tac [!] impI)
fix m
assume C: "n = Suc (Suc m)"
show "bn_valid m (p - bn_count (Suc (Suc m) # ns))
(Suc (Suc (foldl (+) 0 ns + m)) - bn_count (Suc (Suc m) # ns))"
(is "bn_valid _ ?p ?q")
proof (rule bn_valid.cases [of "(m, ?p, ?q)"], simp_all, erule conjE, rule conjI)
fix k
have "bn_count ns \<le> foldl (+) 0 ns"
by (rule bn_count_le)
thus "bn_count ns < Suc (Suc (foldl (+) 0 ns + k))" by simp
next
fix k
assume "m = Suc (Suc k)"
hence "Suc (Suc (foldl (+) 0 ns + k)) - bn_count ns =
foldl (+) 0 ns + n - Suc (Suc (bn_count ns))"
using C by simp
also have "\<dots> \<le> p - Suc (Suc (bn_count ns))"
using B by simp
finally show "Suc (Suc (foldl (+) 0 ns + k)) - bn_count ns \<le>
p - Suc (Suc (bn_count ns))" .
qed
next
fix n' m
assume "n' \<in> set ns"
with A have "case n' of 0 \<Rightarrow> True | Suc 0 \<Rightarrow> True | Suc (Suc m) \<Rightarrow>
bn_valid m (p - bn_count ns) (foldl (+) 0 ns - bn_count ns)" ..
moreover assume "n' = Suc (Suc m)"
ultimately have "bn_valid m (p - bn_count ns) (foldl (+) 0 ns - bn_count ns)"
by simp
thus "bn_valid m (p - bn_count (n # ns))
(foldl (+) 0 ns + n - bn_count (n # ns))"
(is "bn_valid _ ?p ?q")
proof (rule_tac bn_valid.cases [of "(m, ?p, ?q)"], simp_all, (erule_tac conjE)+,
simp)
fix p' q'
assume "bn_count ns < foldl (+) 0 ns"
moreover assume "p - bn_count (n # ns) = p'"
hence "p' = p - bn_count (n # ns)" ..
moreover assume "foldl (+) 0 ns + n - bn_count (n # ns) = q'"
hence "q' = foldl (+) 0 ns + n - bn_count (n # ns)" ..
ultimately show "0 < q' \<and> q' \<le> p'"
using B by (rule_tac bn_count.cases [of "n # ns"], simp_all)
qed
qed
qed
text \<open>
\null
In what follows, the invariance of predicate @{const add_inv} over inductive set @{const gcsort_set}
is then proven as lemma @{text gcsort_add_inv}. It holds under the conditions that the objects'
number is not larger than the counters' upper bound and function @{text index} satisfies predicate
@{const index_less}, and states that, if the counters' sum initially matches the objects' number,
this is still true after any recursive round.
\null
\<close>
lemma bn_comp_fst_ge [rule_format]:
"bn_valid n p q \<longrightarrow> n \<le> fst (bn_comp n p q r)"
proof (induction n p q r rule: bn_comp.induct, simp_all del: mult_Suc,
rule impI, erule conjE)
fix n p r and q :: nat
assume "0 < q"
hence "Suc (Suc n) = Suc (Suc n) * q div q" by simp
also assume "q \<le> p"
hence "Suc (Suc n) * q \<le> Suc (Suc n) * p"
by (rule mult_le_mono2)
hence "Suc (Suc n) * q div q \<le> (Suc (Suc n) * p + r) div q"
by (rule_tac div_le_mono, simp)
finally show "Suc (Suc n) \<le> (Suc (Suc n) * p + r) div q" .
qed
lemma bn_comp_fst_nonzero:
"bn_valid n p q \<Longrightarrow> 0 < n \<Longrightarrow> 0 < fst (bn_comp n p q r)"
by (drule bn_comp_fst_ge [where r = r], simp)
lemma bn_comp_snd_less:
"r < q \<Longrightarrow> snd (bn_comp n p q r) < q"
by (induction n p q r rule: bn_comp.induct, simp_all)
lemma add_replicate:
"foldl (+) k (replicate m n) = k + m * n"
by (induction m arbitrary: k, simp_all)
lemma fill_length:
"length (fill xs ns index key n mi ma) = n"
by (induction xs arbitrary: ns, simp_all add: Let_def)
lemma enum_add [rule_format]:
assumes
A: "index_less index key" and
B: "0 < n"
shows "(\<forall>x \<in> set xs. key x \<in> {mi..ma}) \<longrightarrow>
foldl (+) 0 (enum xs index key n mi ma) = length xs"
proof (induction xs, simp_all add: Let_def add_zeros, rule impI, (erule conjE)+,
simp)
fix x xs
assume "mi \<le> key x" and "key x \<le> ma"
hence "index key x n mi ma < n"
(is "?i < _")
using A and B by (simp add: index_less_def)
hence "?i < length (enum xs index key n mi ma)"
(is "_ < length ?ns")
by (simp add: enum_length)
hence "foldl (+) 0 (?ns[?i := Suc (?ns ! ?i)]) = Suc (foldl (+) 0 ?ns)"
by (rule add_update)
moreover assume "foldl (+) 0 ?ns = length xs"
ultimately show "foldl (+) 0 (?ns[?i := Suc (?ns ! ?i)]) = Suc (length xs)"
by simp
qed
lemma round_add_inv [rule_format]:
"index_less index key \<longrightarrow> bn_inv p q t \<longrightarrow> add_inv n t \<longrightarrow>
add_inv n (round index key p q r t)"
+using [[simproc del: defined_all]]
proof (induction index key p q r t arbitrary: n rule: round.induct, simp_all
add: Let_def split: prod.split, (rule allI)+, (rule impI)+, erule conjE,
(rule_tac [2] allI)+, (rule_tac [2] impI)+, (erule_tac [2] conjE)+,
rule_tac [2] ssubst [OF add_base_zero], simp_all add: add_suc)
fix n ns ns' and xs' :: "'a list"
assume "\<And>n'. foldl (+) 0 ns = n' \<and> n - Suc 0 = n' \<longrightarrow>
foldl (+) 0 ns' = n' \<and> length xs' = n'"
hence "foldl (+) 0 ns = n - Suc 0 \<longrightarrow>
foldl (+) 0 ns' = n - Suc 0 \<and> length xs' = n - Suc 0"
by simp
moreover assume "Suc (foldl (+) 0 ns) = n"
ultimately show "Suc (foldl (+) 0 ns') = n \<and> Suc (length xs') = n" by simp
next
fix index p q r u m m' ns v ms' ws' ns' n
and key :: "'a \<Rightarrow> 'b" and xs :: "'a list" and xs' :: "'a list" and r' :: nat
let ?ws = "take (Suc (Suc m)) xs"
assume
A: "round_suc_suc index key ?ws m m' u = (v, ms', ws')" and
B: "bn_comp m p q r = (m', r')" and
C: "index_less index key" and
D: "bn_valid m p q" and
E: "length xs = n"
assume "\<And>ws a b c d e f g h n'.
ws = ?ws \<Longrightarrow> a = (m', r') \<Longrightarrow> b = m' \<and> c = r' \<Longrightarrow>
d = (v, ms', ws') \<Longrightarrow> e = v \<and> f = (ms', ws') \<Longrightarrow> g = ms' \<and> h = ws' \<Longrightarrow>
foldl (+) 0 ns = n' \<and> n - Suc (Suc m) = n' \<longrightarrow>
foldl (+) 0 ns' = n' \<and> length xs' = n'"
moreover assume "Suc (Suc (foldl (+) m ns)) = n"
hence F: "foldl (+) 0 ns + Suc (Suc m) = n"
by (subst (asm) add_base_zero, simp)
ultimately have
G: "foldl (+) 0 ns' = n - Suc (Suc m) \<and> length xs' = n - Suc (Suc m)"
by simp
from A show "foldl (+) 0 ns' + foldl (+) 0 ms' = n \<and>
length ws' + length xs' = n"
proof (subst (2) add_base_zero, simp add: round_suc_suc_def Let_def split:
if_split_asm, (erule_tac [!] conjE)+, simp_all)
assume "Suc 0 # Suc 0 # replicate m (Suc 0) = ms'"
hence "ms' = Suc 0 # Suc 0 # replicate m (Suc 0)" ..
hence "foldl (+) 0 ms' = Suc (Suc m)"
by (simp add: add_replicate)
hence "foldl (+) 0 ns' + foldl (+) 0 ms' = n"
using F and G by simp
moreover assume "?ws = ws'"
hence "ws' = ?ws" ..
hence "length ws' = Suc (Suc m)"
using F and E by simp
hence "length ws' + length xs' = n"
using F and G by simp
ultimately show ?thesis ..
next
let ?nmi = "mini ?ws key"
let ?nma = "maxi ?ws key"
let ?xmi = "?ws ! ?nmi"
let ?xma = "?ws ! ?nma"
let ?mi = "key ?xmi"
let ?ma = "key ?xma"
let ?k = "case m of 0 \<Rightarrow> m | Suc 0 \<Rightarrow> m | Suc (Suc i) \<Rightarrow> u + m'"
let ?zs = "nths ?ws (- {?nmi, ?nma})"
let ?ms = "enum ?zs index key ?k ?mi ?ma"
assume "Suc 0 # ?ms @ [Suc 0] = ms'"
hence "ms' = Suc 0 # ?ms @ [Suc 0]" ..
moreover assume
"?xmi # map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma) @ [?xma] = ws'"
hence "ws' = ?xmi # map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma)
@ [?xma]" ..
ultimately show ?thesis
proof (simp add: fill_length, subst (2) add_base_zero, simp, cases m)
case 0
moreover from this have "length ?ms = 0"
by (simp add: enum_length)
ultimately show "Suc (Suc (foldl (+) 0 ns' + foldl (+) 0 ?ms)) = n \<and>
Suc (Suc (m + length xs')) = n"
using F and G by simp
next
case Suc
moreover from this have "0 < fst (bn_comp m p q r)"
by (rule_tac bn_comp_fst_nonzero [OF D], simp)
hence "0 < m'"
using B by simp
ultimately have H: "0 < ?k"
by (simp split: nat.split)
have "foldl (+) 0 ?ms = length ?zs"
by (rule enum_add [OF C H], simp, rule conjI,
((rule mini_lb | rule maxi_ub), erule in_set_nthsD)+)
moreover have "length ?ws = Suc (Suc m)"
using F and E by simp
hence "length ?zs = m"
by (simp add: mini_maxi_nths)
ultimately show "Suc (Suc (foldl (+) 0 ns' + foldl (+) 0 ?ms)) = n \<and>
Suc (Suc (m + length xs')) = n"
using F and G by simp
qed
qed
qed
lemma gcsort_add_inv:
assumes A: "index_less index key"
shows "\<lbrakk>t' \<in> gcsort_set index key p t; add_inv n t; n \<le> p\<rbrakk> \<Longrightarrow>
add_inv n t'"
by (erule gcsort_set.induct, simp, rule round_add_inv [OF A], simp_all del:
bn_inv.simps, erule conjE, frule sym, erule subst, rule bn_inv_intro, simp)
subsection "Proof of counters' optimization"
text \<open>
In this section, it is formally proven that the number of the counters (and then of the buckets as
well) used in each recursive round is maximized never exceeding the fixed upper bound.
This property is formalized by theorem @{text round_len}, which holds under the condition that the
objects' number is not larger than the counters' upper bound and states what follows:
\begin{itemize}
\item
While there is some bucket with size larger than two, the sum of the number of the used counters and
the number of the unused ones -- viz. those, if any, left unused due to the presence of some bucket
with size larger than two and equal minimum and maximum keys (cf. section \ref{SEC4}) -- matches the
counters' upper bound.
\\In addition to ensuring the upper bound's enforcement, this implies that the number of the used
counters matches the upper bound unless there is some aforesaid bucket not followed by any other
bucket with size larger than two and distinct minimum and maximum keys.
\item
Once there is no bucket with size larger than two -- in which case a round is executed just in case
there is some bucket with size two --, the number of the used counters matches the objects' number.
\\In fact, the algorithm immediately terminates after such a round since every resulting bucket has
size one, so that increasing the number of the used counters does not matter in this case.
\end{itemize}
\null
\<close>
lemma round_len_less [rule_format]:
"bn_inv p q t \<longrightarrow> r < q \<longrightarrow>
(r + (foldl (+) 0 (fst (snd t)) - bn_count (fst (snd t))) * p) mod q = 0 \<longrightarrow>
(fst (round index key p q r t) +
length (fst (snd (round index key p q r t)))) * q =
(fst t + bn_count (fst (snd t))) * q +
(foldl (+) 0 (fst (snd t)) - bn_count (fst (snd t))) * p + r"
+using [[simproc del: defined_all]]
proof (induction index key p q r t rule: round.induct, simp_all add: Let_def
split: prod.split del: all_simps, ((rule allI)+, (rule impI)+, simp add:
add_suc)+, subst (asm) (3) add_base_zero, subst add_base_zero, erule conjE)
fix index p q r u n ns n' r' v ms' ws' u'
and key :: "'a \<Rightarrow> 'b" and xs :: "'a list" and ns' :: "nat list"
let ?ws = "take (Suc (Suc n)) xs"
assume
A: "round_suc_suc index key ?ws n n' u = (v, ms', ws')" and
B: "bn_comp n p q r = (n', r')" and
C: "bn_valid n p q"
have D: "bn_count ns \<le> foldl (+) 0 ns"
by (rule bn_count_le)
assume "\<And>ws a b c d e f g h.
ws = ?ws \<Longrightarrow> a = (n', r') \<Longrightarrow> b = n' \<and> c = r' \<Longrightarrow>
d = (v, ms', ws') \<Longrightarrow> e = v \<and> f = (ms', ws') \<Longrightarrow> g = ms' \<and> h = ws' \<Longrightarrow>
r' < q \<longrightarrow> (r' + (foldl (+) 0 ns - bn_count ns) * p) mod q = 0 \<longrightarrow>
(u' + length ns') * q =
(v + bn_count ns) * q + (foldl (+) 0 ns - bn_count ns) * p + r'"
moreover assume "r < q"
hence "snd (bn_comp n p q r) < q"
by (rule bn_comp_snd_less)
hence "r' < q"
using B by simp
moreover assume E: "(r + (Suc (Suc (foldl (+) 0 ns + n)) -
bn_count (Suc (Suc n) # ns)) * p) mod q = 0"
from B [symmetric] have "(r' + (foldl (+) 0 ns - bn_count ns) * p) mod q = 0"
proof (rule_tac trans [OF _ E], rule_tac bn_comp.cases [of "(n, p, q, r)"],
simp_all add: add_mult_distrib diff_mult_distrib mod_add_left_eq,
rule_tac arg_cong2 [where f = "(mod)"], simp_all)
fix n p q r
have "bn_count ns * p \<le> foldl (+) 0 ns * p"
using D by (rule mult_le_mono1)
thus "p + (p + (n * p + (foldl (+) 0 ns * p - bn_count ns * p))) =
p + (p + (foldl (+) 0 ns * p + n * p)) - bn_count ns * p"
by arith
qed
ultimately have "(u' + length ns') * q =
(v + bn_count ns) * q + (foldl (+) 0 ns - bn_count ns) * p + r'"
by simp
with A [symmetric] and B [symmetric] show
"(u' + (length ms' + length ns')) * q =
(u + bn_count (Suc (Suc n) # ns)) * q +
(Suc (Suc (foldl (+) 0 ns + n)) - bn_count (Suc (Suc n) # ns)) * p + r"
proof (rule_tac bn_comp.cases [of "(n, p, q, r)"],
simp_all add: round_suc_suc_def Let_def enum_length split: if_split_asm)
fix m p' q' r'
assume
"n = Suc (Suc m)" and
"p = p'" and
"q = q'" and
"r = r'"
moreover have "n \<le> fst (bn_comp n p q r)"
using C by (rule bn_comp_fst_ge)
ultimately have "Suc (Suc m) \<le> (p' + (p' + m * p') + r') div q'"
(is "_ \<le> ?a div _")
by simp
hence F: "Suc (Suc m) * q' \<le> ?a div q' * q'"
by (rule mult_le_mono1)
moreover assume "(u' + length ns') * q' =
(u + ?a div q' - Suc (Suc m) + bn_count ns) * q' +
(foldl (+) 0 ns - bn_count ns) * p' + ?a mod q'"
ultimately have "(u' + length ns') * q' + Suc (Suc m) * q' =
(u + bn_count ns) * q' + (foldl (+) 0 ns - bn_count ns) * p' +
?a div q' * q' + ?a mod q'"
(is "?c = ?d")
proof (simp add: add_mult_distrib diff_mult_distrib)
have "Suc (Suc m) * q' \<le> ?a div q' * q' + ?a mod q'"
using F by arith
hence "q' + (q' + m * q') \<le> ?a"
(is "?b \<le> _")
by simp
thus
"?a + (u * q' + (bn_count ns * q' +
(foldl (+) 0 ns * p' - bn_count ns * p'))) - ?b + ?b =
?a + (u * q' + (bn_count ns * q' +
(foldl (+) 0 ns * p' - bn_count ns * p')))"
by simp
qed
moreover have "?c = q' + (q' + (u' + (m + length ns')) * q')"
(is "_ = ?e")
by (simp add: add_mult_distrib)
moreover have "bn_count ns * p' \<le> foldl (+) 0 ns * p'"
using D by (rule mult_le_mono1)
hence "?d = (u + bn_count ns) * q' +
((Suc (Suc (foldl (+) 0 ns + m)) - bn_count ns) * p' + r')"
(is "_ = ?f")
by (simp (no_asm_simp) add: add_mult_distrib diff_mult_distrib)
ultimately show "?e = ?f" by simp
next
fix m p' q' r'
assume "(u' + length ns') * q' = bn_count ns * q' +
(foldl (+) 0 ns - bn_count ns) * p' + (p' + (p' + m * p') + r') mod q'"
(is "_ = _ + _ + ?a mod _")
hence "(u' + length ns') * q' + (u + ?a div q') * q' =
(u + bn_count ns) * q' + (foldl (+) 0 ns - bn_count ns) * p' + ?a"
(is "?c = ?d")
by (simp add: add_mult_distrib)
moreover have "?c = (u' + (u + ?a div q' + length ns')) * q'"
(is "_ = ?e")
by (simp add: add_mult_distrib)
moreover have "bn_count ns * p' \<le> foldl (+) 0 ns * p'"
using D by (rule mult_le_mono1)
hence "?d = (u + bn_count ns) * q' +
((Suc (Suc (foldl (+) 0 ns + m)) - bn_count ns) * p' + r')"
(is "_ = ?f")
by (simp (no_asm_simp) add: add_mult_distrib diff_mult_distrib)
ultimately show "?e = ?f" by simp
qed
qed
lemma round_len_eq [rule_format]:
"bn_count (fst (snd t)) = foldl (+) 0 (fst (snd t)) \<longrightarrow>
length (fst (snd (round index key p q r t))) = foldl (+) 0 (fst (snd t))"
+using [[simproc del: defined_all]]
proof (induction index key p q r t rule: round.induct, simp_all add: Let_def
split: prod.split del: all_simps, ((rule allI)+, (rule impI)+, simp add:
add_suc)+, subst (asm) (3) add_base_zero, subst add_base_zero)
fix index p q r u n ns n' v ms' ws'
and key :: "'a \<Rightarrow> 'b" and xs :: "'a list" and ns' :: "nat list" and r' :: nat
let ?ws = "take (Suc (Suc n)) xs"
assume
A: "round_suc_suc index key ?ws n n' u = (v, ms', ws')" and
B: "bn_count (Suc (Suc n) # ns) = Suc (Suc (foldl (+) 0 ns + n))"
assume "\<And>ws a b c d e f g h.
ws = ?ws \<Longrightarrow> a = (n', r') \<Longrightarrow> b = n' \<and> c = r' \<Longrightarrow>
d = (v, ms', ws') \<Longrightarrow> e = v \<and> f = (ms', ws') \<Longrightarrow> g = ms' \<and> h = ws' \<Longrightarrow>
bn_count ns = foldl (+) 0 ns \<longrightarrow> length ns' = foldl (+) 0 ns"
moreover have C: "n = 0 \<or> n = Suc 0"
using B by (rule_tac bn_comp.cases [of "(n, p, q, r)"],
insert bn_count_le [of ns], simp_all)
hence "bn_count ns = foldl (+) 0 ns"
using B by (erule_tac disjE, simp_all)
ultimately have "length ns' = foldl (+) 0 ns" by simp
with A [symmetric] show "length ms' + length ns' =
Suc (Suc (foldl (+) 0 ns + n))"
by (rule_tac disjE [OF C], simp_all
add: round_suc_suc_def Let_def enum_length split: if_split_asm)
qed
theorem round_len:
assumes
A: "length xs = foldl (+) 0 ns" and
B: "length xs \<le> p"
shows "if bn_count ns < length xs
then fst (gcsort_round index key p ns xs) +
length (fst (snd (gcsort_round index key p ns xs))) = p
else length (fst (snd (gcsort_round index key p ns xs))) = length xs"
(is "if _ then fst ?t + _ = _ else _")
proof (split if_split, rule conjI, rule_tac [!] impI)
assume C: "bn_count ns < length xs"
moreover have
"bn_inv (p - bn_count ns) (foldl (+) 0 ns - bn_count ns) (0, ns, xs)"
using A and B by (rule_tac bn_inv_intro, simp)
ultimately have
"(fst ?t + length (fst (snd ?t))) * (length xs - bn_count ns) =
bn_count ns * (length xs - bn_count ns) +
(p - bn_count ns) * (length xs - bn_count ns)"
(is "?a * ?b = ?c")
using A by (subst round_len_less, simp_all)
also have "bn_count ns \<le> p"
using B and C by simp
hence "bn_count ns * ?b \<le> p * ?b"
by (rule mult_le_mono1)
hence "?c = p * ?b"
by (simp (no_asm_simp) add: diff_mult_distrib)
finally have "?a * ?b = p * ?b" .
thus "?a = p"
using C by simp
next
assume "\<not> bn_count ns < length xs"
moreover have "bn_count ns \<le> foldl (+) 0 ns"
by (rule bn_count_le)
ultimately show "length (fst (snd ?t)) = length xs"
using A by (subst round_len_eq, simp_all)
qed
end
diff --git a/thys/Generalized_Counting_Sort/Conservation.thy b/thys/Generalized_Counting_Sort/Conservation.thy
--- a/thys/Generalized_Counting_Sort/Conservation.thy
+++ b/thys/Generalized_Counting_Sort/Conservation.thy
@@ -1,2036 +1,2037 @@
(* Title: An Efficient Generalization of Counting Sort for Large, possibly Infinite Key Ranges
Author: Pasquale Noce
Software Engineer at HID Global, Italy
pasquale dot noce dot lavoro at gmail dot com
pasquale dot noce at hidglobal dot com
*)
section "Proof of objects' conservation"
theory Conservation
imports
Algorithm
"HOL-Library.Multiset"
begin
text \<open>
\null
In this section, it is formally proven that GCsort \emph{conserves objects}, viz. that the objects'
list returned by function @{const gcsort} contains as many occurrences of any given object as the
input objects' list.
Here below, steps 5, 6, and 7 of the proof method are accomplished. Particularly, @{text count_inv}
is the predicate that will be shown to be invariant over inductive set @{const gcsort_set}.
\null
\<close>
fun count_inv :: "('a \<Rightarrow> nat) \<Rightarrow> nat \<times> nat list \<times> 'a list \<Rightarrow> bool" where
"count_inv f (u, ns, xs) = (\<forall>x. count (mset xs) x = f x)"
lemma gcsort_count_input:
"count_inv (count (mset xs)) (0, [length xs], xs)"
by simp
lemma gcsort_count_intro:
"count_inv f t \<Longrightarrow> count (mset (gcsort_out t)) x = f x"
by (cases t, simp add: gcsort_out_def)
text \<open>
\null
The main task to be accomplished to prove that GCsort conserves objects is to prove that so does
function @{const fill} in case its input offsets' list is computed via the composition of functions
@{const offs} and @{const enum}, as happens within function @{const round}.
To achieve this result, a multi-step strategy will be adopted. The first step, addressed here below,
opens with the definition of predicate @{text offs_pred}, satisfied by an offsets' list $ns$ and an
objects' list $xs$ just in case each bucket delimited by $ns$ is sufficiently large to accommodate
the corresponding objects in $xs$. Then, lemma @{text offs_pred_cons} shows that this predicate, if
satisfied initially, keeps being true if each object in $xs$ is consumed as happens within function
@{const fill}, viz. increasing the corresponding offset in $ns$ by one.
\null
\<close>
definition offs_num :: "nat \<Rightarrow> 'a list \<Rightarrow> ('a, 'b) index_sign \<Rightarrow>
('a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> nat \<Rightarrow> nat" where
"offs_num n xs index key mi ma i \<equiv>
length [x\<leftarrow>xs. index key x n mi ma = i]"
abbreviation offs_set_next :: "nat list \<Rightarrow> 'a list \<Rightarrow> ('a, 'b) index_sign \<Rightarrow>
('a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> nat \<Rightarrow> nat set" where
"offs_set_next ns xs index key mi ma i \<equiv>
{k. k < length ns \<and> i < k \<and> 0 < offs_num (length ns) xs index key mi ma k}"
abbreviation offs_set_prev :: "nat list \<Rightarrow> 'a list \<Rightarrow> ('a, 'b) index_sign \<Rightarrow>
('a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> nat \<Rightarrow> nat set" where
"offs_set_prev ns xs index key mi ma i \<equiv>
{k. i < length ns \<and> k < i \<and> 0 < offs_num (length ns) xs index key mi ma k}"
definition offs_next :: "nat list \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> ('a, 'b) index_sign \<Rightarrow>
('a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> nat \<Rightarrow> nat" where
"offs_next ns ub xs index key mi ma i \<equiv>
if offs_set_next ns xs index key mi ma i = {}
then ub else ns ! Min (offs_set_next ns xs index key mi ma i)"
definition offs_none :: "nat list \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> ('a, 'b) index_sign \<Rightarrow>
('a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> nat \<Rightarrow> bool" where
"offs_none ns ub xs index key mi ma i \<equiv>
(\<exists>j < length ns. 0 < offs_num (length ns) xs index key mi ma j \<and>
i \<in> {ns ! j + offs_num (length ns) xs index key mi ma j..<
offs_next ns ub xs index key mi ma j}) \<or>
offs_num (length ns) xs index key mi ma 0 = 0 \<and>
i < offs_next ns ub xs index key mi ma 0 \<or>
0 < offs_num (length ns) xs index key mi ma 0 \<and>
i < ns ! 0"
definition offs_pred :: "nat list \<Rightarrow> nat \<Rightarrow> 'a list \<Rightarrow> ('a, 'b) index_sign \<Rightarrow>
('a \<Rightarrow> 'b) \<Rightarrow> 'b \<Rightarrow> 'b \<Rightarrow> bool" where
"offs_pred ns ub xs index key mi ma \<equiv>
\<forall>i < length ns. offs_num (length ns) xs index key mi ma i \<le>
offs_next ns ub xs index key mi ma i - ns ! i"
lemma offs_num_cons:
"offs_num n (x # xs) index key mi ma i =
(if index key x n mi ma = i then Suc else id) (offs_num n xs index key mi ma i)"
by (simp add: offs_num_def)
lemma offs_next_prev:
"(0 < offs_num (length ns) xs index key mi ma i \<and>
offs_set_next ns xs index key mi ma i \<noteq> {} \<and>
Min (offs_set_next ns xs index key mi ma i) = j) =
(0 < offs_num (length ns) xs index key mi ma j \<and>
offs_set_prev ns xs index key mi ma j \<noteq> {} \<and>
Max (offs_set_prev ns xs index key mi ma j) = i)"
(is "?P = ?Q")
proof (rule iffI, (erule_tac [!] conjE)+)
let ?A = "offs_set_next ns xs index key mi ma i"
let ?B = "offs_set_prev ns xs index key mi ma j"
assume
A: "0 < offs_num (length ns) xs index key mi ma i" and
B: "?A \<noteq> {}" and
C: "Min ?A = j"
have "Min ?A \<in> ?A"
using B by (rule_tac Min_in, simp)
hence D: "j \<in> ?A"
using C by simp
hence E: "i \<in> ?B"
using A by simp
show ?Q
proof (rule conjI, rule_tac [2] conjI)
show "0 < offs_num (length ns) xs index key mi ma j"
using D by simp
next
show "?B \<noteq> {}"
using E by blast
next
from E show "Max ?B = i"
proof (subst Max_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI,
(erule_tac conjE)+, rule_tac ccontr, simp)
fix k
assume F: "k < j" and "j < length ns"
hence "k < length ns" by simp
moreover assume
"\<not> k \<le> i" and
"0 < offs_num (length ns) xs index key mi ma k"
ultimately have "k \<in> ?A" by simp
hence "Min ?A \<le> k"
by (rule_tac Min_le, simp)
thus False
using C and F by simp
qed
qed
next
let ?A = "offs_set_prev ns xs index key mi ma j"
let ?B = "offs_set_next ns xs index key mi ma i"
assume
A: "0 < offs_num (length ns) xs index key mi ma j" and
B: "?A \<noteq> {}" and
C: "Max ?A = i"
have "Max ?A \<in> ?A"
using B by (rule_tac Max_in, simp)
hence D: "i \<in> ?A"
using C by simp
hence E: "j \<in> ?B"
using A by simp
show ?P
proof (rule conjI, rule_tac [2] conjI)
show "0 < offs_num (length ns) xs index key mi ma i"
using D by simp
next
show "?B \<noteq> {}"
using E by blast
next
from E show "Min ?B = j"
proof (subst Min_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI,
(erule_tac conjE)+, rule_tac ccontr, simp)
fix k
assume
"j < length ns" and
"\<not> j \<le> k" and
"0 < offs_num (length ns) xs index key mi ma k"
hence "k \<in> ?A" by simp
hence "k \<le> Max ?A"
by (rule_tac Max_ge, simp)
moreover assume "i < k"
ultimately show False
using C by simp
qed
qed
qed
lemma offs_next_cons_eq:
assumes
A: "index key x (length ns) mi ma = i" and
B: "i < length ns" and
C: "0 < offs_num (length ns) (x # xs) index key mi ma j"
shows
"offs_set_prev ns (x # xs) index key mi ma i = {} \<or>
Max (offs_set_prev ns (x # xs) index key mi ma i) \<noteq> j \<Longrightarrow>
offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j =
offs_next ns ub (x # xs) index key mi ma j"
(is "?P \<or> ?Q \<Longrightarrow> _")
proof (simp only: disj_imp, cases "j < i")
let ?A = "offs_set_prev ns (x # xs) index key mi ma i"
let ?B = "offs_set_next (ns[i := Suc (ns ! i)]) xs index key mi ma j"
let ?C = "offs_set_next ns (x # xs) index key mi ma j"
case True
hence D: "j \<in> ?A"
using B and C by simp
hence "j \<le> Max ?A"
by (rule_tac Max_ge, simp)
moreover assume "\<not> ?P \<longrightarrow> ?Q"
hence ?Q
using D by blast
ultimately have E: "j < Max ?A" by simp
have F: "Max ?A \<in> ?A"
using D by (rule_tac Max_in, simp, blast)
have G: "Max ?A \<in> ?B"
proof (simp, rule conjI, rule_tac [2] conjI)
show "Max ?A < length ns"
using F by auto
next
show "j < Max ?A"
using E .
next
have "0 < offs_num (length ns) (x # xs) index key mi ma (Max ?A)"
using F by blast
moreover have "Max ?A < i"
using F by blast
ultimately show "0 < offs_num (length ns) xs index key mi ma (Max ?A)"
using A by (simp add: offs_num_cons)
qed
hence H: "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j =
ns[i := Suc (ns ! i)] ! Min ?B"
by (simp only: offs_next_def split: if_split, blast)
have "Min ?B \<le> Max ?A"
using G by (rule_tac Min_le, simp)
moreover have "Max ?A < i"
using F by blast
ultimately have I: "Min ?B < i" by simp
hence J: "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j =
ns ! Min ?B"
using H by simp
have "Min ?B \<in> ?B"
using G by (rule_tac Min_in, simp, blast)
hence K: "Min ?B \<in> ?C"
using A and I by (simp add: offs_num_cons)
hence L: "Min ?C \<le> Min ?B"
by (rule_tac Min_le, simp)
have "Min ?C \<in> ?C"
using K by (rule_tac Min_in, simp, blast)
moreover have "Min ?C < i"
using L and I by simp
ultimately have "Min ?C \<in> ?B"
using A by (simp add: offs_num_cons)
hence "Min ?B \<le> Min ?C"
using G by (rule_tac Min_le, simp)
hence "Min ?B = Min ?C"
using L by simp
moreover have "offs_next ns ub (x # xs) index key mi ma j = ns ! Min ?C"
using K by (simp only: offs_next_def split: if_split, blast)
ultimately show ?thesis
using J by simp
next
let ?A = "offs_set_next (ns[i := Suc (ns ! i)]) xs index key mi ma j"
let ?B = "offs_set_next ns (x # xs) index key mi ma j"
case False
hence "?A = ?B"
using A by (rule_tac set_eqI, simp add: offs_num_cons)
thus ?thesis
proof (simp only: offs_next_def split: if_split,
(rule_tac conjI, blast, rule_tac impI)+)
assume "?B \<noteq> {}"
hence "Min ?B \<in> ?B"
by (rule_tac Min_in, simp)
hence "i < Min ?B"
using False by simp
thus "ns[i := Suc (ns ! i)] ! Min ?B = ns ! Min ?B" by simp
qed
qed
lemma offs_next_cons_neq:
assumes
A: "index key x (length ns) mi ma = i" and
B: "offs_set_prev ns (x # xs) index key mi ma i \<noteq> {}" and
C: "Max (offs_set_prev ns (x # xs) index key mi ma i) = j"
shows "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j =
(if 0 < offs_num (length ns) xs index key mi ma i
then Suc (ns ! i)
else offs_next ns ub (x # xs) index key mi ma i)"
proof (simp, rule conjI, rule_tac [!] impI)
let ?A = "offs_set_next ns (x # xs) index key mi ma j"
assume "0 < offs_num (length ns) xs index key mi ma i"
with A have "offs_set_next (ns[i := Suc (ns ! i)]) xs index key mi ma j = ?A"
by (rule_tac set_eqI, rule_tac iffI, simp_all add: offs_num_cons split:
if_split_asm)
moreover have "0 < offs_num (length ns) (x # xs) index key mi ma i"
using A by (simp add: offs_num_def)
hence "0 < offs_num (length ns) (x # xs) index key mi ma j \<and> ?A \<noteq> {} \<and>
Min ?A = i"
using B and C by (subst offs_next_prev, simp)
ultimately show "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j =
Suc (ns ! i)"
using B by (simp only: offs_next_def, simp, subst nth_list_update_eq, blast,
simp)
next
let ?A = "offs_set_prev ns (x # xs) index key mi ma i"
assume "offs_num (length ns) xs index key mi ma i = 0"
with A have "offs_set_next (ns[i := Suc (ns ! i)]) xs index key mi ma j =
offs_set_next ns (x # xs) index key mi ma i"
proof (rule_tac set_eqI, rule_tac iffI, simp_all add: offs_num_cons split:
if_split_asm, rule_tac conjI, rule_tac notI, simp, rule_tac impI,
(erule_tac [!] conjE)+, rule_tac ccontr, simp_all)
fix k
have "i < length ns"
using B by blast
moreover assume "i \<noteq> k" and "\<not> i < k"
hence "k < i" by simp
moreover assume "0 < offs_num (length ns) xs index key mi ma k"
ultimately have "k \<in> ?A"
by (simp add: offs_num_cons)
hence "k \<le> Max ?A"
by (rule_tac Max_ge, simp)
moreover assume "j < k"
ultimately show False
using C by simp
next
fix k
have "Max ?A \<in> ?A"
using B by (rule_tac Max_in, simp_all)
hence "j < i"
using C by simp
moreover assume "i < k"
ultimately show "j < k" by simp
qed
with B show "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j =
offs_next ns ub (x # xs) index key mi ma i"
proof (simp only: offs_next_def split: if_split, (rule_tac conjI, rule_tac [!] impI,
simp)+, subst nth_list_update, blast, simp (no_asm_simp))
assume "offs_set_next ns (x # xs) index key mi ma i \<noteq> {}"
hence "Min (offs_set_next ns (x # xs) index key mi ma i)
\<in> offs_set_next ns (x # xs) index key mi ma i"
by (rule_tac Min_in, simp)
thus "i \<noteq> Min (offs_set_next ns (x # xs) index key mi ma i)" by simp
qed
qed
lemma offs_pred_ub_aux [rule_format]:
assumes A: "offs_pred ns ub xs index key mi ma"
shows "i < length ns \<Longrightarrow>
\<forall>j < length ns. i \<le> j \<longrightarrow> 0 < offs_num (length ns) xs index key mi ma j \<longrightarrow>
ns ! j + offs_num (length ns) xs index key mi ma j \<le> ub"
proof (erule strict_inc_induct, rule_tac [!] allI, (rule_tac [!] impI)+,
drule le_less_Suc_eq, simp_all)
fix i
assume B: "length ns = Suc i"
hence "offs_num (Suc i) xs index key mi ma i \<le>
offs_next ns ub xs index key mi ma i - ns ! i"
using A by (simp add: offs_pred_def)
moreover have "offs_next ns ub xs index key mi ma i = ub"
using B by (simp add: offs_next_def)
ultimately have "offs_num (Suc i) xs index key mi ma i \<le> ub - ns ! i" by simp
moreover assume "0 < offs_num (Suc i) xs index key mi ma i"
ultimately show "ns ! i + offs_num (Suc i) xs index key mi ma i \<le> ub" by simp
next
fix i j
assume "j < length ns"
hence "offs_num (length ns) xs index key mi ma j \<le>
offs_next ns ub xs index key mi ma j - ns ! j"
using A by (simp add: offs_pred_def)
moreover assume
B: "\<forall>k < length ns. Suc i \<le> k \<longrightarrow>
0 < offs_num (length ns) xs index key mi ma k \<longrightarrow>
ns ! k + offs_num (length ns) xs index key mi ma k \<le> ub" and
C: "i \<le> j"
have "offs_next ns ub xs index key mi ma j \<le> ub"
proof (simp only: offs_next_def split: if_split, rule conjI, simp, rule impI)
let ?A = "offs_set_next ns xs index key mi ma j"
assume "?A \<noteq> {}"
hence "Min ?A \<in> ?A"
by (rule_tac Min_in, simp)
hence "ns ! Min ?A + offs_num (length ns) xs index key mi ma (Min ?A) \<le> ub"
using B and C by simp
thus "ns ! Min ?A \<le> ub" by simp
qed
ultimately have "offs_num (length ns) xs index key mi ma j \<le> ub - ns ! j"
by simp
moreover assume "0 < offs_num (length ns) xs index key mi ma j"
ultimately show "ns ! j + offs_num (length ns) xs index key mi ma j \<le> ub"
by simp
qed
lemma offs_pred_ub:
"\<lbrakk>offs_pred ns ub xs index key mi ma; i < length ns;
0 < offs_num (length ns) xs index key mi ma i\<rbrakk> \<Longrightarrow>
ns ! i + offs_num (length ns) xs index key mi ma i \<le> ub"
by (drule offs_pred_ub_aux, assumption+, simp)
lemma offs_pred_asc_aux [rule_format]:
assumes A: "offs_pred ns ub xs index key mi ma"
shows "i < length ns \<Longrightarrow>
\<forall>j k. k < length ns \<longrightarrow> i \<le> j \<longrightarrow> j < k \<longrightarrow>
0 < offs_num (length ns) xs index key mi ma j \<longrightarrow>
0 < offs_num (length ns) xs index key mi ma k \<longrightarrow>
ns ! j + offs_num (length ns) xs index key mi ma j \<le> ns ! k"
proof (erule strict_inc_induct, simp, (rule allI)+, (rule impI)+, simp)
fix i j k
assume
B: "k < length ns" and
C: "j < k"
hence "j < length ns" by simp
hence "offs_num (length ns) xs index key mi ma j \<le>
offs_next ns ub xs index key mi ma j - ns ! j"
using A by (simp add: offs_pred_def)
moreover assume
D: "\<forall>j k. k < length ns \<longrightarrow> Suc i \<le> j \<longrightarrow> j < k \<longrightarrow>
0 < offs_num (length ns) xs index key mi ma j \<longrightarrow>
0 < offs_num (length ns) xs index key mi ma k \<longrightarrow>
ns ! j + offs_num (length ns) xs index key mi ma j \<le> ns ! k" and
E: "i \<le> j" and
F: "0 < offs_num (length ns) xs index key mi ma k"
have "offs_next ns ub xs index key mi ma j \<le> ns ! k"
proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI,
rule ccontr, simp)
assume "\<forall>n > j. n < length ns \<longrightarrow>
offs_num (length ns) xs index key mi ma n = 0"
hence "offs_num (length ns) xs index key mi ma k = 0"
using B and C by simp
thus False
using F by simp
next
let ?A = "offs_set_next ns xs index key mi ma j"
have G: "k \<in> ?A"
using B and C and F by simp
hence "Min ?A \<le> k"
by (rule_tac Min_le, simp)
hence "Min ?A < k \<or> Min ?A = k"
by (simp add: le_less)
moreover {
have "Suc i \<le> Min ?A \<longrightarrow> Min ?A < k \<longrightarrow>
0 < offs_num (length ns) xs index key mi ma (Min ?A) \<longrightarrow>
ns ! Min ?A + offs_num (length ns) xs index key mi ma (Min ?A) \<le> ns ! k"
using B and D and F by simp
moreover assume "Min ?A < k"
moreover have "Min ?A \<in> ?A"
using G by (rule_tac Min_in, simp, blast)
ultimately have "ns ! Min ?A < ns ! k"
using E by simp
}
moreover {
assume "Min ?A = k"
hence "ns ! Min ?A = ns ! k" by simp
}
ultimately show "ns ! Min ?A \<le> ns ! k"
by (simp add: le_less, blast)
qed
ultimately have "offs_num (length ns) xs index key mi ma j \<le> ns ! k - ns ! j"
by simp
moreover assume "0 < offs_num (length ns) xs index key mi ma j"
ultimately show "ns ! j + offs_num (length ns) xs index key mi ma j \<le> ns ! k"
by simp
qed
lemma offs_pred_asc:
"\<lbrakk>offs_pred ns ub xs index key mi ma; i < j; j < length ns;
0 < offs_num (length ns) xs index key mi ma i;
0 < offs_num (length ns) xs index key mi ma j\<rbrakk> \<Longrightarrow>
ns ! i + offs_num (length ns) xs index key mi ma i \<le> ns ! j"
by (drule offs_pred_asc_aux, erule less_trans, assumption+, rule order_refl)
lemma offs_pred_next:
assumes
A: "offs_pred ns ub xs index key mi ma" and
B: "i < length ns" and
C: "0 < offs_num (length ns) xs index key mi ma i"
shows "ns ! i < offs_next ns ub xs index key mi ma i"
proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI)
have "ns ! i + offs_num (length ns) xs index key mi ma i \<le> ub"
using A and B and C by (rule offs_pred_ub)
thus "ns ! i < ub"
using C by simp
next
assume "offs_set_next ns xs index key mi ma i \<noteq> {}"
hence "Min (offs_set_next ns xs index key mi ma i)
\<in> offs_set_next ns xs index key mi ma i"
by (rule_tac Min_in, simp)
hence "ns ! i + offs_num (length ns) xs index key mi ma i \<le>
ns ! Min (offs_set_next ns xs index key mi ma i)"
using C by (rule_tac offs_pred_asc [OF A], simp_all)
thus "ns ! i < ns ! Min (offs_set_next ns xs index key mi ma i)"
using C by simp
qed
lemma offs_pred_next_cons_less:
assumes
A: "offs_pred ns ub (x # xs) index key mi ma" and
B: "index key x (length ns) mi ma = i" and
C: "offs_set_prev ns (x # xs) index key mi ma i \<noteq> {}" and
D: "Max (offs_set_prev ns (x # xs) index key mi ma i) = j"
shows "offs_next ns ub (x # xs) index key mi ma j <
offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j"
(is "?M < ?N")
proof -
have E: "0 < offs_num (length ns) (x # xs) index key mi ma i"
using B by (simp add: offs_num_cons)
hence "0 < offs_num (length ns) (x # xs) index key mi ma j \<and>
offs_set_next ns (x # xs) index key mi ma j \<noteq> {} \<and>
Min (offs_set_next ns (x # xs) index key mi ma j) = i"
using C and D by (subst offs_next_prev, simp)
hence F: "?M = ns ! i"
by (simp only: offs_next_def, simp)
have "?N = (if 0 < offs_num (length ns) xs index key mi ma i
then Suc (ns ! i)
else offs_next ns ub (x # xs) index key mi ma i)"
using B and C and D by (rule offs_next_cons_neq)
thus ?thesis
proof (split if_split_asm)
assume "?N = Suc (ns ! i)"
thus ?thesis
using F by simp
next
assume "?N = offs_next ns ub (x # xs) index key mi ma i"
moreover have "ns ! i < offs_next ns ub (x # xs) index key mi ma i"
using C by (rule_tac offs_pred_next [OF A _ E], blast)
ultimately show ?thesis
using F by simp
qed
qed
lemma offs_pred_next_cons:
assumes
A: "offs_pred ns ub (x # xs) index key mi ma" and
B: "index key x (length ns) mi ma = i" and
C: "i < length ns" and
D: "0 < offs_num (length ns) (x # xs) index key mi ma j"
shows "offs_next ns ub (x # xs) index key mi ma j \<le>
offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma j"
(is "?M \<le> ?N")
proof -
let ?P = "offs_set_prev ns (x # xs) index key mi ma i \<noteq> {} \<and>
Max (offs_set_prev ns (x # xs) index key mi ma i) = j"
have "?P \<or> \<not> ?P" by blast
moreover {
assume ?P
hence "?M < ?N"
using B by (rule_tac offs_pred_next_cons_less [OF A], simp_all)
hence ?thesis by simp
}
moreover {
assume "\<not> ?P"
hence "?N = ?M"
by (rule_tac offs_next_cons_eq [OF B C D], blast)
hence ?thesis by simp
}
ultimately show ?thesis ..
qed
lemma offs_pred_cons:
assumes
A: "offs_pred ns ub (x # xs) index key mi ma" and
B: "index key x (length ns) mi ma = i" and
C: "i < length ns"
shows "offs_pred (ns[i := Suc (ns ! i)]) ub xs index key mi ma"
using A
proof (simp add: offs_pred_def, rule_tac allI, rule_tac impI, simp)
let ?ns' = "ns[i := Suc (ns ! i)]"
fix j
assume
"\<forall>j < length ns. offs_num (length ns) (x # xs) index key mi ma j \<le>
offs_next ns ub (x # xs) index key mi ma j - ns ! j" and
"j < length ns"
hence D: "offs_num (length ns) (x # xs) index key mi ma j \<le>
offs_next ns ub (x # xs) index key mi ma j - ns ! j"
by simp
from B and C show "offs_num (length ns) xs index key mi ma j \<le>
offs_next ?ns' ub xs index key mi ma j - ?ns' ! j"
proof (cases "j = i", case_tac [2] "0 < offs_num (length ns) xs index key mi ma j",
simp_all)
assume "j = i"
hence "offs_num (length ns) xs index key mi ma i \<le>
offs_next ns ub (x # xs) index key mi ma i - Suc (ns ! i)"
using B and D by (simp add: offs_num_cons)
moreover have "offs_next ns ub (x # xs) index key mi ma i \<le>
offs_next ?ns' ub xs index key mi ma i"
using B by (rule_tac offs_pred_next_cons [OF A _ C], simp_all add:
offs_num_cons)
ultimately show "offs_num (length ns) xs index key mi ma i \<le>
offs_next ?ns' ub xs index key mi ma i - Suc (ns ! i)"
by simp
next
assume "j \<noteq> i"
hence "offs_num (length ns) xs index key mi ma j \<le>
offs_next ns ub (x # xs) index key mi ma j - ns ! j"
using B and D by (simp add: offs_num_cons)
moreover assume "0 < offs_num (length ns) xs index key mi ma j"
hence "offs_next ns ub (x # xs) index key mi ma j \<le>
offs_next ?ns' ub xs index key mi ma j"
using B by (rule_tac offs_pred_next_cons [OF A _ C], simp_all add:
offs_num_cons)
ultimately show "offs_num (length ns) xs index key mi ma j \<le>
offs_next ?ns' ub xs index key mi ma j - ns ! j"
by simp
qed
qed
text \<open>
\null
The next step consists of proving, as done in lemma @{text fill_count_item} in what follows, that if
certain conditions hold, particularly if offsets' list $ns$ and objects' list $xs$ satisfy predicate
@{const offs_pred}, then function @{const fill} conserves objects if called using $xs$ and $ns$ as
its input arguments.
This lemma is proven by induction on $xs$. Hence, lemma @{thm [source] offs_pred_cons}, proven in
the previous step, is used to remove the antecedent containing predicate @{const offs_pred} from the
induction hypothesis, which has the form of an implication.
\null
\<close>
lemma offs_next_zero:
assumes
A: "i < length ns" and
B: "offs_num (length ns) xs index key mi ma i = 0" and
C: "offs_set_prev ns xs index key mi ma i = {}"
shows "offs_next ns ub xs index key mi ma 0 =
offs_next ns ub xs index key mi ma i"
proof -
have "offs_set_next ns xs index key mi ma 0 =
offs_set_next ns xs index key mi ma i"
proof (rule set_eqI, rule iffI, simp_all, (erule conjE)+, rule ccontr, simp add:
not_less)
fix j
assume D: "0 < offs_num (length ns) xs index key mi ma j"
assume "j \<le> i"
hence "j < i \<or> j = i"
by (simp add: le_less)
moreover {
assume "j < i"
hence "j \<in> offs_set_prev ns xs index key mi ma i"
using A and D by simp
hence False
using C by simp
}
moreover {
assume "j = i"
hence False
using B and D by simp
}
ultimately show False ..
qed
thus ?thesis
by (simp only: offs_next_def)
qed
lemma offs_next_zero_cons_eq:
assumes
A: "index key x (length ns) mi ma = i" and
B: "offs_num (length ns) (x # xs) index key mi ma 0 = 0" and
C: "offs_set_prev ns (x # xs) index key mi ma i \<noteq> {}"
(is "?A \<noteq> _")
shows "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma 0 =
offs_next ns ub (x # xs) index key mi ma 0"
proof -
have D: "Min ?A \<in> ?A"
using C by (rule_tac Min_in, simp)
moreover have E: "0 < Min ?A"
proof (rule ccontr, simp)
assume "Min ?A = 0"
hence "offs_num (length ns) (x # xs) index key mi ma (Min ?A) = 0"
using B by simp
moreover have "0 < offs_num (length ns) (x # xs) index key mi ma (Min ?A)"
using D by auto
ultimately show False by simp
qed
ultimately have "Min ?A \<in> offs_set_next ns (x # xs) index key mi ma 0"
(is "_ \<in> ?B")
by auto
moreover from this have "Min ?B = Min ?A"
proof (subst Min_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI,
(erule_tac conjE)+, rule_tac ccontr, simp add: not_le)
fix j
assume F: "j < Min ?A"
have "i < length ns"
using D by simp
moreover have "Min ?A < i"
using D by auto
hence "j < i"
using F by simp
moreover assume "0 < offs_num (length ns) (x # xs) index key mi ma j"
ultimately have "j \<in> ?A" by simp
hence "Min ?A \<le> j"
by (rule_tac Min_le, simp)
thus False
using F by simp
qed
moreover have "Min ?A \<in> offs_set_next (ns[i := Suc (ns ! i)])
xs index key mi ma 0"
(is "_ \<in> ?C")
using D and E by (auto, simp add: offs_num_cons A [symmetric])
moreover from this have "Min ?C = Min ?A"
proof (subst Min_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI,
(erule_tac conjE)+, rule_tac ccontr, simp add: not_le)
fix j
assume F: "j < Min ?A"
have "i < length ns"
using D by simp
moreover have "Min ?A < i"
using D by auto
hence "j < i"
using F by simp
moreover assume "0 < offs_num (length ns) xs index key mi ma j"
hence "0 < offs_num (length ns) (x # xs) index key mi ma j"
by (simp add: offs_num_cons)
ultimately have "j \<in> ?A" by simp
hence "Min ?A \<le> j"
by (rule_tac Min_le, simp)
thus False
using F by simp
qed
moreover have "Min ?A < i"
using D by auto
ultimately show ?thesis
by (simp only: offs_next_def split: if_split,
(rule_tac conjI, blast, rule_tac impI)+, simp)
qed
lemma offs_next_zero_cons_neq:
assumes
A: "index key x (length ns) mi ma = i" and
B: "i < length ns" and
C: "0 < i" and
D: "offs_set_prev ns (x # xs) index key mi ma i = {}"
shows "offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma 0 =
(if 0 < offs_num (length ns) xs index key mi ma i
then Suc (ns ! i)
else offs_next ns ub (x # xs) index key mi ma i)"
proof (simp, rule conjI, rule_tac [!] impI)
let ?ns' = "ns[i := Suc (ns ! i)]"
assume "0 < offs_num (length ns) xs index key mi ma i"
with A have "offs_set_next ?ns' xs index key mi ma 0 =
offs_set_next ns (x # xs) index key mi ma 0"
by (rule_tac set_eqI, rule_tac iffI, simp_all add: offs_num_cons split:
if_split_asm)
moreover have "i \<in> offs_set_next ns (x # xs) index key mi ma 0"
using A and B and C by (simp add: offs_num_cons)
moreover from this have
"Min (offs_set_next ns (x # xs) index key mi ma 0) = i"
proof (subst Min_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI,
(erule_tac conjE)+, rule_tac ccontr, simp add: not_le)
fix j
assume "j < i" and "0 < offs_num (length ns) (x # xs) index key mi ma j"
hence "j \<in> offs_set_prev ns (x # xs) index key mi ma i"
using B by simp
thus False
using D by simp
qed
ultimately show "offs_next ?ns' ub xs index key mi ma 0 = Suc (ns ! i)"
by (simp only: offs_next_def split: if_split, rule_tac conjI, rule_tac [!] impI,
simp_all)
next
let ?ns' = "ns[i := Suc (ns ! i)]"
assume "offs_num (length ns) xs index key mi ma i = 0"
moreover have "offs_set_prev ?ns' xs index key mi ma i = {}"
using D by (simp add: offs_num_cons split: if_split_asm, blast)
ultimately have "offs_next ?ns' ub xs index key mi ma 0 =
offs_next ?ns' ub xs index key mi ma i"
using B by (rule_tac offs_next_zero, simp_all)
moreover have "offs_next ?ns' ub xs index key mi ma i =
offs_next ns ub (x # xs) index key mi ma i"
using A and B and D by (rule_tac offs_next_cons_eq, simp_all add:
offs_num_cons)
ultimately show "offs_next ?ns' ub xs index key mi ma 0 =
offs_next ns ub (x # xs) index key mi ma i"
by simp
qed
lemma offs_pred_zero_cons_less:
assumes
A: "offs_pred ns ub (x # xs) index key mi ma" and
B: "index key x (length ns) mi ma = i" and
C: "i < length ns" and
D: "0 < i" and
E: "offs_set_prev ns (x # xs) index key mi ma i = {}"
shows "offs_next ns ub (x # xs) index key mi ma 0 <
offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma 0"
(is "?M < ?N")
proof -
have "i \<in> offs_set_next ns (x # xs) index key mi ma 0"
using B and C and D by (simp add: offs_num_cons)
moreover from this have
"Min (offs_set_next ns (x # xs) index key mi ma 0) = i"
proof (subst Min_eq_iff, simp, blast, simp, rule_tac allI, rule_tac impI,
(erule_tac conjE)+, rule_tac ccontr, simp add: not_le)
fix j
assume "j < i" and "0 < offs_num (length ns) (x # xs) index key mi ma j"
hence "j \<in> offs_set_prev ns (x # xs) index key mi ma i"
using C by simp
thus False
using E by simp
qed
ultimately have F: "?M = ns ! i"
by (simp only: offs_next_def split: if_split, blast)
have "?N = (if 0 < offs_num (length ns) xs index key mi ma i
then Suc (ns ! i)
else offs_next ns ub (x # xs) index key mi ma i)"
using B and C and D and E by (rule offs_next_zero_cons_neq)
thus ?thesis
proof (split if_split_asm)
assume "?N = Suc (ns ! i)"
thus ?thesis
using F by simp
next
assume "?N = offs_next ns ub (x # xs) index key mi ma i"
moreover have "ns ! i < offs_next ns ub (x # xs) index key mi ma i"
using B by (rule_tac offs_pred_next [OF A C], simp add: offs_num_cons)
ultimately show ?thesis
using F by simp
qed
qed
lemma offs_pred_zero_cons:
assumes
A: "offs_pred ns ub (x # xs) index key mi ma" and
B: "index key x (length ns) mi ma = i" and
C: "i < length ns" and
D: "offs_num (length ns) (x # xs) index key mi ma 0 = 0"
shows "offs_next ns ub (x # xs) index key mi ma 0 \<le>
offs_next (ns[i := Suc (ns ! i)]) ub xs index key mi ma 0"
(is "?M \<le> ?N")
proof (cases "offs_set_prev ns (x # xs) index key mi ma i = {}")
case True
have "0 < i"
using B and D by (rule_tac ccontr, simp add: offs_num_cons)
hence "?M < ?N"
using True and B by (rule_tac offs_pred_zero_cons_less [OF A _ C])
thus ?thesis by simp
next
case False
hence "?N = ?M"
by (rule offs_next_zero_cons_eq [OF B D])
thus ?thesis by simp
qed
lemma replicate_count:
"count (mset (replicate n x)) x = n"
by (induction n, simp_all)
lemma fill_none [rule_format]:
assumes A: "index_less index key"
shows
"(\<forall>x \<in> set xs. key x \<in> {mi..ma}) \<longrightarrow>
ns \<noteq> [] \<longrightarrow>
offs_pred ns ub xs index key mi ma \<longrightarrow>
offs_none ns ub xs index key mi ma i \<longrightarrow>
fill xs ns index key ub mi ma ! i = None"
proof (induction xs arbitrary: ns, simp add: offs_none_def offs_num_def offs_next_def,
(rule impI)+, simp add: Let_def, (erule conjE)+)
fix x xs and ns :: "nat list"
let ?i' = "index key x (length ns) mi ma"
let ?ns' = "ns[?i' := Suc (ns ! ?i')]"
assume
B: "offs_pred ns ub (x # xs) index key mi ma" and
C: "offs_none ns ub (x # xs) index key mi ma i"
assume
D: "ns \<noteq> []" and "mi \<le> key x" and "key x \<le> ma"
moreover from this have
E: "?i' < length ns"
using A by (simp add: index_less_def)
hence "offs_pred ?ns' ub xs index key mi ma"
by (rule_tac offs_pred_cons [OF B], simp)
moreover assume "\<And>ns. ns \<noteq> [] \<longrightarrow> offs_pred ns ub xs index key mi ma \<longrightarrow>
offs_none ns ub xs index key mi ma i \<longrightarrow>
fill xs ns index key ub mi ma ! i = None"
ultimately have
F: "offs_none ?ns' ub xs index key mi ma i \<longrightarrow>
fill xs ?ns' index key ub mi ma ! i = None"
by simp
show "(fill xs ?ns' index key ub mi ma)[ns ! ?i' := Some x] ! i = None"
proof (insert C, simp add: offs_none_def, erule disjE, erule_tac [2] disjE, simp_all del: subst_all
add: offs_num_cons split: if_split_asm, erule conjE, rule case_split, drule mp,
assumption, simp_all del: subst_all, (erule conjE)+, (erule_tac [2] conjE)+,
erule_tac [3] conjE, erule_tac [5] conjE)
fix j
assume
G: "?i' = j" and
H: "j < length ns" and
I: "Suc (ns ! j + offs_num (length ns) xs index key mi ma j) \<le> i" and
J: "i < offs_next ns ub (x # xs) index key mi ma j"
show "fill xs (ns[j := Suc (ns ! j)]) index key ub mi ma ! i = None"
proof (cases "0 < offs_num (length ns) xs index key mi ma j",
case_tac [2] "offs_set_prev ns (x # xs) index key mi ma j \<noteq> {}",
simp_all only: not_not not_gr0)
have "j < length ns \<and> 0 < offs_num (length ns) xs index key mi ma j \<and>
?ns' ! j + offs_num (length ns) xs index key mi ma j \<le> i \<and>
i < offs_next ?ns' ub xs index key mi ma j \<longrightarrow>
fill xs ?ns' index key ub mi ma ! i = None"
using F by (simp add: offs_none_def, blast)
moreover assume "0 < offs_num (length ns) xs index key mi ma j"
moreover have "?ns' ! j + offs_num (length ns) xs index key mi ma j \<le> i"
using G and H and I by simp
moreover have "0 < offs_num (length ns) (x # xs) index key mi ma j"
using G by (simp add: offs_num_cons)
hence "offs_next ns ub (x # xs) index key mi ma j \<le>
offs_next ?ns' ub xs index key mi ma j"
using G and H by (rule_tac offs_pred_next_cons [OF B], simp_all)
hence "i < offs_next ?ns' ub xs index key mi ma j"
using J by simp
ultimately have "fill xs ?ns' index key ub mi ma ! i = None"
using H by blast
thus ?thesis
using G by simp
next
let ?j' = "Max (offs_set_prev ns (x # xs) index key mi ma j)"
have "?j' < length ns \<and> 0 < offs_num (length ns) xs index key mi ma ?j' \<and>
?ns' ! ?j' + offs_num (length ns) xs index key mi ma ?j' \<le> i \<and>
i < offs_next ?ns' ub xs index key mi ma ?j' \<longrightarrow>
fill xs ?ns' index key ub mi ma ! i = None"
using F by (simp add: offs_none_def, blast)
moreover assume K: "offs_set_prev ns (x # xs) index key mi ma j \<noteq> {}"
hence "?j' \<in> offs_set_prev ns (x # xs) index key mi ma j"
by (rule_tac Max_in, simp)
hence L: "?j' < length ns \<and> ?j' < j \<and>
0 < offs_num (length ns) xs index key mi ma ?j'"
using G by (auto, subst (asm) (2) offs_num_cons, simp)
moreover have "ns ! ?j' + offs_num (length ns) (x # xs)
index key mi ma ?j' \<le> ns ! j"
using G and H and L by (rule_tac offs_pred_asc [OF B], simp_all add:
offs_num_cons)
hence "?ns' ! ?j' + offs_num (length ns) xs index key mi ma ?j' \<le> ns ! j"
using G and H and L by (subst nth_list_update, simp_all add: offs_num_cons)
hence "?ns' ! ?j' + offs_num (length ns) xs index key mi ma ?j' \<le> i"
using I by simp
moreover assume M: "offs_num (length ns) xs index key mi ma j = 0"
have "offs_next (ns[j := Suc (ns ! j)]) ub xs index key mi ma ?j' =
(if 0 < offs_num (length ns) xs index key mi ma j
then Suc (ns ! j)
else offs_next ns ub (x # xs) index key mi ma j)"
using G and K by (rule_tac offs_next_cons_neq, simp_all)
hence "offs_next ?ns' ub xs index key mi ma ?j' =
offs_next ns ub (x # xs) index key mi ma j"
using G and M by simp
hence "i < offs_next ?ns' ub xs index key mi ma ?j'"
using J by simp
ultimately have "fill xs ?ns' index key ub mi ma ! i = None" by blast
thus ?thesis
using G by simp
next
have "offs_num (length ns) xs index key mi ma 0 = 0 \<and>
i < offs_next ?ns' ub xs index key mi ma 0 \<longrightarrow>
fill xs ?ns' index key ub mi ma ! i = None"
using F by (simp add: offs_none_def)
moreover assume
K: "offs_set_prev ns (x # xs) index key mi ma j = {}" and
L: "offs_num (length ns) xs index key mi ma j = 0"
have "offs_set_prev ns (x # xs) index key mi ma j =
offs_set_prev ?ns' xs index key mi ma j"
using G by (rule_tac set_eqI, rule_tac iffI,
simp_all add: offs_num_cons split: if_split_asm)
hence M: "offs_set_prev ?ns' xs index key mi ma j = {}"
using K by simp
hence "offs_num (length ns) xs index key mi ma 0 = 0"
using H and L by (cases j, simp_all)
moreover have N: "offs_next ?ns' ub xs index key mi ma 0 =
offs_next ?ns' ub xs index key mi ma j"
using H and L and M by (rule_tac offs_next_zero, simp_all)
have "offs_next ?ns' ub xs index key mi ma j =
offs_next ns ub (x # xs) index key mi ma j"
using G and H and K by (subst offs_next_cons_eq, simp_all add:
offs_num_cons)
hence "i < offs_next ?ns' ub xs index key mi ma 0"
using J and N by simp
ultimately have "fill xs ?ns' index key ub mi ma ! i = None" by blast
thus ?thesis
using G by simp
qed
next
fix j
assume
G: "?i' \<noteq> j" and
H: "j < length ns" and
I: "ns ! j + offs_num (length ns) xs index key mi ma j \<le> i" and
J: "i < offs_next ns ub (x # xs) index key mi ma j" and
K: "0 < offs_num (length ns) xs index key mi ma j"
from G have "ns ! ?i' \<noteq> i"
proof (rule_tac notI, cases "?i' < j", simp_all add: not_less le_less)
assume "?i' < j"
hence "ns ! ?i' + offs_num (length ns) (x # xs) index key mi ma ?i' \<le> ns ! j"
using H and K by (rule_tac offs_pred_asc [OF B], simp_all add:
offs_num_cons)
moreover assume "ns ! ?i' = i"
ultimately have "i < ns ! j" by (simp add: offs_num_cons)
thus False
using I by simp
next
assume "j < ?i'"
hence L: "?i' \<in> offs_set_next ns (x # xs) index key mi ma j"
(is "_ \<in> ?A")
using E by (simp add: offs_num_cons)
hence "Min ?A \<le> ?i'"
by (rule_tac Min_le, simp)
hence "Min ?A < ?i' \<or> Min ?A = ?i'"
by (simp add: le_less)
hence "ns ! Min ?A \<le> ns ! ?i'"
proof (rule disjE, simp_all)
assume "Min ?A < ?i'"
moreover have "Min ?A \<in> ?A"
using L by (rule_tac Min_in, simp, blast)
hence "0 < offs_num (length ns) (x # xs) index key mi ma (Min ?A)"
by simp
ultimately have "ns ! Min ?A + offs_num (length ns) (x # xs)
index key mi ma (Min ?A) \<le> ns ! ?i'"
using E by (rule_tac offs_pred_asc [OF B], simp_all add: offs_num_cons)
thus ?thesis by simp
qed
hence "offs_next ns ub (x # xs) index key mi ma j \<le> ns ! ?i'"
using L by (simp only: offs_next_def split: if_split, blast)
moreover assume "ns ! ?i' = i"
ultimately show False
using J by simp
qed
thus "(fill xs ?ns' index key ub mi ma)[ns ! ?i' := Some x] ! i = None"
proof simp
have "j < length ns \<and> 0 < offs_num (length ns) xs index key mi ma j \<and>
?ns' ! j + offs_num (length ns) xs index key mi ma j \<le> i \<and>
i < offs_next ?ns' ub xs index key mi ma j \<longrightarrow>
fill xs ?ns' index key ub mi ma ! i = None"
using F by (simp add: offs_none_def, blast)
moreover have "offs_next ns ub (x # xs) index key mi ma j \<le>
offs_next ?ns' ub xs index key mi ma j"
using E and K by (rule_tac offs_pred_next_cons [OF B], simp_all add:
offs_num_cons)
hence "i < offs_next ?ns' ub xs index key mi ma j"
using J by simp
ultimately show "fill xs ?ns' index key ub mi ma ! i = None"
using G and H and I and K by simp
qed
next
assume
G: "0 < ?i'" and
H: "offs_num (length ns) xs index key mi ma 0 = 0" and
I: "i < offs_next ns ub (x # xs) index key mi ma 0"
have "ns ! ?i' \<noteq> i"
proof
have "0 < offs_num (length ns) (x # xs) index key mi ma ?i'"
by (simp add: offs_num_cons)
hence L: "?i' \<in> offs_set_next ns (x # xs) index key mi ma 0"
(is "_ \<in> ?A")
using E and G by simp
hence "Min ?A \<le> ?i'"
by (rule_tac Min_le, simp)
hence "Min ?A < ?i' \<or> Min ?A = ?i'"
by (simp add: le_less)
hence "ns ! Min ?A \<le> ns ! ?i'"
proof (rule disjE, simp_all)
assume "Min ?A < ?i'"
moreover have "Min ?A \<in> ?A"
using L by (rule_tac Min_in, simp, blast)
hence "0 < offs_num (length ns) (x # xs) index key mi ma (Min ?A)"
by simp
ultimately have "ns ! Min ?A + offs_num (length ns) (x # xs)
index key mi ma (Min ?A) \<le> ns ! ?i'"
using E by (rule_tac offs_pred_asc [OF B], simp_all add: offs_num_cons)
thus ?thesis by simp
qed
hence "offs_next ns ub (x # xs) index key mi ma 0 \<le> ns ! ?i'"
using L by (simp only: offs_next_def split: if_split, blast)
moreover assume "ns ! ?i' = i"
ultimately show False
using I by simp
qed
thus "(fill xs ?ns' index key ub mi ma)[ns ! ?i' := Some x] ! i = None"
proof simp
have "offs_num (length ns) xs index key mi ma 0 = 0 \<and>
i < offs_next ?ns' ub xs index key mi ma 0 \<longrightarrow>
fill xs ?ns' index key ub mi ma ! i = None"
using F by (simp add: offs_none_def)
moreover have "offs_next ns ub (x # xs) index key mi ma 0 \<le>
offs_next ?ns' ub xs index key mi ma 0"
using E and G and H by (rule_tac offs_pred_zero_cons [OF B],
simp_all add: offs_num_cons)
hence "i < offs_next ?ns' ub xs index key mi ma 0"
using I by simp
ultimately show "fill xs ?ns' index key ub mi ma ! i = None"
using H by simp
qed
next
assume
G: "?i' = 0" and
H: "i < ns ! 0"
show "fill xs (ns[0 := Suc (ns ! 0)]) index key ub mi ma ! i = None"
proof (cases "0 < offs_num (length ns) xs index key mi ma 0", simp_all)
have "0 < offs_num (length ns) xs index key mi ma 0 \<and> i < ?ns' ! 0 \<longrightarrow>
fill xs ?ns' index key ub mi ma ! i = None"
using F by (simp add: offs_none_def)
moreover assume "0 < offs_num (length ns) xs index key mi ma 0"
moreover have "i < ?ns' ! 0"
using D and G and H by simp
ultimately show ?thesis
using G by simp
next
have "offs_num (length ns) xs index key mi ma 0 = 0 \<and>
i < offs_next ?ns' ub xs index key mi ma 0 \<longrightarrow>
fill xs ?ns' index key ub mi ma ! i = None"
using F by (simp add: offs_none_def)
moreover assume "offs_num (length ns) xs index key mi ma 0 = 0"
moreover have I: "offs_next ?ns' ub xs index key mi ma 0 =
offs_next ns ub (x # xs) index key mi ma 0"
using D and G by (rule_tac offs_next_cons_eq, simp_all add:
offs_num_cons)
have "ns ! 0 < offs_next ns ub (x # xs) index key mi ma 0"
using D and G by (rule_tac offs_pred_next [OF B], simp_all add:
offs_num_cons)
hence "i < offs_next ?ns' ub xs index key mi ma 0"
using H and I by simp
ultimately show ?thesis
using G by simp
qed
next
assume
G: "0 < ?i'" and
H: "0 < offs_num (length ns) xs index key mi ma 0" and
I: "i < ns ! 0"
have "ns ! ?i' \<noteq> i"
proof -
have "ns ! 0 + offs_num (length ns) (x # xs) index key mi ma 0 \<le> ns ! ?i'"
using H by (rule_tac offs_pred_asc [OF B G E], simp_all add:
offs_num_cons)
moreover have "0 < offs_num (length ns) (x # xs) index key mi ma 0"
using H by (simp add: offs_num_cons)
ultimately show ?thesis
using I by simp
qed
thus "(fill xs ?ns' index key ub mi ma)[ns ! ?i' := Some x] ! i = None"
proof simp
have "0 < offs_num (length ns) xs index key mi ma 0 \<and> i < ?ns' ! 0 \<longrightarrow>
fill xs ?ns' index key ub mi ma ! i = None"
using F by (simp add: offs_none_def)
moreover have "i < ?ns' ! 0"
using G and I by simp
ultimately show "fill xs ?ns' index key ub mi ma ! i = None"
using H by simp
qed
qed
qed
lemma fill_index_none [rule_format]:
assumes
A: "index_less index key" and
B: "key x \<in> {mi..ma}" and
C: "ns \<noteq> []" and
D: "offs_pred ns ub (x # xs) index key mi ma"
shows "\<forall>x \<in> set xs. key x \<in> {mi..ma} \<Longrightarrow>
fill xs (ns[(index key x (length ns) mi ma) :=
Suc (ns ! index key x (length ns) mi ma)]) index key ub mi ma !
(ns ! index key x (length ns) mi ma) = None"
(is "_ \<Longrightarrow> fill _ ?ns' _ _ _ _ _ ! (_ ! ?i) = _")
using A and B and C and D
proof (rule_tac fill_none, simp_all, rule_tac offs_pred_cons,
simp_all, simp add: index_less_def, cases "0 < ?i",
cases "offs_set_prev ns (x # xs) index key mi ma ?i = {}",
case_tac [3] "0 < offs_num (length ns) xs index key mi ma 0")
assume
E: "0 < ?i" and
F: "offs_set_prev ns (x # xs) index key mi ma ?i = {}"
have G: "?i < length ns"
using A and B and C by (simp add: index_less_def)
hence "offs_num (length ns) (x # xs) index key mi ma 0 = 0"
using E and F by (rule_tac ccontr, simp)
hence "offs_num (length ns) xs index key mi ma 0 = 0"
by (simp add: offs_num_cons split: if_split_asm)
moreover have "offs_next ?ns' ub xs index key mi ma 0 =
(if 0 < offs_num (length ns) xs index key mi ma ?i
then Suc (ns ! ?i)
else offs_next ns ub (x # xs) index key mi ma ?i)"
using E and F and G by (rule_tac offs_next_zero_cons_neq, simp_all)
hence "ns ! ?i < offs_next ?ns' ub xs index key mi ma 0"
by (simp split: if_split_asm, rule_tac offs_pred_next [OF D G], simp add:
offs_num_cons)
ultimately show "offs_none ?ns' ub xs index key mi ma (ns ! ?i)"
by (simp add: offs_none_def)
next
assume
E: "0 < ?i" and
F: "offs_set_prev ns (x # xs) index key mi ma ?i \<noteq> {}"
(is "?A \<noteq> _")
have G: "?i < length ns"
using A and B and C by (simp add: index_less_def)
have H: "Max ?A \<in> ?A"
using F by (rule_tac Max_in, simp)
hence I: "Max ?A < ?i" by blast
have "Max ?A < length ns"
using H by auto
moreover have "0 < offs_num (length ns) (x # xs) index key mi ma (Max ?A)"
using H by auto
hence "0 < offs_num (length ns) xs index key mi ma (Max ?A)"
using I by (subst (asm) offs_num_cons, split if_split_asm, simp_all)
moreover have "ns ! Max ?A + offs_num (length ns) (x # xs)
index key mi ma (Max ?A) \<le> ns ! ?i"
using G and H by (rule_tac offs_pred_asc [OF D], simp_all add: offs_num_cons)
hence "?ns' ! Max ?A + offs_num (length ns) xs
index key mi ma (Max ?A) \<le> ns ! ?i"
using I by (simp add: offs_num_cons)
moreover have "offs_next ?ns' ub xs index key mi ma (Max ?A) =
(if 0 < offs_num (length ns) xs index key mi ma ?i
then Suc (ns ! ?i)
else offs_next ns ub (x # xs) index key mi ma ?i)"
using F and I by (rule_tac offs_next_cons_neq, simp_all)
hence "ns ! ?i < offs_next ?ns' ub xs index key mi ma (Max ?A)"
by (simp split: if_split_asm, rule_tac offs_pred_next [OF D G], simp add:
offs_num_cons)
ultimately show "offs_none ?ns' ub xs index key mi ma (ns ! ?i)"
by (simp add: offs_none_def, blast)
next
assume "0 < offs_num (length ns) xs index key mi ma 0" and "\<not> 0 < ?i"
moreover have "?i < length ns"
using A and B and C by (simp add: index_less_def)
ultimately show "offs_none ?ns' ub xs index key mi ma (ns ! ?i)"
by (simp add: offs_none_def)
next
assume
E: "\<not> 0 < ?i" and
F: "\<not> 0 < offs_num (length ns) xs index key mi ma 0"
have G: "?i < length ns"
using A and B and C by (simp add: index_less_def)
have "offs_num (length ns) xs index key mi ma 0 = 0"
using F by simp
moreover have "offs_next ?ns' ub xs index key mi ma ?i =
offs_next ns ub (x # xs) index key mi ma ?i"
using E and G by (rule_tac offs_next_cons_eq, simp_all add: offs_num_cons)
hence "ns ! ?i < offs_next ?ns' ub xs index key mi ma ?i"
by (simp, rule_tac offs_pred_next [OF D G], simp add: offs_num_cons)
hence "ns ! ?i < offs_next ?ns' ub xs index key mi ma 0"
using E by simp
ultimately show "offs_none ?ns' ub xs index key mi ma (ns ! ?i)"
by (simp add: offs_none_def)
qed
lemma fill_count_item [rule_format]:
assumes A: "index_less index key"
shows
"(\<forall>x \<in> set xs. key x \<in> {mi..ma}) \<longrightarrow>
ns \<noteq> [] \<longrightarrow>
offs_pred ns ub xs index key mi ma \<longrightarrow>
length xs \<le> ub \<longrightarrow>
count (mset (map the (fill xs ns index key ub mi ma))) x =
count (mset xs) x + (if the None = x then ub - length xs else 0)"
proof (induction xs arbitrary: ns, simp add: replicate_count, (rule impI)+,
simp add: Let_def map_update del: count_add_mset mset_map split del: if_split,
(erule conjE)+, subst add_mset_add_single, simp only: count_single count_union)
fix y xs and ns :: "nat list"
let ?i = "index key y (length ns) mi ma"
let ?ns' = "ns[?i := Suc (ns ! ?i)]"
assume
B: "\<forall>x \<in> set xs. mi \<le> key x \<and> key x \<le> ma" and
C: "mi \<le> key y" and
D: "key y \<le> ma" and
E: "ns \<noteq> []" and
F: "offs_pred ns ub (y # xs) index key mi ma" and
G: "Suc (length xs) \<le> ub"
have H: "?i < length ns"
using A and C and D and E by (simp add: index_less_def)
assume "\<And>ns. ns \<noteq> [] \<longrightarrow> offs_pred ns ub xs index key mi ma \<longrightarrow>
count (mset (map the (fill xs ns index key ub mi ma))) x =
count (mset xs) x + (if the None = x then ub - length xs else 0)"
moreover have "offs_pred ?ns' ub xs index key mi ma"
using F and H by (rule_tac offs_pred_cons, simp_all)
ultimately have "count (mset (map the (fill xs ?ns' index key ub mi ma))) x =
count (mset xs) x + (if the None = x then ub - length xs else 0)"
using E by simp
moreover have "ns ! ?i + offs_num (length ns) (y # xs)
index key mi ma ?i \<le> ub"
using F and H by (rule offs_pred_ub, simp add: offs_num_cons)
hence "ns ! ?i < ub"
by (simp add: offs_num_cons)
ultimately show "count (mset ((map the (fill xs ?ns' index key ub mi ma))
[ns ! ?i := y])) x = count (mset xs) x + (if y = x then 1 else 0) +
(if the None = x then ub - length (y # xs) else 0)"
proof (subst mset_update, simp add: fill_length, subst add_mset_add_single, simp
only: count_diff count_single count_union, subst nth_map, simp add: fill_length,
subst add.assoc, subst (3) add.commute, subst add.assoc [symmetric],
subst add_right_cancel)
have "fill xs ?ns' index key ub mi ma ! (ns ! ?i) = None"
using B and C and D and E by (rule_tac fill_index_none [OF A _ _ F],
simp_all)
thus "count (mset xs) x + (if the None = x then ub - length xs else 0) -
(if the (fill xs ?ns' index key ub mi ma ! (ns ! ?i)) = x then 1 else 0) =
count (mset xs) x + (if the None = x then ub - length (y # xs) else 0)"
using G by simp
qed
qed
text \<open>
\null
Finally, lemma @{text offs_enum_pred} here below proves that, if $ns$ is the offsets' list obtained
by applying the composition of functions @{const offs} and @{const enum} to objects' list $xs$, then
predicate @{const offs_pred} is satisfied by $ns$ and $xs$.
This result is in turn used, together with lemma @{thm [source] fill_count_item}, to prove lemma
@{text fill_offs_enum_count_item}, which states that function @{const fill} conserves objects if its
input offsets' list is computed via the composition of functions @{const offs} and @{const enum}.
\null
\<close>
lemma enum_offs_num:
"i < n \<Longrightarrow> enum xs index key n mi ma ! i = offs_num n xs index key mi ma i"
by (induction xs, simp add: offs_num_def, simp add: Let_def offs_num_cons,
subst nth_list_update_eq, simp_all add: enum_length)
lemma offs_length:
"length (offs ns i) = length ns"
by (induction ns arbitrary: i, simp_all)
lemma offs_add [rule_format]:
"i < length ns \<longrightarrow> offs ns k ! i = foldl (+) k (take i ns)"
by (induction ns arbitrary: i k, simp, simp add: nth_Cons split: nat.split)
lemma offs_mono_aux:
"i \<le> j \<Longrightarrow> j < length ns \<Longrightarrow> offs ns k ! i \<le> offs ns k ! (i + (j - i))"
by (simp only: offs_add take_add, simp add: add_le)
lemma offs_mono:
"i \<le> j \<Longrightarrow> j < length ns \<Longrightarrow> offs ns k ! i \<le> offs ns k ! j"
by (frule offs_mono_aux, simp_all)
lemma offs_update:
"j < length ns \<Longrightarrow>
offs (ns[i := Suc (ns ! i)]) k ! j = (if j \<le> i then id else Suc) (offs ns k ! j)"
by (simp add: offs_add not_le take_update_swap, rule impI, subst nth_take [symmetric],
assumption, subst add_update, simp_all)
lemma offs_equal_suc:
assumes
A: "Suc i < length ns" and
B: "ns ! i = 0"
shows "offs ns m ! i = offs ns m ! Suc i"
proof -
have "offs ns m ! i = foldl (+) m (take i ns)"
using A by (subst offs_add, simp_all)
also have "\<dots> = foldl (+) m (take i ns @ [ns ! i])"
using B by simp
also have "\<dots> = foldl (+) m (take (Suc i) ns)"
using A by (subst take_Suc_conv_app_nth, simp_all)
also have "\<dots> = offs ns m ! Suc i"
using A by (subst offs_add, simp_all)
finally show ?thesis .
qed
lemma offs_equal [rule_format]:
"i < j \<Longrightarrow> j < length ns \<Longrightarrow>
(\<forall>k \<in> {i..<j}. ns ! k = 0) \<longrightarrow> offs ns m ! i = offs ns m ! j"
proof (erule strict_inc_induct, rule_tac [!] impI, simp_all, erule offs_equal_suc, simp)
fix i
assume A: "i < j" and "j < length ns"
hence "Suc i < length ns" by simp
moreover assume "\<forall>k \<in> {i..<j}. ns ! k = 0"
hence "ns ! i = 0"
using A by simp
ultimately have "offs ns m ! i = offs ns m ! Suc i"
by (rule offs_equal_suc)
also assume "\<dots> = offs ns m ! j"
finally show "offs ns m ! i = offs ns m ! j" .
qed
lemma offs_enum_last [rule_format]:
assumes
A: "index_less index key" and
B: "0 < n" and
C: "\<forall>x \<in> set xs. key x \<in> {mi..ma}"
shows "offs (enum xs index key n mi ma) k ! (n - Suc 0) +
offs_num n xs index key mi ma (n - Suc 0) = length xs + k"
proof -
let ?ns = "enum xs index key n mi ma"
from B have D: "last ?ns = offs_num n xs index key mi ma (n - Suc 0)"
by (subst last_conv_nth, subst length_0_conv [symmetric], simp_all add:
enum_length, subst enum_offs_num, simp_all)
have "offs ?ns k ! (n - Suc 0) = foldl (+) k (take (n - Suc 0) ?ns)"
using B by (rule_tac offs_add, simp add: enum_length)
also have "\<dots> = foldl (+) k (butlast ?ns)"
by (simp add: butlast_conv_take enum_length)
finally have "offs ?ns k ! (n - Suc 0) + offs_num n xs index key mi ma
(n - Suc 0) = foldl (+) k (butlast ?ns @ [last ?ns])"
using D by simp
also have "\<dots> = foldl (+) k ?ns"
using B by (subst append_butlast_last_id, subst length_0_conv [symmetric],
simp_all add: enum_length)
also have "\<dots> = foldl (+) 0 ?ns + k"
by (rule add_base_zero)
also have "\<dots> = length xs + k"
using A and B and C by (subst enum_add, simp_all)
finally show ?thesis .
qed
lemma offs_enum_ub [rule_format]:
assumes
A: "index_less index key" and
B: "i < n" and
C: "\<forall>x \<in> set xs. key x \<in> {mi..ma}"
shows "offs (enum xs index key n mi ma) k ! i \<le> length xs + k"
proof -
let ?ns = "enum xs index key n mi ma"
have "offs ?ns k ! i \<le> offs ?ns k ! (n - Suc 0)"
using B by (rule_tac offs_mono, simp_all add: enum_length)
also have "\<dots> \<le> offs ?ns k ! (n - Suc 0) +
offs_num n xs index key mi ma (n - Suc 0)"
by simp
also have "\<dots> = length xs + k"
using A and B and C by (rule_tac offs_enum_last, simp_all)
finally show ?thesis .
qed
lemma offs_enum_next_ge [rule_format]:
assumes
A: "index_less index key" and
B: "i < n"
shows "\<forall>x \<in> set xs. key x \<in> {mi..ma} \<Longrightarrow>
offs (enum xs index key n mi ma) k ! i \<le>
offs_next (offs (enum xs index key n mi ma) k) (length xs + k)
xs index key mi ma i"
(is "_ \<Longrightarrow> offs ?ns _ ! _ \<le> _")
proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI,
rule offs_enum_ub [OF A B], simp)
assume "offs_set_next (offs ?ns k) xs index key mi ma i \<noteq> {}"
(is "?A \<noteq> _")
hence C: "Min ?A \<in> ?A"
by (rule_tac Min_in, simp)
hence "i \<le> Min ?A" by simp
moreover have "Min ?A < length ?ns"
using C by (simp add: offs_length)
ultimately show "offs ?ns k ! i \<le> offs ?ns k ! Min ?A"
by (rule offs_mono)
qed
lemma offs_enum_zero_aux [rule_format]:
"\<lbrakk>index_less index key; 0 < n; \<forall>x \<in> set xs. key x \<in> {mi..ma};
offs_num n xs index key mi ma (n - Suc 0) = 0\<rbrakk> \<Longrightarrow>
offs (enum xs index key n mi ma) k ! (n - Suc 0) = length xs + k"
by (subst offs_enum_last [where key = key and mi = mi and ma = ma,
symmetric], simp+)
lemma offs_enum_zero [rule_format]:
assumes
A: "index_less index key" and
B: "i < n" and
C: "\<forall>x \<in> set xs. key x \<in> {mi..ma}" and
D: "offs_num n xs index key mi ma i = 0"
shows "offs (enum xs index key n mi ma) k ! i =
offs_next (offs (enum xs index key n mi ma) k) (length xs + k)
xs index key mi ma i"
proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI,
cases "i = n - Suc 0", simp)
assume "i = n - Suc 0"
thus "offs (enum xs index key n mi ma) k ! (n - Suc 0) = length xs + k"
using A and B and C and D by (rule_tac offs_enum_zero_aux, simp_all)
next
let ?ns = "enum xs index key n mi ma"
assume E: "offs_set_next (offs ?ns k) xs index key mi ma i = {}"
(is "?A = _")
assume "i \<noteq> n - Suc 0"
hence F: "i < n - Suc 0"
using B by simp
hence "offs ?ns k ! i = offs ?ns k ! (n - Suc 0)"
proof (rule offs_equal, simp_all add: enum_length le_less,
erule_tac conjE, erule_tac disjE, rule_tac ccontr, drule_tac [2] sym, simp_all)
fix j
assume G: "j < n - Suc 0"
hence "j < length (offs ?ns k)"
by (simp add: offs_length enum_length)
moreover assume "i < j"
moreover assume "0 < ?ns ! j"
hence "0 < offs_num (length (offs ?ns k)) xs index key mi ma j"
using G by (subst (asm) enum_offs_num, simp_all add:
offs_length enum_length)
ultimately have "j \<in> ?A" by simp
thus False
using E by simp
next
show "?ns ! i = 0"
using B and D by (subst enum_offs_num, simp_all)
qed
also from A and B and C have "\<dots> = length xs + k"
proof (rule_tac offs_enum_zero_aux, simp_all, rule_tac ccontr, simp)
have "n - Suc 0 < length (offs ?ns k)"
using B by (simp add: offs_length enum_length)
moreover assume "0 < offs_num n xs index key mi ma (n - Suc 0)"
hence "0 < offs_num (length (offs ?ns k)) xs index key mi ma (n - Suc 0)"
by (simp add: offs_length enum_length)
ultimately have "n - Suc 0 \<in> ?A"
using F by simp
thus False
using E by simp
qed
finally show "offs (enum xs index key n mi ma) k ! i = length xs + k" .
next
let ?ns = "enum xs index key n mi ma"
assume "offs_set_next (offs ?ns k) xs index key mi ma i \<noteq> {}"
(is "?A \<noteq> _")
hence "Min ?A \<in> ?A"
by (rule_tac Min_in, simp)
thus "offs ?ns k ! i = offs ?ns k ! Min ?A"
proof (rule_tac offs_equal, simp_all add: le_less, simp add: offs_length,
(erule_tac conjE)+, erule_tac disjE, rule_tac ccontr, drule_tac [2] sym, simp_all)
fix j
assume E: "j < Min ?A" and "Min ?A < length (offs ?ns k)"
hence F: "j < length (offs ?ns k)" by simp
moreover assume "i < j"
moreover assume "0 < ?ns ! j"
hence "0 < offs_num (length (offs ?ns k)) xs index key mi ma j"
using F by (subst (asm) enum_offs_num, simp_all add:
offs_length enum_length)
ultimately have "j \<in> ?A" by simp
hence "Min ?A \<le> j"
by (rule_tac Min_le, simp)
thus False
using E by simp
next
show "?ns ! i = 0"
using B and D by (subst enum_offs_num, simp_all)
qed
qed
lemma offs_enum_next_cons [rule_format]:
assumes
A: "index_less index key" and
B: "\<forall>x \<in> set xs. key x \<in> {mi..ma}"
shows "(if i < index key x n mi ma then (\<le>) else (<))
(offs_next (offs (enum xs index key n mi ma) k)
(length xs + k) xs index key mi ma i)
(offs_next (offs ((enum xs index key n mi ma) [index key x n mi ma :=
Suc (enum xs index key n mi ma ! index key x n mi ma)]) k)
(Suc (length xs + k)) (x # xs) index key mi ma i)"
(is "(if i < ?i' then _ else _)
(offs_next (offs ?ns _) _ _ _ _ _ _ _)
(offs_next (offs ?ns' _) _ _ _ _ _ _ _)")
proof (simp_all only: offs_next_def not_less split: if_split, (rule conjI, rule impI)+,
simp, simp, (rule_tac [!] impI, (rule_tac [!] conjI)?)+, rule_tac [2-3] FalseE,
rule_tac [4] conjI, rule_tac [4-5] impI)
assume
C: "offs_set_next (offs ?ns k) xs index key mi ma i = {}"
(is "?A = _") and
D: "offs_set_next (offs ?ns' k) (x # xs) index key mi ma i \<noteq> {}"
(is "?A' \<noteq> _") and
E: "i < ?i'"
from C have F: "\<forall>j \<noteq> ?i'. j \<notin> ?A'"
by (rule_tac allI, rule_tac impI, rule_tac notI, simp add: enum_length
offs_length offs_num_cons split: if_split_asm, (erule_tac conjE)+, simp)
from D have "Min ?A' \<in> ?A'"
by (rule_tac Min_in, simp)
hence G: "Min ?A' < n"
by (simp add: offs_length enum_length)
have H: "Min ?A' = ?i'"
proof (rule Min_eqI, simp, rule eq_refl, erule contrapos_pp, insert F, simp)
have "\<exists>j. j \<in> ?A'"
using D by blast
then obtain j where "j \<in> ?A'" ..
moreover from this have "j = ?i'"
by (erule_tac contrapos_pp, insert F, simp)
ultimately show "?i' \<in> ?A'" by simp
qed
with G have "offs ?ns' k ! Min ?A' = offs ?ns k ! Min ?A'"
by (subst offs_update, simp_all add: enum_length)
also from A and B and G and H have
"\<dots> = offs_next (offs ?ns k) (length xs + k) xs index key mi ma (Min ?A')"
proof (rule_tac offs_enum_zero, simp_all, rule_tac ccontr, simp)
assume "?i' < n" and "0 < offs_num n xs index key mi ma ?i'"
hence "?i' \<in> ?A"
using E by (simp add: offs_length enum_length)
thus False
using C by simp
qed
also have "\<dots> = length xs + k"
proof (simp only: offs_next_def split: if_split, rule conjI, simp, rule impI,
rule FalseE, simp, erule exE, (erule conjE)+)
fix j
assume "j < length (offs ?ns k)"
moreover assume "Min ?A' < j"
hence "i < j"
using E and H by simp
moreover assume "0 < offs_num (length (offs ?ns k)) xs index key mi ma j"
ultimately have "j \<in> ?A" by simp
thus False
using C by simp
qed
finally show "length xs + k \<le> offs ?ns' k ! Min ?A'" by simp
next
assume
C: "offs_set_next (offs ?ns k) xs index key mi ma i = {}"
(is "?A = _") and
D: "offs_set_next (offs ?ns' k) (x # xs) index key mi ma i \<noteq> {}"
(is "?A' \<noteq> _") and
E: "?i' \<le> i"
have "\<exists>j. j \<in> ?A'"
using D by blast
then obtain j where F: "j \<in> ?A'" ..
hence "j < length (offs ?ns k)"
by (simp add: offs_length)
moreover have "i < j"
using F by simp
moreover from this have
"0 < offs_num (length (offs ?ns k)) xs index key mi ma j"
using E and F by (simp add: offs_length enum_length offs_num_cons)
ultimately have "j \<in> ?A" by simp
thus False
using C by simp
next
assume
C: "offs_set_next (offs ?ns k) xs index key mi ma i \<noteq> {}"
(is "?A \<noteq> _") and
D: "offs_set_next (offs ?ns' k) (x # xs) index key mi ma i = {}"
(is "?A' = _")
have "\<exists>j. j \<in> ?A"
using C by blast
then obtain j where E: "j \<in> ?A" ..
hence "j < length (offs ?ns' k)"
by (simp add: offs_length)
moreover have "i < j"
using E by simp
moreover have "0 < offs_num (length (offs ?ns' k)) (x # xs) index key mi ma j"
using E by (simp add: offs_length enum_length offs_num_cons)
ultimately have "j \<in> ?A'" by simp
thus False
using D by simp
next
assume "offs_set_next (offs ?ns k) xs index key mi ma i \<noteq> {}"
(is "?A \<noteq> _")
hence "Min ?A \<in> ?A"
by (rule_tac Min_in, simp)
hence C: "Min ?A < n"
by (simp add: offs_length enum_length)
assume "offs_set_next (offs ?ns' k) (x # xs) index key mi ma i \<noteq> {}"
(is "?A' \<noteq> _")
hence D: "Min ?A' \<in> ?A'"
by (rule_tac Min_in, simp)
hence E: "Min ?A' < n"
by (simp add: offs_length enum_length)
have "offs ?ns k ! Min ?A \<le> offs ?ns k ! Min ?A'"
proof (cases "offs_num n xs index key mi ma (Min ?A') = 0")
case True
have "offs ?ns k ! Min ?A' =
offs_next (offs ?ns k) (length xs + k) xs index key mi ma (Min ?A')"
using A and B and E and True by (rule_tac offs_enum_zero, simp_all)
also from A and B and C have "\<dots> \<ge> offs ?ns k ! Min ?A"
proof (simp only: offs_next_def split: if_split, rule_tac conjI, rule_tac [!] impI,
rule_tac offs_enum_ub, simp, simp, simp)
assume "offs_set_next (offs ?ns k) xs index key mi ma (Min ?A') \<noteq> {}"
(is "?B \<noteq> _")
hence "Min ?B \<in> ?B"
by (rule_tac Min_in, simp)
hence "Min ?B \<in> ?A"
using D by simp
moreover from this have "Min ?A \<le> Min ?B"
by (rule_tac Min_le, simp)
ultimately show "offs ?ns k ! Min ?A \<le> offs ?ns k ! Min ?B"
by (rule_tac offs_mono, simp_all add: offs_length)
qed
finally show ?thesis .
next
case False
hence "Min ?A' \<in> ?A"
using D by (simp add: offs_length enum_length)
hence "Min ?A \<le> Min ?A'"
by (rule_tac Min_le, simp)
thus ?thesis
by (rule offs_mono, simp_all add: enum_length E)
qed
also have "\<dots> \<le> offs ?ns' k ! Min ?A'"
using E by (subst offs_update, simp_all add: enum_length)
finally show "offs ?ns k ! Min ?A \<le> offs ?ns' k ! Min ?A'" .
next
let ?A = "offs_set_next (offs ?ns k) xs index key mi ma i"
assume "offs_set_next (offs ?ns' k) (x # xs) index key mi ma i \<noteq> {}"
(is "?A' \<noteq> _")
hence C: "Min ?A' \<in> ?A'"
by (rule_tac Min_in, simp)
hence D: "Min ?A' < n"
by (simp add: offs_length enum_length)
assume "?i' \<le> i"
hence E: "?i' < Min ?A'"
using C by simp
hence "0 < offs_num n xs index key mi ma (Min ?A')"
using C by (simp add: offs_length enum_length offs_num_cons)
hence "Min ?A' \<in> ?A"
using C by (simp add: offs_length enum_length)
hence "Min ?A \<le> Min ?A'"
by (rule_tac Min_le, simp)
hence "offs ?ns k ! Min ?A \<le> offs ?ns k ! Min ?A'"
by (rule offs_mono, simp_all add: enum_length D)
also have "\<dots> < offs ?ns' k ! Min ?A'"
using E by (subst offs_update, simp_all add: enum_length D)
finally show "offs ?ns k ! Min ?A < offs ?ns' k ! Min ?A'" .
qed
lemma offs_enum_pred [rule_format]:
assumes A: "index_less index key"
shows "(\<forall>x \<in> set xs. key x \<in> {mi..ma}) \<longrightarrow>
offs_pred (offs (enum xs index key n mi ma) k) (length xs + k)
xs index key mi ma"
proof (induction xs, simp add: offs_pred_def offs_num_def,
simp add: Let_def offs_pred_def offs_length enum_length, rule impI, (erule conjE)+,
simp, rule allI, rule impI, erule allE, drule mp, assumption)
fix x xs i
let ?i' = "index key x n mi ma"
let ?ns = "enum xs index key n mi ma"
let ?ns' = "?ns[?i' := Suc (?ns ! ?i')]"
assume
B: "\<forall>x \<in> set xs. mi \<le> key x \<and> key x \<le> ma" and
C: "i < n" and
D: "offs_num n xs index key mi ma i \<le>
offs_next (offs ?ns k) (length xs + k) xs index key mi ma i - offs ?ns k ! i"
have E: "(if i < ?i' then (\<le>) else (<))
(offs_next (offs ?ns k) (length xs + k) xs index key mi ma i)
(offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i)"
using A and B by (subst offs_enum_next_cons, simp_all)
show "offs_num n (x # xs) index key mi ma i \<le>
offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i -
offs ?ns' k ! i"
proof (subst offs_update, simp add: enum_length C, rule linorder_cases [of i ?i'],
simp_all add: offs_num_cons)
assume "i < ?i'"
hence "offs_next (offs ?ns k) (length xs + k) xs index key mi ma i \<le>
offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i"
using E by simp
thus "offs_num n xs index key mi ma i \<le>
offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i -
offs ?ns k ! i"
using D by arith
next
assume F: "i = ?i'"
hence "Suc (offs_num n xs index key mi ma ?i') \<le>
Suc (offs_next (offs ?ns k) (length xs + k) xs index key mi ma ?i' -
offs ?ns k ! ?i')"
using D by simp
also from A and B and C and F have "\<dots> =
Suc (offs_next (offs ?ns k) (length xs + k) xs index key mi ma ?i') -
offs ?ns k ! ?i'"
by (rule_tac Suc_diff_le [symmetric], rule_tac offs_enum_next_ge, simp_all)
finally have "Suc (offs_num n xs index key mi ma ?i') \<le>
Suc (offs_next (offs ?ns k) (length xs + k) xs index key mi ma ?i') -
offs ?ns k ! ?i'" .
moreover have
"Suc (offs_next (offs ?ns k) (length xs + k) xs index key mi ma ?i') \<le>
offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma ?i'"
using E and F by simp
ultimately show "Suc (offs_num n xs index key mi ma ?i') \<le>
offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma ?i' -
offs ?ns k ! ?i'"
by arith
next
assume "?i' < i"
hence "Suc (offs_next (offs ?ns k) (length xs + k) xs index key mi ma i) \<le>
offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i"
using E by simp
thus "offs_num n xs index key mi ma i \<le>
offs_next (offs ?ns' k) (Suc (length xs + k)) (x # xs) index key mi ma i -
Suc (offs ?ns k ! i)"
using D by arith
qed
qed
lemma fill_offs_enum_count_item [rule_format]:
"\<lbrakk>index_less index key; \<forall>x \<in> set xs. key x \<in> {mi..ma}; 0 < n\<rbrakk> \<Longrightarrow>
count (mset (map the (fill xs (offs (enum xs index key n mi ma) 0)
index key (length xs) mi ma))) x =
count (mset xs) x"
by (subst fill_count_item, simp_all, simp only: length_greater_0_conv [symmetric]
offs_length enum_length, insert offs_enum_pred [of index key xs mi ma n 0], simp)
text \<open>
\null
Using lemma @{thm [source] fill_offs_enum_count_item}, step 9 of the proof method can now be dealt
with. It is accomplished by proving lemma @{text gcsort_count_inv}, which states that the number of
the occurrences of whatever object in the objects' list is still the same after any recursive round.
\null
\<close>
lemma nths_count:
"count (mset (nths xs A)) x =
count (mset xs) x - card {i. i < length xs \<and> i \<notin> A \<and> xs ! i = x}"
proof (induction xs arbitrary: A, simp_all add: nths_Cons)
fix v xs A
let ?B = "{i. i < length xs \<and> Suc i \<notin> A \<and> xs ! i = x}"
let ?C = "\<lambda>v. {i. i < Suc (length xs) \<and> i \<notin> A \<and> (v # xs) ! i = x}"
have A: "\<And>v. ?C v = ?C v \<inter> {0} \<union> ?C v \<inter> {i. \<exists>j. i = Suc j}"
by (subst Int_Un_distrib [symmetric], auto, arith)
have "\<And>v. card (?C v) = card (?C v \<inter> {0}) + card (?C v \<inter> {i. \<exists>j. i = Suc j})"
by (subst A, rule card_Un_disjoint, simp_all, blast)
moreover have "\<And>v. card ?B = card (?C v \<inter> {i. \<exists>j. i = Suc j})"
by (rule bij_betw_same_card [of Suc], auto)
ultimately show
"(0 \<in> A \<longrightarrow>
(v = x \<longrightarrow> Suc (count (mset xs) x - card ?B) =
Suc (count (mset xs) x) - card (?C x)) \<and>
(v \<noteq> x \<longrightarrow> count (mset xs) x - card ?B =
count (mset xs) x - card (?C v))) \<and>
(0 \<notin> A \<longrightarrow>
(v = x \<longrightarrow> count (mset xs) x - card ?B =
Suc (count (mset xs) x) - card (?C x)) \<and>
(v \<noteq> x \<longrightarrow> count (mset xs) x - card ?B =
count (mset xs) x - card (?C v)))"
proof ((rule_tac [!] conjI, rule_tac [!] impI)+, simp_all)
have "card ?B \<le> count (mset xs) x"
by (simp add: count_mset length_filter_conv_card, rule card_mono,
simp, blast)
thus "Suc (count (mset xs) x - card ?B) = Suc (count (mset xs) x) - card ?B"
by (rule Suc_diff_le [symmetric])
qed
qed
lemma round_count_inv [rule_format]:
"index_less index key \<longrightarrow> bn_inv p q t \<longrightarrow> add_inv n t \<longrightarrow> count_inv f t \<longrightarrow>
count_inv f (round index key p q r t)"
proof (induction index key p q r t arbitrary: n f rule: round.induct,
(rule_tac [!] impI)+, simp, simp, simp_all only: simp_thms)
fix index p q r u ns xs n f and key :: "'a \<Rightarrow> 'b"
let ?t = "round index key p q r (u, ns, tl xs)"
assume
"\<And>n f. bn_inv p q (u, ns, tl xs) \<longrightarrow> add_inv n (u, ns, tl xs) \<longrightarrow>
count_inv f (u, ns, tl xs) \<longrightarrow> count_inv f ?t" and
"bn_inv p q (u, Suc 0 # ns, xs)" and
"add_inv n (u, Suc 0 # ns, xs)" and
"count_inv f (u, Suc 0 # ns, xs)"
thus "count_inv f (round index key p q r (u, Suc 0 # ns, xs))"
proof (cases ?t, simp add: add_suc, rule_tac allI, cases xs,
simp_all add: disj_imp split: if_split_asm)
fix y ys x and xs' :: "'a list"
assume "\<And>n f. foldl (+) 0 ns = n \<and> length ys = n \<longrightarrow>
(\<forall>x. count (mset ys) x = f x) \<longrightarrow> (\<forall>x. count (mset xs') x = f x)"
moreover assume "Suc (foldl (+) 0 ns) = n" and "Suc (length ys) = n"
ultimately have "\<And>n f. (\<forall>x. count (mset ys) x = f x) \<longrightarrow>
(\<forall>x. count (mset xs') x = f x)"
by blast
moreover assume A: "\<forall>x. (y = x \<longrightarrow> Suc (count (mset ys) x) = f x) \<and>
(y \<noteq> x \<longrightarrow> count (mset ys) x = f x)"
have "\<forall>x. count (mset ys) x = (f(y := f y - Suc 0)) x"
(is "\<forall>x. _ = ?f' x")
by (simp add: A, insert spec [OF A, where x = y], simp)
ultimately have "\<forall>x. count (mset xs') x = ?f' x" ..
thus "(y = x \<longrightarrow> Suc (count (mset xs') x) = f x) \<and>
(y \<noteq> x \<longrightarrow> count (mset xs') x = f x)"
by (simp, insert spec [OF A, where x = y], rule_tac impI, simp)
qed
next
fix index p q r u m ns n f and key :: "'a \<Rightarrow> 'b" and xs :: "'a list"
let ?ws = "take (Suc (Suc m)) xs"
let ?ys = "drop (Suc (Suc m)) xs"
let ?r = "\<lambda>m'. round_suc_suc index key ?ws m m' u"
let ?t = "\<lambda>r' v. round index key p q r' (v, ns, ?ys)"
assume A: "index_less index key"
assume
"\<And>ws a b c d e g h i n f.
ws = ?ws \<Longrightarrow> a = bn_comp m p q r \<Longrightarrow> (b, c) = bn_comp m p q r \<Longrightarrow>
d = ?r b \<Longrightarrow> (e, g) = ?r b \<Longrightarrow> (h, i) = g \<Longrightarrow>
bn_inv p q (e, ns, ?ys) \<longrightarrow> add_inv n (e, ns, ?ys) \<longrightarrow>
count_inv f (e, ns, ?ys) \<longrightarrow> count_inv f (?t c e)" and
"bn_inv p q (u, Suc (Suc m) # ns, xs)" and
"add_inv n (u, Suc (Suc m) # ns, xs)" and
"count_inv f (u, Suc (Suc m) # ns, xs)"
thus "count_inv f (round index key p q r (u, Suc (Suc m) # ns, xs))"
+ using [[simproc del: defined_all]]
proof (simp split: prod.split, ((rule_tac allI)+, ((rule_tac impI)+)?)+,
(erule_tac conjE)+, subst (asm) (2) add_base_zero, simp)
fix m' r' v ms' ws' xs' x
assume
B: "?r m' = (v, ms', ws')" and
C: "bn_comp m p q r = (m', r')" and
D: "bn_valid m p q" and
E: "Suc (Suc (foldl (+) 0 ns + m)) = n" and
F: "length xs = n"
assume "\<And>ws a b c d e g h i n' f.
ws = ?ws \<Longrightarrow> a = (m', r') \<Longrightarrow> b = m' \<and> c = r' \<Longrightarrow>
d = (v, ms', ws') \<Longrightarrow> e = v \<and> g = (ms', ws') \<Longrightarrow> h = ms' \<and> i = ws' \<Longrightarrow>
foldl (+) 0 ns = n' \<and> n - Suc (Suc m) = n' \<longrightarrow>
(\<forall>x. count (mset ?ys) x = f x) \<longrightarrow> (\<forall>x. count (mset xs') x = f x)"
hence "foldl (+) 0 ns = n - Suc (Suc m) \<longrightarrow>
(\<forall>x. count (mset xs') x = count (mset ?ys) x)"
by simp
hence "count (mset xs') x = count (mset ?ys) x"
using E by simp
moreover assume "\<forall>x. count (mset xs) x = f x"
ultimately have "f x = count (mset ?ws) x + count (mset xs') x"
by (subst (asm) append_take_drop_id [of "Suc (Suc m)", symmetric],
subst (asm) mset_append, simp)
with B [symmetric] show "count (mset ws') x + count (mset xs') x = f x"
proof (simp add: round_suc_suc_def Let_def del: count_add_mset mset_map
split: if_split_asm, subst (1 2) add_mset_add_single, simp
only: count_single count_union)
let ?nmi = "mini ?ws key"
let ?nma = "maxi ?ws key"
let ?xmi = "?ws ! ?nmi"
let ?xma = "?ws ! ?nma"
let ?mi = "key ?xmi"
let ?ma = "key ?xma"
let ?k = "case m of 0 \<Rightarrow> m | Suc 0 \<Rightarrow> m | Suc (Suc i) \<Rightarrow> u + m'"
let ?zs = "nths ?ws (- {?nmi, ?nma})"
let ?ms = "enum ?zs index key ?k ?mi ?ma"
let ?A = "{i. i < Suc (Suc m) \<and> (i = ?nmi \<or> i = ?nma) \<and> ?ws ! i = x}"
have G: "length ?ws = Suc (Suc m)"
using E and F by simp
hence H: "card ?A \<le> count (mset ?ws) x"
by (simp add: count_mset length_filter_conv_card, rule_tac card_mono,
simp, blast)
show "count (mset (map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma))) x
+ (if ?xma = x then 1 else 0) + (if ?xmi = x then 1 else 0) =
count (mset ?ws) x"
proof (cases "m = 0")
case True
hence I: "length ?zs = 0"
using G by (simp add: mini_maxi_nths)
have "count (mset ?zs) x = count (mset ?ws) x - card ?A"
using G by (subst nths_count, simp)
hence J: "count (mset ?ws) x = card ?A"
using H and I by simp
from I show ?thesis
proof (simp, (rule_tac [!] conjI, rule_tac [!] impI)+,
simp_all (no_asm_simp) add: True)
assume "?xmi = x" and "?xma = x"
with G have "?A = {?nmi, ?nma}"
by (rule_tac set_eqI, rule_tac iffI, simp_all, erule_tac disjE,
insert mini_less [of ?ws key], insert maxi_less [of ?ws key],
simp_all)
with G have "card ?A = Suc (Suc 0)"
by (simp, subst card_insert_disjoint, simp_all,
rule_tac mini_maxi_neq, simp)
thus "Suc (Suc 0) = count (mset (take (Suc (Suc 0)) xs)) x"
using True and J by simp
next
assume "?xmi \<noteq> x" and "?xma = x"
with G have "?A = {?nma}"
by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+,
erule_tac disjE, insert maxi_less [of ?ws key], simp_all)
thus "Suc 0 = count (mset (take (Suc (Suc 0)) xs)) x"
using True and J by simp
next
assume "?xmi = x" and "?xma \<noteq> x"
with G have "?A = {?nmi}"
by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+,
erule_tac disjE, insert mini_less [of ?ws key], simp_all)
thus "Suc 0 = count (mset (take (Suc (Suc 0)) xs)) x"
using True and J by simp
next
assume "?xmi \<noteq> x" and "?xma \<noteq> x"
hence "?A = {}"
by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+,
erule_tac disjE, simp_all)
hence "count (mset ?ws) x = 0"
using J by simp
thus "x \<notin> set (take (Suc (Suc 0)) xs)"
using True by simp
qed
next
case False
hence "0 < ?k"
by (simp, drule_tac bn_comp_fst_nonzero [OF D], subst (asm) C,
simp split: nat.split)
hence "count (mset (map the (fill ?zs (offs ?ms 0) index key
(length ?zs) ?mi ?ma))) x = count (mset ?zs) x"
by (rule_tac fill_offs_enum_count_item [OF A], simp, rule_tac conjI,
((rule_tac mini_lb | rule_tac maxi_ub), erule_tac in_set_nthsD)+)
with G show ?thesis
proof (simp, (rule_tac [!] conjI, rule_tac [!] impI)+,
simp_all add: mini_maxi_nths nths_count)
assume "?xmi = x" and "?xma = x"
with G have "?A = {?nmi, ?nma}"
by (rule_tac set_eqI, rule_tac iffI, simp_all, erule_tac disjE,
insert mini_less [of ?ws key], insert maxi_less [of ?ws key],
simp_all)
with G have "card ?A = Suc (Suc 0)"
by (simp, subst card_insert_disjoint, simp_all,
rule_tac mini_maxi_neq, simp)
thus "Suc (Suc (count (mset ?ws) x - card ?A)) = count (mset ?ws) x"
using H by simp
next
assume "?xmi \<noteq> x" and "?xma = x"
with G have "?A = {?nma}"
by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+,
erule_tac disjE, insert maxi_less [of ?ws key], simp_all)
thus "Suc (count (mset ?ws) x - card ?A) = count (mset ?ws) x"
using H by simp
next
assume "?xmi = x" and "?xma \<noteq> x"
with G have "?A = {?nmi}"
by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+,
erule_tac disjE, insert mini_less [of ?ws key], simp_all)
thus "Suc (count (mset ?ws) x - card ?A) = count (mset ?ws) x"
using H by simp
next
assume "?xmi \<noteq> x" and "?xma \<noteq> x"
hence "?A = {}"
by (rule_tac set_eqI, rule_tac iffI, simp_all, (erule_tac conjE)+,
erule_tac disjE, simp_all)
thus "count (mset ?ws) x - card ?A = count (mset ?ws) x"
by (simp (no_asm_simp))
qed
qed
qed
qed
qed
lemma gcsort_count_inv:
assumes
A: "index_less index key" and
B: "add_inv n t" and
C: "n \<le> p"
shows "\<lbrakk>t' \<in> gcsort_set index key p t; count_inv f t\<rbrakk> \<Longrightarrow>
count_inv f t'"
by (erule gcsort_set.induct, simp, drule gcsort_add_inv [OF A _ B C],
rule round_count_inv [OF A], simp_all del: bn_inv.simps, erule conjE,
frule sym, erule subst, rule bn_inv_intro, insert C, simp)
text \<open>
\null
The only remaining task is to address step 10 of the proof method, which is done by proving theorem
@{text gcsort_count}. It holds under the conditions that the objects' number is not larger than the
counters' upper bound and function @{text index} satisfies predicate @{const index_less}, and states
that for any object, function @{const gcsort} leaves unchanged the number of its occurrences within
the input objects' list.
\null
\<close>
theorem gcsort_count:
assumes
A: "index_less index key" and
B: "length xs \<le> p"
shows "count (mset (gcsort index key p xs)) x = count (mset xs) x"
proof -
let ?t = "(0, [length xs], xs)"
have "count_inv (count (mset xs)) (gcsort_aux index key p ?t)"
by (rule gcsort_count_inv [OF A _ B], rule gcsort_add_input,
rule gcsort_aux_set, rule gcsort_count_input)
hence "count (mset (gcsort_out (gcsort_aux index key p ?t))) x =
(count (mset xs)) x"
by (rule gcsort_count_intro)
moreover have "?t = gcsort_in xs"
by (simp add: gcsort_in_def)
ultimately show ?thesis
by (simp add: gcsort_def)
qed
end
diff --git a/thys/Generalized_Counting_Sort/Sorting.thy b/thys/Generalized_Counting_Sort/Sorting.thy
--- a/thys/Generalized_Counting_Sort/Sorting.thy
+++ b/thys/Generalized_Counting_Sort/Sorting.thy
@@ -1,1267 +1,1268 @@
(* Title: An Efficient Generalization of Counting Sort for Large, possibly Infinite Key Ranges
Author: Pasquale Noce
Software Engineer at HID Global, Italy
pasquale dot noce dot lavoro at gmail dot com
pasquale dot noce at hidglobal dot com
*)
section "Proof of objects' sorting"
theory Sorting
imports Conservation
begin
text \<open>
\null
In this section, it is formally proven that GCsort actually sorts objects.
Here below, steps 5, 6, and 7 of the proof method are accomplished. Predicate @{text sort_inv} is
satisfied just in case, for any bucket delimited by the input counters' list $ns$, the keys of the
corresponding objects within the input objects' list $xs$ are not larger than those of the objects,
if any, to the right of that bucket. The underlying idea is that this predicate:
\begin{itemize}
\item
is trivially satisfied by the output of function @{const gcsort_in}, which places all objects into a
single bucket, and
\item
implies that $xs$ is sorted if every bucket delimited by $ns$ has size one, as happens when function
@{const gcsort_aux} terminates.
\end{itemize}
\null
\<close>
fun sort_inv :: "('a \<Rightarrow> 'b::linorder) \<Rightarrow> nat \<times> nat list \<times> 'a list \<Rightarrow> bool" where
"sort_inv key (u, ns, xs) =
(\<forall>i < length ns. \<forall>j < offs ns 0 ! i. \<forall>k \<in> {offs ns 0 ! i..<length xs}.
key (xs ! j) \<le> key (xs ! k))"
lemma gcsort_sort_input:
"sort_inv key (0, [length xs], xs)"
by simp
lemma offs_nth:
assumes
A: "find (\<lambda>n. Suc 0 < n) ns = None" and
B: "foldl (+) 0 ns = n" and
C: "k < n"
shows "\<exists>i < length ns. offs ns 0 ! i = k"
proof (cases ns, insert B C, simp, cases "k = 0", rule exI [of _ 0], simp,
rule ccontr, simp (no_asm_use))
fix m ms
assume
D: "ns = m # ms" and
E: "0 < k" and
F: "\<forall>i < length ns. offs ns 0 ! i \<noteq> k"
have G: "\<forall>n \<in> set ns. n \<le> Suc 0"
using A by (auto simp add: find_None_iff)
let ?A = "{i. i < length ns \<and> offs ns 0 ! i < k}"
have H: "Max ?A \<in> ?A"
using D and E by (rule_tac Max_in, simp_all, rule_tac exI [of _ 0], simp)
hence I: "Max ?A < length ns"
by simp
hence J: "offs ns 0 ! Max ?A = foldl (+) 0 (take (Max ?A) ns)"
by (rule offs_add)
have "Max ?A < length ns - Suc 0 \<or> Max ?A = length ns - Suc 0"
(is "?P \<or> ?Q")
using H by (simp, arith)
moreover {
assume ?P
hence K: "Suc (Max ?A) < length ns" by simp
hence "offs ns 0 ! Suc (Max ?A) = foldl (+) 0 (take (Suc (Max ?A)) ns)"
by (rule offs_add)
moreover have "take (Suc (Max ?A)) ns = take (Max ?A) ns @ [ns ! Max ?A]"
using I by (rule take_Suc_conv_app_nth)
ultimately have "offs ns 0 ! Suc (Max ?A) =
offs ns 0 ! Max ?A + ns ! Max ?A"
using J by simp
moreover have "offs ns 0 ! Max ?A < k"
using H by simp
moreover have "ns ! Max ?A \<in> set ns"
using I by (rule nth_mem)
with G have "ns ! Max ?A \<le> Suc 0" ..
ultimately have "offs ns 0 ! Suc (Max ?A) \<le> k" by simp
moreover have "offs ns 0 ! Suc (Max ?A) \<noteq> k"
using F and K by simp
ultimately have "offs ns 0 ! Suc (Max ?A) < k" by simp
hence "Suc (Max ?A) \<in> ?A"
using K by simp
hence "Suc (Max ?A) \<le> Max ?A"
by (rule_tac Max_ge, simp)
hence False by simp
}
moreover {
assume ?Q
hence "offs ns 0 ! Max ?A = foldl (+) 0 (take (length ns - Suc 0) ns)"
using J by simp
moreover have K: "length ns - Suc 0 < length ns"
using D by simp
hence "take (Suc (length ns - Suc 0)) ns =
take (length ns - Suc 0) ns @ [ns ! (length ns - Suc 0)]"
by (rule take_Suc_conv_app_nth)
hence "foldl (+) 0 ns =
foldl (+) 0 (take (length ns - Suc 0) ns @ [ns ! (length ns - Suc 0)])"
by simp
ultimately have "n = offs ns 0 ! Max ?A + ns ! (length ns - Suc 0)"
using B by simp
moreover have "offs ns 0 ! Max ?A < k"
using H by simp
moreover have "ns ! (length ns - Suc 0) \<in> set ns"
using K by (rule nth_mem)
with G have "ns ! (length ns - Suc 0) \<le> Suc 0" ..
ultimately have "n \<le> k" by simp
with C have False by simp
}
ultimately show False ..
qed
lemma gcsort_sort_intro:
"\<lbrakk>sort_inv key t; add_inv n t; find (\<lambda>n. Suc 0 < n) (fst (snd t)) = None\<rbrakk> \<Longrightarrow>
sorted (map key (gcsort_out t))"
proof (cases t, simp add: sorted_iff_nth_mono_less gcsort_out_def,
erule conjE, (rule allI)+, (rule impI)+)
fix ns xs j k
assume "find (\<lambda>n. Suc 0 < n) ns = None" and "foldl (+) 0 ns = n"
moreover assume A: "k < n"
ultimately have "\<exists>i < length ns. offs ns 0 ! i = k"
by (rule offs_nth)
then obtain i where "i < length ns \<and> offs ns 0 ! i = k" ..
moreover assume "\<forall>i < length ns. \<forall>j < offs ns 0 ! i. \<forall>k \<in> {offs ns 0 ! i..<n}.
key (xs ! j) \<le> key (xs ! k)"
hence "i < length ns \<longrightarrow> j < offs ns 0 ! i \<longrightarrow> k \<in> {offs ns 0 ! i..<n} \<longrightarrow>
key (xs ! j) \<le> key (xs ! k)"
by simp
ultimately have "j < k \<longrightarrow> k < n \<longrightarrow> key (xs ! j) \<le> key (xs ! k)"
by simp
moreover assume "j < k"
ultimately show "key (xs ! j) \<le> key (xs ! k)"
using A by simp
qed
text \<open>
\null
As lemma @{thm [source] gcsort_sort_intro} comprises an additional assumption concerning the form of
the fixed points of function @{const gcsort_aux}, step 8 of the proof method is necessary this time
to prove that such assumption is satisfied.
\null
\<close>
lemma gcsort_sort_form:
"find (\<lambda>n. Suc 0 < n) (fst (snd (gcsort_aux index key p t))) = None"
by (induction index key p t rule: gcsort_aux.induct, simp)
text \<open>
\null
Here below, step 9 of the proof method is accomplished.
In the most significant case of the proof by recursion induction of lemma @{text round_sort_inv},
namely that of a bucket $B$ with size larger than two and distinct minimum and maximum keys, the
following line of reasoning is adopted. Let $x$ be an object contained in a finer-grained bucket
$B'$ resulting from $B$'s partition, and $y$ an object to the right of $B'$. Then:
\begin{itemize}
\item
If $y$ is contained in some other finer-grained bucket resulting from $B$'s partition, inequality
@{term "key x \<le> key y"} holds because predicate @{const sort_inv} is satisfied by a counters' list
generated by function @{const enum} and an objects' list generated by function @{const fill} in case
@{const fill}'s input offsets' list is computed via the composition of functions @{const offs} and
@{const enum}, as happens within function @{const round}.
\\This is proven beforehand in lemma @{text fill_sort_inv}.
\item
Otherwise, inequality @{term "key x \<le> key y"} holds as well because object $x$ was contained in $B$
by lemma @{thm [source] fill_offs_enum_count_item}, object $y$ occurred to the right of $B$ by lemma
@{thm [source] round_count_inv}, and by hypothesis, the key of any object in $B$ was not larger than
that of any object to the right of $B$.
\end{itemize}
Using lemma @{text round_sort_inv}, the invariance of predicate @{const sort_inv} over inductive set
@{const gcsort_set} is then proven in lemma @{text gcsort_sort_inv}.
\null
\<close>
lemma mini_maxi_keys_le:
"x \<in> set xs \<Longrightarrow> key (xs ! mini xs key) \<le> key (xs ! maxi xs key)"
by (frule mini_lb, drule maxi_ub, erule order_trans)
lemma mini_maxi_keys_eq [rule_format]:
"key (xs ! mini xs key) = key (xs ! maxi xs key) \<longrightarrow> x \<in> set xs \<longrightarrow>
key x = key (xs ! maxi xs key)"
by (induction xs, simp_all add: Let_def, (rule_tac [!] impI, (rule_tac [!] conjI)?)+,
rule_tac [2-4] impI, frule_tac [1-3] mini_maxi_keys_le [where key = key], simp_all)
lemma offs_suc:
"i < length ns \<Longrightarrow> offs ns (Suc k) ! i = Suc (offs ns k ! i)"
by (simp add: offs_add add_suc)
lemma offs_base_zero:
"i < length ns \<Longrightarrow> offs ns k ! i = offs ns 0 ! i + k"
by (simp add: offs_add, subst add_base_zero, simp)
lemma offs_append:
"offs (ms @ ns) k = offs ms k @ offs ns (foldl (+) k ms)"
by (induction ms arbitrary: k, simp_all)
lemma offs_enum_next_le [rule_format]:
assumes
A: "index_less index key" and
B: "i < j" and
C: "j < n" and
D: "\<forall>x \<in> set xs. key x \<in> {mi..ma}"
shows "offs_next (offs (enum xs index key n mi ma) k) (length xs + k)
xs index key mi ma i \<le> offs (enum xs index key n mi ma) k ! j"
(is "_ \<le> offs ?ns _ ! _")
proof (rule ccontr, simp add: not_le)
assume E: "offs ?ns k ! j <
offs_next (offs ?ns k) (length xs + k) xs index key mi ma i"
from B have "offs_set_next (offs ?ns k) xs index key mi ma i =
offs_set_next (offs ?ns k) xs index key mi ma j"
proof (rule_tac set_eqI, rule_tac iffI, simp_all, rule_tac ccontr, simp add: not_less)
fix m
assume "m < length (offs ?ns k) \<and> i < m \<and>
0 < offs_num (length (offs ?ns k)) xs index key mi ma m"
hence F: "m \<in> offs_set_next (offs ?ns k) xs index key mi ma i"
(is "_ \<in> ?A")
by simp
hence "Min ?A \<le> m"
by (rule_tac Min_le, simp)
moreover assume "m \<le> j"
ultimately have "offs ?ns k ! Min ?A \<le> offs ?ns k ! j"
using C by (rule_tac offs_mono, simp_all add: enum_length)
hence "offs_next (offs ?ns k) (length xs + k) xs index key mi ma i \<le>
offs ?ns k ! j"
using F by (simp only: offs_next_def split: if_split, rule_tac conjI,
blast, simp)
thus False
using E by simp
qed
hence "offs ?ns k ! j <
offs_next (offs ?ns k) (length xs + k) xs index key mi ma j"
using E by (simp only: offs_next_def)
moreover have "offs_num n xs index key mi ma j = 0"
proof (rule ccontr, simp)
assume "0 < offs_num n xs index key mi ma j"
hence "j \<in> offs_set_next (offs ?ns k) xs index key mi ma i"
(is "_ \<in> ?A")
using B and C by (simp add: offs_length enum_length)
moreover from this have "Min ?A \<le> j"
by (rule_tac Min_le, simp)
hence "offs ?ns k ! Min ?A \<le> offs ?ns k ! j"
using C by (erule_tac offs_mono, simp add: enum_length)
ultimately have "offs_next (offs ?ns k) (length xs + k) xs index key mi ma i \<le>
offs ?ns k ! j"
by (simp only: offs_next_def split: if_split, rule_tac conjI, blast, simp)
thus False
using E by simp
qed
hence "offs ?ns k ! j =
offs_next (offs ?ns k) (length xs + k) xs index key mi ma j"
using A and C and D by (rule_tac offs_enum_zero, simp_all)
ultimately show False by simp
qed
lemma offs_pred_ub_less:
"\<lbrakk>offs_pred ns ub xs index key mi ma; i < length ns;
0 < offs_num (length ns) xs index key mi ma i\<rbrakk> \<Longrightarrow>
ns ! i < ub"
by (drule offs_pred_ub, assumption+, simp)
lemma fill_count_none [rule_format]:
assumes A: "index_less index key"
shows
"(\<forall>x \<in> set xs. key x \<in> {mi..ma}) \<longrightarrow>
ns \<noteq> [] \<longrightarrow>
offs_pred ns ub xs index key mi ma \<longrightarrow>
length xs \<le> ub \<longrightarrow>
count (mset (fill xs ns index key ub mi ma)) None = ub - length xs"
using A
proof (induction xs arbitrary: ns, simp_all add: replicate_count Let_def,
(rule_tac impI)+, (erule_tac conjE)+, subst mset_update, simp add: fill_length,
erule_tac offs_pred_ub_less, simp_all add: index_less_def offs_num_cons del:
not_None_eq, subst conj_commute, rule_tac conjI, rule_tac [!] impI, rotate_tac,
rotate_tac, erule_tac contrapos_np, rule_tac fill_index_none, simp_all)
fix x xs and ns :: "nat list"
let ?i = "index key x (length ns) mi ma"
let ?ns' = "ns[?i := Suc (ns ! ?i)]"
assume "\<And>ns. ns \<noteq> [] \<longrightarrow> offs_pred ns ub xs index key mi ma \<longrightarrow>
count (mset (fill xs ns index key ub mi ma)) None = ub - length xs"
moreover assume B: "ns \<noteq> []"
moreover assume
"offs_pred ns ub (x # xs) index key mi ma" and
"mi \<le> key x" and
"key x \<le> ma"
hence "offs_pred ?ns' ub xs index key mi ma"
using A and B by (rule_tac offs_pred_cons, simp_all add: index_less_def)
ultimately show "count (mset (fill xs ?ns' index key ub mi ma)) None -
Suc 0 = ub - Suc (length xs)"
by simp
qed
lemma fill_offs_enum_count_none [rule_format]:
"\<lbrakk>index_less index key; \<forall>x \<in> set xs. key x \<in> {mi..ma}; 0 < n\<rbrakk> \<Longrightarrow>
count (mset (fill xs (offs (enum xs index key n mi ma) 0)
index key (length xs) mi ma)) None = 0"
by (subst fill_count_none, simp_all, simp only: length_greater_0_conv [symmetric]
offs_length enum_length, insert offs_enum_pred [of index key xs mi ma n 0], simp)
lemma fill_index [rule_format]:
assumes A: "index_less index key"
shows
"(\<forall>x \<in> set xs. key x \<in> {mi..ma}) \<longrightarrow>
offs_pred ns ub xs index key mi ma \<longrightarrow>
i < length ns \<longrightarrow>
0 < offs_num (length ns) xs index key mi ma i \<longrightarrow>
j \<in> {ns ! i..<offs_next ns ub xs index key mi ma i} \<longrightarrow>
fill xs ns index key ub mi ma ! j = Some x \<longrightarrow>
index key x (length ns) mi ma = i"
proof (induction xs arbitrary: ns, simp add: offs_num_def, simp add: Let_def,
(rule impI)+, (erule conjE)+, simp)
fix y xs and ns :: "nat list"
let ?i = "index key x (length ns) mi ma"
let ?i' = "index key y (length ns) mi ma"
let ?ns' = "ns[?i' := Suc (ns ! ?i')]"
let ?P = "\<lambda>ns.
offs_pred ns ub xs index key mi ma \<longrightarrow>
i < length ns \<longrightarrow>
0 < offs_num (length ns) xs index key mi ma i \<longrightarrow>
ns ! i \<le> j \<and> j < offs_next ns ub xs index key mi ma i \<longrightarrow>
fill xs ns index key ub mi ma ! j = Some x \<longrightarrow>
index key x (length ns) mi ma = i"
assume
B: "\<forall>x \<in> set xs. mi \<le> key x \<and> key x \<le> ma" and
C: "mi \<le> key y" and
D: "key y \<le> ma" and
E: "offs_pred ns ub (y # xs) index key mi ma" and
F: "i < length ns" and
G: "0 < offs_num (length ns) (y # xs) index key mi ma i" and
H: "ns ! i \<le> j" and
I: "j < offs_next ns ub (y # xs) index key mi ma i" and
J: "\<And>ns. ?P ns" and
K: "(fill xs ?ns' index key ub mi ma)[ns ! ?i' := Some y] ! j = Some x"
have "0 < length ns"
using F by arith
hence L: "?i' < length ns"
using A and C and D by (simp add: index_less_def)
hence "ns ! ?i' + offs_num (length ns) (y # xs) index key mi ma ?i' \<le> ub"
by (rule_tac offs_pred_ub [OF E], simp_all add: offs_num_cons)
hence "ns ! ?i' < ub"
by (simp add: offs_num_cons)
with K show "?i = i"
proof (cases "j = ns ! ?i'", simp, subst (asm) nth_list_update_eq, simp_all
add: fill_length)
assume
M: "j = ns ! ?i" and
N: "y = x"
show ?thesis
proof (rule ccontr, erule neqE)
assume "?i < i"
hence "ns ! ?i + offs_num (length ns) (y # xs) index key mi ma ?i \<le> ns ! i"
using F and G and N by (rule_tac offs_pred_asc [OF E], simp_all
add: offs_num_cons)
hence "j < ns ! i"
using M and N by (simp add: offs_num_cons)
thus False
using H by simp
next
let ?A = "offs_set_next ns (y # xs) index key mi ma i"
assume "i < ?i"
hence O: "?i \<in> ?A"
using N and L by (simp add: offs_num_cons)
hence P: "Min ?A \<in> ?A"
by (rule_tac Min_in, simp, blast)
have "Min ?A \<le> ?i"
using O by (rule_tac Min_le, simp)
moreover have "?A \<noteq> {}"
using O by blast
ultimately have "offs_next ns ub (y # xs) index key mi ma i \<le> j"
using M proof (simp only: offs_next_def, simp, subst (asm) le_less,
erule_tac disjE, simp_all)
assume "Min ?A < ?i"
hence "ns ! Min ?A + offs_num (length ns) (y # xs) index key mi ma
(Min ?A) \<le> ns ! ?i"
using O and P by (rule_tac offs_pred_asc [OF E], simp_all)
thus "ns ! Min ?A \<le> ns ! ?i" by simp
qed
thus False
using I by simp
qed
next
assume
M: "j \<noteq> ns ! ?i'" and
N: "fill xs ?ns' index key ub mi ma ! j = Some x"
have "?P ?ns'" using J .
moreover from D and F have "offs_pred ?ns' ub xs index key mi ma"
using L by (rule_tac offs_pred_cons [OF E], simp)
moreover have "i < length ?ns'"
using F by simp
moreover have "0 < offs_num (length ?ns') xs index key mi ma i"
proof (rule ccontr, simp)
assume O: "offs_num (length ns) xs index key mi ma i = 0"
hence P: "offs_num (length ns) (y # xs) index key mi ma i = Suc 0"
using G by (simp add: offs_num_cons split: if_split_asm)
hence "i = ?i'"
using O by (simp add: offs_num_cons split: if_split_asm)
hence "ns ! i < j"
using H and M by simp
hence "ns ! i + offs_num (length ns) (y # xs) index key mi ma i \<le> j"
using P by simp
with F and G and I have "offs_none ns ub (y # xs) index key mi ma j"
by (simp add: offs_none_def, rule_tac disjI1, rule_tac exI
[where x = i], simp)
with B and C and D and E and F have
"fill (y # xs) ns index key ub mi ma ! j = None"
by (rule_tac fill_none [OF A], simp_all, erule_tac disjE, simp_all, auto)
thus False
using K by (simp add: Let_def)
qed
moreover have "?ns' ! i \<le> j"
using F and H and M by (cases "?i' = i", simp_all)
moreover have "offs_next ns ub (y # xs) index key mi ma i \<le>
offs_next ?ns' ub xs index key mi ma i"
using G and L by (rule_tac offs_pred_next_cons [OF E], simp)
hence "j < offs_next ?ns' ub xs index key mi ma i"
using I by simp
ultimately show ?thesis
using N by simp
qed
qed
lemma fill_offs_enum_index [rule_format]:
"index_less index key \<Longrightarrow>
\<forall>x \<in> set xs. key x \<in> {mi..ma} \<Longrightarrow>
i < n \<Longrightarrow>
0 < offs_num n xs index key mi ma i \<Longrightarrow>
j \<in> {offs (enum xs index key n mi ma) 0 ! i..<
offs_next (offs (enum xs index key n mi ma) 0) (length xs)
xs index key mi ma i} \<Longrightarrow>
fill xs (offs (enum xs index key n mi ma) 0) index key (length xs)
mi ma ! j = Some x \<Longrightarrow>
index key x n mi ma = i"
by (insert fill_index [of index key xs mi ma "offs (enum xs index key n mi ma) 0"
"length xs" i j x], insert offs_enum_pred [of index key xs mi ma n 0],
simp add: offs_length enum_length)
lemma fill_sort_inv [rule_format]:
assumes
A: "index_less index key" and
B: "index_mono index key" and
C: "\<forall>x \<in> set xs. key x \<in> {mi..ma}"
shows "sort_inv key (u, enum xs index key n mi ma,
map the (fill xs (offs (enum xs index key n mi ma) 0)
index key (length xs) mi ma))"
(is "sort_inv _ (_, ?ns, _)")
proof (simp, (rule allI, rule impI)+, rule ballI, simp add: enum_length fill_length,
erule conjE)
fix i j k
let ?A = "{i. i < n \<and> 0 < ?ns ! i}"
let ?B = "{i. i < n \<and> offs ?ns 0 ! i \<le> j \<and> 0 < offs_num n xs index key mi ma i}"
let ?C = "{i. i < n \<and> offs ?ns 0 ! i \<le> k \<and> 0 < offs_num n xs index key mi ma i}"
assume
D: "i < n" and
E: "j < offs ?ns 0 ! i" and
F: "offs ?ns 0 ! i \<le> k" and
G: "k < length xs"
have H: "\<forall>i < length xs.
\<exists>x. fill xs (offs ?ns 0) index key (length xs) mi ma ! i = Some x"
proof (rule allI, rule impI, rule ccontr, simp)
fix m
assume "fill xs (offs ?ns 0) index key (length xs) mi ma ! m = None"
moreover assume "m < length xs"
hence "fill xs (offs ?ns 0) index key (length xs) mi ma ! m
\<in> set (fill xs (offs ?ns 0) index key (length xs) mi ma)"
by (rule_tac nth_mem, simp add: fill_length)
ultimately have "None \<in> set (fill xs (offs ?ns 0) index key (length xs) mi ma)"
by simp
moreover have
"count (mset (fill xs (offs ?ns 0) index key (length xs) mi ma)) None = 0"
using C and D by (rule_tac fill_offs_enum_count_none [OF A], simp_all)
ultimately show False by simp
qed
have "\<exists>y ys. xs = y # ys"
using G by (rule_tac list.exhaust [of xs], simp_all)
then obtain z and zs where I: "xs = z # zs" by blast
hence "index key z n mi ma < n"
using A and C and D by (simp add: index_less_def)
moreover from this have "0 < ?ns ! index key z n mi ma"
using I by (simp add: Let_def, subst nth_list_update_eq, simp_all add:
enum_length)
ultimately have "index key z n mi ma \<in> ?A" by simp
hence J: "Min ?A \<in> ?A"
by (rule_tac Min_in, simp, blast)
hence "offs ?ns 0 ! 0 = offs ?ns 0 ! Min ?A"
proof (cases "Min ?A = 0", simp_all, erule_tac offs_equal, simp_all add:
enum_length, rule_tac ccontr, erule_tac conjE, simp)
fix m
assume "m < Min ?A" and "Min ?A < n" and "0 < ?ns ! m"
moreover from this have "Min ?A \<le> m"
by (rule_tac Min_le, simp_all)
ultimately show False by simp
qed
moreover have "\<exists>m ms. ?ns = m # ms"
using D by (rule_tac list.exhaust [of ?ns], simp_all,
simp only: length_0_conv [symmetric] enum_length)
then obtain m and ms where "?ns = m # ms" by blast
ultimately have "offs ?ns 0 ! Min ?A = 0" by simp
hence "Min ?A \<in> ?B"
using J by (simp, subst enum_offs_num [symmetric], simp_all)
hence K: "Max ?B \<in> ?B"
by (rule_tac Max_in, simp, blast)
moreover have "j < offs_next (offs ?ns 0) (length xs)
xs index key mi ma (Max ?B)"
proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI,
rule_tac [2] ccontr, simp_all only: not_less)
show "j < length xs"
using E and F and G by simp
next
assume "offs_set_next (offs ?ns 0) xs index key mi ma (Max ?B) \<noteq> {}"
(is "?Z \<noteq> _")
hence L: "Min ?Z \<in> ?Z"
by (rule_tac Min_in, simp)
moreover assume "offs ?ns 0 ! Min ?Z \<le> j"
ultimately have "Min ?Z \<in> ?B"
by (simp add: offs_length enum_length)
hence "Min ?Z \<le> Max ?B"
by (rule_tac Max_ge, simp)
thus False
using L by simp
qed
moreover have
"\<exists>x. fill xs (offs ?ns 0) index key (length xs) mi ma ! j = Some x"
using E and F and G and H by simp
then obtain x where
L: "fill xs (offs ?ns 0) index key (length xs) mi ma ! j = Some x" ..
ultimately have M: "index key x n mi ma = Max ?B"
using C by (rule_tac fill_offs_enum_index [OF A], simp_all)
have N: "Max ?B \<in> ?C"
using E and F and K by simp
hence "Max ?C \<in> ?C"
by (rule_tac Max_in, simp, blast)
moreover have O: "k < offs_next (offs ?ns 0) (length xs)
xs index key mi ma (Max ?C)"
proof (simp only: offs_next_def split: if_split, rule conjI, rule_tac [!] impI,
rule_tac [2] ccontr, simp_all only: not_less)
show "k < length xs" using G .
next
assume "offs_set_next (offs ?ns 0) xs index key mi ma (Max ?C) \<noteq> {}"
(is "?Z \<noteq> _")
hence P: "Min ?Z \<in> ?Z"
by (rule_tac Min_in, simp)
moreover assume "offs ?ns 0 ! Min ?Z \<le> k"
ultimately have "Min ?Z \<in> ?C"
by (simp add: offs_length enum_length)
hence "Min ?Z \<le> Max ?C"
by (rule_tac Max_ge, simp)
thus False
using P by simp
qed
moreover have
"\<exists>x. fill xs (offs ?ns 0) index key (length xs) mi ma ! k = Some x"
using G and H by simp
then obtain y where
P: "fill xs (offs ?ns 0) index key (length xs) mi ma ! k = Some y" ..
ultimately have Q: "index key y n mi ma = Max ?C"
using C by (rule_tac fill_offs_enum_index [OF A], simp_all)
have "Max ?B \<le> Max ?C"
using N by (rule_tac Max_ge, simp)
hence "Max ?B < Max ?C"
proof (rule_tac ccontr, simp)
assume "Max ?B = Max ?C"
hence "offs ?ns 0 ! i < offs_next (offs ?ns 0) (length xs)
xs index key mi ma (Max ?B)"
using F and O by simp
moreover have "offs ?ns 0 ! Max ?B < offs ?ns 0 ! i"
using E and K by simp
with K have "Max ?B < i"
by (erule_tac contrapos_pp, subst not_less, subst (asm) not_less,
erule_tac offs_mono, simp add: enum_length)
hence "offs_next (offs ?ns 0) (length xs + 0) xs index key mi ma (Max ?B) \<le>
offs ?ns 0 ! i"
using C and D by (rule_tac offs_enum_next_le [OF A], simp_all)
ultimately show False by simp
qed
hence R: "index key x n mi ma < index key y n mi ma"
using M and Q by simp
have "count (mset (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))) x = count (mset xs) x"
using C and D by (rule_tac fill_offs_enum_count_item [OF A], simp_all)
moreover have S: "j < length (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))"
using E and F and G by (simp add: fill_length)
hence "map the (fill xs (offs ?ns 0) index key (length xs) mi ma) ! j
\<in> set (map the (fill xs (offs ?ns 0) index key (length xs) mi ma))"
by (rule nth_mem)
hence "0 < count (mset (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))) x"
using L and S by simp
ultimately have T: "x \<in> set xs" by simp
have "count (mset (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))) y = count (mset xs) y"
using C and D by (rule_tac fill_offs_enum_count_item [OF A], simp_all)
moreover have U: "k < length (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))"
using G by (simp add: fill_length)
hence "map the (fill xs (offs ?ns 0) index key (length xs) mi ma) ! k
\<in> set (map the (fill xs (offs ?ns 0) index key (length xs) mi ma))"
by (rule nth_mem)
hence "0 < count (mset (map the (fill xs (offs ?ns 0) index key
(length xs) mi ma))) y"
using P and U by simp
ultimately have V: "y \<in> set xs" by simp
have "key y \<le> key x \<longrightarrow> index key y n mi ma \<le> index key x n mi ma"
using B and C and T and V by (simp add: index_mono_def)
hence "key x < key y"
by (rule_tac contrapos_pp [OF R], simp add: not_less)
thus "key (the (fill xs (offs ?ns 0) index key (length xs) mi ma ! j)) \<le>
key (the (fill xs (offs ?ns 0) index key (length xs) mi ma ! k))"
using L and P by simp
qed
lemma round_sort_inv [rule_format]:
"index_less index key \<longrightarrow> index_mono index key \<longrightarrow> bn_inv p q t \<longrightarrow>
add_inv n t \<longrightarrow> sort_inv key t \<longrightarrow> sort_inv key (round index key p q r t)"
proof (induction index key p q r t arbitrary: n rule: round.induct,
(rule_tac [!] impI)+, simp, simp_all only: simp_thms)
fix index p q r u ns xs n and key :: "'a \<Rightarrow> 'b"
let ?t = "round index key p q r (u, ns, xs)"
assume "\<And>n. bn_inv p q (u, ns, xs) \<longrightarrow> add_inv n (u, ns, xs) \<longrightarrow>
sort_inv key (u, ns, xs) \<longrightarrow> sort_inv key ?t"
hence "bn_inv p q (u, ns, xs) \<longrightarrow> add_inv n (u, ns, xs) \<longrightarrow>
sort_inv key (u, ns, xs) \<longrightarrow> sort_inv key ?t" .
moreover assume
"bn_inv p q (u, 0 # ns, xs)"
"add_inv n (u, 0 # ns, xs)" and
"sort_inv key (u, 0 # ns, xs)"
ultimately show "sort_inv key (round index key p q r (u, 0 # ns, xs))"
by auto
next
fix index p q r u ns xs n and key :: "'a \<Rightarrow> 'b"
let ?t = "round index key p q r (u, ns, tl xs)"
assume
A: "index_less index key" and
B: "bn_inv p q (u, Suc 0 # ns, xs)"
moreover assume
"\<And>n. bn_inv p q (u, ns, tl xs) \<longrightarrow> add_inv n (u, ns, tl xs) \<longrightarrow>
sort_inv key (u, ns, tl xs) \<longrightarrow> sort_inv key ?t" and
"add_inv n (u, Suc 0 # ns, xs)" and
"sort_inv key (u, Suc 0 # ns, xs)"
ultimately show "sort_inv key (round index key p q r (u, Suc 0 # ns, xs))"
proof (cases ?t, cases xs, simp_all add: add_suc, (rule_tac allI, rule_tac impI)+,
rule_tac ballI, simp, (erule_tac conjE)+, simp)
fix i j k y ys u' ns' xs'
assume
C: "round index key p q r (u, ns, ys) = (u', ns', xs')" and
D: "Suc (foldl (+) 0 ns) = n" and
E: "Suc (length ys) = n" and
F: "\<forall>i < Suc (length ns).
\<forall>j < (0 # offs ns (Suc 0)) ! i. \<forall>k \<in> {(0 # offs ns (Suc 0)) ! i..<n}.
key ((y # ys) ! j) \<le> key (ys ! (k - Suc 0))"
assume A': "j < (0 # offs ns' (Suc 0)) ! i"
hence "\<exists>i'. i = Suc i'"
by (rule_tac nat.exhaust [of i], simp_all)
then obtain i' where B': "i = Suc i'" ..
assume "i < Suc (length ns')"
hence G: "i' < length ns'"
using B' by simp
hence H: "j < Suc (offs ns' 0 ! i')"
using A' and B' by (simp add: offs_suc)
assume "(0 # offs ns' (Suc 0)) ! i \<le> k"
hence "Suc (offs ns' 0 ! i') \<le> k"
using B' and G by (simp add: offs_suc)
moreover from this have "\<exists>k'. k = Suc k'"
by (rule_tac nat.exhaust [of k], simp_all)
then obtain k' where I: "k = Suc k'" ..
ultimately have J: "offs ns' 0 ! i' \<le> k'" by simp
assume "k < Suc (length xs')"
hence K: "k' < length xs'"
using I by simp
let ?P = "\<lambda>n. foldl (+) 0 ns = n \<and> length ys = n"
let ?Q = "\<lambda>n. \<forall>i < length ns.
\<forall>j < offs ns 0 ! i. \<forall>k \<in> {offs ns 0 ! i..<n}.
key (ys ! j) \<le> key (ys ! k)"
let ?R = "\<forall>i < length ns'.
\<forall>j < offs ns' 0 ! i. \<forall>k \<in> {offs ns' 0 ! i..<length xs'}.
key (xs' ! j) \<le> key (xs' ! k)"
assume "\<And>n. ?P n \<longrightarrow> ?Q n \<longrightarrow> ?R"
hence "?P (n - Suc 0) \<longrightarrow> ?Q (n - Suc 0) \<longrightarrow> ?R" .
hence "?Q (n - Suc 0) \<longrightarrow> ?R"
using D and E by auto
moreover have "?Q (n - Suc 0)"
proof ((rule allI, rule impI)+, rule ballI, simp, erule conjE)
fix i j k
have "Suc i < Suc (length ns) \<longrightarrow> (\<forall>j < (0 # offs ns (Suc 0)) ! Suc i.
\<forall>k \<in> {(0 # offs ns (Suc 0)) ! Suc i..<n}.
key ((y # ys) ! j) \<le> key (ys ! (k - Suc 0)))"
using F ..
moreover assume "i < length ns"
ultimately have "Suc j < Suc (offs ns 0 ! i) \<longrightarrow>
(\<forall>k \<in> {Suc (offs ns 0 ! i)..<n}.
key ((y # ys) ! Suc j) \<le> key (ys ! (k - Suc 0)))"
by (auto simp add: offs_suc)
moreover assume "j < offs ns 0 ! i"
ultimately have "\<forall>k \<in> {Suc (offs ns 0 ! i)..<n}.
key (ys ! j) \<le> key (ys ! (k - Suc 0))"
by simp
moreover assume "offs ns 0 ! i \<le> k" and "k < n - Suc 0"
hence "Suc k \<in> {Suc (offs ns 0 ! i)..<n}" by simp
ultimately have "key (ys ! j) \<le> key (ys ! (Suc k - Suc 0))" ..
thus "key (ys ! j) \<le> key (ys ! k)" by simp
qed
ultimately have ?R ..
from I show "key ((y # xs') ! j) \<le> key (xs' ! (k - Suc 0))"
proof (cases j, simp_all)
have "bn_inv p q (u, ns, ys)"
using B by simp
moreover have "add_inv (n - Suc 0) (u, ns, ys)"
using D and E by auto
moreover have "count_inv (\<lambda>x. count (mset ys) x) (u, ns, ys)" by simp
ultimately have "count_inv (\<lambda>x. count (mset ys) x)
(round index key p q r (u, ns, ys))"
by (rule round_count_inv [OF A])
hence "count (mset xs') (xs' ! k') = count (mset ys) (xs' ! k')"
using C by simp
moreover have "0 < count (mset xs') (xs' ! k')"
using K by simp
ultimately have "xs' ! k' \<in> set ys" by simp
hence "\<exists>m < length ys. ys ! m = xs' ! k'"
by (simp add: in_set_conv_nth)
then obtain m where L: "m < length ys \<and> ys ! m = xs' ! k'" ..
have "Suc 0 < Suc (length ns) \<longrightarrow> (\<forall>j < (0 # offs ns (Suc 0)) ! Suc 0.
\<forall>k \<in> {(0 # offs ns (Suc 0)) ! Suc 0..<n}.
key ((y # ys) ! j) \<le> key (ys ! (k - Suc 0)))"
using F ..
moreover have M: "0 < length ns"
using D [symmetric] and E and L by (rule_tac ccontr, simp)
ultimately have "\<forall>k \<in> {Suc (offs ns 0 ! 0)..<n}.
key y \<le> key (ys ! (k - Suc 0))"
by (auto simp add: offs_suc)
hence "\<forall>k \<in> {Suc 0..<n}. key y \<le> key (ys ! (k - Suc 0))"
using M by (cases ns, simp_all)
moreover have "Suc m \<in> {Suc 0..<n}"
using E and L by simp
ultimately have "key y \<le> key (ys ! (Suc m - Suc 0))" ..
thus "key y \<le> key (xs' ! k')"
using L by simp
next
case (Suc j')
moreover have "i' < length ns' \<longrightarrow> (\<forall>j < offs ns' 0 ! i'.
\<forall>k \<in> {offs ns' 0 ! i'..<length xs'}. key (xs' ! j) \<le> key (xs' ! k))"
using `?R` ..
hence "j' < offs ns' 0 ! i' \<longrightarrow>
(\<forall>k \<in> {offs ns' 0 ! i'..<length xs'}. key (xs' ! j') \<le> key (xs' ! k))"
using G by simp
ultimately have "\<forall>k \<in> {offs ns' 0 ! i'..<length xs'}.
key (xs' ! j') \<le> key (xs' ! k)"
using H by simp
moreover have "k' \<in> {offs ns' 0 ! i'..<length xs'}"
using J and K by simp
ultimately show "key (xs' ! j') \<le> key (xs' ! k')" ..
qed
qed
next
fix index p q r u m ns n and key :: "'a \<Rightarrow> 'b" and xs :: "'a list"
let ?ws = "take (Suc (Suc m)) xs"
let ?ys = "drop (Suc (Suc m)) xs"
let ?r = "\<lambda>m'. round_suc_suc index key ?ws m m' u"
let ?t = "\<lambda>r' v. round index key p q r' (v, ns, ?ys)"
assume
A: "index_less index key" and
B: "index_mono index key" and
C: "bn_inv p q (u, Suc (Suc m) # ns, xs)"
assume
"\<And>ws a b c d e f g h n.
ws = ?ws \<Longrightarrow> a = bn_comp m p q r \<Longrightarrow> (b, c) = bn_comp m p q r \<Longrightarrow>
d = ?r b \<Longrightarrow> (e, f) = ?r b \<Longrightarrow> (g, h) = f \<Longrightarrow>
bn_inv p q (e, ns, ?ys) \<longrightarrow> add_inv n (e, ns, ?ys) \<longrightarrow>
sort_inv key (e, ns, ?ys) \<longrightarrow> sort_inv key (?t c e)" and
"add_inv n (u, Suc (Suc m) # ns, xs)" and
"sort_inv key (u, Suc (Suc m) # ns, xs)"
with C show "sort_inv key (round index key p q r (u, Suc (Suc m) # ns, xs))"
+ using [[simproc del: defined_all]]
proof (simp split: prod.split, ((rule_tac allI)+, (rule_tac impI)+)+,
rule_tac ballI, simp, (erule_tac conjE)+, subst (asm) (2) add_base_zero, simp)
fix m' r' v ms' ws' u' ns' xs' i j k
let ?nmi = "mini ?ws key"
let ?nma = "maxi ?ws key"
let ?xmi = "?ws ! ?nmi"
let ?xma = "?ws ! ?nma"
let ?mi = "key ?xmi"
let ?ma = "key ?xma"
let ?k = "case m of 0 \<Rightarrow> m | Suc 0 \<Rightarrow> m | Suc (Suc i) \<Rightarrow> u + m'"
let ?zs = "nths ?ws (- {?nmi, ?nma})"
let ?ms = "enum ?zs index key ?k ?mi ?ma"
let ?P = "\<lambda>n'. foldl (+) 0 ns = n' \<and> n - Suc (Suc m) = n'"
let ?Q = "\<lambda>n'. \<forall>i < length ns.
\<forall>j < offs ns 0 ! i. \<forall>k \<in> {offs ns 0 ! i..<n'}.
key (xs ! Suc (Suc (m + j))) \<le> key (xs ! Suc (Suc (m + k)))"
let ?R = "\<forall>i < length ns'.
\<forall>j < offs ns' 0 ! i. \<forall>k \<in> {offs ns' 0 ! i..<length xs'}.
key (xs' ! j) \<le> key (xs' ! k)"
assume
D: "?r m' = (v, ms', ws')" and
E: "?t r' v = (u', ns', xs')" and
F: "bn_comp m p q r = (m', r')" and
G: "bn_valid m p q" and
H: "Suc (Suc (foldl (+) 0 ns + m)) = n" and
I: "length xs = n" and
J: "\<forall>i < Suc (length ns). \<forall>j < (0 # offs ns (Suc (Suc m))) ! i.
\<forall>k \<in> {(0 # offs ns (Suc (Suc m))) ! i..<n}.
key (xs ! j) \<le> key (xs ! k)" and
K: "i < length ms' + length ns'" and
L: "j < offs (ms' @ ns') 0 ! i" and
M: "offs (ms' @ ns') 0 ! i \<le> k" and
N: "k < length ws' + length xs'"
have O: "length ?ws = Suc (Suc m)"
using H and I by simp
with D [symmetric] have P: "length ws' = Suc (Suc m)"
by (simp add: round_suc_suc_def Let_def fill_length split: if_split_asm)
have Q: "\<And>i. i < m \<Longrightarrow>
the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! i) \<in> set ?ws"
proof -
fix i
let ?x = "the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! i)"
assume R: "i < m"
hence "0 < m" by simp
hence "0 < ?k"
by (drule_tac bn_comp_fst_nonzero [OF G], subst (asm) F,
simp split: nat.split)
hence "count (mset (map the (fill ?zs (offs ?ms 0)
index key (length ?zs) ?mi ?ma))) ?x = count (mset ?zs) ?x"
by (rule_tac fill_offs_enum_count_item [OF A], simp, rule_tac conjI,
((rule_tac mini_lb | rule_tac maxi_ub), erule_tac in_set_nthsD)+)
hence "count (mset (map the (fill ?zs (offs ?ms 0)
index key m ?mi ?ma))) ?x = count (mset ?zs) ?x"
using O by (simp add: mini_maxi_nths)
moreover have "0 < count (mset (map the (fill ?zs (offs ?ms 0)
index key m ?mi ?ma))) ?x"
using R by (simp add: fill_length)
ultimately have "?x \<in> set ?zs" by simp
thus "?x \<in> set ?ws"
by (rule in_set_nthsD)
qed
have R: "\<And>i. i < Suc (Suc m) \<Longrightarrow> Suc (Suc m) \<le> k \<Longrightarrow>
key (?ws ! i) \<le> key (xs' ! (k - Suc (Suc m)))"
proof -
fix i
have "bn_inv p q (v, ns, ?ys)"
using C by simp
moreover have "add_inv (n - Suc (Suc m)) (v, ns, ?ys)"
using H and I by simp
moreover have "count_inv (\<lambda>x. count (mset ?ys) x) (v, ns, ?ys)"
by simp
ultimately have "count_inv (\<lambda>x. count (mset ?ys) x) (?t r' v)"
by (rule round_count_inv [OF A])
hence S: "count (mset xs') (xs' ! (k - Suc (Suc m))) =
count (mset ?ys) (xs' ! (k - Suc (Suc m)))"
using E by simp
have "k < Suc (Suc (m + length xs'))"
using N and P by simp
moreover assume "Suc (Suc m) \<le> k"
ultimately have "0 < count (mset xs') (xs' ! (k - Suc (Suc m)))"
by simp
hence "xs' ! (k - Suc (Suc m)) \<in> set ?ys"
using S by simp
hence "\<exists>p < length ?ys. ?ys ! p = xs' ! (k - Suc (Suc m))"
by (simp add: in_set_conv_nth)
then obtain p where
T: "p < length ?ys \<and> ?ys ! p = xs' ! (k - Suc (Suc m))" ..
have "Suc 0 < Suc (length ns) \<longrightarrow>
(\<forall>j < (0 # offs ns (Suc (Suc m))) ! Suc 0.
\<forall>k \<in> {(0 # offs ns (Suc (Suc m))) ! Suc 0..<n}.
key (xs ! j) \<le> key (xs ! k))"
using J ..
moreover have U: "0 < length ns"
using H [symmetric] and I and T by (rule_tac ccontr, simp)
ultimately have "\<forall>j < offs ns (Suc (Suc m)) ! 0.
\<forall>k \<in> {offs ns (Suc (Suc m)) ! 0..<n}. key (xs ! j) \<le> key (xs ! k)"
by simp
moreover assume "i < Suc (Suc m)"
moreover have "offs ns (Suc (Suc m)) ! 0 = Suc (Suc m)"
using U by (subst offs_base_zero, simp, cases ns, simp_all)
ultimately have "\<forall>k \<in> {Suc (Suc m)..<n}. key (?ws ! i) \<le> key (xs ! k)"
by simp
moreover have "Suc (Suc m) + p \<in> {Suc (Suc m)..<n}"
using I and T by (simp, arith)
ultimately have "key (?ws ! i) \<le> key (xs ! (Suc (Suc m) + p))" ..
thus "key (?ws ! i) \<le> key (xs' ! (k - Suc (Suc m)))"
using O and T by simp
qed
assume "\<And>ws a b c d e f g h n'.
ws = ?ws \<Longrightarrow> a = (m', r') \<Longrightarrow> b = m' \<and> c = r' \<Longrightarrow>
d = (v, ms', ws') \<Longrightarrow> e = v \<and> f = (ms', ws') \<Longrightarrow> g = ms' \<and> h = ws' \<Longrightarrow>
?P n' \<longrightarrow> ?Q n' \<longrightarrow> ?R"
hence "?P (n - Suc (Suc m)) \<longrightarrow> ?Q (n - Suc (Suc m)) \<longrightarrow> ?R"
by simp
hence "?Q (n - Suc (Suc m)) \<longrightarrow> ?R"
using H by simp
moreover have "?Q (n - Suc (Suc m))"
proof ((rule allI, rule impI)+, rule ballI, simp, erule conjE,
subst (1 2) append_take_drop_id [of "Suc (Suc m)", symmetric],
simp only: nth_append O, simp)
fix i j k
have "Suc i < Suc (length ns) \<longrightarrow>
(\<forall>j < (0 # offs ns (Suc (Suc m))) ! Suc i.
\<forall>k \<in> {(0 # offs ns (Suc (Suc m))) ! Suc i..<n}.
key ((xs) ! j) \<le> key (xs ! k))"
using J ..
moreover assume "i < length ns"
ultimately have "j < offs ns 0 ! i \<longrightarrow>
(\<forall>k \<in> {offs ns 0 ! i + Suc (Suc m)..<n}.
key (xs ! (Suc (Suc m) + j)) \<le> key (xs ! k))"
by (simp, subst (asm) offs_base_zero, simp,
subst (asm) (2) offs_base_zero, simp_all)
moreover assume "j < offs ns 0 ! i"
ultimately have "\<forall>k \<in> {offs ns 0 ! i + Suc (Suc m)..<n}.
key (?ys ! j) \<le> key (xs ! k)"
using O by simp
moreover assume "offs ns 0 ! i \<le> k" and "k < n - Suc (Suc m)"
hence "Suc (Suc m) + k \<in> {offs ns 0 ! i + Suc (Suc m)..<n}"
by simp
ultimately have "key (?ys ! j) \<le> key (xs ! (Suc (Suc m) + k))" ..
thus "key (?ys ! j) \<le> key (?ys ! k)"
using O by simp
qed
ultimately have ?R ..
show "key ((ws' @ xs') ! j) \<le> key ((ws' @ xs') ! k)"
proof (simp add: nth_append not_less P, (rule_tac [!] conjI,
rule_tac [!] impI)+, rule_tac [3] FalseE)
assume
S: "j < Suc (Suc m)" and
T: "k < Suc (Suc m)"
from D [symmetric] show "key (ws' ! j) \<le> key (ws' ! k)"
proof (simp add: round_suc_suc_def Let_def split: if_split_asm,
(erule_tac [2] conjE)+, simp_all)
assume U: "?mi = ?ma"
have "j < length ?ws"
using S and O by simp
hence "?ws ! j \<in> set ?ws"
by (rule nth_mem)
with U have "key (?ws ! j) = ?ma"
by (rule mini_maxi_keys_eq)
moreover have "k < length ?ws"
using T and O by simp
hence "?ws ! k \<in> set ?ws"
by (rule nth_mem)
with U have "key (?ws ! k) = ?ma"
by (rule mini_maxi_keys_eq)
ultimately show "key (?ws ! j) \<le> key (?ws ! k)" by simp
next
assume U: "ms' = Suc 0 # ?ms @ [Suc 0]"
hence V: "j < (0 # offs (?ms @ Suc 0 # ns') (Suc 0)) ! i"
using L by simp
hence "\<exists>i'. i = Suc i'"
by (rule_tac nat.exhaust [of i], simp_all)
then obtain i' where W: "i = Suc i'" ..
have "i < Suc (Suc (?k + length ns'))"
using K and U by (simp add: enum_length)
hence X: "i' < Suc (?k + length ns')"
using W by simp
hence Y: "j < Suc (offs (?ms @ Suc 0 # ns') 0 ! i')"
using V and W by (simp add: enum_length offs_suc)
have "(0 # offs (?ms @ Suc 0 # ns') (Suc 0)) ! i \<le> k"
using M and U by simp
hence "Suc (offs (?ms @ Suc 0 # ns') 0 ! i') \<le> k"
using W and X by (simp add: enum_length offs_suc)
moreover from this have "\<exists>k'. k = Suc k'"
by (rule_tac nat.exhaust [of k], simp_all)
then obtain k' where Z: "k = Suc k'" ..
ultimately have AA: "offs (?ms @ Suc 0 # ns') 0 ! i' \<le> k'"
by simp
have AB: "j \<le> k'"
using Y and AA by simp
with Z show
"key ((?xmi # map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma) @
[?xma]) ! j) \<le>
key ((?xmi # map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma) @
[?xma]) ! k)"
proof (cases j, case_tac [2] "j < Suc m", simp_all add: nth_append
not_less fill_length, rule_tac [1-2] conjI, rule_tac [1-4] impI,
rule_tac [5] FalseE, simp_all)
assume "k' < m"
hence "the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k') \<in> set ?ws"
by (rule Q)
thus "?mi \<le> key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k'))"
by (rule mini_lb)
next
assume "m \<le> k'"
hence "k' = m"
using T and Z by simp
thus "?mi \<le> key ([?xma] ! (k' - m))"
by (simp, rule_tac mini_maxi_keys_le, rule_tac nth_mem [of 0],
subst O, simp)
next
fix j'
have "sort_inv key (0, ?ms, map the (fill ?zs (offs ?ms 0) index key
(length ?zs) ?mi ?ma))"
by (rule fill_sort_inv [OF A B], simp, rule conjI,
((rule mini_lb | rule maxi_ub), erule in_set_nthsD)+)
hence "sort_inv key (0, ?ms, map the (fill ?zs (offs ?ms 0) index key
m ?mi ?ma))"
using O by (simp add: mini_maxi_nths)
hence "\<forall>i < ?k. \<forall>j < offs ?ms 0 ! i. \<forall>k \<in> {offs ?ms 0 ! i..<m}.
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j)) \<le>
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k))"
by (simp add: enum_length fill_length)
moreover assume AC: "k' < m"
hence AD: "offs (?ms @ Suc 0 # ns') 0 ! i' < m"
using AA by simp
have AE: "i' < ?k"
proof (rule contrapos_pp [OF AD], simp add: not_less offs_append
nth_append offs_length enum_length)
have "0 < m"
using AC by simp
hence "0 < ?k"
by (drule_tac bn_comp_fst_nonzero [OF G], subst (asm) F,
simp split: nat.split)
with A have "foldl (+) 0 ?ms = length ?zs"
by (rule enum_add, simp, rule_tac conjI,
((rule_tac mini_lb | rule_tac maxi_ub), erule_tac in_set_nthsD)+)
hence "foldl (+) 0 ?ms = m"
using O by (simp add: mini_maxi_nths)
with X show "m \<le> (foldl (+) 0 ?ms #
offs ns' (Suc (foldl (+) 0 ?ms))) ! (i' - ?k)"
by (cases "i' - ?k", simp_all, subst offs_base_zero, simp_all)
qed
ultimately have "\<forall>j < offs ?ms 0 ! i'. \<forall>k \<in> {offs ?ms 0 ! i'..<m}.
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j)) \<le>
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k))"
by simp
moreover assume "j = Suc j'"
with Y and AE have "j' < offs ?ms 0 ! i'"
by (simp add: offs_append nth_append offs_length enum_length)
ultimately have "\<forall>k \<in> {offs ?ms 0 ! i'..<m}.
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j')) \<le>
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k))"
by simp
moreover from AA and AE have "offs ?ms 0 ! i' \<le> k'"
by (simp add: offs_append nth_append offs_length enum_length)
hence "k' \<in> {offs ?ms 0 ! i'..<m}"
using AC by simp
ultimately show
"key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j')) \<le>
key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! k'))" ..
next
fix j'
assume "j' < m"
hence "the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j') \<in> set ?ws"
(is "?x \<in> _")
by (rule Q)
moreover assume "m \<le> k'"
hence "k' = m"
using T and Z by simp
ultimately show "key ?x \<le> key ([?xma] ! (k' - m))"
by (simp, rule_tac maxi_ub)
next
fix j'
assume "j = Suc j'" and "m \<le> j'"
hence "Suc m \<le> j" by simp
hence "Suc (Suc m) \<le> k"
using Z and AB by simp
thus False
using T by simp
qed
qed
next
assume
S: "j < Suc (Suc m)" and
T: "Suc (Suc m) \<le> k"
from D [symmetric] show "key (ws' ! j) \<le> key (xs' ! (k - Suc (Suc m)))"
proof (simp add: round_suc_suc_def Let_def split: if_split_asm,
rule_tac R [OF S T], cases j, case_tac [2] "j < Suc m",
simp_all add: nth_append not_less fill_length)
have "?nmi < length ?ws"
using O by (rule_tac mini_less, simp)
hence "?nmi < Suc (Suc m)"
using O by simp
thus "?mi \<le> key (xs' ! (k - Suc (Suc m)))"
using T by (rule R)
next
fix j'
assume "j' < m"
hence U: "the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j') \<in> set ?ws"
by (rule Q)
have "\<exists>p < Suc (Suc m). ?ws ! p =
the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j')"
using O and U by (simp add: in_set_conv_nth)
then obtain p where V: "p < Suc (Suc m) \<and> ?ws ! p =
the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j')" ..
hence "key (?ws ! p) \<le> key (xs' ! (k - Suc (Suc m)))"
using T by (rule_tac R, simp)
thus "key (the (fill ?zs (offs ?ms 0) index key m ?mi ?ma ! j')) \<le>
key (xs' ! (k - Suc (Suc m)))"
using V by simp
next
fix j'
assume "j = Suc j'" and "m \<le> j'"
moreover from this have "Suc m \<le> j" by simp
hence "j = Suc m"
using S by simp
ultimately have "j' = m" by simp
thus "key ([?xma] ! (j' - m)) \<le> key (xs' ! (k - Suc (Suc m)))"
proof simp
have "?nma < length ?ws"
using O by (rule_tac maxi_less, simp)
hence "?nma < Suc (Suc m)"
using O by simp
thus "?ma \<le> key (xs' ! (k - Suc (Suc m)))"
using T by (rule R)
qed
qed
next
assume "k < Suc (Suc m)" and "Suc (Suc m) \<le> j"
hence "k < j" by simp
moreover have "j < k"
using L and M by simp
ultimately show False by simp
next
assume
S: "Suc (Suc m) \<le> j" and
T: "Suc (Suc m) \<le> k"
have U: "0 < length ns' \<and> 0 < length xs'"
proof (rule ccontr, simp, erule disjE)
have "bn_inv p q (v, ns, ?ys)"
using C by simp
moreover have "add_inv (n - Suc (Suc m)) (v, ns, ?ys)"
using H and I by simp
ultimately have "add_inv (n - Suc (Suc m)) (?t r' v)"
by (rule round_add_inv [OF A])
hence "length xs' = foldl (+) 0 ns'"
using E by simp
moreover assume "ns' = []"
ultimately have "length xs' = 0" by simp
hence "k < Suc (Suc m)"
using N and P by simp
thus False
using T by simp
next
assume "xs' = []"
hence "k < Suc (Suc m)"
using N and P by simp
thus False
using T by simp
qed
hence V: "i - length ms' < length ns'"
using K by arith
hence W: "\<forall>j < offs ns' 0 ! (i - length ms').
\<forall>k \<in> {offs ns' 0 ! (i - length ms')..<length xs'}.
key (xs' ! j) \<le> key (xs' ! k)"
using `?R` by simp
from D [symmetric] have X: "foldl (+) 0 ms' = Suc (Suc m)"
proof (simp add: round_suc_suc_def Let_def add_replicate add_suc
split: if_split_asm, cases "m = 0")
case True
hence "?k = 0" by simp
hence "length ?ms = 0"
by (simp add: enum_length)
thus "foldl (+) 0 ?ms = m"
using True by simp
next
case False
hence "0 < ?k"
by (simp, drule_tac bn_comp_fst_nonzero [OF G], subst (asm) F,
simp split: nat.split)
with A have "foldl (+) 0 ?ms = length ?zs"
by (rule enum_add, simp, rule_tac conjI,
((rule_tac mini_lb | rule_tac maxi_ub), erule_tac in_set_nthsD)+)
thus "foldl (+) 0 ?ms = m"
using O by (simp add: mini_maxi_nths)
qed
have "length ms' < i"
proof (rule ccontr, simp add: not_less)
assume "i \<le> length ms'"
moreover have "length ms' < length (ms' @ ns')"
using U by simp
ultimately have "offs (ms' @ ns') 0 ! i \<le>
offs (ms' @ ns') 0 ! (length ms')"
by (rule offs_mono)
also have "\<dots> = Suc (Suc m)"
using U and X by (simp add: offs_append nth_append offs_length,
cases ns', simp_all)
finally have "offs (ms' @ ns') 0 ! i \<le> Suc (Suc m)" .
hence "j < Suc (Suc m)"
using L by simp
thus False
using S by simp
qed
hence Y: "offs (ms' @ ns') 0 ! i =
Suc (Suc m) + offs ns' 0 ! (i - length ms')"
using X by (simp add: offs_append nth_append offs_length,
subst offs_base_zero [OF V], simp)
hence "j < Suc (Suc m) + offs ns' 0 ! (i - length ms')"
using L by simp
moreover from this have "0 < offs ns' 0 ! (i - length ms')"
using S by (rule_tac ccontr, simp)
ultimately have "j - Suc (Suc m) < offs ns' 0 ! (i - length ms')"
by simp
hence Z: "\<forall>k \<in> {offs ns' 0 ! (i - length ms')..<length xs'}.
key (xs' ! (j - Suc (Suc m))) \<le> key (xs' ! k)"
using W by simp
have "offs ns' 0 ! (i - length ms') \<le> k - Suc (Suc m)"
using M and Y by simp
moreover have "k - Suc (Suc m) < length xs'"
using N and P and U by arith
ultimately have "k - Suc (Suc m) \<in>
{offs ns' 0 ! (i - length ms')..<length xs'}"
by simp
with Z show "key (xs' ! (j - Suc (Suc m))) \<le>
key (xs' ! (k - Suc (Suc m)))" ..
qed
qed
qed
lemma gcsort_sort_inv:
assumes
A: "index_less index key" and
B: "index_mono index key" and
C: "add_inv n t" and
D: "n \<le> p"
shows "\<lbrakk>t' \<in> gcsort_set index key p t; sort_inv key t\<rbrakk> \<Longrightarrow>
sort_inv key t'"
by (erule gcsort_set.induct, simp, drule gcsort_add_inv [OF A _ C D],
rule round_sort_inv [OF A B], simp_all del: bn_inv.simps, erule conjE,
frule sym, erule subst, rule bn_inv_intro, insert D, simp)
text \<open>
\null
The only remaining task is to address step 10 of the proof method, which is done by proving theorem
@{text gcsort_sorted}. It holds under the conditions that the objects' number is not larger than the
counters' upper bound and function @{text index} satisfies both predicates @{const index_less} and
@{const index_mono}, and states that function @{const gcsort} is successful in sorting the input
objects' list.
\null
\<close>
theorem gcsort_sorted:
assumes
A: "index_less index key" and
B: "index_mono index key" and
C: "length xs \<le> p"
shows "sorted (map key (gcsort index key p xs))"
proof -
let ?t = "(0, [length xs], xs)"
have "sort_inv key (gcsort_aux index key p ?t)"
by (rule gcsort_sort_inv [OF A B _ C], rule gcsort_add_input,
rule gcsort_aux_set, rule gcsort_sort_input)
moreover have "add_inv (length xs) (gcsort_aux index key p ?t)"
by (rule gcsort_add_inv [OF A _ _ C],
rule gcsort_aux_set, rule gcsort_add_input)
ultimately have "sorted (map key (gcsort_out (gcsort_aux index key p ?t)))"
by (rule gcsort_sort_intro, simp add: gcsort_sort_form)
moreover have "?t = gcsort_in xs"
by (simp add: gcsort_in_def)
ultimately show ?thesis
by (simp add: gcsort_def)
qed
end
diff --git a/thys/Generalized_Counting_Sort/Stability.thy b/thys/Generalized_Counting_Sort/Stability.thy
--- a/thys/Generalized_Counting_Sort/Stability.thy
+++ b/thys/Generalized_Counting_Sort/Stability.thy
@@ -1,763 +1,764 @@
(* Title: An Efficient Generalization of Counting Sort for Large, possibly Infinite Key Ranges
Author: Pasquale Noce
Software Engineer at HID Global, Italy
pasquale dot noce dot lavoro at gmail dot com
pasquale dot noce at hidglobal dot com
*)
section "Proof of algorithm's stability"
theory Stability
imports Sorting
begin
text \<open>
\null
In this section, it is formally proven that GCsort is \emph{stable}, viz. that the sublist of the
output of function @{const gcsort} built by picking out the objects having a given key matches the
sublist of the input objects' list built in the same way.
Here below, steps 5, 6, and 7 of the proof method are accomplished. Particularly, @{text stab_inv}
is the predicate that will be shown to be invariant over inductive set @{const gcsort_set}.
\null
\<close>
fun stab_inv :: "('b \<Rightarrow> 'a list) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> nat \<times> nat list \<times> 'a list \<Rightarrow>
bool" where
"stab_inv f key (u, ns, xs) = (\<forall>k. [x\<leftarrow>xs. key x = k] = f k)"
lemma gcsort_stab_input:
"stab_inv (\<lambda>k. [x\<leftarrow>xs. key x = k]) key (0, [length xs], xs)"
by simp
lemma gcsort_stab_intro:
"stab_inv f key t \<Longrightarrow> [x\<leftarrow>gcsort_out t. key x = k] = f k"
by (cases t, simp add: gcsort_out_def)
text \<open>
\null
In what follows, step 9 of the proof method is accomplished.
First, lemma @{text fill_offs_enum_stable} proves that function @{const fill}, if its input offsets'
list is computed via the composition of functions @{const offs} and @{const enum}, does not modify
the sublist of its input objects' list formed by the objects having a given key. Moreover, lemmas
@{text mini_stable} and @{text maxi_stable} prove that the extraction of the leftmost minimum and
the rightmost maximum from an objects' list through functions @{const mini} and @{const maxi} is
endowed with the same property.
These lemmas are then used to prove lemma @{text gcsort_stab_inv}, which states that the sublist of
the objects having a given key within the objects' list is still the same after any recursive round.
\null
\<close>
lemma fill_stable [rule_format]:
assumes
A: "index_less index key" and
B: "index_same index key"
shows
"(\<forall>x \<in> set xs. key x \<in> {mi..ma}) \<longrightarrow>
ns \<noteq> [] \<longrightarrow>
offs_pred ns ub xs index key mi ma \<longrightarrow>
map the [w\<leftarrow>fill xs ns index key ub mi ma. \<exists>x. w = Some x \<and> key x = k] =
[x\<leftarrow>xs. k = key x]"
proof (induction xs arbitrary: ns, simp_all add: Let_def, rule conjI,
(rule_tac [!] impI)+, (erule_tac [!] conjE)+, simp_all)
fix x xs and ns :: "nat list"
let ?i = "index key x (length ns) mi ma"
let ?ns' = "ns[?i := Suc (ns ! ?i)]"
let ?ws = "[w\<leftarrow>fill xs ?ns' index key ub mi ma.
\<exists>y. w = Some y \<and> key y = key x]"
let ?ws' = "[w\<leftarrow>(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x].
\<exists>y. w = Some y \<and> key y = key x]"
assume
C: "\<forall>x \<in> set xs. mi \<le> key x \<and> key x \<le> ma" and
D: "mi \<le> key x" and
E: "key x \<le> ma" and
F: "ns \<noteq> []" and
G: "offs_pred ns ub (x # xs) index key mi ma"
have H: "?i < length ns"
using A and D and E and F by (simp add: index_less_def)
assume "\<And>ns. ns \<noteq> [] \<longrightarrow> offs_pred ns ub xs index key mi ma \<longrightarrow>
map the [w\<leftarrow>fill xs ns index key ub mi ma.
\<exists>y. w = Some y \<and> key y = key x] =
[y\<leftarrow>xs. key x = key y]"
moreover have I: "offs_pred ?ns' ub xs index key mi ma"
using G and H by (rule_tac offs_pred_cons, simp_all)
ultimately have J: "map the ?ws = [y\<leftarrow>xs. key x = key y]"
using F by simp
have "ns ! ?i + offs_num (length ns) (x # xs) index key mi ma ?i \<le> ub"
using G and H by (rule offs_pred_ub, simp add: offs_num_cons)
hence K: "ns ! ?i < ub"
by (simp add: offs_num_cons)
moreover from this have
L: "(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x] ! (ns ! ?i) = Some x"
by (subst nth_list_update_eq, simp_all add: fill_length)
ultimately have "0 < length ?ws'"
by (simp add: length_filter_conv_card card_gt_0_iff,
rule_tac exI [where x = "ns ! ?i"], simp add: fill_length)
hence "\<exists>w ws. ?ws' = w # ws"
by (rule_tac list.exhaust [of ?ws'], simp_all)
then obtain w and ws where "?ws' = w # ws" by blast
hence "\<exists>us vs y.
(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x] = us @ w # vs \<and>
(\<forall>u \<in> set us. \<forall>y. u = Some y \<longrightarrow> key y \<noteq> key x) \<and>
w = Some y \<and> key y = key x \<and>
ws = [w\<leftarrow>vs. \<exists>y. w = Some y \<and> key y = key x]"
(is "\<exists>us vs y. ?P us vs y")
by (simp add: filter_eq_Cons_iff, blast)
then obtain us and vs and y where M: "?P us vs y" by blast
let ?A = "{i. i < length ns \<and> 0 < offs_num (length ns) xs index key mi ma i \<and>
ns ! i \<le> length us}"
have "length us = ns ! ?i"
proof (rule ccontr, erule neqE, cases "?A = {}",
cases "0 < offs_num (length ns) xs index key mi ma 0", simp_all only: not_gr0)
assume
N: "length us < ns ! ?i" and
O: "?A = {}" and
P: "0 < offs_num (length ns) xs index key mi ma 0"
have "length us < ns ! 0"
proof (rule ccontr, simp add: not_less)
assume "ns ! 0 \<le> length us"
with F and P have "0 \<in> ?A" by simp
with O show False by blast
qed
hence "length us < ?ns' ! 0"
using F by (cases ?i, simp_all)
hence "offs_none ?ns' ub xs index key mi ma (length us)"
using P by (simp add: offs_none_def)
hence "fill xs ?ns' index key ub mi ma ! (length us) = None"
using C and F and I by (rule_tac fill_none [OF A], simp_all)
hence "(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x] ! (length us) = None"
using N by simp
thus False
using M by simp
next
assume
N: "length us < ns ! ?i" and
O: "?A = {}" and
P: "offs_num (length ns) xs index key mi ma 0 = 0"
from K and N have "length us < offs_next ?ns' ub xs index key mi ma 0"
proof (rule_tac ccontr, simp only: not_less offs_next_def split: if_split_asm)
assume "offs_set_next ?ns' xs index key mi ma 0 \<noteq> {}"
(is "?B \<noteq> _")
hence "Min ?B \<in> ?B"
by (rule_tac Min_in, simp)
moreover assume "?ns' ! Min ?B \<le> length us"
hence "ns ! Min ?B \<le> length us"
using H by (cases "Min ?B = ?i", simp_all)
ultimately have "Min ?B \<in> ?A" by simp
with O show False by blast
qed
hence "offs_none ?ns' ub xs index key mi ma (length us)"
using P by (simp add: offs_none_def)
hence "fill xs ?ns' index key ub mi ma ! (length us) = None"
using C and F and I by (rule_tac fill_none [OF A], simp_all)
hence "(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x] ! (length us) = None"
using N by simp
thus False
using M by simp
next
assume N: "length us < ns ! ?i"
assume "?A \<noteq> {}"
hence "Max ?A \<in> ?A"
by (rule_tac Max_in, simp)
moreover from this have O: "Max ?A \<noteq> ?i"
using N by (rule_tac notI, simp)
ultimately have "index key y (length ?ns') mi ma = Max ?A"
using C proof (rule_tac fill_index [OF A _ I, where j = "length us"], simp_all,
rule_tac ccontr, simp only: not_less not_le offs_next_def split: if_split_asm)
assume "ub \<le> length us"
with N have "ub < ns ! ?i" by simp
with K show False by simp
next
let ?B = "offs_set_next ?ns' xs index key mi ma (Max ?A)"
assume "?ns' ! Min ?B \<le> length us"
hence "ns ! Min ?B \<le> length us"
using H by (cases "Min ?B = ?i", simp_all)
moreover assume "?B \<noteq> {}"
hence P: "Min ?B \<in> ?B"
by (rule_tac Min_in, simp)
ultimately have "Min ?B \<in> ?A" by simp
hence "Min ?B \<le> Max ?A"
by (rule_tac Max_ge, simp)
thus False
using P by simp
next
have "fill xs ?ns' index key ub mi ma ! length us =
(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x] ! length us"
using N by simp
thus "fill xs ?ns' index key ub mi ma ! length us = Some y"
using M by simp
qed
moreover have "index key y (length ?ns') mi ma = ?i"
using B and D and E and M by (cases "y = x", simp_all add:
index_same_def)
ultimately show False
using O by simp
next
assume N: "ns ! ?i < length us"
hence "(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x] ! (ns ! ?i) =
us ! (ns ! ?i)"
using M by (simp add: nth_append)
hence "(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x] ! (ns ! ?i) \<in> set us"
using N by simp
thus False
using L and M by blast
qed
hence "take (ns ! ?i) ((fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x]) = us"
using M by simp
hence N: "take (ns ! ?i) (fill xs ?ns' index key ub mi ma) = us" by simp
have "fill xs ?ns' index key ub mi ma =
take (ns ! ?i) (fill xs ?ns' index key ub mi ma) @
fill xs ?ns' index key ub mi ma ! (ns ! ?i) #
drop (Suc (ns ! ?i)) (fill xs ?ns' index key ub mi ma)"
(is "_ = ?ts @ _ # ?ds")
using K by (rule_tac id_take_nth_drop, simp add: fill_length)
moreover have "fill xs ?ns' index key ub mi ma ! (ns ! ?i) = None"
using C and D and E and F and G by (rule_tac fill_index_none [OF A],
simp_all)
ultimately have O: "fill xs ?ns' index key ub mi ma = us @ None # ?ds"
using N by simp
have "(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x] =
?ts @ Some x # ?ds"
using K by (rule_tac upd_conv_take_nth_drop, simp add: fill_length)
hence "(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x] =
us @ Some x # ?ds"
using N by simp
hence "?ws' = Some x # [w\<leftarrow>?ds. \<exists>y. w = Some y \<and> key y = key x]"
using M by simp
also have "\<dots> = Some x # ?ws"
using M by (subst (2) O, simp)
finally show "map the ?ws' = x # [y\<leftarrow>xs. key x = key y]"
using J by simp
next
fix x xs and ns :: "nat list"
let ?i = "index key x (length ns) mi ma"
let ?ns' = "ns[?i := Suc (ns ! ?i)]"
let ?ws = "[w\<leftarrow>fill xs ?ns' index key ub mi ma. \<exists>y. w = Some y \<and> key y = k]"
let ?ws' = "[w\<leftarrow>(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x].
\<exists>y. w = Some y \<and> key y = k]"
assume
C: "\<forall>x \<in> set xs. mi \<le> key x \<and> key x \<le> ma" and
D: "mi \<le> key x" and
E: "key x \<le> ma" and
F: "ns \<noteq> []" and
G: "offs_pred ns ub (x # xs) index key mi ma" and
H: "k \<noteq> key x"
have I: "?i < length ns"
using A and D and E and F by (simp add: index_less_def)
assume "\<And>ns. ns \<noteq> [] \<longrightarrow> offs_pred ns ub xs index key mi ma \<longrightarrow>
map the [w\<leftarrow>fill xs ns index key ub mi ma.
\<exists>y. w = Some y \<and> key y = k] =
[y\<leftarrow>xs. k = key y]"
moreover have "offs_pred ?ns' ub xs index key mi ma"
using G and I by (rule_tac offs_pred_cons, simp_all)
ultimately have J: "map the ?ws = [y\<leftarrow>xs. k = key y]"
using F by simp
have "ns ! ?i + offs_num (length ns) (x # xs) index key mi ma ?i \<le> ub"
using G and I by (rule offs_pred_ub, simp add: offs_num_cons)
hence K: "ns ! ?i < ub"
by (simp add: offs_num_cons)
have "fill xs ?ns' index key ub mi ma =
take (ns ! ?i) (fill xs ?ns' index key ub mi ma) @
fill xs ?ns' index key ub mi ma ! (ns ! ?i) #
drop (Suc (ns ! ?i)) (fill xs ?ns' index key ub mi ma)"
(is "_ = ?ts @ _ # ?ds")
using K by (rule_tac id_take_nth_drop, simp add: fill_length)
moreover have "fill xs ?ns' index key ub mi ma ! (ns ! ?i) = None"
using C and D and E and F and G by (rule_tac fill_index_none [OF A],
simp_all)
ultimately have L: "fill xs ?ns' index key ub mi ma = ?ts @ None # ?ds"
by simp
have "(fill xs ?ns' index key ub mi ma)[ns ! ?i := Some x] =
?ts @ Some x # ?ds"
using K by (rule_tac upd_conv_take_nth_drop, simp add: fill_length)
moreover have "?ws = [w\<leftarrow>?ts. \<exists>y. w = Some y \<and> key y = k] @
[w\<leftarrow>?ds. \<exists>y. w = Some y \<and> key y = k]"
by (subst L, simp)
ultimately have "?ws' = ?ws"
using H by simp
thus "map the ?ws' = [y\<leftarrow>xs. k = key y]"
using J by simp
qed
lemma fill_offs_enum_stable [rule_format]:
assumes
A: "index_less index key" and
B: "index_same index key"
shows
"\<forall>x \<in> set xs. key x \<in> {mi..ma} \<Longrightarrow>
0 < n \<Longrightarrow>
[x\<leftarrow>map the (fill xs (offs (enum xs index key n mi ma) 0)
index key (length xs) mi ma). key x = k] = [x\<leftarrow>xs. k = key x]"
(is "_ \<Longrightarrow> _ \<Longrightarrow> [_\<leftarrow>map the ?ys. _] = _"
is "_ \<Longrightarrow> _ \<Longrightarrow> [_\<leftarrow>map the (fill _ ?ns _ _ _ _ _). _] = _")
proof (subst fill_stable [symmetric, OF A B, of xs mi ma ?ns], simp,
simp only: length_greater_0_conv [symmetric] offs_length enum_length,
rule offs_enum_pred [OF A], simp, subst filter_map,
simp add: filter_eq_nths fill_length)
assume
C: "\<forall>x \<in> set xs. mi \<le> key x \<and> key x \<le> ma" and
D: "0 < n"
have "{i. i < length xs \<and> key (the (?ys ! i)) = k} =
{i. i < length xs \<and> (\<exists>x. ?ys ! i = Some x \<and> key x = k)}"
(is "?A = ?B")
proof (rule set_eqI, rule iffI, simp_all, erule_tac [!] conjE, erule_tac [2] exE,
erule_tac [2] conjE, simp_all)
fix i
assume E: "i < length xs" and F: "key (the (?ys ! i)) = k"
have "\<exists>x. ?ys ! i = Some x"
proof (cases "?ys ! i", simp_all)
assume "?ys ! i = None"
moreover have "?ys ! i \<in> set ?ys"
using E by (rule_tac nth_mem, simp add: fill_length)
ultimately have "None \<in> set ?ys" by simp
moreover have "count (mset ?ys) None = 0"
using C and D by (rule_tac fill_offs_enum_count_none [OF A], simp)
ultimately show False by simp
qed
then obtain x where "?ys ! i = Some x" ..
moreover from this have "key x = k"
using F by simp
ultimately show "\<exists>x. ?ys ! i = Some x \<and> key x = k" by simp
qed
thus "map the (nths ?ys ?A) = map the (nths ?ys ?B)" by simp
qed
lemma mini_first [rule_format]:
"xs \<noteq> [] \<longrightarrow> i < mini xs key \<longrightarrow>
key (xs ! mini xs key) < key (xs ! i)"
by (induction xs arbitrary: i, simp_all add: Let_def, (rule impI)+,
erule conjE, simp add: not_le nth_Cons split: nat.split)
lemma maxi_last [rule_format]:
"xs \<noteq> [] \<longrightarrow> maxi xs key < i \<longrightarrow> i < length xs \<longrightarrow>
key (xs ! i) < key (xs ! maxi xs key)"
by (induction xs arbitrary: i, simp_all add: Let_def, (rule impI)+,
rule le_less_trans, rule maxi_ub, rule nth_mem, simp)
lemma nths_range:
"nths xs A = nths xs (A \<inter> {..<length xs})"
proof (induction xs arbitrary: A, simp_all add: nths_Cons)
fix xs :: "'a list" and A
assume "\<And>A. nths xs A = nths xs (A \<inter> {..<length xs})"
hence "nths xs {i. Suc i \<in> A} = nths xs ({i. Suc i \<in> A} \<inter> {..<length xs})" .
moreover have
"{i. Suc i \<in> A} \<inter> {..<length xs} = {i. Suc i \<in> A \<and> i < length xs}"
by blast
ultimately show
"nths xs {i. Suc i \<in> A} = nths xs {i. Suc i \<in> A \<and> i < length xs}"
by simp
qed
lemma filter_nths_diff:
assumes
A: "i < length xs" and
B: "\<not> P (xs ! i)"
shows "[x\<leftarrow>nths xs (A - {i}). P x] = [x\<leftarrow>nths xs A. P x]"
proof (cases "i \<in> A", simp_all)
case True
have C: "xs = take i xs @ xs ! i # drop (Suc i) xs"
(is "_ = ?ts @ _ # ?ds")
using A by (rule id_take_nth_drop)
have "nths xs A = nths ?ts A @ nths (xs ! i # ?ds) {j. j + length ?ts \<in> A}"
by (subst C, simp add: nths_append)
moreover have D: "length ?ts = i"
using A by simp
ultimately have E: "nths xs A = nths ?ts (A \<inter> {..<i}) @ xs ! i #
nths ?ds {j. Suc j + i \<in> A}"
(is "_ = ?vs @ _ # ?ws")
using True by (simp add: nths_Cons, subst nths_range, simp)
have "nths xs (A - {i}) = nths ?ts (A - {i}) @
nths (xs ! i # ?ds) {j. j + length ?ts \<in> A - {i}}"
by (subst C, simp add: nths_append)
moreover have "(A - {i}) \<inter> {..<i} = A \<inter> {..<i}"
by (rule set_eqI, rule iffI, simp_all)
ultimately have "nths xs (A - {i}) = ?vs @ ?ws"
using D by (simp add: nths_Cons, subst nths_range, simp)
thus ?thesis
using B and E by simp
qed
lemma mini_stable:
assumes
A: "xs \<noteq> []" and
B: "mini xs key \<in> A"
(is "?nmi \<in> _")
shows "[x\<leftarrow>[xs ! ?nmi] @ nths xs (A - {?nmi}). key x = k] =
[x\<leftarrow>nths xs A. key x = k]"
(is "[x\<leftarrow>[?xmi] @ _. _] = _")
proof (simp, rule conjI, rule_tac [!] impI, rule_tac [2] filter_nths_diff,
rule_tac [2] mini_less, simp_all add: A)
assume C: "key ?xmi = k"
moreover have "?nmi < length xs"
using A by (rule_tac mini_less, simp)
ultimately have "0 < length [x\<leftarrow>xs. key x = k]"
by (simp add: length_filter_conv_card card_gt_0_iff, rule_tac exI, rule_tac conjI)
hence "\<exists>y ys. [x\<leftarrow>xs. key x = k] = y # ys"
by (rule_tac list.exhaust [of "[x\<leftarrow>xs. key x = k]"], simp_all)
then obtain y and ys where "[x\<leftarrow>xs. key x = k] = y # ys" by blast
hence "\<exists>us vs. xs = us @ y # vs \<and>
(\<forall>u \<in> set us. key u \<noteq> k) \<and> key y = k \<and> ys = [x\<leftarrow>vs. key x = k]"
(is "\<exists>us vs. ?P us vs")
by (simp add: filter_eq_Cons_iff)
then obtain us and vs where D: "?P us vs" by blast
have E: "length us = ?nmi"
proof (rule ccontr, erule neqE)
assume "length us < ?nmi"
with A have "key ?xmi < key (xs ! length us)"
by (rule mini_first)
also have "\<dots> = k"
using D by simp
finally show False
using C by simp
next
assume F: "?nmi < length us"
hence "?xmi = us ! ?nmi"
using D by (simp add: nth_append)
hence "?xmi \<in> set us"
using F by simp
thus False
using C and D by simp
qed
moreover have "xs ! length us = y"
using D by simp
ultimately have F: "?xmi = y" by simp
have "nths xs A = nths us A @ nths (y # vs) {i. i + ?nmi \<in> A}"
using D and E by (simp add: nths_append)
also have "\<dots> = nths us A @ y # nths vs {i. Suc i + ?nmi \<in> A}"
(is "_ = _ @ _ # ?ws")
using B by (simp add: nths_Cons)
finally have G: "[x\<leftarrow>nths xs A. key x = k] = y # [x\<leftarrow>?ws. key x = k]"
using D by (simp add: filter_empty_conv, rule_tac ballI,
drule_tac in_set_nthsD, simp)
have "nths xs (A - {?nmi}) = nths us (A - {?nmi}) @
nths (y # vs) {i. i + ?nmi \<in> A - {?nmi}}"
using D and E by (simp add: nths_append)
also have "\<dots> = nths us (A - {?nmi}) @ ?ws"
by (simp add: nths_Cons)
finally have H: "[x\<leftarrow>nths xs (A - {?nmi}). key x = k] = [x\<leftarrow>?ws. key x = k]"
using D by (simp add: filter_empty_conv, rule_tac ballI,
drule_tac in_set_nthsD, simp)
show "?xmi # [x\<leftarrow>nths xs (A - {?nmi}). key x = k] =
[x\<leftarrow>nths xs A. key x = k]"
using F and G and H by simp
qed
lemma maxi_stable:
assumes
A: "xs \<noteq> []" and
B: "maxi xs key \<in> A"
(is "?nma \<in> _")
shows "[x\<leftarrow>nths xs (A - {?nma}) @ [xs ! ?nma]. key x = k] =
[x\<leftarrow>nths xs A. key x = k]"
(is "[x\<leftarrow>_ @ [?xma]. _] = _")
proof (simp, rule conjI, rule_tac [!] impI, rule_tac [2] filter_nths_diff,
rule_tac [2] maxi_less, simp_all add: A)
assume C: "key ?xma = k"
moreover have D: "?nma < length xs"
using A by (rule_tac maxi_less, simp)
ultimately have "0 < length [x\<leftarrow>rev xs. key x = k]"
by (simp add: length_filter_conv_card card_gt_0_iff,
rule_tac exI [where x = "length xs - Suc ?nma"], simp add: rev_nth)
hence "\<exists>y ys. [x\<leftarrow>rev xs. key x = k] = y # ys"
by (rule_tac list.exhaust [of "[x\<leftarrow>rev xs. key x = k]"], simp_all)
then obtain y and ys where "[x\<leftarrow>rev xs. key x = k] = y # ys" by blast
hence "\<exists>us vs. rev xs = us @ y # vs \<and>
(\<forall>u \<in> set us. key u \<noteq> k) \<and> key y = k \<and> ys = [x\<leftarrow>vs. key x = k]"
(is "\<exists>us vs. ?P us vs")
by (simp add: filter_eq_Cons_iff)
then obtain us and vs where E: "?P us vs" by blast
hence F: "xs = rev vs @ y # rev us"
by (simp add: rev_swap)
have G: "length us = length xs - Suc ?nma"
proof (rule ccontr, erule neqE)
assume "length us < length xs - Suc ?nma"
hence "?nma < length xs - Suc (length us)" by simp
moreover have "length xs - Suc (length us) < length xs"
using D by simp
ultimately have "key (xs ! (length xs - Suc (length us))) < key ?xma"
by (rule maxi_last [OF A])
moreover have "length us < length xs"
using F by simp
ultimately have "key (rev xs ! length us) < key ?xma"
by (simp add: rev_nth)
moreover have "key (rev xs ! length us) = k"
using E by simp
ultimately show False
using C by simp
next
assume H: "length xs - Suc ?nma < length us"
hence "rev xs ! (length xs - Suc ?nma) = us ! (length xs - Suc ?nma)"
using E by (simp add: nth_append)
hence "rev xs ! (length xs - Suc ?nma) \<in> set us"
using H by simp
hence "?xma \<in> set us"
using D by (simp add: rev_nth)
thus False
using C and E by simp
qed
moreover have "rev xs ! length us = y"
using E by simp
ultimately have H: "?xma = y"
using D by (simp add: rev_nth)
have "length xs = length us + Suc ?nma"
using D and G by simp
hence I: "length vs = ?nma"
using F by simp
hence "nths xs A = nths (rev vs) A @ nths (y # rev us) {i. i + ?nma \<in> A}"
using F by (simp add: nths_append)
also have "\<dots> = nths (rev vs) (A \<inter> {..<?nma}) @ y #
nths (rev us) {i. Suc i + ?nma \<in> A}"
(is "_ = ?ws @ _ # _")
using B and I by (simp add: nths_Cons, subst nths_range, simp)
finally have J: "[x\<leftarrow>nths xs A. key x = k] = [x\<leftarrow>?ws. key x = k] @ [y]"
using E by (simp add: filter_empty_conv, rule_tac ballI,
drule_tac in_set_nthsD, simp)
have "nths xs (A - {?nma}) = nths (rev vs) (A - {?nma}) @
nths (y # rev us) {i. i + ?nma \<in> A - {?nma}}"
using F and I by (simp add: nths_append)
hence "nths xs (A - {?nma}) = nths (rev vs) ((A - {?nma}) \<inter> {..<?nma}) @
nths (rev us) {i. Suc i + ?nma \<in> A}"
using I by (simp add: nths_Cons, subst nths_range, simp)
moreover have "(A - {?nma}) \<inter> {..<?nma} = A \<inter> {..<?nma}"
by (rule set_eqI, rule iffI, simp_all)
ultimately have K: "[x\<leftarrow>nths xs (A - {?nma}). key x = k] =
[x\<leftarrow>?ws. key x = k]"
using E by (simp add: filter_empty_conv, rule_tac ballI,
drule_tac in_set_nthsD, simp)
show "[x\<leftarrow>nths xs (A - {?nma}). key x = k] @ [?xma] =
[x\<leftarrow>nths xs A. key x = k]"
using H and J and K by simp
qed
lemma round_stab_inv [rule_format]:
"index_less index key \<longrightarrow> index_same index key \<longrightarrow> bn_inv p q t \<longrightarrow>
add_inv n t \<longrightarrow> stab_inv f key t \<longrightarrow> stab_inv f key (round index key p q r t)"
proof (induction index key p q r t arbitrary: n f rule: round.induct,
(rule_tac [!] impI)+, simp, simp, simp_all only: simp_thms)
fix index p q r u ns xs n f and key :: "'a \<Rightarrow> 'b"
let ?t = "round index key p q r (u, ns, tl xs)"
assume
"\<And>n f. bn_inv p q (u, ns, tl xs) \<longrightarrow> add_inv n (u, ns, tl xs) \<longrightarrow>
stab_inv f key (u, ns, tl xs) \<longrightarrow> stab_inv f key ?t" and
"bn_inv p q (u, Suc 0 # ns, xs)" and
"add_inv n (u, Suc 0 # ns, xs)" and
"stab_inv f key (u, Suc 0 # ns, xs)"
thus "stab_inv f key (round index key p q r (u, Suc 0 # ns, xs))"
proof (cases ?t, cases xs, simp_all add: add_suc, arith, erule_tac conjE,
rule_tac allI, simp)
fix k y ys xs'
let ?f' = "f(key y := tl (f (key y)))"
assume "\<And>n' f'. foldl (+) 0 ns = n' \<and> length ys = n' \<longrightarrow>
(\<forall>k'. [x\<leftarrow>ys. key x = k'] = f' k') \<longrightarrow> (\<forall>k'. [x\<leftarrow>xs'. key x = k'] = f' k')"
moreover assume "Suc (foldl (+) 0 ns) = n" and "Suc (length ys) = n"
ultimately have "(\<forall>k'. [x\<leftarrow>ys. key x = k'] = ?f' k') \<longrightarrow>
(\<forall>k'. [x\<leftarrow>xs'. key x = k'] = ?f' k')"
by blast
moreover assume A: "\<forall>k'. (if key y = k'
then y # [x\<leftarrow>ys. key x = k'] else [x\<leftarrow>ys. key x = k']) = f k'"
hence B: "f (key y) = y # [x\<leftarrow>ys. key x = key y]"
by (drule_tac spec [where x = "key y"], simp)
hence "\<forall>k'. [x\<leftarrow>ys. key x = k'] = ?f' k'"
proof (rule_tac allI, simp, rule_tac impI)
fix k'
assume "k' \<noteq> key y"
thus "[x\<leftarrow>ys. key x = k'] = f k'"
using A by (drule_tac spec [where x = k'], simp)
qed
ultimately have C: "\<forall>k'. [x\<leftarrow>xs'. key x = k'] = ?f' k'" ..
show "(key y = k \<longrightarrow> y # [x\<leftarrow>xs'. key x = k] = f k) \<and>
(key y \<noteq> k \<longrightarrow> [x\<leftarrow>xs'. key x = k] = f k)"
proof (rule conjI, rule_tac [!] impI)
assume "key y = k"
moreover have "tl (f (key y)) = [x\<leftarrow>xs'. key x = key y]"
using C by simp
hence "f (key y) = y # [x\<leftarrow>xs'. key x = key y]"
using B by (subst hd_Cons_tl [symmetric], simp_all)
ultimately show "y # [x\<leftarrow>xs'. key x = k] = f k" by simp
next
assume "key y \<noteq> k"
thus "[x\<leftarrow>xs'. key x = k] = f k"
using C by simp
qed
qed
next
fix index p q r u m ns n f and key :: "'a \<Rightarrow> 'b" and xs :: "'a list"
let ?ws = "take (Suc (Suc m)) xs"
let ?ys = "drop (Suc (Suc m)) xs"
let ?r = "\<lambda>m'. round_suc_suc index key ?ws m m' u"
let ?t = "\<lambda>r' v. round index key p q r' (v, ns, ?ys)"
assume
A: "index_less index key" and
B: "index_same index key"
assume
"\<And>ws a b c d e g h i n f.
ws = ?ws \<Longrightarrow> a = bn_comp m p q r \<Longrightarrow> (b, c) = bn_comp m p q r \<Longrightarrow>
d = ?r b \<Longrightarrow> (e, g) = ?r b \<Longrightarrow> (h, i) = g \<Longrightarrow>
bn_inv p q (e, ns, ?ys) \<longrightarrow> add_inv n (e, ns, ?ys) \<longrightarrow>
stab_inv f key (e, ns, ?ys) \<longrightarrow> stab_inv f key (?t c e)" and
"bn_inv p q (u, Suc (Suc m) # ns, xs)" and
"add_inv n (u, Suc (Suc m) # ns, xs)" and
"stab_inv f key (u, Suc (Suc m) # ns, xs)"
thus "stab_inv f key (round index key p q r (u, Suc (Suc m) # ns, xs))"
+ using [[simproc del: defined_all]]
proof (simp split: prod.split, ((rule_tac allI)+, ((rule_tac impI)+)?)+,
(erule_tac conjE)+, subst (asm) (2) add_base_zero, simp)
fix m' r' v ms' ws' xs' k
let ?f = "\<lambda>k. [x\<leftarrow>?ys. key x = k]"
let ?P = "\<lambda>f. \<forall>k. [x\<leftarrow>?ys. key x = k] = f k"
let ?Q = "\<lambda>f. \<forall>k. [x\<leftarrow>xs'. key x = k] = f k"
assume
C: "?r m' = (v, ms', ws')" and
D: "bn_comp m p q r = (m', r')" and
E: "bn_valid m p q" and
F: "Suc (Suc (foldl (+) 0 ns + m)) = n" and
G: "length xs = n"
assume "\<And>ws a b c d e g h i n' f.
ws = ?ws \<Longrightarrow> a = (m', r') \<Longrightarrow> b = m' \<and> c = r' \<Longrightarrow>
d = (v, ms', ws') \<Longrightarrow> e = v \<and> g = (ms', ws') \<Longrightarrow> h = ms' \<and> i = ws' \<Longrightarrow>
foldl (+) 0 ns = n' \<and> n - Suc (Suc m) = n' \<longrightarrow> ?P f \<longrightarrow> ?Q f"
hence "\<And>f. foldl (+) 0 ns = n - Suc (Suc m) \<longrightarrow> ?P f \<longrightarrow> ?Q f"
by simp
hence "\<And>f. ?P f \<longrightarrow> ?Q f"
using F by simp
hence "?P ?f \<longrightarrow> ?Q ?f" .
hence "[x\<leftarrow>xs'. key x = k] = [x\<leftarrow>?ys. key x = k]" by simp
moreover assume "\<forall>k. [x\<leftarrow>xs. key x = k] = f k"
hence "f k = [x\<leftarrow>?ws @ ?ys. key x = k]" by simp
ultimately have "f k = [x\<leftarrow>?ws. key x = k] @ [x\<leftarrow>xs'. key x = k]"
by (subst (asm) filter_append, simp)
with C [symmetric] show "[x\<leftarrow>ws'. key x = k] @ [x\<leftarrow>xs'. key x = k] = f k"
proof (simp add: round_suc_suc_def Let_def del: filter.simps
split: if_split_asm)
let ?nmi = "mini ?ws key"
let ?nma = "maxi ?ws key"
let ?xmi = "?ws ! ?nmi"
let ?xma = "?ws ! ?nma"
let ?mi = "key ?xmi"
let ?ma = "key ?xma"
let ?k = "case m of 0 \<Rightarrow> m | Suc 0 \<Rightarrow> m | Suc (Suc i) \<Rightarrow> u + m'"
let ?zs = "nths ?ws (- {?nmi, ?nma})"
let ?ms = "enum ?zs index key ?k ?mi ?ma"
have H: "length ?ws = Suc (Suc m)"
using F and G by simp
hence I: "?nmi \<noteq> ?nma"
by (rule_tac mini_maxi_neq, simp)
have "[x\<leftarrow>(([?xmi] @ map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma))
@ [?xma]). key x = k] = [x\<leftarrow>?ws. key x = k]"
proof (cases "m = 0")
case True
have "?nmi < length ?ws"
using H by (rule_tac mini_less, simp)
hence J: "?nmi < Suc (Suc 0)"
using True by simp
moreover have "?nma < length ?ws"
using H by (rule_tac maxi_less, simp)
hence K: "?nma < Suc (Suc 0)"
using True by simp
ultimately have "card ({..<Suc (Suc 0)} - {?nma} - {?nmi}) = 0"
using I by ((subst card_Diff_singleton, simp, simp)+,
subst card_lessThan, simp)
hence L: "{..<Suc (Suc 0)} - {?nma} - {?nmi} = {}" by simp
have "length (fill ?zs (offs ?ms 0) index key m ?mi ?ma) = 0"
using True by (simp add: fill_length)
hence M: "map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma) =
nths ?ws ({..<Suc (Suc 0)} - {?nma} - {?nmi})"
by (subst L, simp)
show ?thesis
proof (subst M, subst filter_append)
show "[x\<leftarrow>[?xmi] @ nths ?ws ({..<Suc (Suc 0)} - {?nma} - {?nmi}).
key x = k] @ [x\<leftarrow>[?xma]. key x = k] = [x\<leftarrow>?ws. key x = k]"
proof (subst mini_stable, simp only: length_greater_0_conv
[symmetric] H, simp add: I J, subst filter_append [symmetric])
show "[x\<leftarrow>nths ?ws ({..<Suc (Suc 0)} - {?nma}) @ [?xma].
key x = k] = [x\<leftarrow>?ws. key x = k]"
by (subst maxi_stable, simp only: length_greater_0_conv
[symmetric] H, simp add: K, simp add: True)
qed
qed
next
case False
hence "0 < ?k"
by (simp, drule_tac bn_comp_fst_nonzero [OF E], subst (asm) D,
simp split: nat.split)
hence "[x\<leftarrow>map the (fill ?zs (offs ?ms 0) index key (length ?zs) ?mi ?ma).
key x = k] = [x\<leftarrow>?zs. k = key x]"
by (rule_tac fill_offs_enum_stable [OF A B], simp, rule_tac conjI,
((rule_tac mini_lb | rule_tac maxi_ub), erule_tac in_set_nthsD)+)
moreover have "[x\<leftarrow>?zs. k = key x] = [x\<leftarrow>?zs. key x = k]"
by (rule filter_cong, simp, blast)
ultimately have
J: "[x\<leftarrow>map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma).
key x = k] = [x\<leftarrow>?zs. key x = k]"
using H by (simp add: mini_maxi_nths)
show ?thesis
proof (simp only: filter_append J, subst Compl_eq_Diff_UNIV,
subst Diff_insert, subst filter_append [symmetric])
show "[x\<leftarrow>[?xmi] @ nths ?ws (UNIV - {?nma} - {?nmi}). key x = k]
@ [x\<leftarrow>[?xma]. key x = k] = [x\<leftarrow>?ws. key x = k]"
proof (subst mini_stable, simp only: length_greater_0_conv
[symmetric] H, simp add: I, subst filter_append [symmetric])
show "[x\<leftarrow>nths ?ws (UNIV - {?nma}) @ [?xma]. key x = k] =
[x\<leftarrow>?ws. key x = k]"
by (subst maxi_stable, simp only: length_greater_0_conv
[symmetric] H, simp, subst nths_range, subst H, simp)
qed
qed
qed
thus "[x\<leftarrow>?xmi # map the (fill ?zs (offs ?ms 0) index key m ?mi ?ma) @
[?xma]. key x = k] = [x\<leftarrow>?ws. key x = k]"
by simp
qed
qed
qed
lemma gcsort_stab_inv:
assumes
A: "index_less index key" and
B: "index_same index key" and
C: "add_inv n t" and
D: "n \<le> p"
shows "\<lbrakk>t' \<in> gcsort_set index key p t; stab_inv f key t\<rbrakk> \<Longrightarrow>
stab_inv f key t'"
by (erule gcsort_set.induct, simp, drule gcsort_add_inv [OF A _ C D],
rule round_stab_inv [OF A B], simp_all del: bn_inv.simps, erule conjE,
frule sym, erule subst, rule bn_inv_intro, insert D, simp)
text \<open>
\null
The only remaining task is to address step 10 of the proof method, which is done by proving theorem
@{text gcsort_stable}. It holds under the conditions that the objects' number is not larger than the
counters' upper bound and function @{text index} satisfies both predicates @{const index_less} and
@{const index_same}, and states that function @{const gcsort} leaves unchanged the sublist of the
objects having a given key within the input objects' list.
\null
\<close>
theorem gcsort_stable:
assumes
A: "index_less index key" and
B: "index_same index key" and
C: "length xs \<le> p"
shows "[x\<leftarrow>gcsort index key p xs. key x = k] = [x\<leftarrow>xs. key x = k]"
proof -
let ?t = "(0, [length xs], xs)"
have "stab_inv (\<lambda>k. [x\<leftarrow>xs. key x = k]) key (gcsort_aux index key p ?t)"
by (rule gcsort_stab_inv [OF A B _ C], rule gcsort_add_input,
rule gcsort_aux_set, rule gcsort_stab_input)
hence "[x\<leftarrow>gcsort_out (gcsort_aux index key p ?t). key x = k] =
[x\<leftarrow>xs. key x = k]"
by (rule gcsort_stab_intro)
moreover have "?t = gcsort_in xs"
by (simp add: gcsort_in_def)
ultimately show ?thesis
by (simp add: gcsort_def)
qed
end
diff --git a/thys/Graph_Saturation/RuleSemanticsConnection.thy b/thys/Graph_Saturation/RuleSemanticsConnection.thy
--- a/thys/Graph_Saturation/RuleSemanticsConnection.thy
+++ b/thys/Graph_Saturation/RuleSemanticsConnection.thy
@@ -1,544 +1,544 @@
section \<open>Translating terms into Graphs\<close>
text \<open>We define the translation function and its properties.\<close>
theory RuleSemanticsConnection
imports LabeledGraphSemantics RulesAndChains
begin
text \<open>Definition 15.\<close>
fun translation :: "'c allegorical_term \<Rightarrow> ('c, nat) labeled_graph" where
"translation (A_Lbl l) = LG {(l,0,1)} {0,1}" |
"translation (A_Cnv e) = map_graph_fn (translation e) (\<lambda> x. if x<2 then (1-x) else x)" |
"translation (A_Cmp e\<^sub>1 e\<^sub>2)
= (let G\<^sub>1 = translation e\<^sub>1 ; G\<^sub>2 = translation e\<^sub>2
in graph_union (map_graph_fn G\<^sub>1 (\<lambda> x. if x=0 then 0 else x+card(vertices G\<^sub>2)-1))
(map_graph_fn G\<^sub>2 (\<lambda> x. if x=0 then card (vertices G\<^sub>2) else x)))" |
"translation (A_Int e\<^sub>1 e\<^sub>2)
= (let G\<^sub>1 = translation e\<^sub>1 ; G\<^sub>2 = translation e\<^sub>2
in graph_union G\<^sub>1 (map_graph_fn G\<^sub>2 (\<lambda> x. if x<2 then x else x+card(vertices G\<^sub>1)-2)))"
definition inv_translation where
"inv_translation r \<equiv> {0..<card r} = r \<and> {0,1}\<subseteq>r"
lemma inv_translationI4[intro]:
assumes "finite r" "\<And> x. x < card r \<Longrightarrow> x \<in> r"
shows "r={0..<card r}"
proof(insert assms,induct "card r" arbitrary:r)
case (Suc x r)
let ?r = "r - {x}"
from Suc have p:"x = card ?r" "finite ?r" by auto
have p2:"xa < card ?r \<Longrightarrow> xa \<in> ?r" for xa
using Suc.prems(2)[of xa] Suc.hyps(2) unfolding p(1)[symmetric] by auto
from Suc.hyps(1)[OF p p2] have "?r={0..<card ?r}".
with Suc.hyps(2) Suc.prems(1) show ?case
by (metis atLeast0_lessThan_Suc card_Diff_singleton_if insert_Diff n_not_Suc_n p(1))
qed auto
lemma inv_translationI[intro!]:
assumes "finite r" "\<And> x. x < card r \<Longrightarrow> x \<in> r" "0 \<in> r" "Suc 0 \<in> r"
shows "inv_translation r"
proof -
from inv_translationI4[OF assms(1,2),symmetric]
have c:" {0..<card r} = r " by auto
from assms(3,4) have "{0,1} \<subseteq> r" by auto
with c inv_translation_def show ?thesis by auto
qed
lemma verts_in_translation_finite[intro]:
"finite (vertices (translation X))"
"finite (edges (translation X))"
"0 \<in> vertices (translation X)"
"Suc 0 \<in> vertices (translation X)"
proof(atomize(full),induction X)
case (A_Int X1 X2)
then show ?case by (auto simp:Let_def)
next
case (A_Cmp X1 X2)
then show ?case by (auto simp:Let_def)
next
have [simp]:"{x::nat. x < 2} = {0,1}" by auto
case (A_Cnv X)
then show ?case by auto
qed auto
lemma inv_tr_card_min:
assumes "inv_translation r"
shows "card r \<ge> 2"
proof -
note [simp] = inv_translation_def
have "{0..<x} = r \<Longrightarrow> 2 \<le> x \<longleftrightarrow> 0 \<in> r \<and> 1 \<in> r" for x by auto
thus ge2:"card r\<ge>2" using assms by auto
qed
lemma verts_in_translation[intro]:
"inv_translation (vertices (translation X))"
proof(induct X)
{ fix r
assume assms:"inv_translation r"
note [simp] = inv_translation_def
from assms have a1:"finite r"
by (intro card_ge_0_finite) auto
have [simp]:"{0..<Suc x} = {0..<x} \<union> {x}" for x by auto
note ge2 = inv_tr_card_min[OF assms]
from ge2 assms have r0:"r \<inter> {0} = {0}" "r \<inter> {x. x < 2} = {0,1}" by auto
have [intro!]:"\<And>x. x \<in> r \<Longrightarrow> x < card r"
and g6:"\<And>x. x < card r \<longleftrightarrow> x \<in> r"
using assms[unfolded inv_translation_def] atLeastLessThan_iff by blast+
have g4:"r \<inter> {x. \<not> x < 2} = {2..<card r}"
"r \<inter> (Collect ((<) 0)) = {1..<card r}" using assms by fastforce+
have ins:"1 \<in> r" "0 \<in> r" using assms by auto
have d:"Suc (Suc (card r - 2)) = card r"
using ge2 One_nat_def Suc_diff_Suc Suc_pred
numeral_2_eq_2 by presburger
note ge2 ins g4 g6 r0 d
} note inv_translationD[simp] = this
{
fix a b c
assume assm:"b \<le> (a::nat)"
have "(\<lambda>x. x + a - b) ` {b..<c} = {a..<c+a-b}" (is "?lhs = ?rhs")
proof -
from assm have "?lhs = (\<lambda>x. x + (a - b)) ` {b..<c}" by auto
also have "\<dots> = ?rhs"
unfolding linordered_semidom_class.image_add_atLeastLessThan' using assm by auto
finally show ?thesis by auto
qed
} note e[simp] = this
{ fix r z
assume a1: "inv_translation z" and a2: "inv_translation r"
let ?z2 = "card z + card r - 2"
let ?z1 = "card z + card r - Suc 0"
from a1 a2
have le1:"Suc 0 \<le> card r"
by (metis Suc_leD inv_translationD(1) numerals(2))
hence le2: "card r \<le> ?z1"
by (metis Suc_leD a1 inv_translationD(1) numerals(2) ordered_cancel_comm_monoid_diff_class.le_add_diff)
with le1 have b:"{card r ..< ?z1} \<union> {Suc 0 ..< card r} = {Suc 0 ..< ?z1}"
by auto
have a:"(insert (card r) {0..<card z + card r - Suc 0}) = {0..<card z + card r - Suc 0}"
using le1 le2 a1 a2
by (metis Suc_leD add_Suc_right atLeastLessThan_iff diff_Suc_Suc insert_absorb inv_translationD(1) linorder_not_less not_less_eq_eq numerals(2) ordered_cancel_comm_monoid_diff_class.le_add_diff)
from a1 a2
have "card z + card r - 2 \<ge> card (r::nat set)"
by (simp add: ordered_cancel_comm_monoid_diff_class.le_add_diff)
with a2
have c:"card (r \<union> {card r..<?z2}) = ?z2"
by (metis atLeast0LessThan card_atLeastLessThan diff_zero inv_translation_def ivl_disj_un_one(2))+
note a b c
} note [simp] = this
have [simp]:"a < x \<Longrightarrow> insert a {Suc a..<x} = {a..<x}" for a x by auto
{ case (A_Int X1 X2)
let ?v1 = "vertices (translation X1)"
from A_Int have [simp]:"(insert 0 (insert (Suc 0) (?v1 \<union> x))) = ?v1 \<union> x"
for x unfolding inv_translation_def by auto
from A_Int show ?case by (auto simp:Let_def linorder_not_le)
next
case (A_Cmp X1 X2)
hence "2\<le>card (vertices (translation X1))" "2\<le>card (vertices (translation X2))" by auto
hence "1 \<le>card (vertices (translation X1))" "1\<le>card (vertices (translation X2))"
"1 < card (vertices (translation X1)) + card (vertices (translation X2)) - 1"
by auto
from this A_Cmp
show ?case by (auto simp:Let_def)
next
case (A_Cnv X)
thus ?case by (auto simp:Let_def)
}
qed auto
lemma translation_graph[intro]:
"graph (translation X)"
by (induct X, auto simp:Let_def)
lemma graph_rule_translation[intro]: (* remark at the end of Def 15 *)
"graph_rule (translation X, translation (A_Int X Y))"
using verts_in_translation_finite[of X] verts_in_translation_finite[of "A_Int X Y"]
translation_graph[of X] translation_graph[of "A_Int X Y"]
by (auto simp:Let_def subgraph_def2)
lemma graph_hom_translation[intro]:
"graph_homomorphism (LG {} {0,1}) (translation X) (Id_on {0,1})"
using verts_in_translation[of X]
unfolding inv_translation_def graph_homomorphism_def2 by auto
lemma translation_right_to_left:
assumes f:"graph_homomorphism (translation e) G f" "(0, x) \<in> f" "(1, y) \<in> f"
shows "(x, y) \<in> :G:\<lbrakk>e\<rbrakk>"
using f
proof(induct e arbitrary:f x y)
case (A_Int e\<^sub>1 e\<^sub>2 f x y)
let ?f\<^sub>1 = "id"
let ?f\<^sub>2 = "(\<lambda> x. if x < 2 then x else x + card (vertices (translation e\<^sub>1)) - 2)"
let ?G\<^sub>1 = "translation e\<^sub>1"
let ?G\<^sub>2 = "translation e\<^sub>2"
have f1:"(0, x) \<in> on_graph ?G\<^sub>1 ?f\<^sub>1 O f" "(1, y) \<in> on_graph ?G\<^sub>1 ?f\<^sub>1 O f"
and f2:"(0, x) \<in> on_graph ?G\<^sub>2 ?f\<^sub>2 O f" "(1, y) \<in> on_graph ?G\<^sub>2 ?f\<^sub>2 O f"
using A_Int.prems(2,3) by (auto simp:BNF_Def.Gr_def relcomp_def)
from A_Int.prems(1)
have uni:"graph_homomorphism (graph_union ?G\<^sub>1 (map_graph_fn ?G\<^sub>2 ?f\<^sub>2)) G f"
by (auto simp:Let_def)
from graph_homo_union_id(1)[OF uni translation_graph]
have h1:"graph_homomorphism ?G\<^sub>1 (translation (A_Int e\<^sub>1 e\<^sub>2)) (on_graph ?G\<^sub>1 id)"
by (auto simp:Let_def graph_homomorphism_def)
have "graph (map_graph_fn ?G\<^sub>2 ?f\<^sub>2)" by auto
from graph_homo_union_id(2)[OF uni this]
have h2:"graph_homomorphism ?G\<^sub>2 (translation (A_Int e\<^sub>1 e\<^sub>2)) (on_graph ?G\<^sub>2 ?f\<^sub>2)"
by (auto simp:Let_def graph_homomorphism_def)
from A_Int.hyps(1)[OF graph_homomorphism_composes[OF h1 A_Int.prems(1)] f1]
A_Int.hyps(2)[OF graph_homomorphism_composes[OF h2 A_Int.prems(1)] f2]
show ?case by auto
next
case (A_Cmp e\<^sub>1 e\<^sub>2 f x y)
let ?f\<^sub>1 = "(\<lambda> x. if x=0 then 0 else x+card(vertices (translation e\<^sub>2))-1)"
let ?f\<^sub>2 = "(\<lambda> x. if x=0 then card (vertices (translation e\<^sub>2)) else x)"
let ?G\<^sub>1 = "translation e\<^sub>1"
let ?G\<^sub>2 = "translation e\<^sub>2"
let ?v = "card (vertices (translation e\<^sub>2))"
from A_Cmp.prems(1) have "?v \<in> Domain f" by (auto simp:Let_def graph_homomorphism_def)
then obtain v where v:"(?v,v) \<in> f" by auto
have f1:"(0, x) \<in> on_graph ?G\<^sub>1 ?f\<^sub>1 O f" "(1, v) \<in> on_graph ?G\<^sub>1 ?f\<^sub>1 O f"
and f2:"(0, v) \<in> on_graph ?G\<^sub>2 ?f\<^sub>2 O f" "(1, y) \<in> on_graph ?G\<^sub>2 ?f\<^sub>2 O f"
using A_Cmp.prems(2,3) v by auto
from A_Cmp.prems(1)
have uni:"graph_homomorphism (graph_union (map_graph_fn ?G\<^sub>1 ?f\<^sub>1) (map_graph_fn ?G\<^sub>2 ?f\<^sub>2)) G f"
by (auto simp:Let_def)
have "graph (map_graph_fn ?G\<^sub>1 ?f\<^sub>1)" by auto
from graph_homo_union_id(1)[OF uni this]
have h1:"graph_homomorphism ?G\<^sub>1 (translation (A_Cmp e\<^sub>1 e\<^sub>2)) (on_graph ?G\<^sub>1 ?f\<^sub>1)"
by (auto simp:Let_def graph_homomorphism_def2)
have "graph (map_graph_fn ?G\<^sub>2 ?f\<^sub>2)" by auto
from graph_homo_union_id(2)[OF uni this]
have h2:"graph_homomorphism ?G\<^sub>2 (translation (A_Cmp e\<^sub>1 e\<^sub>2)) (on_graph ?G\<^sub>2 ?f\<^sub>2)"
by (auto simp:Let_def graph_homomorphism_def2)
from A_Cmp.hyps(1)[OF graph_homomorphism_composes[OF h1 A_Cmp.prems(1)] f1]
A_Cmp.hyps(2)[OF graph_homomorphism_composes[OF h2 A_Cmp.prems(1)] f2]
show ?case by auto
next
case (A_Cnv e f x y)
let ?f = "(\<lambda> x. if x < 2 then 1 - x else x)"
let ?G = "translation e"
have i:"graph_homomorphism ?G (map_graph_fn ?G ?f) (on_graph ?G ?f)" using A_Cnv by auto
have "(0, y) \<in> on_graph ?G ?f O f" "(1, x) \<in> on_graph ?G ?f O f"
using A_Cnv.prems(3,2) by (auto simp:BNF_Def.Gr_def relcomp_def)
from A_Cnv.hyps(1)[OF graph_homomorphism_composes[OF i] this] A_Cnv.prems(1)
show ?case by auto
next
case (A_Lbl l f x y)
hence "edge_preserving f {(l,0,1)} (edges G)" unfolding graph_homomorphism_def by auto
with A_Lbl(2,3) show ?case by (auto simp:getRel_def edge_preserving_def)
qed
lemma translation_homomorphism:
assumes "graph_homomorphism (translation e) G f"
shows "f `` {0} \<times> f `` {1} \<subseteq> :G:\<lbrakk>e\<rbrakk>" ":G:\<lbrakk>e\<rbrakk> \<noteq> {}"
using translation_right_to_left[OF assms] assms[unfolded graph_homomorphism_def2]
verts_in_translation_finite[of e] by auto
text \<open>Lemma 5.\<close>
lemma translation:
assumes "graph G"
shows "(x, y) \<in> :G:\<lbrakk>e\<rbrakk> \<longleftrightarrow> (\<exists> f. graph_homomorphism (translation e) G f \<and> (0,x) \<in> f \<and> (1,y) \<in> f)"
(is "?lhs = ?rhs")
proof
have [dest]:"y + card (vertices (translation (e::'a allegorical_term))) - 2 < 2 \<Longrightarrow> (y::nat) < 2"
for y e using inv_tr_card_min[OF verts_in_translation,of e] by linarith
{ fix y fix e::"'a allegorical_term"
assume "y + card (vertices (translation e)) - 2 \<in> vertices (translation e)"
hence "y + card (vertices (translation e)) - 2 < card (vertices (translation e))"
using verts_in_translation[of e,unfolded inv_translation_def] by auto
hence "y < 2" using inv_tr_card_min[OF verts_in_translation,of e] by auto
} note [dest!] = this
{ fix y fix e::"'a allegorical_term"
assume "y + card (vertices (translation e)) - Suc 0 \<in> vertices (translation e)"
hence "y + card (vertices (translation e)) - Suc 0 \<in> {0..<card (vertices (translation e))}"
using verts_in_translation[of e,unfolded inv_translation_def] by simp
hence "y = 0" using inv_tr_card_min[OF verts_in_translation,of e] by auto
} note [dest!] = this
{ fix y fix e::"'a allegorical_term"
assume "card (vertices (translation e)) \<in> vertices (translation e)"
hence "card (vertices (translation e)) \<in> {0..<card (vertices (translation e))}"
using verts_in_translation[of e,unfolded inv_translation_def] by auto
hence "False" by auto
} note [dest!] = this
{ fix y fix e::"'a allegorical_term"
assume "y + card (vertices (translation e)) \<le> Suc 0"
hence " card (vertices (translation e)) \<le> Suc 0" by auto
hence "False" using inv_tr_card_min[OF verts_in_translation[of e]] by auto
} note [dest!] = this
assume ?lhs
then show ?rhs
proof(induct e arbitrary:x y)
case (A_Int e\<^sub>1 e\<^sub>2)
from A_Int have assm:"(x, y) \<in> :G:\<lbrakk>e\<^sub>1\<rbrakk>" "(x, y) \<in> :G:\<lbrakk>e\<^sub>2\<rbrakk>" by auto
from A_Int(1)[OF assm(1)] obtain f\<^sub>1 where
f\<^sub>1:"graph_homomorphism (translation e\<^sub>1) G f\<^sub>1" "(0, x) \<in> f\<^sub>1" "(1, y) \<in> f\<^sub>1" by auto
from A_Int(2)[OF assm(2)] obtain f\<^sub>2 where
f\<^sub>2:"graph_homomorphism (translation e\<^sub>2) G f\<^sub>2" "(0, x) \<in> f\<^sub>2" "(1, y) \<in> f\<^sub>2" by auto
from f\<^sub>1 f\<^sub>2 have v:"Domain f\<^sub>1 = vertices (translation e\<^sub>1)" "Domain f\<^sub>2 = vertices (translation e\<^sub>2)"
unfolding graph_homomorphism_def by auto
let ?f\<^sub>2 = "(\<lambda> x. if x < 2 then x else x + card (vertices (translation e\<^sub>1)) - 2)"
let ?tr\<^sub>2 = "on_graph (translation e\<^sub>2) ?f\<^sub>2"
have inj2:"inj_on ?f\<^sub>2 (vertices (translation e\<^sub>2))" unfolding inj_on_def by auto
have "(0,0) \<in> ?tr\<^sub>2\<inverse>" "(1,1) \<in> ?tr\<^sub>2\<inverse>" by auto
from this[THEN relcompI] f\<^sub>2(2,3)
have zero_one:"(0,x) \<in> ?tr\<^sub>2\<inverse> O f\<^sub>2"
"(1,y) \<in> ?tr\<^sub>2\<inverse> O f\<^sub>2" by auto
{ fix yb zb
assume "(yb + card (vertices (translation e\<^sub>1)) - 2, zb) \<in> f\<^sub>1"
hence "yb + card (vertices (translation e\<^sub>1)) - 2 \<in> vertices (translation e\<^sub>1)" using v by auto
} note in_f[dest!] = this
have d_a:"Domain f\<^sub>1 \<inter> Domain (?tr\<^sub>2\<inverse> O f\<^sub>2) = {0,1}"
using zero_one by (auto simp:v)
have d_b:"Domain (f\<^sub>1 \<inter> ?tr\<^sub>2\<inverse> O f\<^sub>2) = {0,1}"
using zero_one f\<^sub>1(2,3) by auto
note cmp2 = graph_homomorphism_composes[OF graph_homo_inv[OF translation_graph inj2] f\<^sub>2(1)]
have "graph_homomorphism (translation (A_Int e\<^sub>1 e\<^sub>2)) G (f\<^sub>1 \<union> ?tr\<^sub>2\<inverse> O f\<^sub>2)"
using graph_homo_union[OF f\<^sub>1(1) cmp2 d_a[folded d_b]]
by (auto simp:Let_def)
thus ?case using zero_one[THEN UnI2[of _ _ "f\<^sub>1"]] by blast
next
case (A_Cmp e\<^sub>1 e\<^sub>2)
from A_Cmp obtain z where assm:"(x, z) \<in> :G:\<lbrakk>e\<^sub>1\<rbrakk>" "(z, y) \<in> :G:\<lbrakk>e\<^sub>2\<rbrakk>" by auto
from A_Cmp(1)[OF assm(1)] obtain f\<^sub>1 where
f\<^sub>1:"graph_homomorphism (translation e\<^sub>1) G f\<^sub>1" "(0, x) \<in> f\<^sub>1" "(1, z) \<in> f\<^sub>1" by auto
from A_Cmp(2)[OF assm(2)] obtain f\<^sub>2 where
f\<^sub>2:"graph_homomorphism (translation e\<^sub>2) G f\<^sub>2" "(0, z) \<in> f\<^sub>2" "(1, y) \<in> f\<^sub>2" by auto
from f\<^sub>1 f\<^sub>2 have v:"Domain f\<^sub>1 = vertices (translation e\<^sub>1)" "Domain f\<^sub>2 = vertices (translation e\<^sub>2)"
unfolding graph_homomorphism_def by auto
let ?f\<^sub>1 = "(\<lambda> x. if x=0 then 0 else x+card(vertices (translation e\<^sub>2))-1)"
let ?f\<^sub>2 = "(\<lambda> x. if x=0 then card (vertices (translation e\<^sub>2)) else x)"
let ?tr\<^sub>1 = "on_graph (translation e\<^sub>1) ?f\<^sub>1"
let ?tr\<^sub>2 = "on_graph (translation e\<^sub>2) ?f\<^sub>2"
have inj1:"inj_on ?f\<^sub>1 (vertices (translation e\<^sub>1))" unfolding inj_on_def by auto
have inj2:"inj_on ?f\<^sub>2 (vertices (translation e\<^sub>2))" unfolding inj_on_def by auto
have "(card (vertices (translation e\<^sub>2)),0) \<in> ?tr\<^sub>2\<inverse>" "(1,1) \<in> ?tr\<^sub>2\<inverse>"
"(0,0) \<in> ?tr\<^sub>1\<inverse>" "(card (vertices (translation e\<^sub>2)),1) \<in> ?tr\<^sub>1\<inverse>" by auto
from this[THEN relcompI] f\<^sub>2(2,3) f\<^sub>1(2,3)
have zero_one:"(card (vertices (translation e\<^sub>2)),z) \<in> ?tr\<^sub>1\<inverse> O f\<^sub>1"
"(0,x) \<in> ?tr\<^sub>1\<inverse> O f\<^sub>1"
"(card (vertices (translation e\<^sub>2)),z) \<in> ?tr\<^sub>2\<inverse> O f\<^sub>2"
"(1,y) \<in> ?tr\<^sub>2\<inverse> O f\<^sub>2" by auto
have [simp]:
"ye \<in> vertices (translation e\<^sub>2) \<Longrightarrow>
(if ye = 0 then card (vertices (translation e\<^sub>2)) else ye) =
(if yd = 0 then 0 else yd + card (vertices (translation e\<^sub>2)) - 1) \<longleftrightarrow> ye = 0 \<and> yd = 1"
for ye yd using v inv_tr_card_min[OF verts_in_translation,of "e\<^sub>2"]
by(cases "ye=0";cases "yd=0";auto)
have d_a:"Domain (?tr\<^sub>1\<inverse> O f\<^sub>1) \<inter> Domain (?tr\<^sub>2\<inverse> O f\<^sub>2) = {card (vertices (translation e\<^sub>2))}"
- using zero_one by (auto simp:v)
+ using zero_one using [[simproc del: defined_all]] by (auto simp: v)
have d_b:"Domain (?tr\<^sub>1\<inverse> O f\<^sub>1 \<inter> ?tr\<^sub>2\<inverse> O f\<^sub>2) = {card (vertices (translation e\<^sub>2))}"
- using zero_one f\<^sub>1(2,3) by auto
+ using zero_one f\<^sub>1(2,3) using [[simproc del: defined_all]] by auto
note cmp1 = graph_homomorphism_composes[OF graph_homo_inv[OF translation_graph inj1] f\<^sub>1(1)]
note cmp2 = graph_homomorphism_composes[OF graph_homo_inv[OF translation_graph inj2] f\<^sub>2(1)]
have "graph_homomorphism (translation (A_Cmp e\<^sub>1 e\<^sub>2)) G (?tr\<^sub>1\<inverse> O f\<^sub>1 \<union> ?tr\<^sub>2\<inverse> O f\<^sub>2)"
unfolding Let_def translation.simps
by (rule graph_homo_union[OF cmp1 cmp2 d_a[folded d_b]])
thus ?case using zero_one by blast
next
case (A_Cnv e)
let ?G = "translation (A_Cnv e)"
from A_Cnv obtain f where
f:"graph_homomorphism (translation e) G f" "(0, y) \<in> f" "(1, x) \<in> f" by auto
hence v:"Domain f = vertices (translation e)"
unfolding graph_homomorphism_def by auto
define n where "n \<equiv> card (vertices (translation e))"
from verts_in_translation f inv_tr_card_min[OF verts_in_translation] v(1)
have n:"vertices (translation e) = {0..<n}" "{0..<n} \<inter> {x. x < 2} = {1,0}"
"Domain f = {0..<n}" "{0..<n} \<inter> {x. \<not> x < 2} = {2..<n}"
and n2: "n \<ge> 2"
by (auto simp:n_def inv_translation_def)
then have [simp]:"insert (Suc 0) {2..<n} = {1..<n}"
- "insert 0 {Suc 0..<n} = {0..<n}" by auto
+ "insert 0 {Suc 0..<n} = {0..<n}" using [[simproc del: defined_all]] by auto
let ?f = "on_graph ?G (\<lambda> x. if x < 2 then 1 - x else x)"
have h:"graph_homomorphism ?G G (?f O f)"
proof(rule graph_homomorphism_composes[OF _ f(1)],rule graph_homomorphismI)
show "vertices ?G = Domain ?f"
by (auto simp:Domain_int_univ)
show "?f `` vertices ?G \<subseteq> vertices (translation e)" using n2 by auto
show "univalent ?f" by auto
show "edge_preserving ?f (edges (translation (A_Cnv e))) (edges (translation e))"
by (rule edge_preserving_on_graphI,auto simp: BNF_Def.Gr_def)
qed (auto intro:assms)
have xy:"(0, x) \<in> ?f O f" "(1, y) \<in> ?f O f" using n2 f(2,3) n(1,2) by auto
with h show ?case by auto
next
case (A_Lbl l)
let ?f = "{(0,x),(1,y)}"
have xy:"x \<in> vertices G" "y \<in> vertices G" using assms A_Lbl by (auto simp:getRel_def)
have "graph_homomorphism (translation (A_Lbl l)) G ?f \<and> (0, x) \<in> ?f \<and> (1, y) \<in> ?f"
using assms A_Lbl xy unfolding graph_homomorphism_def2
by (auto simp:univalent_def getRel_def on_triple_def Image_def graph_union_def insert_absorb)
then show ?case by auto
qed
qed (insert translation_right_to_left,auto)
abbreviation transl_rule ::
"'a sentence \<Rightarrow> ('a, nat) labeled_graph \<times> ('a, nat) labeled_graph" where
"transl_rule R \<equiv> (translation (fst R),translation (snd R))"
text \<open>Lemma 6.\<close>
lemma maintained_holds_iff:
assumes "graph G"
shows "maintained (translation e\<^sub>L,translation (A_Int e\<^sub>L e\<^sub>R)) G \<longleftrightarrow> G \<Turnstile> e\<^sub>L \<sqsubseteq> e\<^sub>R" (is "?rhs = ?lhs")
proof
assume lhs:?lhs
show ?rhs unfolding maintained_def proof(clarify) fix f
assume f:"graph_homomorphism (fst (translation e\<^sub>L, translation (A_Int e\<^sub>L e\<^sub>R))) G f"
then obtain x y where f2:"(0,x) \<in> f" "(1,y) \<in> f" unfolding graph_homomorphism_def
by (metis DomainE One_nat_def prod.sel(1) verts_in_translation_finite(3,4))
with f have "(x,y) \<in> :G:\<lbrakk>fst (e\<^sub>L \<sqsubseteq> e\<^sub>R)\<rbrakk>" unfolding translation[OF assms] by auto
with lhs have "(x,y) \<in> :G:\<lbrakk>snd (e\<^sub>L \<sqsubseteq> e\<^sub>R)\<rbrakk>" by auto
then obtain g where g: "graph_homomorphism (translation (A_Int e\<^sub>L e\<^sub>R)) G g"
and g2: "(0, x) \<in> g" "(1, y) \<in> g" unfolding translation[OF assms] by auto
have v:"vertices (translation (A_Int e\<^sub>L e\<^sub>R)) = Domain g"
"vertices (translation e\<^sub>L) = Domain f" using f g
unfolding graph_homomorphism_def by auto
from subgraph_subset[of "translation e\<^sub>L" "translation (A_Int e\<^sub>L e\<^sub>R)"]
graph_rule_translation[of e\<^sub>L e\<^sub>R]
have dom_sub: "Domain f \<subseteq> Domain g"
using v unfolding prod.sel by argo
hence dom_le:"card (Domain f) \<le> card (Domain g)"
by (metis card.infinite card_mono inv_tr_card_min not_less rel_simps(51) v(1) verts_in_translation)
have c_f:"card (Domain f) \<ge> 2" using inv_tr_card_min[OF verts_in_translation] v by metis
from f[unfolded graph_homomorphism_def]
have ep_f:"edge_preserving f (edges (translation e\<^sub>L)) (edges G)"
and uni_f:"univalent f" by auto
let ?f = "(\<lambda>x. if x < 2 then x else x + card (vertices (translation e\<^sub>L)) - 2)"
define GR where "GR = map_graph_fn (translation e\<^sub>R) ?f"
from g[unfolded graph_homomorphism_def]
have "edge_preserving g (edges (translation (A_Int e\<^sub>L e\<^sub>R))) (edges G)"
and uni_g:"univalent g" by auto
from edge_preserving_subset[OF subset_refl _ this(1)]
have ep_g:"edge_preserving g (edges GR) (edges G)" by (auto simp:Let_def GR_def)
{ fix a assume a:"a \<in> vertices (translation e\<^sub>R)"
hence "?f a \<in> vertices (translation (A_Int e\<^sub>L e\<^sub>R))" by (auto simp:Let_def)
from this[unfolded v] verts_in_translation[of "A_Int e\<^sub>L e\<^sub>R",unfolded inv_translation_def v]
have "\<not> a < 2 \<Longrightarrow> a + card (Domain f) - 2 < card (Domain g)" by auto
} note[intro!] = this
have [intro!]: " \<not> aa < 2 \<Longrightarrow> card (Domain f) \<le> aa + card (Domain f) - 2" for aa by simp
from v(2) restrictD[OF translation_graph[of e\<^sub>L]]
have df[dest]:"xa \<notin> Domain f \<Longrightarrow> (l,xa,xb) \<in> edges (translation e\<^sub>L) \<Longrightarrow> False"
"xa \<notin> Domain f \<Longrightarrow> (l,xb,xa) \<in> edges (translation e\<^sub>L) \<Longrightarrow> False"
for xa l xb unfolding edge_preserving by auto
{ fix l xa xb ya
assume assm: "(l,xa,xb) \<in> edges GR"
with c_f dom_le
have "xa \<in> {0,1} \<union> {card (Domain f)..<card (Domain g)}"
"xb \<in> {0,1} \<union> {card (Domain f)..<card (Domain g)}"
unfolding GR_def v by auto
hence minb:"xa \<in> {0,1} \<or> xa \<ge> card (Domain f)" "xb \<in> {0,1} \<or> xb \<ge> card (Domain f)"
by auto
{ fix z xa assume minb:"xa \<in> {0,1} \<or> xa \<ge> card (Domain f)" and z:"(xa,z) \<in> f"
from z verts_in_translation[of e\<^sub>L,unfolded inv_translation_def v]
have "xa < card(Domain f)" by auto
with minb verts_in_translation[of "A_Int e\<^sub>L e\<^sub>R",unfolded inv_translation_def v]
have x:"xa \<in> {0,1} \<and> xa \<in> Domain g" by auto
then obtain v where g:"(xa,v) \<in> g" by auto
consider "xa = 0 \<and> z = x" | "xa = 1 \<and> z = y"
using x f2[THEN univalentD[OF uni_f]] z by auto
hence "v = z" using g g2[THEN univalentD[OF uni_g]] by metis
hence "(xa,z) \<in> g" using g by auto
}
note minb[THEN this]
}
with f2 g2[THEN univalentD[OF uni_g]]
have dg:"(l,xa,xb) \<in> edges GR \<Longrightarrow> (xa,ya) \<in> f \<Longrightarrow> (xa,ya) \<in> g"
"(l,xb,xa) \<in> edges GR \<Longrightarrow> (xa,ya) \<in> f \<Longrightarrow> (xa,ya) \<in> g"
for xa l xb ya
unfolding edge_preserving by (auto)
have "vertices (translation e\<^sub>L) \<subseteq> vertices (translation (A_Int e\<^sub>L e\<^sub>R))"
by(rule subgraph_subset,insert graph_rule_translation,auto)
hence subdom:"Domain f \<subseteq> Domain g" unfolding v.
let ?g = "f \<union> (Id_on (UNIV - Domain f) O g)"
have [simp]:"Domain ?g = Domain g" using subdom unfolding Domain_Un_eq by auto
have ih:"graph_homomorphism (translation (A_Int e\<^sub>L e\<^sub>R)) G ?g"
proof(rule graph_homomorphismI)
show "?g `` vertices (translation (A_Int e\<^sub>L e\<^sub>R)) \<subseteq> vertices G"
using g[unfolded graph_homomorphism_def] f[unfolded graph_homomorphism_def]
by (auto simp: v simp del:translation.simps)
show "edge_preserving ?g (edges (translation (A_Int e\<^sub>L e\<^sub>R))) (edges G)"
unfolding Let_def translation.simps graph_union_edges proof
show "edge_preserving ?g (edges (translation e\<^sub>L)) (edges G)"
using edge_preserving_atomic[OF ep_f] unfolding edge_preserving by auto
have "edge_preserving ?g (edges GR) (edges G)"
using edge_preserving_atomic[OF ep_g] dg unfolding edge_preserving by (auto;blast)
thus "edge_preserving ?g (edges (map_graph_fn (translation e\<^sub>R) ?f)) (edges G)"
by (auto simp:GR_def)
qed
qed (insert f[unfolded graph_homomorphism_def] g[unfolded graph_homomorphism_def],auto simp:Let_def)
have ie:"agree_on (translation e\<^sub>L) f ?g" unfolding agree_on_def by (auto simp:v)
from ie ih show "extensible (translation e\<^sub>L, translation (A_Int e\<^sub>L e\<^sub>R)) G f"
unfolding extensible_def prod.sel by auto
qed next
assume rhs:?rhs
{ fix x y assume "(x,y) \<in> :G:\<lbrakk>e\<^sub>L\<rbrakk>"
with translation[OF assms] obtain f
where f:"graph_homomorphism (fst (translation e\<^sub>L, translation (A_Int e\<^sub>L e\<^sub>R))) G f"
"(0, x) \<in> f" "(1, y) \<in> f" by auto
with rhs[unfolded maintained_def,rule_format,OF f(1),unfolded extensible_def]
obtain g where g:"graph_homomorphism (translation (A_Int e\<^sub>L e\<^sub>R)) G g"
"agree_on (translation e\<^sub>L) f g" by auto
hence "(x,y) \<in> :G:\<lbrakk>A_Int e\<^sub>L e\<^sub>R\<rbrakk>" using f unfolding agree_on_def translation[OF assms] by auto
}
thus ?lhs by auto
qed
lemma translation_self[intro]:
"(0, 1) \<in> :translation e:\<lbrakk>e\<rbrakk>"
proof(induct e)
case (A_Int e1 e2)
let ?f = "(\<lambda>x. if x < 2 then x else x + card (vertices (translation e1)) - 2)"
have f: "(?f 0,?f 1) \<in>:map_graph_fn (translation e2) ?f:\<lbrakk>e2\<rbrakk>"
using map_graph_in[OF translation_graph A_Int(2),of ?f] by auto
let ?G = "graph_union (translation e1) (map_graph_fn (translation e2) ?f)"
have "{(0,1)} \<subseteq> :(translation e1):\<lbrakk>e1\<rbrakk>" using A_Int by auto
moreover have "{(0,1)} \<subseteq> :map_graph_fn (translation e2) ?f:\<lbrakk>e2\<rbrakk>" using f by auto
moreover have ":map_graph_fn (translation e2) ?f:\<lbrakk>e2\<rbrakk> \<subseteq> :?G:\<lbrakk>e2\<rbrakk>" ":translation e1:\<lbrakk>e1\<rbrakk> \<subseteq> :?G:\<lbrakk>e1\<rbrakk>"
using graph_union_semantics by blast+
ultimately show ?case by (auto simp:Let_def)
next
case (A_Cmp e1 e2)
let ?f1 = "\<lambda>x. if x = 0 then 0 else x + card (vertices (translation e2)) - 1"
have f1: "(?f1 0,?f1 1) \<in>:map_graph_fn (translation e1) ?f1:\<lbrakk>e1\<rbrakk>"
using map_graph_in[OF translation_graph A_Cmp(1),of ?f1] by auto
let ?f2 = "\<lambda>x. if x = 0 then card (vertices (translation e2)) else x"
have f2: "(?f2 0,?f2 1) \<in>:map_graph_fn (translation e2) ?f2:\<lbrakk>e2\<rbrakk>"
using map_graph_in[OF translation_graph A_Cmp(2),of ?f2] by auto
let ?G = "graph_union (map_graph_fn (translation e1) ?f1) (map_graph_fn (translation e2) ?f2)"
have "{(0,1)} = {(0,card (vertices (translation e2)))} O {(card (vertices (translation e2)),1)}"
by auto
also have "{(0,card (vertices (translation e2)))} \<subseteq> :map_graph_fn (translation e1) ?f1:\<lbrakk>e1\<rbrakk>"
using f1 by auto
also have ":map_graph_fn (translation e1) ?f1:\<lbrakk>e1\<rbrakk> \<subseteq> :?G:\<lbrakk>e1\<rbrakk>"
using graph_union_semantics by auto
also have "{(card (vertices (translation e2)),1)} \<subseteq> :map_graph_fn (translation e2) ?f2:\<lbrakk>e2\<rbrakk>"
using f2 by auto
also have ":map_graph_fn (translation e2) ?f2:\<lbrakk>e2\<rbrakk> \<subseteq> :?G:\<lbrakk>e2\<rbrakk>"
using graph_union_semantics by blast
also have "(:?G:\<lbrakk>e1\<rbrakk>) O (:?G:\<lbrakk>e2\<rbrakk>) = :translation (A_Cmp e1 e2):\<lbrakk>A_Cmp e1 e2\<rbrakk>"
by (auto simp:Let_def)
finally show ?case by auto
next
case (A_Cnv e)
from map_graph_in[OF translation_graph this,of "(\<lambda>x. if x < (2::nat) then 1 - x else x)"]
show ?case using map_graph_in[OF translation_graph] by auto
qed (simp add:getRel_def)
text \<open>Lemma 6 is only used on rules of the form @{term "e\<^sub>L \<sqsubseteq> e\<^sub>R"}.
The requirement of G being a graph can be dropped for one direction.\<close>
lemma maintained_holds[intro]:
assumes ":G:\<lbrakk>e\<^sub>L\<rbrakk> \<subseteq> :G:\<lbrakk>e\<^sub>R\<rbrakk>"
shows "maintained (transl_rule (e\<^sub>L \<sqsubseteq> e\<^sub>R)) G"
proof (cases "graph G")
case True
thus ?thesis using assms sentence_iff maintained_holds_iff prod.sel by metis
next
case False
thus ?thesis by (auto simp:maintained_def graph_homomorphism_def)
qed
lemma maintained_holds_subset_iff[simp]:
assumes "graph G"
shows "maintained (transl_rule (e\<^sub>L \<sqsubseteq> e\<^sub>R)) G \<longleftrightarrow> (:G:\<lbrakk>e\<^sub>L\<rbrakk> \<subseteq> :G:\<lbrakk>e\<^sub>R\<rbrakk>)"
using assms maintained_holds_iff sentence_iff prod.sel by metis
end
\ No newline at end of file
diff --git a/thys/Graph_Saturation/StandardRules.thy b/thys/Graph_Saturation/StandardRules.thy
--- a/thys/Graph_Saturation/StandardRules.thy
+++ b/thys/Graph_Saturation/StandardRules.thy
@@ -1,809 +1,810 @@
section \<open>Standard Rules\<close>
text \<open>We define the standard rules here, and prove the relation to standard rules.
This means proving that the graph rules do what they say they do.\<close>
theory StandardRules
imports StandardModels RuleSemanticsConnection
begin
text \<open>Definition 16 makes this remark. We don't have a specific version of Definition 16.\<close>
lemma conflict_free:
":G:\<lbrakk>A_Lbl l\<rbrakk> = {} \<longleftrightarrow> (\<forall> (l',x,y)\<in>edges G. l' \<noteq> l)"
by (auto simp:getRel_def)
text \<open>Definition 17, abstractly.
It's unlikely that we wish to use the top rule for any symbol except top,
but stating it abstractly makes it consistent with the other rules.\<close>
(* Definition 17 *)
definition top_rule :: "'l \<Rightarrow> ('l,nat) Graph_PreRule" where
"top_rule t = (LG {} {0,1},LG {(t,0,1)} {0,1})"
text \<open>Proof that definition 17 does what it says it does.\<close>
lemma top_rule[simp]:
assumes "graph G"
shows "maintained (top_rule r) G \<longleftrightarrow> vertices G \<times> vertices G = getRel r G"
proof
assume a:"maintained (top_rule r) G"
{ fix a b assume "a \<in> vertices G" "b \<in> vertices G"
hence "graph_homomorphism (LG {} {0, 1}) G {(0::nat,a),(1,b)}"
using assms unfolding graph_homomorphism_def univalent_def by auto
with a[unfolded maintained_def top_rule_def] extensible_refl_concr
have "graph_homomorphism (LG {(r, 0, 1)} {0::nat, 1}) G {(0::nat, a), (1, b)}" by simp
hence "(a, b) \<in> getRel r G"
unfolding graph_homomorphism_def2 graph_union_iff getRel_def by auto
}
thus "vertices G \<times> vertices G = getRel r G" using getRel_dom[OF assms] by auto next
assume a:"vertices G \<times> vertices G = getRel r G"
{ fix f assume a2:"graph_homomorphism (fst (top_rule r)) G f"
hence f:"f `` {0, 1} \<subseteq> vertices G" "on_triple f `` {} \<subseteq> edges G"
"univalent f" "Domain f = {0, 1}"
unfolding top_rule_def prod.sel graph_homomorphism_concr_graph[OF assms graph_empty_e]
by argo+
from a2 have ih:"graph_homomorphism (LG {} {0, 1}) G f" unfolding top_rule_def by auto
have "extensible (top_rule r) G f" unfolding top_rule_def extensible_refl_concr[OF ih]
graph_homomorphism_concr_graph[OF assms graph_single]
using f a[unfolded getRel_def] by fastforce
}
thus "maintained (top_rule r) G" unfolding maintained_def by auto
qed
text \<open>Definition 18.\<close>
(* Definition 18 *)
definition nonempty_rule :: "('l,nat) Graph_PreRule" where
"nonempty_rule = (LG {} {},LG {} {0})"
text \<open>Proof that definition 18 does what it says it does.\<close>
lemma nonempty_rule[simp]:
assumes "graph G"
shows "maintained nonempty_rule G \<longleftrightarrow> vertices G \<noteq> {}"
proof -
have "vertices G = {} \<Longrightarrow> graph_homomorphism (LG {} {0}) G x \<Longrightarrow> False"
"v \<in> vertices G \<Longrightarrow> graph_homomorphism (LG {} {0}) G {(0,v)}"
for v::"'b" and x::"(nat \<times> 'b) set"
unfolding graph_homomorphism_concr_graph[OF assms graph_empty_e] univalent_def by blast+
thus ?thesis unfolding nonempty_rule_def maintained_def extensible_def by (auto intro:assms)
qed
text \<open>Definition 19.\<close>
definition reflexivity_rule :: "'l \<Rightarrow> ('l,nat) Graph_PreRule" where
"reflexivity_rule t = (LG {} {0},LG {(t,0,0)} {0})"
definition symmetry_rule :: "'l \<Rightarrow> ('l,nat) Graph_PreRule" where
"symmetry_rule t = (transl_rule (A_Cnv (A_Lbl t) \<sqsubseteq> A_Lbl t))"
definition transitive_rule :: "'l \<Rightarrow> ('l,nat) Graph_PreRule" where
"transitive_rule t = (transl_rule (A_Cmp (A_Lbl t) (A_Lbl t) \<sqsubseteq> A_Lbl t))"
definition congruence_rule :: "'l \<Rightarrow> 'l \<Rightarrow> ('l,nat) Graph_PreRule" where
"congruence_rule t l = (transl_rule (A_Cmp (A_Cmp (A_Lbl t) (A_Lbl l)) (A_Lbl t) \<sqsubseteq> A_Lbl l))"
abbreviation congruence_rules :: "'l \<Rightarrow> 'l set \<Rightarrow> ('l,nat) Graph_PreRule set"
where
"congruence_rules t L \<equiv> congruence_rule t ` L"
lemma are_rules[intro]:
"graph_rule nonempty_rule"
"graph_rule (top_rule t)"
"graph_rule (reflexivity_rule i)"
unfolding reflexivity_rule_def top_rule_def nonempty_rule_def graph_homomorphism_def
by auto
text \<open>Just before Lemma 7, we remark that if I is an identity, it maintains the identity rules.\<close>
lemma ident_rel_refl:
assumes "graph G" "ident_rel idt G"
shows "maintained (reflexivity_rule idt) G"
unfolding reflexivity_rule_def
proof(rule maintainedI) fix f
assume "graph_homomorphism (LG {} {0::nat}) G f"
hence f:"Domain f = {0}" "graph G" "f `` {0} \<subseteq> vertices G" "univalent f"
unfolding graph_homomorphism_def by force+
from assms(2) univalentD[OF f(4)] f(3)
have "edge_preserving f {(idt, 0, 0)} (edges G)" unfolding edge_preserving
by (auto simp:getRel_def set_eq_iff image_def)
with f have "graph_homomorphism (LG {(idt, 0, 0)} {0}) G f"
"agree_on (LG {} {0}) f f" using assms
unfolding graph_homomorphism_def labeled_graph.sel agree_on_def univalent_def
by auto
then show "extensible (LG {} {0}, LG {(idt, 0, 0)} {0}) G f"
unfolding extensible_def prod.sel by auto
qed
lemma
assumes "ident_rel idt G"
shows ident_rel_trans:"maintained (transitive_rule idt) G"
and ident_rel_symm :"maintained (symmetry_rule idt) G"
and ident_rel_cong :"maintained (congruence_rule idt l) G"
unfolding transitive_rule_def symmetry_rule_def congruence_rule_def
by(intro maintained_holds,insert assms,force)+
text \<open>Definition 19.\<close>
definition identity_rules ::
"'a Standard_Constant set \<Rightarrow> (('a Standard_Constant, nat) Graph_PreRule) set" where
"identity_rules L \<equiv> {reflexivity_rule S_Idt,transitive_rule S_Idt,symmetry_rule S_Idt}
\<union> congruence_rules S_Idt L"
lemma identity_rules_graph_rule:
assumes "x \<in> identity_rules L"
shows "graph_rule x"
proof -
from graph_rule_translation
have gr:"\<And> u v . graph_rule (transl_rule (u \<sqsubseteq> v))" by auto
consider "x = reflexivity_rule S_Idt" | "x = transitive_rule S_Idt" | "x = symmetry_rule S_Idt"
| "\<exists> v w. x = congruence_rule v w" using assms unfolding identity_rules_def Un_iff by blast
thus ?thesis using gr are_rules(3)
unfolding congruence_rule_def transitive_rule_def symmetry_rule_def
by cases fast+
qed
text \<open>Definition 19, showing that the properties indeed do what they claim to do.\<close>
lemma
assumes g[intro]:"graph (G :: ('a, 'b) labeled_graph)"
shows reflexivity_rule: "maintained (reflexivity_rule l) G \<Longrightarrow> refl_on (vertices G) (getRel l G)"
and transitive_rule: "maintained (transitive_rule l) G \<Longrightarrow> trans (getRel l G)"
and symmetry_rule: "maintained (symmetry_rule l) G \<Longrightarrow> sym (getRel l G)"
proof -
{ from assms have gr:"getRel l G \<subseteq> vertices G \<times> vertices G" by (auto simp:getRel_def)
assume m:"maintained (reflexivity_rule l) G" (is "maintained ?r G")
note [simp] = reflexivity_rule_def
show r:"refl_on (vertices G) (getRel l G)"
proof(rule refl_onI[OF gr]) fix x
assume assm:"x \<in> vertices G" define f where "f = {(0::nat,x)}"
have "graph_homomorphism (fst ?r) G f" using assm
by (auto simp:graph_homomorphism_def univalent_def f_def)
from m[unfolded maintained_def] this
obtain g::"(nat\<times>'b) set"
where g:"graph_homomorphism (snd ?r) G g"
"agree_on (fst ?r) f g"
unfolding extensible_def by blast
have "\<And> n v. (n,v) \<in> g \<Longrightarrow> (n = 0) \<and> (v = x)" using g unfolding
agree_on_def graph_homomorphism_def f_def by auto
with g(2) have "g = {(0,x)}" unfolding agree_on_def f_def by auto
with g(1) show "(x,x)\<in> getRel l G"
unfolding graph_homomorphism_def edge_preserving getRel_def by auto
qed
}
{ assume m:"maintained (transitive_rule l) G"
from m[unfolded maintained_holds_subset_iff[OF g] transitive_rule_def]
show "trans (getRel l G)" unfolding trans_def by auto
}
{ assume m:"maintained (symmetry_rule l) G"
from m[unfolded maintained_holds_subset_iff[OF g] symmetry_rule_def]
show "sym (getRel l G)" unfolding sym_def by auto
}
qed
lemma finite_identity_rules[intro]:
assumes "finite L"
shows "finite (identity_rules L)"
using assms unfolding identity_rules_def by auto
lemma equivalence:
assumes gr:"graph G" and m:"maintainedA {reflexivity_rule I,transitive_rule I,symmetry_rule I} G"
shows "equiv (vertices G) (getRel I G)"
proof(rule equivI)
show "refl_on (vertices G) (getRel I G)" using m by(intro reflexivity_rule[OF gr],auto)
show "sym (getRel I G)" using m by(intro symmetry_rule[OF gr],auto)
show "trans (getRel I G)" using m by(intro transitive_rule[OF gr],auto)
qed
lemma congruence_rule:
(* Transitivity is not needed for this proof, but it's more convenient to reuse in this form *)
assumes g:"graph G"
and mA:"maintainedA {reflexivity_rule I,transitive_rule I,symmetry_rule I} G"
and m:"maintained (congruence_rule I l) G"
shows "(\<lambda> v. getRel l G `` {v}) respects (getRel I G)" (is "?g1")
and "(\<lambda> v. (getRel l G)\<inverse> `` {v}) respects (getRel I G)" (is "?g2")
proof -
(* Both parts of this lemma are proved using roughly the same proof. *)
note eq = equivalence[OF g mA]
{ fix y z
assume aI:"(y, z)\<in>getRel I G"
hence a2:"(z, y)\<in>getRel I G" using eq[unfolded equiv_def sym_def] by auto
hence a3:"(z, z)\<in>getRel I G" "(y, y)\<in>getRel I G"
using eq[unfolded equiv_def refl_on_def] by auto
{ fix x
{ assume al:"(y,x) \<in> getRel l G"
hence "x \<in> vertices G" using g unfolding getRel_def by auto
hence r:"(x,x) \<in> getRel I G" using eq[unfolded equiv_def refl_on_def] by auto
note relcompI[OF relcompI[OF a2 al] r]
} note yx = this
{ assume al:"(z,x) \<in> getRel l G"
hence "x \<in> vertices G" using g unfolding getRel_def by auto
hence r:"(x,x) \<in> getRel I G" using eq[unfolded equiv_def refl_on_def] by auto
note relcompI[OF relcompI[OF aI al] r]
} note zx = this
from zx yx m[unfolded maintained_holds_subset_iff[OF g] congruence_rule_def]
have "(y,x) \<in> getRel l G \<longleftrightarrow> (z,x) \<in> getRel l G" by auto
} note v1 = this
{ fix x
{ assume al:"(x,y) \<in> getRel l G"
hence "x \<in> vertices G" using g unfolding getRel_def by auto
hence r:"(x,x) \<in> getRel I G" using eq[unfolded equiv_def refl_on_def] by auto
note relcompI[OF relcompI[OF r al] aI]
} note yx = this
{ assume al:"(x,z) \<in> getRel l G"
hence "x \<in> vertices G" using g unfolding getRel_def by auto
hence r:"(x,x) \<in> getRel I G" using eq[unfolded equiv_def refl_on_def] by auto
note relcompI[OF relcompI[OF r al] a2]
} note zx = this
from zx yx m[unfolded maintained_holds_subset_iff[OF g] congruence_rule_def]
have "(x,y) \<in> getRel l G \<longleftrightarrow> (x,z) \<in> getRel l G" by auto
} note v2 = this
from v1 v2
have "getRel l G `` {y} = getRel l G `` {z}"
"(getRel l G)\<inverse> `` {y} = (getRel l G)\<inverse> `` {z}" by auto
}
thus ?g1 ?g2 unfolding congruent_def by force+
qed
text \<open>Lemma 7, strengthened with an extra property to make subsequent proofs easier to carry out.\<close>
lemma identity_rules:
assumes "graph G"
"maintainedA (identity_rules L) G"
"fst ` edges G \<subseteq> L"
shows "\<exists> f. f o f = f
\<and> ident_rel S_Idt (map_graph_fn G f)
\<and> subgraph (map_graph_fn G f) G
\<and> (\<forall> l x y. (l,x,y) \<in> edges G \<longleftrightarrow> (l,f x,f y) \<in> edges G)"
proof -
(* While this proof defines a concrete f, we only expose it using an existential quantifier.
The reason is that the f of our choice is non-constructive,
and its definition relies on the axiom of choice.
In fact, this theorem applies to the infinite case too,
which means that it's probably equivalent to the axiom of choice.
We therefore have no hopes of giving an executable concrete f here.
In the implementation, we will be able to use finiteness of G (which is not required here),
and therefore we can construct an f with these properties again.
Unfortunately, this does mean doing roughly the same proof twice. *)
have ma:"maintainedA {reflexivity_rule S_Idt, transitive_rule S_Idt, symmetry_rule S_Idt} G"
using assms(2) by (auto simp:identity_rules_def)
note equiv = equivalence[OF assms(1) this]
{ fix l x y
assume "(x, y) \<in> getRel l G" hence l:"l \<in> L" using assms(3) unfolding getRel_def by auto
have r1:"(\<lambda>v. getRel l G `` {v}) respects getRel S_Idt G"
apply(intro congruence_rule[OF assms(1) ma])
using assms(2) l unfolding identity_rules_def by auto
have r2:"(\<lambda>v. (getRel l G)\<inverse> `` {v}) respects getRel S_Idt G"
apply(intro congruence_rule[OF assms(1) ma])
using assms(2) l unfolding identity_rules_def by auto
note congr = r1 r2
} note congr = this
define P where P:"P = (\<lambda> x y. y \<in> getRel S_Idt G `` {x})"
{ fix x
assume a:"getRel S_Idt G `` {x} \<noteq> {}"
hence "\<exists> y. P x y" unfolding P by auto
hence p:"P x (Eps (P x))" unfolding some_eq_ex by auto
{ fix y
assume b:"P x y"
hence "(x,y) \<in> getRel S_Idt G" unfolding P by auto
from equiv_class_eq[OF equiv this]
have "getRel S_Idt G `` {x} = getRel S_Idt G `` {y}".
} note u = this[OF p]
have "getRel S_Idt G `` {Eps (P x)} = getRel S_Idt G `` {x}"
unfolding u by (fact refl)
hence "Eps (P (Eps (P x))) = Eps (P x)" unfolding P by auto
} note P_eq = this
define f where f:"f = (\<lambda> x. (if getRel S_Idt G `` {x} = {} then x else (SOME y. P x y)))"
have "(f \<circ> f) x = f x" for x proof(cases "getRel S_Idt G `` {x} = {}")
case False
then show ?thesis using P_eq by (simp add:o_def f)
qed (auto simp:o_def f)
hence idemp: "f o f = f" by auto
from equivE equiv have refl:"refl_on (vertices G) (getRel S_Idt G)" by auto
hence [intro]:"x \<in> vertices G \<Longrightarrow> (x, x) \<in> getRel S_Idt G" for x unfolding refl_on_def by auto
hence vert_P:"x \<in> vertices G \<Longrightarrow> (x, Eps (P x)) \<in> getRel S_Idt G" for x
unfolding P getRel_def by (metis tfl_some Image_singleton_iff getRel_def)
have r1:"x \<in> vertices G \<longleftrightarrow> P x x" for x using refl unfolding refl_on_def P by auto
have r2[simp]:"getRel S_Idt G `` {x} = {} \<longleftrightarrow> x \<notin> vertices G" for x
using refl assms(1) unfolding refl_on_def by auto
{ fix x y assume "(S_Idt,x,y)\<in> edges G"
hence "(x,y) \<in> getRel S_Idt G" unfolding getRel_def by auto
hence "getRel S_Idt G `` {x} = getRel S_Idt G `` {y}"
using equiv_class_eq[OF equiv] by metis
hence "Eps (P x) = Eps (P y)" unfolding P by auto
} note idt_eq = this
have ident:"ident_rel S_Idt (map_graph_fn G f)"
proof(rule ident_relI,goal_cases)
case (1 x) thus ?case unfolding f by auto
next case (2 x y) thus ?case unfolding getRel_def by (auto simp:f intro!:idt_eq)
next case (3 x y) thus ?case unfolding getRel_def by auto
qed
{ fix l x y
assume a:"(l,x,y) \<in> edges G" "x \<in> vertices G" "y \<in> vertices G"
hence f:"(f x, x) \<in> getRel S_Idt G" "(f y, y) \<in> getRel S_Idt G"
using vert_P equivE[OF equiv] sym_def unfolding f by auto
from a have gr:"(x, y) \<in> getRel l G" unfolding getRel_def by auto
from congruentD[OF congr(1)[OF gr] f(1)] congruentD[OF congr(2)[OF gr] f(2)] a(1)
have "(l,f x, f y) \<in> edges G" unfolding set_eq_iff getRel_def by auto
} note gu1 = this
{ fix x assume a: "x \<in> vertices G"
with vert_P have "(x,Eps (P x)) \<in> getRel S_Idt G" by auto
hence "Eps (P x) \<in> vertices G" using assms(1) unfolding getRel_def by auto
hence "f x \<in> vertices G" using a unfolding f by auto
} note gu2 = this
have "graph_union (map_graph_fn G f) G = G"
using gu1 gu2 assms(1) unfolding graph_union_def by(cases G,auto)
hence subg: "subgraph (map_graph_fn G f) G"
unfolding subgraph_def using assms(1) by auto
have congr:"((l, x, y) \<in> edges G) = ((l, f x, f y) \<in> edges G)" for l x y proof
assume a:"((l, f x, f y) \<in> edges G)"
hence gr:"(f x, f y) \<in> getRel l G" unfolding getRel_def by auto
from a have fv:"f x \<in> vertices G" "f y \<in> vertices G" using assms(1) by auto
{ fix x assume a:"f x \<in> vertices G" "x \<notin> vertices G"
with assms(1) have "getRel S_Idt G `` {x} = {}" by auto
with a f have False by auto
}
with fv have v:"x \<in> vertices G" "y \<in> vertices G" by auto
have gx:"(x, f x) \<in> getRel S_Idt G" and gy:"(y, f y) \<in> getRel S_Idt G" by (auto simp: f v vert_P)
from congruentD[OF congr(1)[OF gr] gx] gr
have "(x, f y) \<in> getRel l G" by auto
with congruentD[OF congr(2)[OF gr] gy]
have "(x, y) \<in> getRel l G" by auto
thus "((l, x, y) \<in> edges G)" unfolding getRel_def by auto next
assume e:"((l, x, y) \<in> edges G)"
hence "x \<in> vertices G" "y \<in> vertices G" using assms(1) by auto
from gu1[OF e this]
show "((l, f x, f y) \<in> edges G)".
qed
from idemp ident subg congr show ?thesis by auto
qed
text \<open>The idempotency property of Lemma 7 suffices to show that 'maintained' is preserved.\<close>
lemma idemp_embedding_maintained_preserved:
assumes subg:"subgraph (map_graph_fn G f) G" and f:"\<And> x. x\<in>vertices G \<Longrightarrow> (f o f) x = f x"
and maint:"maintained r G"
shows "maintained r (map_graph_fn G f)"
proof -
{ fix h assume hom_h:"graph_homomorphism (fst r) (map_graph_fn G f) h"
from subgraph_preserves_hom[OF subg this] maint[unfolded maintained_def extensible_def]
obtain g where g:"graph_homomorphism (snd r) G g"
"agree_on (fst r) h g" by blast
{ fix v x
have subs:"h `` {v} \<subseteq> vertices (map_graph_fn G f)"
using hom_h[unfolded graph_homomorphism_def] by auto
assume "v\<in>vertices (fst r)" and x:"(v, x) \<in> g"
hence "g `` {v} = h `` {v}" using g(2)[unfolded agree_on_def,rule_format,of v] by auto
hence "g `` {v} \<subseteq> vertices (map_graph_fn G f)" using subs by auto
hence x2:"x \<in> vertices (map_graph_fn G f)" using x by auto
then obtain y where "x = f y" "y \<in> vertices G" by auto
hence f:"f x = x" using f x2 unfolding o_def by metis
from x2 subgraph_subset[OF subg] have "(x, f x) \<in> on_graph G f" by auto
with x have "(v, x) \<in> g O on_graph G f" "f x = x" unfolding f by auto
}
hence agr:"agree_on (fst r) h (g O on_graph G f)"
using g(2) unfolding agree_on_def by auto
have "extensible r (map_graph_fn G f) h"
unfolding extensible_def using graph_homomorphism_on_graph[OF g(1)] agr by blast
}
thus ?thesis unfolding maintained_def by blast
qed
text \<open>Definition 20.\<close>
definition const_exists where
"const_exists c \<equiv> transl_rule (\<top> \<sqsubseteq> A_Cmp (A_Cmp \<top> (A_Lbl (S_Const c))) \<top>)"
definition const_exists_rev where
"const_exists_rev c \<equiv> transl_rule (A_Cmp (A_Cmp (A_Lbl (S_Const c)) \<top>) (A_Lbl (S_Const c)) \<sqsubseteq> A_Lbl (S_Const c))"
definition const_prop where
"const_prop c \<equiv> transl_rule (A_Lbl (S_Const c) \<sqsubseteq> \<one>)"
definition const_disj where
"const_disj c\<^sub>1 c\<^sub>2 \<equiv> transl_rule (A_Cmp (A_Lbl (S_Const c\<^sub>1)) (A_Lbl (S_Const c\<^sub>2)) \<sqsubseteq> \<bottom>)"
lemma constant_rules:
assumes "standard' C G" "c \<in> C"
shows "maintained (const_exists c) G"
"maintained (const_exists_rev c) G"
"maintained (const_prop c) G"
"c' \<in> C \<Longrightarrow> c \<noteq> c' \<Longrightarrow> maintained (const_disj c c') G"
proof -
note a = assms[unfolded standard_def]
from a have g:"graph G" by auto
from a
have gr_c:"getRel (S_Const c) G = {(Inl c, Inl c)}"
"getRel S_Idt G = Id_on (vertices G)" "getRel S_Bot G = {}"
"getRel S_Top G = vertices G \<times> vertices G" by auto
with g have inlc:"Inl c \<in> vertices G" by (metis getRel_dom(1) singletonI)
thus "maintained (const_exists c) G" "maintained (const_exists_rev c) G"
"maintained (const_prop c) G"
unfolding const_prop_def const_exists_rev_def const_exists_def maintained_holds_subset_iff[OF g]
by (auto simp:gr_c relcomp_unfold)
assume "c' \<in> C"
with a have gr_c':"getRel (S_Const c') G = {(Inl c', Inl c')}" by auto
thus "c \<noteq> c' \<Longrightarrow> maintained (const_disj c c') G"
unfolding const_disj_def maintained_holds_subset_iff[OF g] using gr_c by auto
qed
definition constant_rules where
"constant_rules C \<equiv> const_exists ` C \<union> const_exists_rev ` C \<union> const_prop ` C
\<union> {const_disj c\<^sub>1 c\<^sub>2 | c\<^sub>1 c\<^sub>2. c\<^sub>1 \<in> C \<and> c\<^sub>2 \<in> C \<and> c\<^sub>1 \<noteq> c\<^sub>2}"
lemma constant_rules_graph_rule:
assumes "x \<in> constant_rules C"
shows "graph_rule x"
proof -
from graph_rule_translation
have gr:"\<And> u v . graph_rule (transl_rule (u \<sqsubseteq> v))" by auto
consider "\<exists> v. x = const_exists v" | "\<exists> v. x = const_exists_rev v" | "\<exists> v. x = const_prop v"
| "\<exists> v w. x = const_disj v w" using assms unfolding constant_rules_def Un_iff by blast
thus ?thesis using gr
unfolding const_exists_def const_exists_rev_def const_prop_def const_disj_def
by cases fast+
qed
lemma finite_constant[intro]:
assumes "finite C"
shows "finite (constant_rules C)"
proof -
have "{const_disj c\<^sub>1 c\<^sub>2 | c\<^sub>1 c\<^sub>2. c\<^sub>1 \<in> C \<and> c\<^sub>2 \<in> C \<and> c\<^sub>1 \<noteq> c\<^sub>2} \<subseteq> case_prod const_disj ` (C \<times> C)"
by auto
moreover have "finite \<dots>" using assms by auto
ultimately have "finite {const_disj c\<^sub>1 c\<^sub>2 | c\<^sub>1 c\<^sub>2. c\<^sub>1 \<in> C \<and> c\<^sub>2 \<in> C \<and> c\<^sub>1 \<noteq> c\<^sub>2}"
by(rule finite_subset)
thus ?thesis unfolding constant_rules_def using assms by blast
qed
lemma standard_maintains_constant_rules:
assumes "standard' C G" "R\<in>constant_rules C"
shows "maintained R G"
proof -
from assms(2)[unfolded constant_rules_def]
consider "\<exists> c \<in> C. R = const_exists c"
| "\<exists> c \<in> C. R = const_exists_rev c"
| "\<exists> c \<in> C. R = const_prop c"
| "\<exists> c\<^sub>1 c\<^sub>2. c\<^sub>1 \<in> C \<and> c\<^sub>2 \<in> C \<and> c\<^sub>1 \<noteq> c\<^sub>2 \<and> R = const_disj c\<^sub>1 c\<^sub>2" by blast
from this assms(1) show ?thesis by(cases,auto simp:constant_rules)
qed
lemma constant_rules_empty[simp]:
"constant_rules {} = {}"
by (auto simp:constant_rules_def)
text \<open>Definition 20, continued.\<close>
definition standard_rules :: "'a set \<Rightarrow> 'a Standard_Constant set \<Rightarrow> (('a Standard_Constant, nat) labeled_graph \<times> ('a Standard_Constant, nat) labeled_graph) set"
where
"standard_rules C L \<equiv> constant_rules C \<union> identity_rules L \<union> {top_rule S_Top,nonempty_rule}"
lemma constant_rules_mono:
assumes "C\<^sub>1 \<subseteq> C\<^sub>2"
shows "constant_rules C\<^sub>1 \<subseteq> constant_rules C\<^sub>2"
using assms unfolding constant_rules_def
by(intro Un_mono,auto) (* also works with just auto, this is faster *)
lemma identity_rules_mono:
assumes "C\<^sub>1 \<subseteq> C\<^sub>2"
shows "identity_rules C\<^sub>1 \<subseteq> identity_rules C\<^sub>2"
using assms unfolding identity_rules_def by auto
lemma standard_rules_mono:
assumes "C\<^sub>1 \<subseteq> C\<^sub>2" "L\<^sub>1 \<subseteq> L\<^sub>2"
shows "standard_rules C\<^sub>1 L\<^sub>1 \<subseteq> standard_rules C\<^sub>2 L\<^sub>2"
using constant_rules_mono[OF assms(1)] identity_rules_mono[OF assms(2)]
unfolding standard_rules_def by auto
lemma maintainedA_invmono:
assumes "C\<^sub>1 \<subseteq> C\<^sub>2" "L\<^sub>1 \<subseteq> L\<^sub>2"
shows "maintainedA (standard_rules C\<^sub>2 L\<^sub>2) G \<Longrightarrow> maintainedA (standard_rules C\<^sub>1 L\<^sub>1) G"
using standard_rules_mono[OF assms] by auto
lemma maintained_preserved_by_isomorphism:
assumes "\<And> x. x \<in> vertices G \<Longrightarrow> (f \<circ> g) x = x" "graph G"
and "maintained r (map_graph_fn G g)"
shows "maintained r G"
proof(cases r)
case (Pair L R)
show ?thesis unfolding Pair proof(standard,goal_cases)
case (1 h)
from assms(3)[unfolded maintained_def Pair] graph_homomorphism_on_graph[OF this, of g]
have "extensible (L, R) (map_graph_fn G g) (h O on_graph G g)" by auto
then obtain h2
where h2:"graph_homomorphism R (map_graph_fn G g) h2" "agree_on L (h O on_graph G g) h2"
unfolding extensible_def by auto
from 1 have h_id:"h O Id_on (vertices G) = h" unfolding graph_homomorphism_def by auto
let ?h = "h2 O on_graph (map_graph_fn G g) f"
from assms(1) have "on_graph G (f \<circ> g) = Id_on (vertices G)" by auto
hence "map_graph_fn G (f \<circ> g) = G" using assms(2) map_graph_fn_id by auto
with graph_homomorphism_on_graph[OF h2(1),of f]
have igh:"graph_homomorphism R G ?h" by auto
have "g x = g xa \<Longrightarrow> x \<in> (vertices G) \<Longrightarrow> xa \<in> (vertices G) \<Longrightarrow> x = xa"
for x xa using assms(1) o_def by metis
hence "g x = g xa \<Longrightarrow> x \<in> (vertices G) \<Longrightarrow> xa \<in> (vertices G) \<Longrightarrow> (x, xa) \<in> Id_on (vertices G)"
for x xa by auto
hence id:"(on_graph G g) O on_graph (map_graph_fn G g) f = Id_on (vertices G)"
using assms(1) by auto
from agree_on_ext[OF h2(2),of "on_graph (map_graph_fn G g) f",unfolded O_assoc]
have agh:"agree_on L h ?h" unfolding agree_on_def id h_id.
from igh agh show ?case unfolding extensible_def by auto
qed
qed
lemma standard_identity_rules:
assumes "standard' C G"
shows "maintained (reflexivity_rule S_Idt) G"
"maintained (transitive_rule S_Idt) G"
"maintained (symmetry_rule S_Idt) G"
"maintained (congruence_rule S_Idt l) G"
proof -
note a = assms[unfolded standard_def]
from a have g:"graph G" by auto
from a
have gr:"getRel S_Idt G = Id_on (vertices G)" "getRel S_Bot G = {}"
"getRel S_Top G = vertices G \<times> vertices G"
and v_gr:"\<forall>a b. ((S_Idt, a, b) \<in> edges G) = (a \<in> vertices G \<and> b = a)"
unfolding getRel_def by auto
thus "maintained (transitive_rule S_Idt) G" "maintained (symmetry_rule S_Idt) G"
"maintained (congruence_rule S_Idt l) G"
unfolding transitive_rule_def symmetry_rule_def congruence_rule_def
maintained_holds_subset_iff[OF g]
by (auto simp:gr relcomp_unfold)
{ fix f :: "(nat \<times> ('a + 'b)) set"
assume "graph_homomorphism (LG {} {0}) G f"
hence u:"univalent f" and d:"Domain f = {0}"
and r:"f `` {0} \<subseteq> vertices G" unfolding graph_homomorphism_def by simp+
from d obtain v where v:"(0,v) \<in> f" by auto
hence f:"f = {(0,v)}"
using d insert_iff mk_disjoint_insert all_not_in_conv old.prod.exhaust
u[unfolded univalent_def] Domain.intros[of _ _ f,unfolded d,THEN singletonD]
by (metis (no_types))
from v r have v:"v \<in> vertices G" by auto
with v_gr have "(S_Idt, v, v) \<in> edges G" by auto
hence "edge_preserving {(0, v)} {(S_Idt, 0, 0)} (edges G)" unfolding edge_preserving by auto
hence "graph_homomorphism (LG {(S_Idt, 0, 0)} {0}) G f" unfolding f
graph_homomorphism_def using g v by (auto simp:univalent_def)
}
thus "maintained (reflexivity_rule S_Idt) G"
unfolding reflexivity_rule_def maintained_def by auto
qed
lemma standard_maintains_identity_rules:
assumes "standard' C G" "x\<in>identity_rules L"
shows "maintained x G"
proof -
consider "x = reflexivity_rule S_Idt" | "x = transitive_rule S_Idt" | "x = symmetry_rule S_Idt"
| "\<exists> l. x = congruence_rule S_Idt l" using assms unfolding identity_rules_def Un_iff by blast
thus ?thesis using standard_identity_rules[OF assms(1)] by(cases,auto)
qed
lemma standard_maintains_rules:
assumes "standard' C G"
shows "maintainedA (standard_rules C L) G"
proof fix R
assume "R \<in> standard_rules C L"
then consider "R \<in> constant_rules C" | "R \<in> identity_rules L"
| "R = top_rule S_Top" | "R = nonempty_rule" by (auto simp:standard_rules_def)
thus "maintained R G"
using assms standard_maintains_constant_rules[OF assms]
standard_maintains_identity_rules[OF assms] by (cases,auto simp:standard_def)
qed
text \<open>A case-split rule.\<close>
lemma standard_rules_edges:
assumes "(lhs, rhs) \<in> standard_rules C L" "(l, x, y) \<in> edges rhs"
shows "(l = S_Bot \<Longrightarrow> thesis) \<Longrightarrow>
(l = S_Top \<Longrightarrow> thesis) \<Longrightarrow>
(l = S_Idt \<Longrightarrow> thesis) \<Longrightarrow>
(l \<in> S_Const ` C \<Longrightarrow> thesis) \<Longrightarrow>
(l \<in> L \<Longrightarrow> thesis) \<Longrightarrow> thesis"
- using assms
- by (auto simp:Let_def standard_rules_def constant_rules_def identity_rules_def
+ using assms [[simproc del: defined_all]]
+ by (auto simp: Let_def standard_rules_def constant_rules_def identity_rules_def
const_exists_def const_exists_rev_def const_prop_def const_disj_def
reflexivity_rule_def transitive_rule_def symmetry_rule_def congruence_rule_def
top_rule_def nonempty_rule_def)
+
text \<open>Lemma 8.
This is a slightly stronger version of Lemma 8:
we reason about maintained rather than holds,
and the quantification for maintained happens within the existential quantifier, rather than outside.
Due to the type system of Isabelle, we construct the concrete type @{term std_graph} for G.
This in contrast to arguing that 'there exists a type large enough', as in the paper.\<close>
lemma maintained_standard_noconstants:
assumes mnt:"maintainedA (standard_rules C L) G'"
and gr:"graph (G'::('V Standard_Constant, 'V') labeled_graph)"
"fst ` edges G' \<subseteq> L" (* Graph on labels L *)
and cf:"getRel S_Bot G' = {}" (* Conflict free *)
shows "\<exists> f g (G::('V, 'V') std_graph). G = map_graph_fn G (f o g) \<and> G = map_graph_fn G' f
\<and> subgraph (map_graph_fn G g) G'
\<and> standard' C G
\<and> (\<forall> r. maintained r G' \<longrightarrow> maintained r G)
\<and> (\<forall> x y e. x \<in> vertices G' \<longrightarrow> y \<in> vertices G' \<longrightarrow>
(g (f x), g (f y)) \<in> :map_graph_fn G g:\<lbrakk>e\<rbrakk> \<longrightarrow> (x,y) \<in> :G':\<lbrakk>e\<rbrakk>)"
proof -
note mnt = mnt[unfolded standard_rules_def]
from mnt have "maintainedA (identity_rules L) G'" by auto
from identity_rules[OF gr(1) this gr(2)] obtain h where
h:"h \<circ> h = h" "ident_rel S_Idt (map_graph_fn G' h)" "subgraph (map_graph_fn G' h) G'"
"((l, x, y) \<in> edges G') = ((l, h x, h y) \<in> edges G')" for l x y by blast
have mg:"\<And> r. maintained r G' \<Longrightarrow> maintained r (map_graph_fn G' h)"
using idemp_embedding_maintained_preserved[OF h(3)] h(1) by auto
from mnt have tr:"maintained (top_rule S_Top) G'" and ne:"maintained nonempty_rule G'" by auto
from nonempty_rule[OF gr(1)] ne obtain x where x:"x \<in> vertices G'" by blast
from tr[unfolded top_rule[OF gr(1)]] x have top_nonempty:"(x, x) \<in> getRel S_Top G'" by auto
have "\<And> c. c \<in> C \<Longrightarrow> \<exists>v. (v, v) \<in> getRel (S_Const c) (map_graph_fn G' h)" proof(goal_cases)
case (1 c)
with mnt have cr5: "maintained (const_exists c) G'"
and cr7: "maintained (const_prop c) G'" unfolding constant_rules_def by blast+
from top_nonempty cr5[unfolded maintained_holds_subset_iff[OF gr(1)] const_exists_def]
obtain y z where yz:"(y,z) \<in> getRel (S_Const c) G'" by auto
from this gr(1) have yzv:"y \<in> vertices G'" "z \<in> vertices G'" by (auto simp:getRel_def)
from getRel_hom[OF yz yzv]
have hi:"(h y,h z) \<in> getRel (S_Const c) (map_graph_fn G' h)".
with h(2) cr7[THEN mg,unfolded maintained_holds_subset_iff[OF map_graph_fn_graphI] const_prop_def]
have "h y = h z" by force
thus "\<exists> v. (v,v) \<in> getRel (S_Const c) (map_graph_fn G' h)" using hi by auto
qed
hence "\<forall> c. \<exists> v. c \<in> C \<longrightarrow> (v, v) \<in> getRel (S_Const c) (map_graph_fn G' h)" by blast
from choice[OF this] obtain m
where m:"\<And> x. x \<in> C \<Longrightarrow> (m x, m x) \<in> getRel (S_Const x) (map_graph_fn G' h)" by blast
let ?m' = "\<lambda> x. if x \<in> m ` C then Inl (the_inv_into C m x) else Inr x"
define f where "f \<equiv> ?m' o h"
have "\<And> x y. x \<in> C \<Longrightarrow> y \<in> C \<Longrightarrow> m x = m y \<Longrightarrow> x = y" proof(goal_cases)
case (1 x y)
with m have "(m x,m x) \<in> getRel (S_Const y) (map_graph_fn G' h)"
"(m x,m x) \<in> getRel (S_Const x) (map_graph_fn G' h)" by metis+
hence mx: "(m x,m x) \<in> getRel (S_Const y) G'"
"(m x,m x) \<in> getRel (S_Const x) G'" using h(3) by force+
from 1(1,2) mnt have cr8:"x \<noteq> y \<Longrightarrow> maintained (const_disj x y) G'"
unfolding constant_rules_def by blast
from cr8[unfolded maintained_holds_subset_iff[OF gr(1)] const_disj_def] mx
have "x\<noteq>y\<Longrightarrow>(m x,m x) \<in> :G':\<lbrakk>\<bottom>\<rbrakk>" by auto
thus "x = y" using cf by auto
qed
hence "univalent (converse (BNF_Def.Gr C m))" unfolding univalent_def by auto
hence inj_m:"inj_on m C" unfolding inj_on_def by auto
from inj_on_the_inv_into[OF inj_m] have inj_m':"inj ?m'" unfolding inj_on_def by auto
define G where "G = map_graph_fn G' f"
hence G:"graph G" "f x \<in> vertices G" "getRel S_Bot G = {}" using x cf unfolding getRel_def
by force+
from comp_inj_on[OF inj_on_the_inv_into[OF inj_m] inj_Inl, unfolded o_def] inj_Inr
have inj_m':"inj_on ?m' (vertices G')" unfolding inj_on_def by auto
define g where "g = the_inv_into (vertices G') ?m'"
have gf_h:"\<And> x. x \<in> vertices G' \<Longrightarrow> (g o f) x = h x" unfolding g_def f_def o_def
apply(rule the_inv_into_f_f[OF inj_m']) using h unfolding subgraph_def graph_union_iff by auto
have mg_eq:"map_graph_fn G' (g \<circ> f) = map_graph_fn G' h"
by (rule map_graph_fn_eqI[OF gf_h])
have "\<And> x. x \<in> vertices G' \<Longrightarrow> h x \<in> vertices G'" using h(3)
unfolding subgraph_def graph_union_iff by(cases G',auto)
hence gf_id:"\<And> x. x \<in> vertices G' \<Longrightarrow> (g o f) (h x) = (h x)"
using h(1) gf_h unfolding o_def by metis
{ fix x assume "x \<in> vertices G"
then obtain y where y:"f y = x" "y \<in> vertices G'" unfolding G_def by auto
from gf_h[OF y(2)] have "(f o g) (f y) = f (h y)" unfolding o_def by auto
also have "\<dots> = f y" using h(1) unfolding f_def o_def by metis
finally have "(f o g) x = x" unfolding y.
} note fg_id = this
have fg_inv:"map_graph_fn G (f o g) = G"
using h(1) G_def f_def mg_eq map_graph_fn_comp by (metis (no_types, lifting))
have ir:"ident_rel S_Idt G" unfolding set_eq_iff proof(standard,standard,goal_cases)
case (1 x)
from this[unfolded G_def]
obtain v1 v2 where v:"(v1,v2) \<in> getRel S_Idt G'" "x = (f v1,f v2)"
unfolding getRel_def map_graph_def on_triple_def by auto
hence vv:"v1 \<in> vertices G'" "v2 \<in> vertices G'" using gr unfolding getRel_def by auto
with h(2) v(1) have "h v1 = h v2" unfolding image_def by blast
hence x:"x = (f v1,f v1)" unfolding f_def v by auto
from vv(1) show ?case unfolding x G_def by auto
next
case (2 x)
hence x:"fst x = snd x" "fst x \<in> vertices G" by auto
hence "(fst x) \<in> f ` vertices G'" unfolding G_def o_def by auto
then obtain v where v:"v \<in> vertices G'" "f v = fst x" by auto
hence hv:"h v \<in> vertices (map_graph_fn G' h)" by simp
hence "(h v,h v) \<in> getRel S_Idt (map_graph_fn G' h)" unfolding h(2) by auto
from getRel_hom[OF this hv hv]
have "(?m' (h v),?m' (h v)) \<in> getRel S_Idt (map_graph_fn G' (?m' o h))"
unfolding map_graph_fn_comp by fast
hence "(f v,f v) \<in> getRel S_Idt (map_graph_fn G' f)" unfolding f_def by auto
hence "(fst x,snd x) \<in> getRel S_Idt G" unfolding x v G_def by auto
thus ?case unfolding G_def by auto
qed
from tr[unfolded top_rule[OF gr(1)]]
have tr0:"getRel S_Top (map_graph_fn G' h)
= {(x,y). x \<in> vertices (map_graph_fn G' h) \<and> y \<in> vertices (map_graph_fn G' h)}"
and tr:"getRel S_Top G = {(x, y). x \<in> vertices G \<and> y \<in> vertices G}"
unfolding G_def getRel_def on_triple_def map_graph_def by auto
have m:"\<And> x. x \<in> C \<Longrightarrow> {(m x, m x)} = getRel (S_Const x) (map_graph_fn G' h)" proof fix x
assume x:"x \<in> C"
{ fix y z assume a:"(y,z) \<in> getRel (S_Const x) (map_graph_fn G' h)"
let ?t = "getRel S_Top (map_graph_fn G' h)"
let ?r = "getRel (S_Const x) (map_graph_fn G' h)"
have mx:"(m x,m x) \<in> getRel (S_Const x) (map_graph_fn G' h)" using m x by auto
with a have v:"y \<in> vertices (map_graph_fn G' h)"
"z \<in> vertices (map_graph_fn G' h)"
"m x \<in> vertices (map_graph_fn G' h)" unfolding getRel_def by force+
with tr0 have "(m x,y) \<in> ?t" "(z,m x) \<in> ?t" by auto
with a mx have lhs:"(m x,z) \<in> ?r O ?t O ?r" "(y,m x) \<in> ?r O ?t O ?r" by auto
from x mnt have "maintained (const_exists_rev x) G'"
and "maintained (const_prop x) G'" unfolding constant_rules_def by blast+
hence cr6:"maintained (const_exists_rev x) (map_graph_fn G' h)"
and cr7:"maintained (const_prop x) (map_graph_fn G' h)"
by (intro mg, force)+
hence "(m x,z) \<in> getRel S_Idt (map_graph_fn G' h)"
"(y,m x) \<in> getRel S_Idt (map_graph_fn G' h)" using lhs
unfolding maintained_holds_subset_iff[OF map_graph_fn_graphI]
const_exists_rev_def const_prop_def by auto
hence "y = m x" "z = m x" using h(2) by auto
}
thus "getRel (S_Const x) (map_graph_fn G' h) \<subseteq> {(m x, m x)}" by auto
qed (insert m,auto)
from mg_eq have mg_eq:"map_graph_fn G g = map_graph_fn G' h" unfolding G_def map_graph_fn_comp.
{ fix l fix v::"'V + 'V'"
assume a:"(l, v)\<in>(\<lambda>c. (S_Const c, Inl c)) ` C"
hence "getRel l G = {(v, v)}" using m proof(cases l)
case (S_Const x) hence x:"l = S_Const x" "v = Inl x" "x \<in> C" using a by auto
hence mx:"m x \<in> m ` C" by auto
from m[OF x(3)] have "(m x,m x) \<in> getRel (S_Const x) (map_graph_fn G' h)" by auto
hence "(S_Const x,m x,m x) \<in> edges (map_graph_fn G' h)" unfolding getRel_def by auto
hence "m x \<in> vertices (map_graph_fn G' h)" unfolding map_graph_def Image_def by auto
then obtain x' where x':"m x = h x'" "x' \<in> vertices G'" by auto
from h(1) have hmx[simp]:"h (m x) = m x" unfolding x' o_def by metis
hence fmx:"f (m x) = v" unfolding x f_def
using the_inv_into_f_f[OF inj_m] inj_m[unfolded inj_on_def,rule_format,OF x(3)] mx by auto
have "{(f (m x), f (m x))} = getRel l (map_graph_fn G (f \<circ> g))"
unfolding map_graph_fn_comp getRel_hom_map[OF map_graph_fn_graphI]
m[OF x(3),folded mg_eq x(1),symmetric] by auto
hence gr:"getRel l G = {(f (m x), f (m x))}" unfolding fg_inv by blast
show ?thesis unfolding gr fmx by (fact refl)
qed auto
} note cr = this
have sg:"subgraph (map_graph_fn G g) G'" unfolding mg_eq using h(3).
have std:"standard' C G" unfolding standard_def using G ir tr cr by blast
have mtd:"\<And>r. maintained r G' \<Longrightarrow> maintained r G" proof(goal_cases)
case (1 r) from mg[OF 1,folded mg_eq] maintained_preserved_by_isomorphism[OF fg_id G(1)]
show ?case by metis
qed
{ fix x y e
assume "x \<in> vertices G'" "y \<in> vertices G'"
"(g (f x), g (f y)) \<in> :map_graph_fn (map_graph_fn G' f) g:\<lbrakk>e\<rbrakk>"
hence "(x,y) \<in> :G':\<lbrakk>e\<rbrakk>"
proof(induct e arbitrary: x y)
case (A_Cmp e1 e2)
then obtain z where z:"(g (f x), z) \<in> :map_graph_fn (map_graph_fn G' f) g:\<lbrakk>e1\<rbrakk>"
"(z, g (f y)) \<in> :map_graph_fn (map_graph_fn G' f) g:\<lbrakk>e2\<rbrakk>" by auto
hence "z \<in> vertices (map_graph_fn (map_graph_fn G' f) g)"
using semantics_in_vertices(1)[OF map_graph_fn_graphI] by metis
then obtain z' where z':"z = g (f z')" "z' \<in> vertices G'" by auto
with A_Cmp(1)[OF A_Cmp(3) z'(2) z(1)[unfolded z']]
A_Cmp(2)[OF z'(2) A_Cmp(4) z(2)[unfolded z']]
have "(x, y) \<in> (:G':\<lbrakk>e1\<rbrakk>) O (:G':\<lbrakk>e2\<rbrakk>)" by auto
then show ?case by auto
next
case (A_Lbl l)
hence "(l, g (f x), g (f y)) \<in> edges (map_graph_fn G g)"
by (auto simp:getRel_def G_def)
then obtain x' y'
where "(l, x', y') \<in> edges G" "g (f x) = g x'" "g (f y) = g y'" by auto
then obtain x' y'
where xy:"(l, x', y') \<in> edges G'" "g (f x) = g (f x')" "g (f y) = g (f y')"
unfolding G_def by auto
hence "x' \<in> vertices G'" "y' \<in> vertices G'" using gr(1) by auto
from this[THEN gf_h,unfolded o_def] A_Lbl(1,2)[THEN gf_h,unfolded o_def]
have "h x = h x'" "h y = h y'" using xy(2,3) by auto
hence "(l, x, y) \<in> edges G'" using h(4)[of l x y] h(4)[of l x' y'] xy(1) by auto
then show ?case by (simp add:getRel_def)
qed auto
}
hence cons:"(\<forall> x y e. x \<in> vertices G' \<longrightarrow> y \<in> vertices G' \<longrightarrow> (g (f x), g (f y)) \<in> :map_graph_fn G g:\<lbrakk>e\<rbrakk> \<longrightarrow> (x,y) \<in> :G':\<lbrakk>e\<rbrakk>)"
unfolding G_def by auto
show ?thesis using cons G_def fg_inv[symmetric] sg std mtd by blast
qed
end
\ No newline at end of file
diff --git a/thys/Group-Ring-Module/Algebra1.thy b/thys/Group-Ring-Module/Algebra1.thy
--- a/thys/Group-Ring-Module/Algebra1.thy
+++ b/thys/Group-Ring-Module/Algebra1.thy
@@ -1,7227 +1,7183 @@
(** Algebra1
author Hidetsune Kobayashi
Department of Mathematics
Nihon University
hikoba@math.cst.nihon-u.ac.jp
May 3, 2004.
April 6, 2007 (revised)
chapter 0. Preliminaries
section 1. Natural numbers and Integers
section 2. Sets
section 3. Functions
section 4. Nsets, set of natural numbers
section 4'. Lower bounded set of integers
section 5. Augmented integer: integer and \<infinity> -\<infinity>
section 6. amin, amax
section 7. cardinality of sets
Note. Some lemmas in this chapter are already formalized by L. Paulson
and others.
chapter 1. Ordered Set
section 1. Basic Concepts of Ordered Sets
**)
theory Algebra1
imports Main "HOL-Library.FuncSet"
begin
chapter "Preliminaries"
text\<open>Some of the lemmas of this section are proved in src/HOL/Integ
of Isabelle version 2003.\<close>
section "Lemmas for logical manipulation"
lemma True_then:"True \<longrightarrow> P \<Longrightarrow> P"
by simp
lemma ex_conjI:"\<lbrakk>P c; Q c\<rbrakk> \<Longrightarrow> \<exists>c. P c \<and> Q c"
by blast
lemma forall_spec:"\<lbrakk> \<forall>b. P b \<longrightarrow> Q b; P a\<rbrakk> \<Longrightarrow> Q a"
by simp
lemma a_b_exchange:"\<lbrakk>a; a = b\<rbrakk> \<Longrightarrow> b"
by simp
lemma eq_prop:"\<lbrakk> P; P = Q\<rbrakk> \<Longrightarrow> Q"
by simp
lemma forball_contra:"\<lbrakk>\<forall>y\<in>A. P x y \<longrightarrow> \<not> Q y; \<forall>y\<in>A. Q y \<or> R y\<rbrakk> \<Longrightarrow>
\<forall>y\<in>A. (\<not> P x y) \<or> R y"
by blast
lemma forball_contra1:"\<lbrakk>\<forall>y\<in>A. P x y \<longrightarrow> Q y; \<forall>y\<in>A. \<not> Q y\<rbrakk> \<Longrightarrow> \<forall>y\<in>A. \<not> P x y"
by blast
section "Natural numbers and Integers"
text\<open>Elementary properties of natural numbers and integers\<close>
lemma nat_nonzero_pos:"(a::nat) \<noteq> 0 \<Longrightarrow> 0 < a"
by simp
lemma add_both:"(a::nat) = b \<Longrightarrow> a + c = b + c"
by simp
lemma add_bothl:"a = b \<Longrightarrow> c + a = c + b"
by simp
lemma diff_Suc:"(n::nat) \<le> m \<Longrightarrow> m - n + Suc 0 = Suc m - n"
by arith
lemma le_convert:"\<lbrakk>a = b; a \<le> c\<rbrakk> \<Longrightarrow> b \<le> c"
by simp
lemma ge_convert:"\<lbrakk>a = b; c \<le> a\<rbrakk> \<Longrightarrow> c \<le> b"
by simp
lemma less_convert:"\<lbrakk> a = b; c < b \<rbrakk> \<Longrightarrow> c < a"
by auto
lemma ineq_conv1:"\<lbrakk>a = b; a < c\<rbrakk> \<Longrightarrow> b < c"
by simp
lemma diff_Suc_pos:"0 < a - Suc 0 \<Longrightarrow> 0 < a"
by simp
lemma minus_SucSuc:"a - Suc (Suc 0) = a - Suc 0 - Suc 0"
by simp
lemma Suc_Suc_Tr:"Suc (Suc 0) \<le> n \<Longrightarrow> Suc (n - Suc (Suc 0)) = n - Suc 0"
by arith
lemma Suc_Suc_less:"Suc 0 < a \<Longrightarrow> Suc (a - Suc (Suc 0)) < a"
by arith
lemma diff_zero_eq:"n = (0::nat) \<Longrightarrow> m = m - n"
by simp
lemma Suc_less_le:"x < Suc n \<Longrightarrow> x \<le> n"
by auto
lemma less_le_diff:"x < n \<Longrightarrow> x \<le> n - Suc 0"
by arith
lemma le_pre_le:"x \<le> n - Suc 0 \<Longrightarrow> x \<le> n"
by arith
lemma nat_not_less:"\<not> (m::nat) < n \<Longrightarrow> n \<le> m"
by (rule contrapos_pp, simp+)
lemma less_neq:"n < (m::nat) \<Longrightarrow> n \<noteq> m"
by (simp add:nat_neq_iff[THEN sym, of "n" "m"])
lemma less_le_diff1:"n \<noteq> 0 \<Longrightarrow> ((m::nat) < n) = (m \<le> (n - Suc 0))"
by arith
lemma nat_not_less1:"n \<noteq> 0 \<Longrightarrow> (\<not> (m::nat) < n) = (\<not> m \<le> (n - Suc 0))"
by arith
lemma nat_eq_le:"m = (n::nat) \<Longrightarrow> m \<le> n"
by simp
subsection "Integers"
lemma non_zero_int:" (n::int) \<noteq> 0 \<Longrightarrow> 0 < n \<or> n < 0"
by arith
lemma zgt_0_zge_1:"(0::int) < z \<Longrightarrow> 1 \<le> z"
by arith
lemma not_zle:"(\<not> (n::int) \<le> m) = (m < n)"
by auto
lemma not_zless:"(\<not> (n::int) < m) = (m \<le> n)"
by auto
lemma zle_imp_zless_or_eq:"(n::int) \<le> m \<Longrightarrow> n < m \<or> n = m"
by arith
lemma zminus_zadd_cancel:" - z + (z + w) = (w::int)"
by simp
lemma int_neq_iff:"((w::int) \<noteq> z) = (w < z) \<or> (z < w)"
by auto
lemma zless_imp_zle:"(z::int) < z' \<Longrightarrow> z \<le> z'"
by simp
lemma zdiff:"z - (w::int) = z + (- w)"
by simp
lemma zle_zless_trans:"\<lbrakk> (i::int) \<le> j; j < k\<rbrakk> \<Longrightarrow> i < k"
by arith
lemma zless_zle_trans:"\<lbrakk> (i::int) < j; j \<le> k\<rbrakk> \<Longrightarrow> i < k"
by arith
lemma zless_neq:"(i::int) < j \<Longrightarrow> i \<noteq> j"
by simp
lemma int_mult_mono:"\<lbrakk> i < j; (0::int) < k \<rbrakk> \<Longrightarrow> k * i < k * j"
apply (frule zmult_zless_mono2_lemma [of "i" "j" "nat k"])
apply simp apply simp
done
lemma int_mult_le:"\<lbrakk>i \<le> j; (0::int) \<le> k\<rbrakk> \<Longrightarrow> k * i \<le> k * j"
apply (simp add:order_le_less)
apply (case_tac "i < j")
apply (case_tac "0 < k")
apply simp
apply simp
apply simp
done
lemma int_mult_le1:"\<lbrakk>i \<le> j; (0::int) \<le> k\<rbrakk> \<Longrightarrow> i * k \<le> j * k"
apply (simp add:mult.commute[of _ "k"])
apply (simp add:int_mult_le)
done
lemma zmult_zminus_right:"(w::int) * (- z) = - (w * z)"
apply (insert distrib_left[of "w" "z" "-z"])
apply simp
done
lemma zmult_zle_mono1_neg:"\<lbrakk>(i::int) \<le> j; k \<le> 0\<rbrakk> \<Longrightarrow> j * k \<le> i * k"
apply (subgoal_tac "0 \<le> - k") prefer 2 apply simp
apply (frule int_mult_le [of "i" "j" "- k"], assumption+)
apply (simp add:mult.commute)
done
lemma zmult_zless_mono_neg:"\<lbrakk>(i::int) < j; k < 0\<rbrakk> \<Longrightarrow> j * k < i * k"
apply (subgoal_tac "0 < -k",
frule int_mult_mono[of "i" "j" "-k"], assumption+,
simp add:mult.commute, simp)
done
lemma zmult_neg_neg:"\<lbrakk>i < (0::int); j < 0 \<rbrakk> \<Longrightarrow> 0 < i * j"
apply (frule zmult_zless_mono_neg[of "i" "0" "j"], assumption)
apply simp
done
lemma zmult_pos_pos:"\<lbrakk>(0::int) < i; 0 < j\<rbrakk> \<Longrightarrow> 0 < i * j"
apply (frule int_mult_mono[of "0" "i" "j"], assumption+)
apply (simp add:mult.commute)
done
lemma zmult_pos_neg:"\<lbrakk>(0::int) < i; j < 0\<rbrakk> \<Longrightarrow> i * j < 0"
apply (frule zmult_zless_mono_neg[of "0" "i" "j"], assumption+, simp)
done
lemma zmult_neg_pos:"\<lbrakk>i < (0::int); 0 < j\<rbrakk> \<Longrightarrow> i * j < 0"
apply (frule int_mult_mono[of "i" "0" "j"], assumption+,
simp add:mult.commute)
done
lemma zle:"((z::int) \<le> w) = (\<not> (w < z))"
by auto
lemma times_1_both:"\<lbrakk>(0::int) < z; z * z' = 1\<rbrakk> \<Longrightarrow> z = 1 \<and> z' = 1"
apply (subgoal_tac "0 < z'")
apply (frule zgt_0_zge_1[of "z'"],
subgoal_tac "z' = 1", simp,
subgoal_tac "1 < z' \<or> 1 = z'", thin_tac "1 \<le> z'", thin_tac "0 < z'")
apply (rule contrapos_pp, simp+,
frule int_mult_mono[of "1" "z'" "z"], assumption+, simp, arith)
apply (rule contrapos_pp, simp+, simp add:zle[THEN sym],
frule zless_imp_zle[of "0" "z"], frule int_mult_le[of "z'" "0" "z"],
assumption+, simp)
done
lemma zminus_minus:"i - - (j::int) = i + j"
by simp
lemma zminus_minus_pos:"(n::int) < 0 \<Longrightarrow> 0 < - n"
by simp
lemma zadd_zle_mono:"\<lbrakk>w' \<le> w; z' \<le> (z::int)\<rbrakk> \<Longrightarrow> w' + z' \<le> w + z"
by simp
lemma zmult_zle_mono:"\<lbrakk>i \<le> (j::int); 0 < k\<rbrakk> \<Longrightarrow> k * i \<le> k * j"
apply (case_tac "i = j") apply simp
apply (frule zle_imp_zless_or_eq[of "i" "j"])
apply (thin_tac "i \<le> j") apply simp
done
lemma zmult_zle_mono_r:"\<lbrakk>i \<le> (j::int); 0 < k\<rbrakk> \<Longrightarrow> i * k \<le> j * k"
apply (frule zmult_zle_mono[of "i" "j" "k"], assumption)
apply (simp add:mult.commute)
done
lemma pos_zmult_pos:"\<lbrakk> 0 \<le> (a::int); 0 < (b::int)\<rbrakk> \<Longrightarrow> a \<le> a * b"
apply (case_tac "a = 0") apply simp
apply (frule zle_imp_zless_or_eq[of "0" "a"]) apply (thin_tac "0 \<le> a")
apply simp
done
lemma pos_mult_l_gt:"\<lbrakk>(0::int) < w; i \<le> j; 0 \<le> i\<rbrakk> \<Longrightarrow> i \<le> w * j"
by (metis not_zless pos_zmult_pos order_trans mult.commute)
lemma pos_mult_r_gt:"\<lbrakk>(0::int) < w; i \<le> j; 0 \<le> i\<rbrakk> \<Longrightarrow> i \<le> j * w"
apply (frule pos_mult_l_gt[of "w" "i" "j"], assumption+)
apply (simp add:mult.commute[of "w" "j"])
done
lemma mult_pos_iff:"\<lbrakk>(0::int) < i; 0 \<le> i * j \<rbrakk> \<Longrightarrow> 0 \<le> j"
apply (rule contrapos_pp, simp+)
apply (cut_tac linorder_linear[of "0" "j"]) apply simp
apply (simp add:not_zle)
apply (frule int_mult_mono[of "j" "0" "i"], assumption+) apply simp
done
lemma zmult_eq:"\<lbrakk>(0::int) < w; z = z'\<rbrakk> \<Longrightarrow> w * z = w * z'"
by simp
lemma zmult_eq_r:"\<lbrakk>(0::int) < w; z = z'\<rbrakk> \<Longrightarrow> z * w = z' * w"
by simp
lemma zdiv_eq_l:"\<lbrakk>(0::int) < w; z * w = z' * w \<rbrakk> \<Longrightarrow> z = z'"
by simp
lemma zdiv_eq_r:"\<lbrakk>(0::int) < w; w * z = w * z' \<rbrakk> \<Longrightarrow> z = z'"
by simp
lemma int_nat_minus:"0 < (n::int) \<Longrightarrow> nat (n - 1) = (nat n) - 1"
by arith
lemma int_nat_add:"\<lbrakk>0 < (n::int); 0 < (m::int)\<rbrakk> \<Longrightarrow> (nat (n - 1)) + (nat (m - 1)) + (Suc 0) = nat (n + m - 1)"
by arith
lemma int_equation:"(x::int) = y + z \<Longrightarrow> x - y = z"
by simp
lemma int_pos_mult_monor:"\<lbrakk> 0 < (n::int); 0 \<le> n * m \<rbrakk> \<Longrightarrow> 0 \<le> m"
by (rule mult_pos_iff, assumption+)
lemma int_pos_mult_monol:"\<lbrakk> 0 < (m::int); 0 \<le> n * m \<rbrakk> \<Longrightarrow> 0 \<le> n"
apply (rule int_pos_mult_monor, assumption+)
apply (simp add:mult.commute)
done
lemma zdiv_positive:"\<lbrakk>(0::int) \<le> a; 0 < b\<rbrakk> \<Longrightarrow> 0 \<le> a div b"
apply (frule_tac a = 0 and a' = a and b = b in zdiv_mono1, assumption+)
apply simp
done
lemma zdiv_pos_mono_r:"\<lbrakk> (0::int) < w; w * z \<le> w * z'\<rbrakk> \<Longrightarrow> z \<le> z'"
apply (rule contrapos_pp, simp+)
done (** zmult_div_mono to rename **)
lemma zdiv_pos_mono_l:"\<lbrakk> (0::int) < w; z * w \<le> z' * w\<rbrakk> \<Longrightarrow> z \<le> z'"
apply (simp add:mult.commute)
done
lemma zdiv_pos_pos_l:"\<lbrakk> (0::int) < w; 0 \<le> z * w\<rbrakk> \<Longrightarrow> 0 \<le> z"
by (simp add:mult.commute, frule zdiv_pos_mono_r[of "w" "0" "z"], simp,
assumption)
section "Sets"
(* Preliminary properties of sets are proved here. Some of them are
already proved by L. Paulson and others. *)
subsection "A short notes for proof steps"
subsection "Sets"
lemma inEx:"x \<in> A \<Longrightarrow> \<exists>y\<in>A. y = x"
by simp
lemma inEx_rev:" \<exists>y\<in>A. y = x \<Longrightarrow> x \<in> A"
by blast
lemma nonempty_ex:"A \<noteq> {} \<Longrightarrow> \<exists>x. x \<in> A"
by blast
lemma ex_nonempty:"\<exists>x. x \<in> A \<Longrightarrow> A \<noteq> {}"
by blast
lemma not_eq_outside:"a \<notin> A \<Longrightarrow> \<forall>b\<in>A. b \<noteq> a"
by blast
lemma ex_nonempty_set:"\<exists>a. P a \<Longrightarrow> {x. P x} \<noteq> {}"
by blast
lemma nonempty: "x \<in> A \<Longrightarrow> A \<noteq> {}"
by blast
lemma subset_self:"A \<subseteq> A"
by simp
lemma conditional_subset:"{x\<in>A. P x} \<subseteq> A"
by blast
lemma bsubsetTr:"{x. x \<in> A \<and> P x} \<subseteq> A"
by blast
lemma sets_not_eq:"\<lbrakk>A \<noteq> B; B \<subseteq> A\<rbrakk> \<Longrightarrow> \<exists>a\<in>A. a \<notin> B"
by blast
lemma diff_nonempty:"\<lbrakk>A \<noteq> B; B \<subseteq> A\<rbrakk> \<Longrightarrow> A - B \<noteq> {}"
by blast
lemma sub_which1:"\<lbrakk>A \<subseteq> B \<or> B \<subseteq> A; x \<in> A; x \<notin> B\<rbrakk> \<Longrightarrow> B \<subseteq> A"
by blast
lemma sub_which2:"\<lbrakk>A \<subseteq> B \<or> B \<subseteq> A; x \<notin> A; x \<in> B\<rbrakk> \<Longrightarrow> A \<subseteq> B"
by blast
lemma nonempty_int: "A \<inter> B \<noteq> {} \<Longrightarrow> \<exists>x. x \<in> A \<inter> B "
by blast
lemma no_meet1:"A \<inter> B = {}\<Longrightarrow> \<forall>a \<in> A. a \<notin> B"
by blast
lemma no_meet2:"A \<inter> B = {}\<Longrightarrow> \<forall>a \<in> B. a \<notin> A"
by blast
lemma elem_some:"x \<in> A \<Longrightarrow> \<exists>y\<in>A. x = y"
by blast
lemma singleton_sub:"a \<in> A \<Longrightarrow> {a} \<subseteq> A"
by blast
lemma eq_elem_in: "\<lbrakk> a \<in> A; a = b \<rbrakk> \<Longrightarrow> b \<in> A"
by simp
lemma eq_set_inc: "\<lbrakk> a \<in> A; A = B \<rbrakk> \<Longrightarrow> a \<in> B"
by simp
lemma eq_set_not_inc:"\<lbrakk>a \<notin> A; A = B \<rbrakk> \<Longrightarrow> a \<notin> B"
by simp
lemma int_subsets: "\<lbrakk> A1 \<subseteq> A; B1 \<subseteq> B \<rbrakk> \<Longrightarrow> A1 \<inter> B1 \<subseteq> A \<inter> B"
by blast
lemma inter_mono:"A \<subseteq> B \<Longrightarrow> A \<inter> C \<subseteq> B \<inter> C"
by blast
lemma sub_Un1:"B \<subseteq> B \<union> C"
by blast
lemma sub_Un2:"C \<subseteq> B \<union> C"
by blast
lemma subset_contr:"\<lbrakk> A \<subset> B; B \<subseteq> A \<rbrakk> \<Longrightarrow> False"
by blast
lemma psubset_contr:"\<lbrakk> A \<subset> B; B \<subset> A \<rbrakk> \<Longrightarrow> False"
by blast
lemma eqsets_sub:"A = B \<Longrightarrow> A \<subseteq> B"
by simp
lemma not_subseteq:" \<not> A \<subseteq> B \<Longrightarrow> \<exists>a \<in> A. a \<notin> B"
by blast
lemma in_un1:"\<lbrakk> x \<in> A \<union> B; x \<notin> B \<rbrakk> \<Longrightarrow> x \<in> A"
by blast
lemma proper_subset:"\<lbrakk>A \<subseteq> B; x \<notin> A; x \<in> B\<rbrakk> \<Longrightarrow> A \<noteq> B"
by blast
lemma in_un2:"\<lbrakk> x \<in> A \<union> B; x \<notin> A \<rbrakk> \<Longrightarrow> x \<in> B"
by simp
lemma diff_disj:"x \<notin> A \<Longrightarrow> A - {x} = A"
by auto
lemma in_diff:"\<lbrakk>x \<noteq> a; x \<in> A\<rbrakk> \<Longrightarrow> x \<in> A - {a}"
by simp
lemma in_diff1:"x \<in> A - {a} \<Longrightarrow> x \<noteq> a"
by simp
lemma sub_inserted1:"\<lbrakk>Y \<subseteq> insert a X; \<not> Y \<subseteq> X\<rbrakk> \<Longrightarrow> a \<notin> X \<and> a \<in> Y"
by blast
lemma sub_inserted2:"\<lbrakk>Y \<subseteq> insert a X; \<not> Y \<subseteq> X\<rbrakk> \<Longrightarrow> Y = (Y - {a}) \<union> {a}"
by blast
lemma insert_sub:"\<lbrakk> A \<subseteq> B; a \<in> B\<rbrakk> \<Longrightarrow> (insert a A) \<subseteq> B"
by blast
lemma insert_diff:"A \<subseteq> (insert b B) \<Longrightarrow> A - {b} \<subseteq> B"
by blast
lemma insert_inc1:"A \<subseteq> insert a A"
by blast
lemma insert_inc2:"a \<in> insert a A"
by simp
lemma nonempty_some:"A \<noteq> {} \<Longrightarrow> (SOME x. x \<in> A) \<in> A"
apply (frule nonempty_ex[of "A"])
apply (rule someI2_ex) apply simp+
done
lemma mem_family_sub_Un:"A \<in> C \<Longrightarrow> A \<subseteq> \<Union> C"
by blast
lemma sub_Union:"\<exists>X\<in>C. A \<subseteq> X \<Longrightarrow> A \<subseteq> \<Union> C"
by blast
lemma family_subset_Un_sub:"\<forall>A\<in>C. A \<subseteq> B \<Longrightarrow> \<Union> C \<subseteq> B"
by blast
lemma in_set_with_P:"P x \<Longrightarrow> x \<in> {y. P y}"
by blast
lemma sub_single:"\<lbrakk>A \<noteq> {}; A \<subseteq> {a}\<rbrakk> \<Longrightarrow> A = {a}"
by blast
lemma not_sub_single:"\<lbrakk>A \<noteq> {}; A \<noteq> {a}\<rbrakk> \<Longrightarrow> \<not> A \<subseteq> {a}"
by blast
lemma not_sub:"\<not> A \<subseteq> B \<Longrightarrow> \<exists>a. a\<in>A \<and> a \<notin> B"
by blast
section "Functions"
definition
cmp :: "['b \<Rightarrow> 'c, 'a \<Rightarrow> 'b] \<Rightarrow> ('a \<Rightarrow> 'c)" where
"cmp g f = (\<lambda>x. g (f x))"
definition
idmap :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a)" where
"idmap A = (\<lambda>x\<in>A. x)"
definition
constmap :: "['a set, 'b set] \<Rightarrow> ('a \<Rightarrow>'b)" where
"constmap A B = (\<lambda>x\<in>A. SOME y. y \<in> B)"
definition
invfun :: "['a set, 'b set, 'a \<Rightarrow> 'b] \<Rightarrow> ('b \<Rightarrow> 'a)" where
"invfun A B (f :: 'a \<Rightarrow> 'b) = (\<lambda>y\<in>B.(SOME x. (x \<in> A \<and> f x = y)))"
abbreviation
INVFUN :: "['a \<Rightarrow> 'b, 'b set, 'a set] \<Rightarrow> ('b \<Rightarrow> 'a)" ("(3_\<inverse>\<^bsub>_,_\<^esub>)" [82,82,83]82) where
"f\<inverse>\<^bsub>B,A\<^esub> == invfun A B f"
lemma eq_fun:"\<lbrakk> f \<in> A \<rightarrow> B; f = g \<rbrakk> \<Longrightarrow> g \<in> A \<rightarrow> B"
by simp
lemma eq_fun_eq_val:" f = g \<Longrightarrow> f x = g x"
by simp
lemma eq_elems_eq_val:"x = y \<Longrightarrow> f x = f y"
by simp
lemma cmp_fun:"\<lbrakk>f \<in> A \<rightarrow> B; g \<in> B \<rightarrow> C \<rbrakk> \<Longrightarrow> cmp g f \<in> A \<rightarrow> C"
by (auto simp add:cmp_def)
lemma cmp_fun_image:"\<lbrakk>f \<in> A \<rightarrow> B; g \<in> B \<rightarrow> C \<rbrakk> \<Longrightarrow>
(cmp g f) ` A = g ` (f ` A)"
apply (rule equalityI)
apply (rule subsetI, simp add:image_def)
apply (erule bexE, simp add:cmp_def, blast)
apply (rule subsetI, simp add:image_def[of g])
apply (erule bexE, simp)
apply (simp add:image_def cmp_def)
apply blast
done
lemma cmp_fun_sub_image:"\<lbrakk>f \<in> A \<rightarrow> B; g \<in> B \<rightarrow> C; A1 \<subseteq> A\<rbrakk> \<Longrightarrow>
(cmp g f) ` A1 = g ` (f ` A1)"
apply (rule equalityI)
apply (rule subsetI, simp add:image_def)
apply (erule bexE, simp add:cmp_def, blast)
apply (rule subsetI, simp add:image_def[of g])
apply (erule bexE, simp)
apply (simp add:image_def cmp_def)
apply blast
done
lemma restrict_fun_eq:"\<forall>x\<in>A. f x = g x \<Longrightarrow> (\<lambda>x\<in>A. f x) = (\<lambda>x\<in>A. g x)"
apply (simp add:fun_eq_iff)
done
lemma funcset_mem: "\<lbrakk>f \<in> A \<rightarrow> B; x \<in> A\<rbrakk> \<Longrightarrow> f x \<in> B"
apply (simp add: Pi_def)
done
lemma img_subset:"f \<in> A \<rightarrow> B \<Longrightarrow> f ` A \<subseteq> B"
apply (rule subsetI)
apply (simp add:image_def, erule bexE, simp)
apply (simp add:funcset_mem)
done
lemma funcset_mem1:"\<lbrakk>\<forall>l\<in>A. f l \<in> B; x \<in> A\<rbrakk> \<Longrightarrow> f x \<in> B"
apply simp
done
lemma func_to_img:"f \<in> A \<rightarrow> B \<Longrightarrow> f \<in> A \<rightarrow> f ` A"
by (simp add:Pi_def)
lemma restrict_in_funcset: "\<forall>x\<in> A. f x \<in> B \<Longrightarrow>
(\<lambda>x\<in>A. f x)\<in> A \<rightarrow> B"
apply (simp add:Pi_def restrict_def)
done
lemma funcset_eq:"\<lbrakk> f \<in> extensional A; g \<in> extensional A; \<forall>x\<in>A. f x = g x \<rbrakk> \<Longrightarrow> f = g"
apply (simp add:extensionalityI)
done
lemma eq_funcs:"\<lbrakk>f \<in> A \<rightarrow> B; g \<in> A \<rightarrow> B; f = g; x \<in> A\<rbrakk> \<Longrightarrow> f x = g x"
by simp
lemma restriction_of_domain:"\<lbrakk> f \<in> A \<rightarrow> B; A1 \<subseteq> A \<rbrakk> \<Longrightarrow>
restrict f A1 \<in> A1 \<rightarrow> B"
by blast
lemma restrict_restrict:"\<lbrakk> restrict f A \<in> A \<rightarrow> B; A1 \<subseteq> A \<rbrakk> \<Longrightarrow>
restrict (restrict f A) A1 = restrict f A1"
apply (rule funcset_eq[of _ "A1"])
apply (simp add:restrict_def extensional_def)
apply (simp add:restrict_def extensional_def)
apply (rule ballI) apply simp
apply (simp add:subsetD)
done
lemma restr_restr_eq:"\<lbrakk> restrict f A \<in> A \<rightarrow> B; restrict f A = restrict g A;
A1 \<subseteq> A \<rbrakk> \<Longrightarrow> restrict f A1 = restrict g A1"
apply (subst restrict_restrict[THEN sym, of "f" "A" "B" "A1"], assumption+)
apply (simp add:restrict_restrict[THEN sym, of "g" "A" "B" "A1"])
done
lemma funcTr:"\<lbrakk> f \<in> A \<rightarrow> B; g \<in> A \<rightarrow> B; f = g; a \<in> A\<rbrakk> \<Longrightarrow> f a = g a"
apply simp
done
lemma funcTr1:"\<lbrakk>f = g; a \<in> A\<rbrakk> \<Longrightarrow> f a = g a"
apply simp
done
lemma restrictfun_im:"\<lbrakk> (restrict f A) \<in> A \<rightarrow> B; A1 \<subseteq> A \<rbrakk> \<Longrightarrow>
(restrict f A) ` A1 = f ` A1"
apply (subgoal_tac "\<forall>x\<in>A1. x \<in> A")
apply (simp add:image_def)
apply (rule ballI) apply (simp add:subsetD)
done
lemma mem_in_image:"\<lbrakk> f \<in> A \<rightarrow> B; a \<in> A\<rbrakk> \<Longrightarrow> f a \<in> f ` A "
apply (simp add:image_def)
apply blast
done
lemma mem_in_image1:"\<lbrakk> \<forall>l\<in>A. f l \<in> B; a \<in> A\<rbrakk> \<Longrightarrow> f a \<in> f ` A "
apply simp
done
lemma mem_in_image2:"a \<in> A \<Longrightarrow> f a \<in> f ` A"
apply simp
done
lemma mem_in_image3:"b \<in> f ` A \<Longrightarrow> \<exists>a \<in> A. b = f a"
by (simp add:image_def)
lemma elem_in_image2: "\<lbrakk> f \<in> A \<rightarrow> B; A1 \<subseteq> A; x \<in> A1\<rbrakk> \<Longrightarrow> f x \<in> f` A1"
apply (simp add:image_def)
apply blast
done
lemma funcs_nonempty:"\<lbrakk> A \<noteq> {}; B \<noteq> {} \<rbrakk> \<Longrightarrow> (A \<rightarrow> B) \<noteq> {}"
apply (subgoal_tac "constmap A B \<in> A \<rightarrow> B") apply (simp add:nonempty)
apply (simp add:Pi_def)
apply (rule allI) apply (rule impI)
apply (simp add:constmap_def)
apply (frule nonempty_ex[of "B"])
apply (rule someI2_ex) apply assumption+
done
lemma idmap_funcs: "idmap A \<in> A \<rightarrow> A"
apply (simp add:Pi_def restrict_def idmap_def)
done
lemma l_idmap_comp: "\<lbrakk>f \<in> extensional A; f \<in> A \<rightarrow> B\<rbrakk> \<Longrightarrow>
compose A (idmap B) f = f"
apply (rule funcset_eq[of _ "A"])
apply (simp add:compose_def)
apply assumption
apply (rule ballI)
apply (simp add:funcset_mem[of "f" "A" "B"] compose_def idmap_def)
done
lemma r_idmap_comp:"\<lbrakk>f \<in> extensional A; f \<in> A \<rightarrow> B\<rbrakk> \<Longrightarrow>
compose A f (idmap A) = f"
apply (rule funcset_eq[of _ "A"])
apply (simp add:compose_def)
apply assumption
apply (rule ballI)
apply (simp add:funcset_mem[of "f" "A" "B"] compose_def idmap_def)
done
lemma extend_fun: "\<lbrakk> f \<in> A \<rightarrow> B; B \<subseteq> B1 \<rbrakk> \<Longrightarrow> f \<in> A \<rightarrow> B1"
apply (simp add:Pi_def restrict_def)
apply (rule allI) apply (rule impI)
apply (simp add:subsetD)
done
lemma restrict_fun: "\<lbrakk> f \<in> A \<rightarrow> B; A1 \<subseteq> A \<rbrakk> \<Longrightarrow> restrict f A1 \<in> A1 \<rightarrow> B"
apply (simp add:Pi_def restrict_def)
apply (rule allI) apply (rule impI)
apply (simp add:subsetD)
done
lemma set_of_hom: "\<forall>x \<in> A. f x \<in> B \<Longrightarrow> restrict f A \<in> A \<rightarrow> B"
apply (simp add:Pi_def restrict_def)
done
lemma composition : "\<lbrakk> f \<in> A \<rightarrow> B; g \<in> B \<rightarrow> C\<rbrakk> \<Longrightarrow> (compose A g f) \<in> A \<rightarrow> C"
apply (simp add:Pi_def restrict_def compose_def)
done
lemma comp_assoc:"\<lbrakk>f \<in> A \<rightarrow> B; g \<in> B \<rightarrow> C; h \<in> C \<rightarrow> D \<rbrakk> \<Longrightarrow>
compose A h (compose A g f) = compose A (compose B h g) f"
apply (rule funcset_eq[of _ "A"])
apply (simp add:compose_def)
apply (simp add:compose_def)
apply (rule ballI)
apply (simp add:funcset_mem[of "f" "A" "B"] compose_def)
done
lemma restrictfun_inj: "\<lbrakk> inj_on f A; A1 \<subseteq> A \<rbrakk> \<Longrightarrow> inj_on (restrict f A1) A1"
apply (simp add:inj_on_def)
apply (simp add:subsetD)
done
lemma restrict_inj:"\<lbrakk>inj_on f A; A1 \<subseteq> A\<rbrakk> \<Longrightarrow> inj_on f A1"
apply (simp add:inj_on_def)
apply ((rule ballI)+, rule impI)
apply (frule_tac c = x in subsetD[of "A1" "A"], assumption+,
frule_tac c = y in subsetD[of "A1" "A"], assumption+)
apply simp
done
lemma injective:"\<lbrakk> inj_on f A; x \<in> A; y \<in> A; x \<noteq> y \<rbrakk> \<Longrightarrow> f x \<noteq> f y"
apply (rule contrapos_pp, simp+)
apply (simp add:inj_on_def)
done
lemma injective_iff:"\<lbrakk> inj_on f A; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow>
(x = y) = (f x = f y)"
apply (rule iffI, simp)
apply (rule contrapos_pp, simp+)
apply (frule injective[of "f" "A" "x" "y"], assumption+)
apply simp
done
lemma injfun_elim_image:"\<lbrakk>f \<in> A \<rightarrow> B; inj_on f A; x \<in> A\<rbrakk> \<Longrightarrow>
f ` (A - {x}) = (f ` A) - {f x}"
apply (rule equalityI)
apply (rule subsetI, simp add:image_def, erule bexE)
apply (simp, (erule conjE)+)
apply (rule contrapos_pp, simp+)
apply (erule disjE, simp add:inj_on_def, blast)
apply (frule_tac x = xaa and y = x in injective[of f A ], assumption+,
blast)
apply (rule subsetI, simp add:image_def)
apply (rule contrapos_pp, simp+, erule conjE, erule bexE)
apply (frule_tac x = xaa in bspec)
apply (simp, rule contrapos_pp, simp+)
done
lemma cmp_inj:"\<lbrakk>f \<in> A \<rightarrow> B; g \<in> B \<rightarrow> C; inj_on f A; inj_on g B \<rbrakk> \<Longrightarrow>
inj_on (cmp g f) A"
apply (simp add:inj_on_def [of "cmp g f"])
apply (rule ballI)+
apply (simp add:cmp_def) apply (rule impI)
apply (subgoal_tac "f x = f y")
apply (simp add:inj_on_def [of "f"])
apply (frule_tac x = x in funcset_mem [of "f" "A" "B"], assumption+)
apply (frule_tac x = y in funcset_mem [of "f" "A" "B"], assumption+)
apply (simp add:inj_on_def [of "g"])
done
lemma cmp_assoc:"\<lbrakk>f \<in> A \<rightarrow> B; g \<in> B \<rightarrow> C; h \<in> C \<rightarrow> D; x \<in> A\<rbrakk> \<Longrightarrow>
(cmp h (cmp g f)) x = (cmp (cmp h g) f) x"
apply (simp add:cmp_def)
done
lemma bivar_fun: "\<lbrakk> f \<in> A \<rightarrow> (B \<rightarrow> C); a \<in> A \<rbrakk> \<Longrightarrow> f a \<in> B \<rightarrow> C"
by (simp add:Pi_def)
lemma bivar_fun_mem: "\<lbrakk> f \<in> A \<rightarrow> (B \<rightarrow> C); a \<in> A; b \<in> B \<rbrakk> \<Longrightarrow> f a b \<in> C"
apply (frule funcset_mem[of "f" "A" "B \<rightarrow> C"], assumption+)
apply (rule funcset_mem[of "f a" "B" "C"], assumption+)
done
lemma bivar_func_eq:"\<lbrakk>\<forall>a\<in>A. \<forall>b\<in>B. f a b = g a b \<rbrakk> \<Longrightarrow>
(\<lambda>x\<in>A. \<lambda>y\<in>B. f x y) = (\<lambda>x\<in>A. \<lambda>y\<in>B. g x y)"
apply (subgoal_tac "\<forall>x\<in>A. (\<lambda>y\<in>B. f x y) = (\<lambda>y\<in>B. g x y)")
apply (rule funcset_eq [of _ "A"])
apply (simp add:extensional_def restrict_def)
apply (simp add:extensional_def restrict_def)
apply (rule ballI)
apply simp
apply (rule ballI)
apply (rule funcset_eq [of _ "B"])
apply (simp add:restrict_def extensional_def)
apply (simp add:restrict_def extensional_def)
apply (rule ballI) apply simp
done
lemma set_image: "\<lbrakk> f \<in> A \<rightarrow> B; A1 \<subseteq> A; A2 \<subseteq> A \<rbrakk> \<Longrightarrow>
f`(A1 \<inter> A2) \<subseteq> (f` A1) \<inter> (f` A2)"
apply (simp add: image_def)
apply auto
done
lemma image_sub: "\<lbrakk>f \<in> A \<rightarrow> B; A1 \<subseteq> A \<rbrakk> \<Longrightarrow> (f`A1) \<subseteq> B"
by (auto simp add:image_def)
lemma image_sub0: "f \<in> A \<rightarrow> B \<Longrightarrow> (f`A) \<subseteq> B"
by (simp add:image_sub[of "f" "A" "B" "A"])
lemma image_nonempty:"\<lbrakk>f \<in> A \<rightarrow> B; A1 \<subseteq> A; A1 \<noteq> {} \<rbrakk> \<Longrightarrow> f ` A1 \<noteq> {}"
by (frule nonempty_some[of "A1"],
frule elem_in_image2[of "f" "A" "B" "A1" "SOME x. x \<in> A1"],
assumption+, simp add:nonempty)
lemma im_set_mono: "\<lbrakk>f \<in>A \<rightarrow> B; A1 \<subseteq> A2; A2 \<subseteq> A \<rbrakk> \<Longrightarrow> (f ` A1) \<subseteq> (f ` A2)"
apply (simp add:image_def)
apply auto
done
lemma im_set_un:"\<lbrakk> f\<in>A \<rightarrow> B; A1 \<subseteq> A; A2 \<subseteq> A \<rbrakk> \<Longrightarrow>
f`(A1 \<union> A2) = (f`A1) \<union> (f`A2)"
apply (simp add:image_def)
apply auto
done
lemma im_set_un1:"\<lbrakk>\<forall>l\<in>A. f l \<in> B; A = A1 \<union> A2\<rbrakk> \<Longrightarrow>
f `(A1 \<union> A2) = f `(A1) \<union> f `(A2)"
apply (rule equalityI,
rule subsetI, simp add:image_def, erule bexE)
apply blast
apply (rule subsetI,
simp add:image_def, erule disjE, erule bexE, blast)
apply (erule bexE) apply blast
done
lemma im_set_un2:"A = A1 \<union> A2 \<Longrightarrow> f `A = f `(A1) \<union> f `(A2)"
apply (rule equalityI,
rule subsetI, simp add:image_def, erule bexE, blast)
apply (rule subsetI,
simp add:image_def, erule disjE, erule bexE, blast, erule bexE, blast)
done
definition
invim::"['a \<Rightarrow> 'b, 'a set, 'b set] \<Rightarrow> 'a set" where
"invim f A B = {x. x\<in>A \<and> f x \<in> B}"
lemma invim: "\<lbrakk> f:A \<rightarrow> B; B1 \<subseteq> B \<rbrakk> \<Longrightarrow> invim f A B1 \<subseteq> A"
by (auto simp add:invim_def)
lemma setim_cmpfn: "\<lbrakk> f:A \<rightarrow> B; g:B \<rightarrow> C; A1 \<subseteq> A \<rbrakk> \<Longrightarrow>
(compose A g f)` A1 = g`(f` A1)"
apply (simp add:image_def compose_def)
apply auto
done
definition
surj_to :: "['a \<Rightarrow> 'b, 'a set, 'b set] \<Rightarrow> bool" where
"surj_to f A B \<longleftrightarrow> f`A = B"
lemma surj_to_test:"\<lbrakk> f \<in> A \<rightarrow> B; \<forall>b\<in>B. \<exists>a\<in>A. f a = b \<rbrakk> \<Longrightarrow>
surj_to f A B"
by (auto simp add:surj_to_def image_def)
lemma surj_to_image:"f \<in> A \<rightarrow> B \<Longrightarrow> surj_to f A (f ` A)"
apply (rule surj_to_test[of "f" "A" "f ` A"])
apply (rule func_to_img[of "f" "A" "B"], assumption)
apply (rule ballI, simp add:image_def, erule bexE, simp)
apply blast
done
lemma surj_to_el:"\<lbrakk> f \<in> A \<rightarrow> B; surj_to f A B \<rbrakk> \<Longrightarrow> \<forall>b\<in>B. \<exists>a\<in>A. f a = b"
apply (simp add:surj_to_def image_def)
apply auto
done
lemma surj_to_el1:"\<lbrakk> f \<in> A \<rightarrow> B; surj_to f A B; b\<in>B\<rbrakk> \<Longrightarrow> \<exists>a\<in>A. f a = b"
apply (simp add:surj_to_el)
done
lemma surj_to_el2:"\<lbrakk>surj_to f A B; b \<in> B\<rbrakk> \<Longrightarrow> \<exists>a\<in>A. f a = b"
apply (simp add:surj_to_def image_def)
apply (frule sym, thin_tac "{y. \<exists>x\<in>A. y = f x} = B", simp)
apply (erule bexE, simp, blast)
done
lemma compose_surj: "\<lbrakk>f:A \<rightarrow> B; surj_to f A B; g : B \<rightarrow> C; surj_to g B C \<rbrakk>
\<Longrightarrow> surj_to (compose A g f) A C "
apply (simp add:surj_to_def compose_def image_def)
apply auto
done
lemma cmp_surj: "\<lbrakk>f:A \<rightarrow> B; surj_to f A B; g : B \<rightarrow> C; surj_to g B C \<rbrakk>
\<Longrightarrow> surj_to (cmp g f) A C "
apply (rule surj_to_test, simp add:cmp_fun)
apply (rule ballI, simp add:surj_to_def [of "g"], frule sym,
thin_tac "g ` B = C", simp, simp add:image_def,
simp add:cmp_def)
apply auto
apply (simp add:surj_to_def, frule sym,
thin_tac " f ` A = B", simp add:image_def)
apply auto
done
lemma inj_onTr0:"\<lbrakk> f \<in> A \<rightarrow> B; x \<in> A; y \<in> A; inj_on f A; f x = f y\<rbrakk> \<Longrightarrow> x = y"
apply (simp add:inj_on_def)
done
lemma inj_onTr1:"\<lbrakk>inj_on f A; x \<in> A; y \<in> A; f x = f y\<rbrakk> \<Longrightarrow> x = y"
apply (simp add:inj_on_def)
done
lemma inj_onTr2:"\<lbrakk>inj_on f A; x \<in> A; y \<in> A; f x \<noteq> f y\<rbrakk> \<Longrightarrow> x \<noteq> y"
apply (rule contrapos_pp, simp+)
done (* premis inj_on can be changed to some condition indicating f to be
a function *)
lemma comp_inj: "\<lbrakk> f \<in> A \<rightarrow> B; inj_on f A; g \<in> B \<rightarrow> C; inj_on g B \<rbrakk>
\<Longrightarrow> inj_on (compose A g f) A "
apply (simp add:inj_on_def [of "compose A g f"])
apply (rule ballI)+ apply (rule impI)
apply (rule inj_onTr0 [of "f" "A" "B"], assumption+)
apply (frule funcset_mem [of "f" "A" "B" _], assumption+)
apply (rotate_tac -3)
apply (frule funcset_mem [of "f" "A" "B" _], assumption+)
apply (rule inj_onTr0 [of "g" "B" "C" _], assumption+)
apply (simp add:compose_def)
done
lemma cmp_inj_1: "\<lbrakk> f \<in> A \<rightarrow> B; inj_on f A; g \<in> B \<rightarrow> C; inj_on g B \<rbrakk>
\<Longrightarrow> inj_on (cmp g f) A "
apply (simp add:inj_on_def [of "cmp g f"])
apply (rule ballI)+ apply (rule impI)
apply (simp add:cmp_def)
apply (frule_tac x = x in funcset_mem [of "f" "A" "B"], assumption+)
apply (frule_tac x = y in funcset_mem [of "f" "A" "B"], assumption+)
apply (frule_tac x = "f x" and y = "f y" in inj_onTr1 [of "g" "B"],
assumption+)
apply (rule_tac x = x and y = y in inj_onTr1 [of "f" "A"], assumption+)
done
lemma cmp_inj_2: "\<lbrakk>\<forall>l\<in>A. f l \<in> B; inj_on f A; \<forall>k\<in>B. g k \<in> C; inj_on g B \<rbrakk>
\<Longrightarrow> inj_on (cmp g f) A "
apply (simp add:inj_on_def [of "cmp g f"])
apply (rule ballI)+ apply (rule impI)
apply (simp add:cmp_def)
apply (frule_tac x = x in funcset_mem1 [of "A" "f" "B"], assumption+)
apply (frule_tac x = y in funcset_mem1 [of "A" "f" "B"], assumption+)
apply (frule_tac x = "f x" and y = "f y" in inj_onTr1 [of "g" "B"],
assumption+)
apply (rule_tac x = x and y = y in inj_onTr1 [of "f" "A"], assumption+)
done
lemma invfun_mem:"\<lbrakk> f \<in> A \<rightarrow> B; inj_on f A; surj_to f A B; b \<in> B \<rbrakk>
\<Longrightarrow> (invfun A B f) b \<in> A"
apply (simp add:invfun_def)
apply (simp add:surj_to_def image_def) apply (frule sym)
apply (thin_tac "{y. \<exists>x\<in>A. y = f x} = B") apply simp
apply (thin_tac "B = {y. \<exists>x\<in>A. y = f x}") apply auto
apply (rule someI2_ex)
apply blast apply simp
done
lemma inv_func:"\<lbrakk> f \<in> A \<rightarrow> B; inj_on f A; surj_to f A B\<rbrakk>
\<Longrightarrow> (invfun A B f) \<in> B \<rightarrow> A"
apply (simp add:Pi_def)
apply (rule allI) apply (rule impI)
apply (rule invfun_mem) apply (rule funcsetI)
apply simp+
done
lemma invfun_r:"\<lbrakk> f\<in>A \<rightarrow> B; inj_on f A; surj_to f A B; b \<in> B \<rbrakk>
\<Longrightarrow> f ((invfun A B f) b) = b"
apply (simp add:invfun_def)
apply (rule someI2_ex)
apply (simp add:surj_to_def image_def)
apply auto
done
lemma invfun_l:"\<lbrakk>f \<in> A \<rightarrow> B; inj_on f A; surj_to f A B; a \<in> A\<rbrakk>
\<Longrightarrow> (invfun A B f) (f a) = a"
apply (simp add:invfun_def Pi_def restrict_def)
apply (rule someI2_ex) apply auto
apply (simp add:inj_on_def)
done
lemma invfun_inj:"\<lbrakk>f \<in> A \<rightarrow> B; inj_on f A; surj_to f A B\<rbrakk>
\<Longrightarrow> inj_on (invfun A B f) B"
apply (simp add:inj_on_def [of "invfun A B f" "B"] )
apply auto
apply (frule_tac b = y in invfun_r [of "f" "A" "B"], assumption+)
apply (frule_tac b = x in invfun_r [of "f" "A" "B"], assumption+)
apply simp
done
lemma invfun_surj:"\<lbrakk>f \<in> A \<rightarrow> B; inj_on f A; surj_to f A B\<rbrakk>
\<Longrightarrow> surj_to (invfun A B f) B A "
apply (simp add:surj_to_def [of "invfun A B f" "B" "A"] image_def)
apply (rule equalityI)
apply (rule subsetI) apply (simp add:CollectI)
apply auto
apply (simp add:invfun_mem)
apply (frule funcset_mem [of "f" "A" "B"], assumption+)
apply (frule_tac t = x in invfun_l [of "f" "A" "B", THEN sym], assumption+)
apply auto
done
definition
bij_to :: "['a \<Rightarrow> 'b, 'a set, 'b set] \<Rightarrow> bool" where
"bij_to f A B \<longleftrightarrow> surj_to f A B \<and> inj_on f A"
lemma idmap_bij:"bij_to (idmap A) A A"
apply (simp add:bij_to_def)
apply (rule conjI)
apply (simp add:surj_to_def, simp add:image_def, simp add:idmap_def)
apply (simp add:inj_on_def, simp add:idmap_def)
done
lemma bij_invfun:"\<lbrakk>f \<in> A \<rightarrow> B; bij_to f A B\<rbrakk> \<Longrightarrow>
bij_to (invfun A B f) B A"
apply (simp add:bij_to_def)
apply (simp add:invfun_inj invfun_surj)
done
lemma l_inv_invfun:"\<lbrakk> f \<in> A \<rightarrow> B; inj_on f A; surj_to f A B\<rbrakk>
\<Longrightarrow> compose A (invfun A B f) f = idmap A"
apply (rule ext)
apply (simp add:compose_def idmap_def)
apply (rule impI)
apply (simp add:invfun_l)
done
lemma invfun_mem1:"\<lbrakk>f \<in> A \<rightarrow> B; bij_to f A B; b \<in> B\<rbrakk> \<Longrightarrow>
(invfun A B f) b \<in> A"
apply (simp add:bij_to_def, erule conjE)
apply (simp add:invfun_mem)
done
lemma invfun_r1:"\<lbrakk> f\<in>A \<rightarrow> B; bij_to f A B; b \<in> B \<rbrakk>
\<Longrightarrow> f ((invfun A B f) b) = b"
apply (simp add:bij_to_def, erule conjE)
apply (rule invfun_r, assumption+)
done
lemma invfun_l1:"\<lbrakk>f \<in> A \<rightarrow> B; bij_to f A B; a \<in> A\<rbrakk>
\<Longrightarrow> (invfun A B f) (f a) = a"
apply (simp add:bij_to_def, erule conjE)
apply (rule invfun_l, assumption+)
done
lemma compos_invfun_r:"\<lbrakk>f \<in> A \<rightarrow> B; bij_to f A B; g \<in> A \<rightarrow> C; h \<in> B \<rightarrow> C;
g \<in> extensional A; compose B g (invfun A B f) = h\<rbrakk> \<Longrightarrow>
g = compose A h f"
apply (rule funcset_eq[of g A "compose A h f"], assumption)
apply (simp add:compose_def extensional_def)
apply (rule ballI)
apply (frule sym, thin_tac "compose B g (invfun A B f) = h", simp)
apply (simp add:compose_def Pi_def invfun_l1)
done
lemma compos_invfun_l:"\<lbrakk>f \<in> A \<rightarrow> B; bij_to f A B; g \<in> C \<rightarrow> B; h \<in> C \<rightarrow> A;
compose C (invfun A B f) g = h; g \<in> extensional C \<rbrakk> \<Longrightarrow>
g = compose C f h"
apply (rule funcset_eq[of g C "compose C f h"], assumption)
apply (simp add:compose_def extensional_def)
apply (rule ballI)
apply (frule sym, thin_tac "compose C (invfun A B f) g = h",
simp add:compose_def)
apply (frule_tac x = x in funcset_mem[of g C B], assumption)
apply (simp add:invfun_r1)
done
lemma invfun_set:"\<lbrakk>f \<in> A \<rightarrow> B; bij_to f A B; C \<subseteq> B\<rbrakk> \<Longrightarrow>
f ` ((invfun A B f)` C) = C"
apply (rule equalityI)
apply (rule subsetI)
apply (simp add:image_def, erule exE,
erule conjE, erule bexE, simp,
frule_tac c = xb in subsetD[of "C" "B"], assumption+)
apply (simp add:bij_to_def, erule conjE,
simp add:invfun_r)
apply (rule subsetI, simp add:image_def)
apply (frule_tac c = x in subsetD[of "C" "B"], assumption+,
simp add:bij_to_def, erule conjE,
frule_tac b = x in invfun_r[of "f" "A" "B"], assumption+)
apply (frule sym, thin_tac "f (invfun A B f x) = x")
apply blast
done
lemma compos_bij:"\<lbrakk>f \<in> A \<rightarrow> B; bij_to f A B; g \<in> B \<rightarrow> C; bij_to g B C\<rbrakk> \<Longrightarrow>
bij_to (compose A g f) A C"
apply (simp add:bij_to_def, (erule conjE)+)
apply (simp add:comp_inj[of "f" "A" "B" "g" "C"])
apply (simp add:compose_surj)
done
section "Nsets"
(* NSet is the set of natural numbers, and "Nset n" is the set of
natural numbers from 0 through n *)
definition
nset :: "[nat, nat] \<Rightarrow> (nat) set" where
"nset i j = {k. i \<le> k \<and> k \<le> j}"
definition
slide :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"slide i j == i + j"
definition
sliden :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"sliden i j == j - i"
definition
jointfun :: "[nat, nat \<Rightarrow> 'a, nat, nat \<Rightarrow> 'a] \<Rightarrow> (nat \<Rightarrow> 'a)" where
"(jointfun n f m g) = (\<lambda>i. if i \<le> n then f i else g ((sliden (Suc n)) i))"
definition
skip :: "nat \<Rightarrow> (nat \<Rightarrow> nat)" where
"skip i = (\<lambda>x. (if i = 0 then Suc x else
(if x \<in> {j. j \<le> (i - Suc 0)} then x else Suc x)))"
lemma nat_pos:"0 \<le> (l::nat)"
apply simp
done
lemma Suc_pos:"Suc k \<le> r \<Longrightarrow> 0 < r"
apply simp
done
lemma nat_pos2:"(k::nat) < r \<Longrightarrow> 0 < r"
apply simp
done
lemma eq_le_not:"\<lbrakk>(a::nat) \<le> b; \<not> a < b \<rbrakk> \<Longrightarrow> a = b"
apply auto
done
lemma im_of_constmap:"(constmap {0} {a}) ` {0} = {a}"
apply (simp add:constmap_def)
done
lemma noteq_le_less:"\<lbrakk> m \<le> (n::nat); m \<noteq> n \<rbrakk> \<Longrightarrow> m < n"
apply auto
done
lemma nat_not_le_less:"(\<not> (n::nat) \<le> m) = (m < n)"
by (simp add: not_le)
lemma self_le:"(n::nat) \<le> n"
apply simp
done
lemma n_less_Suc:"(n::nat) < Suc n"
apply simp
done
lemma less_diff_pos:"i < (n::nat) \<Longrightarrow> 0 < n - i"
apply auto
done
lemma less_diff_Suc:"i < (n::nat) \<Longrightarrow> n - (Suc i) = (n - i) - (Suc 0)"
apply auto
done
lemma less_pre_n:"0 < n \<Longrightarrow> n - Suc 0 < n"
apply simp
done
lemma Nset_inc_0:"(0::nat) \<in> {i. i \<le> n}"
apply simp
done
lemma Nset_1:"{i. i \<le> Suc 0} = {0, Suc 0}"
apply auto
done
lemma Nset_1_1:"(k \<le> Suc 0) = (k = 0 \<or> k = Suc 0)"
apply (rule iffI)
apply (case_tac "k = 0", simp+)
apply (erule disjE, simp+)
done
lemma Nset_2:"{i, j} = {j, i}"
apply auto
done
lemma Nset_nonempty:"{i. i \<le> (n::nat)} \<noteq> {}"
apply (subgoal_tac "0 \<in> {i. i \<le> n}")
apply (rule nonempty[of 0], assumption)
apply simp
done
lemma Nset_le:"x \<in> {i. i \<le> n} \<Longrightarrow> x \<le> n"
apply simp
done
lemma n_in_Nsetn:"(n::nat) \<in> {i. i \<le> n}"
apply simp
done
lemma Nset_pre:"\<lbrakk> (x::nat) \<in> {i. i \<le> (Suc n)}; x \<noteq> Suc n \<rbrakk> \<Longrightarrow> x \<in> {i. i \<le> n}"
apply simp
done
lemma Nset_pre1:"{i. i \<le> (Suc n)} - {Suc n} = {i. i \<le> n}"
apply (rule equalityI)
apply (rule subsetI, simp)+
done
lemma le_Suc_mem_Nsetn:"x \<le> Suc n \<Longrightarrow> x - Suc 0 \<in> {i. i \<le> n}"
apply (frule diff_le_mono[of x "Suc n" "Suc 0"],
thin_tac "x \<le> Suc n", simp)
done
lemma le_Suc_diff_le:"x \<le> Suc n \<Longrightarrow> x - Suc 0 \<le> n"
apply (frule diff_le_mono[of x "Suc n" "Suc 0"],
thin_tac "x \<le> Suc n", simp)
done
lemma Nset_not_pre:"\<lbrakk> x \<notin> {i. i \<le> n}; x \<in> {i. i \<le> (Suc n)}\<rbrakk> \<Longrightarrow> x = Suc n"
by simp
lemma mem_of_Nset:"x \<le> (n::nat) \<Longrightarrow> x \<in> {i. i \<le> n}"
apply simp
done
lemma less_mem_of_Nset:"x < (n::nat) \<Longrightarrow> x \<in> {i. i \<le> n}"
apply (frule less_imp_le [of "x" "n"])
apply simp
done
lemma Nset_nset:"{i. i \<le> (Suc (n + m))} = {i. i \<le> n} \<union>
nset (Suc n) (Suc (n + m))"
apply (rule equalityI)
apply (rule subsetI)
apply (simp add:nset_def)
apply (auto simp add: nset_def)
done
lemma Nset_nset_1:"\<lbrakk>0 < n; i < n\<rbrakk> \<Longrightarrow> {j. j \<le> n} = {j. j \<le> i} \<union>
nset (Suc i) n"
apply auto
apply (simp add:nset_def)
apply (simp add:nset_def)
done
lemma Nset_img0:"\<lbrakk>f \<in> {j. j \<le> Suc n} \<rightarrow> B; (f (Suc n)) \<in> f ` {j. j \<le> n}\<rbrakk> \<Longrightarrow>
f ` {j. j \<le> Suc n} = f ` {j. j \<le> n}"
-apply (rule equalityI)
- apply (rule subsetI)
- apply (simp add:image_def)
- apply (erule exE, erule conjE)
- apply (erule exE, simp, erule conjE)
- apply (case_tac "xb = Suc n", simp, blast)
- apply (frule_tac m = xb and n = "Suc n" in noteq_le_less, assumption,
- thin_tac "xb \<le> Suc n",
- frule_tac x = xb and n = "Suc n" in less_le_diff,
- thin_tac "xb < Suc n", simp, blast)
- apply (rule subsetI, simp add:image_def, (erule exE)+, (erule conjE)+)
- apply (simp,
- frule_tac x = xb and y = n and z = "Suc n" in le_less_trans,
- simp,
- frule_tac x = xb and y = "Suc n" in less_imp_le,
- blast)
-done
+ by (auto simp add: le_Suc_eq)
lemma Nset_img:"f \<in> {j. j \<le> Suc n} \<rightarrow> B \<Longrightarrow>
insert (f (Suc n)) (f ` {j. j \<le> n}) = f ` {j. j \<le> Suc n}"
-apply (rule equalityI)
- apply (rule subsetI)
- apply (simp add:image_def,
- erule disjE, blast,
- erule exE, erule conjE,
- frule_tac x = xa and y = n and z = "Suc n" in le_less_trans,
- simp+,
- frule_tac x = xa and y = "Suc n" in less_imp_le, blast)
- apply (rule subsetI,
- simp add:image_def,
- erule exE, simp, erule conjE,
- case_tac "xa = Suc n", simp)
- apply (frule_tac m = xa and n = "Suc n" in noteq_le_less, assumption,
- frule_tac x = xa and n = "Suc n" in less_le_diff,
- thin_tac "xa \<le> Suc n", thin_tac "xa < Suc n", simp)
- apply blast
-done
+ by (auto elim: le_SucE)
primrec nasc_seq :: "[nat set, nat, nat] \<Rightarrow> nat"
where
dec_seq_0: "nasc_seq A a 0 = a"
| dec_seq_Suc: "nasc_seq A a (Suc n) =
(SOME b. ((b \<in> A) \<and> (nasc_seq A a n) < b))"
lemma nasc_seq_mem:"\<lbrakk>(a::nat) \<in> A; \<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. x \<le> m))\<rbrakk> \<Longrightarrow>
(nasc_seq A a n) \<in> A"
apply (induct n)
apply (simp_all add: not_le)
apply (subgoal_tac "\<exists>x\<in>A. (nasc_seq A a n) < x") prefer 2 apply blast
apply (thin_tac "\<forall>m. m \<in> A \<longrightarrow> (\<exists>x\<in>A. m < x)",
rule someI2_ex, blast, simp)
done
lemma nasc_seqn:"\<lbrakk>(a::nat) \<in> A; \<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. x \<le> m))\<rbrakk> \<Longrightarrow>
(nasc_seq A a n) < (nasc_seq A a (Suc n))"
apply (simp,
frule nasc_seq_mem [of "a" "A" "n"], simp)
apply (simp add: not_le,
subgoal_tac "\<exists>x\<in>A. (nasc_seq A a n) < x") prefer 2 apply simp
apply (thin_tac "\<forall>m. m \<in> A \<longrightarrow> (\<exists>x\<in>A. m < x)",
rule someI2_ex, blast, simp)
done
lemma nasc_seqn1:"\<lbrakk>(a::nat) \<in> A; \<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. x \<le> m))\<rbrakk> \<Longrightarrow>
Suc (nasc_seq A a n) \<le> (nasc_seq A a (Suc n))"
apply (frule nasc_seqn [of "a" "A" "n"], assumption+)
apply simp
done
lemma ubs_ex_n_maxTr:"\<lbrakk>(a::nat) \<in> A; \<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. x \<le> m))\<rbrakk>
\<Longrightarrow> (a + n) \<le> (nasc_seq A a n)"
apply (induct_tac n)
apply simp
apply (frule_tac n = n in nasc_seqn1[of "a" "A"], assumption+,
subgoal_tac "Suc (a + n) \<le> Suc (nasc_seq A a n)",
frule_tac i = "Suc (a + n)" and j = "Suc (nasc_seq A a n)" and
k = "nasc_seq A a (Suc n)" in le_trans, assumption+,
simp, thin_tac "Suc (nasc_seq A a n) \<le> nasc_seq A a (Suc n)",
subst Suc_le_mono, assumption+)
done
lemma ubs_ex_n_max:"\<lbrakk>A \<noteq> {}; A \<subseteq> {i. i \<le> (n::nat)}\<rbrakk> \<Longrightarrow>
\<exists>!m. m\<in>A \<and> (\<forall>x\<in>A. x \<le> m)"
apply (frule nonempty_ex[of "A"])
apply (thin_tac "A \<noteq> {}")
apply (erule exE)
apply (rename_tac a)
apply (rule ex_ex1I)
prefer 2
apply (erule conjE)+
apply (frule_tac x = y in bspec, assumption+,
thin_tac "\<forall>x\<in>A. x \<le> m",
frule_tac x = m in bspec, assumption+,
thin_tac "\<forall>x\<in>A. x \<le> y", simp)
apply (rule contrapos_pp, simp+)
apply (frule_tac a = a and A = A and n = "n + 1" in ubs_ex_n_maxTr, simp)
apply (frule_tac a = a in nasc_seq_mem[of _ "A" "n + 1"], simp)
apply (frule_tac c = "nasc_seq A a (n + 1)" in subsetD[of "A" "{i. i \<le> n}"],
assumption+, simp)
done
definition
n_max :: "nat set \<Rightarrow> nat" where
"n_max A = (THE m. m \<in> A \<and> (\<forall>x\<in>A. x \<le> m))"
lemma n_max:"\<lbrakk>A \<subseteq> {i. i \<le> (n::nat)}; A \<noteq> {}\<rbrakk> \<Longrightarrow>
(n_max A) \<in> A \<and> (\<forall>x\<in>A. x \<le> (n_max A))"
apply (simp add:n_max_def)
apply (frule ubs_ex_n_max[of "A" "n"], assumption)
apply (rule theI')
apply assumption
done
lemma n_max_eq_sets:"\<lbrakk>A = B; A \<noteq> {}; \<exists>n. A \<subseteq> {j. j \<le> n}\<rbrakk> \<Longrightarrow>
n_max A = n_max B"
by simp
(* n_max has no meaning unless conditions A \<noteq> {}; \<exists>n. A \<subseteq> {j. j \<le> n} *)
lemma skip_mem:"l \<in> {i. i \<le> n} \<Longrightarrow> (skip i l) \<in> {i. i \<le> (Suc n)}"
apply (case_tac "i = 0")
apply (simp add:skip_def)
apply (simp)+
apply (simp add:skip_def)
done
lemma skip_fun:"(skip i) \<in> {i. i \<le> n} \<rightarrow> {i. i \<le> (Suc n)}"
apply (rule Pi_I)
apply (rule skip_mem, assumption)
done
lemma skip_im_Tr0:"x \<in> {i. i \<le> n} \<Longrightarrow> skip 0 x = Suc x"
apply (simp add:skip_def)
done
lemma skip_im_Tr0_1:"0 < y \<Longrightarrow> skip 0 (y - Suc 0) = y"
apply (simp add:skip_def)
done
lemma skip_im_Tr1:"\<lbrakk> i \<in> {i. i \<le> (Suc n)}; 0 < i; x \<le> i - Suc 0 \<rbrakk> \<Longrightarrow>
skip i x = x"
by (simp add:skip_def)
lemma skip_im_Tr1_1:"\<lbrakk> i \<in> {i. i \<le> (Suc n)}; 0 < i; x < i\<rbrakk> \<Longrightarrow>
skip i x = x"
apply (frule less_le_diff[of x i])
apply (simp add:skip_def)
done
lemma skip_im_Tr1_2:"\<lbrakk> i \<le> (Suc n); x < i\<rbrakk> \<Longrightarrow> skip i x = x"
apply (rule skip_im_Tr1_1[of i n x], simp+)
done
lemma skip_im_Tr2:"\<lbrakk> 0 < i; i \<in> {i. i \<le> (Suc n)}; i \<le> x\<rbrakk> \<Longrightarrow>
skip i x = Suc x"
by (simp add:skip_def)
lemma skip_im_Tr2_1:"\<lbrakk>i \<in> {i. i \<le> (Suc n)}; i \<le> x\<rbrakk> \<Longrightarrow>
skip i x = Suc x"
apply (case_tac "i = 0")
apply (simp add:skip_def)
apply (simp, rule skip_im_Tr2, assumption+, simp+)
done
lemma skip_im_Tr3:"x \<in> {i. i \<le> n} \<Longrightarrow> skip (Suc n) x = x"
apply (simp add:skip_def)
done
lemma skip_im_Tr4:"\<lbrakk>x \<le> Suc n; 0 < x\<rbrakk> \<Longrightarrow> x - Suc 0 \<le> n"
apply (simp add:Suc_le_mono [of "x - Suc 0" "n", THEN sym])
done
lemma skip_fun_im:"i \<in> {j. j \<le> (Suc n)} \<Longrightarrow>
(skip i) ` {j. j \<le> n} = ({j. j \<le> (Suc n)} - {i})"
apply (rule equalityI)
apply (rule subsetI)
apply (case_tac "i = 0", simp)
apply (simp add:image_def, erule exE, erule conjE)
apply (cut_tac x = xa in skip_im_Tr0[of _ n], simp, simp)
apply (simp add:image_def, erule exE, erule conjE, simp)
apply (case_tac "xa < i")
apply (frule_tac x = xa in skip_im_Tr1_2[of i n], simp+)
apply (cut_tac m1 = xa and n1 = i in nat_not_le_less[THEN sym], simp)
apply (cut_tac x = xa and n = n in skip_im_Tr2_1[of i], simp+)
apply (rule subsetI, simp, erule conjE)
apply (cut_tac x = x and y = i in less_linear, simp)
apply (erule disjE)
apply (simp add:image_def)
apply (frule_tac x = x in skip_im_Tr1_2[of i n], assumption,
frule_tac x = x and y = i and z = "Suc n" in less_le_trans,
assumption+,
frule_tac m = x and n = "Suc n" in Suc_leI,
simp only:Suc_le_mono,
frule sym, thin_tac "skip i x = x", blast)
apply (cut_tac x = "x - Suc 0" in skip_im_Tr2_1[of i n],
simp, simp add:less_le_diff)
apply (cut_tac x = i and n = x in less_le_diff, assumption,
simp add:image_def)
apply (frule_tac m = x and n = "Suc n" and l = "Suc 0" in diff_le_mono,
simp)
apply (frule sym, thin_tac "skip i (x - Suc 0) = x", blast)
done
lemma skip_fun_im1:"\<lbrakk>i \<in> {j. j \<le> (Suc n)}; x \<in> {j. j \<le> n}\<rbrakk> \<Longrightarrow>
(skip i) x \<in> ({j. j \<le> (Suc n)} - {i})"
by (subst skip_fun_im[THEN sym], assumption,
simp add:image_def, blast)
lemma skip_id:"l < i \<Longrightarrow> skip i l = l"
apply (simp add:skip_def )
done
lemma Suc_neq:"\<lbrakk>0 < i; i - Suc 0 < l\<rbrakk> \<Longrightarrow> i \<noteq> Suc l"
by (rule contrapos_pp, simp+)
lemma skip_il_neq_i:"skip i l \<noteq> i"
apply (auto simp add:skip_def)
done
lemma skip_inj:"\<lbrakk>i \<in> {k. k \<le> n}; j \<in> {k. k \<le> n}; i \<noteq> j\<rbrakk> \<Longrightarrow>
skip k i \<noteq> skip k j"
apply (simp add:skip_def)
done
lemma le_imp_add_int:" i \<le> (j::nat) \<Longrightarrow> \<exists>k. j = i + k"
apply (case_tac "i = j")
apply simp
apply (frule le_imp_less_or_eq) apply (thin_tac "i \<le> j")
apply simp
apply (insert less_imp_add_positive [of "i" "j"])
apply simp
apply blast
done
lemma jointfun_hom0:"\<lbrakk> f \<in> {j. j \<le> n} \<rightarrow> A; g \<in> {k. k \<le> m} \<rightarrow> B \<rbrakk> \<Longrightarrow>
(jointfun n f m g) \<in> {l. l \<le> (Suc (n + m))} \<rightarrow> (A \<union> B)"
by (simp add:jointfun_def sliden_def Pi_def)
lemma jointfun_mem:"\<lbrakk>\<forall>j \<le> (n::nat). f j \<in> A; \<forall>j \<le> m. g j \<in> B;
l \<le> (Suc (n + m))\<rbrakk> \<Longrightarrow> (jointfun n f m g) l \<in> (A \<union> B)"
apply (rule funcset_mem[of "jointfun n f m g" "{j. j \<le> Suc (n + m)}" "A \<union> B"
l])
apply (rule jointfun_hom0)
apply simp+
done
lemma jointfun_inj:"\<lbrakk>f \<in> {j. j \<le> n} \<rightarrow> B; inj_on f {j. j \<le> n};
b \<notin> f ` {j. j \<le> n}\<rbrakk> \<Longrightarrow>
inj_on (jointfun n f 0 (\<lambda>k\<in>{0::nat}. b)) {j. j \<le> Suc n}"
apply (simp add:inj_on_def, (rule allI, rule impI)+, rule impI)
apply (case_tac "x = Suc n", simp)
apply (case_tac "y = Suc n", simp)
apply (frule_tac m = y and n = "Suc n" in noteq_le_less, assumption)
apply (
frule_tac x = y and n = "Suc n" in less_le_diff,
thin_tac "y < Suc n", thin_tac "y \<le> Suc n",
simp add:jointfun_def sliden_def)
apply (case_tac "y = Suc n", simp,
frule_tac m = x and n = "Suc n" in noteq_le_less, assumption,
frule_tac x = x and n = "Suc n" in less_le_diff,
thin_tac "x < Suc n", thin_tac "x \<le> Suc n",
simp add:jointfun_def sliden_def)
apply (rotate_tac -3, frule sym, thin_tac " f x = b", simp)
apply (frule_tac m = x and n = "Suc n" in noteq_le_less, assumption,
frule_tac x = x and n = "Suc n" in less_le_diff,
thin_tac "x < Suc n", thin_tac "x \<le> Suc n", simp,
frule_tac m = y and n = "Suc n" in noteq_le_less, assumption,
frule_tac x = y and n = "Suc n" in less_le_diff,
thin_tac "y < Suc n", thin_tac "y \<le> Suc n", simp,
simp add:jointfun_def)
done
lemma slide_hom:"i \<le> j \<Longrightarrow> (slide i) \<in> {l. l \<le> (j - i)} \<rightarrow> nset i j"
apply (simp add:Pi_def restrict_def)
apply (rule allI) apply (rule impI)
apply (simp add:slide_def)
apply (simp add:nset_def)
done
lemma slide_mem:"\<lbrakk> i \<le> j; l \<in> {k. k \<le> (j - i)}\<rbrakk> \<Longrightarrow> slide i l \<in> nset i j"
apply (frule slide_hom)
apply (rule funcset_mem, assumption+)
done
lemma slide_iM:"(slide i) ` {l. 0 \<le> l} = {k. i \<le> k}"
apply (simp add:image_def slide_def)
apply (rule equalityI)
apply (rule subsetI)
apply simp
apply auto
apply (rule le_imp_add_int)
apply assumption
done
lemma jointfun_hom:"\<lbrakk> f \<in> {i. i \<le> n} \<rightarrow> A; g \<in> {j. j \<le> m} \<rightarrow> B \<rbrakk> \<Longrightarrow>
(jointfun n f m g) \<in> {j. j \<le> (Suc (n + m))} \<rightarrow> A \<union> B"
by (simp add:sliden_def Pi_def jointfun_def)
lemma im_jointfunTr1:"(jointfun n f m g) ` {i. i \<le> n} = f ` {i. i \<le> n}"
apply auto
apply (simp add:jointfun_def)
apply (simp add:jointfun_def)
done
lemma im_jointfunTr2:"(jointfun n f m g) ` (nset (Suc n) (Suc (n + m))) =
g ` ({j. j \<le> m})"
apply auto
apply (simp add:nset_def) apply auto
apply (frule_tac m = xa and n = "Suc (n + m)" and l = "Suc n" in diff_le_mono)
apply simp
apply (simp add:jointfun_def sliden_def)
apply (simp add:image_def)
apply (cut_tac le_add1[of "n" "m"],
simp only:Suc_le_mono[THEN sym, of "n" "n+m"])
apply (frule_tac l = xa in slide_mem[of "Suc n" "Suc (n + m)"])
apply simp
apply (subst jointfun_def)
apply (subgoal_tac "\<forall>i\<in>nset (Suc n) (Suc (n+m)). \<not> (i \<le> n) ")
apply simp
apply (thin_tac "\<forall>i\<in>nset (Suc n) (Suc (n + m)). \<not> i \<le> n")
apply (subgoal_tac "g xa = g (sliden (Suc n) (slide (Suc n) xa))")
apply blast
apply (simp add:slide_def sliden_def)
apply (auto simp add: nset_def)
done
lemma im_jointfun:"\<lbrakk>f \<in> {j. j \<le> n} \<rightarrow> A; g \<in> {j. j \<le> m} \<rightarrow> B\<rbrakk> \<Longrightarrow>
(jointfun n f m g) `({j. j \<le> (Suc (n + m))}) =
f `{j. j \<le> n} \<union> g `{j. j \<le> m}"
apply (cut_tac im_set_un1 [of "{j. j \<le> (Suc (n + m))}" "jointfun n f m g"
"A \<union> B" "{i. i \<le> n}" "nset (Suc n) (Suc (n + m))"])
apply (simp add:Nset_nset[THEN sym, of n m],
simp add:im_jointfunTr1[of n f m g],
simp add:im_jointfunTr2[of n f m g])
apply (rule ballI)
apply (simp add:jointfun_def,
case_tac "l \<le> n", simp add:Pi_def,
simp add:sliden_def,
simp add:nat_not_le_less,
frule_tac m = n and n = l in Suc_leI,
frule_tac m = l and n = "Suc (n + m)" and l = "Suc n" in diff_le_mono,
thin_tac "l \<le> Suc (n + m)", simp,
simp add:Pi_def)
apply (simp add:Nset_nset[of n m])
done
lemma im_jointfun1:"(jointfun n f m g) `({j. j \<le> (Suc (n + m))}) =
f `{j. j \<le> n} \<union> g ` {j. j \<le> m}"
apply (cut_tac Nset_nset[of "n" "m"])
apply (subst im_set_un2[of "{j. j \<le> (Suc (n + m))}" "{j. j \<le> n}"
"nset (Suc n) (Suc (n + m))" "jointfun n f m g"], assumption)
apply (simp add:im_jointfunTr1 im_jointfunTr2)
done
lemma jointfun_surj:"\<lbrakk>f \<in> {j. j \<le> n} \<rightarrow> A; surj_to f {j. j \<le> (n::nat)} A;
g \<in> {j. j \<le> (m::nat)} \<rightarrow> B; surj_to g {j. j \<le> m} B\<rbrakk> \<Longrightarrow>
surj_to (jointfun n f m g) {j. j \<le> Suc (n + m)} (A \<union> B)"
apply (simp add:surj_to_def [of "jointfun n f m g"])
apply (simp add:im_jointfun)
apply (simp add:surj_to_def)
done
lemma Nset_un:"{j. j \<le> (Suc n)} = {j. j \<le> n} \<union> {Suc n}"
apply (rule equalityI)
apply (rule subsetI)
apply simp
apply auto
done
lemma Nsetn_sub: "{j. j \<le> n} \<subseteq> {j. j \<le> (Suc n)}"
apply (rule subsetI)
apply simp
done
lemma Nset_pre_sub:"(0::nat) < k \<Longrightarrow> {j. j \<le> (k - Suc 0)} \<subseteq> {j. j \<le> k}"
apply (rule subsetI)
apply simp
done
lemma Nset_pre_un:"(0::nat) < k \<Longrightarrow> {j. j \<le> k} = {j. j \<le> (k - Suc 0)} \<union> {k}"
apply (insert Nset_un [of "k - Suc 0"])
apply simp
done
lemma Nsetn_sub_mem:" l \<in> {j. j \<le> n} \<Longrightarrow> l \<in> {j. j \<le> (Suc n)}"
apply simp
done
lemma Nsetn_sub_mem1:"\<forall>j. j \<in> {j. j \<le> n} \<longrightarrow> j \<in> {j. j \<le> (Suc n)}"
by (simp add:Nsetn_sub_mem)
lemma Nset_Suc:"{j. j \<le> (Suc n)} = insert (Suc n) {j. j \<le> n}"
apply (rule equalityI)
apply (rule subsetI)
apply simp
apply auto
done
lemma nsetnm_sub_mem:"\<forall>j. j \<in>nset n (n + m) \<longrightarrow> j \<in> nset n (Suc (n + m))"
by (rule allI, simp add:nset_def)
lemma Nset_0:"{j. j \<le> (0::nat)} = {0}"
by simp
lemma Nset_Suc0:"{i. i \<le> (Suc 0)} = {0, Suc 0}"
apply (rule equalityI)
apply (rule subsetI, simp)
apply (case_tac "x = 0", simp)
apply simp+
done
lemma Nset_Suc_Suc:"Suc (Suc 0) \<le> n \<Longrightarrow>
{j. j \<le> (n - Suc (Suc 0))} = {j. j \<le> n} - {n - Suc 0, n}"
apply (insert Nset_un [of "n - (Suc 0)"])
apply (insert Nset_un [of "n - Suc (Suc 0)"])
apply (subgoal_tac "{j. j \<le> (Suc (n - Suc (Suc 0)))} = {j. j \<le> (n - Suc 0)}")
apply (simp,
thin_tac "{j. j \<le> n} =
insert n (insert (Suc (n - Suc (Suc 0))) {j. j \<le> n - Suc (Suc 0)})",
thin_tac " {j. j \<le> n - Suc 0} =
insert (Suc (n - Suc (Suc 0))) {j. j \<le> n - Suc (Suc 0)}",
thin_tac "{j. j \<le> Suc (n - Suc (Suc 0))} =
insert (Suc (n - Suc (Suc 0))) {j. j \<le> n - Suc (Suc 0)}")
apply (simp add:Suc_Suc_Tr)
apply (auto )
done
lemma func_pre:"f \<in> {j. j \<le> (Suc n)} \<rightarrow> A \<Longrightarrow> f \<in> {j. j \<le> n} \<rightarrow> A"
by (simp add:Pi_def)
lemma image_Nset_Suc:"f ` ({j. j \<le> (Suc n)}) =
insert (f (Suc n)) (f ` {j. j \<le> n})"
apply (cut_tac Nset_un[of "n"])
apply (frule im_set_un2[of "{j. j \<le> (Suc n)}" "{j. j \<le> n}" "{Suc n}" "f"])
apply (simp add:Un_commute)
done
definition
Nleast :: "nat set \<Rightarrow> nat" where
"Nleast A = (THE a. (a \<in> A \<and> (\<forall>x\<in>A. a \<le> x)))"
definition
Nlb :: "[nat set, nat] \<Rightarrow> bool" where
"Nlb A n \<longleftrightarrow> (\<forall>a\<in>A. n \<le> a)"
primrec ndec_seq :: "[nat set, nat, nat] \<Rightarrow> nat" where
ndec_seq_0 :"ndec_seq A a 0 = a"
| ndec_seq_Suc:"ndec_seq A a (Suc n) =
(SOME b. ((b \<in> A) \<and> b < (ndec_seq A a n)))"
lemma ndec_seq_mem:"\<lbrakk>a \<in> (A::nat set); \<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. m \<le> x))\<rbrakk> \<Longrightarrow>
(ndec_seq A a n) \<in> A"
apply (induct_tac n)
apply simp apply simp
apply (simp add: not_less [symmetric])
apply (subgoal_tac "\<exists>x\<in>A. x < (ndec_seq A a n)") prefer 2 apply blast
apply (thin_tac "\<forall>m. m \<in> A \<longrightarrow> (\<exists>x\<in>A. x < m)")
apply (rule someI2_ex) apply blast
apply simp
done
lemma ndec_seqn:"\<lbrakk>a \<in> (A::nat set);\<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. m \<le> x))\<rbrakk> \<Longrightarrow>
(ndec_seq A a (Suc n)) < (ndec_seq A a n)"
apply (frule ndec_seq_mem [of "a" "A" "n"], assumption+)
apply simp
apply (simp add: not_less [symmetric])
apply (subgoal_tac "\<exists>x\<in>A. x < (ndec_seq A a n)") prefer 2 apply simp
apply (thin_tac "\<forall>m. m \<in> A \<longrightarrow> (\<exists>x\<in>A. x < m)")
apply (rule someI2_ex) apply blast
apply simp
done
lemma ndec_seqn1:"\<lbrakk>a \<in> (A::nat set); \<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. m \<le> x))\<rbrakk> \<Longrightarrow>
(ndec_seq A a (Suc n)) \<le> (ndec_seq A a n) - 1"
apply (frule ndec_seqn [of "a" "A" "n"], assumption+,
thin_tac "\<not> (\<exists>m. m \<in> A \<and> (\<forall>x\<in>A. m \<le> x))")
apply (simp del:ndec_seq_Suc)
done
lemma ex_NleastTr:"\<lbrakk>a \<in> (A::nat set); \<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. m \<le> x))\<rbrakk> \<Longrightarrow>
(ndec_seq A a n) \<le> (a - n)"
apply (induct_tac n)
apply simp
apply (frule_tac n = n in ndec_seqn1[of "a" "A"], assumption+)
apply (subgoal_tac "ndec_seq A a n - 1 \<le> (a - n) - 1") prefer 2
apply arith
apply arith
done
lemma nat_le:"((a::nat) - (a + 1)) \<le> 0"
apply arith
done
lemma ex_Nleast:"(A::nat set) \<noteq> {} \<Longrightarrow> \<exists>!m. m\<in>A \<and> (\<forall>x\<in>A. m \<le> x)"
apply (frule nonempty_ex[of "A"], thin_tac "A \<noteq> {}",
erule exE, rename_tac a)
apply (case_tac "0 \<in> A")
apply (rule ex_ex1I, subgoal_tac "\<forall>x\<in>A. 0 \<le> a", blast,
rule ballI, simp)
apply ((erule conjE)+,
subgoal_tac "m \<le> 0", thin_tac "\<forall>x\<in>A. m \<le> x",
subgoal_tac "y \<le> 0", thin_tac "\<forall>x\<in>A. y \<le> x",
simp, blast, blast)
apply (rule ex_ex1I)
prefer 2 apply (erule conjE)+
apply (subgoal_tac "m \<le> y", thin_tac "\<forall>x\<in>A. m \<le> x",
subgoal_tac "y \<le> m", thin_tac "\<forall>x\<in>A. y \<le> x",
simp, blast, blast)
apply (rule contrapos_pp, simp,
frule_tac a = a and A = A and n = "a + 1" in ex_NleastTr, assumption+)
apply (subgoal_tac "(a - (a + 1)) \<le> 0")
prefer 2 apply (rule nat_le)
apply (frule_tac i = "ndec_seq A a (a + 1)" and j = "a - (a + 1)" and k = 0 in le_trans, assumption+,
frule_tac a = a and n = "a + 1" in ndec_seq_mem [of _ "A"],
assumption+)
apply (thin_tac "\<not> (\<exists>m. m \<in> A \<and> (\<forall>x\<in>A. m \<le> x))",
thin_tac "ndec_seq A a (a + 1) \<le> a - (a + 1)",
thin_tac "a - (a + 1) \<le> 0")
apply simp
done
lemma Nleast:"(A::nat set) \<noteq> {} \<Longrightarrow> Nleast A \<in> A \<and> (\<forall>x\<in>A. (Nleast A) \<le> x)"
apply (frule ex_Nleast [of "A"])
apply (simp add:Nleast_def)
apply (rule theI')
apply simp
done
subsection "Lemmas for existence of reduced chain."
(* Later some of these lemmas should be removed. *)
lemma jointgd_tool1:" 0 < i \<Longrightarrow> 0 \<le> i - Suc 0"
by arith
lemma jointgd_tool2:" 0 < i \<Longrightarrow> i = Suc (i - Suc 0)"
by arith
lemma jointgd_tool3:"\<lbrakk>0 < i; i \<le> m\<rbrakk> \<Longrightarrow> i - Suc 0 \<le> (m - Suc 0)"
by arith
lemma jointgd_tool4:"n < i \<Longrightarrow> i - n = Suc( i - Suc n)"
by arith
lemma pos_prec_less:"0 < i \<Longrightarrow> i - Suc 0 < i"
by arith
lemma Un_less_Un:"\<lbrakk>f \<in> {j. j \<le> (Suc n)} \<rightarrow> (X::'a set set);
A \<subseteq> \<Union>(f ` {j. j \<le> (Suc n)});
i \<in> {j. j \<le> (Suc n)}; j \<in> {l. l \<le> (Suc n)}; i \<noteq> j \<and> f i \<subseteq> f j\<rbrakk>
\<Longrightarrow> A \<subseteq> \<Union>(compose {j. j \<le> n} f (skip i) ` {j. j \<le> n})"
apply (simp add:compose_def)
apply (rule subsetI, simp)
apply (frule_tac c = x and A = A and B = "\<Union>x\<in>{j. j \<le> Suc n}. f x" in
subsetD, assumption+, simp)
apply (erule exE, (erule conjE)+)
apply (case_tac "xa = i", simp,
frule_tac c = x in subsetD[of "f i" "f j"], assumption+)
apply (cut_tac less_linear[of i j], simp, erule disjE,
frule less_le_diff[of i j],
cut_tac skip_im_Tr2_1[of i n "j - Suc 0"],
simp,
frule eq_elems_eq_val[THEN sym, of "skip i (j - Suc 0)" j f],
cut_tac a = x in eq_set_inc[of _ "f j" "f (skip i (j - Suc 0))"],
assumption+,
frule le_Suc_diff_le[of j n], blast, simp, assumption, simp)
apply (frule skip_im_Tr1_2[of i n j], assumption,
frule eq_elems_eq_val[THEN sym, of "skip i j" j f])
apply (cut_tac a = x in eq_set_inc[of _ "f j" "f (skip i j)"],
assumption+)
apply (frule_tac x = j and y = i and z = "Suc n" in less_le_trans,
assumption+,
frule Suc_less_le[of j n], blast)
apply (cut_tac x = xa and y = i in less_linear, simp,
erule disjE,
frule_tac x = xa in skip_im_Tr1_2[of i n], assumption)
apply (frule_tac x1 = "skip i xa" and y1 = xa and f1 = f in
eq_elems_eq_val[THEN sym],
frule_tac a = x and A = "f xa" and B = "f (skip i xa)" in eq_set_inc,
assumption,
frule_tac x = xa and y = i and z = "Suc n" in less_le_trans,
assumption+,
frule_tac x = xa and n = n in Suc_less_le, blast)
apply (frule_tac x = i and n = xa in less_le_diff,
cut_tac x = "xa - Suc 0" and n = n in skip_im_Tr2_1 [of i],
simp, assumption,
simp,
frule_tac x1 = "skip i (xa - Suc 0)" and y1 = xa and f1 = f in
eq_elems_eq_val[THEN sym],
frule_tac a = x and A = "f xa" and B = "f (skip i (xa - Suc 0))" in
eq_set_inc, assumption,
frule_tac x = xa and n = n in le_Suc_diff_le)
apply blast
done
section "Lower bounded set of integers"
(* In this section. I prove that a lower bounded set of integers
has the minimal element *)
definition "Zset = {x. \<exists>(n::int). x = n}"
definition
Zleast :: "int set \<Rightarrow> int" where
"Zleast A = (THE a. (a \<in> A \<and> (\<forall>x\<in>A. a \<le> x)))"
definition
LB :: "[int set, int] \<Rightarrow> bool" where
"LB A n = (\<forall>a\<in>A. n \<le> a)"
lemma linorder_linear1:"(m::int) < n \<or> n \<le> m"
apply (subgoal_tac "m < n \<or> n = m \<or> n < m")
apply (case_tac "m < n") apply simp apply simp
apply (subgoal_tac "m < n \<or> m = n \<or> n < m")
apply blast
apply (simp add:less_linear)
done
primrec dec_seq :: "[int set, int, nat] \<Rightarrow> int"
where
dec_seq_0: "dec_seq A a 0 = a"
| dec_seq_Suc: "dec_seq A a (Suc n) = (SOME b. ((b \<in> A) \<and> b < (dec_seq A a n)))"
lemma dec_seq_mem:"\<lbrakk>a \<in> A; A \<subseteq> Zset;\<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. m \<le> x))\<rbrakk> \<Longrightarrow>
(dec_seq A a n) \<in> A"
apply (induct_tac n)
apply simp apply simp apply (simp add:not_zle)
apply (subgoal_tac "\<exists>x\<in>A. x < (dec_seq A a n)") prefer 2 apply blast
apply (thin_tac "\<forall>m. m \<in> A \<longrightarrow> (\<exists>x\<in>A. x < m)")
apply (rule someI2_ex) apply blast
apply simp
done
lemma dec_seqn:"\<lbrakk>a \<in> A; A \<subseteq> Zset;\<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. m \<le> x))\<rbrakk> \<Longrightarrow>
(dec_seq A a (Suc n)) < (dec_seq A a n)"
apply simp
apply (frule dec_seq_mem [of "a" "A" "n"], assumption+)
apply simp
apply (simp add:not_zle)
apply (subgoal_tac "\<exists>x\<in>A. x < (dec_seq A a n)") prefer 2 apply simp
apply (thin_tac "\<forall>m. m \<in> A \<longrightarrow> (\<exists>x\<in>A. x < m)")
apply (rule someI2_ex) apply blast
apply simp
done
lemma dec_seqn1:"\<lbrakk>a \<in> A; A \<subseteq> Zset;\<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. m \<le> x))\<rbrakk> \<Longrightarrow>
(dec_seq A a (Suc n)) \<le> (dec_seq A a n) - 1"
apply (frule dec_seqn [of "a" "A" "n"], assumption+)
apply simp
done
lemma lbs_ex_ZleastTr:"\<lbrakk>a \<in> A; A \<subseteq> Zset;\<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. m \<le> x))\<rbrakk> \<Longrightarrow>
(dec_seq A a n) \<le> (a - int(n))"
apply (induct_tac n)
apply simp
apply (frule_tac n = n in dec_seqn1[of "a" "A"], assumption+)
apply (subgoal_tac "dec_seq A a n - 1 \<le> a - (int n) - 1") prefer 2
apply simp apply (thin_tac "dec_seq A a n \<le> a - int n")
apply (frule_tac x = "dec_seq A a (Suc n)" and y = "dec_seq A a n - 1" and
z = "a - int n - 1" in order_trans, assumption+)
apply (thin_tac "\<not> (\<exists>m. m \<in> A \<and> (\<forall>x\<in>A. m \<le> x))")
apply (thin_tac "dec_seq A a (Suc n) \<le> dec_seq A a n - 1")
apply (thin_tac "dec_seq A a n - 1 \<le> a - int n - 1")
apply (subgoal_tac "a - int n - 1 = a - int (Suc n)") apply simp
apply (thin_tac "dec_seq A a (Suc n) \<le> a - int n - 1")
apply simp
done
lemma big_int_less:"a - int(nat(abs(a) + abs(N) + 1)) < N"
apply (simp add:zabs_def)
done
lemma lbs_ex_Zleast:"\<lbrakk>A \<noteq> {}; A \<subseteq> Zset; LB A n\<rbrakk> \<Longrightarrow> \<exists>!m. m\<in>A \<and> (\<forall>x\<in>A. m \<le> x)"
apply (frule nonempty_ex[of "A"])
apply (thin_tac "A \<noteq> {}")
apply (erule exE)
apply (rename_tac a)
apply (rule ex_ex1I)
prefer 2
apply (thin_tac "LB A n") apply (erule conjE)+
apply (subgoal_tac "m \<le> y") prefer 2 apply simp
apply (subgoal_tac "y \<le> m") prefer 2 apply simp
apply (thin_tac "\<forall>x\<in>A. m \<le> x") apply (thin_tac "\<forall>x\<in>A. y \<le> x")
apply simp
apply (rule contrapos_pp) apply simp
apply (frule_tac a = a and A = A and n = "nat(abs(a) + abs(n) + 1)" in lbs_ex_ZleastTr, assumption+)
apply (subgoal_tac "a - int(nat(abs(a) + abs(n) + 1)) < n")
prefer 2 apply (rule big_int_less)
apply (frule_tac x = "dec_seq A a (nat (\<bar>a\<bar> + \<bar>n\<bar> + 1))" and y = "a - int (nat (\<bar>a\<bar> + \<bar>n\<bar> + 1))" and z = n in order_le_less_trans, assumption+)
apply (frule_tac a = a and n = "nat (\<bar>a\<bar> + \<bar>n\<bar> + 1)" in dec_seq_mem [of _ "A"], assumption+)
apply (thin_tac "\<not> (\<exists>m. m \<in> A \<and> (\<forall>x\<in>A. m \<le> x))")
apply (thin_tac "dec_seq A a (nat (\<bar>a\<bar> + \<bar>n\<bar> + 1))
\<le> a - int (nat (\<bar>a\<bar> + \<bar>n\<bar> + 1))")
apply (thin_tac "a - int (nat (\<bar>a\<bar> + \<bar>n\<bar> + 1)) < n")
apply (simp add:LB_def)
apply (subgoal_tac "n \<le> dec_seq A a (nat (\<bar>a\<bar> + \<bar>n\<bar> + 1))")
apply (thin_tac "\<forall>a\<in>A. n \<le> a") apply (simp add:not_zle)
apply blast
done
lemma Zleast:"\<lbrakk>A \<noteq> {}; A \<subseteq> Zset; LB A n\<rbrakk> \<Longrightarrow> Zleast A \<in> A \<and>
(\<forall>x\<in>A. (Zleast A) \<le> x)"
apply (frule lbs_ex_Zleast [of "A" "n"], assumption+)
apply (simp add:Zleast_def)
apply (rule theI')
apply simp
done
lemma less_convert1:"\<lbrakk> a = c; a < b \<rbrakk> \<Longrightarrow> c < b"
apply auto
done
lemma less_convert2:"\<lbrakk>a = b; b < c\<rbrakk> \<Longrightarrow> a < c"
apply auto
done
section \<open>Augmented integer: integer and \<open>\<infinity>-\<infinity>\<close>\<close>
definition
zag :: "(int * int) set" where
"zag = {(x,y) | x y. x * y = (0::int) \<and> (y = -1 \<or> y = 0 \<or> y = 1)}"
definition
zag_pl::"[(int * int), (int * int)] \<Rightarrow> (int * int)" where
"zag_pl x y == if (snd x + snd y) = 2 then (0, 1)
else if (snd x + snd y) = 1 then (0, 1)
else if (snd x + snd y) = 0 then (fst x + fst y, 0)
else if (snd x + snd y) = -1 then (0, -1)
else if (snd x + snd y) = -2 then (0, -1) else undefined"
definition
zag_t :: "[(int * int), (int * int)] \<Rightarrow> (int * int)" where
"zag_t x y = (if (snd x)*(snd y) = 0 then
(if 0 < (fst x)*(snd y) + (snd x)*(fst y) then (0,1)
else (if (fst x)*(snd y) + (snd x)*(fst y) = 0
then ((fst x)*(fst y), 0) else (0, -1)))
else (if 0 < (snd x)*(snd y) then (0, 1) else (0, -1)))"
definition "Ainteg = zag"
typedef ant = Ainteg
morphisms Rep_Ainteg Abs_Ainteg
unfolding Ainteg_def
proof
show "(1, 0) \<in> zag" unfolding zag_def by auto
qed
definition
ant :: "int \<Rightarrow> ant" where
"ant m = Abs_Ainteg( (m, 0))"
definition
tna :: "ant \<Rightarrow> int" where
"tna z = (if Rep_Ainteg(z) \<noteq> (0,1) \<and> Rep_Ainteg(z) \<noteq> (0,-1) then
fst (Rep_Ainteg(z)) else undefined)"
instantiation ant :: "{zero, one, plus, uminus, minus, times, ord}"
begin
definition
Zero_ant_def : "0 == ant 0"
definition
One_ant_def : "1 == ant 1"
definition
add_ant_def:
"z + w ==
Abs_Ainteg (zag_pl (Rep_Ainteg z) (Rep_Ainteg w))"
definition
minus_ant_def : "- z ==
Abs_Ainteg((- (fst (Rep_Ainteg z)), - (snd (Rep_Ainteg z))))"
definition
diff_ant_def: "z - (w::ant) == z + (-w)"
definition
mult_ant_def:
"z * w ==
Abs_Ainteg (zag_t (Rep_Ainteg z) (Rep_Ainteg w))"
definition
le_ant_def:
"(z::ant) \<le> w == if (snd (Rep_Ainteg w)) = 1 then True
else (if (snd (Rep_Ainteg w)) = 0 then (if (snd (Rep_Ainteg z)) = 1
then False else (if (snd (Rep_Ainteg z)) = 0 then
(fst (Rep_Ainteg z)) \<le> (fst (Rep_Ainteg w)) else True))
else (if snd (Rep_Ainteg z) = -1 then True else False))"
definition
less_ant_def: "((z::ant) < (w::ant)) == (z \<le> w \<and> z \<noteq> w)"
instance ..
end
definition
inf_ant :: ant ("\<infinity>") where
"\<infinity> = Abs_Ainteg((0,1))"
definition
an :: "nat \<Rightarrow> ant" where
"an m = ant (int m)"
definition
na :: "ant \<Rightarrow> nat" where
"na x = (if (x < 0) then 0 else
if x \<noteq> \<infinity> then (nat (tna x)) else undefined)"
definition
UBset :: "ant \<Rightarrow> ant set" where
"UBset z = {(x::ant). x \<le> z}"
definition
LBset :: "ant \<Rightarrow> ant set" where
"LBset z = {(x::ant). z \<le> x}"
lemma ant_z_in_Ainteg:"(z::int, 0) \<in> Ainteg"
apply (simp add:Ainteg_def zag_def)
done
lemma ant_inf_in_Ainteg:"((0::int), 1) \<in> Ainteg"
apply (simp add:Ainteg_def zag_def)
done
lemma ant_minf_in_Ainteg:"((0::int), -1) \<in> Ainteg"
apply (simp add:Ainteg_def zag_def)
done
lemma ant_0_in_Ainteg:"((0::int), 0) \<in> Ainteg"
apply (simp add:Ainteg_def zag_def)
done
lemma an_0[simp]:"an 0 = 0"
by (simp add:an_def Zero_ant_def)
lemma an_1[simp]:"an 1 = 1"
by (simp add:an_def One_ant_def)
lemma mem_ant:"(a::ant) = -\<infinity> \<or> (\<exists>(z::int). a = ant z) \<or> a = \<infinity>"
apply (case_tac "a = -\<infinity> \<or> a = \<infinity>")
apply blast
apply (simp, simp add:ant_def,
cut_tac Rep_Ainteg[of "a"],
simp add:Ainteg_def zag_def,
erule conjE, simp add:inf_ant_def,
simp add:minus_ant_def,
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply auto
apply (cut_tac Rep_Ainteg[of "a"],
subgoal_tac "Abs_Ainteg (Rep_Ainteg a) = Abs_Ainteg ((0,-1))",
thin_tac "Rep_Ainteg a = (0, -1)",
simp add:Rep_Ainteg_inverse, simp)
apply (cut_tac Rep_Ainteg[of "a"],
subgoal_tac "Abs_Ainteg (Rep_Ainteg a) = Abs_Ainteg ((0,0))",
thin_tac "Rep_Ainteg a = (0, 0)",
simp add:Rep_Ainteg_inverse, blast, simp)
apply (cut_tac Rep_Ainteg[of "a"],
subgoal_tac "Abs_Ainteg (Rep_Ainteg a) = Abs_Ainteg ((0,1))",
thin_tac "Rep_Ainteg a = (0, 1)",
simp add:Rep_Ainteg_inverse, simp)
apply (cut_tac Rep_Ainteg[of "a"],
subgoal_tac "Abs_Ainteg (Rep_Ainteg a) = Abs_Ainteg ((x,0))",
thin_tac "Rep_Ainteg a = (x, 0)",
simp add:Rep_Ainteg_inverse, blast, simp)
done
lemma minf:"-\<infinity> = Abs_Ainteg((0,-1))"
apply (simp add:inf_ant_def minus_ant_def,
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
done
lemma z_neq_inf[simp]:"(ant z) \<noteq> \<infinity> "
apply (rule contrapos_pp, simp+)
apply (simp add:ant_def inf_ant_def)
apply (subgoal_tac "Rep_Ainteg (Abs_Ainteg (z,0)) =
Rep_Ainteg (Abs_Ainteg (0,1))",
thin_tac "Abs_Ainteg (z, 0) = Abs_Ainteg (0, 1)",
cut_tac ant_z_in_Ainteg[of "z"],
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply simp
done
lemma z_neq_minf[simp]:"(ant z) \<noteq> -\<infinity>"
apply (rule contrapos_pp, simp+)
apply (subgoal_tac "ant (-z) = \<infinity>")
apply (cut_tac z_neq_inf[of "- z"], simp)
apply (simp add:ant_def inf_ant_def minus_ant_def)
apply (cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply (subgoal_tac "- Abs_Ainteg (z, 0) = - Abs_Ainteg (0, -1)",
thin_tac "Abs_Ainteg (z, 0) = Abs_Ainteg (0, -1)",
simp add:minus_ant_def,
cut_tac ant_z_in_Ainteg[of "z"],
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply simp
done
lemma minf_neq_inf[simp]:"-\<infinity> \<noteq> \<infinity>"
apply (cut_tac ant_inf_in_Ainteg,
simp add:inf_ant_def minus_ant_def Abs_Ainteg_inverse)
apply (rule contrapos_pp, simp+,
subgoal_tac "Rep_Ainteg (Abs_Ainteg (0,-1)) =
Rep_Ainteg (Abs_Ainteg (0,1))",
thin_tac "Abs_Ainteg (0, -1) = Abs_Ainteg (0, 1)",
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply simp
done
lemma a_ipi[simp]:"\<infinity> + \<infinity> = \<infinity>"
apply (simp add:add_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
simp add:zag_pl_def)
done
lemma a_zpi[simp]:"(ant z) + \<infinity> = \<infinity>"
apply (simp add:add_ant_def inf_ant_def ant_def,
cut_tac ant_z_in_Ainteg[of "z"],
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
simp add:zag_pl_def)
done
lemma a_ipz[simp]:" \<infinity> + (ant z) = \<infinity>"
apply (simp add:add_ant_def inf_ant_def ant_def,
cut_tac ant_z_in_Ainteg[of "z"],
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
simp add:zag_pl_def)
done
lemma a_zpz:"(ant m) + (ant n) = ant (m + n)"
apply (simp add:add_ant_def inf_ant_def ant_def,
cut_tac ant_z_in_Ainteg[of "m"],
cut_tac ant_z_in_Ainteg[of "n"],
simp add:Abs_Ainteg_inverse,
simp add:zag_pl_def)
done
lemma a_mpi[simp]:"-\<infinity> + \<infinity> = 0"
apply (simp add:add_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:minus_ant_def,
simp add:Abs_Ainteg_inverse,
simp add:Zero_ant_def ant_def zag_pl_def)
done
lemma a_ipm[simp]:"\<infinity> + (-\<infinity>) = 0"
apply (simp add:add_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:minus_ant_def,
simp add:Abs_Ainteg_inverse,
simp add:Zero_ant_def ant_def zag_pl_def)
done
lemma a_mpm[simp]:"-\<infinity> + (-\<infinity>) = -\<infinity>"
apply (simp add:add_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:minus_ant_def,
simp add:Abs_Ainteg_inverse,
simp add:Zero_ant_def ant_def zag_pl_def)
done
lemma a_mpz[simp]:"-\<infinity> + (ant m) = -\<infinity>"
apply (simp add:add_ant_def minus_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
simp add:ant_def,
cut_tac ant_z_in_Ainteg[of "m"],
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_pl_def)
done
lemma a_zpm[simp]:"(ant m) + (-\<infinity>) = -\<infinity>"
apply (simp add:add_ant_def minus_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
simp add:ant_def,
cut_tac ant_z_in_Ainteg[of "m"],
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_pl_def)
done
lemma a_mdi[simp]:"-\<infinity> - \<infinity> = - \<infinity>"
apply (simp add:diff_ant_def minus_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply (simp add:add_ant_def,
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse, simp add:zag_pl_def)
done
lemma a_zdz:"(ant m) - (ant n) = ant (m - n)"
apply (simp add:diff_ant_def minus_ant_def ant_def,
cut_tac ant_z_in_Ainteg[of "n"],
simp add:Abs_Ainteg_inverse)
apply (simp add:add_ant_def,
cut_tac ant_z_in_Ainteg[of "m"],
cut_tac ant_z_in_Ainteg[of "-n"],
simp add:Abs_Ainteg_inverse zag_pl_def)
done
lemma a_i_i[simp]:"\<infinity> * \<infinity> = \<infinity>"
apply (simp add:mult_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_t_def)
done
lemma a_0_i[simp]:"0 * \<infinity> = 0"
by (simp add:mult_ant_def inf_ant_def Zero_ant_def, simp add:ant_def,
cut_tac ant_inf_in_Ainteg, cut_tac ant_0_in_Ainteg,
simp add:Abs_Ainteg_inverse, simp add:zag_t_def)
lemma a_i_0[simp]:"\<infinity> * 0 = 0"
by (simp add:mult_ant_def inf_ant_def Zero_ant_def, simp add:ant_def,
cut_tac ant_inf_in_Ainteg, cut_tac ant_0_in_Ainteg,
simp add:Abs_Ainteg_inverse, simp add:zag_t_def)
lemma a_0_m[simp]:"0 * (-\<infinity>) = 0"
by (simp add:mult_ant_def inf_ant_def Zero_ant_def, simp add:ant_def,
cut_tac ant_inf_in_Ainteg, cut_tac ant_0_in_Ainteg,
simp add:Abs_Ainteg_inverse, simp add:zag_t_def)
lemma a_m_0[simp]:"(-\<infinity>) * 0 = 0"
by (simp add:mult_ant_def inf_ant_def Zero_ant_def, simp add:ant_def,
cut_tac ant_inf_in_Ainteg, cut_tac ant_0_in_Ainteg,
simp add:Abs_Ainteg_inverse, simp add:zag_t_def)
lemma a_m_i[simp]:"(-\<infinity>) * \<infinity> = -\<infinity>"
by (simp add:mult_ant_def inf_ant_def minus_ant_def,
cut_tac ant_inf_in_Ainteg, cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse, simp add:zag_t_def)
lemma a_i_m[simp]:"\<infinity> * (-\<infinity>) = - \<infinity>"
by (simp add:mult_ant_def inf_ant_def minus_ant_def,
cut_tac ant_inf_in_Ainteg, cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse, simp add:zag_t_def)
lemma a_pos_i[simp]:"0 < m \<Longrightarrow> (ant m) * \<infinity> = \<infinity>"
apply (simp add:mult_ant_def inf_ant_def ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_z_in_Ainteg[of "m"],
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_t_def)
done
lemma a_i_pos[simp]:"0 < m \<Longrightarrow> \<infinity> * (ant m) = \<infinity>"
apply (simp add:mult_ant_def inf_ant_def ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_z_in_Ainteg[of "m"],
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_t_def)
done
lemma a_neg_i[simp]:"m < 0 \<Longrightarrow> (ant m) * \<infinity> = -\<infinity>"
apply (simp add:mult_ant_def inf_ant_def ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
cut_tac ant_z_in_Ainteg[of "m"],
simp add:minus_ant_def,
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_t_def)
done
lemma a_i_neg[simp]:"m < 0 \<Longrightarrow> \<infinity> * (ant m) = -\<infinity>"
apply (simp add:mult_ant_def inf_ant_def ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
cut_tac ant_z_in_Ainteg[of "m"],
simp add:minus_ant_def,
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_t_def)
done
lemma a_z_z:"(ant m) * (ant n) = ant (m*n)"
apply (simp add:mult_ant_def ant_def,
cut_tac ant_z_in_Ainteg[of "m"],
cut_tac ant_z_in_Ainteg[of "n"],
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_t_def)
done
lemma a_pos_m[simp]:"0 < m \<Longrightarrow> (ant m) * (-\<infinity>) = -\<infinity>"
apply (simp add:mult_ant_def inf_ant_def minus_ant_def ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
cut_tac ant_z_in_Ainteg[of "m"],
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_t_def)
done
lemma a_m_pos[simp]:"0 < m \<Longrightarrow> (-\<infinity>) * (ant m) = -\<infinity>"
apply (simp add:mult_ant_def inf_ant_def minus_ant_def ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
cut_tac ant_z_in_Ainteg[of "m"],
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_t_def)
done
lemma a_neg_m[simp]:"m < 0 \<Longrightarrow> (ant m) * (-\<infinity>) = \<infinity>"
apply (simp add:mult_ant_def inf_ant_def minus_ant_def ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
cut_tac ant_z_in_Ainteg[of "m"],
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_t_def)
done
lemma neg_a_m[simp]:"m < 0 \<Longrightarrow> (-\<infinity>) * (ant m) = \<infinity>"
apply (simp add:mult_ant_def inf_ant_def minus_ant_def ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
cut_tac ant_z_in_Ainteg[of "m"],
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_t_def)
done
lemma a_m_m[simp]:"(-\<infinity>) * (-\<infinity>) = \<infinity>"
apply (simp add:mult_ant_def inf_ant_def minus_ant_def ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply (simp add:zag_t_def)
done
lemma inj_on_Abs_Ainteg:"inj_on Abs_Ainteg Ainteg"
apply (simp add:inj_on_def)
apply (rule ballI)+
apply (rule impI,
subgoal_tac "Rep_Ainteg (Abs_Ainteg x) = Rep_Ainteg (Abs_Ainteg y)",
thin_tac "Abs_Ainteg x = Abs_Ainteg y",
simp add:Abs_Ainteg_inverse, simp)
done
lemma an_Suc:"an (Suc n) = an n + 1"
apply (subst an_1[THEN sym])
apply (simp del:an_1 add:an_def)
apply (simp del:an_1 add:a_zpz, simp add:add.commute)
done
lemma aeq_zeq [iff]: "(ant m = ant n) = (m = n)"
apply (rule iffI)
apply (subgoal_tac "Rep_Ainteg (ant m) = Rep_Ainteg (ant n)",
thin_tac "ant m = ant n",
cut_tac ant_z_in_Ainteg[of "m"],
cut_tac ant_z_in_Ainteg[of "n"],
simp add:ant_def Abs_Ainteg_inverse)
apply simp+
done
lemma aminus:"- ant m = ant (-m)"
apply (simp add:ant_def minus_ant_def,
cut_tac ant_z_in_Ainteg[of "m"],
simp add:Abs_Ainteg_inverse)
done
lemma aminusZero:"- ant 0 = ant 0"
apply (simp add:aminus)
done
lemma ant_0: "ant 0 = (0::ant)"
by (simp add: Zero_ant_def)
lemma inf_neq_0[simp]:"\<infinity> \<noteq> 0"
apply (cut_tac z_neq_inf[of "0"], frule not_sym)
apply (simp add:ant_0)
done
lemma zero_neq_inf[simp]:"0 \<noteq> \<infinity>"
by (cut_tac inf_neq_0, frule not_sym, simp)
lemma minf_neq_0[simp]:"-\<infinity> \<noteq> 0"
apply (cut_tac z_neq_minf[of "0"], frule not_sym)
apply (simp add:ant_0)
done
lemma zero_neq_minf[simp]:"0 \<noteq> -\<infinity>"
by (cut_tac minf_neq_0, frule not_sym, simp)
lemma a_minus_zero[simp]:"-(0::ant) = 0"
by (cut_tac aminusZero, simp add:ant_0)
lemma a_minus_minus: "- (- z) = (z::ant)"
apply (cut_tac mem_ant[of "z"])
apply (erule disjE, simp add:minf, simp add: minus_ant_def,
cut_tac ant_minf_in_Ainteg,
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply (erule disjE) apply (erule exE, simp add:aminus)
apply (simp add:minf, simp add: minus_ant_def,
cut_tac ant_minf_in_Ainteg,
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
simp add:inf_ant_def)
done
lemma aminus_0: "- (- 0) = (0::ant)"
apply (simp add:a_minus_minus)
done
lemma a_a_z_0:"\<lbrakk> 0 < z; a * ant z = 0\<rbrakk> \<Longrightarrow> a = 0"
by (rule contrapos_pp, simp+, cut_tac mem_ant[of "a"], erule disjE,
simp, erule disjE, erule exE, simp add:a_z_z,
simp only:ant_0[THEN sym], simp, simp)
lemma adiv_eq:"\<lbrakk> z \<noteq> 0; a * (ant z) = b * (ant z)\<rbrakk> \<Longrightarrow> a = b"
apply (cut_tac mem_ant[of "a"], cut_tac mem_ant[of "b"],
(erule disjE)+, simp, erule disjE, erule exE,
cut_tac less_linear[of "z" "0"], erule disjE, simp add:a_z_z,
frule sym, thin_tac "\<infinity> = ant (za * z)", simp,
simp add:a_z_z, frule sym, thin_tac "- \<infinity> = ant (za * z)", simp,
cut_tac less_linear[of "z" "0"], erule disjE, simp,
simp, erule disjE, erule exE)
apply (erule disjE,
cut_tac less_linear[of "z" "0"], simp,
erule disjE, simp add:a_z_z, simp add:a_z_z,
erule disjE, erule exE, simp add:a_z_z,
cut_tac less_linear[of "z" "0"], simp,
erule disjE, simp add:a_z_z, simp add:a_z_z,
erule disjE,
cut_tac less_linear[of "z" "0"], simp,
erule disjE, simp+)
apply (erule disjE, erule exE, simp add:a_z_z,
cut_tac less_linear[of "z" "0"], simp, erule disjE, simp,
frule sym, thin_tac "- \<infinity> = ant (za * z)", simp,
simp, frule sym, thin_tac "\<infinity> = ant (za * z)", simp,
cut_tac less_linear[of "z" "0"], simp)
done
lemma aminus_add_distrib: "- (z + w) = (- z) + (- w::ant)"
apply (cut_tac mem_ant[of "z"], cut_tac mem_ant[of "w"],
(erule disjE)+, simp add:a_minus_minus,
erule disjE, erule exE, simp,
simp add:a_minus_minus aminus, simp add:a_minus_minus)
apply ((erule disjE)+, erule exE,
simp add:a_minus_minus, simp add:aminus,
simp add:a_minus_minus)
apply ((erule disjE)+, (erule exE)+, simp add:a_zpz aminus,
erule exE, simp add:aminus,
erule disjE, erule exE, simp add:aminus, simp)
done
lemma aadd_commute:"(x::ant) + y = y + x"
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"])
apply (erule disjE, erule disjE, simp,
erule disjE, erule exE, simp+,
(erule disjE)+, erule exE, simp+)
apply ((erule disjE)+, (erule exE)+, simp add:a_zpz,
erule exE, simp, erule disjE, erule exE, simp+)
done
definition
aug_inf :: "ant set" ("Z\<^sub>\<infinity>") where
"Z\<^sub>\<infinity> = {(z::ant). z \<noteq> -\<infinity> }"
definition
aug_minf :: "ant set" ("Z\<^sub>-\<^sub>\<infinity>") where
"Z\<^sub>-\<^sub>\<infinity> = {(z::ant). z \<noteq> \<infinity> }"
lemma z_in_aug_inf:"ant z \<in> Z\<^sub>\<infinity>"
apply (simp add:aug_inf_def)
done
lemma Zero_in_aug_inf:"0 \<in> Z\<^sub>\<infinity>"
by (simp only:Zero_ant_def, simp add: aug_inf_def)
lemma z_in_aug_minf:"ant z \<in> Z\<^sub>-\<^sub>\<infinity>"
by (simp add:aug_minf_def)
lemma mem_aug_minf:"a \<in> Z\<^sub>-\<^sub>\<infinity> \<Longrightarrow> a = - \<infinity> \<or> (\<exists>z. a = ant z)"
by (cut_tac mem_ant[of a], simp add:aug_minf_def)
lemma minus_an_in_aug_minf:" - an n \<in> Z\<^sub>-\<^sub>\<infinity>"
apply (simp add:an_def)
apply (simp add:aminus)
apply (simp add:z_in_aug_minf)
done
lemma Zero_in_aug_minf:"0 \<in> Z\<^sub>-\<^sub>\<infinity>"
by (simp add:Zero_ant_def aug_minf_def)
lemma aadd_assoc_i: "\<lbrakk>x \<in> Z\<^sub>\<infinity>; y \<in> Z\<^sub>\<infinity>; z \<in> Z\<^sub>\<infinity>\<rbrakk> \<Longrightarrow> (x + y) + z = x + (y + z)"
apply (cut_tac mem_ant[of "x"],
cut_tac mem_ant[of "y"],
cut_tac mem_ant[of "z"], simp add:aug_inf_def,
(erule disjE)+, (erule exE)+, (simp add:a_zpz)+,
(erule exE)+, simp add:a_zpz)
apply ((erule disjE)+, (erule exE)+, simp,
erule exE, simp,
(erule disjE)+, (erule exE)+, simp add:a_zpz,
erule exE, simp, erule disjE, erule exE, simp)
apply simp
done
lemma aadd_assoc_m: "\<lbrakk>x \<in> Z\<^sub>-\<^sub>\<infinity>; y \<in> Z\<^sub>-\<^sub>\<infinity>; z \<in> Z\<^sub>-\<^sub>\<infinity>\<rbrakk> \<Longrightarrow>
(x + y) + z = x + (y + z)"
apply (cut_tac mem_ant[of "x"],
cut_tac mem_ant[of "y"],
cut_tac mem_ant[of "z"], simp add:aug_minf_def )
apply ((erule disjE)+, simp, erule exE, simp,
erule disjE, erule exE, simp, (erule exE)+, simp add:a_zpz)
apply ((erule disjE)+, erule exE, simp, (erule exE)+, simp,
erule disjE, erule exE, simp, erule exE, simp add:a_zpz)
apply ((erule exE)+, simp add:a_zpz)
done
lemma aadd_0_r: "x + (0::ant) = x"
apply (cut_tac mem_ant[of "x"], simp add:Zero_ant_def)
apply ((erule disjE)+, simp)
apply (erule disjE, erule exE, simp add:a_zpz,
simp)
done
lemma aadd_0_l: "(0::ant) + x = x"
apply (cut_tac mem_ant[of "x"], simp add:Zero_ant_def)
apply ((erule disjE)+, simp)
apply (erule disjE, erule exE, simp, simp add:a_zpz, simp)
done
lemma aadd_minus_inv: "(- x) + x = (0::ant)" (** \<longrightarrow> aadd_minus_l **)
apply (cut_tac mem_ant[of "x"],
erule disjE, simp add:a_minus_minus,
erule disjE, erule exE, simp add:aminus, simp add:a_zpz,
simp add:Zero_ant_def, simp)
done
lemma aadd_minus_r: "x + (- x) = (0::ant)"
apply (cut_tac aadd_minus_inv[of "x"])
apply (simp add:aadd_commute)
done
lemma ant_minus_inj:"ant z \<noteq> ant w \<Longrightarrow> - ant z \<noteq> - ant w"
by (simp add:aminus)
lemma aminus_mult_minus: "(- (ant z)) * (ant w) = - ((ant z) * (ant w))"
apply (simp add:ant_def minus_ant_def,
cut_tac ant_z_in_Ainteg[of "z"],
cut_tac ant_z_in_Ainteg[of "-z"],
cut_tac ant_z_in_Ainteg[of "w"],
simp add:Abs_Ainteg_inverse)
apply (simp add:mult_ant_def) apply (simp add:Abs_Ainteg_inverse,
simp add:zag_t_def,
cut_tac ant_z_in_Ainteg[of "z * w"])
apply (simp add:Abs_Ainteg_inverse)
done
lemma amult_commute: "(x::ant) * y = y * x"
apply (cut_tac mem_ant[of "x"],
cut_tac mem_ant[of "y"])
apply (erule disjE, erule disjE, simp)
apply (erule disjE, erule exE, simp)
apply (cut_tac x = 0 and y = z in less_linear)
apply (erule disjE, simp)
apply (erule disjE, rotate_tac -1, frule sym, thin_tac "0 = z", simp)
apply (simp add:inf_ant_def ant_def, simp add:minus_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_z_in_Ainteg[of "0"],
cut_tac ant_z_in_Ainteg[of "-1"],
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply (simp add:mult_ant_def, simp add:Abs_Ainteg_inverse,
simp add:zag_t_def, simp)
apply (simp add:inf_ant_def)
apply (simp add:mult_ant_def minus_ant_def,
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse, simp add:zag_t_def)
apply (erule disjE, erule disjE, simp)
apply (erule exE,
cut_tac x = 0 and y = z in less_linear)
apply (erule disjE, simp)
apply (erule disjE, rotate_tac -1, thin_tac "0 = z", simp add:mult_ant_def,
simp add:ant_def inf_ant_def minus_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac z = z in ant_z_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse, simp add:zag_t_def,
simp)
apply (simp add:inf_ant_def minus_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac z = z in ant_z_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
simp add:mult_ant_def,
simp add:Abs_Ainteg_inverse, simp add:zag_t_def)
apply ((erule disjE)+, (erule exE)+, simp add:a_z_z)
apply (erule exE,
cut_tac x = 0 and y = z in less_linear,
erule disjE, simp)
apply (erule disjE, rotate_tac -1, frule sym, thin_tac "0 = z", simp,
simp add:mult_ant_def ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_z_in_Ainteg[of "0"],
simp add:Abs_Ainteg_inverse, simp add:zag_t_def,
simp)
apply (erule disjE, erule exE,
cut_tac x = 0 and y = z in less_linear,
erule disjE, simp,
erule disjE, rotate_tac -1, frule sym, thin_tac "0 = z", simp,
simp add:mult_ant_def ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_z_in_Ainteg[of "0"],
simp add:Abs_Ainteg_inverse, simp add:zag_t_def,
simp+)
done
lemma z_le_i[simp]:"(ant x) \<le> \<infinity> "
apply (simp add:le_ant_def ant_def,
cut_tac ant_z_in_Ainteg[of "0"],
cut_tac ant_z_in_Ainteg[of "x"],
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
simp add:inf_ant_def,
simp add:Abs_Ainteg_inverse)
done
lemma z_less_i[simp]:"(ant x) < \<infinity> "
apply (cut_tac z_le_i[of "x"],
cut_tac z_neq_inf[of "x"],
simp add:less_ant_def)
done
lemma m_le_z:"-\<infinity> \<le> (ant x) "
apply (simp add:le_ant_def ant_def,
cut_tac ant_z_in_Ainteg[of "0"],
cut_tac ant_z_in_Ainteg[of "x"],
cut_tac ant_minf_in_Ainteg,
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
simp add:inf_ant_def,
simp add:minus_ant_def,
simp add:Abs_Ainteg_inverse)
done
lemma m_less_z[simp]:"-\<infinity> < (ant x)"
apply (cut_tac m_le_z[of "x"],
cut_tac z_neq_minf[of "x"],
frule not_sym, thin_tac "ant x \<noteq> - \<infinity>",
simp add:less_ant_def)
done
lemma noninf_mem_Z:"\<lbrakk>x \<in> Z\<^sub>\<infinity>; x \<noteq> \<infinity>\<rbrakk> \<Longrightarrow> \<exists>(z::int). x = ant z"
apply (simp add:aug_inf_def)
apply (cut_tac mem_ant[of "x"], simp)
done
lemma z_mem_Z:"ant z \<in> Z\<^sub>\<infinity>"
by (simp add:aug_inf_def)
lemma inf_ge_any[simp]:"x \<le> \<infinity>"
apply (cut_tac mem_ant[of "x"], erule disjE)
apply (simp add:inf_ant_def minus_ant_def,
cut_tac ant_minf_in_Ainteg,
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
simp add:le_ant_def, simp add:Abs_Ainteg_inverse)
apply (erule disjE, erule exE, simp)
apply (simp add:inf_ant_def,
cut_tac ant_inf_in_Ainteg,
simp add:le_ant_def, simp add:Abs_Ainteg_inverse)
done
lemma zero_lt_inf:"0 < \<infinity>"
by (simp add:less_ant_def)
lemma minf_le_any[simp]:"-\<infinity> \<le> x"
apply (cut_tac mem_ant[of "x"], erule disjE)
apply (simp add:inf_ant_def minus_ant_def,
cut_tac ant_minf_in_Ainteg,
cut_tac ant_inf_in_Ainteg,
simp add:Abs_Ainteg_inverse,
simp add:le_ant_def, simp add:Abs_Ainteg_inverse)
apply (erule disjE, erule exE, simp)
apply (simp add:inf_ant_def, simp add:minus_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:le_ant_def, simp add:Abs_Ainteg_inverse)
apply simp
done
lemma minf_less_0:"-\<infinity> < 0"
by (simp add:less_ant_def)
lemma ale_antisym[simp]:"\<lbrakk>(x::ant) \<le> y; y \<le> x \<rbrakk> \<Longrightarrow> x = y"
apply (rule contrapos_pp, simp+)
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"])
apply (erule disjE, erule disjE, simp)
apply (erule disjE, erule exE, simp, simp add:ant_def,
simp add:minus_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
cut_tac z = z in ant_z_in_Ainteg, simp add:Abs_Ainteg_inverse,
simp add:le_ant_def Abs_Ainteg_inverse)
apply (thin_tac "x \<le> y",
simp add:le_ant_def ant_def minus_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply (erule disjE, erule disjE, erule exE,
thin_tac "y \<le> x",
simp add:le_ant_def ant_def minus_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
cut_tac z = z in ant_z_in_Ainteg, simp add:Abs_Ainteg_inverse)
apply (thin_tac "y \<le> x",
simp add:le_ant_def ant_def minus_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
simp add:Abs_Ainteg_inverse)
apply ((erule disjE)+, (erule exE)+,
cut_tac z = z in ant_z_in_Ainteg,
cut_tac z = za in ant_z_in_Ainteg,
simp add:le_ant_def ant_def,
simp add:Abs_Ainteg_inverse)
apply (erule exE,
simp add:le_ant_def ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac z = z in ant_z_in_Ainteg, simp add:Abs_Ainteg_inverse)
apply (erule disjE, erule exE, thin_tac "y \<le> x",
simp add:le_ant_def ant_def minus_ant_def inf_ant_def,
cut_tac ant_inf_in_Ainteg,
cut_tac ant_minf_in_Ainteg,
cut_tac z = z in ant_z_in_Ainteg, simp add:Abs_Ainteg_inverse)
apply simp
done
lemma x_gt_inf[simp]:"\<infinity> \<le> x \<Longrightarrow> x = \<infinity>"
apply (cut_tac inf_ge_any[of "x"],
rule ale_antisym[of "x" "\<infinity>"], assumption+)
done
lemma Zinf_pOp_closed:"\<lbrakk>x \<in> Z\<^sub>\<infinity>; y \<in> Z\<^sub>\<infinity>\<rbrakk> \<Longrightarrow> x + y \<in> Z\<^sub>\<infinity>"
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"],
simp add:aug_inf_def,
(erule disjE)+, (erule exE)+, simp add:a_zpz,
cut_tac z = "-(z + za)" in z_neq_inf,
rule contrapos_pp, simp+,
cut_tac m1 = "z+za" in aminus[THEN sym], simp add:a_minus_minus,
erule exE, simp, simp add:minf_neq_inf[THEN not_sym],
erule disjE, erule exE, simp,
simp add:minf_neq_inf[THEN not_sym],
simp)
done
lemma Zminf_pOp_closed:"\<lbrakk>x \<in> Z\<^sub>-\<^sub>\<infinity>; y \<in> Z\<^sub>-\<^sub>\<infinity>\<rbrakk> \<Longrightarrow> x + y \<in> Z\<^sub>-\<^sub>\<infinity>"
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"],
simp add:aug_minf_def,
(erule disjE)+, simp, erule exE, simp,
erule disjE, erule exE, simp,
(erule exE)+, simp add:a_zpz)
done
lemma amult_distrib1:"(ant z) \<noteq> 0 \<Longrightarrow>
(a + b) * (ant z) = a * (ant z) + b * (ant z)"
apply (cut_tac mem_ant[of "a"], cut_tac mem_ant[of "b"],
(erule disjE)+, simp, cut_tac less_linear[of "z" "0"],
erule disjE, simp, erule disjE, simp, simp add:ant_0, simp,
erule disjE, erule exE, simp,
cut_tac less_linear[of "z" "0"],
erule disjE, simp add:a_z_z, erule disjE, simp add:ant_0,
simp add:a_z_z,
cut_tac less_linear[of "z" "0"], simp,
erule disjE, simp add:ant_0[THEN sym] a_z_z)
apply (erule disjE, simp add:ant_0[THEN sym],
simp, simp add:ant_0[THEN sym], simp add:a_z_z,
(erule disjE)+, (erule exE)+, cut_tac less_linear[of "z" "0"], simp,
erule disjE, simp add:a_z_z,
erule disjE, simp add:ant_0, simp add:a_z_z,
cut_tac less_linear[of "z" "0"],
erule disjE, simp add:ant_0[THEN sym])
apply (simp add:a_z_z, simp,
erule disjE, simp add:ant_0, simp add:ant_0[THEN sym] a_z_z,
(erule disjE)+, (erule exE)+, simp add:a_zpz a_z_z,
simp add: distrib_right, erule exE, simp add:a_z_z,
cut_tac less_linear[of "z" "0"], erule disjE, simp,
erule disjE, simp add:ant_0, simp)
apply (erule disjE, erule exE, simp,
cut_tac less_linear[of "z" "0"], erule disjE, simp add:a_z_z,
erule disjE, simp add:ant_0, simp add:a_z_z,
cut_tac less_linear[of "z" "0"], erule disjE, simp,
erule disjE, simp add:ant_0, simp)
done
lemma amult_0_r:"(ant z) * 0 = 0"
by (simp add:ant_0[THEN sym] a_z_z)
lemma amult_0_l:"0 * (ant z) = 0"
by (simp add:ant_0[THEN sym] a_z_z)
definition
asprod :: "[int, ant] \<Rightarrow> ant" (infixl "*\<^sub>a" 200) where
"m *\<^sub>a x ==
if x = \<infinity> then (if 0 < m then \<infinity> else (if m < 0 then -\<infinity> else
if m = 0 then 0 else undefined))
else (if x = -\<infinity> then
(if 0 < m then -\<infinity> else (if m < 0 then \<infinity> else
if m = 0 then 0 else undefined))
else (ant m) * x)"
lemma asprod_pos_inf[simp]:"0 < m \<Longrightarrow> m *\<^sub>a \<infinity> = \<infinity>"
apply (simp add:asprod_def)
done
lemma asprod_neg_inf[simp]:"m < 0 \<Longrightarrow> m *\<^sub>a \<infinity> = -\<infinity>"
apply (simp add:asprod_def)
done
lemma asprod_pos_minf[simp]:"0 < m \<Longrightarrow> m *\<^sub>a (-\<infinity>) = (-\<infinity>)"
apply (simp add:asprod_def)
done
lemma asprod_neg_minf[simp]:"m < 0 \<Longrightarrow> m *\<^sub>a (-\<infinity>) = \<infinity>"
apply (simp add:asprod_def)
done
lemma asprod_mult:" m *\<^sub>a (ant n) = ant(m * n)"
apply (cut_tac z_neq_inf[of "n"],
cut_tac z_neq_minf[of "n"],
simp add:asprod_def, simp add:a_z_z)
done
lemma asprod_1:"1 *\<^sub>a x = x"
by (cut_tac mem_ant[of "x"], erule disjE, simp,
erule disjE, erule exE, simp add:asprod_mult, simp)
(** atode asprod_1_x to awaseru **)
lemma agsprod_assoc_a:"m *\<^sub>a (n *\<^sub>a (ant x)) = (m * n) *\<^sub>a (ant x)"
apply (simp add:asprod_mult)
done
lemma agsprod_assoc:"\<lbrakk>m \<noteq> 0; n \<noteq> 0\<rbrakk> \<Longrightarrow> m *\<^sub>a (n *\<^sub>a x) = (m * n) *\<^sub>a x"
apply (cut_tac less_linear[of "m" "0"], cut_tac less_linear[of "n" "0"],
cut_tac mem_ant[of "x"],
(erule disjE)+, simp,
frule zmult_neg_neg[of "m" "n"], assumption+, simp)
apply (erule disjE, erule exE, simp add:asprod_mult,
frule zmult_neg_neg[of "m" "n"], assumption+, simp+,
erule disjE, simp,
frule zmult_neg_pos[of "m" "n"], assumption+, simp,
erule disjE, erule exE, simp,
frule zmult_neg_pos[of "m" "n"], assumption+, simp add:asprod_mult,
frule zmult_neg_pos[of "m" "n"], assumption+, simp)
apply (simp, (erule disjE)+,
frule zmult_pos_neg[of "m" "n"], assumption+,
simp,
erule disjE, erule exE, simp add:asprod_mult,
frule zmult_pos_neg[of "m" "n"], assumption+, simp)
apply (frule zmult_pos_pos[of "m" "n"], assumption+,
erule disjE, simp,
erule disjE, erule exE, simp add:asprod_mult, simp)
done
lemma asprod_distrib1:"m \<noteq> 0 \<Longrightarrow> m *\<^sub>a (x + y) = (m *\<^sub>a x) + (m *\<^sub>a y)"
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"])
apply (cut_tac less_linear[of "m" "0"],
erule disjE,
erule disjE, erule disjE, simp,
erule disjE, simp add:asprod_def add_ant_def, simp,
simp, (erule disjE)+, erule exE, simp add:asprod_mult,
simp add:Zero_ant_def asprod_mult)
apply (erule disjE, erule exE, simp add:asprod_mult,
simp add: Zero_ant_def asprod_mult,
erule disjE, erule disjE, erule disjE, erule exE,
simp add:asprod_mult,
simp add:Zero_ant_def asprod_mult,
erule disjE, erule exE, simp add:asprod_mult,
simp add:Zero_ant_def asprod_mult)
apply (simp, erule disjE, erule exE, simp,
(erule disjE)+, erule exE, simp add:asprod_mult,
simp add:a_zpz, simp add:asprod_mult distrib_left,
simp add:asprod_mult)
apply (erule disjE, erule exE, simp add:a_zpz asprod_mult,
simp add: distrib_left, simp add:asprod_mult,
(erule disjE)+, erule exE, simp add:asprod_mult, simp,
erule disjE, erule exE, simp add:asprod_mult, simp)
done
lemma asprod_0_x[simp]:"0 *\<^sub>a x = 0"
apply (simp add:asprod_def, (rule impI)+,
cut_tac mem_ant[of "x"], simp, erule exE,
simp add:asprod_def a_z_z, simp add:ant_0)
done
lemma asprod_n_0:"n *\<^sub>a 0 = 0"
apply (simp add:Zero_ant_def asprod_mult)
done
lemma asprod_distrib2:"\<lbrakk>0 < i; 0 < j\<rbrakk> \<Longrightarrow> (i + j) *\<^sub>a x = (i *\<^sub>a x) + (j *\<^sub>a x)"
by (cut_tac mem_ant[of "x"], erule disjE, simp,
erule disjE, erule exE, simp add:asprod_mult,
simp add: distrib_right a_zpz, simp)
lemma asprod_minus:"x \<noteq> -\<infinity> \<and> x \<noteq> \<infinity> \<Longrightarrow> - z *\<^sub>a x = z *\<^sub>a (- x)"
apply (cut_tac mem_ant[of "x"], erule disjE, simp+)
apply (erule exE, simp add:asprod_mult aminus)
done
lemma asprod_div_eq:"\<lbrakk>n \<noteq> 0; n *\<^sub>a x = n *\<^sub>a y\<rbrakk> \<Longrightarrow> x = y"
apply (cut_tac less_linear[of "n" "0"], simp)
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"])
apply ((erule disjE)+, simp,
erule disjE, erule exE, rule contrapos_pp, simp+,
simp add:asprod_mult)
apply (cut_tac z1 = "n * z" in z_neq_inf[THEN not_sym], simp+)
apply ((erule disjE)+, erule exE, simp add:asprod_mult,
cut_tac z = "n * z" in z_neq_inf,
rule contrapos_pp, simp, simp,
(erule disjE)+, (erule exE)+, simp add:asprod_mult,
erule exE, simp add: asprod_mult)
apply (erule disjE, erule exE, simp add:asprod_mult,
simp add:z_neq_minf[THEN not_sym], simp)
apply ((erule disjE)+, simp,
erule disjE, erule exE, rule contrapos_pp, simp+,
simp add:asprod_mult,
cut_tac z1 = "n * z" in z_neq_minf[THEN not_sym], simp,
rule contrapos_pp, simp+)
apply ((erule disjE)+, (erule exE)+, simp add:asprod_mult,
erule exE, simp add:asprod_mult,
erule disjE, erule exE, simp add:asprod_mult
z_neq_inf[THEN not_sym], simp)
apply (erule disjE, simp, erule disjE, erule exE, simp add:asprod_mult
z_neq_inf[THEN not_sym], simp)
done
lemma asprod_0:"\<lbrakk>z \<noteq> 0; z *\<^sub>a x = 0 \<rbrakk> \<Longrightarrow> x = 0"
by (rule asprod_div_eq[of "z" "x" "0"], assumption, simp add:asprod_n_0)
lemma asp_z_Z:"z *\<^sub>a ant x \<in> Z\<^sub>\<infinity>"
by (simp add:asprod_mult z_in_aug_inf)
lemma tna_ant:" tna (ant z) = z"
apply (cut_tac z_neq_minf[of "z"], cut_tac z_neq_inf[of "z"],
simp add:ant_def tna_def)
apply (cut_tac ant_z_in_Ainteg[of "z"], simp add:Abs_Ainteg_inverse)
done
lemma ant_tna:"x \<noteq> \<infinity> \<and> x \<noteq> -\<infinity> \<Longrightarrow> ant (tna x) = x"
apply (cut_tac mem_ant[of "x"], simp, erule exE)
apply (simp add:ant_def tna_def)
apply (cut_tac z = z in ant_z_in_Ainteg, simp add:Abs_Ainteg_inverse)
done
lemma ant_sol:"\<lbrakk>a \<in> Z\<^sub>\<infinity>; b \<in> Z\<^sub>\<infinity>; c \<in> Z\<^sub>\<infinity>; b \<noteq> \<infinity>; a = b + c\<rbrakk> \<Longrightarrow> a - b = c"
apply (subgoal_tac "-b \<in> Z\<^sub>\<infinity>", simp add:diff_ant_def,
subgoal_tac "a + (-b) = b + c + (-b)",
subst aadd_commute[of "b" "c"], subst aadd_assoc_i, assumption+,
simp add:aadd_minus_r, simp add:aadd_0_r, simp)
apply (cut_tac mem_ant[of "b"], simp add:aug_inf_def,
erule exE, simp add:aminus)
done
subsection "Ordering of integers and ordering nats"
subsection \<open>The \<open>\<le>\<close> Ordering\<close>
lemma zneq_aneq:"(n \<noteq> m) = ((ant n) \<noteq> (ant m))"
apply (rule iffI)
apply (rule contrapos_pp, simp+)
done
lemma ale:"(n \<le> m) = ((ant n) \<le>(ant m))"
apply (rule iffI)
apply (simp add:ant_def le_ant_def,
cut_tac ant_z_in_Ainteg[of "n"],
cut_tac ant_z_in_Ainteg[of "m"],
simp add:Abs_Ainteg_inverse)+
done
lemma aless:"(n < m) = ((ant n) < (ant m))"
apply (simp add:less_ant_def,
cut_tac ale[of "n" "m"], arith)
done
lemma ale_refl: "w \<le> (w::ant)"
apply (cut_tac mem_ant[of "w"],
erule disjE, simp,
erule disjE, erule exE, simp,
subst ale[THEN sym], simp+)
done
lemma aeq_ale:"(a::ant) = b \<Longrightarrow> a \<le> b"
by (simp add:ale_refl)
lemma ale_trans: "\<lbrakk> (i::ant) \<le> j; j \<le> k \<rbrakk> \<Longrightarrow> i \<le> k"
apply (cut_tac mem_ant[of "i"], cut_tac mem_ant[of "j"],
cut_tac mem_ant[of "k"],
(erule disjE)+, simp add:ale_refl, erule disjE, erule exE, simp+,
(erule disjE)+, simp add:ale_refl, simp add:ale_refl)
apply ((erule disjE)+, erule exE, simp+,
erule exE, simp,
cut_tac x = "ant z" in minf_le_any,
frule_tac x = "ant z" in ale_antisym[of _ "-\<infinity>"], assumption+,
simp+,
cut_tac minf_le_any[of "\<infinity>"], frule ale_antisym[of "-\<infinity>" "\<infinity>"],
simp+)
apply (erule disjE, simp,
(erule disjE)+, (erule exE)+, simp,
cut_tac x = "ant za" in minf_le_any,
frule_tac x = "ant za" in ale_antisym[of _ "-\<infinity>"], assumption+,
simp, erule exE,
cut_tac x = "ant z" in minf_le_any, simp)
apply (cut_tac minf_le_any[of "\<infinity>"],
frule_tac ale_antisym[of "-\<infinity>" "\<infinity>"], assumption+,
simp, erule disjE, erule exE, simp,
cut_tac x = "ant z" in inf_ge_any,
frule_tac x = "ant z" in ale_antisym[of _ "\<infinity>"], assumption+,
simp)
apply (cut_tac minf_le_any[of "\<infinity>"], frule ale_antisym[of "-\<infinity>" "\<infinity>"],
simp+,
(erule disjE)+, (erule exE)+, simp add:ale[THEN sym],
simp, (erule disjE)+, (erule exE)+,
cut_tac x = "ant za" in inf_ge_any,
frule_tac x = "ant za" in ale_antisym[of _ "\<infinity>"],
simp+)
apply (erule disjE, erule exE,
cut_tac inf_ge_any[of "j"],
frule ale_antisym[of "j" "\<infinity>"], assumption+,
cut_tac x = "ant z" in inf_ge_any, simp+)
done
(* Axiom 'order_aless_le_not_le' of class 'order': *)
lemma aless_le_not_le: "((w::ant) < z) = (w \<le> z \<and> \<not> z \<le> w)"
by (auto simp add: less_ant_def)
instance ant :: order
proof qed
(assumption |
rule ale_refl ale_trans ale_antisym aless_le_not_le)+
(* Axiom 'linorder_linear' of class 'linorder': *)
lemma ale_linear: "(z::ant) \<le> w \<or> w \<le> z"
apply (cut_tac mem_ant[of "z"], cut_tac mem_ant[of "w"],
erule disjE, simp,
erule disjE, simp)
apply ((erule disjE)+, (erule exE)+, simp add:ale[THEN sym],
simp add:linorder_linear)
apply simp+
done
instance ant :: linorder
proof qed (rule ale_linear)
lemmas aless_linear = less_linear [where 'a = ant]
lemma ant_eq_0_conv [simp]: "(ant n = 0) = (n = 0)"
apply (simp add:Zero_ant_def)
done
lemma aless_zless: "(ant m < ant n) = (m<n)"
by (simp add: ale ant_def linorder_not_le [symmetric])
lemma a0_less_int_conv [simp]: "(0 < ant n) = (0 < n)"
apply (simp add:Zero_ant_def)
apply (simp add:aless[THEN sym])
done
lemma a0_less_1: "0 < (1::ant)"
apply (simp add:Zero_ant_def One_ant_def)
apply (subst aless_zless) apply simp
done
lemma a0_neq_1 [simp]: "0 \<noteq> (1::ant)"
by (simp only:Zero_ant_def One_ant_def, subst zneq_aneq[THEN sym], simp)
lemma ale_zle [simp]: "((ant i) \<le> (ant j)) = (i\<le>j)"
by (subst ale[of "i" "j"], simp)
lemma ant_1 [simp]: "ant 1 = 1"
by (simp add: One_ant_def)
lemma zpos_apos:"(0 \<le> n) = (0 \<le> (ant n))"
apply (simp only:ale[of "0" "n"], simp only:ant_0[THEN sym])
done
lemma zposs_aposss:"(0 < n) = (0 < (ant n))"
apply (rule iffI)
apply (unfold Zero_ant_def,
subst aless[THEN sym, of "0" "n"], simp,
subst aless[of "0" "n"], simp)
done
lemma an_nat_pos[simp]:"0 \<le> an n"
by (simp add:ant_0[THEN sym] an_def)
lemma amult_one_l:" 1 * (x::ant) = x"
by (cut_tac mem_ant[of "x"], erule disjE, simp
only:ant_1[THEN sym], simp del:ant_1,
erule disjE, erule exE, simp only:ant_1[THEN sym],
simp del:ant_1 add:a_z_z,
simp only:ant_1[THEN sym], simp del:ant_1)
lemma amult_one_r:"(x::ant)* 1 = x"
by (cut_tac amult_one_l[of "x"], simp add:amult_commute)
lemma amult_eq_eq_r:"\<lbrakk>z \<noteq> 0; a * ant z = b * ant z\<rbrakk> \<Longrightarrow> a = b"
apply (cut_tac less_linear[of "z" "0"], simp,
cut_tac mem_ant[of "a"], cut_tac mem_ant[of "b"],
(erule disjE)+, simp,
erule disjE, erule exE, simp add:a_z_z,
frule sym, thin_tac "\<infinity> = ant (za * z)", simp,
simp, (erule disjE)+, simp, erule exE, simp add:a_z_z, simp)
apply ((erule disjE)+, (erule exE)+, simp add:a_z_z,
erule exE, simp add:a_z_z, erule disjE, erule exE,
simp add:a_z_z,
frule sym, thin_tac "- \<infinity> = ant (za * z)", simp, simp,
(erule disjE)+, simp, erule disjE, erule exE, simp add:a_z_z,
frule sym, thin_tac "- \<infinity> = ant (za * z)", simp, simp)
apply ((erule disjE)+, erule exE, simp add:a_z_z, simp,
(erule disjE)+, (erule exE)+, simp add:a_z_z,
erule exE, simp add:a_z_z, erule disjE, erule exE, simp add:a_z_z,
frule sym, thin_tac "\<infinity> = ant (za * z)", simp, simp)
done
lemma amult_eq_eq_l:"\<lbrakk>z \<noteq> 0; (ant z) * a = (ant z) * b\<rbrakk> \<Longrightarrow> a = b"
by (simp add:amult_commute, rule amult_eq_eq_r, assumption+)
lemma amult_pos:"\<lbrakk>0 < b; 0 \<le> x\<rbrakk> \<Longrightarrow> x \<le> (b *\<^sub>a x)"
apply (cut_tac mem_ant[of "x"], erule disjE, simp,
erule disjE, erule exE, simp add:asprod_mult,
simp add:zpos_apos[THEN sym],
frule_tac a = z and b = b in pos_zmult_pos, assumption+,
simp add:mult.commute, simp)
done
lemma asprod_amult:"0 < z \<Longrightarrow> z *\<^sub>a x = (ant z) * x"
apply (simp add:asprod_def)
done
lemma amult_pos1:"\<lbrakk>0 < b; 0 \<le> x\<rbrakk> \<Longrightarrow> x \<le> ((ant b) * x)"
by (frule amult_pos[of "b" "x"], assumption, simp add:asprod_amult)
lemma amult_pos_mono_l:"0 < w \<Longrightarrow> (((ant w) * x) \<le> ((ant w) * y)) = (x \<le> y)"
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"],
(erule disjE)+, simp, erule disjE, erule exE, simp, simp,
(erule disjE)+, erule exE, simp add:a_z_z)
apply (rule iffI,
cut_tac x = "ant (w * z)" in minf_le_any, frule_tac x = "ant (w * z)"
in ale_antisym, assumption+, simp,
cut_tac x = "ant z" in minf_le_any, frule_tac x = "ant z"
in ale_antisym, assumption+, simp)
apply simp
apply ((erule disjE)+, (erule exE)+, simp add:a_z_z)
apply (erule exE, simp add:a_z_z)
apply (erule disjE, erule exE, simp add:a_z_z,
rule iffI,
cut_tac x = "ant (w * z)" in inf_ge_any,
frule_tac x = "ant (w * z)" in ale_antisym[of _ "\<infinity>"], assumption+,
simp,
cut_tac x = "ant z" in inf_ge_any,
frule_tac x = "ant z" in ale_antisym[of _ "\<infinity>"], assumption+,
simp, simp)
done
lemma amult_pos_mono_r:"0 < w \<Longrightarrow> ((x * (ant w)) \<le> (y * (ant w))) = (x \<le> y)"
apply (simp add:amult_commute[of _ "ant w"])
apply (rule amult_pos_mono_l, assumption)
done
lemma apos_neq_minf:"0 \<le> a \<Longrightarrow> a \<noteq> -\<infinity>"
by (rule contrapos_pp, simp+,
cut_tac minf_le_any[of "0"],
frule ale_antisym[of "0" "-\<infinity>"], assumption+, simp)
lemma asprod_pos_mono:"0 < w \<Longrightarrow> ((w *\<^sub>a x) \<le> (w *\<^sub>a y)) = (x \<le> y)"
by (simp add:asprod_amult, simp add:amult_pos_mono_l)
lemma a_inv:"(a::ant) + b = 0 \<Longrightarrow> a = -b"
apply (cut_tac mem_ant[of "a"], cut_tac mem_ant[of "b"],
(erule disjE)+, frule sym, thin_tac "a + b = 0",
simp add:ant_0[THEN sym])
apply (erule disjE, erule exE, simp, simp,
(erule disjE)+, erule exE, simp, simp,
simp add:a_minus_minus,
(erule disjE)+, (erule exE)+, simp add:aminus a_zpz,
erule exE, simp,
erule disjE, erule exE, simp, simp)
done
lemma asprod_pos_pos:"0 \<le> x \<Longrightarrow> 0 \<le> int n *\<^sub>a x"
apply (cases "n = 0")
apply simp_all
using asprod_pos_mono [THEN sym, of "int n" "0" "x"]
apply (simp add:asprod_n_0)
done
lemma asprod_1_x[simp]:"1 *\<^sub>a x = x"
apply (simp add:asprod_def)
apply (rule impI)+
apply (cut_tac mem_ant[of "x"], simp, erule exE, simp add:a_z_z)
apply (simp only:ant_1[THEN sym], simp del:ant_1 add:a_z_z)
done
lemma asprod_n_1[simp]:"n *\<^sub>a 1 = ant n"
apply (simp only:ant_1[THEN sym]) apply (simp only:asprod_mult)
apply simp
done
subsection "Aug ordering"
lemma aless_imp_le:" x < (y::ant) \<Longrightarrow> x \<le> y"
by (simp add:less_ant_def)
lemma gt_a0_ge_1:"(0::ant) < x \<Longrightarrow> 1 \<le> x"
apply (cut_tac mem_ant[of "x"],
erule disjE, unfold Zero_ant_def, simp)
apply (cut_tac less_ant_def[of "0" "-\<infinity>"], simp add:ant_0,
cut_tac minf_le_any[of "0"],
frule ale_antisym[of "0" "-\<infinity>"], assumption+,
simp add:ant_0[THEN sym], blast)
apply (erule disjE, erule exE, unfold One_ant_def, simp del:ant_1,
simp add:aless_zless, simp)
done
lemma gt_a0_ge_aN:"\<lbrakk>0 < x; N \<noteq> 0\<rbrakk> \<Longrightarrow> (ant (int N)) \<le> (int N) *\<^sub>a x"
apply (cut_tac mem_ant[of "x"], erule disjE, simp)
apply (cut_tac aless_imp_le[of "0" "-\<infinity>"],
cut_tac minf_le_any[of "0"],
frule ale_antisym[of "0" "-\<infinity>"], simp,
simp only: Zero_ant_def, simp)
apply (erule disjE, erule exE, simp add:asprod_mult, simp)
done
lemma aless_le_trans:"\<lbrakk>(x::ant) < y; y \<le> z\<rbrakk> \<Longrightarrow> x < z"
by auto
lemma ale_less_trans:"\<lbrakk>(x::ant) \<le> y; y < z\<rbrakk> \<Longrightarrow> x < z"
by auto
lemma aless_trans:"\<lbrakk>(x::ant) < y; y < z\<rbrakk> \<Longrightarrow> x < z"
by auto
lemma ale_neq_less:"\<lbrakk> (x::ant)\<le> y; x \<noteq> y\<rbrakk> \<Longrightarrow> x < y"
apply (simp add:less_ant_def)
done
lemma aneg_le:"(\<not> (x::ant) \<le> y) = (y < x)"
apply (cut_tac ale_linear[of "y" "x"])
apply (rule iffI, simp)
apply (rule contrapos_pp, simp+)
done
lemma aneg_less:"(\<not> x < (y::ant)) = (y \<le> x)"
by auto
lemma aadd_le_mono:"x \<le> (y::ant) \<Longrightarrow> x + z \<le> y + z"
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"],
cut_tac mem_ant[of "z"],
(erule disjE)+, simp, erule disjE, erule exE, simp+,
(erule disjE)+, erule exE, simp+,
(erule disjE)+, (erule exE)+, simp, erule exE, simp,
erule disjE, erule exE, simp+, (erule disjE)+, simp,
erule exE, simp+,
cut_tac minf_le_any[of "\<infinity>"], frule ale_antisym[of "-\<infinity>" "\<infinity>"],
assumption+, simp, (erule disjE)+, (erule exE)+, simp+,
cut_tac x = "ant za" in minf_le_any,
frule_tac x = "ant za" in ale_antisym[of _ "-\<infinity>"], assumption+, simp)
apply (erule exE, simp,
cut_tac x = "ant za" in minf_le_any,
frule_tac x = "ant za" in ale_antisym[of _ "-\<infinity>"], assumption+, simp,
erule disjE, erule exE, simp+,
cut_tac minf_le_any[of "\<infinity>"], frule ale_antisym[of "-\<infinity>" "\<infinity>"],
assumption+, simp, (erule disjE)+, (erule exE)+, simp+,
erule exE, simp, erule disjE, erule exE, simp+)
apply (cut_tac x = "ant za" in inf_ge_any, frule_tac x = "ant za" in
ale_antisym[of _ "\<infinity>"], assumption+, simp+,
(erule disjE)+, (erule exE)+, simp add:a_zpz,
(erule exE)+, simp add:a_zpz, (erule disjE)+, (erule exE)+,
simp add:a_zpz, erule exE, simp,
(erule disjE)+, (erule exE)+, simp add:a_zpz)
apply (cut_tac x = "ant za" in inf_ge_any, frule_tac x = "ant za" in
ale_antisym[of _ "\<infinity>"], assumption+, simp+,
erule exE, simp, erule disjE, erule exE, simp+)
done
lemma aadd_less_mono_z:"(x::ant) < y \<Longrightarrow> (x + (ant z)) < (y + (ant z))"
apply (simp add:less_ant_def, simp add:aadd_le_mono)
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"])
apply auto
apply (metis a_inv a_ipi a_ipz a_zpz aadd_minus_r less_le diff_ant_def minf_less_0)
apply (metis a_inv a_ipi a_ipz a_zpz aadd_minus_r less_le diff_ant_def minf_less_0)
apply (metis a_zpz add_right_cancel aeq_zeq)
apply (metis a_zpz less_le z_less_i)
done
lemma aless_le_suc[simp]:"(a::ant) < b \<Longrightarrow> a + 1 \<le> b"
apply (cut_tac mem_ant[of "b"])
apply (erule disjE,
frule aless_imp_le[of "a" "b"], simp,
cut_tac minf_le_any[of "a"], frule ale_antisym[of "a" "-\<infinity>"],
assumption, simp)
apply (erule disjE, erule exE, cut_tac mem_ant[of "a"], erule disjE,
unfold One_ant_def, simp del:ant_1,
erule disjE, erule exE, simp del:ant_1 add:a_zpz, simp only:aless_zless,
frule aless_imp_le[of "a" "b"], simp del:ant_1, simp)
done
lemma aposs_le_1:"(0::ant) < x \<Longrightarrow> 1 \<le> x"
apply (frule aless_le_suc[of "0" "x"],
simp add:aadd_0_l)
done
lemma pos_in_aug_inf:"(0::ant) \<le> x \<Longrightarrow> x \<in> Z\<^sub>\<infinity>"
apply (simp add:aug_inf_def)
apply (rule contrapos_pp, simp+)
apply (cut_tac minf_le_any[of "0"],
frule ale_antisym[of "0" "-\<infinity>"], assumption+,
unfold Zero_ant_def,
simp )
done
lemma aug_inf_noninf_is_z:"\<lbrakk>x \<in> Z\<^sub>\<infinity>; x \<noteq> \<infinity>\<rbrakk> \<Longrightarrow> \<exists>z. x = ant z"
apply (cut_tac mem_ant[of "x"], simp add:aug_inf_def)
done
lemma aadd_two_pos:"\<lbrakk>0 \<le> (x::ant); 0 \<le> y\<rbrakk> \<Longrightarrow> 0 \<le> x + y"
apply (cut_tac Zero_in_aug_inf,
cut_tac pos_in_aug_inf[of "x"],
cut_tac pos_in_aug_inf[of "y"])
apply (cut_tac aadd_le_mono[of "0" "x" "y"], simp add:aadd_0_l,
assumption+)
done
lemma aadd_pos_poss:"\<lbrakk>(0::ant) \<le> x; 0 < y\<rbrakk> \<Longrightarrow> 0 < (x + y)"
apply (frule aless_imp_le[of "0" "y"],
subst less_ant_def, rule conjI, simp add:aadd_two_pos,
rule contrapos_pp, simp+)
apply (cut_tac Zero_in_aug_inf,
cut_tac pos_in_aug_inf[of "x"],
cut_tac pos_in_aug_inf[of "y"],
case_tac "y = \<infinity>", simp,
cut_tac mem_ant[of "x"], erule disjE,
simp add:aug_inf_def)
apply (erule disjE, erule exE, simp, simp,
case_tac "x = \<infinity>", unfold Zero_ant_def,
frule aug_inf_noninf_is_z[of "y"], assumption, erule exE,
simp, frule sym, thin_tac "\<infinity> = ant 0", simp)
apply (thin_tac "ant 0 \<le> y",
frule aug_inf_noninf_is_z[of "x"], assumption, erule exE,
frule aug_inf_noninf_is_z[of "y"], assumption, erule exE,
simp add:a_zpz, simp add: aless_zless)
apply (simp add:aless_imp_le)+
done
lemma aadd_poss_pos:"\<lbrakk>(0::ant) < x; 0 \<le> y\<rbrakk> \<Longrightarrow> 0 < (x + y)"
apply (subst aadd_commute, rule aadd_pos_poss, assumption+)
done
lemma aadd_pos_le:"0 \<le> (a::ant) \<Longrightarrow> b \<le> a + b"
apply (cut_tac mem_ant[of "a"], (erule disjE)+,
simp, cut_tac minf_le_any[of "0"], frule ale_antisym[of "0" "-\<infinity>"],
assumption+, simp)
apply (erule disjE, erule exE,
simp, thin_tac "a = ant z", cut_tac mem_ant[of "b"],
erule disjE, simp,
erule disjE, erule exE, simp add:a_zpz, simp only:ant_0[THEN sym],
simp only:ale, simp+)
apply (cut_tac mem_ant[of "b"],
erule disjE, simp,
erule disjE, erule exE, simp, simp)
done
lemma aadd_poss_less:"\<lbrakk>b \<noteq> \<infinity>; b \<noteq> -\<infinity>; 0 < a\<rbrakk> \<Longrightarrow> b < a + b"
apply (cut_tac mem_ant[of "b"], simp)
apply (erule exE,
cut_tac mem_ant[of "a"], erule disjE, simp,
thin_tac "a = - \<infinity>",
cut_tac minf_le_any[of "0"],
frule aless_imp_le[of "0" "-\<infinity>"],
frule ale_antisym[of "0" "-\<infinity>"], assumption+,
simp only:ant_0[THEN sym], simp)
apply (erule disjE, erule exE, simp add:a_zpz,
subst aless[THEN sym], simp, simp)
done
lemma ale_neg:"(0::ant) \<le> x \<Longrightarrow> (- x) \<le> 0"
apply (frule pos_in_aug_inf[of "x"])
apply (case_tac "x = \<infinity>", simp,
frule aug_inf_noninf_is_z[of "x"], assumption, erule exE,
simp add:aminus, unfold Zero_ant_def,
simp only:ale_zle)
done
lemma ale_diff_pos:"(x::ant) \<le> y \<Longrightarrow> 0 \<le> (y - x)"
apply (case_tac "y = -\<infinity>", simp,
cut_tac minf_le_any[of "x"],
frule ale_antisym[of "x" "-\<infinity>"], assumption+,
simp add:diff_ant_def a_minus_minus,
cut_tac mem_ant[of "y"], simp, thin_tac "y \<noteq> - \<infinity>",
erule disjE, erule exE)
apply (case_tac "x = \<infinity>", simp,
cut_tac x = "ant z" in inf_ge_any,
frule_tac x = "ant z" in ale_antisym[of _ "\<infinity>"], simp+,
cut_tac mem_ant[of "x"], simp+, erule disjE,
simp add:diff_ant_def a_minus_minus)
apply (erule exE, simp add:a_zdz, unfold Zero_ant_def,
simp only:ale_zle,
cut_tac mem_ant[of "x"], erule disjE,
simp add:diff_ant_def a_minus_minus,
erule disjE, erule exE, simp add:diff_ant_def aminus,
simp add:diff_ant_def ant_0)
done
lemma aless_diff_poss:"(x::ant) < y \<Longrightarrow> 0 < (y - x)"
apply (case_tac "y = -\<infinity>", simp,
cut_tac minf_le_any[of "x"],
frule less_imp_le[of "x" "-\<infinity>"],
frule antisym[of "x" "-\<infinity>"], assumption+,
cut_tac less_le[of "x" "-\<infinity>"], simp)
apply (case_tac "x = -\<infinity>", simp,
case_tac "y = \<infinity>", simp add:diff_ant_def a_minus_minus,
simp add:zero_lt_inf,
cut_tac mem_ant[of "y"], simp, erule exE, simp add:diff_ant_def
a_minus_minus, simp add:zero_lt_inf)
apply (case_tac "x = \<infinity>", simp,
frule aless_imp_le[of "\<infinity>" "y"],
cut_tac inf_ge_any[of "y"], frule ale_antisym[of "y" "\<infinity>"],
assumption+, simp,
cut_tac mem_ant[of "x"], simp, erule exE,
case_tac "y = \<infinity>", simp add:diff_ant_def aminus,
simp add:zero_lt_inf)
apply (cut_tac mem_ant[of "y"], simp, erule exE, simp,
simp add:diff_ant_def, simp add:aminus a_zpz,
simp add:aless_zless)
done
lemma ale_minus:" (x::ant) \<le> y \<Longrightarrow> - y \<le> - x"
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"])
apply ((erule disjE)+, simp, erule disjE, erule exE,
simp add:aminus a_minus_minus, simp add:a_minus_minus,
(erule disjE)+, (erule exE)+,
simp, cut_tac x = "ant z" in minf_le_any, frule_tac x = "ant z" in
ale_antisym[of _ "-\<infinity>"], assumption+, simp,
simp, cut_tac x = \<infinity> in minf_le_any,
frule_tac x = \<infinity> in ale_antisym[of _ "-\<infinity>"], assumption+, simp)
apply ((erule disjE)+, (erule exE)+, simp add:aminus, erule exE, simp,
erule disjE, erule exE, simp, cut_tac x = "ant z" in inf_ge_any,
frule_tac x = "ant z" in ale_antisym[of _ "\<infinity>"], assumption+, simp,
simp)
done
lemma aless_minus:"(x::ant) < y \<Longrightarrow> - y < - x"
by (simp add:less_ant_def, erule conjE, simp add:ale_minus,
rule not_sym, rule contrapos_pp, simp+,
cut_tac a_minus_minus[of "x"], simp add:a_minus_minus)
lemma aadd_minus_le:"(a::ant) \<le> 0 \<Longrightarrow> a + b \<le> b"
apply (frule ale_minus[of "a" "0"],
cut_tac aadd_pos_le[of "-a" "-b"], simp add:aminus_0)
apply (frule ale_minus[of "-b" "-a + -b"], simp add:aminus_add_distrib,
simp add:a_minus_minus, simp add:aminus_0)
done
lemma aadd_minus_less:"\<lbrakk>b \<noteq> -\<infinity> \<and> b \<noteq> \<infinity>; (a::ant) < 0\<rbrakk> \<Longrightarrow> a + b < b"
apply (simp add:less_ant_def, erule conjE,
simp add:aadd_minus_le)
apply (rule contrapos_pp, simp+,
cut_tac mem_ant[of "a"], cut_tac mem_ant[of "b"],
simp, erule disjE, erule exE, simp,
frule sym, thin_tac "- \<infinity> = ant z", simp,
erule disjE, (erule exE)+, simp add:a_zpz,
erule exE, simp, frule sym, thin_tac "\<infinity> = ant z", simp)
done
lemma an_inj:"an n = an m \<Longrightarrow> n = m"
apply (simp add:an_def)
done
lemma nat_eq_an_eq:"n = m \<Longrightarrow> an n = an m"
apply simp
done
lemma aneq_natneq:"(an n \<noteq> an m) = (n \<noteq> m)"
apply (simp add:an_def)
done
lemma ale_natle:" (an n \<le> an m) = (n \<le> m)"
apply (simp add:an_def)
done
lemma aless_natless:"(an n < an m) = (n < m)"
apply (simp add:an_def)
apply (simp add:aless_zless)
done
lemma na_an:"na (an n) = n"
by (simp only:na_def an_def,
subgoal_tac "\<not> ant (int n) < 0", simp,
simp add:tna_ant, subst aneg_less[of "ant (int n)" "0"],
simp only:ant_0[THEN sym], subst ale_zle[of "0" "int n"], simp)
lemma asprod_ge:
"0 < b \<Longrightarrow> N \<noteq> 0 \<Longrightarrow> an N \<le> int N *\<^sub>a b"
apply (frule aposs_le_1[of "b"])
apply simp
using asprod_pos_mono [THEN sym, of "int N" "1" "b"]
apply (simp only: ant_1 [THEN sym], simp add: asprod_amult, simp add:an_def)
done
lemma an_npn:"an (n + m) = an n + an m"
by (unfold an_def, simp add:a_zpz)
lemma an_ndn:"n \<le> m \<Longrightarrow> an (m - n) = an m - an n"
apply (cut_tac an_npn[of "m - n" n], simp)
apply (unfold an_def)
apply (simp add:a_zpz[of "int (m - n)" "int n"])
apply (subst a_zdz[of "int (m - n) + int n" "int n"], simp)
done
section "Amin, amax"
definition
amin :: "[ant, ant] \<Rightarrow> ant" where
"amin x y = (if (x \<le> y) then x else y)"
definition
amax :: "[ant, ant] \<Rightarrow> ant" where
"amax x y = (if (x \<le> y) then y else x)"
primrec Amin :: "[nat, nat \<Rightarrow> ant] \<Rightarrow> ant"
where
Amin_0 : "Amin 0 f = (f 0)"
| Amin_Suc :"Amin (Suc n) f = amin (Amin n f) (f (Suc n))"
primrec Amax :: "[nat, nat \<Rightarrow> ant] \<Rightarrow> ant"
where
Amax_0 : "Amax 0 f = f 0"
| Amax_Suc :"Amax (Suc n) f = amax (Amax n f) (f (Suc n))"
lemma amin_ge:"x \<le> amin x y \<or> y \<le> amin x y"
apply (simp add:amin_def)
done
lemma amin_le_l:"amin x y \<le> x"
apply (simp add:amin_def, cut_tac ale_linear[of "x" "y"],
rule impI, simp)
done
lemma amin_le_r:"amin x y \<le> y"
apply (simp add:amin_def)
done
lemma amax_le:"amax x y \<le> x \<or> amax x y \<le> y"
apply (simp add:amax_def)
done
lemma amax_le_n:"\<lbrakk>x \<le> n; y \<le> n\<rbrakk> \<Longrightarrow> amax x y \<le> n"
by (simp add:amax_def)
lemma amax_ge_l:"x \<le> amax x y"
apply (simp add:amax_def)
done
lemma amax_ge_r:"y \<le> amax x y"
apply (simp add:amax_def, cut_tac ale_linear[of "x" "y"],
rule impI, simp)
done
lemma amin_mem_i:"\<lbrakk>x \<in> Z\<^sub>\<infinity>; y \<in> Z\<^sub>\<infinity>\<rbrakk> \<Longrightarrow> amin x y \<in> Z\<^sub>\<infinity>"
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"], simp add:aug_inf_def,
(erule disjE)+, (erule exE)+, cut_tac amin_ge[of "x" "y"],
rule contrapos_pp, simp+,
erule disjE,
cut_tac x = "ant z" in minf_le_any,
frule_tac x = "ant z" in ale_antisym[of _ "-\<infinity>"], assumption+, simp,
cut_tac x = "ant za" in minf_le_any,
frule_tac x = "ant za" in ale_antisym[of _ "-\<infinity>"], assumption+, simp)
apply (erule exE, simp add:amin_def, erule disjE,
erule exE, simp add:amin_def, simp add:amin_def)
done
lemma amax_mem_m:"\<lbrakk>x \<in> Z\<^sub>-\<^sub>\<infinity>; y \<in> Z\<^sub>-\<^sub>\<infinity>\<rbrakk> \<Longrightarrow> amax x y \<in> Z\<^sub>-\<^sub>\<infinity>"
apply (cut_tac mem_ant[of "x"], cut_tac mem_ant[of "y"],
simp add:aug_minf_def)
apply ((erule disjE)+, simp add:amax_def,
erule exE, simp add:amax_def,
erule disjE, erule exE, simp add:amax_def)
apply ((erule exE)+, cut_tac amax_le[of "x" "y"],
rule contrapos_pp, simp+) apply (erule disjE,
cut_tac x = "ant z" in inf_ge_any,
frule_tac x = "ant z" in ale_antisym[of _ "\<infinity>"], assumption+, simp,
cut_tac x = "ant za" in inf_ge_any,
frule_tac x = "ant za" in ale_antisym[of _ "\<infinity>"], assumption+, simp)
done
lemma amin_commute:"amin x y = amin y x"
apply (cut_tac ale_linear[of "x" "y"], erule disjE, simp add:amin_def)
apply (simp add:amin_def)
done
lemma amin_mult_pos:"0 < z \<Longrightarrow> amin (z *\<^sub>a x) (z *\<^sub>a y) = z *\<^sub>a amin x y"
by (simp add:amin_def, simp add:asprod_pos_mono)
lemma amin_amult_pos:"0 < z \<Longrightarrow>
amin ((ant z) * x) ((ant z) * y) = (ant z) * amin x y"
by (simp add:asprod_amult[THEN sym], simp add:amin_mult_pos)
lemma times_amin:"\<lbrakk>0 < a; amin (x * (ant a)) (y * (ant a)) \<le> z * (ant a)\<rbrakk> \<Longrightarrow>
amin x y \<le> z"
by (frule amin_mult_pos[of "a" "x" "y"], simp add:asprod_amult,
simp add:amult_commute[of "ant a"], simp add:amult_pos_mono_r)
lemma Amin_memTr:"f \<in> {i. i \<le> n} \<rightarrow> Z\<^sub>\<infinity> \<longrightarrow> Amin n f \<in> Z\<^sub>\<infinity>"
apply (induct_tac n,
simp add:Pi_def)
apply (rule impI,
frule_tac func_pre[of "f" _ "Z\<^sub>\<infinity>"],
simp, rule amin_mem_i, assumption+,
simp add:Pi_def)
done
lemma Amin_mem:"f \<in> {i. i \<le> n} \<rightarrow> Z\<^sub>\<infinity> \<Longrightarrow> Amin n f \<in> Z\<^sub>\<infinity>"
apply (simp add:Amin_memTr)
done
lemma Amax_memTr:"f \<in> {i. i \<le> n} \<rightarrow> Z\<^sub>-\<^sub>\<infinity> \<longrightarrow> Amax n f \<in> Z\<^sub>-\<^sub>\<infinity>"
apply (induct_tac n,
simp add:Pi_def)
apply (rule impI,
frule_tac func_pre[of "f" _ "Z\<^sub>-\<^sub>\<infinity>"],
simp, rule amax_mem_m, assumption+,
simp add:Pi_def)
done
lemma Amax_mem:"f \<in> {i. i \<le> n} \<rightarrow> Z\<^sub>-\<^sub>\<infinity> \<Longrightarrow> Amax n f \<in> Z\<^sub>-\<^sub>\<infinity>"
apply (simp add:Amax_memTr)
done
lemma Amin_mem_mem:"\<forall>j\<le> n. f j \<in> Z\<^sub>\<infinity> \<Longrightarrow> Amin n f \<in> Z\<^sub>\<infinity>"
by (rule Amin_mem, simp)
lemma Amax_mem_mem:"\<forall>j \<le> n. f j \<in> Z\<^sub>-\<^sub>\<infinity> \<Longrightarrow> Amax n f \<in> Z\<^sub>-\<^sub>\<infinity>"
by (rule Amax_mem, simp)
lemma Amin_leTr:"f \<in> {i. i \<le> n} \<rightarrow> Z\<^sub>\<infinity> \<longrightarrow> (\<forall>j\<in>{i. i \<le> n}. Amin n f \<le> (f j))"
apply (induct_tac n,
rule impI, rule ballI,
simp)
apply (rule impI, rule ballI,
frule func_pre, simp)
apply (case_tac "j = Suc n", simp, rule amin_le_r)
apply (cut_tac x = j and n = n in Nset_pre, simp, assumption,
cut_tac x = "Amin n f" and y = "f (Suc n)" in amin_le_l,
thin_tac "j \<le> Suc n", simp)
apply (frule_tac x = j in spec,
thin_tac "\<forall>j\<le>n. Amin n f \<le> f j", simp)
done
lemma Amin_le:"\<lbrakk>f \<in> {j. j \<le> n} \<rightarrow> Z\<^sub>\<infinity>; j \<in> {k. k \<le> n}\<rbrakk> \<Longrightarrow> Amin n f \<le> (f j)"
apply (simp add:Amin_leTr)
done
lemma Amax_geTr:"f \<in> {j. j \<le> n} \<rightarrow> Z\<^sub>-\<^sub>\<infinity> \<longrightarrow> (\<forall>j\<in>{j. j \<le> n}. (f j) \<le> Amax n f)"
apply (induct_tac n,
rule impI, rule ballI,
simp)
apply (rule impI, rule ballI,
frule func_pre, simp,
case_tac "j = Suc n", simp, rule amax_ge_r,
cut_tac x = j and n = n in Nset_pre, simp, assumption,
thin_tac "j \<le> Suc n",
simp)
apply (cut_tac x = "Amax n f" and y = "f (Suc n)" in amax_ge_l,
drule_tac x = j in spec, simp)
done
lemma Amax_ge:"\<lbrakk>f \<in> {j. j \<le> n} \<rightarrow> Z\<^sub>-\<^sub>\<infinity>; j \<in> {j. j \<le> n}\<rbrakk> \<Longrightarrow>
(f j) \<le> (Amax n f)"
apply (simp add:Amax_geTr)
done
lemma Amin_mem_le:"\<lbrakk>\<forall>j \<le> n. (f j) \<in> Z\<^sub>\<infinity>; j \<in> {j. j \<le> n}\<rbrakk> \<Longrightarrow>
(Amin n f) \<le> (f j)"
by (rule Amin_le, simp, simp)
lemma Amax_mem_le:"\<lbrakk>\<forall>j \<le> n. (f j) \<in> Z\<^sub>-\<^sub>\<infinity>; j \<in> {j. j \<le> n}\<rbrakk> \<Longrightarrow>
(f j) \<le> (Amax n f)"
by (rule Amax_ge, simp, simp)
lemma amin_ge1:"\<lbrakk>(z::ant) \<le> x; z \<le> y \<rbrakk> \<Longrightarrow> z \<le> amin x y"
by (simp add:amin_def)
lemma amin_gt:"\<lbrakk>(z::ant) < x; z < y\<rbrakk> \<Longrightarrow> z < amin x y"
apply (simp add:less_ant_def, (erule conjE)+,
rule conjI, simp add:amin_ge1)
apply (rule contrapos_pp, simp+,
case_tac "x \<le> y", simp add:amin_def, simp add:amin_def)
done
lemma Amin_ge1Tr:"(\<forall>j\<le>(Suc n). (f j) \<in> Z\<^sub>\<infinity> \<and> z \<le> (f j)) \<longrightarrow>
z \<le> (Amin (Suc n) f)"
apply (induct_tac n)
apply (rule impI)
apply (frule_tac x = 0 in spec,
frule_tac x = "Suc 0" in spec,
thin_tac "\<forall>j\<le>Suc 0. f j \<in> Z\<^sub>\<infinity> \<and> z \<le> f j", simp, (erule conjE)+,
simp add:amin_ge1)
apply (rule impI,
simp,
frule_tac a = "Suc (Suc n)" in forall_spec,
thin_tac "\<forall>j\<le>Suc (Suc n). f j \<in> Z\<^sub>\<infinity> \<and> z \<le> f j", simp,
thin_tac "\<forall>j\<le>Suc (Suc n). f j \<in> Z\<^sub>\<infinity> \<and> z \<le> f j", erule conjE)
apply (rule amin_ge1, assumption+)
done
lemma Amin_ge1:"\<lbrakk> \<forall>j \<le> (Suc n). f j \<in> Z\<^sub>\<infinity>; \<forall>j \<le> (Suc n). z \<le> (f j)\<rbrakk> \<Longrightarrow>
z \<le> (Amin (Suc n) f)"
apply (simp del:Amin_Suc add:Amin_ge1Tr)
done
lemma amin_trans1:"\<lbrakk>x \<in> Z\<^sub>\<infinity>; y \<in> Z\<^sub>\<infinity>; z \<in> Z\<^sub>\<infinity>; z \<le> x \<rbrakk> \<Longrightarrow>
amin z y \<le> amin x y"
apply (case_tac "z \<le> y", simp add:amin_def)
apply (simp add:amin_def)
apply (simp only:aneg_le[of "z" "y"], frule aless_imp_le[of "y" "z"],
frule ale_trans[of "y" "z" "x"], assumption+, rule impI,
frule ale_antisym[of "y" "x"], assumption+)
done
lemma inf_in_aug_inf:"\<infinity> \<in> Z\<^sub>\<infinity>"
apply (simp add:aug_inf_def, simp add:not_sym)
done
subsection "Maximum element of a set of ants"
primrec aasc_seq :: "[ant set, ant, nat] \<Rightarrow> ant"
where
aasc_seq_0 : "aasc_seq A a 0 = a"
| aasc_seq_Suc : "aasc_seq A a (Suc n) =
(SOME b. ((b \<in> A) \<and> (aasc_seq A a n) < b))"
lemma aasc_seq_mem:"\<lbrakk>a \<in> A; \<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. x \<le> m))\<rbrakk> \<Longrightarrow>
(aasc_seq A a n) \<in> A"
apply (induct_tac n)
apply (simp, simp)
apply (simp add:aneg_le,
frule_tac a = "aasc_seq A a n" in forall_spec, assumption+,
thin_tac "\<forall>m. m \<in> A \<longrightarrow> (\<exists>x\<in>A. m < x)",
rule someI2_ex, blast, simp)
done
lemma aasc_seqn:"\<lbrakk>a \<in> A; \<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. x \<le> m))\<rbrakk> \<Longrightarrow>
(aasc_seq A a n) < (aasc_seq A a (Suc n))"
apply (frule aasc_seq_mem [of "a" "A" "n"], assumption+,
simp add:aneg_le,
frule_tac a = "aasc_seq A a n" in forall_spec, assumption+,
thin_tac "\<forall>m. m \<in> A \<longrightarrow> (\<exists>x\<in>A. m < x)", rule someI2_ex, blast, simp)
done
lemma aasc_seqn1:"\<lbrakk>a \<in> A; \<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. x \<le> m))\<rbrakk> \<Longrightarrow>
(aasc_seq A a n) + 1 \<le> (aasc_seq A a (Suc n))"
by (frule aasc_seqn [of "a" "A" "n"], assumption+, simp)
lemma aubs_ex_n_maxTr:"\<lbrakk>a \<in> A; \<not> (\<exists>m. m\<in>A \<and> (\<forall>x\<in>A. x \<le> m))\<rbrakk> \<Longrightarrow>
(a + an n) \<le> (aasc_seq A a n)"
apply (induct_tac n)
apply (simp add:aadd_0_r,
frule_tac n = n in aasc_seqn1[of "a" "A"], assumption+,
cut_tac x = "a + an n" and y = "aasc_seq A a n" in
aadd_le_mono[of _ _ "1"], assumption, simp,
frule_tac i = "a + an n + 1" and j = "aasc_seq A a n + 1" and
k = "(SOME b. b \<in> A \<and> aasc_seq A a n < b)" in ale_trans, assumption+)
apply (simp add:an_Suc,
case_tac "a = -\<infinity>",
subst ant_1[THEN sym], simp del:ant_1 add:a_zpz an_def,
subgoal_tac "a \<in> Z\<^sub>\<infinity>", subgoal_tac "an n \<in> Z\<^sub>\<infinity>",
subgoal_tac "1 \<in> Z\<^sub>\<infinity>",
subst aadd_assoc_i[THEN sym], assumption+)
apply (subst ant_1[THEN sym], simp del:ant_1 add:aug_inf_def,
(simp add:aug_inf_def an_def)+)
done
lemma aubs_ex_AMax:"\<lbrakk>A \<subseteq> UBset (ant z); A \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists>!m. m\<in>A \<and> (\<forall>x\<in>A. x \<le> m)"
apply (case_tac "A = {-\<infinity>}", simp,
frule not_sub_single[of "A" "-\<infinity>"], assumption+,
frule not_sub[of "A" "{-\<infinity>}"],
erule exE, erule conjE, simp, rename_tac a, rule ex_ex1I)
prefer 2
apply ((erule conjE)+,
frule_tac x = y in bspec, assumption+,
thin_tac "\<forall>x\<in>A. x \<le> m",
frule_tac x = m in bspec, assumption+,
thin_tac "\<forall>x\<in>A. x \<le> y", simp)
apply (rule contrapos_pp, simp,
subgoal_tac "\<exists>w. a = ant w", erule exE,
frule_tac a = a and A = A and n = "nat ((abs w) + (abs z) + 1)" in
aubs_ex_n_maxTr, simp,
frule_tac a = a and n = "nat ((abs w) + (abs z) + 1)" in
aasc_seq_mem[of _ "A"], assumption,
thin_tac "\<not> (\<exists>m. m \<in> A \<and> (\<forall>x\<in>A. x \<le> m))",
simp add:UBset_def)
apply (frule_tac c = "aasc_seq A (ant w) (nat (\<bar>w\<bar> + \<bar>z\<bar> + 1))" in
subsetD[of "A" "{x. x \<le> ant z}"], assumption+,
simp)
apply(frule_tac i = "ant w + an (nat (\<bar>w\<bar> + \<bar>z\<bar> + 1))" and
j = "aasc_seq A (ant w) (nat (\<bar>w\<bar> + \<bar>z\<bar> + 1))" and
k = "ant z" in ale_trans, assumption+)
apply (thin_tac "ant w + an (nat (\<bar>w\<bar> + \<bar>z\<bar> + 1))
\<le> aasc_seq A (ant w) (nat (\<bar>w\<bar> + \<bar>z\<bar> + 1))",
thin_tac "aasc_seq A (ant w) (nat (\<bar>w\<bar> + \<bar>z\<bar> + 1)) \<in> A",
thin_tac "aasc_seq A (ant w) (nat (\<bar>w\<bar> + \<bar>z\<bar> + 1)) \<le> ant z",
simp add:an_def a_zpz)
apply (cut_tac a = a in mem_ant, erule disjE, simp, erule disjE, erule exE,
simp, simp add:UBset_def, frule subsetD[of "A" "{x. x \<le> ant z}" "\<infinity>"],
assumption+, simp, cut_tac inf_ge_any[of "ant z"],
frule_tac ale_antisym[of "ant z" "\<infinity>"], assumption+, simp)
done
definition
AMax :: "ant set \<Rightarrow> ant" where
"AMax A = (THE m. m \<in> A \<and> (\<forall>x\<in>A. x \<le> m))"
definition
AMin::"ant set \<Rightarrow> ant" where
"AMin A = (THE m. m \<in> A \<and> (\<forall>x\<in>A. m \<le> x))"
definition
rev_o :: "ant \<Rightarrow> ant" where
"rev_o x = - x"
lemma AMax:"\<lbrakk>A \<subseteq> UBset (ant z); A \<noteq> {}\<rbrakk> \<Longrightarrow>
(AMax A) \<in> A \<and> (\<forall>x\<in>A. x \<le> (AMax A))"
apply (simp add:AMax_def)
apply (frule aubs_ex_AMax[of "A" "z"], assumption)
apply (rule theI')
apply assumption
done
lemma AMax_mem:"\<lbrakk>A \<subseteq> UBset (ant z); A \<noteq> {}\<rbrakk> \<Longrightarrow> (AMax A) \<in> A"
apply (simp add:AMax[of "A" "z"])
done
lemma rev_map_nonempty:"A \<noteq> {} \<Longrightarrow> rev_o ` A \<noteq> {}"
by (rule contrapos_pp, simp+)
lemma rev_map:"rev_o \<in> LBset (ant (-z)) \<rightarrow> UBset (ant z)"
by (rule Pi_I, simp add:UBset_def LBset_def rev_o_def,
frule_tac x = "ant (-z)" and y = x in ale_minus, simp add:aminus)
lemma albs_ex_AMin:"\<lbrakk>A \<subseteq> LBset (ant z); A \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists>!m. m\<in>A \<and> (\<forall>x\<in>A. m \<le> x)"
apply (rule ex_ex1I)
prefer 2 apply ((erule conjE)+,
frule_tac x = y in bspec, assumption+,
thin_tac "\<forall>x\<in>A. m \<le> x",
frule_tac x = m in bspec, assumption+,
thin_tac "\<forall>x\<in>A. y \<le> x", simp)
apply (subgoal_tac "- AMax (rev_o ` A) \<in> A \<and>
(\<forall>x \<in> A. (- AMax (rev_o ` A)) \<le> x)", blast,
cut_tac rev_map[of "-z"], simp add:a_minus_minus,
frule rev_map_nonempty[of "A"],
frule image_sub[of "rev_o" "LBset (ant z)" "UBset (ant (-z))" "A"],
assumption+, frule AMax[of "rev_o ` A" "-z"], assumption+,
erule conjE,
rule conjI, thin_tac "\<forall>x\<in>rev_o ` A. x \<le> AMax (rev_o ` A)",
thin_tac "rev_o \<in> LBset (ant z) \<rightarrow> UBset (ant (- z))",
thin_tac "rev_o ` A \<noteq> {}",
thin_tac "rev_o ` A \<subseteq> UBset (ant (- z))")
apply (simp add:image_def rev_o_def,
erule bexE, simp add:a_minus_minus, rule ballI,
subgoal_tac "rev_o x \<in> rev_o ` A",
frule_tac x = "rev_o x" in bspec, assumption+,
thin_tac "\<forall>x\<in>rev_o ` A. x \<le> AMax (rev_o ` A)",
thin_tac "rev_o \<in> LBset (ant z) \<rightarrow> UBset (ant (- z))",
thin_tac "rev_o ` A \<noteq> {}",
thin_tac "rev_o ` A \<subseteq> UBset (ant (- z))")
apply (simp add:image_def rev_o_def, erule bexE, simp add:a_minus_minus,
frule_tac x = "-x" and y = "-xa" in ale_minus, simp add:a_minus_minus,
simp add:image_def, blast)
done
lemma AMin:"\<lbrakk>A \<subseteq> LBset (ant z); A \<noteq> {}\<rbrakk> \<Longrightarrow>
(AMin A) \<in> A \<and> (\<forall>x\<in>A. (AMin A) \<le> x)"
apply (simp add:AMin_def)
apply (frule albs_ex_AMin[of "A" "z"], assumption)
apply (rule theI')
apply assumption
done
lemma AMin_mem:"\<lbrakk>A \<subseteq> LBset (ant z); A \<noteq> {}\<rbrakk> \<Longrightarrow> (AMin A) \<in> A"
apply (simp add:AMin)
done
primrec ASum :: "(nat \<Rightarrow> ant) \<Rightarrow> nat \<Rightarrow> ant"
where
ASum_0: "ASum f 0 = f 0"
| ASum_Suc: "ASum f (Suc n) = (ASum f n) + (f (Suc n))"
lemma age_plus:"\<lbrakk>0 \<le> (a::ant); 0 \<le> b; a + b \<le> c\<rbrakk> \<Longrightarrow> a \<le> c"
apply (frule aadd_le_mono[of "0" "b" "a"])
apply (simp add:aadd_commute[of "b" "a"] aadd_0_l)
done
lemma age_diff_le:"\<lbrakk>(a::ant) \<le> c; 0 \<le> b\<rbrakk> \<Longrightarrow> a - b \<le> c"
apply (frule ale_minus[of "0" "b"], thin_tac "0 \<le> b", simp)
apply (frule aadd_le_mono[of "a" "c" "-b"])
apply (frule aadd_le_mono[of "-b" "0" "c"])
apply (thin_tac "a \<le> c", thin_tac "- b \<le> 0",
simp add:aadd_commute[of "-b" "c"] aadd_0_l)
apply (simp add:diff_ant_def)
done
lemma adiff_le_adiff:"a \<le> (a'::ant) \<Longrightarrow> a - b \<le> a' - b"
apply (simp add:diff_ant_def)
apply (rule aadd_le_mono[of "a" "a'" "-b"], assumption+)
done
lemma aplus_le_aminus:"\<lbrakk> a \<in> Z\<^sub>-\<^sub>\<infinity>; b \<in> Z\<^sub>-\<^sub>\<infinity>; c \<in> Z\<^sub>-\<^sub>\<infinity>; -b \<in> Z\<^sub>-\<^sub>\<infinity>\<rbrakk> \<Longrightarrow>
((a + b) \<le> (c::ant)) = (a \<le> c - b)"
apply (rule iffI)
apply (frule aadd_le_mono[of "a + b" "c" "-b"])
apply (simp add:aadd_assoc_m, simp add:aadd_minus_r)
apply (simp add:aadd_0_r, simp add:diff_ant_def)
apply (frule aadd_le_mono[of "a" "c - b" "b"])
apply (simp add:diff_ant_def)
apply (simp add:aadd_assoc_m)
apply (simp add:aadd_minus_inv[of "b"])
apply (simp add: aadd_0_r)
done
section "Cardinality of sets"
text \<open>cardinality is defined for the finite sets only\<close>
lemma card_eq:"A = B \<Longrightarrow> card A = card B"
apply simp
done
lemma card0:"card {} = 0"
by simp
lemma card_nonzero:"\<lbrakk>finite A; card A \<noteq> 0\<rbrakk> \<Longrightarrow> A \<noteq> {}"
by (rule contrapos_pp, simp+)
lemma finite1:"finite {a}"
by simp
lemma card1:"card {a} = 1"
by simp
lemma nonempty_card_pos:"\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> 0 < card A"
apply (frule nonempty_ex [of "A"], erule exE,
frule_tac a = x and A = A in singleton_sub)
apply (frule_tac B = A and A = "{x}" in card_mono, assumption+,
simp add:card1)
done
lemma nonempty_card_pos1:"\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> Suc 0 \<le> card A"
apply (frule nonempty_card_pos[of "A"], assumption+)
apply (rule Suc_leI[of "0" "card A"], assumption)
done
lemma card1_tr0:"\<lbrakk> finite A; card A = Suc 0; a \<in> A \<rbrakk> \<Longrightarrow> {a} = A"
apply (cut_tac card1[of "a"])
apply (rule card_seteq[of "A" "{a}"], assumption)
apply (rule singleton_sub[of "a" "A"], assumption)
apply simp
done
lemma card1_tr1:"(constmap {0::nat} {x}) \<in> {0} \<rightarrow> {x} \<and>
surj_to (constmap {0::nat} {x}) {0} {x}"
apply (rule conjI, simp add:constmap_def Pi_def,
simp add:surj_to_def image_def constmap_def)
done
lemma card1_Tr2:"\<lbrakk>finite A; card A = Suc 0\<rbrakk> \<Longrightarrow>
\<exists>f. f \<in> {0::nat} \<rightarrow> A \<and> surj_to f {0} A"
apply (frule card_nonzero[of "A"], simp)
apply (cut_tac nonempty_ex[of "A"], erule exE)
apply (frule_tac a = x in card1_tr0[of "A"], assumption+)
apply (rotate_tac -1, frule sym, thin_tac "{x} = A", simp)
apply (cut_tac x = x in card1_tr1, blast, simp)
done
lemma card2:"\<lbrakk> finite A; a \<in> A; b \<in> A; a \<noteq> b \<rbrakk> \<Longrightarrow> Suc (Suc 0) \<le> card A"
apply (cut_tac card1[of "a"])
apply (frule singleton_sub[of "b" "A"])
apply (frule finite_subset[of "{b}" "A"], assumption)
apply (frule card_insert_disjoint[of "{b}" "a"])
apply simp
apply (simp only:card1)
apply (frule insert_sub[of "{b}" "A" "a"], assumption+)
apply (frule card_mono [of "A" "{a, b}"], assumption)
apply simp
done
lemma card2_inc_two:"\<lbrakk>0 < (n::nat); x \<in> {j. j \<le> n}\<rbrakk> \<Longrightarrow>
\<exists>y \<in> {j. j \<le> n}. x \<noteq> y"
apply (rule contrapos_pp, simp+)
apply (frule_tac m = 0 and n = n in Suc_leI) apply (
frule_tac a = "Suc 0" in forall_spec, assumption)
apply (frule_tac a = 0 in forall_spec)
apply (rule less_imp_le, assumption)
apply simp
done
lemma Nset2_prep1:"\<lbrakk>finite A; card A = Suc (Suc n) \<rbrakk> \<Longrightarrow> \<exists>x. x\<in>A"
apply (frule card_nonzero[of "A"])
apply simp
apply (simp add:nonempty_ex)
done
lemma ex_least_set:"\<lbrakk>A = {H. finite H \<and> P H}; H \<in> A\<rbrakk> \<Longrightarrow>
\<exists>K \<in> A. (LEAST j. j \<in> (card ` A)) = card K"
(* proof by L. C. Paulson *)
by (simp add:image_def, rule LeastI, rule_tac x = "H" in exI, simp)
lemma Nset2_prep2:"x \<in> A \<Longrightarrow> A - {x} \<union> {x} = A"
by auto
lemma Nset2_finiteTr:"\<forall>A. (finite A \<and>(card A = Suc n) \<longrightarrow>
(\<exists>f. f \<in> {i. i \<le> n} \<rightarrow> A \<and> surj_to f {i. i \<le> n} A))"
apply (induct_tac n, rule allI, rule impI, erule conjE)
apply (simp add: card1_Tr2 del: Pi_split_insert_domain)
(* n *)
apply (rule allI, rule impI, erule conjE, frule Nset2_prep1, assumption+)
apply (erule exE)
apply(drule_tac a = "A - {x}" in forall_spec)
apply simp
apply (erule exE)
apply (cut_tac x = x in card1_tr1, (erule conjE)+)
apply (frule_tac f = f and n = n and A = "A - {x}" and
g = "constmap {0} {x}" and m = 0 and B = "{x}" in jointfun_surj,
assumption+)
apply simp+
apply (frule_tac f = f and n = n and A = "A - {x}" and
g = "constmap {0} {x}" and m = 0 and B = "{x}" in jointfun_hom0,
simp,
frule_tac x = x and A = A in Nset2_prep2, simp, blast)
done
lemma Nset2_finite:"\<lbrakk> finite A; card A = Suc n\<rbrakk> \<Longrightarrow>
\<exists>f. f \<in> {i. i \<le> n} \<rightarrow> A \<and> surj_to f {i. i \<le> n} A "
by (simp add:Nset2_finiteTr)
lemma Nset2finite_inj_tr0:"j \<in> {i. i \<le> (n::nat)} \<Longrightarrow>
card ({i. i \<le> n} - {j}) = n"
by simp
lemma Nset2finite_inj_tr1:"\<lbrakk> i \<le> (n::nat); j \<le> n; f i = f j; i \<noteq> j \<rbrakk> \<Longrightarrow>
f ` ({i. i \<le> n} - {j}) = f ` {i. i \<le> n}"
apply (simp add:image_def, rule equalityI, rule subsetI, simp add:CollectI,
erule bexE, case_tac "xa = j", frule sym, thin_tac "f i = f j",
simp, blast)
apply (rule subsetI, simp, erule exE, case_tac "xa = j", frule sym,
thin_tac "f i = f j", blast, blast)
done
lemma Nset2finite_inj:"\<lbrakk>finite A; card A = Suc n; surj_to f {i. i \<le> n} A \<rbrakk> \<Longrightarrow>
inj_on f {i. i \<le> n}"
by (metis card_Collect_le_nat eq_card_imp_inj_on finite_Collect_le_nat surj_to_def)
definition
zmax :: "[int, int] \<Rightarrow> int" where
"zmax x y = (if (x \<le> y) then y else x)"
primrec Zmax :: "[nat, nat \<Rightarrow> int] \<Rightarrow> int"
where
Zmax_0 : "Zmax 0 f = f 0"
| Zmax_Suc :"Zmax (Suc n) f = zmax (Zmax n f) (f (Suc n))"
lemma Zmax_memTr:"f \<in> {i. i \<le> (n::nat)} \<rightarrow> (UNIV::int set) \<longrightarrow>
Zmax n f \<in> f ` {i. i \<le> n}"
apply (induct_tac n)
apply simp
apply (rule impI)
apply (frule func_pre)
apply (frule_tac f = f and A = "{i. i \<le> Suc n}" and B = UNIV and
?A1.0 = "{i. i \<le> n}" and ?A2.0 = "{i. i \<le> Suc n}" in im_set_mono)
apply (rule subsetI, simp, simp, simp)
apply (case_tac "(Zmax n f) \<le> (f (Suc n))", simp add:zmax_def)
apply (simp add:zmax_def)
apply (simp add:subsetD)
done
lemma zmax_ge_r:"y \<le> zmax x y"
by (simp add:zmax_def)
lemma zmax_ge_l:"x \<le> zmax x y"
by (simp add:zmax_def)
lemma Zmax_geTr:"f \<in> {j. j \<le> (n::nat)} \<rightarrow> (UNIV::int set) \<longrightarrow>
(\<forall>j\<in>{j. j \<le> n}. (f j) \<le> Zmax n f)"
apply (induct_tac n,
rule impI, rule ballI,
simp)
apply (rule impI, rule ballI,
frule func_pre, simp,
case_tac "j = Suc n", simp, rule zmax_ge_r,
cut_tac x = j and n = n in Nset_pre, simp, assumption,
thin_tac "j \<le> Suc n",
simp)
apply (cut_tac x = "Zmax n f" and y = "f (Suc n)" in zmax_ge_l,
frule_tac x = j in spec,
thin_tac "\<forall>j\<le>n. f j \<le> Zmax n f")
apply simp
done
lemma Zmax_plus1:"f \<in> {j. j \<le> (n::nat)} \<rightarrow> (UNIV::int set) \<Longrightarrow>
((Zmax n f) + 1) \<notin> f ` {j. j \<le> n}"
apply (cut_tac Zmax_geTr[of f n])
apply (rule contrapos_pp, simp+)
apply (simp add:image_def, erule exE, erule conjE)
apply (frule_tac a = x in forall_spec, assumption,
thin_tac "\<forall>j\<le>n. f j \<le> Zmax n f")
apply (frule sym, thin_tac "Zmax n f + 1 = f x", simp)
done
lemma image_Nsetn_card_pos:" 0 < card (f ` {i. i \<le> (n::nat)})"
apply(rule nonempty_card_pos)
apply auto
done
lemma card_image_Nsetn_Suc
:"\<lbrakk>f \<in> {j. j \<le> Suc n} \<rightarrow> B;
f (Suc n) \<notin> f ` {j. j \<le> n}\<rbrakk> \<Longrightarrow>
card (f ` {j. j \<le> Suc n}) - Suc 0 =
Suc (card (f ` {j. j \<le> n}) - Suc 0)"
apply (simp add:image_Nset_Suc)
apply (cut_tac image_Nsetn_card_pos[of f n], simp)
done
-lemma slide_surj:"i < (j::nat) \<Longrightarrow>
- surj_to (slide i) {l. l \<le> (j - i)} (nset i j)"
-proof -
- assume p1:"i < j"
- from p1 show ?thesis
- apply (simp add:surj_to_def image_def)
- apply (rule equalityI,
- rule subsetI, simp, erule exE, simp add:slide_def nset_def,
- frule less_imp_le [of i j], erule conjE,
- thin_tac "i < j", frule add_le_mono [of _ "j - i" "i" "i"],
- simp+, rule subsetI, simp)
- apply (simp add:nset_def slide_def, erule conjE,
- frule_tac m = x and n = j and l = i in diff_le_mono,
- subgoal_tac "x = i + (x - i)", blast, simp)
- done
-qed
+lemma slide_surj:
+ \<open>surj_to (slide i) {l. l \<le> (j - i)} (nset i j)\<close> if \<open>i < j\<close> for i j :: nat
+ using that
+ by (auto simp add: surj_to_def image_def slide_def nset_def) presburger
lemma slide_inj:"i < j \<Longrightarrow> inj_on (slide i) {k. k \<le> (j - i)}"
apply (simp add:inj_on_def, (rule allI)+)
apply (rule impI, rule allI, rule impI, rule impI)
apply (simp add:slide_def)
done
lemma card_nset:"i < (j :: nat) \<Longrightarrow> card (nset i j) = Suc (j - i)"
apply (frule slide_inj [of "i" "j"])
apply (frule card_image [of "slide i" "{k. k \<le> (j - i)}"])
apply (simp, frule slide_surj [of "i" "j"], simp add:surj_to_def)
done
lemma sliden_hom:"i < j \<Longrightarrow> (sliden i) \<in> nset i j \<rightarrow> {k. k \<le> (j - i)}"
by (simp add:Pi_def, rule allI, rule impI, simp add:sliden_def,
simp add:nset_def, erule conjE, simp add:diff_le_mono)
lemma slide_sliden:"(sliden i) (slide i k) = k"
by (simp add:sliden_def slide_def)
lemma sliden_surj:"i < j \<Longrightarrow> surj_to (sliden i) (nset i j) {k. k \<le> (j - i)}"
apply (simp add:surj_to_def image_def, rule equalityI)
apply (rule subsetI, simp, erule bexE, simp add:nset_def sliden_def,
erule conjE, rule_tac m = xa in diff_le_mono[of _ "j" "i"],
assumption+)
apply (rule subsetI, simp add:nset_def sliden_def,
frule_tac i = x in add_le_mono[of _ "j - i" "i" "i"], simp,
simp, subgoal_tac "i \<le> x + i", subgoal_tac "x = (x + i) - i",
blast) apply simp+
done
lemma sliden_inj: "i < j \<Longrightarrow> inj_on (sliden i) (nset i j)"
apply (simp add:inj_on_def, (rule ballI)+, rule impI, simp add:sliden_def)
apply (simp add:nset_def, (erule conjE)+,
subgoal_tac "(x - i = y - i) = (x = y)", blast)
apply (rule eq_diff_iff, assumption+)
done
definition
transpos :: "[nat, nat] \<Rightarrow> (nat \<Rightarrow> nat)" where
"transpos i j = (\<lambda>k. if k = i then j else if k = j then i else k)"
lemma transpos_id:"\<lbrakk> i \<le> n; j \<le> n; i \<noteq> j ; x \<in> {k. k \<le> n} - {i, j} \<rbrakk>
\<Longrightarrow> transpos i j x = x"
proof -
assume p1:"i \<le> n" and p2:"j \<le> n" and p3:" i \<noteq> j" and
p4:"x \<in> {k. k \<le> n} - {i, j}"
from p1 and p2 and p3 and p4 show ?thesis
apply (simp add:transpos_def)
done
qed
lemma transpos_id_1:"\<lbrakk>i \<le> n; j \<le> n; i \<noteq> j; x \<le> n; x \<noteq> i; x \<noteq> j\<rbrakk> \<Longrightarrow>
transpos i j x = x"
proof -
assume p1:"i \<le> n" and p2:"j \<le> n" and p3:"i \<noteq> j" and p4:"x \<le> n" and p5:"x \<noteq> i" and p6:"x \<noteq> j"
from p1 and p2 and p3 and p4 and p5 and p6 show ?thesis
apply (simp add:transpos_def)
done
qed
lemma transpos_id_2:"i \<le> n \<Longrightarrow> transpos i n (Suc n) = Suc n"
by (simp add:transpos_def)
lemma transpos_ij_1:"\<lbrakk>i \<le> n; j \<le> n; i \<noteq> j \<rbrakk> \<Longrightarrow>
transpos i j i = j"
by (simp add:transpos_def)
lemma transpos_ij_2:"\<lbrakk>i \<le> n; j \<le> n; i \<noteq> j\<rbrakk> \<Longrightarrow> transpos i j j = i"
by (simp add:transpos_def)
lemma transpos_hom:"\<lbrakk>i \<le> n; j \<le> n; i \<noteq> j\<rbrakk> \<Longrightarrow>
(transpos i j) \<in> {i. i \<le> n} \<rightarrow> {i. i \<le> n}"
apply (simp add:Pi_def, rule allI, rule impI)
apply (case_tac "x = i", simp add:transpos_def)
apply (case_tac "x = j", simp add:transpos_def,
subst transpos_id, assumption+, simp, assumption)
done
lemma transpos_mem:"\<lbrakk>i \<le> n; j \<le> n; i \<noteq> j; l \<le> n\<rbrakk> \<Longrightarrow>
(transpos i j l) \<le> n"
apply (frule transpos_hom [of "i" "n" "j"], assumption+,
cut_tac funcset_mem[of "transpos i j" "{i. i \<le> n}" "{i. i \<le> n}" l])
apply simp+
done
lemma transpos_inj:"\<lbrakk>i \<le> n; j \<le> n; i \<noteq> j\<rbrakk>
\<Longrightarrow> inj_on (transpos i j) {i. i \<le> n}"
apply (simp add:inj_on_def, (rule allI, rule impI)+, rule impI,
case_tac "x = i", case_tac "y = j",
simp add:transpos_def)
apply (simp add:transpos_ij_1, rule contrapos_pp, simp+,
frule_tac x = y in transpos_id [of "i" "n" "j"], assumption+,
simp+)
apply (case_tac "x = j", simp,
simp add:transpos_ij_2, rule contrapos_pp, simp+,
frule_tac x = y in transpos_id [of "i" "n" "j"], assumption+,
simp, rule contrapos_pp, simp+, simp add:transpos_ij_1)
apply (simp, simp add:transpos_ij_1, simp add:transpos_id_1,
thin_tac "x = transpos i j y",
case_tac "y = i", simp add:transpos_ij_1,
case_tac "y = j", simp add:transpos_ij_2)
apply (simp add:transpos_id_1)
done
lemma transpos_surjec:"\<lbrakk>i \<le> n; j \<le> n; i \<noteq> j\<rbrakk>
\<Longrightarrow> surj_to (transpos i j) {i. i \<le> n} {i. i \<le> n}"
apply (simp add:surj_to_def,
frule transpos_hom [of "i" "n" "j"], assumption+,
frule image_sub [of "transpos i j" "{i. i \<le> n}" "{i. i \<le> n}"
"{i. i \<le> n}"], simp)
apply (frule transpos_inj [of "i" "n" "j"], assumption+,
frule card_image [of "transpos i j" "{i. i \<le> n}"],
simp add:card_seteq)
done
lemma comp_transpos:"\<lbrakk>i \<le> n; j \<le> n; i \<noteq> j\<rbrakk> \<Longrightarrow>
\<forall>k \<le> n. (compose {i. i \<le> n} (transpos i j) (transpos i j)) k = k"
proof -
assume p1:"i \<le> n" and p2:"j \<le> n" and p3:"i \<noteq> j"
from p1 and p2 and p3 show ?thesis
apply (simp add:compose_def)
apply (rule allI)
apply (case_tac "k = i") apply simp
apply (subst transpos_ij_1, assumption+)
apply (rule transpos_ij_2, simp+)
apply (rule impI)
apply (case_tac "k = j") apply simp
apply (subst transpos_ij_2, simp+)
apply (rule transpos_ij_1, simp+)
apply (subst transpos_id_1, assumption+)
apply (simp add:transpos_mem)
apply (simp add:transpos_id_1)+
done
qed
lemma comp_transpos_1:"\<lbrakk>i \<le> n; j \<le> n; i \<noteq> j; k \<le> n\<rbrakk> \<Longrightarrow>
(transpos i j) ((transpos i j) k) = k"
apply (frule comp_transpos [of "i" "n" "j"], assumption+)
apply (simp add:compose_def)
done
lemma cmp_transpos1:"\<lbrakk>i \<le> n; j \<le> n; i \<noteq> j; k \<le> n\<rbrakk> \<Longrightarrow>
(cmp (transpos i j) (transpos i j)) k = k"
apply (simp add:cmp_def)
apply (simp add:comp_transpos_1)
done
lemma cmp_transpos:"\<lbrakk>i \<le> n; i \<noteq> n; a \<le> (Suc n)\<rbrakk> \<Longrightarrow>
(cmp (transpos i n) (cmp (transpos n (Suc n)) (transpos i n))) a =
transpos i (Suc n) a"
apply (simp add:cmp_def)
apply (case_tac "a = Suc n", simp)
apply (simp add:transpos_id_2)
apply (cut_tac transpos_ij_2[of n "Suc n" "Suc n"], simp,
cut_tac transpos_ij_2[of i "Suc n" "Suc n"], simp,
cut_tac transpos_ij_2[of i n n], simp+)
apply (frule le_imp_less_or_eq[of a "Suc n"],
thin_tac "a \<le> Suc n", simp,
frule Suc_less_le[of a n])
apply (case_tac "a = n", simp,
cut_tac transpos_ij_2[of i n n], simp,
cut_tac transpos_id[of i "Suc n" "Suc n" n], simp,
cut_tac transpos_id[of n "Suc n" "Suc n" i], simp,
cut_tac transpos_ij_1[of i n n], simp+)
apply (case_tac "a = i", simp,
cut_tac transpos_ij_1[of i n n], simp+,
cut_tac transpos_ij_1[of i "Suc n" "Suc n"], simp,
cut_tac transpos_ij_1[of n "Suc n" "Suc n"], simp,
cut_tac transpos_id[of i "Suc n" n "Suc n"], simp+)
apply (cut_tac transpos_id[of i n n a], simp,
cut_tac transpos_id[of i "Suc n" "Suc n" a], simp,
cut_tac transpos_id[of n "Suc n" "Suc n" a], simp+)
done
lemma im_Nset_Suc:"insert (f (Suc n)) (f ` {i. i \<le> n}) = f ` {i. i\<le>(Suc n)}"
apply (simp add:image_def)
apply (rule equalityI)
apply (rule subsetI, simp)
apply (erule disjE, blast)
apply (erule exE, erule conjE, simp,
frule_tac i = xa and j = n and k = "Suc n" in le_trans,
simp)
apply blast
apply (rule subsetI, simp, erule exE, erule conjE)
apply (case_tac "xa = Suc n", simp)
apply (metis le_SucE)
done
lemma Nset_injTr0:"\<lbrakk>f \<in> {i. i \<le> (Suc n)} \<rightarrow> {i. i \<le> (Suc n)};
inj_on f {i. i \<le> (Suc n)}; f (Suc n) = Suc n\<rbrakk> \<Longrightarrow>
f \<in> {i. i \<le> n} \<rightarrow> {i. i \<le> n} \<and> inj_on f {i. i \<le> n}"
proof -
assume p1:"f \<in> {i. i \<le> (Suc n)} \<rightarrow> {i. i \<le> (Suc n)}" and
p2:"inj_on f {i. i \<le> (Suc n)}" and p3:"f (Suc n) = Suc n"
have q1:"\<forall>l \<le> n. l \<le> (Suc n)" apply simp done
from p1 and p2 and p3 and q1 have q2:"f \<in> {i. i \<le> n} \<rightarrow> {i. i \<le> n}"
apply (simp add:Pi_def)
apply (rule allI, rule impI)
apply (frule_tac a = x in forall_spec, simp,
thin_tac "\<forall>x\<le>Suc n. f x \<le> Suc n")
apply (rule contrapos_pp, simp+)
apply (simp add:nat_not_le_less)
apply (frule_tac n = "f x" in Suc_leI[of n], thin_tac "n < (f x)")
apply (frule_tac m = "Suc n" and n = "f x" in le_antisym, assumption)
apply(unfold inj_on_def)
apply (frule_tac x = x in bspec, simp,
thin_tac "\<forall>x\<in>{i. i \<le> Suc n}. \<forall>y\<in>{i. i \<le> Suc n}. f x = f y \<longrightarrow> x = y",
frule_tac x = "Suc n" in bspec, simp)
apply (frule_tac r = "f (Suc n)" and s = "Suc n" and t = "f x" in trans,
assumption,
thin_tac "f (Suc n) = Suc n", thin_tac "Suc n = f x",
thin_tac "\<forall>y\<in>{i. i \<le> Suc n}. f x = f y \<longrightarrow> x = y")
apply simp
done
from p2 have q3:"inj_on f {i. i \<le> n}"
apply (simp add:inj_on_def) done
from q2 and q3 show ?thesis apply simp done
qed
lemma inj_surj:"\<lbrakk>f \<in> {i. i \<le> (n::nat)} \<rightarrow> {i. i \<le> n};
inj_on f {i. i \<le> (n::nat)}\<rbrakk> \<Longrightarrow> f ` {i. i \<le> n} = {i. i \<le> n}"
proof -
assume p1:"f \<in> {i. i \<le> n} \<rightarrow> {i. i \<le> n}" and p2:"inj_on f {i. i \<le> n}"
have q1:"0 < Suc 0" apply simp done
from p1 and p2 and q1 show ?thesis
apply simp
apply (frule image_sub [of "f" "{i. i \<le> n}" "{i. i \<le> n}" "{i. i \<le> n}"])
apply simp+
apply (cut_tac card_image [of "f" "{i. i \<le> n}"])
apply (simp add:card_seteq) apply assumption
done
qed
lemma Nset_pre_mem:"\<lbrakk>f:{i. i\<le>(Suc n)} \<rightarrow>{i. i\<le>(Suc n)};
inj_on f {i. i\<le>(Suc n)}; f (Suc n) = Suc n; k \<le> n\<rbrakk> \<Longrightarrow> f k \<in> {i. i\<le>n}"
apply (frule Nset_injTr0[of f n], assumption+, erule conjE)
apply (frule_tac x = k in funcset_mem[of f "{i. i \<le> n}" "{i. i \<le> n}"],
simp, assumption)
done
lemma Nset_injTr1:"\<lbrakk> \<forall>l \<le>(Suc n). f l \<le> (Suc n); inj_on f {i. i \<le> (Suc n)};
f (Suc n) = Suc n \<rbrakk> \<Longrightarrow> inj_on f {i. i \<le> n}"
by (cut_tac Nset_injTr0[of f n], simp, simp)
lemma Nset_injTr2:"\<lbrakk>\<forall>l\<le> (Suc n). f l \<le> (Suc n); inj_on f {i. i \<le> (Suc n)};
f (Suc n) = Suc n\<rbrakk> \<Longrightarrow> \<forall>l \<le> n. f l \<le> n"
apply (rule allI, rule impI)
apply (cut_tac k = l in Nset_pre_mem[of f n])
apply (simp+)
done
lemma TR_inj_inj:"\<lbrakk>\<forall>l\<le> (Suc n). f l \<le> (Suc n); inj_on f {i. i \<le> (Suc n)};
i \<le> (Suc n); j \<le> (Suc n); i < j \<rbrakk> \<Longrightarrow>
inj_on (compose {i. i \<le> (Suc n)} (transpos i j) f) {i. i \<le> (Suc n)}"
apply (frule transpos_inj[of i "Suc n" j], assumption+,
simp )
apply (rule comp_inj [of f "{i. i \<le> (Suc n)}" "{i. i \<le> (Suc n)}"
"transpos i j" "{i. i \<le> (Suc n)}"])
apply (simp, assumption,
rule transpos_hom[of i "Suc n" j], simp+)
done
lemma enumeration:"\<lbrakk>f \<in> {i. i \<le> (n::nat)} \<rightarrow> {i. i \<le> m}; inj_on f {i. i \<le> n}\<rbrakk>
\<Longrightarrow> n \<le> m"
apply (frule image_sub[of f "{i. i \<le> n}" "{i. i \<le> m}" "{i. i \<le> n}"])
apply simp
apply (frule card_image[of f "{i. i \<le> n}"])
apply(drule card_mono[OF finite_Collect_le_nat])
apply simp
done
lemma enumerate_1:"\<lbrakk>\<forall>j \<le> (n::nat). f j \<in> A; \<forall>j \<le> (m::nat). g j \<in> A;
inj_on f {i. i \<le> n}; inj_on g {j. j \<le> m}; f `{j. j \<le> n} = A;
g ` {j. j \<le> m} = A \<rbrakk> \<Longrightarrow> n = m"
apply (frule card_image[of f "{i. i \<le> n}"],
frule card_image[of g "{i. i \<le> m}"])
apply simp
done
definition
ninv :: "[nat, (nat \<Rightarrow> nat)] \<Rightarrow> (nat \<Rightarrow> nat)" where
"ninv n f = (\<lambda>y\<in>{i. i \<le> n}. (SOME x. (x \<le> n \<and> y = f x)))"
lemma ninv_hom:"\<lbrakk>f \<in> {i. i \<le> n} \<rightarrow> {i. i \<le> n}; inj_on f {i. i \<le> n}\<rbrakk> \<Longrightarrow>
ninv n f \<in> {i. i \<le> n} \<rightarrow> {i. i \<le> n}"
apply (rule Pi_I)
apply (simp add:ninv_def)
apply (frule inj_surj[of f n], assumption+,
frule_tac x = x in funcset_mem[of f "{i. i \<le> n}" "{i. i \<le> n}"],
simp)
apply (frule sym, thin_tac "f ` {i. i \<le> n} = {i. i \<le> n}",
cut_tac a = x and A = "{i. i \<le> n}" and B = "f ` {i. i \<le> n}" in
eq_set_inc, simp, assumption,
thin_tac "f x \<in> {i. i \<le> n}", thin_tac "{i. i \<le> n} = f ` {i. i \<le> n}",
simp add:image_def, rule someI2_ex)
apply blast+
done
lemma ninv_r_inv:"\<lbrakk>f \<in> {i. i \<le> (n::nat)} \<rightarrow> {i. i \<le> n}; inj_on f {i. i \<le> n};
b \<le> n\<rbrakk> \<Longrightarrow> f (ninv n f b) = b "
apply (simp add:ninv_def)
apply (frule inj_surj, assumption+)
apply (cut_tac a = b in eq_set_inc[of _ "{i. i \<le> n}" "f ` {i. i \<le> n}"])
apply (simp, rule sym, assumption)
apply (thin_tac "f ` {i. i \<le> n} = {i. i \<le> n}", simp add:image_def,
erule exE, erule conjE, frule sym, thin_tac "b = f x")
apply (rule someI2_ex, blast)
apply (erule conjE, rule sym, assumption)
done
lemma ninv_inj:"\<lbrakk>f \<in> {i. i \<le> n} \<rightarrow> {i. i \<le> n}; inj_on f {i. i \<le> n}\<rbrakk> \<Longrightarrow>
inj_on (ninv n f) {i. i \<le> n}"
apply (subst inj_on_def, simp)
apply ((rule allI, rule impI)+, rule impI)
apply (frule ninv_hom[of f n], assumption,
frule_tac x = x in funcset_mem[of "ninv n f" "{i. i \<le> n}" "{i. i \<le> n}"], simp,
frule_tac x = y in funcset_mem[of "ninv n f" "{i. i \<le> n}" "{i. i \<le> n}"],
simp,
frule_tac b = x in ninv_r_inv [of f n], assumption+)
apply (simp add:ninv_r_inv)
done
subsection "Lemmas required in Algebra6.thy"
lemma ge2_zmult_pos:
"2 \<le> m \<Longrightarrow> 0 < z \<Longrightarrow> 1 < int m * z"
proof -
assume a1: "0 < z"
assume a2: "2 \<le> m"
have "int m + - 1 * (int m * z) \<le> 0"
using a1 by (simp add: pos_zmult_pos)
then show ?thesis
using a2 by linarith
qed
lemma zmult_pos_mono:"\<lbrakk> (0::int) < w; w * z \<le> w * z'\<rbrakk> \<Longrightarrow> z \<le> z'"
apply (rule contrapos_pp, simp+)
done
lemma zmult_pos_mono_r:
"\<lbrakk>(0::int) < w; z * w \<le> z' * w\<rbrakk> \<Longrightarrow> z \<le> z'"
apply (simp add:mult.commute)
done
lemma an_neq_inf:"an n \<noteq> \<infinity>"
by (simp add:an_def)
lemma an_neq_minf:"an n \<noteq> -\<infinity>"
by (simp add:an_def)
lemma aeq_mult:"\<lbrakk>z \<noteq> 0; a = b\<rbrakk> \<Longrightarrow> a * ant z = b * ant z"
by simp
lemma tna_0[simp]:"tna 0 = 0"
by (simp add:ant_0[THEN sym] tna_ant)
lemma ale_nat_le:"(an n \<le> an m) = (n \<le> m)"
by (simp add:an_def)
lemma aless_nat_less:"(an n < an m) = (n < m)"
by (simp add:an_def, subst aless_zless[of "int n" "int m"], simp)
lemma apos_natpos:"\<lbrakk>a \<noteq> \<infinity>; 0 \<le> a\<rbrakk> \<Longrightarrow> 0 \<le> na a"
by (cut_tac ale_nat_le[of "0" "na a"], simp add:na_def an_def)
lemma apos_tna_pos:"\<lbrakk>n \<noteq> \<infinity>; 0 \<le> n\<rbrakk> \<Longrightarrow> 0 \<le> tna n"
by (subst tna_0[THEN sym],
subst ale_zle[THEN sym, of "tna 0" "tna n"],
frule apos_neq_minf[of "n"],
simp add:ant_tna ant_0)
lemma apos_na_pos:"\<lbrakk>n \<noteq> \<infinity>; 0 \<le> n\<rbrakk> \<Longrightarrow> 0 \<le> na n"
by (frule apos_tna_pos[of "n"], assumption,
cut_tac tna_0[THEN sym], simp del:tna_0)
lemma aposs_tna_poss:"\<lbrakk>n \<noteq> \<infinity>; 0 < n\<rbrakk> \<Longrightarrow> 0 < tna n"
apply (subst tna_0[THEN sym],
subst aless_zless[THEN sym, of "tna 0" "tna n"],
frule aless_imp_le[of "0" "n"],
frule apos_neq_minf[of "n"],
simp add:ant_tna ant_0)
done
lemma aposs_na_poss:"\<lbrakk>n \<noteq> \<infinity>; 0 < n\<rbrakk> \<Longrightarrow> 0 < na n"
apply (frule aless_imp_le[of "0" "n"],
simp add:aneg_less[THEN sym, of "0" "n"],
simp add:na_def)
apply (rule aposs_tna_poss, assumption+)
done
lemma nat_0_le: "0 \<le> z ==> int (nat z) = z"
apply simp
done
lemma int_eq:"m = n \<Longrightarrow> int m = int n"
by simp
lemma box_equation:"\<lbrakk>a = b; a = c\<rbrakk> \<Longrightarrow> b = c"
apply simp
done
lemma aeq_nat_eq:"\<lbrakk>n \<noteq> \<infinity>; 0 \<le> n; m \<noteq> \<infinity>; 0 \<le> m\<rbrakk> \<Longrightarrow>
(n = m) = (na n = na m)"
apply (rule iffI, simp)
apply (cut_tac aneg_less[THEN sym, of "0" "n"],
cut_tac aneg_less[THEN sym, of "0" "m"], simp,
simp add:na_def,
frule apos_neq_minf[of "n"],
frule apos_neq_minf[of "m"])
apply (cut_tac mem_ant[of "m"],
cut_tac mem_ant[of "n"], simp,
(erule exE)+, simp,
simp add:tna_ant,
simp only:ant_0[THEN sym],
simp only:ale_zle)
done
lemma na_minf:"na (-\<infinity>) = 0"
apply (simp add:na_def, rule impI,
cut_tac minf_less_0, simp)
done
lemma an_na:"\<lbrakk>a \<noteq> \<infinity>; 0 \<le> a\<rbrakk> \<Longrightarrow> an (na a) = a"
apply (frule apos_tna_pos[of "a"], assumption,
frule apos_neq_minf[of "a"],
cut_tac mem_ant[of "a"], simp, erule exE,
simp, simp add:an_def na_def)
apply (cut_tac y = 0 and x = "ant z" in aneg_less, simp,
simp only:ant_0[THEN sym],
simp only:ale_zle, simp add:tna_ant)
done
lemma not_na_le_minf:"\<not> (an n \<le> -\<infinity> )"
apply (rule contrapos_pp, simp+)
apply (cut_tac minf_le_any[of "an n"], frule ale_antisym[of "an n" "-\<infinity>"],
assumption+, simp add:an_def)
done
lemma not_na_less_minf:"\<not> (an n < -\<infinity>)"
apply (simp add:aneg_less)
done
lemma not_na_ge_inf:"\<not> \<infinity> \<le> (an n)"
apply (simp add:aneg_le, unfold an_def)
apply (simp add:z_less_i[of "int n"])
done
lemma an_na_le:"j \<le> an n \<Longrightarrow> na j \<le> n"
apply (case_tac "j = -\<infinity>", simp add:na_minf)
apply (simp add:na_def)
apply (case_tac "j = \<infinity>", simp, rule impI)
apply (cut_tac not_na_ge_inf[of n], simp)
apply simp
apply (rule impI, simp add:aneg_less)
apply (frule an_na[of j], assumption)
apply (subgoal_tac "nat (tna j) = na j", simp,
thin_tac "nat (tna j) = na j")
apply (cut_tac ale_trans[of "an (na j)" j "an n"], thin_tac "j \<le> an n",
thin_tac "an (na j) = j", simp add:ale_nat_le[of "na j" n],
simp add:ale_refl[of j], assumption)
apply (thin_tac "an (na j) = j", simp add:na_def,
rule impI)
apply (simp add:aneg_le[THEN sym, of j 0])
done
lemma aless_neq :"(x::ant) < y \<Longrightarrow> x \<noteq> y"
by (rule contrapos_pp, simp+)
chapter "Ordered Set"
(* In this chapter, I prove Zorn's lemma in general form. *)
section "Basic Concepts of Ordered Sets"
record 'a carrier =
carrier :: "'a set"
record 'a Order = "'a carrier" +
rel :: "('a \<times> 'a) set"
locale Order =
fixes D (structure)
assumes closed: "rel D \<subseteq> carrier D \<times> carrier D"
and refl: "a \<in> carrier D \<Longrightarrow> (a, a) \<in> rel D"
and antisym: "\<lbrakk>a \<in> carrier D; b \<in> carrier D; (a, b) \<in> rel D;
(b, a) \<in> rel D\<rbrakk> \<Longrightarrow> a = b"
and trans: "\<lbrakk>a \<in> carrier D; b \<in> carrier D; c \<in> carrier D;
(a, b) \<in> rel D; (b, c) \<in> rel D\<rbrakk> \<Longrightarrow> (a, c) \<in> rel D"
(* print_locale Order *)
definition
ole :: "_ \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<preceq>\<index>" 60) where
"a \<preceq>\<^bsub>D\<^esub> b \<longleftrightarrow> (a, b) \<in> rel D"
definition
oless :: "_ \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" (infix "\<prec>\<index>" 60) where
"a \<prec>\<^bsub>D\<^esub> b \<equiv> a \<preceq>\<^bsub>D\<^esub> b \<and> a \<noteq> b"
lemma Order_component:"(E::'a Order) = \<lparr> carrier = carrier E, rel = rel E \<rparr>"
by simp (** An ordered set consists of two components **)
lemma Order_comp_eq:"\<lbrakk>carrier (E::'a Order) = carrier (F::'a Order);
rel E = rel F\<rbrakk> \<Longrightarrow> E = F"
by simp (* components coincide then ordered sets coincide. *)
lemma (in Order) le_rel:"\<lbrakk>a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<preceq> b) = ((a, b) \<in> rel D)"
by (simp add:ole_def)
lemma (in Order) less_imp_le:
"\<lbrakk>a \<in> carrier D; b \<in> carrier D; a \<prec> b \<rbrakk> \<Longrightarrow> a \<preceq> b"
by (simp add:oless_def)
lemma (in Order) le_refl:"a \<in> carrier D \<Longrightarrow> a \<preceq> a"
apply (unfold ole_def)
apply (rule refl, assumption)
done
lemma (in Order) le_antisym:"\<lbrakk>a \<in> carrier D; b \<in> carrier D;
a \<preceq> b; b \<preceq> a \<rbrakk> \<Longrightarrow> a = b"
apply (unfold ole_def)
apply (rule antisym)
apply assumption+
done
lemma (in Order) le_trans:"\<lbrakk>a \<in> carrier D; b \<in> carrier D; c \<in> carrier D;
a \<preceq> b; b \<preceq> c \<rbrakk> \<Longrightarrow> a \<preceq> c"
apply (unfold ole_def)
apply (rule_tac a = a and b = b and c = c in trans)
apply assumption+
done
lemma (in Order) less_trans:"\<lbrakk>a \<in> carrier D; b \<in> carrier D; c \<in> carrier D;
a \<prec> b; b \<prec> c \<rbrakk> \<Longrightarrow> a \<prec> c"
apply (unfold oless_def)
apply (erule conjE)+
apply (simp add:le_trans[of a b c])
apply (rule contrapos_pp, simp+)
apply (frule_tac le_antisym[of b c], assumption+)
apply simp
done
lemma (in Order) le_less_trans:"\<lbrakk>a \<in> carrier D; b \<in> carrier D; c \<in> carrier D;
a \<preceq> b; b \<prec> c \<rbrakk> \<Longrightarrow> a \<prec> c"
apply (simp add:oless_def)
apply (erule conjE)
apply (simp add:le_trans[of a b c])
apply (rule contrapos_pp, simp+)
apply (frule le_antisym[of "b" "c"])
apply assumption+
apply simp
done
lemma (in Order) less_le_trans:"\<lbrakk>a \<in> carrier D; b \<in> carrier D; c \<in> carrier D;
a \<prec> b; b \<preceq> c \<rbrakk> \<Longrightarrow> a \<prec> c"
apply (simp add:oless_def)
apply ( erule conjE)
apply (simp add:le_trans[of a b c])
apply (rule contrapos_pp, simp+)
apply (frule le_antisym[of "b" "c"])
apply assumption+
apply simp
done
lemma (in Order) le_imp_less_or_eq:
"\<lbrakk>a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow> (a \<preceq> b) = (a \<prec> b \<or> a = b)"
apply (simp add:oless_def)
apply (rule iffI)
apply simp
apply (erule disjE)
apply simp
apply simp
apply (rule le_refl)
apply assumption
done
lemma (in Order) less_neq: "a \<prec> b \<Longrightarrow> a \<noteq> b"
by (simp add: oless_def)
lemma (in Order) le_neq_less: "\<lbrakk>a \<preceq> b; a \<noteq> b\<rbrakk> \<Longrightarrow> a \<prec> b"
by (simp add: oless_def)
lemma (in Order) less_irrefl: "\<lbrakk>a \<in> carrier D; a \<prec> a\<rbrakk> \<Longrightarrow> C"
by (simp add:oless_def)
lemma (in Order) less_irrefl': "a \<in> carrier D \<Longrightarrow> \<not> a \<prec> a"
by (simp add:oless_def)
lemma (in Order) less_asym:
"a \<in> carrier D \<Longrightarrow> b \<in> carrier D \<Longrightarrow> a \<prec> b \<Longrightarrow> b \<prec> a \<Longrightarrow> C"
apply (simp add:oless_def)
apply (erule conjE)+
apply (frule le_antisym[of "a" "b"])
apply assumption+
apply simp
done
lemma (in Order) less_asym':
"a \<in> carrier D \<Longrightarrow> b \<in> carrier D \<Longrightarrow> a \<prec> b \<Longrightarrow> \<not> b \<prec> a"
apply (rule contrapos_pp, simp+)
apply (simp add:oless_def)
apply (erule conjE)+
apply (frule le_antisym[of "a" "b"])
apply assumption+
apply simp
done
lemma (in Order) gt_than_any_outside:"\<lbrakk>A \<subseteq> carrier D; b \<in> carrier D;
\<forall>x\<in>A. x \<prec> b\<rbrakk> \<Longrightarrow> b \<notin> A"
apply (rule contrapos_pp, simp+)
apply (frule_tac x = b in bspec)
apply (assumption,
thin_tac "\<forall>x\<in>A. x \<prec> b", simp add:oless_def)
done
definition
Iod :: "_ \<Rightarrow> 'a set \<Rightarrow> _" where
"Iod D T =
D \<lparr>carrier := T, rel := {(a, b). (a, b) \<in> rel D \<and> a \<in> T \<and> b \<in> T}\<rparr>"
definition
SIod :: "'a Order \<Rightarrow> 'a set \<Rightarrow> 'a Order" where
"SIod D T = \<lparr>carrier = T, rel = {(a, b). (a, b) \<in> rel D \<and> a \<in> T \<and> b \<in> T}\<rparr>"
lemma (in Order) Iod_self: "D = Iod D (carrier D)"
apply (unfold Iod_def)
apply (cases D)
apply (insert closed)
apply (simp add:Iod_def)
apply (rule equalityI)
apply (rule subsetI)
apply auto
done
lemma SIod_self:"Order D \<Longrightarrow> D = SIod D (carrier D)"
apply (unfold SIod_def)
apply (cases D)
apply (cut_tac Order.closed[of "D"])
apply auto
done
lemma (in Order) Od_carrier:"carrier (D\<lparr>carrier := S, rel := R\<rparr>) = S"
by simp
lemma (in Order) Od_rel:"rel (D\<lparr>carrier := S, rel := R\<rparr>) = R"
by simp
lemma (in Order) Iod_carrier:
"T \<subseteq> carrier D \<Longrightarrow> carrier (Iod D T) = T"
by (simp add: Iod_def)
lemma SIod_carrier:"\<lbrakk>Order D; T \<subseteq> carrier D\<rbrakk> \<Longrightarrow> carrier (SIod D T) = T"
by (simp add:SIod_def)
lemma (in Order) Od_compare:"(S = S' \<and> R = R') = (D\<lparr>carrier := S, rel := R\<rparr> = D\<lparr>carrier := S', rel := R'\<rparr>)"
apply (rule iffI)
apply simp
apply (cut_tac Od_carrier[of R S], cut_tac Od_carrier[of R' S'], simp)
apply (cut_tac Od_rel[of R S], cut_tac Od_rel[of R' S'])
apply (thin_tac "S' = S")
apply simp
done
lemma (in Order) Iod_le:
"\<lbrakk>T \<subseteq> carrier D; a \<in> T; b \<in> T\<rbrakk> \<Longrightarrow> (a \<preceq>\<^bsub>Iod D T\<^esub> b) = (a \<preceq> b)"
apply (simp add: Iod_def)
apply (simp add:ole_def)
done
lemma SIod_le:"\<lbrakk>T \<subseteq> carrier D; a \<in> T; b \<in> T\<rbrakk> \<Longrightarrow>
(a \<preceq>\<^bsub>SIod D T\<^esub> b) = (a \<preceq>\<^bsub>D\<^esub> b)"
apply (simp add:SIod_def)
apply (simp add:ole_def)
done
lemma (in Order) Iod_less:
"\<lbrakk>T \<subseteq> carrier D; a \<in> T; b \<in> T\<rbrakk> \<Longrightarrow> (a \<prec>\<^bsub>Iod D T\<^esub> b) = (a \<prec> b)"
apply (simp add:oless_def)
apply (simp add:Iod_le)
done
lemma SIod_less:"\<lbrakk>T \<subseteq> carrier D; a \<in> T; b \<in> T\<rbrakk> \<Longrightarrow>
(a \<prec>\<^bsub>SIod D T\<^esub> b) = (a \<prec>\<^bsub>D\<^esub> b)"
by (simp add:oless_def SIod_le)
lemma (in Order) Iod_Order:
"T \<subseteq> carrier D \<Longrightarrow> Order (Iod D T)"
apply (rule Order.intro)
apply (simp add:Iod_def)
apply (rule subsetI)
apply (unfold split_paired_all)
apply simp
apply (simp add:Iod_carrier)
apply (simp add:Iod_def)
apply (rule refl)
apply (rule subsetD, assumption+)
apply (simp add:Iod_carrier)
apply (simp add:Iod_def)
apply (rule_tac a = a and b = b in antisym)
apply (simp add:subsetD[of "T" "carrier D"])+
apply (simp add:Iod_def)
apply (rule_tac a = a and b = b and c = c in trans)
apply (simp add:subsetD[of "T" "carrier D"])+
done
lemma SIod_Order:"\<lbrakk> Order D; T \<subseteq> carrier D\<rbrakk> \<Longrightarrow> Order (SIod D T)"
apply (rule Order.intro)
apply (rule subsetI)
apply (simp add:SIod_def)
apply (unfold split_paired_all)
apply simp
apply (simp add:SIod_def)
apply (frule_tac c = a in subsetD[of T "carrier D"], assumption+)
apply (simp add:Order.refl[of D])
apply (simp add:SIod_def)
apply (rule Order.antisym[of D], assumption+)
apply (simp add:subsetD)+
apply (simp add:SIod_def)
apply (frule_tac c = a in subsetD[of T "carrier D"], assumption+,
frule_tac c = b in subsetD[of T "carrier D"], assumption+,
frule_tac c = c in subsetD[of T "carrier D"], assumption+)
apply (rule_tac a = a and b = b and c = c in Order.trans[of D], assumption+)
done
lemma (in Order) emptyset_Iod:"Order (Iod D {})"
apply (rule Iod_Order)
apply simp
done
lemma (in Order) Iod_sub_sub:
"\<lbrakk>S \<subseteq> T; T \<subseteq> carrier D\<rbrakk> \<Longrightarrow> Iod (Iod D T) S = Iod D S"
apply (simp add:Iod_def)
apply (subst Od_compare[THEN sym])
apply simp
apply blast
done
lemma SIod_sub_sub:
"\<lbrakk>S \<subseteq> T; T \<subseteq> carrier D\<rbrakk> \<Longrightarrow> SIod (SIod D T) S = SIod D S"
apply (simp add:SIod_def)
apply blast
done
lemma rel_SIod:"\<lbrakk>Order D; Order E; carrier E \<subseteq> carrier D;
\<forall>a\<in>carrier E. \<forall>b\<in>carrier E. (a \<preceq>\<^bsub>E\<^esub> b) = (a \<preceq>\<^bsub>D\<^esub> b)\<rbrakk> \<Longrightarrow>
rel E = rel (SIod D (carrier E))"
apply (rule equalityI) (* show the equality of the sets *)
apply (rule subsetI)
apply (unfold split_paired_all)
apply (simp add:ole_def)
apply (simp add:SIod_def)
apply (cut_tac Order.closed[of "E"])
apply blast
apply assumption
apply (rule subsetI)
apply (unfold split_paired_all)
apply (simp add:SIod_def)
apply (simp add:ole_def)
done
lemma SIod_self_le:"\<lbrakk>Order D; Order E;
carrier E \<subseteq> carrier D;
\<forall>a\<in>carrier E. \<forall>b\<in>carrier E. (a \<preceq>\<^bsub>E\<^esub> b) = (a \<preceq>\<^bsub>D\<^esub> b) \<rbrakk> \<Longrightarrow>
E = SIod D (carrier E)"
apply (rule Order_comp_eq[of "E" "SIod D (carrier E)"])
apply (simp add:SIod_carrier)
apply (rule rel_SIod[of "D" "E"], assumption+)
done
subsection \<open>Total ordering\<close>
locale Torder = Order +
assumes le_linear: "\<lbrakk>a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
a \<preceq> b \<or> b \<preceq> a"
lemma (in Order) Iod_empty_Torder:"Torder (Iod D {})"
apply (rule Torder.intro)
apply(simp add:emptyset_Iod)
apply (rule Torder_axioms.intro)
apply (simp add:Iod_carrier)
done
lemma (in Torder) le_cases:
"\<lbrakk>a \<in> carrier D; b \<in> carrier D; (a \<preceq> b \<Longrightarrow> C); (b \<preceq> a \<Longrightarrow> C)\<rbrakk> \<Longrightarrow> C"
by (cut_tac le_linear[of "a" "b"], blast, assumption+)
lemma (in Torder) Order:"Order D"
apply (rule Order_axioms)
done
lemma (in Torder) less_linear:
"a \<in> carrier D \<Longrightarrow> b \<in> carrier D \<Longrightarrow> a \<prec> b \<or> a = b \<or> b \<prec> a"
apply (simp add:oless_def)
apply (rule le_cases[of "a" "b"])
apply assumption+
apply blast
apply blast
done
lemma (in Torder) not_le_less:
"a \<in> carrier D \<Longrightarrow> b \<in> carrier D \<Longrightarrow>
(\<not> a \<preceq> b) = (b \<prec> a)"
apply (unfold oless_def)
apply (cut_tac le_linear[of a b])
apply (rule iffI)
apply simp
apply (rule contrapos_pp, simp+)
apply (rule contrapos_pp, simp+)
apply (erule conjE)
apply (frule le_antisym[of b a])
apply assumption+
apply simp+
done
lemma (in Torder) not_less_le:
"a \<in> carrier D \<Longrightarrow> b \<in> carrier D \<Longrightarrow>
(\<not> a \<prec> b) = (b \<preceq> a)"
apply (unfold oless_def)
apply (rule iffI)
apply (simp only:de_Morgan_conj[of "a \<preceq> b" "a \<noteq> b"])
apply (simp only:not_le_less[of "a" "b"])
apply (erule disjE)
apply (simp add:less_imp_le)
apply (simp add:le_imp_less_or_eq)
apply (rule contrapos_pp, simp+)
apply (erule conjE)
apply (frule le_antisym[of "a" "b"])
apply assumption+
apply simp
done
lemma (in Order) Iod_not_le_less:"\<lbrakk>T \<subseteq> carrier D; a \<in> T; b \<in> T;
Torder (Iod D T)\<rbrakk> \<Longrightarrow> (\<not> a \<preceq>\<^bsub>(Iod D T)\<^esub> b) = b \<prec>\<^bsub>(Iod D T)\<^esub> a"
apply (subst Torder.not_le_less)
apply assumption+
apply (simp add:Iod_carrier)+
done
lemma (in Order) Iod_not_less_le:"\<lbrakk>T \<subseteq> carrier D; a \<in> T; b \<in> T;
Torder (Iod D T)\<rbrakk> \<Longrightarrow> (\<not> a \<prec>\<^bsub>(Iod D T)\<^esub> b) = b \<preceq>\<^bsub>(Iod D T)\<^esub> a"
apply (subst Torder.not_less_le)
apply assumption+
apply (simp add:Iod_carrier)+
done
subsection \<open>Two ordered sets\<close>
definition
Order_Pow :: "'a set \<Rightarrow> 'a set Order" ("(po _)" [999] 1000) where
"po A =
\<lparr>carrier = Pow A,
rel = {(X, Y). X \<in> Pow A \<and> Y \<in> Pow A \<and> X \<subseteq> Y}\<rparr>"
interpretation order_Pow: Order "po A"
apply (unfold Order_Pow_def)
apply (rule Order.intro)
apply (rule subsetI)
apply (unfold split_paired_all)
apply simp
apply simp
apply simp
apply simp
done
definition
Order_fs :: "'a set \<Rightarrow> 'b set \<Rightarrow> ('a set * ('a \<Rightarrow> 'b)) Order" where
"Order_fs A B =
\<lparr>carrier = {Z. \<exists>A1 f. A1 \<in> Pow A \<and> f \<in> A1 \<rightarrow> B \<and>
f \<in> extensional A1 \<and> Z = (A1, f)},
rel = {Y. Y \<in> ({Z. \<exists>A1 f. A1 \<in> Pow A \<and> f \<in> A1 \<rightarrow> B \<and> f \<in> extensional A1
\<and> Z = (A1, f)}) \<times> ({Z. \<exists>A1 f. A1 \<in> Pow A \<and> f \<in> A1 \<rightarrow> B \<and> f \<in> extensional A1
\<and> Z = (A1, f)}) \<and> fst (fst Y) \<subseteq> fst (snd Y) \<and>
(\<forall>a\<in> (fst (fst Y)). (snd (fst Y)) a = (snd (snd Y)) a)}\<rparr>"
lemma Order_fs:"Order (Order_fs A B)"
apply (simp add:Order_fs_def)
apply (rule Order.intro)
apply (rule subsetI)
apply (unfold split_paired_all)
apply (auto intro: funcset_eq)
done
subsection \<open>Homomorphism of ordered sets\<close>
definition
ord_inj :: "[('a, 'm0) Order_scheme, ('b, 'm1) Order_scheme,
'a \<Rightarrow> 'b] \<Rightarrow> bool" where
"ord_inj D E f \<longleftrightarrow> f \<in> extensional (carrier D) \<and>
f \<in> (carrier D) \<rightarrow> (carrier E) \<and>
(inj_on f (carrier D)) \<and>
(\<forall>a\<in>carrier D. \<forall>b\<in>carrier D. (a \<prec>\<^bsub>D\<^esub> b) = ((f a) \<prec>\<^bsub>E\<^esub> (f b)))"
definition
ord_isom :: "[('a, 'm0) Order_scheme, ('b, 'm1) Order_scheme,
'a \<Rightarrow> 'b] \<Rightarrow> bool" where
"ord_isom D E f \<longleftrightarrow> ord_inj D E f \<and>
(surj_to f (carrier D) (carrier E))"
lemma (in Order) ord_inj_func:"\<lbrakk>Order E; ord_inj D E f\<rbrakk> \<Longrightarrow>
f \<in> carrier D \<rightarrow> carrier E"
by (simp add:ord_inj_def)
lemma (in Order) ord_isom_func:"\<lbrakk>Order E; ord_isom D E f\<rbrakk> \<Longrightarrow>
f \<in> carrier D \<rightarrow> carrier E"
by (simp add:ord_isom_def ord_inj_func)
lemma (in Order) ord_inj_restrict_isom:"\<lbrakk>Order E; ord_inj D E f; T \<subseteq> carrier D\<rbrakk>
\<Longrightarrow> ord_isom (Iod D T) (Iod E (f ` T)) (restrict f T)"
apply (subst ord_isom_def) (* The following two lemmas are preliminaries. *)
apply (frule ord_inj_func[of E f], assumption,
frule image_sub[of f "carrier D" "carrier E" "T"], assumption+)
apply (rule conjI)
apply (subst ord_inj_def)
apply (simp add:Iod_carrier Order.Iod_carrier)
apply (rule conjI)
apply (rule restrict_inj[of f "carrier D" "T"])
apply (simp add:ord_inj_def, assumption+)
apply (rule ballI)+
apply (frule_tac x = a in elem_in_image2[of f "carrier D" "carrier E" T],
assumption+,
frule_tac x = b in elem_in_image2[of f "carrier D" "carrier E" T],
assumption+)
apply (simp add:Iod_less Order.Iod_less)
apply (frule_tac c = a in subsetD[of T "carrier D"], assumption+,
frule_tac c = b in subsetD[of T "carrier D"], assumption+)
apply (simp add:ord_inj_def)
apply (subst surj_to_def)
apply (simp add:Iod_carrier Order.Iod_carrier)
done
lemma ord_inj_Srestrict_isom:"\<lbrakk>Order D; Order E; ord_inj D E f; T \<subseteq> carrier D\<rbrakk>
\<Longrightarrow> ord_isom (SIod D T) (SIod E (f ` T)) (restrict f T)"
apply (subst ord_isom_def)
apply (frule Order.ord_inj_func[of D E f], assumption+,
frule image_sub[of f "carrier D" "carrier E" "T"], assumption+)
apply (rule conjI)
apply (subst ord_inj_def)
apply (simp add:SIod_carrier)
apply (rule conjI)
apply (rule restrict_inj[of f "carrier D" "T"])
apply (simp add:ord_inj_def, assumption+)
apply (rule ballI)+
apply (frule_tac x = a in elem_in_image2[of f "carrier D" "carrier E" T],
assumption+,
frule_tac x = b in elem_in_image2[of f "carrier D" "carrier E" T],
assumption+)
apply (simp add:SIod_less)
apply (frule_tac c = a in subsetD[of T "carrier D"], assumption+,
frule_tac c = b in subsetD[of T "carrier D"], assumption+)
apply (simp add:ord_inj_def)
apply (simp add:SIod_carrier)
apply (simp add:surj_to_def)
done
lemma (in Order) id_ord_isom:"ord_isom D D (idmap (carrier D))"
apply (simp add:ord_isom_def)
apply (cut_tac idmap_bij[of "carrier D"])
apply (simp add:bij_to_def)
apply (simp add:ord_inj_def)
apply (simp add:idmap_def[of "carrier D"])
done
lemma (in Order) ord_isom_bij_to:"\<lbrakk>Order E; ord_isom D E f\<rbrakk> \<Longrightarrow>
bij_to f (carrier D) (carrier E)"
by (simp add:bij_to_def ord_isom_def,
simp add:ord_inj_def)
lemma (in Order) ord_inj_mem:"\<lbrakk>Order E; ord_inj D E f; a \<in> carrier D\<rbrakk> \<Longrightarrow>
(f a) \<in> carrier E"
apply (simp add:ord_inj_def, (erule conjE)+)
apply (simp add:Pi_def)
done
lemma (in Order) ord_isom_mem:"\<lbrakk>Order E; ord_isom D E f; a \<in> carrier D\<rbrakk> \<Longrightarrow>
(f a) \<in> carrier E"
apply (simp add:ord_isom_def, (erule conjE)+)
apply (simp add:ord_inj_mem)
done
lemma (in Order) ord_isom_surj:"\<lbrakk>Order E; ord_isom D E f; b \<in> carrier E\<rbrakk> \<Longrightarrow>
\<exists>a\<in>carrier D. b = f a"
apply (simp add:ord_isom_def, (erule conjE)+)
apply (simp add:surj_to_def image_def)
apply (frule sym, thin_tac "{y. \<exists>x\<in>carrier D. y = f x} = carrier E",
simp)
done
lemma (in Order) ord_isom_surj_forall:"\<lbrakk>Order E; ord_isom D E f\<rbrakk> \<Longrightarrow>
\<forall>b \<in> carrier E. \<exists>a\<in>carrier D. b = f a"
apply (rule ballI)
apply (rule ord_isom_surj[of "E" "f"], assumption+)
done
lemma (in Order) ord_isom_onto:"\<lbrakk>Order E; ord_isom D E f\<rbrakk> \<Longrightarrow>
f ` (carrier D) = carrier E "
apply (frule ord_isom_bij_to[of "E" "f"], assumption+)
apply(simp add:bij_to_def surj_to_def)
done
lemma (in Order) ord_isom_inj_on:"\<lbrakk>Order E; ord_isom D E f\<rbrakk> \<Longrightarrow>
inj_on f (carrier D)"
by (simp add:ord_isom_def ord_inj_def)
lemma (in Order) ord_isom_inj:"\<lbrakk>Order E; ord_isom D E f;
a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow> (a = b) = ((f a) = (f b))"
apply (frule ord_isom_inj_on[of E f], assumption)
apply (simp add:injective_iff)
done
lemma (in Order) ord_isom_surj_to:"\<lbrakk>Order E; ord_isom D E f\<rbrakk> \<Longrightarrow>
surj_to f (carrier D) (carrier E)"
by (simp add:ord_isom_def)
lemma (in Order) ord_inj_less:"\<lbrakk>Order E; ord_inj D E f; a \<in> carrier D;
b \<in> carrier D\<rbrakk> \<Longrightarrow> (a \<prec>\<^bsub>D\<^esub> b) = ((f a) \<prec>\<^bsub>E\<^esub> (f b))"
by (simp add:ord_inj_def)
lemma (in Order) ord_isom_less:"\<lbrakk>Order E; ord_isom D E f;
a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow> (a \<prec>\<^bsub>D\<^esub> b) = ((f a) \<prec>\<^bsub>E\<^esub> (f b))"
by (simp add:ord_isom_def ord_inj_less)
lemma (in Order) ord_isom_less_forall:"\<lbrakk>Order E; ord_isom D E f\<rbrakk> \<Longrightarrow>
\<forall>a \<in> carrier D. \<forall> b \<in> carrier D. (a \<prec>\<^bsub>D\<^esub> b) = ((f a) \<prec>\<^bsub>E\<^esub> (f b))"
by ((rule ballI)+,
simp add:ord_isom_less)
lemma (in Order) ord_isom_le:"\<lbrakk>Order E; ord_isom D E f;
a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow> (a \<preceq>\<^bsub>D\<^esub> b) = ((f a) \<preceq>\<^bsub>E\<^esub> (f b))"
apply (frule_tac a = a in ord_isom_mem[of "E" "f"], assumption+,
frule_tac a = b in ord_isom_mem[of "E" "f"], assumption+)
apply (simp add:le_imp_less_or_eq Order.le_imp_less_or_eq[of "E"])
apply (simp add:ord_isom_less ord_isom_inj)
done
lemma (in Order) ord_isom_le_forall:"\<lbrakk>Order E; ord_isom D E f\<rbrakk> \<Longrightarrow>
\<forall>a \<in> carrier D. \<forall> b \<in> carrier D. (a \<preceq> b) = ((f a) \<preceq>\<^bsub>E\<^esub> (f b))"
by ((rule ballI)+,
rule ord_isom_le, assumption+)
lemma (in Order) ord_isom_convert:"\<lbrakk>Order E; ord_isom D E f;
x \<in> carrier D; a \<in> carrier D\<rbrakk> \<Longrightarrow> (\<forall>y\<in>carrier D. (x \<prec> y \<longrightarrow> \<not> y \<prec> a)) =
(\<forall>z\<in>carrier E. ((f x) \<prec>\<^bsub>E\<^esub> z \<longrightarrow> \<not> z \<prec>\<^bsub>E\<^esub> (f a)))"
apply (rule iffI)
apply (rule ballI, rule impI)
apply (frule_tac b = z in ord_isom_surj[of "E" "f"], assumption+,
erule bexE)
apply ( simp add:ord_isom_less[THEN sym, of "E" "f"])
apply (rule ballI, rule impI)
apply (simp add:ord_isom_less[of "E" "f"])
apply (frule_tac a = y in ord_isom_mem[of "E" "f"], assumption+)
apply simp
done
lemma (in Order) ord_isom_sym:"\<lbrakk>Order E; ord_isom D E f\<rbrakk> \<Longrightarrow>
ord_isom E D (invfun (carrier D) (carrier E) f)"
apply (frule ord_isom_func[of E f], assumption+,
frule ord_isom_inj_on[of E f], assumption+,
frule ord_isom_surj_to[of E f], assumption+)
apply (subst ord_isom_def, subst ord_inj_def)
apply (simp add:inv_func)
apply (simp add:invfun_inj)
apply (simp add:invfun_surj)
apply (rule conjI)
apply (simp add:invfun_def extensional_def)
apply (rule ballI)+
apply (frule_tac b = a in invfun_mem[of "f" "carrier D" "carrier E"],
assumption+,
frule_tac b = b in invfun_mem[of "f" "carrier D" "carrier E"],
assumption+)
apply (frule_tac a = "(f\<inverse>\<^bsub>carrier E,carrier D\<^esub>) a" and b = "(f\<inverse>\<^bsub>carrier E,carrier D\<^esub>) b"
in ord_isom_less[of E f], assumption+)
apply (simp add:invfun_r)
done
lemma (in Order) ord_isom_trans:"\<lbrakk>Order E; Order F; ord_isom D E f;
ord_isom E F g \<rbrakk> \<Longrightarrow> ord_isom D F (compose (carrier D) g f)"
apply (frule ord_isom_func[of E f], assumption+,
frule ord_isom_inj_on[of E f], assumption+,
frule ord_isom_surj_to[of E f], assumption+,
frule Order.ord_isom_func[of E F g], assumption+,
frule Order.ord_isom_inj_on[of E F g], assumption+,
frule Order.ord_isom_surj_to[of E F g], assumption+)
(* lemmas concerning compose require assumptions given above *)
apply (subst ord_isom_def, subst ord_inj_def)
apply (simp add:composition)
apply (simp add:comp_inj[of "f" "carrier D" "carrier E" "g" "carrier F"])
apply (simp add:compose_surj)
apply (rule ballI)+
apply (frule_tac x = a in funcset_mem[of f "carrier D" "carrier E"],
assumption+,
frule_tac x = b in funcset_mem[of f "carrier D" "carrier E"],
assumption+)
apply (frule_tac a = a and b = b in ord_isom_less[of E f], assumption+,
frule_tac a = "f a" and b = "f b" in Order.ord_isom_less[of E F g],
assumption+)
apply (simp add:compose_def)
done
definition
ord_equiv :: "[_, ('b, 'm1) Order_scheme] \<Rightarrow> bool" where
"ord_equiv D E \<longleftrightarrow> (\<exists>f. ord_isom D E f)"
lemma (in Order) ord_equiv:"\<lbrakk>Order E; ord_isom D E f\<rbrakk> \<Longrightarrow> ord_equiv D E"
by (simp add:ord_equiv_def, blast)
lemma (in Order) ord_equiv_isom:"\<lbrakk>Order E; ord_equiv D E\<rbrakk> \<Longrightarrow>
\<exists>f. ord_isom D E f"
by (simp add:ord_equiv_def)
lemma (in Order) ord_equiv_reflex:"ord_equiv D D"
apply (simp add:ord_equiv_def)
apply (cut_tac id_ord_isom, blast)
done
lemma (in Order) eq_ord_equiv:"\<lbrakk>Order E; D = E\<rbrakk> \<Longrightarrow> ord_equiv D E"
apply (frule sym, thin_tac "D = E")
apply ( simp add:ord_equiv_reflex)
done
lemma (in Order) ord_equiv_sym:"\<lbrakk>Order E; ord_equiv D E \<rbrakk> \<Longrightarrow> ord_equiv E D"
apply (simp add:ord_equiv_def)
apply (erule exE,
frule_tac E = E and f = f in ord_isom_sym, assumption+, blast)
done
lemma (in Order) ord_equiv_trans:"\<lbrakk>Order E; Order F; ord_equiv D E;
ord_equiv E F\<rbrakk> \<Longrightarrow> ord_equiv D F"
apply (simp add:ord_equiv_def)
apply (erule exE)+
apply (frule_tac f = f and g = fa in ord_isom_trans [of "E" "F"],
assumption+, blast)
done
lemma (in Order) ord_equiv_box:"\<lbrakk>Order E; Order F; ord_equiv D E;
ord_equiv D F\<rbrakk> \<Longrightarrow> ord_equiv E F"
apply (rule Order.ord_equiv_trans[of E D F])
apply assumption
apply (rule Order_axioms)
apply assumption
apply (rule ord_equiv_sym) apply assumption+
done
lemma SIod_isom_Iod:"\<lbrakk>Order D; T \<subseteq> carrier D \<rbrakk> \<Longrightarrow>
ord_isom (SIod D T) (Iod D T) (\<lambda>x\<in>T. x)"
apply (simp add:ord_isom_def ord_inj_def)
apply (simp add:SIod_carrier Order.Iod_carrier)
apply (rule conjI)
apply (fold idmap_def[of T])
apply (simp add:SIod_less Order.Iod_less)
apply (cut_tac A = T in idmap_bij,
simp add:bij_to_def)
done
definition
minimum_elem :: "[_ , 'a set, 'a] \<Rightarrow> bool" where
"minimum_elem = (\<lambda>D X a. a \<in> X \<and> (\<forall>x\<in>X. a \<preceq>\<^bsub>D\<^esub> x))"
locale Worder = Torder +
assumes ex_minimum: "\<forall>X. X \<subseteq> (carrier D) \<and> X \<noteq> {} \<longrightarrow>
(\<exists>x. minimum_elem D X x)"
lemma (in Worder) Order:"Order D"
by (rule Order)
lemma (in Worder) Torder:"Torder D"
apply (rule Torder_axioms)
done
lemma (in Worder) Worder:"Worder D"
apply (rule Worder_axioms)
done
lemma (in Worder) equiv_isom:"\<lbrakk>Worder E; ord_equiv D E\<rbrakk> \<Longrightarrow>
\<exists>f. ord_isom D E f"
by (insert Order, frule Worder.Order[of "E"], simp add:ord_equiv_def)
lemma (in Order) minimum_elem_mem:"\<lbrakk>X \<subseteq> carrier D; minimum_elem D X a\<rbrakk>
\<Longrightarrow> a \<in> X"
by (simp add:minimum_elem_def)
lemma (in Order) minimum_elem_unique:"\<lbrakk>X \<subseteq> carrier D; minimum_elem D X a1;
minimum_elem D X a2\<rbrakk> \<Longrightarrow> a1 = a2"
apply (frule minimum_elem_mem[of "X" "a1"], assumption+,
frule minimum_elem_mem[of "X" "a2"], assumption+)
apply (simp add:minimum_elem_def)
apply (drule_tac x = a2 in bspec, assumption)
apply (drule_tac x = a1 in bspec, assumption)
apply (rule le_antisym[of a1 a2])
apply (simp add:subsetD)+
done
lemma (in Order) compare_minimum_elements:"\<lbrakk>S \<subseteq> carrier D; T \<subseteq> carrier D;
S \<subseteq> T; minimum_elem D S s; minimum_elem D T t \<rbrakk> \<Longrightarrow> t \<preceq> s"
apply (frule minimum_elem_mem[of "S" "s"], assumption+)
apply (frule subsetD[of "S" "T" "s"], assumption+)
apply (simp add:minimum_elem_def)
done
lemma (in Order) minimum_elem_sub:"\<lbrakk>T \<subseteq> carrier D; X \<subseteq> T\<rbrakk>
\<Longrightarrow> minimum_elem D X a = minimum_elem (Iod D T) X a"
apply (simp add:minimum_elem_def)
apply (simp add:subset_eq[of X T])
apply (rule iffI, erule conjE)
apply simp
apply (rule ballI)
apply (simp add:Iod_le)
apply simp
apply (rule ballI)
apply (erule conjE)
apply (simp add:Iod_le)
done
lemma minimum_elem_Ssub:"\<lbrakk>Order D; T \<subseteq> carrier D; X \<subseteq> T\<rbrakk>
\<Longrightarrow> minimum_elem D X a = minimum_elem (SIod D T) X a"
apply (simp add:minimum_elem_def)
apply (rule iffI)
apply simp
apply (rule ballI, erule conjE)
apply (drule_tac x = x in bspec, assumption)
apply (frule_tac c = x in subsetD[of "X" "T"], assumption+,
frule_tac c = a in subsetD[of "X" "T"], assumption+)
apply (simp add:SIod_le)
apply simp
apply (rule ballI, erule conjE)
apply (drule_tac x = x in bspec, assumption)
apply (frule_tac c = x in subsetD[of "X" "T"], assumption+,
frule_tac c = a in subsetD[of "X" "T"], assumption+)
apply (simp add:SIod_le)
done
lemma (in Order) augmented_set_minimum:"\<lbrakk>X \<subseteq> carrier D; a \<in> carrier D;
Y - {a} \<subseteq> X; y - {a} \<noteq> {}; minimum_elem (Iod D X) (Y - {a}) x;
\<forall>x\<in>X. x \<preceq> a\<rbrakk> \<Longrightarrow> minimum_elem (Iod D (insert a X)) Y x"
apply (frule insert_mono[of "Y - {a}" "X" "a"])
apply simp
apply (frule insert_sub[of X "carrier D" a], assumption+)
apply (simp add:minimum_elem_sub[THEN sym, of "insert a X" Y],
simp add:minimum_elem_sub[THEN sym, of X "Y - {a}"])
apply (simp add:subset_eq[of "Y - {a}" X])
apply (simp add:minimum_elem_def, (erule conjE)+)
apply (rule ballI)
apply blast
done
lemma augmented_Sset_minimum:"\<lbrakk>Order D; X \<subseteq> carrier D; a \<in> carrier D;
Y - {a} \<subseteq> X; y - {a} \<noteq> {}; minimum_elem (SIod D X) (Y - {a}) x;
\<forall>x\<in>X. x \<preceq>\<^bsub>D\<^esub> a\<rbrakk> \<Longrightarrow> minimum_elem (SIod D (insert a X)) Y x"
apply (frule insert_mono[of "Y - {a}" "X" "a"])
apply simp
apply (frule insert_sub[of X "carrier D" a], assumption+)
apply (simp add:minimum_elem_Ssub[THEN sym, of D "insert a X" Y],
simp add:minimum_elem_Ssub[THEN sym, of D X "Y - {a}"])
apply (simp add:subset_eq[of "Y - {a}" X])
apply (simp add:minimum_elem_def, (erule conjE)+)
apply (rule ballI)
apply blast
done
lemma (in Order) ord_isom_minimum:"\<lbrakk>Order E; ord_isom D E f;
S \<subseteq> carrier D; a \<in> carrier D; minimum_elem D S a\<rbrakk> \<Longrightarrow>
minimum_elem E (f`S) (f a)"
apply (subst minimum_elem_def,
frule minimum_elem_mem[of "S" "a"], assumption+)
apply (simp add:ord_isom_mem)
apply (rule ballI)
apply (simp add:minimum_elem_def)
apply (frule_tac x = x in bspec, assumption,
thin_tac "Ball S ((\<preceq>) a)")
apply (frule_tac b = x in ord_isom_le[of E f a], assumption+)
apply (simp add:subsetD)
apply simp
done
lemma (in Worder) pre_minimum:"\<lbrakk>T \<subseteq> carrier D; minimum_elem D T t;
s \<in> carrier D; s \<prec>\<^bsub>D\<^esub> t \<rbrakk> \<Longrightarrow> \<not> s \<in> T"
apply (rule contrapos_pp, simp+)
apply (simp add:minimum_elem_def, (erule conjE)+)
apply (frule_tac x = s in bspec, assumption+,
thin_tac "\<forall>x\<in>T. t \<preceq>\<^bsub>D\<^esub> x")
apply (simp add:oless_def, erule conjE)
apply (frule le_antisym[of s t])
apply (simp add:subsetD[of "T" "carrier D"], assumption+)
apply simp
done
lemma bex_nonempty_subset:"\<exists>a. a \<in> A \<and> P a \<Longrightarrow>
{x. x \<in> A \<and> P x} \<subseteq> A \<and> {x. x \<in> A \<and> P x} \<noteq> {}"
apply (erule exE, rule conjI)
apply (rule subsetI, simp)
apply (rule_tac A = "{x \<in> A. P x}" in nonempty, simp)
done
lemma (in Worder) to_subset:"\<lbrakk>T \<subseteq> carrier D; ord_isom D (Iod D T) f\<rbrakk> \<Longrightarrow>
\<forall>a. a \<in> carrier D \<longrightarrow> a \<preceq> (f a)"
apply (rule contrapos_pp, simp+)
apply (cut_tac ex_minimum)
apply (drule_tac a = "{a. a \<in> carrier D \<and> \<not> a \<preceq> f a}" in forall_spec) (*
thin_tac "\<forall>X. X \<subseteq> carrier D \<and> X \<noteq> {} \<longrightarrow> (\<exists>x. minimum_elem D X x)") *)
apply (rule conjI)
apply (rule subsetI, simp)
apply (rule ex_nonempty, simp)
(*
apply (thin_tac "\<forall>X. X \<subseteq> carrier D \<and> X \<noteq> {} \<longrightarrow> (\<exists>x. minimum_elem D X x)",
thin_tac "\<exists>a. a \<in> carrier D \<and> \<not> a \<preceq> f a") *)
apply ((erule exE)+, simp add:minimum_elem_def, (erule conjE)+)
apply (frule Iod_Order[of "T"],
frule_tac a = x in ord_isom_mem[of "Iod D T" "f"], assumption+)
apply (frule_tac a = x and b = "f x" in ord_isom_le[of "Iod D T" "f"],
assumption+)
apply (simp add:Iod_carrier subsetD)
apply (frule Iod_carrier[of "T"],
frule_tac a = "f x" in eq_set_inc[of _ "carrier (Iod D T)" "T"],
assumption+)
apply (frule_tac c = "f x" in subsetD[of "T" "carrier D"], assumption+)
apply (frule_tac a = "f x" in ord_isom_mem[of "Iod D T" "f"], assumption+)
apply (frule_tac a = "f (f x)" in eq_set_inc[of _ "carrier (Iod D T)" "T"],
assumption+)
apply (drule_tac a = "f x" in forall_spec)
(* thin_tac "\<forall>xa. xa \<in> carrier D \<and> \<not> xa \<preceq> f xa \<longrightarrow> x \<preceq> xa") *)
apply (simp add:subsetD Iod_le)
apply simp
done
lemma to_subsetS:"\<lbrakk>Worder D; T \<subseteq> carrier D; ord_isom D (SIod D T) f\<rbrakk> \<Longrightarrow>
\<forall>a. a \<in> carrier D \<longrightarrow> a \<preceq>\<^bsub>D\<^esub> (f a)"
apply (frule Worder.Order[of "D"],
frule SIod_isom_Iod[of "D" "T"], assumption+,
frule Order.ord_isom_trans[of "D" "SIod D T" "Iod D T" f "\<lambda>x\<in>T. x"])
apply (simp add:SIod_Order, simp add:Order.Iod_Order, assumption+)
apply (frule_tac D = D and T = T and f = "compose (carrier D) (\<lambda>x\<in>T. x) f"
in Worder.to_subset, assumption+)
apply (rule allI, rule impI)
apply (drule_tac a = a in forall_spec, simp)
(* thin_tac "\<forall>a. a \<in> carrier D \<longrightarrow>
a \<preceq>\<^bsub>D\<^esub> compose (carrier D) (\<lambda>x\<in>T. x) f a") *)
apply (frule_tac a = a in Order.ord_isom_mem[of "D" "SIod D T" "f"])
apply (simp add:SIod_Order, assumption+)
apply (simp add:SIod_carrier)
apply (simp add:compose_def)
done
lemma (in Worder) isom_Worder:"\<lbrakk>Order T; ord_isom D T f\<rbrakk> \<Longrightarrow> Worder T"
apply (rule Worder.intro)
apply (rule Torder.intro)
apply assumption
apply (rule Torder_axioms.intro)
apply (frule_tac b = a in ord_isom_surj[of T f], assumption+,
frule_tac b = b in ord_isom_surj[of T f], assumption+,
(erule bexE)+)
apply (cut_tac Torder_axioms, simp add:Torder_axioms_def)
apply (meson le_cases ord_isom_le)
apply (rule Worder_axioms.intro)
apply (rule allI, rule impI, erule conjE)
apply (frule ord_isom_func[of "T" "f"], assumption+)
apply (frule ord_isom_bij_to[of "T" "f"], assumption+)
apply (frule ord_isom_sym[of "T" "f"], assumption+,
frule Order.ord_isom_func[of "T" "D"
"invfun (carrier D) (carrier T) f"])
apply (rule Order, assumption)
apply (frule_tac ?A1.0 = X in image_sub[of
"invfun (carrier D) (carrier T) f" "carrier T" "carrier D"],
assumption+,
frule_tac ?A1.0 = X in image_nonempty[of "invfun (carrier D)
(carrier T) f" "carrier T" "carrier D"], assumption+)
apply (cut_tac ex_minimum) (** Because D is well ordered **)
apply (drule_tac a = "invfun (carrier D) (carrier T) f ` X" in forall_spec,
(* thin_tac "\<forall>X. X \<subseteq> carrier D \<and> X \<noteq> {} \<longrightarrow> (\<exists>x. minimum_elem D X x)", *)
simp) apply (
(* thin_tac "\<forall>X. X \<subseteq> carrier D \<and> X \<noteq> {} \<longrightarrow> (\<exists>x. minimum_elem D X x)", *)
erule exE)
apply (frule_tac S = "invfun (carrier D) (carrier T) f ` X" and a = x in
ord_isom_minimum[of "T" "f"], assumption+)
apply (frule_tac X = "invfun (carrier D) (carrier T) f ` X" and a = x in
minimum_elem_mem, assumption+)
apply (simp add:subsetD) apply assumption
apply (simp add:invfun_set, blast)
done
lemma (in Worder) equiv_Worder:"\<lbrakk>Order T; ord_equiv D T\<rbrakk> \<Longrightarrow> Worder T"
by (simp add:ord_equiv_def,
erule exE, simp add:isom_Worder)
lemma (in Worder) equiv_Worder1:"\<lbrakk>Order T; ord_equiv T D\<rbrakk> \<Longrightarrow> Worder T"
apply (cut_tac Worder,
frule Worder.Order[of D],
frule Order.ord_equiv_sym[of T D], assumption+)
apply (rule equiv_Worder, assumption+)
done
lemma (in Worder) ord_isom_self_id:"ord_isom D D f \<Longrightarrow> f = idmap (carrier D)"
apply (cut_tac Order,
frule ord_isom_sym [of "D" "f"], assumption+,
frule ord_isom_func[of "D" "f"], assumption+)
apply (rule funcset_eq[of "f" "carrier D" "idmap (carrier D)"])
apply (simp add:ord_isom_def ord_inj_def, simp add:idmap_def)
apply (rule ballI)
apply (simp add:idmap_def)
apply (cut_tac subset_self[of "carrier D"],
frule to_subset [of "carrier D" "f"],
simp add:Iod_self[THEN sym])
apply (drule_tac a = x in forall_spec, assumption
(* thin_tac "\<forall>a. a \<in> carrier D \<longrightarrow> a \<preceq> (f a)" *))
apply (frule to_subset [of "carrier D" "invfun (carrier D) (carrier D) f"])
apply (simp add:Iod_self[THEN sym])
apply (drule_tac a = x in forall_spec, assumption) (*,
thin_tac "\<forall>a. a \<in> carrier D \<longrightarrow>
a \<preceq> (invfun (carrier D) (carrier D) f a)") *)
apply (frule_tac x = x in funcset_mem [of "f" "carrier D" "carrier D"],
assumption+)
apply (frule_tac a = x in ord_isom_mem[of "D"
"invfun (carrier D) (carrier D) f"], assumption+)
apply (frule_tac a = x and b = "invfun (carrier D) (carrier D) f x" in
ord_isom_le[of "D" "f"], assumption+)
apply simp
apply (frule ord_isom_bij_to[of "D" "f"], assumption+,
simp add:bij_to_def, erule conjE)
apply (simp add:invfun_r[of "f" "carrier D" "carrier D"])
apply (rule_tac a = "f x" and b = x in le_antisym,
assumption+)
done
lemma (in Worder) isom_unique:"\<lbrakk>Worder E; ord_isom D E f; ord_isom D E g\<rbrakk>
\<Longrightarrow> f = g"
apply (frule Worder.Order[of "E"])
apply (insert Order,
frule ord_isom_sym[of "E" "g"], assumption+,
frule ord_isom_trans [of "E" "D" "f"
"invfun (carrier D) (carrier E) g"], assumption+,
frule ord_isom_func[of "D"
"compose (carrier D) (invfun (carrier D) (carrier E) g) f"], assumption+)
apply (frule ord_isom_self_id [of
"compose (carrier D) (invfun (carrier D) (carrier E) g) f"])
apply (thin_tac "ord_isom E D (invfun (carrier D) (carrier E) g)")
apply (cut_tac id_ord_isom, insert Order,
frule ord_isom_func[of "D" "idmap (carrier D)"], assumption+)
apply (rule funcset_eq[of "f" "carrier D" "g"])
apply (simp add:ord_isom_def ord_inj_def)
apply (simp add:ord_isom_def ord_inj_def)
apply (rule ballI)
apply (frule_tac x = x in eq_funcs[of
"compose (carrier D) (invfun (carrier D) (carrier E) g) f"
"carrier D" "carrier D" "idmap (carrier D)"], assumption+)
apply (frule_tac a = x in ord_isom_mem [of "E" "f"], assumption+,
thin_tac " compose (carrier D) (invfun (carrier D) (carrier E) g) f =
idmap (carrier D)",
simp add:idmap_def compose_def)
apply (simp add:ord_isom_def[of _ "E" "g"] ord_inj_def, (erule conjE)+)
apply (frule_tac b = "f x" in invfun_r[of "g" "carrier D" "carrier E"],
assumption+)
apply simp
done
definition
segment :: "_ \<Rightarrow> 'a \<Rightarrow> 'a set" where
"segment D a = (if a \<notin> carrier D then carrier D else
{x. x \<prec>\<^bsub>D\<^esub> a \<and> x \<in> carrier D})"
definition
Ssegment :: "'a Order \<Rightarrow> 'a \<Rightarrow> 'a set" where
"Ssegment D a = (if a \<notin> carrier D then carrier D else
{x. x \<prec>\<^bsub>D\<^esub> a \<and> x \<in> carrier D})"
lemma (in Order) segment_sub:"segment D a \<subseteq> carrier D"
apply (rule subsetI, simp add:segment_def)
apply (case_tac "a \<notin> carrier D", simp)
apply ( simp add:segment_def)
done
lemma Ssegment_sub:"Ssegment D a \<subseteq> carrier D"
by (rule subsetI, simp add:Ssegment_def,
case_tac "a \<notin> carrier D", simp, simp add:Ssegment_def)
lemma (in Order) segment_free:"a \<notin> carrier D \<Longrightarrow>
segment D a = carrier D"
by (simp add:segment_def)
lemma Ssegment_free:"a \<notin> carrier D \<Longrightarrow>
Ssegment D a = carrier D"
by (simp add:Ssegment_def)
lemma (in Order) segment_sub_sub:"\<lbrakk>S \<subseteq> carrier D; d \<in> S\<rbrakk> \<Longrightarrow>
segment (Iod D S) d \<subseteq> segment D d"
apply (rule subsetI)
apply (frule_tac c = d in subsetD[of "S" "carrier D"], assumption+)
apply (simp add:segment_def)
apply (simp add:Iod_carrier)
apply (erule conjE, simp add:Iod_less[of "S"])
apply (simp add:subsetD)
done
lemma Ssegment_sub_sub:"\<lbrakk>Order D; S \<subseteq> carrier D; d \<in> S\<rbrakk> \<Longrightarrow>
Ssegment (SIod D S) d \<subseteq> Ssegment D d"
apply (rule subsetI)
apply (frule_tac c = d in subsetD[of "S" "carrier D"], assumption+)
apply (simp add:Ssegment_def)
apply (simp add:SIod_carrier, erule conjE, simp add:SIod_less[of "S"])
apply (simp add:subsetD)
done
lemma (in Order) a_notin_segment:"a \<notin> segment D a"
by (simp add:segment_def oless_def)
lemma a_notin_Ssegment:"a \<notin> Ssegment D a"
by (simp add:Ssegment_def oless_def)
lemma (in Order) Iod_carr_segment:
"carrier (Iod D (segment D a)) = segment D a"
by (cut_tac segment_sub[of "a"], simp add:Iod_carrier)
lemma SIod_carr_Ssegment:"Order D \<Longrightarrow>
carrier (SIod D (Ssegment D a)) = Ssegment D a"
apply (cut_tac Ssegment_sub[of "D" "a"])
apply (simp add:SIod_carrier)
done
lemma (in Order) segment_inc:"\<lbrakk>a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<prec> b) = (a \<in> segment D b)"
by (simp add:segment_def)
lemma Ssegment_inc:"\<lbrakk>Order D; a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<prec>\<^bsub>D\<^esub> b) = (a \<in> Ssegment D b)"
by (simp add:Ssegment_def)
lemma (in Order) segment_inc1:"b \<in> carrier D \<Longrightarrow>
(a \<prec> b \<and> a \<in> carrier D) = (a \<in> segment D b)"
by (simp add:segment_def)
lemma Ssegment_inc1:"\<lbrakk>Order D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<prec>\<^bsub>D\<^esub> b \<and> a \<in> carrier D) = (a \<in> Ssegment D b)"
by (simp add:Ssegment_def)
lemma (in Order) segment_inc_if:"\<lbrakk>b \<in> carrier D;a \<in> segment D b\<rbrakk> \<Longrightarrow>
a \<prec> b"
by (simp add:segment_def)
lemma Ssegment_inc_if:"\<lbrakk>Order D; b \<in> carrier D; a \<in> Ssegment D b\<rbrakk> \<Longrightarrow>
a \<prec>\<^bsub>D\<^esub> b"
by (simp add:Ssegment_def)
lemma (in Order) segment_inc_less:"\<lbrakk>W \<subseteq> carrier D; a \<in> carrier D;
y \<in> W; x \<in> segment (Iod D W) a; y \<prec> x\<rbrakk> \<Longrightarrow> y \<in> segment (Iod D W) a"
apply (frule Iod_Order[of "W"],
frule Order.segment_sub[of "Iod D W" "a"],
frule subsetD[of "segment (Iod D W) a" "carrier (Iod D W)" x],
assumption+, simp add:Iod_carrier)
apply (case_tac "a \<in> carrier (Iod D W)")
apply (subst Order.segment_inc[THEN sym, of "Iod D W" "y" "a"], assumption,
simp add:Iod_carrier, simp add:Iod_carrier)
apply (simp add:Iod_carrier, simp add:Iod_less)
apply (rule less_trans[of y x a], (simp add:subsetD)+)
apply (frule Order.segment_inc[THEN sym, of "Iod D W" "x" "a"],
(simp add:Iod_carrier)+,
frule_tac Order.segment_sub[of "Iod D W" x],
frule subsetD[of "segment (Iod D W) a" "W" "x"], assumption+,
simp add:Iod_carrier,
frule_tac subsetD[of "segment (Iod D W) a" W x], assumption+,
simp add:Iod_less)
apply (simp add:Order.segment_free[of "Iod D W" a], simp add:Iod_carrier)
done
lemma (in Order) segment_order_less:"\<forall>b\<in>carrier D. \<forall>x\<in> segment D b. \<forall>y\<in> segment D b. (x \<prec> y) = (x \<prec>\<^bsub>(Iod D (segment D b))\<^esub> y)"
by ((rule ballI)+,
cut_tac a = b in segment_sub, simp add:Iod_less)
lemma Ssegment_order_less:"Order D \<Longrightarrow>
\<forall>b\<in>carrier D. \<forall>x\<in> Ssegment D b. \<forall>y\<in> Ssegment D b.
(x \<prec>\<^bsub>D\<^esub> y) = (x \<prec>\<^bsub>(SIod D (Ssegment D b))\<^esub> y)"
by ((rule ballI)+,
cut_tac a = b in Ssegment_sub[of "D"], simp add:SIod_less)
lemma (in Order) segment_order_le:"\<forall>b\<in>carrier D. \<forall>x\<in> segment D b.
\<forall>y\<in> segment D b. (x \<preceq> y) = (x \<preceq>\<^bsub>(Iod D (segment D b))\<^esub> y)"
by ((rule ballI)+,
cut_tac a = b in segment_sub, simp add:Iod_le)
lemma Ssegment_order_le:"\<forall>b\<in>carrier D. \<forall>x\<in> Ssegment D b.
\<forall>y\<in> Ssegment D b. (x \<preceq>\<^bsub>D\<^esub> y) = (x \<preceq>\<^bsub>(SIod D (Ssegment D b))\<^esub> y)"
by ((rule ballI)+,
cut_tac a = b in Ssegment_sub[of "D"], simp add:SIod_le)
lemma (in Torder) Iod_Torder:"X \<subseteq> carrier D \<Longrightarrow> Torder (Iod D X)"
apply (rule Torder.intro)
apply (simp add:Iod_Order)
apply (rule Torder_axioms.intro)
apply (simp add:Iod_carrier Iod_le)
apply (meson contra_subsetD le_cases)
done
lemma SIod_Torder:"\<lbrakk>Torder D; X \<subseteq> carrier D\<rbrakk> \<Longrightarrow> Torder (SIod D X)"
apply (simp add:Torder_def, simp add:SIod_Order, simp add:Torder_axioms_def)
apply ((rule allI, rule impI)+,
simp add:SIod_carrier SIod_le) apply (erule conjE)
apply (frule_tac c = a in subsetD[of "X" "carrier D"], assumption+,
frule_tac c = b in subsetD[of "X" "carrier D"], assumption+)
apply blast
done
lemma (in Order) segment_not_inc:"\<lbrakk>a \<in> carrier D; b \<in> carrier D;
a \<prec> b\<rbrakk> \<Longrightarrow> b \<notin> segment D a"
apply (rule contrapos_pp, simp+, simp add:segment_def)
apply (simp add:oless_def, (erule conjE)+)
apply (frule le_antisym[of "a" "b"], assumption+, simp)
done
lemma Ssegment_not_inc:"\<lbrakk>Order D; a \<in> carrier D; b \<in> carrier D; a \<prec>\<^bsub>D\<^esub> b\<rbrakk> \<Longrightarrow>
b \<notin> Ssegment D a"
apply (rule contrapos_pp, simp+, simp add:Ssegment_def)
apply (simp add:oless_def, (erule conjE)+)
apply (frule Order.le_antisym[of "D" "a" "b"], assumption+, simp)
done
lemma (in Torder) segment_not_inc_iff:"\<lbrakk>a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<preceq> b) = (b \<notin> segment D a)"
apply (rule iffI)
apply (simp add:le_imp_less_or_eq,
erule disjE, simp add:segment_not_inc, simp add:a_notin_segment)
apply (simp add:segment_def, simp add:not_less_le[THEN sym, of "b" "a"])
done
lemma Ssegment_not_inc_iff:"\<lbrakk>Torder D; a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<preceq>\<^bsub>D\<^esub> b) = (b \<notin> Ssegment D a)"
apply (rule iffI)
apply (frule Torder.Order[of "D"])
apply (simp add:Order.le_imp_less_or_eq,
erule disjE, rule Ssegment_not_inc, assumption+)
apply (simp add: a_notin_Ssegment)
apply (simp add:Ssegment_def)
apply ( simp add:Torder.not_less_le[THEN sym, of "D" "b" "a"])
done
lemma (in Torder) minimum_segment_of_sub:"\<lbrakk>X \<subseteq> carrier D;
minimum_elem D (segment (Iod D X) d) m \<rbrakk> \<Longrightarrow> minimum_elem D X m"
apply (case_tac "d \<notin> carrier (Iod D X)")
apply (simp add:segment_def)
apply (simp add:Iod_carrier)
apply (simp add:Iod_carrier)
apply (subst minimum_elem_def)
apply (frule Iod_Order[of "X"],
frule Order.segment_sub[of "Iod D X" "d"],
simp add:Iod_carrier,
frule subset_trans[of "segment (Iod D X) d" "X" "carrier D"],
assumption+,
frule minimum_elem_mem[of "segment (Iod D X) d" m], assumption)
apply (simp add:subsetD[of "segment (Iod D X) d" "X" m])
apply (rule ballI)
apply (simp add:minimum_elem_def)
apply (case_tac "x \<in> segment (Iod D X) d")
apply (frule_tac a1 = x in Order.segment_inc[THEN sym, of "Iod D X" _ d])
apply (simp add:Iod_carrier subsetD)
apply (simp add:Iod_carrier)
apply (simp add:Iod_less)
apply (frule Iod_Torder[of "X"])
apply (frule_tac b1 = x in Torder.segment_not_inc_iff[THEN sym,
of "Iod D X" d])
apply (simp add:Iod_carrier)
apply (simp add:Iod_carrier)
apply simp
apply (frule Order.segment_inc[THEN sym, of "Iod D X" m d],
thin_tac "x \<notin> segment (Iod D X) d",
frule Order.segment_sub[of "Iod D X" "d"])
apply (simp add:Iod_carrier subsetD)
apply (simp add:Iod_carrier)
apply simp
apply (frule subsetD[of "segment (Iod D X) d" "X" m], assumption)
apply (simp add:Iod_le Iod_less)
apply (frule subsetD[of X "carrier D" m], assumption+,
frule subsetD[of X "carrier D" d], assumption+,
frule_tac c = x in subsetD[of X "carrier D"], assumption+)
apply (frule_tac c = x in less_le_trans[of m d], assumption+)
apply (simp add:less_imp_le)
done
lemma (in Torder) segment_out:"\<lbrakk>a \<in> carrier D; b \<in> carrier D;
a \<prec> b\<rbrakk> \<Longrightarrow> segment (Iod D (segment D a)) b = segment D a"
apply (subst segment_def[of "Iod D (segment D a)"])
apply (frule segment_not_inc[of "a" "b"], assumption+)
apply (cut_tac segment_sub[of "a"])
apply (simp add:Iod_carrier)
done
lemma (in Torder) segment_minimum_minimum:"\<lbrakk>X \<subseteq> carrier D; d \<in> X;
minimum_elem (Iod D (segment D d)) (X \<inter> (segment D d)) m\<rbrakk> \<Longrightarrow>
minimum_elem D X m"
apply (cut_tac segment_sub[of d])
apply (subst minimum_elem_def)
apply (cut_tac Order.minimum_elem_mem[of "Iod D (segment D d)"
"X \<inter> (segment D d)" m])
apply (cut_tac Int_lower1[of X "segment D d"],
frule_tac subsetD[of "X \<inter> segment D d" X m], assumption+, simp)
apply (rule ballI)
apply (case_tac "x \<in> segment D d")
apply (simp add:minimum_elem_def)
apply (drule_tac x = x in bspec,
(* thin_tac "Ball (X \<inter> segment D d) ((\<preceq>\<^bsub>Iod) D (segment D d)\<^esub> m)", *)
simp) apply (
simp add:Iod_le)
apply (frule subsetD[of X "carrier D" d], assumption+,
frule subsetD[of X "carrier D" m], assumption+,
frule_tac c = x in subsetD[of X "carrier D"], assumption+)
apply (simp add:segment_inc[THEN sym, of _ d],
simp add:not_less_le)
apply (frule_tac c = x in less_le_trans[of m d], assumption+)
apply (simp add:less_imp_le)
apply (simp add:Iod_Order)
apply (simp add:Iod_carrier)
apply (simp add:Int_lower2)
done
lemma (in Torder) segment_mono:"\<lbrakk>a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<prec> b) = (segment D a \<subset> segment D b)"
apply (rule iffI)
apply (rule psubsetI, rule subsetI)
apply (simp add:segment_def, erule conjE)
apply (rule_tac a = x and b = a and c = b in less_trans,
assumption+)
apply (cut_tac a_notin_segment[of "a"],
simp add:segment_inc[of "a" "b"], blast)
apply (simp add:psubset_eq, erule conjE,
frule not_sym[of "segment D a" "segment D b"],
thin_tac "segment D a \<noteq> segment D b",
frule sets_not_eq[of "segment D b" "segment D a"], assumption+)
apply (erule bexE)
apply (thin_tac "segment D a \<subseteq> segment D b",
thin_tac "segment D b \<noteq> segment D a")
apply (simp add:segment_def, (erule conjE)+)
apply (frule_tac a = aa and b = a in not_less_le, assumption+,
simp, simp add:oless_def, (erule conjE)+)
apply (frule_tac a = a and b = aa and c = b in le_trans,
assumption+, simp)
apply (rule contrapos_pp, simp+)
done
lemma Ssegment_mono:"\<lbrakk>Torder D; a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<prec>\<^bsub>D\<^esub> b) = (Ssegment D a \<subset> Ssegment D b)"
apply (frule Torder.Order)
apply (rule iffI)
apply (rule psubsetI, rule subsetI)
apply (simp add:Ssegment_def, erule conjE)
apply (rule_tac a = x and b = a and c = b in Order.less_trans,
assumption+)
apply (cut_tac a_notin_Ssegment[of "a"],
simp add:Ssegment_inc[of "D" "a" "b"], blast)
apply (simp add:psubset_eq, erule conjE,
frule not_sym[of "Ssegment D a" "Ssegment D b"],
thin_tac "Ssegment D a \<noteq> Ssegment D b",
frule sets_not_eq[of "Ssegment D b" "Ssegment D a"], assumption+)
apply (erule bexE)
apply (thin_tac "Ssegment D a \<subseteq> Ssegment D b",
thin_tac "Ssegment D b \<noteq> Ssegment D a")
apply (simp add:Ssegment_def, (erule conjE)+)
apply (frule_tac a = aa and b = a in Torder.not_less_le, assumption+,
simp, simp add:oless_def, (erule conjE)+)
apply (frule_tac a = a and b = aa and c = b in Order.le_trans,
assumption+, simp)
apply (rule contrapos_pp, simp+)
done
lemma (in Torder) segment_le_mono:"\<lbrakk>a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<preceq> b) = (segment D a \<subseteq> segment D b)"
apply (simp add:le_imp_less_or_eq[of "a" "b"])
apply (rule iffI)
apply (erule disjE)
apply (simp add:segment_mono[of "a" "b"], simp)
apply (frule segment_mono[THEN sym, of "a" "b"], assumption+)
apply (simp add:psubset_eq)
apply (case_tac "segment D a \<noteq> segment D b", simp)
apply simp
apply (rule contrapos_pp, simp+,
frule less_linear[of "a" "b"], assumption+, simp,
simp add:segment_mono[of "b" "a"])
done
lemma Ssegment_le_mono:"\<lbrakk>Torder D; a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<preceq>\<^bsub>D\<^esub> b) = (Ssegment D a \<subseteq> Ssegment D b)"
apply (cut_tac Torder.Order[of "D"])
apply (simp add:Order.le_imp_less_or_eq[of "D" "a" "b"])
apply (rule iffI)
apply (erule disjE)
apply (simp add: Ssegment_mono[of "D" "a" "b"])
apply (frule Ssegment_mono[THEN sym, of "D" "a" "b"], assumption+)
apply (simp add:psubset_eq)
apply (case_tac "Ssegment D a \<noteq> Ssegment D b")
apply (cut_tac Ssegment_mono[THEN sym, of "D" "a" "b"])
apply (simp add:psubset_eq, assumption+)
apply simp
apply (cut_tac a_notin_Ssegment[of "a" "D"], simp)
apply (simp add:Ssegment_not_inc_iff[THEN sym, of "D" "b" "a"])
apply (frule sym, thin_tac "Ssegment D a = Ssegment D b")
apply (cut_tac a_notin_Ssegment[of "b" "D"], simp)
apply (simp add:Ssegment_not_inc_iff[THEN sym, of "D" "a" "b"])
apply (frule Order.le_antisym[of "D" "a" "b"], assumption+, simp+)
done
lemma (in Torder) segment_inj:"\<lbrakk>a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a = b) = (segment D a = segment D b)"
apply (rule iffI)
apply simp
apply (rule equalityE[of "segment D a" "segment D b"], assumption)
apply (thin_tac "segment D a = segment D b")
apply (simp add:segment_le_mono[THEN sym, of "a" "b"])
apply (simp add:segment_le_mono[THEN sym, of "b" "a"])
apply (simp add:le_antisym)
done
lemma Ssegment_inj:"\<lbrakk>Torder D; a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a = b) = (Ssegment D a = Ssegment D b)"
apply (rule iffI)
apply simp
apply (rule equalityE[of "Ssegment D a" "Ssegment D b"], assumption)
apply (thin_tac "Ssegment D a = Ssegment D b")
apply (simp add:Ssegment_le_mono[THEN sym, of "D" "a" "b"])
apply (simp add:Ssegment_le_mono[THEN sym, of "D" "b" "a"])
apply (cut_tac Torder.Order[of "D"])
apply (simp add:Order.le_antisym, assumption)
done
lemma (in Torder) segment_inj_neq:"\<lbrakk>a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<noteq> b) = (segment D a \<noteq> segment D b)"
by (simp add:segment_inj)
lemma Ssegment_inj_neq:"\<lbrakk>Torder D; a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<noteq> b) = (Ssegment D a \<noteq> Ssegment D b)"
by (simp add:Ssegment_inj)
lemma (in Order) segment_inc_psub:"\<lbrakk>x \<in> segment D a\<rbrakk> \<Longrightarrow>
segment D x \<subset> segment D a"
apply (simp add:psubset_eq)
apply (rule conjI, rule subsetI)
apply (simp add:segment_def)
apply (case_tac "a \<notin> carrier D", simp)
apply (simp, (erule conjE)+)
apply (rule_tac a = xa and b = x and c = a in less_trans, assumption+)
apply (cut_tac a_notin_segment[of "x"]) apply blast
done
lemma Ssegment_inc_psub:"\<lbrakk>Order D; x \<in> Ssegment D a\<rbrakk> \<Longrightarrow>
Ssegment D x \<subset> Ssegment D a"
apply (simp add:psubset_eq)
apply (rule conjI, rule subsetI)
apply (simp add:Ssegment_def)
apply (case_tac "a \<notin> carrier D", simp)
apply (simp, (erule conjE)+)
apply (rule_tac a = xa and b = x and c = a in Order.less_trans[of "D"],
assumption+)
apply (cut_tac a_notin_Ssegment[of "x"]) apply blast
done
lemma (in Order) segment_segment:"\<lbrakk>b \<in> carrier D; a \<in> segment D b\<rbrakk> \<Longrightarrow>
segment (Iod D (segment D b)) a = segment D a"
apply (rule equalityI)
apply (rule subsetI)
apply (simp add:segment_def[of "Iod D (segment D b)" "a"])
apply (cut_tac segment_sub[of "b"], simp add:Iod_carrier)
apply (erule conjE)
apply (simp add:Iod_less)
apply (frule_tac c = x in subsetD[of "segment D b" "carrier D"], assumption+,
frule_tac c = a in subsetD[of "segment D b" "carrier D"], assumption+)
apply (simp add:segment_inc[of _ "a"])
apply (rule subsetI)
apply (simp add:segment_def[of "Iod D (segment D b)" "a"])
apply (cut_tac segment_sub[of "b"], simp add:Iod_carrier)
apply (frule segment_inc_psub[of "a" "b"],
frule psubset_imp_subset[of "segment D a" "segment D b"],
thin_tac "segment D a \<subset> segment D b",
frule_tac c = x in subsetD[of "segment D a" "segment D b"],
assumption+)
apply (simp add:Iod_less) apply (simp add:segment_def)
done
lemma Ssegment_Ssegment:"\<lbrakk>Order D; b \<in> carrier D; a \<in> Ssegment D b\<rbrakk> \<Longrightarrow>
Ssegment (SIod D (Ssegment D b)) a = Ssegment D a"
apply (rule equalityI)
apply (rule subsetI)
apply (simp add:Ssegment_def[of "SIod D (Ssegment D b)" "a"])
apply (cut_tac Ssegment_sub[of "D" "b"], simp add:SIod_carrier)
apply (erule conjE)
apply (simp add:SIod_less)
apply (frule_tac c = x in subsetD[of "Ssegment D b" "carrier D"], assumption+,
frule_tac c = a in subsetD[of "Ssegment D b" "carrier D"], assumption+)
apply (simp add:Ssegment_inc[of "D"_ "a"])
apply (rule subsetI)
apply (simp add:Ssegment_def[of "SIod D (Ssegment D b)" "a"])
apply (cut_tac Ssegment_sub[of "D" "b"], simp add:SIod_carrier)
apply (frule Ssegment_inc_psub[of "D" "a" "b"], assumption,
frule psubset_imp_subset[of "Ssegment D a" "Ssegment D b"],
thin_tac "Ssegment D a \<subset> Ssegment D b",
frule_tac c = x in subsetD[of "Ssegment D a" "Ssegment D b"],
assumption+)
apply (simp add:SIod_less) apply (simp add:Ssegment_def)
done
lemma (in Order) Iod_segment_segment:"a \<in> carrier (Iod D (segment D b)) \<Longrightarrow>
Iod (Iod D (segment D b)) (segment (Iod D (segment D b)) a) =
Iod D (segment D a)"
apply (case_tac "b \<in> carrier D")
apply (cut_tac segment_sub[of "b"])
apply (simp add:Iod_carrier)
apply (frule segment_inc_psub[of "a" "b"],
frule psubset_imp_subset[of "segment D a" "segment D b"],
thin_tac "segment D a \<subset> segment D b")
apply (simp add:segment_segment[of "b" "a"])
apply (simp add:Iod_sub_sub[of "segment D a" "segment D b"])
apply (simp add:segment_def[of D b])
apply (simp add:Iod_self[THEN sym])
done
lemma SIod_Ssegment_Ssegment:"\<lbrakk>Order D; a \<in> carrier (SIod D (Ssegment D b))\<rbrakk>
\<Longrightarrow>
SIod (SIod D (Ssegment D b)) (Ssegment (SIod D (Ssegment D b)) a) =
SIod D (Ssegment D a)"
apply (case_tac "b \<in> carrier D")
apply (cut_tac Ssegment_sub[of "D" "b"])
apply (simp add:SIod_carrier[of "D"])
apply (frule Ssegment_inc_psub[of "D" "a" "b"], simp add:subsetD) apply (
frule psubset_imp_subset[of "Ssegment D a" "Ssegment D b"],
thin_tac "Ssegment D a \<subset> Ssegment D b")
apply (simp add:Ssegment_Ssegment[of "D" "b" "a"])
apply (simp add:SIod_sub_sub[of "Ssegment D a" "Ssegment D b"])
apply (simp add:Ssegment_def[of D b], simp add:SIod_self[THEN sym])
done
lemma (in Order) ord_isom_segment_mem:"\<lbrakk>Order E;
ord_isom D E f; a \<in> carrier D; x \<in> segment D a \<rbrakk> \<Longrightarrow>
(f x) \<in> segment E (f a)"
apply (frule segment_inc_if[of "a" "x"], assumption+)
apply (frule ord_isom_less[of "E" "f" "x" "a"], assumption+)
apply (simp add:segment_def, assumption, simp)
apply (frule ord_isom_mem[of "E" "f" "x"], assumption+,
simp add:segment_def,
frule ord_isom_mem[of "E" "f" "a"], assumption+)
apply (simp add:Order.segment_inc[of "E" "f x" "f a"])
done
lemma ord_isom_Ssegment_mem:"\<lbrakk>Order D; Order E;
ord_isom D E f; a \<in> carrier D; x \<in> Ssegment D a\<rbrakk> \<Longrightarrow>
(f x) \<in> Ssegment E (f a)"
apply (frule Ssegment_inc_if[of "D" "a" "x"], assumption+)
apply (frule Order.ord_isom_less[of "D" "E" "f" "x" "a"], assumption+)
apply (simp add:Ssegment_def, assumption, simp)
apply (frule Order.ord_isom_mem[of "D" "E" "f" "x"], assumption+,
simp add:Ssegment_def,
frule Order.ord_isom_mem[of "D" "E" "f" "a"], assumption+)
apply (simp add:Ssegment_def)
done
lemma (in Order) ord_isom_segment_segment:"\<lbrakk>Order E;
ord_isom D E f; a \<in> carrier D \<rbrakk> \<Longrightarrow>
ord_isom (Iod D (segment D a)) (Iod E (segment E (f a)))
(\<lambda>x\<in>carrier (Iod D (segment D a)). f x)"
apply (frule ord_isom_inj_on[of E f], assumption+)
apply (cut_tac segment_sub[of a])
apply (frule restrict_inj[of f "carrier D" "segment D a"], assumption)
apply (frule ord_isom_surj_to[of E f], assumption+)
apply (subst ord_isom_def, subst ord_inj_def)
apply (simp add:Iod_carr_segment Order.Iod_carr_segment)
apply (subgoal_tac "restrict f (segment D a) \<in>
segment D a \<rightarrow> segment E (f a)", simp)
defer
apply (simp add:ord_isom_segment_mem)
apply (rule conjI)
defer
apply (rule surj_to_test, assumption+)
apply (rule ballI, simp)
apply (frule ord_isom_func[of E f], assumption+)
apply (frule surj_to_el[of f "carrier D" "carrier E"], assumption+,
frule ord_isom_mem[of E f a], assumption+,
frule Order.segment_sub[of E "f a"],
frule_tac c = b in subsetD[of "segment E (f a)" "carrier E"],
assumption+,
drule_tac x = b in bspec, assumption, (*
thin_tac "\<forall>b\<in>carrier E. \<exists>a\<in>carrier D. f a = b", *)
erule bexE)
apply (simp add:Order.segment_inc[THEN sym, of E _ "f a"],
rotate_tac -1, frule sym, thin_tac "f aa = b", simp,
frule_tac a1 = aa and b1 = a in ord_isom_less[THEN sym, of E f],
assumption+, simp,
simp add:segment_inc[of _ a], blast)
apply (rule ballI)+
apply (frule ord_isom_mem[of E f a], assumption+,
frule Order.segment_sub[of E "f a"])
apply (frule_tac x = aa in ord_isom_segment_mem[of E f a], assumption+,
frule_tac x = b in ord_isom_segment_mem[of E f a], assumption+,
simp add:Iod_less Order.Iod_less,
subst ord_isom_less[of E f], assumption+, (simp add:subsetD)+)
done
lemma ord_isom_Ssegment_Ssegment:"\<lbrakk>Order D; Order E;
ord_isom D E f; a \<in> carrier D \<rbrakk> \<Longrightarrow>
ord_isom (SIod D (Ssegment D a)) (SIod E (Ssegment E (f a)))
(\<lambda>x\<in>carrier (SIod D (Ssegment D a)). f x)"
apply (frule_tac a = a in Order.ord_isom_mem[of D E f], assumption+)
apply (cut_tac Ssegment_sub[of D a],
cut_tac Ssegment_sub[of "E" "f a"])
apply (subst ord_isom_def, simp add:ord_inj_def)
apply (rule conjI)
apply (rule Pi_I)
apply (simp add:SIod_carrier)
apply (frule_tac c = x in subsetD[of "Ssegment D a" "carrier D"], assumption+)
apply (frule_tac a = x in Order.ord_isom_mem[of D E f], assumption+)
apply (subst Ssegment_inc[THEN sym, of "E" _ "f a"], assumption+)
apply (subst Order.ord_isom_less[THEN sym, of D E f _ a], assumption+)
apply (subst Ssegment_inc[of D _ a], assumption+)
apply (rule conjI)
apply (simp add:SIod_carrier)
apply (simp add:ord_isom_def bij_to_def, (erule conjE)+)
apply (simp add:ord_inj_def, (erule conjE)+)
apply (rule restrict_inj[of "f" "carrier D" "Ssegment D a"], assumption+)
apply (rule conjI)
apply (rule ballI)+
apply (simp add:SIod_carrier)
apply (frule_tac c = aa in subsetD[of "Ssegment D a" "carrier D"],
assumption+,
frule_tac c = b in subsetD[of "Ssegment D a" "carrier D"], assumption+)
apply (frule_tac a1 = aa and b1 = a in Ssegment_inc[THEN sym], assumption+,
frule_tac a1 = b and b1 = a in Ssegment_inc[THEN sym], assumption+,
simp)
apply (simp add:Order.ord_isom_less[of D E f])
apply (frule_tac a = a in Order.ord_isom_mem[of D E f], assumption+,
frule_tac a = aa in Order.ord_isom_mem[of D E f], assumption+,
frule_tac a = b in Order.ord_isom_mem[of D E f], assumption+)
apply (simp add:Ssegment_inc[of E])
apply (simp add:SIod_less Order.ord_isom_less)
apply (simp add:surj_to_def,
simp add:SIod_carrier)
apply (rule equalityI)
apply (rule subsetI, simp add:image_def, erule bexE)
apply (frule_tac c = xa in subsetD[of "Ssegment D a" "carrier D"],
assumption+)
apply (frule_tac a = xa in Ssegment_inc[of D _ a], assumption+, simp)
apply (simp add:Order.ord_isom_less[of D E f _ a])
apply (frule_tac a = xa in Order.ord_isom_mem[of D E f], assumption+)
apply (subst Ssegment_inc[THEN sym], assumption+)
apply (rule subsetI)
apply (frule_tac c = x in subsetD[of "Ssegment E (f a)" "carrier E"],
assumption+)
apply (simp add:Ssegment_inc[THEN sym])
apply (frule_tac b = x in Order.ord_isom_surj[of D E f], assumption+,
erule bexE, simp, thin_tac "x = f aa")
apply (simp add:Order.ord_isom_less[THEN sym])
apply (simp add:Ssegment_inc[of D])
done
lemma (in Order) ord_equiv_segment_segment:
"\<lbrakk>Order E; ord_equiv D E; a \<in> carrier D\<rbrakk>
\<Longrightarrow> \<exists>t\<in>carrier E. ord_equiv (Iod D (segment D a)) (Iod E (segment E t))"
apply (simp add:ord_equiv_def, erule exE)
apply (frule_tac f = f in ord_isom_segment_segment[of E _ a], assumption+)
apply (frule_tac f = f in ord_isom_mem[of E _ a], assumption+)
apply blast
done
lemma ord_equiv_Ssegment_Ssegment:
"\<lbrakk>Order D; Order E; ord_equiv D E; a \<in> carrier D\<rbrakk>
\<Longrightarrow> \<exists>t\<in>carrier E. ord_equiv (SIod D (Ssegment D a)) (SIod E (Ssegment E t))"
apply (simp add:ord_equiv_def, erule exE)
apply (frule_tac f = f in ord_isom_Ssegment_Ssegment[of "D" "E" _ "a"],
assumption+)
apply (frule_tac f = f in Order.ord_isom_mem[of D E _ a], assumption+)
apply blast
done
lemma (in Order) ord_isom_restricted:
"\<lbrakk>Order E; ord_isom D E f; D1 \<subseteq> carrier D\<rbrakk> \<Longrightarrow>
ord_isom (Iod D D1) (Iod E (f ` D1)) (\<lambda>x\<in>D1. f x)"
apply (simp add:ord_isom_def[of D E f], erule conjE)
apply (simp add:ord_inj_restrict_isom[of E f D1])
done
lemma ord_isom_restrictedS:
"\<lbrakk>Order D; Order E; ord_isom D E f; D1 \<subseteq> carrier D\<rbrakk> \<Longrightarrow>
ord_isom (SIod D D1) (SIod E (f ` D1)) (\<lambda>x\<in>D1. f x)"
apply (simp add:ord_isom_def[of D E f], erule conjE)
apply (simp add:ord_inj_Srestrict_isom[of D E f D1])
done
lemma (in Order) ord_equiv_induced:
"\<lbrakk>Order E; ord_isom D E f; D1 \<subseteq> carrier D \<rbrakk> \<Longrightarrow>
ord_equiv (Iod D D1) (Iod E (f ` D1))"
apply (simp add:ord_equiv_def)
apply (frule ord_isom_restricted [of "E" "f" "D1"], assumption+)
apply blast
done
lemma ord_equiv_inducedS:
"\<lbrakk>Order D; Order E; ord_isom D E f; D1 \<subseteq> carrier D \<rbrakk> \<Longrightarrow>
ord_equiv (SIod D D1) (SIod E (f ` D1))"
apply (simp add:ord_equiv_def)
apply (frule ord_isom_restrictedS [of "D" "E" "f" "D1"], assumption+)
apply blast
done
lemma (in Order) equiv_induced_by_inj:"\<lbrakk>Order E; ord_inj D E f;
D1 \<subseteq> carrier D\<rbrakk> \<Longrightarrow> ord_equiv (Iod D D1) (Iod E (f ` D1))"
apply (simp add:ord_equiv_def)
apply (frule ord_inj_restrict_isom [of E f D1], assumption+)
apply blast
done
lemma equiv_induced_by_injS:"\<lbrakk>Order D; Order E; ord_inj D E f;
D1 \<subseteq> carrier D\<rbrakk> \<Longrightarrow> ord_equiv (SIod D D1) (SIod E (f ` D1))"
apply (simp add:ord_equiv_def)
apply (frule ord_inj_Srestrict_isom[of D E f D1], assumption+)
apply blast
done
lemma (in Torder) le_segment_segment:"\<lbrakk>a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<preceq> b) = (segment (Iod D (segment D b)) a = segment D a)"
apply (cut_tac segment_sub[of b],
frule Iod_Order[of "segment D b"])
apply (case_tac "a = b") apply simp
apply (simp add:le_refl)
apply ( cut_tac a_notin_segment[of "b"])
apply (subst Order.segment_free[of "Iod D (segment D b)" b], assumption)
apply (simp add:Iod_carrier)
apply (simp add:Iod_carrier)
apply (subst le_imp_less_or_eq[of "a" "b"], assumption+, simp)
apply (rule iffI)
apply (rule equalityI)
apply (rule subsetI)
apply (frule_tac a1 = x in Order.segment_inc[THEN sym,
of "Iod D (segment D b)" _ a])
apply (frule_tac Order.segment_sub[of "Iod D (segment D b)" a])
apply (rule subsetD, assumption+)
apply (simp add:Iod_carrier) apply (simp add:segment_inc)
apply simp
apply (subst segment_inc[THEN sym])
apply (simp add:segment_def Iod_def) apply assumption
apply (simp add:segment_inc)
apply (frule Order.segment_sub[of "Iod D (segment D b)" a])
apply (simp add:Iod_carrier)
apply (simp add:subsetD Iod_less)
apply (rule subsetI)
apply (subst Order.segment_inc[THEN sym, of "Iod D (segment D b)"],
assumption+)
apply (simp add:Iod_carrier)
apply (simp add:segment_mono[of a b] psubset_eq, erule conjE)
apply (rule subsetD[of "segment D a" "segment D b"], assumption+)
apply (simp add:Iod_carrier segment_inc)
apply (frule segment_inc[of a b], assumption, simp)
apply (frule segment_mono[of a b], assumption, simp)
apply (simp add:psubset_eq, (erule conjE)+)
apply (frule_tac c = x in subsetD[of "segment D a" "segment D b"],
assumption+)
apply (simp add:Iod_less)
apply (subst segment_inc) apply (simp add:subsetD) apply assumption+
apply (rule contrapos_pp, simp+)
apply (simp add:not_less_le)
apply (simp add:le_imp_less_or_eq)
apply (frule segment_not_inc[of b a], assumption+)
apply (frule Order.segment_free[of "Iod D (segment D b)" a])
apply (simp add:Iod_carrier)
apply (simp add:Iod_carrier)
apply (simp add:segment_inj[THEN sym, of b a])
done
lemma le_Ssegment_Ssegment:"\<lbrakk>Torder D; a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(a \<preceq>\<^bsub>D\<^esub> b) = (Ssegment (SIod D (Ssegment D b)) a = Ssegment D a)"
apply (frule Torder.Order[of "D"])
apply (case_tac "a = b") apply simp
apply (simp add:Order.le_refl)
apply (cut_tac Ssegment_sub[of "D" "b"])
apply (frule SIod_Order[of "D" "Ssegment D b"], assumption)
apply (cut_tac a_notin_Ssegment[of "b" "D"])
apply (frule SIod_carrier[THEN sym, of "D" "Ssegment D b"], assumption+)
apply (frule eq_set_not_inc[of "b" "Ssegment D b"
"carrier (SIod D (Ssegment D b))"], assumption+)
apply (thin_tac "b \<notin> Ssegment D b",
thin_tac "Ssegment D b = carrier (SIod D (Ssegment D b))")
apply (cut_tac Ssegment_free[of "b" "SIod D (Ssegment D b)" ])
apply (simp add:SIod_carrier) apply assumption+
apply (subst Order.le_imp_less_or_eq[of "D" "a" "b"], assumption+)
apply simp
apply (cut_tac Ssegment_sub[of "D" "b"])
apply (subst Ssegment_def[of "SIod D (Ssegment D b)"],
subst SIod_carrier[of "D" "Ssegment D b"], assumption+)
apply (subst Ssegment_inc[of "D" "a" "b"], assumption+)
apply (rule iffI) apply simp
apply (simp add:SIod_carrier)
apply (rule equalityI)
apply (rule subsetI)
apply (simp, erule conjE)
apply (simp add:SIod_less)
apply (subst Ssegment_def, simp add:Ssegment_def)
apply (rule subsetI, simp)
apply (simp add:Ssegment_inc[THEN sym, of "D" "a" "b"])
apply (cut_tac a1 = x in Ssegment_inc[THEN sym, of "D" _ "a"], assumption+)
apply (simp add:Ssegment_def, assumption, simp)
apply (cut_tac a = x in Order.less_trans[of "D" _ "a" "b"], assumption)
apply (simp add:Ssegment_def, assumption+)
apply (cut_tac a = x in Ssegment_inc[of "D" _ "b"], assumption)
apply (simp add:Ssegment_def)
apply assumption+
apply simp
apply (cut_tac a = a in Ssegment_inc[of "D" _ "b"])
apply assumption+
apply simp
apply (simp add:SIod_less)
apply (rule contrapos_pp, simp+)
apply (simp add:SIod_carrier)
apply (frule sym, thin_tac "Ssegment D b = Ssegment D a", simp)
apply (simp add:Ssegment_inc[THEN sym, of "D" "a" "b"])
apply (simp add:Torder.not_less_le[of "D" "a" "b"])
apply (frule not_sym, thin_tac "a \<noteq> b")
apply (simp add:Order.le_imp_less_or_eq[of "D" "b" "a"])
apply (simp add:Ssegment_inc[of "D" "b" "a"])
apply (simp add:a_notin_Ssegment[of "b" "D"])
done
lemma (in Torder) inc_segment_segment:"\<lbrakk>b \<in> carrier D;
a \<in> segment D b\<rbrakk> \<Longrightarrow> segment (Iod D (segment D b)) a = segment D a"
apply (cut_tac segment_sub[of "b"],
frule subsetD[of "segment D b" "carrier D" "a"], assumption)
apply (subst le_segment_segment[THEN sym, of "a" "b"],
assumption+)
apply (simp add:segment_inc[THEN sym])
apply (simp add:less_imp_le)
done
lemma (in Torder) segment_segment:"\<lbrakk>a \<in> carrier D; b \<in> carrier D\<rbrakk> \<Longrightarrow>
(segment (Iod D (segment D b)) a = segment D a) =
((segment D a) \<subseteq> (segment D b))"
apply (subst le_segment_segment[THEN sym, of "a" "b"],
assumption+)
apply (simp add:segment_le_mono[of "a" "b"])
done
lemma (in Torder) less_in_Iod:"\<lbrakk>a \<in> carrier D; b \<in> carrier D; a \<prec> b\<rbrakk>
\<Longrightarrow> (a \<prec> b) = (a \<in> carrier (Iod D (segment D b)))"
apply (simp add:Iod_def segment_inc)
done
definition
SS :: "_ \<Rightarrow> 'a set Order" where
"SS D = \<lparr>carrier = {X. \<exists>a\<in>carrier D. X = segment D a}, rel =
{XX. XX \<in> {X. \<exists>a\<in>carrier D. X = segment D a} \<times>
{X. \<exists>a\<in>carrier D. X = segment D a} \<and> ((fst XX) \<subseteq> (snd XX))} \<rparr>"
(** Ordered set consisting of segments **)
definition
segmap::"_ \<Rightarrow> 'a \<Rightarrow> 'a set" where
"segmap D = (\<lambda>x\<in>(carrier D). segment D x)"
lemma segmap_func:"segmap D \<in> carrier D \<rightarrow> carrier (SS D)"
by (simp add:SS_def segmap_def Pi_def) blast
lemma (in Worder) ord_isom_segmap:" ord_isom D (SS D) (segmap D)"
apply (simp add:ord_isom_def)
apply (rule conjI)
apply (simp add:ord_inj_def)
apply (rule conjI)
apply (simp add:segmap_def)
apply (rule conjI)
apply (simp add:segmap_func)
apply (rule conjI)
apply (simp add:inj_on_def)
apply ((rule ballI)+, rule impI, simp add:segmap_def,
simp add:segment_inj[THEN sym])
apply (rule ballI)+
apply (simp add:oless_def[of "SS D"]) apply (simp add:ole_def SS_def)
apply (rule iffI)
apply (simp add:oless_def, erule conjE)
apply (frule_tac a = a and b = b in segment_le_mono, assumption+)
apply (simp add:segment_inj segmap_def)
apply blast
apply (erule conjE)+
apply (thin_tac "\<exists>aa\<in>carrier D. segmap D a = segment D aa",
thin_tac " \<exists>a\<in>carrier D. segmap D b = segment D a")
apply (simp add:segmap_def segment_inj[THEN sym])
apply (simp add:segment_le_mono[THEN sym])
apply (simp add:oless_def)
apply (rule surj_to_test[of "segmap D" "carrier D" "carrier (SS D)"])
apply (simp add:segmap_func)
apply (rule ballI)
apply (simp add:SS_def, erule bexE, simp)
apply (simp add:segmap_def, blast)
done
lemma (in Worder) nonequiv_segment:"a \<in> carrier D \<Longrightarrow>
\<not> ord_equiv D (Iod D (segment D a))"
apply (rule contrapos_pp, simp+)
apply (simp add:ord_equiv_def)
apply (erule exE)
apply (cut_tac segment_sub[of "a"])
apply (frule Iod_Order[of "segment D a"])
apply (frule_tac f = f in ord_isom_func[of "Iod D (segment D a)"],
assumption+)
apply (frule_tac f = f and a = a in ord_isom_mem[of "Iod D (segment D a)"]
, assumption+)
apply (frule_tac f = f in to_subset [of "segment D a"], assumption+)
apply (drule_tac a = a in forall_spec, assumption) (*
apply (thin_tac "\<forall>a. a \<in> carrier D \<longrightarrow> a \<preceq> (f a)") *)
apply (simp add:Iod_carrier)
apply (frule_tac c = "f a" in subsetD[of "segment D a" "carrier D" ],
assumption+)
apply (simp add:segment_inc[THEN sym])
apply (simp add:not_le_less[THEN sym, of "a" _])
done
lemma nonequiv_Ssegment:"\<lbrakk>Worder D; a \<in> carrier D\<rbrakk> \<Longrightarrow>
\<not> ord_equiv D (SIod D (Ssegment D a))"
apply (frule Worder.Order[of "D"], frule Worder.Torder[of "D"])
apply (rule contrapos_pp, simp+)
apply (simp add:ord_equiv_def)
apply (erule exE)
apply (cut_tac Ssegment_sub[of "D" "a"])
apply (frule SIod_Order[of "D" "Ssegment D a"], assumption)
apply (frule_tac f = f in Order.ord_isom_func[of "D" "SIod D (Ssegment D a)"],
assumption+,
frule_tac f = f and a = a in Order.ord_isom_mem[of "D"
"SIod D (Ssegment D a)"], assumption+)
apply (frule_tac f = f in to_subsetS [of "D" "Ssegment D a"], assumption+)
apply (drule_tac a = a in forall_spec, assumption) (*
thin_tac "\<forall>a. a \<in> carrier D \<longrightarrow> a \<preceq>\<^bsub>D\<^esub> f a") *)
apply (simp add:SIod_carrier)
apply (frule_tac c = "f a" in subsetD[of "Ssegment D a" "carrier D"],
assumption+)
apply (simp add:Ssegment_inc[THEN sym])
apply (simp add:Torder.not_le_less[THEN sym, of "D" "a" _])
done
lemma (in Worder) subset_Worder:" T \<subseteq> carrier D \<Longrightarrow>
Worder (Iod D T)"
apply (rule Worder.intro)
apply (simp add: Iod_Torder)
apply (rule Worder_axioms.intro)
apply (rule allI, rule impI)
apply (simp add:Iod_carrier, erule conjE)
apply (cut_tac ex_minimum)
apply (frule_tac A = X and B = T and C = "carrier D" in subset_trans,
assumption+)
apply (frule_tac a = X in forall_spec, simp,
thin_tac "\<forall>X. X \<subseteq> carrier D \<and> X \<noteq> {} \<longrightarrow> (\<exists>x. minimum_elem D X x)")
apply (erule exE)
apply (simp add:minimum_elem_sub)
apply blast
done
lemma SIod_Worder:"\<lbrakk>Worder D; T \<subseteq> carrier D\<rbrakk> \<Longrightarrow> Worder (SIod D T)"
apply (frule Worder.Order[of "D"],
frule Worder.Torder[of "D"])
apply (rule Worder.intro)
apply (simp add: SIod_Torder)
apply (rule Worder_axioms.intro)
apply (rule allI, rule impI, erule conjE, simp add:SIod_carrier)
apply (frule Worder.ex_minimum)
apply (frule_tac A = X and B = T and C = "carrier D" in subset_trans,
assumption+)
apply (frule_tac a = X in forall_spec, simp,
thin_tac "\<forall>X. X \<subseteq> carrier D \<and> X \<noteq> {} \<longrightarrow> (\<exists>x. minimum_elem D X x)")
apply (simp add:minimum_elem_Ssub)
done
lemma (in Worder) segment_Worder:"Worder (Iod D (segment D a))"
apply (rule subset_Worder [of "segment D a"])
apply (rule segment_sub[of a])
done
lemma Ssegment_Worder:"Worder D \<Longrightarrow>Worder (SIod D (Ssegment D a))"
apply (rule SIod_Worder, assumption)
apply (rule Ssegment_sub[of "D" "a"])
done
lemma (in Worder) segment_unique1:"\<lbrakk>a \<in> carrier D; b \<in> carrier D; a \<prec> b\<rbrakk> \<Longrightarrow>
\<not> ord_equiv (Iod D (segment D b)) (Iod D (segment D a))"
apply (cut_tac segment_Worder[of b],
cut_tac segment_sub[of b],
frule segment_mono[of a b], assumption, simp add:psubset_eq,
erule conjE)
apply (simp add:segment_inc,
frule Worder.nonequiv_segment[of "Iod D (segment D b)" a],
simp add:Iod_carrier)
apply (frule segment_segment[THEN sym, of a b], assumption, simp)
apply (simp add:Iod_sub_sub[of "segment D a" "segment D b"])
done
lemma Ssegment_unique1:"\<lbrakk>Worder D; a \<in> carrier D; b \<in> carrier D; a \<prec>\<^bsub>D\<^esub> b\<rbrakk> \<Longrightarrow>
\<not> ord_equiv (SIod D (Ssegment D b)) (SIod D (Ssegment D a))"
apply (frule Worder.Order[of "D"], frule Worder.Torder[of "D"],
frule Ssegment_inc[of "D" "a" "b"], assumption+, simp,
frule Ssegment_Worder [of "D" "b"])
apply (cut_tac Ssegment_sub[of "D" "b"]) apply (
frule Ssegment_mono[of D a b], assumption+, simp)
apply (frule nonequiv_Ssegment[of "SIod D (Ssegment D b)" "a"])
apply (simp add:SIod_carrier)
apply (frule le_Ssegment_Ssegment[of D a b], assumption+)
apply (simp add:oless_def psubset_eq, (erule conjE)+)
apply (simp add:SIod_sub_sub[of "Ssegment D a" "Ssegment D b"])
done
lemma (in Worder) segment_unique:"\<lbrakk>a \<in> carrier D; b \<in> carrier D;
ord_equiv (Iod D (segment D a)) (Iod D (segment D b)) \<rbrakk> \<Longrightarrow> a = b"
apply (cut_tac segment_sub[of a],
frule_tac Iod_Order[of "segment D a"],
cut_tac segment_sub[of b],
frule_tac Iod_Order[of "segment D b"])
apply (rule contrapos_pp, simp+)
apply (frule less_linear[of "a" "b"], assumption+)
apply simp
apply (erule disjE)
apply (frule segment_unique1[of "a" "b"], assumption+)
apply (simp add:Order.ord_equiv_sym[of "Iod D (segment D a)"
"Iod D (segment D b)"])
apply (simp add:segment_unique1[of "b" "a"])
done
lemma Ssegment_unique:"\<lbrakk>Worder D; a \<in> carrier D; b \<in> carrier D;
ord_equiv (SIod D (Ssegment D a)) (SIod D (Ssegment D b)) \<rbrakk> \<Longrightarrow> a = b"
apply (frule Worder.Order[of "D"], frule Worder.Torder[of "D"],
cut_tac Ssegment_sub[of "D" "b"],
cut_tac Ssegment_sub[of "D" "a"],
frule SIod_Order[of "D" "Ssegment D a"], assumption,
frule SIod_Order[of "D" "Ssegment D b"], assumption)
apply (rule contrapos_pp, simp+)
apply (frule Torder.less_linear[of "D" "a" "b"], assumption+)
apply simp
apply (erule disjE)
apply (frule Ssegment_unique1[of "D" "a" "b"], assumption+)
apply (simp add:Order.ord_equiv_sym[of "SIod D (Ssegment D a)"
"SIod D (Ssegment D b)"])
apply (simp add:Ssegment_unique1[of "D" "b" "a"])
done
lemma (in Worder) subset_segment:"\<lbrakk>T \<subseteq> carrier D;
\<forall>b\<in>T. \<forall>x. x \<prec> b \<and> x \<in> carrier D \<longrightarrow> x \<in> T;
minimum_elem D (carrier D - T) a\<rbrakk> \<Longrightarrow> T = segment D a"
apply (cut_tac Diff_subset[of "carrier D" T],
frule minimum_elem_mem [of "carrier D - T" a], assumption,
simp, erule conjE)
apply (rule equalityI)
apply (rule subsetI)
apply (frule_tac c = x in subsetD[of T "carrier D"], assumption+)
apply (subst segment_inc[THEN sym], assumption+)
apply (frule_tac x = x in bspec, assumption,
thin_tac "\<forall>b\<in>T. \<forall>x. x \<prec> b \<and> x \<in> carrier D \<longrightarrow> x \<in> T")
apply (rule contrapos_pp, simp+)
apply (frule_tac a = x and b = a in not_less_le, assumption+)
apply (simp add:le_imp_less_or_eq, thin_tac "\<not> x \<prec> a")
apply (erule disjE)
apply (frule_tac a = a in forall_spec) apply (
thin_tac "\<forall>xa. xa \<prec> x \<and> xa \<in> carrier D \<longrightarrow> xa \<in> T")
apply simp apply simp apply simp
apply (rule subsetI)
apply (cut_tac a = a in segment_sub)
apply (frule_tac c = x and A = "segment D a" in subsetD[of _ "carrier D"],
assumption+)
apply (thin_tac "\<forall>b\<in>T. \<forall>x. x \<prec> b \<and> x \<in> carrier D \<longrightarrow> x \<in> T")
apply (rule contrapos_pp, simp+)
apply (simp add:minimum_elem_def)
apply (frule_tac x = x in bspec, simp)
apply (simp add:segment_inc[THEN sym])
apply (simp add:not_le_less[THEN sym])
done
lemma subset_Ssegment:"\<lbrakk>Worder D; T \<subseteq> carrier D;
\<forall>b\<in>T. \<forall>x. x \<prec>\<^bsub>D\<^esub> b \<and> x \<in> carrier D \<longrightarrow> x \<in> T;
minimum_elem D (carrier D - T) a\<rbrakk> \<Longrightarrow> T = Ssegment D a"
apply (cut_tac Diff_subset[of "carrier D" T],
frule Worder.Torder[of D],
frule Worder.Order[of D],
frule Order.minimum_elem_mem [of D "carrier D - T" a], assumption+,
simp, erule conjE)
apply (rule equalityI)
apply (rule subsetI)
apply (frule_tac c = x in subsetD[of T "carrier D"], assumption+)
apply (subst Ssegment_inc[THEN sym], assumption+)
apply (frule_tac x = x in bspec, assumption,
thin_tac "\<forall>b\<in>T. \<forall>x. x \<prec>\<^bsub>D\<^esub> b \<and> x \<in> carrier D \<longrightarrow> x \<in> T")
apply (rule contrapos_pp, simp+)
apply (frule_tac a = x and b = a in Torder.not_less_le, assumption+)
apply (simp add:Order.le_imp_less_or_eq, thin_tac "\<not> x \<prec>\<^bsub>D\<^esub> a")
apply (erule disjE)
apply (frule_tac a = a in forall_spec) apply (
thin_tac "\<forall>xa. xa \<prec>\<^bsub>D\<^esub> x \<and> xa \<in> carrier D \<longrightarrow> xa \<in> T")
apply simp apply simp apply simp
apply (rule subsetI)
apply (cut_tac a = a in Ssegment_sub[of D])
apply (frule_tac c = x and A = "Ssegment D a" in subsetD[of _ "carrier D"],
assumption+)
apply (thin_tac "\<forall>b\<in>T. \<forall>x. x \<prec>\<^bsub>D\<^esub> b \<and> x \<in> carrier D \<longrightarrow> x \<in> T")
apply (rule contrapos_pp, simp+)
apply (simp add:minimum_elem_def)
apply (frule_tac x = x in bspec, simp,
thin_tac "Ball (carrier D - T) ((\<preceq>\<^bsub>D\<^esub>) a)")
apply (simp add:Ssegment_inc[THEN sym])
apply (simp add:Torder.not_le_less[THEN sym])
done
lemma (in Worder) segmentTr:"\<lbrakk>T \<subseteq> carrier D;
\<forall>b \<in> T. (\<forall>x. (x \<prec> b \<and> x \<in> (carrier D) \<longrightarrow> x \<in> T))\<rbrakk> \<Longrightarrow>
(T = carrier D) \<or> (\<exists>a. a \<in> (carrier D) \<and> T = segment D a)"
apply (case_tac "T = carrier D")
apply simp
apply simp
apply (frule not_sym, thin_tac "T \<noteq> carrier D",
frule diff_nonempty[of "carrier D" "T"], assumption)
apply (cut_tac ex_minimum)
apply (frule_tac a = "carrier D - T" in forall_spec, simp)
apply (thin_tac "\<forall>X. X \<subseteq> carrier D \<and> X \<noteq> {} \<longrightarrow> (\<exists>x. minimum_elem D X x)")
apply (erule exE, rename_tac a)
apply (thin_tac "carrier D \<noteq> T", thin_tac "carrier D - T \<noteq> {}")
apply (cut_tac Diff_subset[of "carrier D" "T"])
apply (frule_tac a = a in minimum_elem_mem[of "carrier D - T"],
assumption+,
thin_tac "carrier D - T \<subseteq> carrier D")
apply (simp only:Diff_iff, erule conjE)
apply (frule_tac a = a in subset_segment[of T], assumption+)
apply blast
done
lemma SsegmentTr:"\<lbrakk>Worder D; T \<subseteq> carrier D;
\<forall>b \<in> T. (\<forall>x. (x \<prec>\<^bsub>D\<^esub> b \<and> x \<in> (carrier D) \<longrightarrow> x \<in> T))\<rbrakk> \<Longrightarrow>
(T = carrier D) \<or> (\<exists>a. a \<in> (carrier D) \<and> T = Ssegment D a)"
apply (case_tac "T = carrier D")
apply simp
apply simp
apply (frule not_sym, thin_tac "T \<noteq> carrier D",
frule diff_nonempty[of "carrier D" "T"], assumption)
apply (cut_tac Worder.ex_minimum[of D])
apply (frule_tac a = "carrier D - T" in forall_spec, simp)
apply (thin_tac "\<forall>X. X \<subseteq> carrier D \<and> X \<noteq> {} \<longrightarrow> (\<exists>x. minimum_elem D X x)")
apply (erule exE, rename_tac a)
apply (thin_tac "carrier D \<noteq> T", thin_tac "carrier D - T \<noteq> {}")
apply (cut_tac Diff_subset[of "carrier D" "T"])
apply (frule Worder.Order[of D])
apply (frule_tac a = a in Order.minimum_elem_mem[of D "carrier D - T"],
assumption+,
thin_tac "carrier D - T \<subseteq> carrier D")
apply (simp only:Diff_iff, erule conjE)
apply (subgoal_tac "T = Ssegment D a")
apply blast
apply (rule equalityI)
apply (rule subsetI)
apply (frule_tac c = x in subsetD[of T "carrier D"], assumption+)
apply (subst Ssegment_inc[THEN sym], assumption+)
apply (frule_tac x = x in bspec, assumption,
thin_tac "\<forall>b\<in>T. \<forall>x. x \<prec>\<^bsub>D\<^esub> b \<and> x \<in> carrier D \<longrightarrow> x \<in> T")
apply (rule contrapos_pp, simp+)
apply (frule Worder.Torder[of D],
frule_tac a = x and b = a in Torder.not_less_le[of D], assumption+)
apply (simp add:Order.le_imp_less_or_eq, thin_tac "\<not> x \<prec>\<^bsub>D\<^esub> a")
apply (erule disjE)
apply (frule_tac a = a in forall_spec) apply (
thin_tac "\<forall>xa. xa \<prec>\<^bsub>D\<^esub> x \<and> xa \<in> carrier D \<longrightarrow> xa \<in> T")
apply simp apply simp apply simp
apply (rule subsetI)
apply (frule Worder.Torder[of D],
frule Torder.Order[of D])
apply (cut_tac a = a in Ssegment_sub[of D])
apply (frule_tac c = x and A = "Ssegment D a" in subsetD[of _ "carrier D"],
assumption+)
apply (thin_tac "\<forall>b\<in>T. \<forall>x. x \<prec>\<^bsub>D\<^esub> b \<and> x \<in> carrier D \<longrightarrow> x \<in> T")
apply (rule contrapos_pp, simp+)
apply (simp add:minimum_elem_def)
apply (frule_tac x = x in bspec, simp)
apply (simp add:Ssegment_inc[THEN sym])
apply (simp add:Torder.not_le_less[THEN sym])
apply assumption
done
lemma (in Worder) ord_isom_segment_segment:"\<lbrakk>Worder E;
ord_isom D E f; a \<in> carrier D \<rbrakk> \<Longrightarrow>
ord_isom (Iod D (segment D a)) (Iod E (segment E (f a)))
(\<lambda>x\<in>carrier (Iod D (segment D a)). f x)"
by (frule Worder.Order[of "E"],
rule ord_isom_segment_segment[of "E" "f" "a"], assumption+)
definition
Tw :: "[_ , ('b, 'm1) Order_scheme] \<Rightarrow> 'a \<Rightarrow> 'b" ("(2Tw\<^bsub>_,_\<^esub>)" [60,61]60) where
"Tw\<^bsub>D,T\<^esub> = (\<lambda>a\<in> carrier D. SOME x. x\<in>carrier T \<and>
ord_equiv (Iod D (segment D a)) (Iod T (segment T x)))"
lemma (in Worder) Tw_func:"\<lbrakk>Worder T;
\<forall>a\<in>carrier D. \<exists>b\<in>carrier T. ord_equiv (Iod D (segment D a))
(Iod T (segment T b))\<rbrakk> \<Longrightarrow> Tw\<^bsub>D,T\<^esub> \<in> carrier D \<rightarrow> carrier T"
apply (rule Pi_I)
apply (simp add:Tw_def)
apply (rule someI2_ex) apply blast apply simp
done
lemma (in Worder) Tw_mem:"\<lbrakk>Worder E; x \<in> carrier D;
\<forall>a\<in>carrier D. \<exists>b\<in>carrier E. ord_equiv (Iod D (segment D a))
(Iod E (segment E b))\<rbrakk> \<Longrightarrow> (Tw\<^bsub>D,E\<^esub>) x \<in> carrier E"
by (frule Tw_func[of E], assumption,
simp add:Pi_def)
lemma (in Worder) Tw_equiv:"\<lbrakk>Worder T;
\<forall>a\<in>carrier D. \<exists>b\<in>carrier T. ord_equiv (Iod D (segment D a))
(Iod T (segment T b)); x \<in> carrier D \<rbrakk> \<Longrightarrow>
ord_equiv (Iod D (segment D x)) (Iod T (segment T ((Tw\<^bsub>D,T\<^esub>) x)))"
apply (frule_tac x = x in bspec, assumption+,
thin_tac "\<forall>a\<in>carrier D.
\<exists>b\<in>carrier T. ord_equiv (Iod D (segment D a)) (Iod T (segment T b))")
apply (simp add:Tw_def)
apply (rule someI2_ex)
apply blast apply simp
done
lemma (in Worder) Tw_inj:"\<lbrakk>Worder E;
\<forall>a\<in>carrier D. \<exists>b\<in>carrier E. ord_equiv (Iod D (segment D a))
(Iod E (segment E b))\<rbrakk> \<Longrightarrow> inj_on (Tw\<^bsub>D,E\<^esub>) (carrier D)"
apply (simp add:inj_on_def)
apply (rule ballI)+ apply (rule impI)
apply (frule_tac x = x in Tw_equiv [of "E"], assumption+)
apply simp
apply (frule Tw_func[of "E"], assumption)
apply (frule_tac x = x in funcset_mem[of "Tw D E" "carrier D" "carrier E"],
assumption+,
frule_tac x = y in funcset_mem[of "Tw D E" "carrier D" "carrier E"],
assumption+)
apply (frule Worder.Order[of "E"],
cut_tac a = x in segment_sub,
cut_tac a = y in segment_sub,
cut_tac a = "Tw D E y" in Order.segment_sub[of "E"], assumption)
apply (frule_tac T = "segment D x" in Iod_Order,
frule_tac T = "segment D y" in Iod_Order,
frule_tac T = "segment E (Tw D E y)" in Order.Iod_Order[of "E"],
assumption)
apply (thin_tac "Tw D E x = Tw D E y")
apply (frule_tac x = y in Tw_equiv[of "E"], assumption+)
apply (frule_tac D = "Iod D (segment D y)" and
E = "Iod E (segment E (Tw D E y))" in Order.ord_equiv_sym,
assumption+,
thin_tac "ord_equiv (Iod D (segment D y))
(Iod E (segment E (Tw D E y)))")
apply (frule_tac D = "Iod D (segment D x)" and
E = "Iod E (segment E (Tw D E y))" and
F = "Iod D (segment D y)" in Order.ord_equiv_trans, assumption+)
apply (simp add:segment_unique)
done
lemma (in Worder) Tw_eq_ord_isom:"\<lbrakk>Worder E;
\<forall>a\<in>carrier D. \<exists>b\<in>carrier E.
ord_equiv (Iod D (segment D a)) (Iod E (segment E b)); a \<in> carrier D;
ord_isom (Iod D (segment D a)) (Iod E (segment E (Tw D E a))) f;
x \<in> segment D a \<rbrakk> \<Longrightarrow> f x = Tw D E x"
apply (cut_tac segment_sub[of a])
apply (frule_tac c = x in subsetD[of "segment D a" "carrier D"], assumption+,
frule Tw_equiv[of E x], assumption+)
apply (frule Worder.Torder[of E],
frule Torder.Order[of E])
apply (cut_tac a = x in segment_Worder,
frule_tac D = "Iod D (segment D x)" in Worder.Torder,
frule_tac D = "Iod D (segment D x)" in Worder.Order)
apply (frule_tac T = "segment D a" in Iod_Order)
apply (frule_tac x = a in Tw_mem[of E], assumption+)
apply (frule_tac a = "Tw D E x" in Order.segment_sub[of E])
apply (frule_tac a = "Tw D E a" in Worder.segment_Worder,
frule_tac D = "Iod E (segment E (Tw D E a))" in Worder.Order)
apply (frule_tac f = f and a = x in Order.ord_isom_segment_segment[of
"Iod D (segment D a)" "Iod E (segment E (Tw D E a))"], assumption+)
apply (simp add:Iod_carrier)
apply (frule_tac a = x and b = a in segment_le_mono, assumption+)
apply (frule_tac a1 = x and b1 = a in segment_inc[THEN sym], assumption+)
apply (simp add:oless_def)
apply (frule_tac a1 = x and b1 = a in segment_segment[THEN sym], assumption+)
apply simp
apply (simp add:Iod_sub_sub)
apply (frule_tac f = f and a = x in Order.ord_isom_mem[of
"Iod D (segment D a)" "Iod E (segment E (Tw D E a))"],
simp add:Iod_carrier,
frule Order.segment_sub[of E "Tw D E a"],
simp add:Order.Iod_carrier, simp add:Iod_carrier,
frule Order.segment_sub[of E "Tw D E a"],
simp add:Order.Iod_carrier[of E],
frule_tac c = "f x" in subsetD[of "segment E (Tw D E a)"
"carrier E"], assumption+)
apply (frule_tac a1 = "f x" in Order.segment_inc[THEN sym, of E _
"Tw D E a"], assumption+, simp)
apply (simp add:oless_def, (erule conjE)+)
apply (frule_tac a = "f x" and b = "Tw D E a" in
Torder.segment_le_mono [of E], assumption+, simp)
apply (frule_tac a = "f x" and b = "Tw D E a" in
Order.segment_segment[of E], assumption+)
apply simp
apply (simp add:Order.Iod_sub_sub)
apply (frule_tac D = "Iod D (segment D x)" in Torder.Order)
apply (frule_tac D = "Iod D (segment D x)" and E = "Iod E (segment E (f x))"
and F = "Iod E (segment E (Tw D E x))" in Order.ord_equiv_box)
apply (frule_tac a = "f x" in Order.segment_sub[of E])
apply (frule_tac T = "segment E (f x)" in Order.Iod_Order[of E], assumption+)
apply (frule_tac a = "f x" in Order.segment_sub[of E])
apply (frule Tw_mem[of E x], assumption+)
apply (frule Order.segment_sub[of E "Tw D E x"])
apply (rule Order.Iod_Order[of E], assumption+)
apply (simp add:ord_equiv_def, blast)
apply assumption
apply (frule_tac a = "f x" and b = "Tw D E x" in
Worder.segment_unique[of E], assumption+)
apply (frule_tac x = x in Tw_mem[of E], assumption+)
done
lemma (in Worder) Tw_ord_injTr:"\<lbrakk>Worder E;
\<forall>a\<in>carrier D. \<exists>b\<in>carrier E.
ord_equiv (Iod D (segment D a)) (Iod E (segment E b));
a \<in> carrier D; b \<in> carrier D; a \<prec> b\<rbrakk> \<Longrightarrow>
Tw D E a \<prec>\<^bsub>E\<^esub> (Tw D E b)"
apply (frule_tac x = b in Tw_equiv [of "E"], assumption+)
apply (simp add:segment_inc)
apply (simp add:ord_equiv_def, erule exE, fold ord_equiv_def)
apply (frule_tac f = f in Tw_eq_ord_isom[of E b _ a], assumption+)
apply (cut_tac segment_sub[of b])
apply (frule Iod_Order[of "segment D b"])
apply (frule Worder.Order[of E],
frule Tw_mem[of E b], assumption+,
frule Order.segment_sub[of E "Tw D E b"],
frule Order.Iod_Order[of E "segment E (Tw D E b)"], assumption)
apply (frule_tac f = f and a = a in Order.ord_isom_mem[of
"Iod D (segment D b)" "Iod E (segment E (Tw D E b))"], assumption+)
apply (simp add:Iod_carrier)
apply (simp add:Order.Iod_carrier)
apply (subst Order.segment_inc[of E], assumption+)
apply (simp add:Tw_mem)+
done
lemma (in Worder) Tw_ord_inj:"\<lbrakk>Worder E;
\<forall>a\<in>carrier D. \<exists>b\<in>carrier E. ord_equiv (Iod D (segment D a))
(Iod E (segment E b))\<rbrakk> \<Longrightarrow> ord_inj D E (Tw D E)"
apply (simp add:ord_inj_def)
apply (rule conjI)
apply (simp add:Tw_def extensional_def)
apply (simp add:Tw_func)
apply (rule conjI)
apply (simp add:Tw_inj)
apply (rule ballI)+
apply (rule iffI)
apply (simp add:Tw_ord_injTr)
apply (rule contrapos_pp, simp+)
apply (simp add:not_less_le)
apply (simp add:le_imp_less_or_eq)
apply (erule disjE)
apply (frule_tac a = b and b = a in Tw_ord_injTr[of "E"], assumption+)
apply (frule Tw_func [of "E"], assumption+)
apply (frule_tac x = a in funcset_mem[of "Tw D E" "carrier D" "carrier E"],
assumption+,
frule_tac x = b in funcset_mem[of "Tw D E" "carrier D" "carrier E"],
assumption+)
apply (frule Worder.Torder[of "E"],
frule_tac a1 = "Tw D E b" and b1 = "Tw D E a" in
Torder.not_le_less[THEN sym, of "E"], assumption+, simp)
apply (frule Worder.Order[of "E"],
frule_tac a = "Tw D E b" and b = "Tw D E a" in
Order.less_imp_le[of "E"], assumption+, simp)
apply (simp add:oless_def)
done
lemma (in Worder) ord_isom_restricted_by_Tw:"\<lbrakk>Worder E;
\<forall>a\<in>carrier D. \<exists>b\<in>carrier E.
ord_equiv (Iod D (segment D a)) (Iod E (segment E b));
D1 \<subseteq> carrier D\<rbrakk> \<Longrightarrow>
ord_isom (Iod D D1) (Iod E ((Tw D E) ` D1))
(restrict (Tw D E) D1)"
apply (frule Tw_ord_inj [of "E"], assumption+)
apply (frule Worder.Order[of E])
apply (rule ord_inj_restrict_isom [of E "Tw D E" "D1"], assumption+)
done
lemma (in Worder) Tw_segment_segment:"\<lbrakk>Worder E;
\<forall>a\<in>carrier D.\<exists>b\<in>carrier E.
ord_equiv (Iod D (segment D a)) (Iod E (segment E b)); a \<in> carrier D\<rbrakk>
\<Longrightarrow> Tw D E ` (segment D a) = segment E (Tw D E a)"
apply (rule equalityI)
apply (rule subsetI)
apply (simp add:image_def, erule bexE)
apply (frule Tw_equiv[of "E" "a"], assumption+)
apply (simp add:ord_equiv_def, erule exE, fold ord_equiv_def)
apply (frule_tac x = xa in Tw_eq_ord_isom[of E a], assumption+)
apply (rotate_tac -1, frule sym, thin_tac "f xa = Tw D E xa", simp)
apply (cut_tac segment_sub[of a],
frule Iod_Order[of "segment D a"])
apply (frule Worder.Order[of E],
frule_tac a = "Tw D E a" in Order.segment_sub[of E],
frule Tw_mem[of E a], assumption+,
frule Order.segment_sub[of E "Tw D E a"])
apply (frule_tac T = "segment E (Tw D E a)" in Order.Iod_Order[of E],
assumption+)
apply (frule_tac a = xa and f = f and D = "Iod D (segment D a)" and
E = "Iod E (segment E (Tw D E a))" in Order.ord_isom_mem,
assumption+)
apply (simp add:Iod_carrier)
apply (simp add:Order.Iod_carrier)
apply (rule subsetI)
apply (simp add:image_def)
apply (frule Tw_equiv[of "E" "a"], assumption+)
apply (simp add:ord_equiv_def, erule exE, fold ord_equiv_def)
apply (cut_tac segment_sub[of a],
frule Iod_Order[of "segment D a"])
apply (frule Worder.Order[of E],
frule_tac a = "Tw D E a" in Order.segment_sub[of E],
frule Tw_mem[of E a], assumption+,
frule Order.segment_sub[of E "Tw D E a"])
apply (frule_tac T = "segment E (Tw D E a)" in Order.Iod_Order[of E],
assumption+)
apply (frule Iod_Order[of "segment D a"])
apply (frule_tac b = x in Order.ord_isom_surj [of "Iod D (segment D a)"
"Iod E (segment E (Tw D E a))"], assumption+)
apply (simp add:Order.Iod_carrier)
apply (erule bexE, simp add:Iod_carrier)
apply (frule_tac f = f and x = aa in Tw_eq_ord_isom[of E a], assumption+)
apply (simp, blast)
done
lemma (in Worder) ord_isom_Tw_segment:"\<lbrakk>Worder E;
\<forall>a\<in>carrier D. \<exists>b\<in>carrier E.
ord_equiv (Iod D (segment D a)) (Iod E (segment E b)); a\<in>carrier D\<rbrakk> \<Longrightarrow>
ord_isom (Iod D (segment D a)) (Iod E (segment E (Tw D E a)))
(restrict (Tw D E) (segment D a))"
apply (cut_tac segment_sub[of "a"],
frule ord_isom_restricted_by_Tw[of "E" "segment D a"], assumption+,
simp add:Tw_segment_segment[of "E" "a"])
done
lemma (in Worder) well_ord_compare1:"\<lbrakk>Worder E;
\<forall>a\<in>carrier D. \<exists>b\<in>carrier E.
ord_equiv (Iod D (segment D a)) (Iod E (segment E b))\<rbrakk> \<Longrightarrow>
(ord_equiv D E) \<or> (\<exists>c\<in>carrier E. ord_equiv D (Iod E (segment E c)))"
apply (frule Tw_ord_inj [of "E"], assumption+)
apply (frule Tw_func[of "E"], assumption+)
apply (frule ord_isom_restricted_by_Tw [of "E" "carrier D"], assumption+,
simp)
apply (simp add:Iod_self[THEN sym])
apply (frule image_sub0[of "Tw D E" "carrier D" "carrier E"],
frule Worder.segmentTr [of "E" "(Tw D E) ` (carrier D)"],
assumption)
apply (rule ballI, rule allI, rule impI, erule conjE)
apply (thin_tac "ord_isom D (Iod E (Tw D E ` carrier D))
(restrict (Tw D E) (carrier D))")
apply (thin_tac "Tw D E ` carrier D \<subseteq> carrier E",
simp add:image_def, erule bexE)
apply (frule_tac a = xa in ord_isom_Tw_segment[of "E"], assumption+)
apply (rename_tac b x c)
apply (frule_tac x = c in funcset_mem[of "Tw D E" "carrier D" "carrier E"],
assumption, simp, thin_tac "b = Tw D E c")
apply (frule Worder.Order[of "E"],
frule_tac a = "Tw D E c" in Order.segment_sub[of "E"],
cut_tac a = c in segment_Worder,
cut_tac a = "Tw D E c" in Worder.segment_Worder[of "E"],
assumption,
frule_tac D = "Iod D (segment D c)" in Worder.Order,
frule_tac D = "Iod E (segment E (Tw D E c))" in Worder.Order)
apply (frule_tac D = "Iod D (segment D c)" and
E = "Iod E (segment E (Tw D E c))" and
f = "restrict (Tw D E) (segment D c)" and b = x in
Order.ord_isom_surj, assumption+)
apply (simp add:Order.Iod_carrier[of "E"])
apply (frule_tac a = x and b = "Tw D E c" in Order.segment_inc[of "E"],
assumption+, simp)
apply (insert Order,
cut_tac a = c in segment_sub,
simp add:Iod_carrier, erule bexE, blast)
apply (erule disjE)
apply simp
apply (frule Worder.Order[of "E"],
simp add:Order.Iod_self[THEN sym, of "E"],
simp add:ord_equiv)
apply (erule exE, erule conjE, simp,
frule Worder.Order[of "E"],
frule_tac a = a in Order.segment_sub[of "E"],
cut_tac a = a in Worder.segment_Worder[of "E"],
assumption,
frule_tac D = "Iod E (segment E a)" in Worder.Order,
frule_tac E = "Iod E (segment E a)" in ord_equiv, simp, blast)
done
lemma bex_nonempty_set:"\<exists>x \<in> A. P x \<Longrightarrow> {x. x \<in> A \<and> P x } \<noteq> {}"
by blast
lemma nonempty_set_sub:"{x. x \<in> A \<and> P x } \<noteq> {} \<Longrightarrow>
{x. x \<in> A \<and> P x} \<subseteq> A"
by (rule subsetI, simp)
lemma (in Torder) less_minimum:"\<lbrakk>minimum_elem D {x. x \<in> carrier D \<and> P x} d\<rbrakk>
\<Longrightarrow> \<forall>a. (((a \<prec> d) \<and> a \<in> carrier D) \<longrightarrow> \<not> (P a))"
apply (rule allI, rule impI, erule conjE)
apply (rule contrapos_pp, simp+)
apply (simp add:minimum_elem_def, (erule conjE)+)
apply (frule_tac a = a in forall_spec, simp,
thin_tac "\<forall>x. x \<in> carrier D \<and> P x \<longrightarrow> d \<preceq> x")
apply (simp add:not_le_less[THEN sym, of "d"])
done
lemma (in Torder) segment_minimum_empty:"\<lbrakk>X \<subseteq> carrier D; d \<in> X\<rbrakk> \<Longrightarrow>
(minimum_elem D X d) = (segment (Iod D X) d = {})"
apply (rule iffI)
apply (rule contrapos_pp, simp+)
apply (frule nonempty_ex[of "segment (Iod D X) d"], erule exE,
thin_tac "segment (Iod D X) d \<noteq> {}",
frule minimum_elem_mem[of "X" "d"], assumption+,
frule_tac c = d in subsetD[of "X" "carrier D"], assumption+)
apply (simp add:segment_def,
simp add:Iod_carrier, erule conjE,
simp add:Iod_less[of "X"])
apply (simp add:minimum_elem_def,
frule_tac x = x in bspec, assumption,
frule_tac c = x in subsetD[of "X" "carrier D"], assumption+,
frule_tac a1 = x and b1 = d in not_less_le[THEN sym], assumption+)
apply simp
apply (rule contrapos_pp, simp+)
apply (simp add:minimum_elem_def)
apply (erule bexE)
apply (frule_tac c = d in subsetD[of "X" "carrier D"], assumption+,
frule_tac c = x in subsetD[of "X" "carrier D"], assumption+,
simp add:not_le_less)
apply (simp add:segment_def Iod_carrier,
simp add:Iod_less[THEN sym, of "X"])
done
end
diff --git a/thys/Group-Ring-Module/Algebra5.thy b/thys/Group-Ring-Module/Algebra5.thy
--- a/thys/Group-Ring-Module/Algebra5.thy
+++ b/thys/Group-Ring-Module/Algebra5.thy
@@ -1,5060 +1,5061 @@
(** Algebra5
author Hidetsune Kobayashi
Group You Santo
Department of Mathematics
Nihon University
h_coba@math.cst.nihon-u.ac.jp
May 3, 2004.
April 6, 2007 (revised).
chapter 4. Ring theory (continued)
section 6. operation of ideals
section 7. direct product1, general case
section 8. Chinese remainder theorem
section 9. addition of finite elements of a ring and ideal_multiplication
section 10. extension and contraction
section 11. complete system of representatives
section 12. Polynomial ring
section 13. addition and multiplication of polyn_exprs
section 14. the degree of a polynomial
**)
theory Algebra5 imports Algebra4 begin
section "Operation of ideals"
lemma (in Ring) ideal_sumTr1:"\<lbrakk>ideal R A; ideal R B\<rbrakk> \<Longrightarrow>
A \<minusplus> B = \<Inter> {J. ideal R J \<and> (A \<union> B) \<subseteq> J}"
apply (frule sum_ideals[of "A" "B"], assumption,
frule sum_ideals_la1[of "A" "B"], assumption,
frule sum_ideals_la2[of "A" "B"], assumption)
apply (rule equalityI)
(* A \<minusplus>\<^sub>R B \<subseteq> \<Inter>{J. ideal R J \<and> A \<union> B \<subseteq> J} *)
apply (rule_tac A = "{J. ideal R J \<and> (A \<union> B) \<subseteq> J}" and C = "A \<minusplus> B" in
Inter_greatest)
apply (simp, (erule conjE)+)
apply (rule_tac A = A and B = B and I = X in sum_ideals_cont,
simp add:ideal_subset1, simp add:ideal_subset1, assumption+)
apply (rule_tac B = "A \<minusplus> B" and A = "{J. ideal R J \<and> (A \<union> B) \<subseteq> J}" in
Inter_lower)
apply simp
done
lemma (in Ring) sum_ideals_commute:"\<lbrakk>ideal R A; ideal R B\<rbrakk> \<Longrightarrow>
A \<minusplus> B = B \<minusplus> A"
apply (frule ideal_sumTr1 [of "B"], assumption+,
frule ideal_sumTr1 [of "A" "B"], assumption+)
apply (simp add:Un_commute[of "B" "A"])
done
lemma (in Ring) ideal_prod_mono1:"\<lbrakk>ideal R A; ideal R B; ideal R C;
A \<subseteq> B \<rbrakk> \<Longrightarrow> A \<diamondsuit>\<^sub>r C \<subseteq> B \<diamondsuit>\<^sub>r C"
apply (frule ideal_prod_ideal[of "B" "C"], assumption+)
apply (rule ideal_prod_subTr[of "A" "C" "B \<diamondsuit>\<^sub>r C"], assumption+)
apply (rule ballI)+
apply (frule_tac c = i in subsetD[of "A" "B"], assumption+)
apply (rule_tac i = i and j = j in prod_mem_prod_ideals[of "B" "C"],
assumption+)
done
lemma (in Ring) ideal_prod_mono2:"\<lbrakk>ideal R A; ideal R B; ideal R C;
A \<subseteq> B \<rbrakk> \<Longrightarrow> C \<diamondsuit>\<^sub>r A \<subseteq> C \<diamondsuit>\<^sub>r B"
apply (frule ideal_prod_mono1[of "A" "B" "C"], assumption+)
apply (simp add:ideal_prod_commute)
done
lemma (in Ring) cont_ideal_prod:"\<lbrakk>ideal R A; ideal R B; ideal R C;
A \<subseteq> C; B \<subseteq> C \<rbrakk> \<Longrightarrow> A \<diamondsuit>\<^sub>r B \<subseteq> C"
apply (simp add:ideal_prod_def)
apply (rule subsetI, simp)
apply (frule ideal_prod_ideal[of "A" "B"], assumption,
frule_tac a = "A \<diamondsuit>\<^sub>r B" in forall_spec,
thin_tac "\<forall>xa. ideal R xa \<and> {x. \<exists>i\<in>A. \<exists>j\<in>B. x = i \<cdot>\<^sub>r j} \<subseteq> xa \<longrightarrow> x \<in> xa",
simp)
apply (rule subsetI, simp, (erule bexE)+, simp add:prod_mem_prod_ideals)
apply (frule ideal_prod_la1[of "A" "B"], assumption,
frule_tac c = x in subsetD[of "A \<diamondsuit>\<^sub>r B" "A"], assumption+,
simp add:subsetD)
done
lemma (in Ring) ideal_distrib:"\<lbrakk>ideal R A; ideal R B; ideal R C\<rbrakk> \<Longrightarrow>
A \<diamondsuit>\<^sub>r (B \<minusplus> C) = A \<diamondsuit>\<^sub>r B \<minusplus> A \<diamondsuit>\<^sub>r C"
apply (frule sum_ideals[of "B" "C"], assumption,
frule ideal_prod_ideal[of "A" "B \<minusplus> C"], assumption+,
frule ideal_prod_ideal[of "A" "B"], assumption+,
frule ideal_prod_ideal[of "A" "C"], assumption+,
frule sum_ideals[of "A \<diamondsuit>\<^sub>r B" "A \<diamondsuit>\<^sub>r C"], assumption)
apply (rule equalityI)
apply (rule ideal_prod_subTr[of "A" "B \<minusplus> C" "A \<diamondsuit>\<^sub>r B \<minusplus> A \<diamondsuit>\<^sub>r C"], assumption+)
apply (rule ballI)+
apply (frule_tac x = j in ideals_set_sum[of B C], assumption+,
(erule bexE)+, simp) apply (
thin_tac "j = h \<plusminus> k",
frule_tac h = i in ideal_subset[of "A"], assumption+,
frule_tac h = h in ideal_subset[of "B"], assumption+,
frule_tac h = k in ideal_subset[of "C"], assumption+)
apply (simp add:ring_distrib1)
apply (frule_tac i = i and j = h in prod_mem_prod_ideals[of "A" "B"],
assumption+,
frule_tac i = i and j = k in prod_mem_prod_ideals[of "A" "C"],
assumption+)
apply (frule sum_ideals_la1[of "A \<diamondsuit>\<^sub>r B" "A \<diamondsuit>\<^sub>r C"], assumption+,
frule sum_ideals_la2[of "A \<diamondsuit>\<^sub>r B" "A \<diamondsuit>\<^sub>r C"], assumption+)
apply (frule_tac c = "i \<cdot>\<^sub>r h" in subsetD[of "A \<diamondsuit>\<^sub>r B" "A \<diamondsuit>\<^sub>r B \<minusplus> A \<diamondsuit>\<^sub>r C"],
assumption+,
frule_tac c = "i \<cdot>\<^sub>r k" in subsetD[of "A \<diamondsuit>\<^sub>r C" "A \<diamondsuit>\<^sub>r B \<minusplus> A \<diamondsuit>\<^sub>r C"],
assumption+)
apply (simp add:ideal_pOp_closed)
apply (rule sum_ideals_cont[of "A \<diamondsuit>\<^sub>r (B \<minusplus> C)" "A \<diamondsuit>\<^sub>r B" "A \<diamondsuit>\<^sub>r C" ],
assumption+)
apply (rule ideal_prod_subTr[of "A" "B" "A \<diamondsuit>\<^sub>r (B \<minusplus> C)"], assumption+)
apply (rule ballI)+
apply (frule sum_ideals_la1[of "B" "C"], assumption+,
frule_tac c = j in subsetD[of "B" "B \<minusplus> C"], assumption+,
rule_tac i = i and j = j in prod_mem_prod_ideals[of "A" "B \<minusplus> C"],
assumption+)
apply (rule ideal_prod_subTr[of "A" "C" "A \<diamondsuit>\<^sub>r (B \<minusplus> C)"], assumption+)
apply (rule ballI)+
apply (frule sum_ideals_la2[of "B" "C"], assumption+,
frule_tac c = j in subsetD[of "C" "B \<minusplus> C"], assumption+,
rule_tac i = i and j = j in prod_mem_prod_ideals[of "A" "B \<minusplus> C"],
assumption+)
done
definition
coprime_ideals::"[_, 'a set, 'a set] \<Rightarrow> bool" where
"coprime_ideals R A B \<longleftrightarrow> A \<minusplus>\<^bsub>R\<^esub> B = carrier R"
lemma (in Ring) coprimeTr:"\<lbrakk>ideal R A; ideal R B\<rbrakk> \<Longrightarrow>
coprime_ideals R A B = (\<exists>a \<in> A. \<exists>b \<in> B. a \<plusminus> b = 1\<^sub>r)"
apply (rule iffI)
apply (simp add:coprime_ideals_def)
apply (cut_tac ring_one, frule sym, thin_tac "A \<minusplus> B = carrier R", simp,
thin_tac "carrier R = A \<minusplus> B", frule ideals_set_sum[of A B],
assumption+, (erule bexE)+, frule sym, blast)
apply (erule bexE)+
apply (frule ideal_subset1[of A], frule ideal_subset1[of B])
apply (frule_tac a = a and b = b in set_sum_mem[of _ A _ B], assumption+)
apply (simp add:coprime_ideals_def)
apply (frule sum_ideals[of "A" "B"], assumption+,
frule ideal_inc_one[of "A \<minusplus> B"], assumption)
apply (simp add:coprime_ideals_def)
done
lemma (in Ring) coprime_int_prod:"\<lbrakk>ideal R A; ideal R B; coprime_ideals R A B\<rbrakk>
\<Longrightarrow> A \<inter> B = A \<diamondsuit>\<^sub>r B"
apply (frule ideal_prod_la1[of "A" "B"], assumption+,
frule ideal_prod_la2[of "A" "B"], assumption+)
apply (rule equalityI)
defer
(** A \<diamondsuit>\<^bsub>R\<^esub> B \<subseteq> A \<inter> B **)
apply simp
apply (simp add:coprime_ideals_def)
apply (frule int_ideal[of "A" "B"], assumption)
apply (frule idealprod_whole_r[of "A \<inter> B"])
apply (frule sym, thin_tac "A \<minusplus> B = carrier R", simp)
apply (simp add:ideal_distrib)
apply (simp add:ideal_prod_commute[of "A \<inter> B" "A"])
apply (cut_tac Int_lower1[of "A" "B"], cut_tac Int_lower2[of "A" "B"])
apply (frule ideal_prod_mono2[of "A \<inter> B" "B" "A"], assumption+,
frule ideal_prod_mono1[of "A \<inter> B" "A" "B"], assumption+)
apply (frule ideal_prod_ideal[of "A \<inter> B" "B"], assumption+,
frule ideal_prod_ideal[of "A" "A \<inter> B"], assumption+,
frule ideal_subset1[of "A \<diamondsuit>\<^sub>r (A \<inter> B)"],
frule ideal_subset1[of "(A \<inter> B) \<diamondsuit>\<^sub>r B"])
apply (frule ideal_prod_ideal[of A B], assumption+,
frule sum_ideals_cont[of "A \<diamondsuit>\<^sub>r B" "A \<diamondsuit>\<^sub>r (A \<inter> B)" "(A \<inter> B) \<diamondsuit>\<^sub>r B"],
assumption+)
apply simp
done
lemma (in Ring) coprime_elems:"\<lbrakk>ideal R A; ideal R B; coprime_ideals R A B\<rbrakk> \<Longrightarrow>
\<exists>a\<in>A. \<exists>b\<in>B. a \<plusminus> b = 1\<^sub>r"
by (simp add:coprimeTr)
lemma (in Ring) coprime_elemsTr:"\<lbrakk>ideal R A; ideal R B; a\<in>A; b\<in>B; a \<plusminus> b = 1\<^sub>r\<rbrakk>
\<Longrightarrow> pj R A b = 1\<^sub>r\<^bsub>(qring R A)\<^esub> \<and> pj R B a = 1\<^sub>r\<^bsub>(qring R B)\<^esub>"
apply (frule pj_Hom [OF Ring, of "A"],
frule pj_Hom [OF Ring, of "B"])
apply (frule qring_ring[of "A"], frule qring_ring[of "B"])
apply (cut_tac ring_is_ag,
frule Ring.ring_is_ag[of "R /\<^sub>r A"],
frule Ring.ring_is_ag[of "R /\<^sub>r B"])
apply (frule_tac a = a and b = b in aHom_add[of "R" "R /\<^sub>r A" "pj R A"],
assumption+,
simp add:rHom_def, simp add:ideal_subset,
simp add:ideal_subset, simp)
apply (frule ideal_subset[of "A" "a"], assumption,
frule ideal_subset[of "B" "b"], assumption+)
apply (simp only:pj_zero[OF Ring, THEN sym, of "A" "a"],
frule rHom_mem[of "pj R A" "R" "qring R A" "b"], assumption+,
simp add:aGroup.l_zero[of "R /\<^sub>r A"])
apply (simp add:rHom_one[OF Ring, of "qring R A"])
apply (frule_tac a = a and b = b in aHom_add[of "R" "R /\<^sub>r B" "pj R B"],
assumption+,
simp add:rHom_def, simp add:ideal_subset,
simp add:ideal_subset, simp)
apply (simp only:pj_zero[OF Ring, THEN sym, of "B" "b"],
frule rHom_mem[of "pj R B" "R" "qring R B" "a"], assumption+,
simp add:aGroup.ag_r_zero[of "R /\<^sub>r B"])
apply (simp add:rHom_one[OF Ring, of "qring R B"])
done
lemma (in Ring) partition_of_unity:"\<lbrakk>ideal R A; a \<in> A; b \<in> carrier R;
a \<plusminus> b = 1\<^sub>r; u \<in> carrier R; v \<in> carrier R\<rbrakk> \<Longrightarrow>
pj R A (a \<cdot>\<^sub>r v \<plusminus> b \<cdot>\<^sub>r u ) = pj R A u"
apply (frule pj_Hom [OF Ring, of "A"],
frule ideal_subset [of "A" "a"], assumption+,
frule ring_tOp_closed[of "a" "v"], assumption+,
frule ring_tOp_closed[of "b" "u"], assumption+,
frule qring_ring[of "A"])
apply (simp add:ringhom1[OF Ring, of "qring R A" "a \<cdot>\<^sub>r v" "b \<cdot>\<^sub>r u" "pj R A"])
apply (frule ideal_ring_multiple1[of "A" "a" "v"], assumption+,
simp add:pj_zero[OF Ring, THEN sym, of "A" "a \<cdot>\<^sub>r v"],
frule rHom_mem[of "pj R A" "R" "qring R A" "b \<cdot>\<^sub>r u"], assumption+,
simp add:Ring.l_zero, simp add:rHom_tOp[OF Ring])
apply (frule ringhom1[OF Ring, of "qring R A" "a" "b" "pj R A"], assumption+,
simp,
simp add:pj_zero[OF Ring, THEN sym, of "A" "a"],
frule rHom_mem[of "pj R A" "R" "qring R A" "b"], assumption+,
simp add:Ring.l_zero)
apply (simp add:rHom_one[OF Ring, of "qring R A" "pj R A"],
rotate_tac -2, frule sym, thin_tac "1\<^sub>r\<^bsub>R /\<^sub>r A\<^esub> = pj R A b",
simp,
rule Ring.ring_l_one[of "qring R A" "pj R A u"], assumption)
apply (simp add:rHom_mem)
done
lemma (in Ring) coprimes_commute:"\<lbrakk>ideal R A; ideal R B; coprime_ideals R A B \<rbrakk>
\<Longrightarrow> coprime_ideals R B A"
apply (simp add:coprime_ideals_def)
apply (simp add:sum_ideals_commute)
done
lemma (in Ring) coprime_surjTr:"\<lbrakk>ideal R A; ideal R B; coprime_ideals R A B;
X \<in> carrier (qring R A); Y \<in> carrier (qring R B) \<rbrakk> \<Longrightarrow>
\<exists>r\<in>carrier R. pj R A r = X \<and> pj R B r = Y"
apply (frule qring_ring [of "A"],
frule qring_ring [of "B"],
frule coprime_elems [of "A" "B"], assumption+,
frule pj_Hom [OF Ring, of "A"],
frule pj_Hom [OF Ring, of "B"])
apply (erule bexE)+
apply (simp add:qring_carrier[of "A"],
simp add:qring_carrier[of "B"], (erule bexE)+,
rename_tac a b u v)
apply (rotate_tac -1, frule sym, thin_tac "v \<uplus>\<^bsub>R\<^esub> B = Y",
rotate_tac -3, frule sym, thin_tac "u \<uplus>\<^bsub>R\<^esub> A = X", simp)
apply (frule_tac h = a in ideal_subset[of "A"], assumption+,
frule_tac h = b in ideal_subset[of "B"], assumption+,
frule_tac x = a and y = v in ring_tOp_closed, assumption+,
frule_tac x = b and y = u in ring_tOp_closed, assumption+,
cut_tac ring_is_ag,
frule_tac x = "a \<cdot>\<^sub>r v" and y = "b \<cdot>\<^sub>r u" in aGroup.ag_pOp_closed[of "R"],
assumption+)
apply (frule_tac a = a and b = b and u = u and v = v in
partition_of_unity[of "A"], assumption+,
simp add:pj_mem[OF Ring, of "A"])
apply (frule_tac a = b and b = a and u = v and v = u in
partition_of_unity[of "B"], assumption+,
subst aGroup.ag_pOp_commute[of "R"], assumption+,
simp add:pj_mem[OF Ring, of "B"])
apply (frule_tac x = "b \<cdot>\<^sub>r u" and y = "a \<cdot>\<^sub>r v" in
aGroup.ag_pOp_commute[of "R"], assumption+, simp)
apply (simp add:pj_mem[OF Ring], blast)
done
lemma (in Ring) coprime_n_idealsTr0:"\<lbrakk>ideal R A; ideal R B; ideal R C;
coprime_ideals R A C; coprime_ideals R B C \<rbrakk> \<Longrightarrow>
coprime_ideals R (A \<diamondsuit>\<^sub>r B) C"
apply (simp add:coprimeTr[of "A" "C"],
simp add:coprimeTr[of "B" "C"], (erule bexE)+)
apply (cut_tac ring_is_ag,
frule_tac h = a in ideal_subset[of "A"], assumption+,
frule_tac h = aa in ideal_subset[of "B"], assumption+,
frule_tac h = b in ideal_subset[of "C"], assumption+,
frule_tac h = ba in ideal_subset[of "C"], assumption+,
frule_tac x = a and y = b in aGroup.ag_pOp_closed, assumption+)
apply (frule_tac x = "a \<plusminus> b" and y = aa and z = ba in ring_distrib1,
assumption+) apply (
rotate_tac 6, frule sym, thin_tac "a \<plusminus> b = 1\<^sub>r",
frule sym, thin_tac "aa \<plusminus> ba = 1\<^sub>r")
apply (simp only:ring_distrib2)
apply (rotate_tac -1, frule sym, thin_tac "1\<^sub>r = a \<plusminus> b", simp,
thin_tac "aa \<plusminus> ba = 1\<^sub>r")
apply (simp add:ring_r_one,
thin_tac "a \<cdot>\<^sub>r aa \<plusminus> b \<cdot>\<^sub>r aa \<plusminus> (a \<cdot>\<^sub>r ba \<plusminus> b \<cdot>\<^sub>r ba) \<in> carrier R",
thin_tac "a \<plusminus> b = a \<cdot>\<^sub>r aa \<plusminus> b \<cdot>\<^sub>r aa \<plusminus> (a \<cdot>\<^sub>r ba \<plusminus> b \<cdot>\<^sub>r ba)")
apply (frule_tac x = a and y = aa in ring_tOp_closed, assumption+,
frule_tac x = b and y = aa in ring_tOp_closed, assumption+,
frule_tac x = a and y = ba in ring_tOp_closed, assumption+,
frule_tac x = b and y = ba in ring_tOp_closed, assumption+,
frule_tac x = "a \<cdot>\<^sub>r ba" and y = "b \<cdot>\<^sub>r ba" in aGroup.ag_pOp_closed,
assumption+)
apply (simp add:aGroup.ag_pOp_assoc,
frule sym, thin_tac "1\<^sub>r = a \<cdot>\<^sub>r aa \<plusminus> (b \<cdot>\<^sub>r aa \<plusminus> (a \<cdot>\<^sub>r ba \<plusminus> b \<cdot>\<^sub>r ba))")
apply (frule_tac i = a and j = aa in prod_mem_prod_ideals[of "A" "B"],
assumption+)
apply (frule_tac x = ba and r = a in ideal_ring_multiple[of "C"],
assumption+,
frule_tac x = ba and r = b in ideal_ring_multiple[of "C"],
assumption+,
frule_tac x = b and r = aa in ideal_ring_multiple1[of "C"],
assumption+)
apply (frule_tac I = C and x = "a \<cdot>\<^sub>r ba" and y = "b \<cdot>\<^sub>r ba" in
ideal_pOp_closed, assumption+,
frule_tac I = C and x = "b \<cdot>\<^sub>r aa" and y = "a \<cdot>\<^sub>r ba \<plusminus> b \<cdot>\<^sub>r ba" in
ideal_pOp_closed, assumption+)
apply (frule ideal_prod_ideal[of "A" "B"], assumption)
apply (subst coprimeTr, assumption+, blast)
done
lemma (in Ring) coprime_n_idealsTr1:"ideal R C \<Longrightarrow>
(\<forall>k \<le> n. ideal R (J k)) \<and> (\<forall>i \<le> n. coprime_ideals R (J i) C) \<longrightarrow>
coprime_ideals R (i\<Pi>\<^bsub>R,n\<^esub> J) C"
apply (induct_tac n)
apply (rule impI)
apply (erule conjE)
apply simp
apply (rule impI)
apply (erule conjE)
apply (cut_tac n = "Suc n" in n_in_Nsetn)
apply (cut_tac n = n in Nset_Suc) apply simp
apply (cut_tac n = n and J = J in n_prod_ideal,
rule allI, simp)
apply (rule_tac A = "i\<Pi>\<^bsub>R,n\<^esub> J" and B = "J (Suc n)" in
coprime_n_idealsTr0[of _ _ "C"], simp+)
done
lemma (in Ring) coprime_n_idealsTr2:"\<lbrakk>ideal R C; (\<forall>k \<le> n. ideal R (J k));
(\<forall>i \<le> n. coprime_ideals R (J i) C) \<rbrakk> \<Longrightarrow>
coprime_ideals R (i\<Pi>\<^bsub>R,n\<^esub> J) C"
by (simp add:coprime_n_idealsTr1)
lemma (in Ring) coprime_n_idealsTr3:"(\<forall>k \<le> (Suc n). ideal R (J k)) \<and>
(\<forall>i \<le> (Suc n). \<forall>j \<le> (Suc n). i \<noteq> j \<longrightarrow>
coprime_ideals R (J i) (J j)) \<longrightarrow> coprime_ideals R (i\<Pi>\<^bsub>R,n\<^esub> J) (J (Suc n))"
apply (rule impI, erule conjE)
apply (case_tac "n = 0", simp)
apply (simp add:Nset_1)
apply (cut_tac nat_eq_le[of "Suc n" "Suc n"],
frule_tac a = "Suc n" in forall_spec, assumption)
apply (rotate_tac 1, frule_tac a = "Suc n" in forall_spec, assumption,
thin_tac "\<forall>j\<le>Suc n. Suc n \<noteq> j \<longrightarrow> coprime_ideals R (J (Suc n)) (J j)")
apply (rule_tac C = "J (Suc n)" and n = n and J = J in coprime_n_idealsTr2,
assumption, rule allI, simp)
apply (rule allI, rule impI)
apply (frule_tac a = i in forall_spec, simp,
thin_tac "\<forall>i\<le>Suc n. \<forall>j\<le>Suc n. i \<noteq> j \<longrightarrow> coprime_ideals R (J i) (J j)",
rotate_tac -1,
frule_tac a = "Suc n" in forall_spec, assumption+)
apply simp+
done
lemma (in Ring) coprime_n_idealsTr4:"\<lbrakk>(\<forall>k \<le> (Suc n). ideal R (J k)) \<and>
(\<forall>i \<le> (Suc n). \<forall>j \<le> (Suc n). i \<noteq> j \<longrightarrow>
coprime_ideals R (J i) (J j))\<rbrakk> \<Longrightarrow> coprime_ideals R (i\<Pi>\<^bsub>R,n\<^esub> J) (J (Suc n))"
apply (simp add:coprime_n_idealsTr3)
done
section "Direct product1, general case"
definition
prod_tOp :: "['i set, 'i \<Rightarrow> ('a, 'm) Ring_scheme] \<Rightarrow>
('i \<Rightarrow> 'a) \<Rightarrow> ('i \<Rightarrow> 'a) \<Rightarrow> ('i \<Rightarrow> 'a)" where
"prod_tOp I A = (\<lambda>f\<in>carr_prodag I A. \<lambda>g\<in>carr_prodag I A.
\<lambda>x\<in>I. (f x) \<cdot>\<^sub>r\<^bsub>(A x)\<^esub> (g x))"
(** Let x \<in> I, then A x is a ring, {A x | x \<in> I} is a set of rings. **)
definition
prod_one::"['i set, 'i \<Rightarrow> ('a, 'm) Ring_scheme] \<Rightarrow> ('i \<Rightarrow> 'a)" where
"prod_one I A == \<lambda>x\<in>I. 1\<^sub>r\<^bsub>(A x)\<^esub>"
(** 1\<^sub>(A x) is the unit of a ring A x. **)
definition
prodrg :: "['i set, 'i \<Rightarrow> ('a, 'more) Ring_scheme] \<Rightarrow> ('i \<Rightarrow> 'a) Ring" where
"prodrg I A = \<lparr>carrier = carr_prodag I A, pop = prod_pOp I A, mop =
prod_mOp I A, zero = prod_zero I A, tp = prod_tOp I A,
un = prod_one I A \<rparr>"
(** I is the index set **)
abbreviation
PRODRING ("(r\<Pi>\<^bsub>_\<^esub>/ _)" [72,73]72) where
"r\<Pi>\<^bsub>I\<^esub> A == prodrg I A"
definition
augm_func :: "[nat, nat \<Rightarrow> 'a,'a set, nat, nat \<Rightarrow> 'a, 'a set] \<Rightarrow> nat \<Rightarrow> 'a" where
"augm_func n f A m g B = (\<lambda>i\<in>{j. j \<le> (n + m)}. if i \<le> n then f i else
if (Suc n) \<le> i \<and> i \<le> n + m then g ((sliden (Suc n)) i) else undefined)"
(* Remark. g is a function of Nset (m - 1) \<rightarrow> B *)
definition
ag_setfunc :: "[nat, nat \<Rightarrow> ('a, 'more) Ring_scheme, nat,
nat \<Rightarrow> ('a, 'more) Ring_scheme] \<Rightarrow> (nat \<Rightarrow> 'a) set \<Rightarrow> (nat \<Rightarrow> 'a) set
\<Rightarrow> (nat \<Rightarrow> 'a) set" where
"ag_setfunc n B1 m B2 X Y =
{f. \<exists>g. \<exists>h. (g\<in>X) \<and>(h\<in>Y) \<and>(f = (augm_func n g (Un_carrier {j. j \<le> n} B1)
m h (Un_carrier {j. j \<le> (m - 1)} B2)))}"
(* Later we consider X = ac_Prod_Rg (Nset n) B1 and Y = ac_Prod_Rg (Nset (m - 1)) B2 *)
primrec
ac_fProd_Rg :: "[nat, nat \<Rightarrow> ('a, 'more) Ring_scheme] \<Rightarrow>
(nat \<Rightarrow> 'a) set"
where
fprod_0: "ac_fProd_Rg 0 B = carr_prodag {0::nat} B"
| frpod_n: "ac_fProd_Rg (Suc n) B = ag_setfunc n B (Suc 0) (compose {0::nat}
B (slide (Suc n))) (carr_prodag {j. j \<le> n} B) (carr_prodag {0} (compose {0} B (slide (Suc n))))"
definition
prodB1 :: "[('a, 'm) Ring_scheme, ('a, 'm) Ring_scheme] \<Rightarrow>
(nat \<Rightarrow> ('a, 'm) Ring_scheme)" where
"prodB1 R S = (\<lambda>k. if k=0 then R else if k=Suc 0 then S else
undefined)"
definition
Prod2Rg :: "[('a, 'm) Ring_scheme, ('a, 'm) Ring_scheme]
\<Rightarrow> (nat \<Rightarrow> 'a) Ring" (infixl "\<Oplus>\<^sub>r" 80) where
"A1 \<Oplus>\<^sub>r A2 = prodrg {0, Suc 0} (prodB1 A1 A2)"
text \<open>Don't try \<open>(Prod_ring (Nset n) B) \<Oplus>\<^sub>r (B (Suc n))\<close>\<close>
lemma carr_prodrg_mem_eq:"\<lbrakk>f \<in> carrier (r\<Pi>\<^bsub>I\<^esub> A); g \<in> carrier (r\<Pi>\<^bsub>I\<^esub> A);
\<forall>i\<in>I. f i = g i \<rbrakk> \<Longrightarrow> f = g"
by (simp add:prodrg_def carr_prodag_def, (erule conjE)+,
rule funcset_eq[of _ I], assumption+)
lemma prod_tOp_mem:"\<lbrakk>\<forall>k\<in>I. Ring (A k); X \<in> carr_prodag I A;
Y \<in> carr_prodag I A\<rbrakk> \<Longrightarrow> prod_tOp I A X Y \<in> carr_prodag I A"
apply (subst carr_prodag_def, simp)
apply (rule conjI)
apply (simp add:prod_tOp_def restrict_def extensional_def)
apply (rule conjI)
apply (rule Pi_I)
apply (simp add:Un_carrier_def prod_tOp_def)
apply (simp add:carr_prodag_def, (erule conjE)+)
apply (blast dest: Ring.ring_tOp_closed del:PiE)
apply (rule ballI)
apply (simp add:prod_tOp_def)
apply (simp add:carr_prodag_def, (erule conjE)+)
apply (simp add:Ring.ring_tOp_closed)
done
lemma prod_tOp_func:"\<forall>k\<in>I. Ring (A k) \<Longrightarrow>
prod_tOp I A \<in> carr_prodag I A \<rightarrow> carr_prodag I A \<rightarrow> carr_prodag I A"
by (simp add:prod_tOp_mem)
lemma prod_one_func:"\<forall>k\<in>I. Ring (A k) \<Longrightarrow>
prod_one I A \<in> carr_prodag I A"
apply (simp add:prod_one_def carr_prodag_def)
apply (rule conjI)
apply (rule Pi_I)
apply (simp add:Un_carrier_def)
apply (blast dest: Ring.ring_one)
apply (simp add: Ring.ring_one)
done
lemma prodrg_carrier:"\<forall>k\<in>I. Ring (A k) \<Longrightarrow>
carrier (prodrg I A) = carrier (prodag I A)"
by (simp add:prodrg_def prodag_def)
lemma prodrg_ring:"\<forall>k\<in>I. Ring (A k) \<Longrightarrow> Ring (prodrg I A)"
apply (rule Ring.intro)
apply (simp add:prodrg_def)
apply (rule prod_pOp_func,
rule ballI, simp add:Ring.ring_is_ag)
apply (simp add:prodrg_def, rule prod_pOp_assoc,
simp add:Ring.ring_is_ag, assumption+)
apply (simp add:prodrg_def, rule prod_pOp_commute,
simp add:Ring.ring_is_ag, assumption+)
apply (simp add:prodrg_def, rule prod_mOp_func,
simp add:Ring.ring_is_ag)
apply (simp add:prodrg_def)
apply (cut_tac X = a in prod_mOp_mem[of "I" "A"])
apply (simp add:Ring.ring_is_ag, assumption)
apply (cut_tac X = "prod_mOp I A a" and Y = a in prod_pOp_mem[of "I" "A"],
simp add:Ring.ring_is_ag, assumption+,
cut_tac prod_zero_func[of "I" "A"])
apply (rule carr_prodag_mem_eq[of "I" "A"],
simp add:Ring.ring_is_ag, assumption+,
rule ballI, simp add:prod_pOp_def,
subst prod_mOp_mem_i, simp add:Ring.ring_is_ag, assumption+,
subst prod_zero_i, simp add:Ring.ring_is_ag, assumption+,
rule aGroup.l_m, simp add:Ring.ring_is_ag,
simp add:prodag_comp_i, simp add:Ring.ring_is_ag)
apply (simp add:prodrg_def,
rule prod_zero_func, simp add:Ring.ring_is_ag)
apply (simp add:prodrg_def,
cut_tac prod_zero_func[of "I" "A"],
cut_tac X = "prod_zero I A" and Y = a in prod_pOp_mem[of "I" "A"],
simp add:Ring.ring_is_ag, assumption+,
rule carr_prodag_mem_eq[of "I" "A"],
simp add:Ring.ring_is_ag, assumption+)
apply (rule ballI)
apply (simp add:prod_pOp_def prod_zero_def)
apply (rule aGroup.ag_l_zero, simp add:Ring.ring_is_ag)
apply (simp add:prodag_comp_i, simp add:Ring.ring_is_ag)
apply (simp add:prodrg_def,
rule prod_tOp_func, simp add:Ring.ring_is_ag)
apply (simp add:prodrg_def)
apply (frule_tac X = a and Y = b in prod_tOp_mem[of "I" "A"], assumption+,
frule_tac X = "prod_tOp I A a b" and Y = c in
prod_tOp_mem[of "I" "A"], assumption+,
frule_tac X = b and Y = c in prod_tOp_mem[of "I" "A"], assumption+,
frule_tac X = a and Y = "prod_tOp I A b c" in
prod_tOp_mem[of "I" "A"], assumption+)
apply (rule carr_prodag_mem_eq[of "I" "A"],
simp add:Ring.ring_is_ag, assumption+, rule ballI)
apply (simp add:prod_tOp_def)
apply (rule Ring.ring_tOp_assoc, simp, (simp add:prodag_comp_i)+)
apply (simp add:prodrg_def)
apply (frule_tac X = a and Y = b in prod_tOp_mem[of "I" "A"], assumption+,
frule_tac X = b and Y = a in prod_tOp_mem[of "I" "A"], assumption+,
rule carr_prodag_mem_eq[of "I" "A"],
simp add:Ring.ring_is_ag, assumption+)
apply (rule ballI, simp add:prod_tOp_def)
apply (rule Ring.ring_tOp_commute, (simp add:prodag_comp_i)+)
apply (simp add:prodrg_def, rule prod_one_func, assumption)
apply (simp add:prodrg_def)
apply (cut_tac X = b and Y = c in prod_pOp_mem[of "I" "A"],
simp add:Ring.ring_is_ag, assumption+,
cut_tac X = a and Y = b in prod_tOp_mem[of "I" "A"], assumption+,
cut_tac X = a and Y = c in prod_tOp_mem[of "I" "A"], assumption+,
cut_tac X = "prod_tOp I A a b" and Y = "prod_tOp I A a c" in
prod_pOp_mem[of "I" "A"], simp add:Ring.ring_is_ag, assumption+)
apply (rule carr_prodag_mem_eq[of "I" "A"],
simp add:Ring.ring_is_ag, rule prod_tOp_mem[of "I" "A"], assumption+)
apply (rule ballI, simp add:prod_tOp_def prod_pOp_def)
apply (rule Ring.ring_distrib1, (simp add:prodag_comp_i)+)
apply (simp add:prodrg_def,
cut_tac prod_one_func[of "I" "A"],
cut_tac X = "prod_one I A" and Y = a in prod_tOp_mem[of "I" "A"],
assumption+)
apply (rule carr_prodag_mem_eq[of "I" "A"],
simp add:Ring.ring_is_ag, assumption+,
rule ballI, simp add:prod_tOp_def prod_one_def,
rule Ring.ring_l_one, simp, simp add:prodag_comp_i)
apply simp
done
lemma prodrg_elem_extensional:"\<lbrakk>\<forall>k\<in>I. Ring (A k); f \<in> carrier (prodrg I A)\<rbrakk>
\<Longrightarrow> f \<in> extensional I"
apply (simp add:prodrg_carrier)
apply (simp add:prodag_def carr_prodag_def)
done
lemma prodrg_pOp:"\<forall>k\<in>I. Ring (A k) \<Longrightarrow>
pop (prodrg I A) = prod_pOp I A"
apply (simp add:prodrg_def)
done
lemma prodrg_mOp:"\<forall>k\<in>I. Ring (A k) \<Longrightarrow>
mop (prodrg I A) = prod_mOp I A"
apply (simp add:prodrg_def)
done
lemma prodrg_zero:"\<forall>k\<in>I. Ring (A k) \<Longrightarrow>
zero (prodrg I A) = prod_zero I A"
apply (simp add:prodrg_def)
done
lemma prodrg_tOp:"\<forall>k\<in>I. Ring (A k) \<Longrightarrow>
tp (prodrg I A) = prod_tOp I A"
apply (simp add:prodrg_def)
done
lemma prodrg_one:"\<forall>k\<in>I. Ring (A k) \<Longrightarrow>
un (prodrg I A) = prod_one I A"
apply (simp add:prodrg_def)
done
lemma prodrg_sameTr5:"\<lbrakk>\<forall>k\<in>I. Ring (A k); \<forall>k\<in>I. A k = B k\<rbrakk>
\<Longrightarrow> prod_tOp I A = prod_tOp I B"
apply (frule prod_tOp_func)
apply (subgoal_tac "carr_prodag I A = carr_prodag I B", simp)
apply (frule prod_tOp_func [of "I" "B"])
apply (rule funcset_eq [of _ "carr_prodag I B" _])
apply (simp add:prod_tOp_def extensional_def)
apply (simp add:prod_tOp_def extensional_def)
apply (rule ballI)
apply (frule_tac x = x in funcset_mem [of "prod_tOp I A" "carr_prodag I B"
"carr_prodag I B \<rightarrow> carr_prodag I B"], assumption+)
apply (frule_tac x = x in funcset_mem [of "prod_tOp I B" "carr_prodag I B"
"carr_prodag I B \<rightarrow> carr_prodag I B"], assumption+)
apply (thin_tac " prod_tOp I A
\<in> carr_prodag I B \<rightarrow> carr_prodag I B \<rightarrow> carr_prodag I B")
apply (thin_tac "prod_tOp I B
\<in> carr_prodag I B \<rightarrow> carr_prodag I B \<rightarrow> carr_prodag I B")
apply (rule funcset_eq [of _ "carr_prodag I B"])
apply (simp add:prod_tOp_def extensional_def)
apply (simp add:prod_tOp_def extensional_def)
apply (rule ballI)
apply (frule_tac f = "prod_tOp I A x" and A = "carr_prodag I B" and
x = xa in funcset_mem, assumption+)
apply (frule_tac f = "prod_tOp I B x" and A = "carr_prodag I B" and
x = xa in funcset_mem, assumption+)
apply (subgoal_tac "\<forall>k\<in>I. aGroup (B k)")
apply (rule carr_prodag_mem_eq, assumption+)
apply (rule ballI, simp add:prod_tOp_def)
apply (rule ballI, rule Ring.ring_is_ag, simp)
apply (subgoal_tac "\<forall>k\<in>I. aGroup (A k)")
apply (simp add:prodag_sameTr1)
apply (rule ballI, rule Ring.ring_is_ag, simp)
done
lemma prodrg_sameTr6:"\<lbrakk>\<forall>k\<in>I. Ring (A k); \<forall>k\<in>I. A k = B k\<rbrakk>
\<Longrightarrow> prod_one I A = prod_one I B"
apply (frule prod_one_func [of "I" "A"])
apply (cut_tac prodag_sameTr1[of "I" "A" "B"])
apply (rule carr_prodag_mem_eq[of I A "prod_one I A" "prod_one I B"])
apply (simp add:Ring.ring_is_ag, assumption)
apply (cut_tac prod_one_func [of "I" "B"], simp)
apply simp
apply (rule ballI, simp add:prod_one_def)
apply (simp add:Ring.ring_is_ag, simp)
done
lemma prodrg_same:"\<lbrakk>\<forall>k\<in>I. Ring (A k); \<forall>k\<in>I. A k = B k\<rbrakk>
\<Longrightarrow> prodrg I A = prodrg I B"
apply (subgoal_tac "\<forall>k\<in>I. aGroup (A k)")
apply (frule prodag_sameTr1, assumption+)
apply (frule prodag_sameTr2, assumption+)
apply (frule prodag_sameTr3, assumption+)
apply (frule prodag_sameTr4, assumption+)
apply (frule prodrg_sameTr5, assumption+)
apply (frule prodrg_sameTr6, assumption+)
apply (simp add:prodrg_def)
apply (rule ballI, rule Ring.ring_is_ag, simp)
done
lemma prodrg_component:"\<lbrakk>f \<in> carrier (prodrg I A); i \<in> I\<rbrakk> \<Longrightarrow>
f i \<in> carrier (A i)"
by (simp add:prodrg_def carr_prodag_def)
lemma project_rhom:"\<lbrakk>\<forall>k\<in>I. Ring (A k); j \<in> I\<rbrakk> \<Longrightarrow>
PRoject I A j \<in> rHom ( prodrg I A) (A j)"
apply (simp add:rHom_def)
apply (rule conjI)
apply (simp add:aHom_def)
apply (rule conjI)
apply (rule Pi_I)
apply (simp add:prodrg_carrier)
apply (cut_tac prodag_carrier[of I A], simp)
apply (simp add:PRoject_def)
apply (cut_tac prodag_carrier[of I A], simp,
thin_tac "carrier (a\<Pi>\<^bsub>I\<^esub> A) = carr_prodag I A")
apply (simp add:prodag_comp_i)
apply (simp add:Ring.ring_is_ag)
apply (simp add:Ring.ring_is_ag)
apply (subgoal_tac "\<forall>k\<in>I. aGroup (A k)")
apply (frule project_aHom [of "I" "A" "j"], assumption+)
apply (rule conjI)
apply (simp add:prodrg_carrier aHom_def)
apply (simp add:prodrg_carrier)
apply (simp add:prodrg_pOp)
apply (simp add:prodag_pOp[THEN sym])
apply (simp add:aHom_def)
apply (rule ballI, simp add:Ring.ring_is_ag)
apply (rule conjI)
apply (rule ballI)+
apply (simp add:prodrg_carrier)
apply (cut_tac prodag_carrier[of I A], simp)
apply (frule_tac X = x and Y = y in prod_tOp_mem[of I A], assumption+)
apply (simp add:prodrg_tOp)
apply (simp add:PRoject_def)
apply (simp add:prod_tOp_def)
apply (rule ballI)
apply (simp add:Ring.ring_is_ag)
apply (frule prodrg_ring [of "I" "A"])
apply (frule Ring.ring_one[of "r\<Pi>\<^bsub>I\<^esub> A"])
apply (simp add:prodrg_carrier)
apply (cut_tac prodag_carrier[of I A], simp)
apply (simp add:PRoject_def) apply (simp add:prodrg_def)
apply (fold prodrg_def) apply (simp add:prod_one_def)
apply (rule ballI)
apply (simp add:Ring.ring_is_ag)
done
lemma augm_funcTr:"\<lbrakk>\<forall>k \<le>(Suc n). Ring (B k);
f \<in> carr_prodag {i. i \<le> (Suc n)} B\<rbrakk> \<Longrightarrow>
f = augm_func n (restrict f {i. i \<le> n}) (Un_carrier {i. i \<le> n} B) (Suc 0)
(\<lambda>x\<in>{0::nat}. f (x + Suc n))
(Un_carrier {0} (compose {0} B (slide (Suc n))))"
apply (rule carr_prodag_mem_eq [of "{i. i \<le> (Suc n)}" "B" "f"
"augm_func n (restrict f {i. i \<le> n}) (Un_carrier {i. i \<le> n} B) (Suc 0)
(\<lambda>x\<in>{0}. f (x + Suc n)) (Un_carrier {0} (compose {0} B (slide (Suc n))))"])
apply (rule ballI, simp add:Ring.ring_is_ag)
apply (simp add:augm_func_def)
apply (simp add:carr_prodag_def)
apply (rule conjI)
apply (simp add:augm_func_def)
apply (rule conjI)
apply (rule Pi_I)
apply (simp add:augm_func_def sliden_def)
apply (erule conjE)+
apply (frule_tac x = x in funcset_mem[of f "{i. i \<le> Suc n}"
"Un_carrier {i. i \<le> Suc n} B"]) apply simp
apply simp
apply (rule impI)
apply (rule_tac x = "Suc n" in funcset_mem[of f "{i. i \<le> Suc n}"
"Un_carrier {i. i \<le> Suc n} B"], assumption) apply simp
apply (rule allI, (erule conjE)+)
apply (simp add:augm_func_def)
apply (case_tac "i \<le> n", simp add:sliden_def)
apply (simp add:sliden_def, rule impI)
apply (simp add:nat_not_le_less,
frule_tac m = n and n = i in Suc_leI,
frule_tac m = i and n = "Suc n" in Nat.le_antisym, assumption+,
simp)
apply (rule ballI, simp)
apply (simp add:augm_func_def sliden_def)
apply (rule impI, simp add:nat_not_le_less)
apply (frule_tac n = l in Suc_leI[of n _])
apply (frule_tac m = l and n = "Suc n" in Nat.le_antisym, assumption+,
simp)
done
lemma A_to_prodag_mem:"\<lbrakk>Ring A; \<forall>k\<in>I. Ring (B k); \<forall>k\<in>I. (S k) \<in>
rHom A (B k); x \<in> carrier A \<rbrakk> \<Longrightarrow> A_to_prodag A I S B x \<in> carr_prodag I B"
apply (simp add:carr_prodag_def)
apply (rule conjI)
apply (simp add:A_to_prodag_def extensional_def)
apply (rule conjI)
apply (rule Pi_I)
apply (simp add: A_to_prodag_def)
apply (subgoal_tac "(S xa) \<in> rHom A (B xa)") prefer 2 apply simp
apply (thin_tac "\<forall>k\<in>I. S k \<in> rHom A (B k)")
apply (frule_tac f = "S xa" and A = A and R = "B xa" in rHom_mem, assumption+)
apply (simp add:Un_carrier_def) apply blast
apply (rule ballI)
apply (simp add:A_to_prodag_def)
apply (rule_tac f = "S i" and A = A and R = "B i" and a = x in rHom_mem)
apply simp
apply assumption
done
lemma A_to_prodag_rHom:"\<lbrakk>Ring A; \<forall>k\<in>I. Ring (B k); \<forall>k\<in>I. (S k) \<in>
rHom A (B k) \<rbrakk> \<Longrightarrow> A_to_prodag A I S B \<in> rHom A (r\<Pi>\<^bsub>I\<^esub> B)"
apply (simp add:rHom_def [of "A" "r\<Pi>\<^bsub>I\<^esub> B"])
apply (rule conjI)
apply (cut_tac A_to_prodag_aHom[of A I B S])
apply (subst aHom_def, simp)
apply (simp add:prodrg_carrier)
apply (simp add:aHom_def)
apply (simp add:prodrg_def)
apply (cut_tac prodag_pOp[of I B], simp)
apply (rule ballI, simp add:Ring.ring_is_ag,
simp add:Ring.ring_is_ag,
rule ballI, simp add:Ring.ring_is_ag)
apply (rule ballI)
apply (simp add:rHom_def)
apply (rule conjI)
apply (rule ballI)+
apply (frule_tac x = x and y = y in Ring.ring_tOp_closed[of A], assumption+)
apply (frule_tac x = "x \<cdot>\<^sub>r\<^bsub>A\<^esub> y" in A_to_prodag_mem[of A I B S], assumption+,
frule_tac x = x in A_to_prodag_mem[of A I B S], assumption+,
frule_tac x = y in A_to_prodag_mem[of A I B S], assumption+)
apply (simp add:prodrg_tOp[of I B])
apply (frule_tac X = "A_to_prodag A I S B x " and Y = "A_to_prodag A I S B y" in prod_tOp_mem[of I B], assumption+)
apply (cut_tac X = "A_to_prodag A I S B (x \<cdot>\<^sub>r\<^bsub>A\<^esub> y)" and Y = "prod_tOp I B (A_to_prodag A I S B x) (A_to_prodag A I S B y)" in carr_prodag_mem_eq[of I B])
apply (rule ballI, simp add:Ring.ring_is_ag) apply assumption+
apply (rule ballI, simp add:prod_tOp_def A_to_prodag_def)
apply (frule_tac x = l in bspec, assumption,
thin_tac "\<forall>k\<in>I. Ring (B k)",
frule_tac x = l in bspec, assumption,
thin_tac "\<forall>k\<in>I. S k \<in> rHom A (B k)")
apply (simp add:rHom_tOp) apply simp
apply (simp add:prodrg_one[of I B])
apply (frule prod_one_func[of I B])
apply (frule Ring.ring_one[of A],
frule_tac x = "1\<^sub>r\<^bsub>A\<^esub>" in A_to_prodag_mem[of A I B S], assumption+)
apply (cut_tac X = "A_to_prodag A I S B 1\<^sub>r\<^bsub>A\<^esub>" and Y = "prod_one I B" in
carr_prodag_mem_eq[of I B])
apply (rule ballI, simp add:Ring.ring_is_ag)
apply assumption+
apply (rule ballI)
apply (subst A_to_prodag_def, simp add:prod_one_def)
apply (frule_tac x = l in bspec, assumption,
thin_tac "\<forall>k\<in>I. Ring (B k)",
frule_tac x = l in bspec, assumption,
thin_tac "\<forall>k\<in>I. S k \<in> rHom A (B k)")
apply (simp add:rHom_one)
apply assumption
done
lemma ac_fProd_ProdTr1:"\<forall>k \<le> (Suc n). Ring (B k) \<Longrightarrow>
ag_setfunc n B (Suc 0) (compose {0::nat} B (slide (Suc n)))
(carr_prodag {i. i \<le> n} B) (carr_prodag {0}
(compose {0} B (slide (Suc n)))) \<subseteq> carr_prodag {i. i \<le> (Suc n)} B"
+supply [[simproc del: defined_all]]
apply (rule subsetI)
apply (simp add:ag_setfunc_def)
apply (erule exE, erule conjE, erule exE, erule conjE)
apply (simp,
thin_tac "x =
augm_func n g (Un_carrier {j. j \<le> n} B) (Suc 0) h
(Un_carrier {0} (compose {0} B (slide (Suc n))))")
apply (simp add:carr_prodag_def [of "{j. j \<le> (Suc n)}" "B"])
apply (rule conjI)
apply (simp add:augm_func_def)
apply (rule conjI)
apply (simp add:Pi_def) apply (rule allI) apply (rule impI)
apply (simp add:augm_func_def)
apply (case_tac "x \<le> n")
apply simp apply (simp add:carr_prodag_def)
apply (erule conjE)+ apply (frule_tac x = x in mem_of_Nset [of _ "n"])
apply (frule_tac f = g and x = x in funcset_mem[of _ "{j. j \<le> n}"
"Un_carrier {j. j \<le> n} B"], assumption+)
apply (simp add:Un_carrier_def,
erule exE, erule conjE, erule exE, simp, erule conjE,
frule_tac x = i and y = n and z = "Suc n" in le_less_trans,
simp,
frule_tac x = i and y = "Suc n" in less_imp_le, blast)
apply (simp add:sliden_def)
apply (simp add:carr_prodag_def Un_carrier_def, (erule conjE)+)
apply (simp add:compose_def slide_def)
apply (cut_tac n_in_Nsetn[of "Suc n"], blast)
apply (rule allI, rule impI)
apply (simp add:augm_func_def)
apply (case_tac "i \<le> n", simp)
apply (simp add:carr_prodag_def [of "{i. i \<le> n}" _])
apply simp apply (thin_tac "g \<in> carr_prodag {i. i \<le> n} B")
apply (simp add: not_less [symmetric, of _ n],
frule_tac n = i in Suc_leI[of n],
frule_tac m = i and n = "Suc n" in le_antisym, assumption+, simp)
apply (simp add:carr_prodag_def compose_def slide_def sliden_def)
done
lemma ac_fProd_Prod:"\<forall>k \<le> n. Ring (B k) \<Longrightarrow>
ac_fProd_Rg n B = carr_prodag {j. j \<le> n} B"
apply (case_tac "n = 0")
apply simp
apply (subgoal_tac "\<exists>m. n = Suc m")
apply (subgoal_tac "\<forall>m. n = Suc m \<longrightarrow>
ac_fProd_Rg n B = carr_prodag {j. j \<le> n} B")
apply blast apply (thin_tac "\<exists>m. n = Suc m")
apply (rule allI, rule impI) apply (simp, thin_tac "n = Suc m")
apply (rule equalityI)
apply (simp add:ac_fProd_ProdTr1)
apply (rule subsetI)
apply (rename_tac m f)
apply (frule augm_funcTr, assumption+)
apply (simp add:ag_setfunc_def)
apply (subgoal_tac "(restrict f {j. j \<le> m}) \<in> carr_prodag {j. j \<le> m} B")
apply (subgoal_tac "(\<lambda>x\<in>{0::nat}. f (Suc (x + m))) \<in> carr_prodag {0}
(compose {0} B (slide (Suc m)))")
apply blast
apply (thin_tac "f =
augm_func m (restrict f {i. i \<le> m}) (Un_carrier {i. i \<le> m} B)
(Suc 0) (\<lambda>x\<in>{0}. f (Suc (x + m)))
(Un_carrier {0} (compose {0} B (slide (Suc m))))")
apply (simp add:carr_prodag_def)
apply (rule conjI)
apply (simp add:Pi_def restrict_def)
apply (simp add:Un_carrier_def compose_def slide_def)
apply (simp add:compose_def slide_def)
apply (thin_tac "f =
augm_func m (restrict f {i. i \<le> m}) (Un_carrier {i. i \<le> m} B)
(Suc 0) (\<lambda>x\<in>{0}. f (Suc (x + m)))
(Un_carrier {0} (compose {0} B (slide (Suc m))))")
apply (simp add:carr_prodag_def)
apply (simp add:Un_carrier_def)
apply (simp add:Pi_def)
apply (rule allI) apply (rule impI)
apply (erule conjE)+
apply (rotate_tac 1)
apply (frule_tac a = x in forall_spec, simp)
apply (erule exE,
thin_tac "\<forall>x\<le>Suc m. \<exists>xa. (\<exists>i\<le>Suc m. xa = carrier (B i)) \<and> f x \<in> xa")
apply (frule_tac a = x in forall_spec, simp)
apply blast
apply (cut_tac t = n in Suc_pred[THEN sym], simp)
apply blast
done
text\<open>A direct product of a finite number of rings defined with
\<open>ac_fProd_Rg\<close> is equal to that defined by using \<open>carr_prodag\<close>.\<close>
definition
fprodrg :: "[nat, nat \<Rightarrow> ('a, 'more) Ring_scheme] \<Rightarrow>
\<lparr>carrier:: (nat \<Rightarrow> 'a) set, pop::[(nat \<Rightarrow> 'a), (nat \<Rightarrow> 'a)]
\<Rightarrow> (nat \<Rightarrow> 'a), mop:: (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a), zero::(nat \<Rightarrow> 'a),
tp :: [(nat \<Rightarrow> 'a), (nat \<Rightarrow> 'a)] \<Rightarrow> (nat \<Rightarrow> 'a), un :: (nat \<Rightarrow> 'a) \<rparr>" where
"fprodrg n B = \<lparr> carrier = ac_fProd_Rg n B,
pop = \<lambda>f. \<lambda>g. prod_pOp {i. i \<le> n} B f g, mop = \<lambda>f. prod_mOp {i. i \<le> n} B f,
zero = prod_zero {i. i \<le> n} B, tp = \<lambda>f. \<lambda>g. prod_tOp {i. i \<le> n} B f g,
un = prod_one {i. i \<le> n} B \<rparr>"
definition
fPRoject ::"[nat, nat \<Rightarrow> ('a, 'more) Ring_scheme, nat]
\<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> 'a" where
"fPRoject n B x = (\<lambda>f\<in>ac_fProd_Rg n B. f x)"
lemma fprodrg_ring:"\<forall>k \<le> n. Ring (B k) \<Longrightarrow> Ring (fprodrg n B)"
apply (simp add:fprodrg_def)
apply (frule ac_fProd_Prod)
apply simp
apply (fold prodrg_def)
apply (simp add:prodrg_ring)
done
section "Chinese remainder theorem"
lemma Chinese_remTr1:"\<lbrakk>Ring A; \<forall>k \<le> (n::nat). ideal A (J k);
\<forall>k \<le> n. B k = qring A (J k); \<forall>k \<le> n. S k = pj A (J k) \<rbrakk> \<Longrightarrow>
ker\<^bsub>A,(r\<Pi>\<^bsub>{j. j \<le> n}\<^esub> B)\<^esub> (A_to_prodag A {j. j \<le> n} S B) =
\<Inter> {I. \<exists>k\<in>{j. j \<le> n}. I = (J k)}"
apply (rule equalityI)
apply (rule subsetI)
apply (simp add:ker_def)
apply (rule allI, rule impI)
apply (rename_tac a K, erule conjE)
apply (simp add:prodrg_def, simp add:A_to_prodag_def prod_one_def)
apply (erule exE, erule conjE)
apply (subgoal_tac "(\<lambda>k\<in>{j. j \<le> n}. S k a) k = (\<lambda>x\<in>{j. j \<le> n}. \<zero>\<^bsub>B x\<^esub>) k")
apply (thin_tac "(\<lambda>k\<in>{j. j \<le> n}. S k a) = prod_zero {j. j \<le> n} B")
apply simp apply (frule_tac I = "J k" in Ring.qring_zero [of "A"])
apply simp
apply (frule_tac I = "J k" and x = a in pj_mem [of "A"]) apply simp
apply assumption apply simp
apply (frule_tac I = "J k" and a = a in Ring.a_in_ar_coset [of "A"])
apply simp apply assumption apply simp
apply (simp add:prod_zero_def)
apply (rule subsetI)
apply (simp add:CollectI ker_def)
apply (cut_tac Nset_inc_0[of n])
apply (frule_tac a = "J 0" in forall_spec, blast)
apply (frule_tac x = 0 in spec, simp)
apply (frule_tac h = x in Ring.ideal_subset [of "A" "J 0"], simp+)
apply (thin_tac "x \<in> J 0")
apply (simp add:A_to_prodag_def prodrg_def)
apply (simp add:prod_zero_def)
apply (rule funcset_eq [of _ "{j. j \<le> n}"])
apply (simp add:extensional_def restrict_def)+
apply (rule allI, rule impI)
apply (simp add:Ring.qring_zero)
apply (frule_tac a = xa in forall_spec, assumption,
thin_tac "\<forall>k \<le> n. ideal A (J k)")
apply (subst pj_mem [of "A"], assumption+)
apply (frule_tac I = "J xa" and a = x in Ring.a_in_ar_coset [of "A"],
assumption+)
apply (rule_tac a = x and I = "J xa" in Ring.Qring_fix1 [of "A"], assumption+)
apply blast
done
lemma (in Ring) coprime_prod_int2Tr:
"((\<forall>k \<le> (Suc n). ideal R (J k)) \<and>
(\<forall>i \<le> (Suc n). \<forall>j \<le> (Suc n). (i \<noteq>j \<longrightarrow> coprime_ideals R (J i) (J j))))
\<longrightarrow> (\<Inter> {I. \<exists>k \<le> (Suc n). I = (J k)} = ideal_n_prod R (Suc n) J)"
apply (induct_tac n)
apply (rule impI)
apply (erule conjE)
apply (simp add:Nset_1)
apply (subst coprime_int_prod [THEN sym, of "J 0" "J (Suc 0)"], simp+)
apply (rule equalityI, rule subsetI)
apply (simp, blast)
apply (rule subsetI, simp, rule allI, rule impI, erule exE, (erule conjE)+)
apply simp
apply (simp add:Nset_1_1, erule disjE, (simp del:ideal_n_prodSn)+)
(** n **)
apply (rule impI)
apply (subgoal_tac "\<Inter>{I. \<exists>k \<le> (Suc (Suc n)). I = J k} =
(\<Inter>{I. \<exists>k \<le> (Suc n). I = J k}) \<inter> (J (Suc (Suc n)))")
apply (subgoal_tac "\<Inter>{I. \<exists>k \<le> (Suc n). I = J k} = (i\<Pi>\<^bsub>R,(Suc n)\<^esub> J)")
(* apply (simp del:ideal_n_prodSn)*)
apply (frule_tac n = "Suc n" and J = J in coprime_n_idealsTr4)
apply (simp del:ideal_n_prodSn)
apply (subst coprime_int_prod)
apply (rule n_prod_ideal)
apply (rule allI, simp, simp, assumption)
apply simp apply (cut_tac n = "Suc n" in Nsetn_sub_mem1)
apply simp
apply (thin_tac "(\<forall>k\<le>Suc n. ideal R (J k)) \<and>
(\<forall>i\<le>Suc n. \<forall>j\<le>Suc n. i \<noteq> j \<longrightarrow> coprime_ideals R (J i) (J j)) \<longrightarrow>
\<Inter>{I. \<exists>k\<le>Suc n. I = J k} = i\<Pi>\<^bsub>R,Suc n\<^esub> J",
thin_tac "(\<forall>k\<le>Suc (Suc n). ideal R (J k)) \<and>
(\<forall>i\<le>Suc (Suc n).
\<forall>j\<le>Suc (Suc n). i \<noteq> j \<longrightarrow> coprime_ideals R (J i) (J j))")
apply (rule equalityI, rule subsetI, simp)
apply (rule conjI,
rule allI, rule impI, erule exE, erule conjE, simp,
frule_tac a = xa in forall_spec,
frule_tac x = k and y = "Suc n" and z = "Suc (Suc n)" in
le_less_trans, simp,
frule_tac x = k and y = "Suc (Suc n)" in less_imp_le, blast)
apply simp
apply (frule_tac a = "J (Suc (Suc n))" in forall_spec,
cut_tac n = "Suc (Suc n)" in le_refl, blast, simp)
apply (rule subsetI, simp, rule allI, rule impI)
apply (erule exE, erule conjE)
apply (erule conjE,
case_tac "k = Suc (Suc n)", simp)
apply (frule_tac m = k and n = "Suc (Suc n)" in noteq_le_less, assumption,
thin_tac "k \<le> Suc (Suc n)")
apply (frule_tac x = k and n = "Suc n" in Suc_less_le)
apply (frule_tac a = xa in forall_spec,
blast,
thin_tac "\<forall>xa. (\<exists>k\<le>Suc n. xa = J k) \<longrightarrow> x \<in> xa",
simp)
done
lemma (in Ring) coprime_prod_int2:"\<lbrakk> \<forall>k \<le> (Suc n). ideal R (J k);
\<forall>i \<le> (Suc n). \<forall>j \<le> (Suc n). (i \<noteq>j \<longrightarrow> coprime_ideals R (J i) (J j))\<rbrakk>
\<Longrightarrow> (\<Inter> {I. \<exists>k \<le> (Suc n). I = (J k)} = ideal_n_prod R (Suc n) J)"
apply (simp add:coprime_prod_int2Tr)
done
lemma (in Ring) coprime_2_n:"\<lbrakk>ideal R A; ideal R B\<rbrakk> \<Longrightarrow>
(qring R A) \<Oplus>\<^sub>r (qring R B) = r\<Pi>\<^bsub>{j. j \<le> (Suc 0)}\<^esub> (prodB1 (qring R A) (qring R B))"
apply (simp add:Prod2Rg_def Nset_1)
done
text\<open>In this and following lemmata, ideals A and B are of type
\<open>('a, 'more) RingType_scheme\<close>. Don't try
\<open>(r\<Pi>\<^sub>(Nset n) B) \<Oplus>\<^sub>r B (Suc n)\<close>\<close>
lemma (in Ring) A_to_prodag2_hom:"\<lbrakk>ideal R A; ideal R B; S 0 = pj R A;
S (Suc 0) = pj R B\<rbrakk> \<Longrightarrow>
A_to_prodag R {j. j \<le> (Suc 0)} S (prodB1 (qring R A) (qring R B)) \<in>
rHom R (qring R A \<Oplus>\<^sub>r qring R B)"
apply (subst coprime_2_n [of "A" "B"], assumption+)
apply (rule A_to_prodag_rHom, rule Ring_axioms)
apply (rule ballI)
apply (case_tac "k = 0")
apply (simp add:prodB1_def)
apply (simp add:qring_ring)
apply (simp)
apply (frule_tac n = k in Suc_leI[of 0], thin_tac "0 < k")
apply (frule_tac m = k and n = "Suc 0" in le_antisym, assumption)
apply (simp, simp add:prodB1_def, simp add:qring_ring)
apply (rule ballI)
apply (simp add:Nset_1)
apply (erule disjE)
apply (simp add:prodB1_def, rule pj_Hom, rule Ring_axioms, assumption)
apply (simp, simp add:prodB1_def)
apply (rule pj_Hom, rule Ring_axioms, assumption+)
done
lemma (in Ring) A2coprime_rsurjecTr:"\<lbrakk>ideal R A; ideal R B; S 0 = pj R A;
S (Suc 0) = pj R B\<rbrakk> \<Longrightarrow>
(carrier (qring R A \<Oplus>\<^sub>r qring R B)) =
carr_prodag {j. j \<le> (Suc 0)} (prodB1 (qring R A) (qring R B))"
apply (simp add:Prod2Rg_def prodrg_def Nset_1)
done
lemma (in Ring) A2coprime_rsurjec:"\<lbrakk>ideal R A; ideal R B; S 0 = pj R A;
S (Suc 0) = pj R B; coprime_ideals R A B\<rbrakk> \<Longrightarrow>
surjec\<^bsub>R,((qring R A) \<Oplus>\<^sub>r (qring R B))\<^esub>
(A_to_prodag R {j. j\<le>(Suc 0)} S (prodB1 (qring R A) (qring R B)))"
apply (frule A_to_prodag2_hom [of "A" "B" "S"], assumption+)
apply (simp add:surjec_def)
apply (rule conjI,
simp add:rHom_def)
apply (rule surj_to_test[of "A_to_prodag R {j. j \<le> (Suc 0)} S
(prodB1 (qring R A) (qring R B))" "carrier R"
"carrier (qring R A \<Oplus>\<^sub>r qring R B)"])
apply (simp add:rHom_def aHom_def)
apply (rule ballI)
apply (frule rHom_func[of "A_to_prodag R {j. j \<le> (Suc 0)} S
(prodB1 (R /\<^sub>r A) (R /\<^sub>r B))" R],
thin_tac "A_to_prodag R {j. j \<le> (Suc 0)} S (prodB1 (R /\<^sub>r A) (R /\<^sub>r B))
\<in> rHom R (R /\<^sub>r A \<Oplus>\<^sub>r R /\<^sub>r B)")
apply (simp add:A2coprime_rsurjecTr[of A B S])
apply (simp add:Nset_1)
apply (frule_tac X = "b 0" and Y = "b (Suc 0)" in
coprime_surjTr[of A B], assumption+)
apply (simp add:carr_prodag_def prodB1_def,
simp add:carr_prodag_def prodB1_def)
apply (erule bexE)
apply (frule_tac x = r in funcset_mem[of "A_to_prodag R {0, Suc 0} S
(prodB1 (R /\<^sub>r A) (R /\<^sub>r B))"
"carrier R" "carr_prodag {0, Suc 0} (prodB1 (R /\<^sub>r A) (R /\<^sub>r B))"],
assumption+)
apply (cut_tac X = "A_to_prodag R {0, Suc 0} S (prodB1 (R /\<^sub>r A) (R /\<^sub>r B)) r"
and Y = b in
carr_prodag_mem_eq[of "{0, Suc 0}" "prodB1 (R /\<^sub>r A) (R /\<^sub>r B)"])
apply (rule ballI)
apply (simp, erule disjE)
apply (simp add:prodB1_def, fold prodB1_def,
simp add:qring_ring Ring.ring_is_ag)
apply (simp add:prodB1_def, fold prodB1_def,
simp add:qring_ring Ring.ring_is_ag)
apply assumption+
apply (rule ballI, simp, erule disjE, simp)
apply (subst A_to_prodag_def, simp)
apply (subst A_to_prodag_def, simp)
apply blast
done
lemma (in Ring) prod2_n_Tr1:"\<lbrakk>\<forall>k \<le> (Suc 0). ideal R (J k);
\<forall>k \<le> (Suc 0). B k = qring R (J k);
\<forall>k \<le> (Suc 0). S k = pj R (J k) \<rbrakk> \<Longrightarrow>
A_to_prodag R {j. j \<le> (Suc 0)} S
(prodB1 (qring R (J 0)) (qring R (J (Suc 0)))) =
A_to_prodag R {j. j \<le> (Suc 0)} S B"
apply (subgoal_tac "\<forall>k \<le> (Suc 0). (prodB1 (qring R (J 0)) (qring R (J (Suc 0)))) k = B k")
apply (simp add:A_to_prodag_def)
apply (rule allI, rule impI)
apply (case_tac "k = 0", simp add:Nset_1_1)
apply (simp add:prodB1_def)
apply (simp add:Nset_1_1)
apply (simp add:prodB1_def)
done
lemma (in aGroup) restrict_prod_Suc:"\<lbrakk>\<forall>k \<le> (Suc (Suc n)). ideal R (J k);
\<forall>k \<le> (Suc (Suc n)). B k = R /\<^sub>r J k;
\<forall>k \<le> (Suc (Suc n)). S k = pj R (J k);
f \<in> carrier (r\<Pi>\<^bsub>{j. j \<le> (Suc (Suc n))}\<^esub> B)\<rbrakk> \<Longrightarrow>
restrict f {j. j \<le> (Suc n)} \<in> carrier (r\<Pi>\<^bsub>{j. j \<le> (Suc n)}\<^esub> B)"
apply (cut_tac Nsetn_sub_mem1[of "Suc n"])
apply (simp add:prodrg_def)
apply (simp add:carr_prodag_def, (erule conjE)+)
apply (simp add:Un_carrier_def)
apply (rule Pi_I)
apply simp
apply (frule_tac x = x in funcset_mem[of f "{j. j \<le> (Suc (Suc n))}"
"\<Union>{X. \<exists>i \<le> (Suc (Suc n)). X = carrier (R /\<^sub>r J i)}"],
simp)
apply simp
apply (erule exE, erule conjE, erule exE, erule conjE, simp)
apply (rotate_tac -5)
apply (frule_tac a = x in forall_spec) apply simp
apply blast
done
lemma (in Ring) Chinese_remTr2:"(\<forall>k \<le> (Suc n). ideal R (J k)) \<and>
(\<forall>k\<le>(Suc n). B k = qring R (J k)) \<and>
(\<forall>k\<le>(Suc n). S k = pj R (J k)) \<and>
(\<forall>i\<le>(Suc n). \<forall>j\<le> (Suc n). (i \<noteq>j \<longrightarrow>
coprime_ideals R (J i) (J j))) \<longrightarrow>
surjec\<^bsub>R,(r\<Pi>\<^bsub>{j. j\<le> (Suc n)}\<^esub> B)\<^esub>
(A_to_prodag R {j. j\<le>(Suc n)} S B)"
apply (cut_tac Ring)
apply (induct_tac n)
(* case n = 0, i.e. two coprime_ideals *) (** use coprime_surjTr **)
apply (rule impI) apply (erule conjE)+
apply (frule A_to_prodag_rHom [of R "{j. j \<le> Suc 0}" "B" "S"])
apply (rule ballI, simp add:Ring.qring_ring)
apply (rule ballI, simp add:pj_Hom)
apply (frule rHom_func[of "A_to_prodag R {j. j \<le> (Suc 0)} S B" R
"r\<Pi>\<^bsub>{j. j \<le> (Suc 0)}\<^esub> B"])
apply (simp add:surjec_def)
apply (rule conjI)
apply (simp add:rHom_def)
apply (rule surj_to_test, assumption+)
apply (rule ballI) apply (simp add:Nset_1)
apply (cut_tac coprime_elems [of "J 0" "J (Suc 0)"])
apply (rename_tac f)
apply (erule bexE, erule bexE)
apply (simp add:prodrg_def) apply (fold prodrg_def)
apply (cut_tac X = "f 0" and Y = "f (Suc 0)" in
coprime_surjTr[of "J 0" "J (Suc 0)"], simp+)
apply (simp add:carr_prodag_def, simp add:carr_prodag_def)
apply (erule bexE, (erule conjE)+)
apply (frule_tac x = r in funcset_mem[of "A_to_prodag R {0, Suc 0} S B"
"carrier R" "carr_prodag {0, Suc 0} B"], assumption+)
apply (cut_tac X = "A_to_prodag R {0, Suc 0} S B r" and Y = f in
carr_prodag_mem_eq[of "{0, Suc 0}" B])
apply (rule ballI, simp, erule disjE, simp add:qring_ring
Ring.ring_is_ag,
simp add:Ring.qring_ring Ring.ring_is_ag)
apply assumption+
apply (rule ballI, simp, erule disjE, simp)
apply (simp add:A_to_prodag_def, simp add:A_to_prodag_def)
apply blast apply simp+
apply (rule impI, (erule conjE)+)
(**** n ****)
apply (cut_tac n = "Suc n" in Nsetn_sub_mem1)
apply (subgoal_tac "\<forall>k\<in>{i. i \<le> Suc (Suc n)}. Ring (B k)")
apply (frule_tac I = "{i. i \<le> Suc (Suc n)}" in A_to_prodag_rHom [of "R" _ "B" "S"])
apply assumption
apply (rule ballI)
apply (simp add:pj_Hom)
apply simp
apply (subst surjec_def, rule conjI,
simp add:rHom_def)
apply (cut_tac n = "Suc n" in coprime_n_idealsTr4[of _ J])
apply blast
apply (frule_tac f = "A_to_prodag R {j. j \<le> (Suc (Suc n))} S B" and
A = R in rHom_func)
apply (rule_tac f = "A_to_prodag R {j. j \<le> (Suc (Suc n))} S B" and
A = "carrier R" and B = "carrier (r\<Pi>\<^bsub>{j. j \<le> (Suc (Suc n))}\<^esub> B)" in
surj_to_test, assumption+)
apply (rule ballI)
apply (cut_tac n = "Suc n" in n_prod_ideal[of _ J])
apply (rule allI, simp)
apply (frule_tac A = "i\<Pi>\<^bsub>R,(Suc n)\<^esub> J" and B = "J (Suc (Suc n))" in
coprime_elems,
cut_tac n = "Suc (Suc n)" in n_in_Nsetn,
blast, assumption)
apply (erule bexE, erule bexE) apply (rename_tac n f a b)
apply (thin_tac " coprime_ideals R (i\<Pi>\<^bsub>R,(Suc n)\<^esub> J) (J (Suc (Suc n)))")
apply (cut_tac n = "Suc n" and a = a and J = J in ele_n_prod,
rule allI, simp, assumption)
apply (cut_tac ring_is_ag)
apply (frule_tac n = n and f = f in aGroup.restrict_prod_Suc[of R _ R J B S],
assumption+)
apply (frule_tac S = "r\<Pi>\<^bsub>{j. j \<le> (Suc n)}\<^esub> B" and
f = "A_to_prodag R {j. j \<le> (Suc n)} S B" in surjec_surj_to[of R])
apply (frule_tac f = "A_to_prodag R {j. j \<le> (Suc n)} S B" and A = "carrier R"
and B = "carrier (r\<Pi>\<^bsub>{j. j \<le> (Suc n)}\<^esub> B)" and
b = "restrict f {j. j \<le> (Suc n)}" in surj_to_el2, assumption)
apply (erule bexE, rename_tac n f a b u)
apply (cut_tac n = "Suc (Suc n)" in n_in_Nsetn,
frule_tac f = f and I = "{j. j \<le> (Suc (Suc n))}" and A = B and
i = "Suc (Suc n)" in prodrg_component, assumption)
apply simp
apply (frule_tac J = "J (Suc (Suc n))" and X = "f (Suc (Suc n))" in
pj_surj_to[of R], simp, assumption)
apply (erule bexE, rename_tac n f a b u v)
apply (frule_tac a = "Suc (Suc n)" in forall_spec, simp,
frule_tac I = "J (Suc (Suc n))" and h = b in Ring.ideal_subset[of R],
assumption+,
cut_tac I = "i\<Pi>\<^bsub>R,n\<^esub> J \<diamondsuit>\<^sub>r\<^bsub>R\<^esub> J (Suc n)" and h = a in
Ring.ideal_subset[of R], assumption+)
apply (frule_tac x = b and y = u in Ring.ring_tOp_closed[of R], assumption+,
frule_tac x = a and y = v in Ring.ring_tOp_closed[of R], assumption+,
frule Ring.ring_is_ag[of R],
frule_tac x = "b \<cdot>\<^sub>r\<^bsub>R\<^esub> u" and y = "a \<cdot>\<^sub>r\<^bsub>R\<^esub> v" in aGroup.ag_pOp_closed[of R],
assumption+)
apply (frule_tac f = "A_to_prodag R {j. j \<le> (Suc (Suc n))} S B" and
A = "carrier R" and B = "carrier (r\<Pi>\<^bsub>{j. j \<le> (Suc (Suc n))}\<^esub> B)" and
x = "b \<cdot>\<^sub>r\<^bsub>R\<^esub> u \<plusminus>\<^bsub>R\<^esub> a \<cdot>\<^sub>r\<^bsub>R\<^esub> v" in funcset_mem, assumption+)
apply (frule_tac f = "A_to_prodag R {j. j \<le> (Suc (Suc n))} S B
(b \<cdot>\<^sub>r\<^bsub>R\<^esub> u \<plusminus>\<^bsub>R\<^esub> a \<cdot>\<^sub>r\<^bsub>R\<^esub> v)" and I = "{j. j \<le> (Suc (Suc n))}"
and g = f in carr_prodrg_mem_eq, simp)
apply (rule ballI)
apply (subst A_to_prodag_def, simp)
apply (frule_tac I = "J i" in pj_Hom[of R], simp)
apply (simp add: rHom_add)
apply (frule_tac a = i in forall_spec, assumption,
thin_tac "\<forall>k \<le> (Suc (Suc n)). ideal R (J k)",
frule_tac I = "J i" in Ring.qring_ring[of R], assumption,
thin_tac "\<forall>k \<le> (Suc (Suc n)). Ring (R /\<^sub>r J k)",
frule_tac R = "R /\<^sub>r (J i)" and x = b and y = u and f = "pj R (J i)" in
rHom_tOp[of R], assumption+, simp,
thin_tac "pj R (J i) (b \<cdot>\<^sub>r\<^bsub>R\<^esub> u) = pj R (J i) b \<cdot>\<^sub>r\<^bsub>R /\<^sub>r J i\<^esub> pj R (J i) u",
frule_tac R = "R /\<^sub>r (J i)" and x = a and y = v and f = "pj R (J i)" in
rHom_tOp[of R], simp add:Ring.qring_ring, assumption+)
apply (frule_tac f = "pj R (J i)" and R = "R /\<^sub>r J i" and a = v in
rHom_mem[of _ R], assumption+,
frule_tac f = "pj R (J i)" and R = "R /\<^sub>r J i" and a = u in
rHom_mem[of _ R], assumption+,
frule_tac f = "pj R (J i)" and R = "R /\<^sub>r J i" and a = b in
rHom_mem[of _ R], assumption+,
frule_tac f = "pj R (J i)" and R = "R /\<^sub>r J i" and a = a in
rHom_mem[of _ R], assumption+)
apply (frule_tac R = "R /\<^sub>r J i" in Ring.ring_is_ag)
apply (case_tac "i \<le> (Suc n)")
apply (frule_tac I1 = "J i" and x1 = a in pj_zero[THEN sym, of R ],
assumption+, simp,
thin_tac "pj R (J i) (a \<cdot>\<^sub>r\<^bsub>R\<^esub> v) = \<zero>\<^bsub>R /\<^sub>r J i\<^esub> \<cdot>\<^sub>r\<^bsub>R /\<^sub>r J i\<^esub> pj R (J i) v")
apply (simp add:Ring.ring_times_0_x)
apply (frule_tac f = "pj R (J i)" and A = R and R = "R /\<^sub>r J i" and
x = a and y = b in rHom_add, assumption+, simp,
thin_tac "A_to_prodag R {j. j \<le> Suc (Suc n)} S B
(b \<cdot>\<^sub>r u) \<plusminus>\<^bsub>r\<Pi>\<^bsub>{i. i \<le> Suc (Suc n)}\<^esub> B\<^esub>
A_to_prodag R {j. j \<le> Suc (Suc n)} S B (a \<cdot>\<^sub>r v)
\<in> carrier (r\<Pi>\<^bsub>{j. j \<le> Suc (Suc n)}\<^esub> B)")
apply (simp add:aGroup.ag_l_zero)
apply (rotate_tac -1, frule sym, thin_tac " pj R (J i) 1\<^sub>r\<^bsub>R\<^esub> = pj R (J i) b",
simp add:rHom_one) apply (simp add:Ring.ring_l_one)
apply (simp add:aGroup.ag_r_zero)
apply (frule_tac f = "A_to_prodag R {j. j \<le> (Suc n)} S B u" and
g = "restrict f {j. j \<le> (Suc n)}" and x = i in eq_fun_eq_val,
thin_tac "A_to_prodag R {j. j\<le>(Suc n)} S B u = restrict f {j. j\<le>(Suc n)}")
apply (simp add:A_to_prodag_def)
apply simp
apply (frule_tac y = i and x = "Suc n" in not_le_imp_less,
frule_tac m = "Suc n" and n = i in Suc_leI,
frule_tac m = i and n = "Suc (Suc n)" in Nat.le_antisym, assumption+,
simp)
apply (frule_tac I1 = "J (Suc (Suc n))" and x1 = b in pj_zero[THEN sym, of
R ], assumption+, simp add:Ring.ring_times_0_x)
apply (frule_tac f = "pj R (J (Suc (Suc n)))" and A = R and
R = "R /\<^sub>r J (Suc (Suc n))" and
x = a and y = b in rHom_add, assumption+, simp)
apply (simp add:aGroup.ag_r_zero)
apply (rotate_tac -1, frule sym,
thin_tac "pj R (J (Suc (Suc n))) 1\<^sub>r\<^bsub>R\<^esub> = pj R (J (Suc (Suc n))) a",
simp add:rHom_one,
simp add:Ring.ring_l_one)
apply (simp add:aGroup.ag_l_zero)
apply blast
apply (rule ballI, simp add:Ring.qring_ring)
done
lemma (in Ring) Chinese_remTr3:"\<lbrakk>\<forall>k \<le> (Suc n). ideal R (J k);
\<forall>k \<le> (Suc n). B k = qring R (J k); \<forall>k\<le> (Suc n). S k = pj R (J k);
\<forall>i \<le> (Suc n). \<forall>j \<le> (Suc n). (i \<noteq>j \<longrightarrow> coprime_ideals R (J i) (J j))\<rbrakk> \<Longrightarrow>
surjec\<^bsub>R,(r\<Pi>\<^bsub>{j. j \<le> (Suc n)}\<^esub> B)\<^esub>
(A_to_prodag R {j. j \<le> (Suc n)} S B)"
apply (simp add:Chinese_remTr2 [of "n" "J" "B" "S"])
done
lemma (in Ring) imset:"\<lbrakk>\<forall>k\<le> (Suc n). ideal R (J k)\<rbrakk>
\<Longrightarrow> {I. \<exists>k\<le> (Suc n). I = J k} = {J k| k. k \<in> {j. j \<le> (Suc n)}}"
apply blast
done
theorem (in Ring) Chinese_remThm:"\<lbrakk>(\<forall>k \<le> (Suc n). ideal R (J k));
\<forall>k\<le>(Suc n). B k = qring R (J k); \<forall>k \<le> (Suc n). S k = pj R (J k);
\<forall>i \<le> (Suc n). \<forall>j \<le> (Suc n). (i \<noteq>j \<longrightarrow> coprime_ideals R (J i) (J j))\<rbrakk>
\<Longrightarrow> bijec\<^bsub>(qring R (\<Inter> {J k | k. k\<in>{j. j \<le> (Suc n)}})),(r\<Pi>\<^bsub>{j. j \<le> (Suc n)}\<^esub> B)\<^esub>
((A_to_prodag R {j. j \<le> (Suc n)} S B)\<degree>\<^bsub>R,(prodrg {j. j \<le> (Suc n)} B)\<^esub>)"
apply (frule Chinese_remTr3, assumption+)
apply (cut_tac I = "{j. j \<le> (Suc n)}" and A = B in prodrg_ring)
apply (rule ballI, simp add:qring_ring)
apply (cut_tac surjec_ind_bijec [of "R" "r\<Pi>\<^bsub>{j. j \<le> (Suc n)}\<^esub> B"
"A_to_prodag R {j. j \<le> (Suc n)} S B"])
apply (cut_tac Ring,
simp add:Chinese_remTr1 [of "R" "Suc n" "J" "B" "S"])
apply (simp add:imset, rule Ring_axioms, assumption+)
apply (rule A_to_prodag_rHom, rule Ring_axioms)
apply (rule ballI)
apply (simp add:qring_ring)
apply (rule ballI, simp, rule pj_Hom, rule Ring_axioms, simp)
apply assumption
done
lemma (in Ring) prod_prime:"\<lbrakk>ideal R A; \<forall>k\<le>(Suc n). prime_ideal R (P k);
\<forall>l\<le>(Suc n). \<not> (A \<subseteq> P l);
\<forall>k\<le> (Suc n). \<forall>l\<le> (Suc n). k = l \<or> \<not> (P k) \<subseteq> (P l)\<rbrakk> \<Longrightarrow>
\<forall>i \<le> (Suc n). (nprod R (ppa R P A i) n \<in> A \<and>
(\<forall>l \<in> {j. j\<le>(Suc n)} - {i}. nprod R (ppa R P A i) n \<in> P l) \<and>
(nprod R (ppa R P A i) n \<notin> P i))"
apply (rule allI, rule impI)
apply (rule conjI)
apply (frule_tac i = i in prod_primeTr1[of n P A], assumption+)
apply (frule_tac n = n and f = "ppa R P A i" in ideal_nprod_inc[of A])
apply (rule allI, rule impI)
apply (rotate_tac -2,
frule_tac a = ia in forall_spec, assumption,
thin_tac "\<forall>l \<le> n.
ppa R P A i l \<in> A \<and>
ppa R P A i l \<in> P (skip i l) \<and> ppa R P A i l \<notin> P i",
simp add:ideal_subset)
apply (rotate_tac -1,
frule_tac a = n in forall_spec, simp,
thin_tac "\<forall>l\<le> n.
ppa R P A i l \<in> A \<and>
ppa R P A i l \<in> P (skip i l) \<and> ppa R P A i l \<notin> P i",
(erule conjE)+,
blast, assumption)
apply (frule_tac i = i in prod_primeTr1[of n P A], assumption+)
apply (rule conjI)
apply (rule ballI)
apply (frule_tac a = l in forall_spec, simp,
frule_tac I = "P l" in prime_ideal_ideal)
apply (frule_tac n = n and f = "ppa R P A i" and A = "P l" in ideal_nprod_inc)
apply (rule allI) apply (simp add:ppa_mem, simp)
apply (case_tac "l < i",
thin_tac "\<forall>l\<le> (Suc n). \<not> A \<subseteq> P l",
thin_tac "\<forall>k\<le> (Suc n). \<forall>l \<le> (Suc n). k = l \<or> \<not> P k \<subseteq> P l")
apply (erule conjE,
frule_tac x = l and y = i and z = "Suc n" in less_le_trans,
assumption,
frule_tac x = l and n = n in Suc_less_le)
apply (rotate_tac 2,
frule_tac a = l in forall_spec, assumption,
thin_tac "\<forall>l\<le>n. ppa R P A i l \<in> A \<and>
ppa R P A i l \<in> P (skip i l) \<and> ppa R P A i l \<notin> P i",
thin_tac "l < Suc n")
apply (simp only:skip_im_Tr1_2, blast)
apply (frule_tac x = l and y = i in leI,
thin_tac "\<not> l < i",
cut_tac x = l and A = "{j. j \<le> (Suc n)}" and a = i in in_diff1)
apply simp
apply (erule conjE,
frule not_sym, thin_tac "l \<noteq> i",
frule_tac x = i and y = l in le_imp_less_or_eq,
erule disjE, thin_tac "i \<le> l",
frule_tac x = i and n = l in less_le_diff)
apply (cut_tac i = i and n = n and x = "l - Suc 0" in skip_im_Tr2_1,
simp, assumption+, simp,
frule_tac x = l and n = n in le_Suc_diff_le)
apply (rotate_tac -7,
frule_tac a = "l - Suc 0" in forall_spec, assumption,
thin_tac "\<forall>l\<le>n. ppa R P A i l \<in> A \<and>
ppa R P A i l \<in> P (skip i l) \<and> ppa R P A i l \<notin> P i",
simp, (erule conjE)+)
apply blast
apply simp
apply assumption
apply (frule_tac a = i in forall_spec, assumption,
thin_tac "\<forall>k\<le> (Suc n). prime_ideal R (P k)")
apply (rule_tac P = "P i" and n = n and f = "ppa R P A i" in
prime_nprod_exc, assumption+)
apply (rule allI, rule impI)
apply (rotate_tac -3,
frule_tac a = ia in forall_spec, assumption,
thin_tac "\<forall>l \<le> n.
ppa R P A i l \<in> A \<and>
ppa R P A i l \<in> P (skip i l) \<and> ppa R P A i l \<notin> P i",
simp add:ideal_subset)
apply (rule allI, rule impI) apply (
rotate_tac 4,
frule_tac a = l in forall_spec, assumption,
thin_tac "\<forall>l\<le> n.
ppa R P A i l \<in> A \<and>
ppa R P A i l \<in> P (skip i l) \<and> ppa R P A i l \<notin> P i",
simp)
done
lemma skip_im1:"\<lbrakk>i \<le> (Suc n); P \<in> {j. j \<le> (Suc n)} \<rightarrow> Collect (prime_ideal R)\<rbrakk>
\<Longrightarrow>
compose {j. j \<le> n} P (skip i) ` {j. j \<le> n} = P ` ({j. j \<le> (Suc n)} - {i})"
apply (cut_tac skip_fun[of i n])
apply (subst setim_cmpfn[of _ _ _ _ "{X. prime_ideal R X}"], assumption+)
apply simp
apply (simp add:skip_fun_im)
done
lemma (in Ring) mutch_aux1:"\<lbrakk>ideal R A; i \<le> (Suc n);
P \<in> {j. j \<le> (Suc n)} \<rightarrow> Collect (prime_ideal R)\<rbrakk> \<Longrightarrow>
compose {j. j \<le> n} P (skip i) \<in> {j. j \<le> n} \<rightarrow> Collect (prime_ideal R)"
apply (cut_tac skip_fun[of i n])
apply (simp add:composition[of "skip i" "{j. j \<le> n}" "{j. j \<le> (Suc n)}" P
"Collect (prime_ideal R)"])
done
lemma (in Ring) prime_ideal_cont1Tr:"ideal R A \<Longrightarrow>
\<forall>P. ((P \<in> {j. j \<le> (n::nat)} \<rightarrow> {X. prime_ideal R X}) \<and>
(A \<subseteq> \<Union> (P ` {j. j \<le> n}))) \<longrightarrow> (\<exists>i\<le> n. A \<subseteq> (P i))"
apply (induct_tac n)
apply (rule allI, rule impI)
apply (erule conjE)
apply simp
(** n **)
apply (rule allI, rule impI)
apply (erule conjE)+
apply (case_tac "\<exists>i \<le> (Suc n). \<exists>j\<le> (Suc n). (i \<noteq> j \<and> P i \<subseteq> P j)")
apply ((erule exE, erule conjE)+, erule conjE)
apply (frule_tac f = P and n = n and X = "{X. prime_ideal R X}" and
A = A and i = i and j = j in Un_less_Un, assumption+, simp+)
apply (frule mutch_aux1, assumption+)
apply (frule_tac a = "compose {j. j \<le> n} P (skip i)" in forall_spec,
simp, erule exE)
apply (cut_tac i = i and n = n and x = ia in skip_fun_im1,
simp+, erule conjE, simp add:compose_def,blast)
(** last_step induction assumption is of no use **)
apply (thin_tac "\<forall>P. P \<in> {j. j \<le> n} \<rightarrow> {X. prime_ideal R X} \<and>
A \<subseteq> \<Union>(P ` {j. j \<le> n}) \<longrightarrow>
(\<exists>i\<le>n. A \<subseteq> P i)",
rule contrapos_pp, simp+)
apply (cut_tac n = n and P = P in prod_prime [of A], assumption)
apply (rule allI, rule impI,
frule_tac f = P and A = "{j. j \<le> (Suc n)}" and B = "{X. prime_ideal R X}"
and x = k in funcset_mem, simp, simp, assumption+)
apply (frule_tac n = "Suc n" and
f = "\<lambda>i\<in>{j. j \<le> (Suc n)}. (nprod R (ppa R P A i) n)" in
nsum_ideal_inc[of A], rule allI, rule impI, simp)
apply (subgoal_tac "(nsum R (\<lambda>i\<in>{j. j \<le> (Suc n)}. nprod R (ppa R P A i) n)
(Suc n)) \<notin> (\<Union>x\<in>{j. j \<le> (Suc n)}. P x)")
apply blast
apply (simp del:nsum_suc)
apply (rule allI, rule impI) apply (rename_tac n P l)
apply (frule_tac f = P and A = "{j. j \<le> (Suc n)}" and
B = "{X. prime_ideal R X}"
and x = l in funcset_mem, simp, simp del:nsum_suc,
frule_tac I = "P l" in prime_ideal_ideal)
apply (rule_tac A = "P l" and n = "Suc n" and
f = "\<lambda>i\<in>{j. j \<le> (Suc n)}. (nprod R (ppa R P A i) n)" in
nsum_ideal_exc, simp+, rule allI, simp add:ideal_subset)
apply (rule contrapos_pp, simp+)
apply (rotate_tac -1,
frule_tac a = l in forall_spec, simp,
thin_tac "\<forall>j\<le>Suc n.
(\<exists>la\<in>{i. i \<le> Suc n} - {j}. e\<Pi>\<^bsub>R,n\<^esub> ppa R P A la \<notin> P l) \<or>
e\<Pi>\<^bsub>R,n\<^esub> ppa R P A j \<in> P l",
thin_tac "\<forall>i\<le>Suc n. \<forall>j\<le>Suc n. i = j \<or> \<not> P i \<subseteq> P j",
thin_tac "\<forall>i\<le>Suc n. \<not> A \<subseteq> P i")
apply (erule disjE, erule bexE)
apply (frule_tac a = la in forall_spec, simp,
thin_tac "\<forall>i\<le>Suc n.
e\<Pi>\<^bsub>R,n\<^esub> ppa R P A i \<in> A \<and>
(\<forall>l\<in>{j. j \<le> Suc n} - {i}. e\<Pi>\<^bsub>R,n\<^esub> ppa R P A i \<in> P l) \<and>
e\<Pi>\<^bsub>R,n\<^esub> ppa R P A i \<notin> P i",
(erule conjE)+)
apply blast
apply blast
done
lemma (in Ring) prime_ideal_cont1:"\<lbrakk>ideal R A; \<forall>i \<le> (n::nat).
prime_ideal R (P i); A \<subseteq> \<Union> {X. (\<exists>i \<le> n. X = (P i))} \<rbrakk> \<Longrightarrow>
\<exists>i\<le> n. A\<subseteq>(P i)"
apply (frule prime_ideal_cont1Tr[of A n])
apply (frule_tac a = P in forall_spec,
thin_tac "\<forall>P. P \<in> {j. j \<le> n} \<rightarrow> {X. prime_ideal R X} \<and>
A \<subseteq> \<Union>(P ` {j. j \<le> n}) \<longrightarrow> (\<exists>i\<le>n. A \<subseteq> P i)")
apply (rule conjI, simp,
rule subsetI, simp,
frule_tac c = x in subsetD[of A "\<Union>{X. \<exists>i\<le>n. X = P i}"], assumption+,
simp, blast)
apply assumption
done
lemma (in Ring) prod_n_ideal_contTr0:"(\<forall>l\<le> n. ideal R (J l)) \<longrightarrow>
i\<Pi>\<^bsub>R,n\<^esub> J \<subseteq> \<Inter>{X. (\<exists>k\<le>n. X = (J k))}"
apply (induct_tac n)
apply simp
apply (rule impI)
apply (cut_tac n = n in Nsetn_sub_mem1,
simp)
apply (cut_tac n = n in n_prod_ideal[of _ J], simp)
apply (cut_tac I = "i\<Pi>\<^bsub>R,n\<^esub> J" and J = "J (Suc n)" in
ideal_prod_sub_Int) apply assumption apply simp
apply (frule_tac A = "i\<Pi>\<^bsub>R,n\<^esub> J \<diamondsuit>\<^sub>r\<^bsub>R\<^esub> J (Suc n)" and
B = "i\<Pi>\<^bsub>R,n\<^esub> J \<inter> J (Suc n)" and
C = "\<Inter>{X. \<exists>k\<le> n. X = J k} \<inter> J (Suc n)" in subset_trans)
apply (rule_tac A = "i\<Pi>\<^bsub>R,n\<^esub> J" and B = "\<Inter>{X. \<exists>k\<le>n. X = J k}" and
C = "J (Suc n)" in inter_mono, assumption)
apply (rule_tac A = "i\<Pi>\<^bsub>R,n\<^esub> J \<diamondsuit>\<^sub>r J (Suc n)" and
B = "\<Inter>{X. \<exists>k\<le> n. X = J k} \<inter> J (Suc n)" and
C = "\<Inter>{X. \<exists>k\<le> (Suc n). X = J k}" in subset_trans,
assumption)
apply (rule subsetI)
apply simp
apply (rule allI, rule impI)
apply (erule exE, (erule conjE)+)
apply (case_tac "k = Suc n", simp)
apply (frule_tac m = k and n = "Suc n" in noteq_le_less, assumption)
apply (thin_tac " k \<le> Suc n")
apply (frule_tac x = k and n = "Suc n" in less_le_diff,
thin_tac "k < Suc n", simp, thin_tac "\<forall>l\<le>Suc n. ideal R (J l)")
apply (frule_tac a = xa in forall_spec, blast,
thin_tac "\<forall>xa. (\<exists>k\<le>n. xa = J k) \<longrightarrow> x \<in> xa",
simp)
done
lemma (in Ring) prod_n_ideal_contTr:"\<lbrakk>\<forall>l\<le> n. ideal R (J l)\<rbrakk> \<Longrightarrow>
i\<Pi>\<^bsub>R,n\<^esub> J \<subseteq> \<Inter>{X. (\<exists>k \<le> n. X = (J k))}"
apply (simp add:prod_n_ideal_contTr0)
done
lemma (in Ring) prod_n_ideal_cont2:"\<lbrakk>\<forall>l\<le> (n::nat). ideal R (J l);
prime_ideal R P; \<Inter>{X. (\<exists>k\<le> n. X = (J k))} \<subseteq> P\<rbrakk> \<Longrightarrow>
\<exists>l\<le> n. (J l) \<subseteq> P"
apply (frule prod_n_ideal_contTr[of n J])
apply (frule_tac A = "i\<Pi>\<^bsub>R,n\<^esub> J" and B = "\<Inter>{X. \<exists>k\<le> n. X = J k}" and C = P
in subset_trans, assumption+)
apply (thin_tac "\<Inter>{X. \<exists>k\<le> n. X = J k} \<subseteq> P",
thin_tac "i\<Pi>\<^bsub>R,n\<^esub> J \<subseteq> \<Inter>{X. \<exists>k\<le> n. X = J k}")
apply (simp add:ideal_n_prod_prime)
done
lemma (in Ring) prod_n_ideal_cont3:"\<lbrakk>\<forall>l\<le> (n::nat). ideal R (J l);
prime_ideal R P; \<Inter>{X. (\<exists>k\<le> n. X = (J k))} = P\<rbrakk> \<Longrightarrow>
\<exists>l\<le> n. (J l) = P"
apply (frule prod_n_ideal_cont2[of n J P], assumption+)
apply simp
apply (erule exE)
apply (subgoal_tac "J l = P")
apply blast
apply (rule equalityI, simp)
apply (rule subsetI)
apply (rotate_tac -4, frule sym, thin_tac "\<Inter>{X. \<exists>k\<le> n. X = J k} = P")
apply simp
apply blast
done
definition
ideal_quotient :: "[_ , 'a set, 'a set] \<Rightarrow> 'a set" where
"ideal_quotient R A B = {x| x. x \<in> carrier R \<and> (\<forall>b\<in>B. x \<cdot>\<^sub>r\<^bsub>R\<^esub> b \<in> A)}"
abbreviation
IDEALQT ("(3_/ \<dagger>\<^sub>_/ _)" [82,82,83]82) where
"A \<dagger>\<^sub>R B == ideal_quotient R A B"
lemma (in Ring) ideal_quotient_is_ideal:
"\<lbrakk>ideal R A; ideal R B\<rbrakk> \<Longrightarrow> ideal R (ideal_quotient R A B)"
apply (rule ideal_condition)
apply (rule subsetI)
apply (simp add:ideal_quotient_def CollectI)
apply (simp add:ideal_quotient_def)
apply (cut_tac ring_zero)
apply (subgoal_tac "\<forall>b\<in>B. \<zero> \<cdot>\<^sub>r b \<in> A")
apply blast
apply (rule ballI)
apply (frule_tac h = b in ideal_subset[of B], assumption)
apply (frule_tac x = b in ring_times_0_x )
apply (simp add:ideal_zero)
apply (rule ballI)+
apply (simp add:ideal_quotient_def, (erule conjE)+,
rule conjI)
apply (rule ideal_pOp_closed)
apply (simp add:whole_ideal, assumption+)
apply (cut_tac ring_is_ag)
apply (simp add:aGroup.ag_mOp_closed)
apply (rule ballI)
apply (subst ring_distrib2)
apply (simp add:ideal_subset, assumption)
apply (cut_tac ring_is_ag, simp add: aGroup.ag_mOp_closed)
apply (frule_tac a1 = y and b1 = b in ring_inv1_1 [THEN sym])
apply (simp add:ideal_subset, simp)
apply (rule ideal_pOp_closed, assumption+, simp)
apply (rule ideal_inv1_closed, assumption+, simp)
apply (rule ballI)+
apply (simp add:ideal_quotient_def)
apply (rule conjI)
apply (erule conjE)
apply (simp add:ring_tOp_closed)
apply (erule conjE)
apply (rule ballI)
apply (subst ring_tOp_assoc, assumption+, simp add:ideal_subset)
apply (simp add:ideal_ring_multiple [of "A"])
done
section \<open>Addition of finite elements of a ring and \<open>ideal_multiplication\<close>\<close>
text\<open>We consider sum in an abelian group\<close>
lemma (in aGroup) nsum_mem1Tr:" A +> J \<Longrightarrow>
(\<forall>j \<le> n. f j \<in> J) \<longrightarrow> nsum A f n \<in> J"
apply (induct_tac n)
apply (rule impI)
apply simp
apply (rule impI)
apply simp
apply (rule asubg_pOp_closed, assumption+)
apply simp
done
lemma (in aGroup) fSum_mem:"\<lbrakk>\<forall>j \<in> nset (Suc n) m. f j \<in> carrier A; n < m\<rbrakk> \<Longrightarrow>
fSum A f (Suc n) m \<in> carrier A"
apply (simp add:fSum_def)
apply (rule nsum_mem)
apply (rule allI, simp add:cmp_def slide_def)
apply (rule impI)
apply (frule_tac x = "Suc (n + j)" in bspec)
apply (simp add:nset_def, arith)
done
lemma (in aGroup) nsum_mem1:"\<lbrakk>A +> J; \<forall>j \<le> n. f j \<in> J\<rbrakk> \<Longrightarrow> nsum A f n \<in> J"
apply (simp add:nsum_mem1Tr)
done
lemma (in aGroup) nsum_eq_i:"\<lbrakk>\<forall>j\<le>n. f j \<in> carrier A; \<forall>j\<le>n. g j \<in> carrier A;
i \<le> n; \<forall>l \<le> i. f l = g l\<rbrakk> \<Longrightarrow> nsum A f i = nsum A g i"
apply (rule nsum_eq)
apply (rule allI, rule impI, simp)+
done
lemma (in aGroup) nsum_cmp_eq:"\<lbrakk>f \<in> {j. j\<le>(n::nat)} \<rightarrow> carrier A;
h1 \<in> {j. j \<le> n} \<rightarrow> {j. j \<le> n}; h2 \<in> {j. j \<le> n} \<rightarrow> {j. j \<le> n}; i \<le> n\<rbrakk> \<Longrightarrow>
nsum A (cmp f (cmp h2 h1)) i = nsum A (cmp (cmp f h2) h1) i"
apply (rule nsum_eq_i [of n "cmp f (cmp h2 h1)" "cmp (cmp f h2) h1" i])
apply (rule allI, rule impI, simp add:cmp_def)
apply ((rule funcset_mem, assumption)+, simp)
apply (rule allI, rule impI, simp add:cmp_def,
(rule funcset_mem, assumption)+, simp+)
apply (rule allI, rule impI, simp add:cmp_def)
done
lemma (in aGroup) nsum_cmp_eq_transpos:"\<lbrakk> \<forall>j\<le>(Suc n). f j \<in> carrier A;
i \<le> n;i \<noteq> n \<rbrakk> \<Longrightarrow>
nsum A (cmp f (cmp (transpos i n) (cmp (transpos n (Suc n)) (transpos i n))))
(Suc n) = nsum A (cmp f (transpos i (Suc n))) (Suc n)"
apply (rule nsum_eq [of "Suc n" "cmp f (cmp (transpos i n)
(cmp (transpos n (Suc n)) (transpos i n)))"
"cmp f (transpos i (Suc n))"])
apply (rule allI, rule impI)
apply (simp add:cmp_def)
apply (cut_tac i = i and n = "Suc n" and j = n and l = j in transpos_mem,
simp+)
apply (cut_tac i = n and n = "Suc n" and j = "Suc n" and l = "transpos i n j"
in transpos_mem, simp+)
apply (cut_tac i = i and n = "Suc n" and j = n and
l = "transpos n (Suc n) (transpos i n j)" in transpos_mem,
simp+)
apply (rule allI, rule impI, simp add:cmp_def)
apply (cut_tac i = i and n = "Suc n" and j = "Suc n" and l = j in transpos_mem,
simp+)
apply (rule allI, rule impI)
apply (simp add:cmp_def)
apply (thin_tac "\<forall>j\<le>Suc n. f j \<in> carrier A",
rule eq_elems_eq_val[of _ _ f])
apply (simp add:transpos_def)
done
lemma transpos_Tr_n1:"Suc (Suc 0) \<le> n \<Longrightarrow>
transpos (n - Suc 0) n n = n - Suc 0"
apply (simp add:transpos_def)
done
lemma transpos_Tr_n2:"Suc (Suc 0) \<le> n \<Longrightarrow>
transpos (n - (Suc 0)) n (n - (Suc 0)) = n"
apply (simp add:transpos_def)
done
lemma (in aGroup) additionTr0:"\<lbrakk>0 < n; \<forall>j \<le> n. f j \<in> carrier A\<rbrakk>
\<Longrightarrow> nsum A (cmp f (transpos (n - 1) n)) n = nsum A f n"
apply (case_tac "n \<le> 1")
apply simp
apply (frule Suc_leI [of "0" "n"])
apply (frule le_antisym [of "n" "Suc 0"], assumption+, simp)
apply (simp add:cmp_def)
apply (subst transpos_ij_1[of 0 "Suc 0"], simp+)
apply (subst transpos_ij_2[of 0 "Suc 0"], simp+)
apply (rule ag_pOp_commute, simp+)
apply (frule not_le_imp_less[of n "Suc 0"])
apply (frule_tac Suc_leI [of "Suc 0" "n"],
thin_tac "\<not> n \<le> Suc 0")
apply (cut_tac nsum_suc[of A f "n - Suc 0"], simp)
apply (cut_tac nsum_suc[of A "cmp f (transpos (n - Suc 0) n)" "n - Suc 0"],
simp,
thin_tac "\<Sigma>\<^sub>e A f n = \<Sigma>\<^sub>e A f (n - Suc 0) \<plusminus> f n",
thin_tac "\<Sigma>\<^sub>e A (cmp f (transpos (n - Suc 0) n)) n =
\<Sigma>\<^sub>e A (cmp f (transpos (n - Suc 0) n)) (n - Suc 0) \<plusminus>
(cmp f (transpos (n - Suc 0) n) n)")
apply (case_tac "n = Suc (Suc 0)", simp)
apply (cut_tac transpos_id_1[of "Suc 0" "Suc (Suc 0)" "Suc (Suc 0)" 0],
cut_tac transpos_ij_1[of "Suc 0" "Suc (Suc 0)" "Suc (Suc 0)"],
cut_tac transpos_ij_2[of "Suc 0" "Suc (Suc 0)" "Suc (Suc 0)"],
simp add:cmp_def,
thin_tac "n = Suc (Suc 0)",
thin_tac "transpos (Suc 0) (Suc (Suc 0)) 0 = 0",
thin_tac "transpos (Suc 0) (Suc (Suc 0)) (Suc 0) = Suc (Suc 0)",
thin_tac "transpos (Suc 0) (Suc (Suc 0)) (Suc (Suc 0)) = Suc 0")
apply (subst ag_pOp_assoc, simp+)
apply (subst ag_pOp_commute[of "f (Suc (Suc 0))" "f (Suc 0)"], simp+)
apply (subst ag_pOp_assoc[THEN sym], simp+)
apply (frule not_sym)
apply (frule noteq_le_less[of "Suc (Suc 0)" n], assumption,
thin_tac "Suc (Suc 0) \<le> n")
apply (cut_tac nsum_suc[of A f "n - Suc 0 - Suc 0"])
apply (cut_tac Suc_pred[of "n - Suc 0"], simp del:nsum_suc)
apply (cut_tac nsum_suc[of A "cmp f (transpos (n - Suc 0) n)"
"n - Suc (Suc 0)"], simp del:nsum_suc,
thin_tac "\<Sigma>\<^sub>e A f (n - Suc 0) = \<Sigma>\<^sub>e A f (n - Suc (Suc 0)) \<plusminus> f (n - Suc 0)",
thin_tac "Suc (n - Suc (Suc 0)) = n - Suc 0",
thin_tac "\<Sigma>\<^sub>e A (cmp f (transpos (n - Suc 0) n)) (n - Suc 0) =
\<Sigma>\<^sub>e A (cmp f (transpos (n - Suc 0) n)) (n - Suc (Suc 0)) \<plusminus>
(cmp f (transpos (n - Suc 0) n)) (n - Suc 0)")
apply (cut_tac nsum_eq_i[of n "cmp f (transpos (n - Suc 0) n)" f
"n - Suc (Suc 0)"], simp,
thin_tac "\<Sigma>\<^sub>e A (cmp f (transpos (n - Suc 0) n)) (n - Suc (Suc 0)) =
\<Sigma>\<^sub>e A f (n - Suc (Suc 0))")
apply (simp add:cmp_def)
apply (cut_tac transpos_ij_1[of "n - Suc 0" n n], simp)
apply (cut_tac transpos_ij_2[of "n - Suc 0" n n], simp)
apply (subst ag_pOp_assoc,
rule nsum_mem, rule allI, rule impI)
apply (frule_tac x = j and y = "n - Suc (Suc 0)" and z = n in
le_less_trans, simp, frule_tac x = j and y = n in less_imp_le)
apply simp+
apply (subst ag_pOp_commute[of "f n"], simp, simp)
apply (subst ag_pOp_assoc[THEN sym],
rule nsum_mem, rule allI, rule impI,
frule_tac x = j and y = "n - Suc (Suc 0)" and z = n in
le_less_trans, simp, frule_tac x = j and y = n in less_imp_le)
apply simp+
apply (rule allI, rule impI, simp add:cmp_def)
apply (cut_tac i = "n - Suc 0" and n = n and j = n and l = j in transpos_mem,
simp+)
apply (rule allI, rule impI)
apply (simp add:cmp_def)
apply (cut_tac i = "n - Suc 0" and n = n and j = n and x = l in transpos_id,
simp+)
apply (cut_tac x = l and y = "n - Suc (Suc 0)" and z = n in le_less_trans,
assumption) apply simp
apply arith
apply simp
apply arith
done
lemma (in aGroup) additionTr1:"\<lbrakk> \<forall>f. \<forall>h. f \<in> {j. j\<le>(Suc n)} \<rightarrow> carrier A \<and>
h \<in> {j. j\<le>(Suc n)} \<rightarrow> {j. j\<le>(Suc n)} \<and> inj_on h {j. j\<le>(Suc n)} \<longrightarrow>
nsum A (cmp f h) (Suc n) = nsum A f (Suc n);
f \<in> {j. j\<le>(Suc (Suc n))} \<rightarrow> carrier A;
h \<in> {j. j\<le>(Suc (Suc n))} \<rightarrow> {j. j\<le>(Suc (Suc n))};
inj_on h {j. j\<le>(Suc (Suc n))}; h (Suc (Suc n)) = Suc (Suc n)\<rbrakk>
\<Longrightarrow> nsum A (cmp f h) (Suc (Suc n)) = nsum A f (Suc (Suc n))"
apply (subgoal_tac "f \<in> {j. j\<le>(Suc n)} \<rightarrow> carrier A")
apply (subgoal_tac "h \<in> {j. j\<le>(Suc n)} \<rightarrow> {j. j\<le>(Suc n)}")
apply (subgoal_tac "inj_on h {j. j\<le>(Suc n)}")
apply (subgoal_tac "nsum A (cmp f h) (Suc n) = nsum A f (Suc n)")
apply (thin_tac "\<forall>f. \<forall>h. f \<in> {j. j\<le>(Suc n)} \<rightarrow> carrier A \<and>
h \<in> {j. j\<le>(Suc n)} \<rightarrow> {j. j\<le>(Suc n)} \<and> inj_on h {j. j\<le>(Suc n)} \<longrightarrow>
nsum A (cmp f h) (Suc n) = nsum A f (Suc n)")
prefer 2 apply simp
apply simp
apply (thin_tac "nsum A (cmp f h) n \<plusminus> (cmp f h (Suc n)) = nsum A f n \<plusminus> (f (Suc n))")
apply (simp add:cmp_def)
apply (thin_tac "\<forall>f h. (f \<in> {j. j \<le> Suc n} \<rightarrow> carrier A) \<and>
(h \<in> {j. j\<le>Suc n} \<rightarrow> {j. j\<le>Suc n}) \<and> (inj_on h {j. j\<le>Suc n}) \<longrightarrow>
\<Sigma>\<^sub>e A (cmp f h) (Suc n) = \<Sigma>\<^sub>e A f (Suc n)")
apply (frule Nset_injTr0 [of "h" "Suc n"], assumption+, simp)
apply (frule Nset_injTr0 [of "h" "Suc n"], assumption+, simp)
apply (simp add:Pi_def)
done
lemma (in aGroup) additionTr1_1:"\<lbrakk>\<forall>f. \<forall>h. f \<in> {j. j\<le>Suc n} \<rightarrow> carrier A \<and>
h \<in> {j. j\<le>Suc n} \<rightarrow> {j. j\<le>Suc n} \<and> inj_on h {j. j\<le>Suc n} \<longrightarrow>
nsum A (cmp f h) (Suc n) = nsum A f (Suc n);
f \<in> {j. j\<le>Suc (Suc n)} \<rightarrow> carrier A; i \<le> n\<rbrakk> \<Longrightarrow>
nsum A (cmp f (transpos i (Suc n))) (Suc (Suc n)) = nsum A f (Suc (Suc n))"
apply (rule additionTr1 [of "n" "f" "transpos i (Suc n)"], assumption+)
apply (rule transpos_hom [of "i" "Suc (Suc n)" "Suc n"])
apply simp+
apply (rule transpos_inj [of "i" "Suc (Suc n)" "Suc n"])
apply simp+
apply (subst transpos_id[of i "Suc (Suc n)" "Suc n" "Suc (Suc n)"])
apply simp+
done
lemma (in aGroup) additionTr1_2:"\<lbrakk>\<forall>f. \<forall>h. f \<in> {j. j\<le>Suc n} \<rightarrow> carrier A \<and>
h \<in> {j. j\<le>Suc n} \<rightarrow> {j. j\<le>Suc n} \<and>
inj_on h {j. j\<le>Suc n} \<longrightarrow>
nsum A (cmp f h) (Suc n) = nsum A f (Suc n);
f \<in> {j. j\<le> Suc (Suc n)} \<rightarrow> carrier A; i \<le> (Suc n)\<rbrakk> \<Longrightarrow>
nsum A (cmp f (transpos i (Suc (Suc n)))) (Suc (Suc n)) =
nsum A f (Suc (Suc n))"
apply (case_tac "i = Suc n")
apply (simp del:nsum_suc)
apply (cut_tac additionTr0 [of "Suc (Suc n)" "f"], simp, simp,
rule allI, rule impI, rule funcset_mem[of f "{j. j \<le> Suc (Suc n)}"
"carrier A"], (simp del:nsum_suc)+)
apply (subst nsum_cmp_eq_transpos [THEN sym, of "Suc n" f i],
rule allI, rule impI, rule funcset_mem[of f "{j. j \<le> Suc (Suc n)}"
"carrier A"], assumption+,
simp, assumption+)
apply (subst nsum_cmp_eq [of "f" "Suc (Suc n)"
"cmp (transpos (Suc n) (Suc(Suc n))) (transpos i (Suc n))"
"transpos i (Suc n)" "Suc (Suc n)"], assumption+,
rule Pi_I, simp add:cmp_def,
rule transpos_mem, (simp del:nsum_suc)+,
rule transpos_mem, (simp del:nsum_suc)+,
rule Pi_I, simp,
rule transpos_mem, (simp del:nsum_suc)+)
apply (subst nsum_cmp_eq [of "cmp f (transpos i (Suc n))" "Suc (Suc n)"
"(transpos i (Suc n))" "transpos (Suc n) (Suc (Suc n))" "Suc (Suc n)"],
rule Pi_I, simp add:cmp_def,
rule funcset_mem[of f "{j. j \<le> Suc (Suc n)}" "carrier A"], assumption,
simp,
rule transpos_mem, (simp del:nsum_suc)+,
(rule Pi_I, simp,
rule transpos_mem, (simp del:nsum_suc)+)+)
apply (subst additionTr1_1 [of "n" "cmp (cmp f (transpos i (Suc n)))
(transpos (Suc n) (Suc (Suc n)))" "i"], assumption+,
rule cmp_fun [of _ "{j. j \<le> (Suc (Suc n))}"
"{j. j \<le> (Suc (Suc n))}" _ "carrier A"],
rule transpos_hom, simp, simp, simp,
rule cmp_fun [of _ "{j. j \<le> (Suc (Suc n))}"
"{j. j \<le> (Suc (Suc n))}" "f" "carrier A"],
rule transpos_hom, simp, simp, assumption+, arith)
apply (cut_tac additionTr0 [of "Suc (Suc n)" "cmp f (transpos i (Suc n))"],
simp del:nsum_suc,
thin_tac "nsum A (cmp
(cmp f (transpos i (Suc n))) (transpos (Suc n) (Suc (Suc n)))) (Suc (Suc n))
= nsum A (cmp f (transpos i (Suc n))) (Suc (Suc n))")
apply (rule additionTr1_1, assumption+, arith, simp,
rule allI, rule impI, simp add:cmp_def,
rule funcset_mem[of f "{j. j \<le> Suc (Suc n)}" "carrier A"],
assumption)
apply (simp add:transpos_mem)
done
lemma (in aGroup) additionTr2:" \<forall>f. \<forall>h. f \<in> {j. j \<le> (Suc n)} \<rightarrow> carrier A \<and>
h \<in> {j. j \<le> (Suc n)} \<rightarrow> {j. j \<le> (Suc n)} \<and>
inj_on h {j. j \<le> (Suc n)} \<longrightarrow>
nsum A (cmp f h) (Suc n) = nsum A f (Suc n)"
apply (induct_tac n)
apply (rule allI)+
apply (rule impI, (erule conjE)+)
apply (simp add:cmp_def)
apply (case_tac "h 0 = 0")
apply (simp add:Nset_1)
apply (simp add:Nset_1 ag_pOp_commute)
(************* n *****************)
apply (rule allI)+
apply (rule impI, (erule conjE)+)
apply (case_tac "h (Suc (Suc n)) = Suc (Suc n)")
apply (rule additionTr1, assumption+)
apply (frule_tac f = h and n = "Suc (Suc n)" in inj_surj, assumption+)
apply (frule sym, thin_tac "h ` {i. i \<le> Suc (Suc n)} = {i. i \<le> Suc (Suc n)}")
apply (cut_tac n = "Suc (Suc n)" in n_in_Nsetn)
apply (frule_tac a = "Suc (Suc n)" and A = "{i. i \<le> Suc (Suc n)}" and
B = "h ` {i. i \<le> Suc (Suc n)}" in eq_set_inc, assumption+)
apply (thin_tac "{i. i \<le> Suc (Suc n)} = h ` {i. i \<le> Suc (Suc n)}")
apply (simp del:nsum_suc add:image_def)
apply (erule exE, erule conjE)
apply (frule sym, thin_tac "Suc (Suc n) = h x")
apply (frule_tac i = x and n = "Suc (Suc n)" and j = "Suc (Suc n)" in
transpos_ij_2, simp del:nsum_suc add:n_in_Nsetn)
apply (rule contrapos_pp, (simp del:nsum_suc)+)
apply (frule_tac x = "transpos x (Suc (Suc n)) (Suc (Suc n))" and y = x and
f = h in eq_elems_eq_val,
thin_tac "transpos x (Suc (Suc n)) (Suc (Suc n)) = x",
simp del:nsum_suc)
apply (frule_tac f = h and A = "{i. i \<le> Suc (Suc n)}" and x = x and
y = "Suc (Suc n)" in inj_onTr2, simp, simp,
frule not_sym, simp)
apply (cut_tac f1 = "cmp f h" and n1 = n and i1 = x in
additionTr1_2[THEN sym], assumption)
apply (rule cmp_fun, simp, assumption, arith)
apply (simp del:nsum_suc,
thin_tac "\<Sigma>\<^sub>e A (cmp f h) (Suc (Suc n)) =
\<Sigma>\<^sub>e A (cmp (cmp f h) (transpos x (Suc (Suc n)))) (Suc (Suc n))")
apply (frule_tac f = f and n = "Suc n" and A = "carrier A" in func_pre)
apply (cut_tac f = "cmp h (transpos x (Suc (Suc n)))" and A = "{j. j \<le> (Suc ( Suc n))}" and ?A1.0 = "{j. j \<le> (Suc n)}" in restrict_inj)
apply (rule_tac f = "transpos x (Suc (Suc n))" and A = "{j. j \<le> Suc (Suc n)}"
and B = "{j. j \<le> Suc (Suc n)}" and g = h and C = "{j. j \<le> Suc (Suc n)}" in
cmp_inj, simp,
rule transpos_hom, simp, simp, assumption+,
rule transpos_inj, simp, simp, assumption+,
rule subsetI, simp)
apply (subst nsum_cmp_assoc,
rule allI, rule impI, simp add:Pi_def,
rule transpos_hom, assumption, simp, assumption+)
apply (cut_tac f = "cmp f (cmp h (transpos x (Suc (Suc n))))" and n = "Suc n"
in nsum_suc[of A ], simp del:nsum_suc,
thin_tac "\<Sigma>\<^sub>e A (cmp f (cmp h (transpos x (Suc (Suc n))))) (Suc (Suc n)) =
\<Sigma>\<^sub>e A (cmp f (cmp h (transpos x (Suc (Suc n))))) (Suc n) \<plusminus>
cmp f (cmp h (transpos x (Suc (Suc n)))) (Suc (Suc n))")
apply (frule_tac x = f in spec,
thin_tac "\<forall>f h. f \<in> {j. j \<le> Suc n} \<rightarrow> carrier A \<and>
h \<in> {j. j \<le> Suc n} \<rightarrow> {j. j \<le> Suc n} \<and>
inj_on h {j. j \<le> Suc n} \<longrightarrow>
\<Sigma>\<^sub>e A (cmp f h) (Suc n) = \<Sigma>\<^sub>e A f (Suc n)")
apply (frule_tac a = "cmp h (transpos x (Suc (Suc n)))" in forall_spec,
thin_tac "\<forall>h. f \<in> {j. j \<le> Suc n} \<rightarrow> carrier A \<and>
h \<in> {j. j \<le> Suc n} \<rightarrow> {j. j \<le> Suc n} \<and> inj_on h {j. j \<le> Suc n} \<longrightarrow>
\<Sigma>\<^sub>e A (cmp f h) (Suc n) = \<Sigma>\<^sub>e A f (Suc n)")
apply simp
apply (rule Pi_I)
apply (simp add:cmp_def)
apply (case_tac "xa = x", simp)
apply (cut_tac i = x and n = "Suc (Suc n)" and j = "Suc (Suc n)" in
transpos_ij_1, simp, simp, simp, simp,
frule_tac x = "Suc (Suc n)" and f = h and A = "{j. j \<le> Suc (Suc n)}"
and B = "{j. j \<le> Suc (Suc n)}" in funcset_mem, simp,
thin_tac "h \<in> {j. j \<le> Suc (Suc n)} \<rightarrow> {j. j \<le> Suc (Suc n)}")
apply (cut_tac m = "h (Suc (Suc n))" and n = "Suc (Suc n)" in noteq_le_less,
simp, simp,
rule_tac x = "h (Suc (Suc n))" and n = "Suc n" in Suc_less_le,
assumption)
apply (subst transpos_id, simp, simp, simp, simp,
frule_tac x = xa and f = h and A = "{j. j \<le> Suc (Suc n)}" and
B = "{j. j \<le> Suc (Suc n)}" in funcset_mem, simp)
apply (frule_tac f = h and A = "{j. j \<le> Suc (Suc n)}" and x = xa and y = x
in injective, simp, simp, assumption)
apply (cut_tac m = "h xa" and n = "Suc (Suc n)" in noteq_le_less, simp,
simp)
apply (rule Suc_less_le, assumption,
thin_tac "\<forall>h. f \<in> {j. j \<le> Suc n} \<rightarrow> carrier A \<and>
h \<in> {j. j \<le> Suc n} \<rightarrow> {j. j \<le> Suc n} \<and> inj_on h {j. j \<le> Suc n} \<longrightarrow>
\<Sigma>\<^sub>e A (cmp f h) (Suc n) = \<Sigma>\<^sub>e A f (Suc n)")
apply (simp del:nsum_suc add:cmp_def)
apply simp
done
lemma (in aGroup) addition2:"\<lbrakk>f \<in> {j. j \<le> (Suc n)} \<rightarrow> carrier A;
h \<in> {j. j \<le> (Suc n)} \<rightarrow> {j. j \<le> (Suc n)}; inj_on h {j. j \<le> (Suc n)}\<rbrakk> \<Longrightarrow>
nsum A (cmp f h) (Suc n) = nsum A f (Suc n)"
apply (simp del:nsum_suc add:additionTr2)
done
lemma (in aGroup) addition21:"\<lbrakk>f \<in> {j. j \<le> n} \<rightarrow> carrier A;
h \<in> {j. j \<le> n} \<rightarrow> {j. j \<le> n}; inj_on h {j. j \<le> n}\<rbrakk> \<Longrightarrow>
nsum A (cmp f h) n = nsum A f n"
apply (case_tac "n = 0")
apply (simp add: cmp_def)
apply (cut_tac f = f and n = "n - Suc 0" and h = h in addition2)
apply simp+
done
lemma (in aGroup) addition3:"\<lbrakk>\<forall>j \<le> (Suc n). f j \<in> carrier A; j \<le> (Suc n);
j \<noteq> Suc n\<rbrakk> \<Longrightarrow> nsum A f (Suc n) = nsum A (cmp f (transpos j (Suc n))) (Suc n)"
apply (rule addition2 [THEN sym,of "f" "n" "transpos j (Suc n)"])
apply (simp)
apply (rule transpos_hom, assumption+, simp, assumption)
apply (rule transpos_inj, simp+)
done
lemma (in aGroup) nsum_splitTr:"(\<forall>j \<le> (Suc (n + m)). f j \<in> carrier A) \<longrightarrow>
nsum A f (Suc (n + m)) = nsum A f n \<plusminus> (nsum A (cmp f (slide (Suc n))) m)"
apply (induct_tac m)
apply (rule impI) apply (simp add:slide_def cmp_def)
apply (rule impI, simp del:nsum_suc)
apply (cut_tac n = "Suc (n + na)" in nsum_suc[of A f],
simp del:nsum_suc,
thin_tac "\<Sigma>\<^sub>e A f (Suc (Suc (n + na))) =
\<Sigma>\<^sub>e A f n \<plusminus> \<Sigma>\<^sub>e A (cmp f (slide (Suc n))) na \<plusminus> f (Suc (Suc (n + na)))")
apply (cut_tac f = "cmp f (slide (Suc n))" and n = na in nsum_suc[of A],
simp del:nsum_suc)
apply (subst ag_pOp_assoc)
apply (rule nsum_mem, rule allI, simp)
apply (rule_tac n = na in nsum_mem,
thin_tac "\<Sigma>\<^sub>e A (cmp f (slide (Suc n))) (Suc na) =
\<Sigma>\<^sub>e A (cmp f (slide (Suc n))) na \<plusminus> (cmp f (slide (Suc n)) (Suc na))")
apply (rule allI, rule impI,
simp add:cmp_def slide_def, simp)
apply (simp add:cmp_def slide_def)
done
lemma (in aGroup) nsum_split:"\<forall>j \<le> (Suc (n + m)). f j \<in> carrier A \<Longrightarrow>
nsum A f (Suc (n + m)) = nsum A f n \<plusminus> (nsum A (cmp f (slide (Suc n))) m)"
by (simp del:nsum_suc add:nsum_splitTr)
lemma (in aGroup) nsum_split1:"\<lbrakk>\<forall>j \<le> m. f j \<in> carrier A; n < m\<rbrakk> \<Longrightarrow>
nsum A f m = nsum A f n \<plusminus> (fSum A f (Suc n) m)"
apply (cut_tac nsum_split[of n "m - n - Suc 0" f])
apply simp
apply (simp add:fSum_def)
apply simp
done
lemma (in aGroup) nsum_minusTr:" (\<forall>j \<le> n. f j \<in> carrier A) \<longrightarrow>
-\<^sub>a (nsum A f n) = nsum A (\<lambda>x\<in>{j. j \<le> n}. -\<^sub>a (f x)) n"
apply (induct_tac n)
apply (rule impI, simp)
apply (rule impI)
apply (subst nsum_suc, subst nsum_suc)
apply (subst ag_p_inv)
apply (rule_tac n = n in nsum_mem [of _ f],
rule allI, simp, simp)
apply (subgoal_tac "\<forall>j\<le>n. f j \<in> carrier A", simp)
apply (rule_tac a = "\<Sigma>\<^sub>e A (\<lambda>u. if u \<le> n then -\<^sub>a (f u) else undefined) n"
and b = "\<Sigma>\<^sub>e A (\<lambda>x\<in>{j. j \<le> (Suc n)}. -\<^sub>a (f x)) n" and c = "-\<^sub>a (f (Suc n))"
in ag_pOp_add_r,
rule_tac n = n in nsum_mem,
rule allI, rule impI, simp,
rule ag_mOp_closed, simp)
apply (rule_tac n = n in nsum_mem,
rule allI, rule impI, simp,
rule ag_mOp_closed, simp,
rule ag_mOp_closed, simp)
apply (rule_tac f = "\<lambda>u. if u \<le> n then -\<^sub>a (f u) else undefined" and
n = n and g = "\<lambda>x\<in>{j. j \<le> (Suc n)}. -\<^sub>a (f x)" in nsum_eq,
rule allI, rule impI,
simp, rule ag_mOp_closed, simp,
rule allI, simp, rule impI, rule ag_mOp_closed, simp)
apply (rule allI, simp)
apply (rule allI, simp)
done
lemma (in aGroup) nsum_minus:"\<forall>j \<le> n. f j \<in> carrier A \<Longrightarrow>
-\<^sub>a (nsum A f n) = nsum A (\<lambda>x\<in>{j. j \<le> n}. -\<^sub>a (f x)) n"
apply (simp add:nsum_minusTr)
done
lemma (in aGroup) ring_nsum_zeroTr:"(\<forall>j \<le> (n::nat). f j \<in> carrier A) \<and>
(\<forall>j \<le> n. f j = \<zero>) \<longrightarrow> nsum A f n = \<zero>"
apply (induct_tac n)
apply (rule impI) apply (erule conjE)+ apply simp
apply (rule impI, (erule conjE)+)
apply (cut_tac n = n in Nsetn_sub_mem1, simp)
apply (simp add:ag_inc_zero)
apply (cut_tac ag_inc_zero,
simp add:ag_r_zero)
done
lemma (in aGroup) ring_nsum_zero:"\<forall>j \<le> (n::nat). f j = \<zero> \<Longrightarrow> \<Sigma>\<^sub>e A f n = \<zero>"
apply (cut_tac ring_nsum_zeroTr[of n f])
apply (simp add:ag_inc_zero)
done
lemma (in aGroup) ag_nsum_1_nonzeroTr:
"\<forall>f. (\<forall>j \<le> n. f j \<in> carrier A) \<and>
(l \<le> n \<and> (\<forall>j \<in> {j. j \<le> n} - {l}. f j = \<zero>))
\<longrightarrow> nsum A f n = f l"
apply (induct_tac n)
apply simp
apply (rule allI,
rule impI, (erule conjE)+)
apply (case_tac "l = Suc n")
apply simp
apply (subgoal_tac "{j. j \<le> Suc n} - {Suc n} = {j. j \<le> n}", simp,
frule ring_nsum_zero, simp)
apply (rule ag_l_zero, simp)
apply (rule equalityI, rule subsetI, simp,
rule subsetI, simp)
apply (frule_tac m = l and n = "Suc n" in noteq_le_less, assumption,
thin_tac "l \<le> Suc n",
frule_tac x = l and n = n in Suc_less_le)
apply (cut_tac n = n in Nsetn_sub_mem1, simp)
apply (thin_tac "\<forall>f. (\<forall>j\<le>n. f j \<in> carrier A) \<and>
(\<forall>j\<in>{j. j \<le> n} - {l}. f j = \<zero>) \<longrightarrow>
\<Sigma>\<^sub>e A f n = f l",
frule_tac a = l in forall_spec, simp)
apply (simp add:ag_r_zero)
done
lemma (in aGroup) ag_nsum_1_nonzero:"\<lbrakk>\<forall>j \<le> n. f j \<in> carrier A; l \<le> n;
\<forall>j\<in>({j. j \<le> n} - {l}). f j = \<zero> \<rbrakk> \<Longrightarrow> nsum A f n = f l"
apply (simp add:ag_nsum_1_nonzeroTr[of n l])
done
definition
set_mult :: "[_ , 'a set, 'a set] \<Rightarrow> 'a set" where
"set_mult R A B = {z. \<exists>x\<in>A. \<exists>y\<in>B. x \<cdot>\<^sub>r\<^bsub>R\<^esub> y = z}"
definition
sum_mult :: "[_ , 'a set, 'a set] \<Rightarrow> 'a set" where
"sum_mult R A B = {x. \<exists>n. \<exists>f \<in> {j. j \<le> (n::nat)}
\<rightarrow> set_mult R A B. nsum R f n = x}"
(*
zero_fini::"[_ , 'a set, 'a set] \<Rightarrow> (nat \<Rightarrow> 'a)"
"zero_fini R A B i == \<zero> \<cdot>\<^sub>r \<zero>" *)
lemma (in Ring) set_mult_sub:"\<lbrakk>A \<subseteq> carrier R; B \<subseteq> carrier R\<rbrakk> \<Longrightarrow>
set_mult R A B \<subseteq> carrier R"
apply (rule subsetI, simp add:set_mult_def, (erule bexE)+,
frule sym, thin_tac "xa \<cdot>\<^sub>r y = x", simp)
apply (rule ring_tOp_closed, (simp add:subsetD)+)
done
lemma (in Ring) set_mult_mono:"\<lbrakk>A1 \<subseteq> carrier R; A2 \<subseteq> carrier R; A1 \<subseteq> A2;
B \<subseteq> carrier R\<rbrakk> \<Longrightarrow> set_mult R A1 B \<subseteq> set_mult R A2 B"
apply (rule subsetI)
apply (simp add:set_mult_def, (erule bexE)+)
apply (frule_tac c = xa in subsetD[of A1 A2], assumption+)
apply blast
done
lemma (in Ring) sum_mult_Tr1:"\<lbrakk>A \<subseteq> carrier R; B \<subseteq> carrier R\<rbrakk> \<Longrightarrow>
(\<forall>j \<le> n. f j \<in> set_mult R A B) \<longrightarrow> nsum R f n \<in> carrier R"
apply (cut_tac ring_is_ag)
apply (induct_tac n)
apply (rule impI, simp)
apply (frule set_mult_sub[of A B], assumption, simp add:subsetD)
apply (rule impI)
apply (cut_tac n = n in Nsetn_sub_mem1, simp)
apply (frule set_mult_sub[of A B], assumption)
apply (frule_tac a = "Suc n" in forall_spec, simp,
frule_tac c = "f (Suc n)" in subsetD[of "set_mult R A B" "carrier R"],
assumption)
apply (rule aGroup.ag_pOp_closed, assumption+)
done
lemma (in Ring) sum_mult_mem:"\<lbrakk>A \<subseteq> carrier R; B \<subseteq> carrier R;
\<forall>j \<le> n. f j \<in> set_mult R A B\<rbrakk> \<Longrightarrow> nsum R f n \<in> carrier R"
apply (cut_tac ring_is_ag)
apply (simp add:sum_mult_Tr1)
done
lemma (in Ring) sum_mult_mem1:"\<lbrakk>A \<subseteq> carrier R; B \<subseteq> carrier R;
x \<in> sum_mult R A B\<rbrakk> \<Longrightarrow>
\<exists>n. \<exists>f\<in>{j. j \<le> (n::nat)} \<rightarrow> set_mult R A B. nsum R f n = x"
by (simp add:sum_mult_def)
lemma (in Ring) sum_mult_subR:"\<lbrakk>A \<subseteq> carrier R; B \<subseteq> carrier R\<rbrakk> \<Longrightarrow>
sum_mult R A B \<subseteq> carrier R"
apply (rule subsetI)
apply (frule_tac x = x in sum_mult_mem1[of A B], assumption+)
apply (erule exE, erule bexE, frule sym, thin_tac "\<Sigma>\<^sub>e R f n = x", simp)
apply (cut_tac ring_is_ag)
apply (rule aGroup.nsum_mem[of R], assumption)
apply (rule allI, rule impI)
apply (frule_tac f = f and A = "{j. j \<le> n}" and B = "set_mult R A B" and
x = j in funcset_mem, simp)
apply (frule set_mult_sub[of A B], assumption)
apply (simp add:subsetD)
done
lemma (in Ring) times_mem_sum_mult:"\<lbrakk>A \<subseteq> carrier R; B \<subseteq> carrier R;
a \<in> A; b \<in> B \<rbrakk> \<Longrightarrow> a \<cdot>\<^sub>r b \<in> sum_mult R A B"
apply (simp add:sum_mult_def)
apply (subgoal_tac "(\<lambda>i\<in>{j. j \<le> (0::nat)}. a \<cdot>\<^sub>r b) \<in> {j. j \<le> 0} \<rightarrow> set_mult R A B")
apply (subgoal_tac "nsum R (\<lambda>i\<in>{j. j \<le> (0::nat)}. a \<cdot>\<^sub>r b) 0 = a \<cdot>\<^sub>r b")
apply blast
apply simp
apply (rule Pi_I, simp add:set_mult_def, blast)
done
lemma (in Ring) mem_minus_sum_multTr2:"\<lbrakk>A \<subseteq> carrier R; B \<subseteq> carrier R;
\<forall>j \<le> n. f j \<in> set_mult R A B; i \<le> n\<rbrakk> \<Longrightarrow> f i \<in> carrier R"
apply (frule_tac a = i in forall_spec, simp)
apply (frule set_mult_sub[of A B], assumption, simp add:subsetD)
done
lemma (in aGroup) nsum_jointfun:"\<lbrakk>\<forall>j \<le> n. f j \<in> carrier A;
\<forall>j \<le> m. g j \<in> carrier A\<rbrakk> \<Longrightarrow>
\<Sigma>\<^sub>e A (jointfun n f m g) (Suc (n + m)) = \<Sigma>\<^sub>e A f n \<plusminus> (\<Sigma>\<^sub>e A g m)"
apply (subst nsum_split)
apply (rule allI, rule impI)
apply (frule_tac f = f and n = n and A = "carrier A" and g = g and m = m
and B = "carrier A" in jointfun_mem, assumption+, simp)
apply (subgoal_tac "nsum A (jointfun n f m g) n = nsum A f n")
apply (subgoal_tac "nsum A (cmp (jointfun n f m g) (slide (Suc n))) m =
nsum A g m")
apply simp
apply (thin_tac "nsum A (jointfun n f m g) n = nsum A f n")
apply (rule nsum_eq)
apply (rule allI, rule impI,
simp add:cmp_def jointfun_def slide_def sliden_def,
assumption)
apply (rule allI, simp add:cmp_def jointfun_def slide_def sliden_def)
apply (rule nsum_eq)
apply (rule allI, rule impI,
simp add:jointfun_def, assumption)
apply (rule allI, rule impI)
apply (simp add:jointfun_def)
done
lemma (in Ring) sum_mult_pOp_closed:"\<lbrakk>A \<subseteq> carrier R; B \<subseteq> carrier R;
a \<in> sum_mult R A B; b \<in> sum_mult R A B \<rbrakk> \<Longrightarrow> a \<plusminus>\<^bsub>R\<^esub> b \<in> sum_mult R A B"
apply (cut_tac ring_is_ag)
apply (simp add:sum_mult_def)
apply ((erule exE)+, (erule bexE)+)
apply (rename_tac n m f g)
apply (frule sym, thin_tac "\<Sigma>\<^sub>e R f n = a", frule sym,
thin_tac "\<Sigma>\<^sub>e R g m = b", simp)
apply (frule set_mult_sub[of A B], assumption)
apply (subst aGroup.nsum_jointfun[THEN sym, of R], assumption)
apply (rule allI, rule impI,
frule_tac f = f and A = "{j. j \<le> n}" and B = "set_mult R A B" and
x = j in funcset_mem, simp, simp add:subsetD)
apply (rule allI, rule impI,
frule_tac f = g and A = "{j. j \<le> m}" and B = "set_mult R A B" and
x = j in funcset_mem, simp, simp add:subsetD)
apply (frule_tac f = f and n = n and A = "set_mult R A B" and g = g and m = m
and B = "set_mult R A B" in jointfun_hom, assumption+)
apply (simp del:nsum_suc)
apply blast
done
lemma (in Ring) set_mult_mOp_closed:"\<lbrakk>A \<subseteq> carrier R; ideal R B;
x \<in> set_mult R A B\<rbrakk> \<Longrightarrow> -\<^sub>a x \<in> set_mult R A B"
apply (cut_tac ring_is_ag,
simp add:set_mult_def,
(erule bexE)+, frule sym, thin_tac "xa \<cdot>\<^sub>r y = x", simp,
frule_tac c = xa in subsetD[of A "carrier R"], assumption+,
frule ideal_subset1[of B],
frule_tac c = y in subsetD[of B "carrier R"], assumption+,
simp add:ring_inv1_2,
frule_tac I = B and x = y in ideal_inv1_closed,
assumption+)
apply blast
done
lemma (in Ring) set_mult_ring_times_closed:"\<lbrakk>A \<subseteq> carrier R; ideal R B;
x \<in> set_mult R A B; r \<in> carrier R\<rbrakk> \<Longrightarrow> r \<cdot>\<^sub>r x \<in> set_mult R A B"
apply (cut_tac ring_is_ag,
simp add:set_mult_def,
(erule bexE)+, frule sym, thin_tac "xa \<cdot>\<^sub>r y = x", simp,
frule_tac c = xa in subsetD[of A "carrier R"], assumption+,
frule ideal_subset1[of B],
frule_tac c = y in subsetD[of B "carrier R"], assumption,
frule_tac x = r and y = "xa \<cdot>\<^sub>r y" in ring_tOp_commute,
simp add:ring_tOp_closed, simp,
subst ring_tOp_assoc, assumption+)
apply (frule_tac x = y and y = r in ring_tOp_commute, assumption+,
simp,
frule_tac x = y and r = r in ideal_ring_multiple [of B], assumption+)
apply blast
done
lemma (in Ring) set_mult_sub_sum_mult:"\<lbrakk>A \<subseteq> carrier R; ideal R B\<rbrakk> \<Longrightarrow>
set_mult R A B \<subseteq> sum_mult R A B"
apply (rule subsetI)
apply (simp add:sum_mult_def)
apply (cut_tac f = "(\<lambda>i\<in>{j. j \<le> (0::nat)}. x)" in nsum_0[of R])
apply (cut_tac n_in_Nsetn[of 0],
simp del:nsum_0)
apply (cut_tac f = "\<lambda>i\<in>{j. j \<le> (0::nat)}. x" and B = "%_. set_mult R A B" in
Pi_I[of "{j. j \<le> 0}"],
simp)
apply (subgoal_tac "\<Sigma>\<^sub>e R (\<lambda>i\<in>{j. j \<le> 0}. x) 0 = x")
apply blast
apply simp
done
lemma (in Ring) sum_mult_pOp_closedn:"\<lbrakk>A \<subseteq> carrier R; ideal R B\<rbrakk> \<Longrightarrow>
(\<forall>j \<le> n. f j \<in> set_mult R A B) \<longrightarrow> \<Sigma>\<^sub>e R f n \<in> sum_mult R A B"
apply (induct_tac n)
apply (rule impI, simp)
apply (frule set_mult_sub_sum_mult[of A B], assumption+, simp add:subsetD)
apply (rule impI)
apply simp
apply (frule_tac a = "Suc n" in forall_spec, simp)
apply (frule set_mult_sub_sum_mult[of A B], assumption+,
frule_tac c= "f (Suc n)" in
subsetD[of "set_mult R A B" "sum_mult R A B"], assumption+)
apply (rule sum_mult_pOp_closed, assumption,
simp add:ideal_subset1, assumption+)
done
lemma (in Ring) mem_minus_sum_multTr4:"\<lbrakk>A \<subseteq> carrier R; ideal R B\<rbrakk> \<Longrightarrow>
(\<forall>j \<le> n. f j \<in> set_mult R A B) \<longrightarrow> -\<^sub>a (nsum R f n) \<in> sum_mult R A B"
apply (cut_tac ring_is_ag)
apply (induct_tac n)
apply (rule impI)
apply (cut_tac n_in_Nsetn[of 0])
apply (frule_tac x = "f 0" in set_mult_mOp_closed[of A B], assumption+)
apply simp
apply (frule set_mult_sub_sum_mult[of A B], assumption+,
simp add:subsetD)
apply (rule impI)
apply (cut_tac n = n in Nsetn_sub_mem1, simp)
apply (frule sum_mult_subR[of A B], simp add:ideal_subset1)
apply (frule_tac n = n and f = f in sum_mult_pOp_closedn[of A B], assumption,
cut_tac n = n in Nsetn_sub_mem1, simp)
apply (frule_tac c = "\<Sigma>\<^sub>e R f n" in subsetD[of "sum_mult R A B" "carrier R"],
assumption+,
frule_tac a = "Suc n" in forall_spec, simp,
thin_tac "\<forall>j\<le>Suc n. f j \<in> set_mult R A B",
frule set_mult_sub[of A B], simp add:ideal_subset1,
frule_tac c = "f (Suc n)" in subsetD[of "set_mult R A B" "carrier R"],
assumption+ )
apply (frule_tac x = "\<Sigma>\<^sub>e R f n" and y = "f (Suc n)" in aGroup.ag_p_inv[of R],
assumption+, simp)
apply (rule_tac a = "-\<^sub>a (\<Sigma>\<^sub>e R f n)" and b = "-\<^sub>a (f (Suc n))" in
sum_mult_pOp_closed[of A B], assumption+,
simp add:ideal_subset1, assumption)
apply (frule_tac x = "f (Suc n)" in set_mult_mOp_closed[of A B], assumption+,
frule set_mult_sub_sum_mult[of A B], assumption+)
apply (simp add:subsetD)
done
lemma (in Ring) sum_mult_iOp_closed:"\<lbrakk>A \<subseteq> carrier R; ideal R B;
x \<in> sum_mult R A B \<rbrakk> \<Longrightarrow> -\<^sub>a x \<in> sum_mult R A B"
apply (frule sum_mult_mem1 [of A B x],
simp add:ideal_subset1, assumption)
apply (erule exE, erule bexE, frule sym, thin_tac "\<Sigma>\<^sub>e R f n = x")
apply simp
apply (frule_tac n = n and f = f in mem_minus_sum_multTr4[of A B],
assumption+)
apply (simp add:Pi_def)
done
lemma (in Ring) sum_mult_ring_multiplicationTr:
"\<lbrakk>A \<subseteq> carrier R; ideal R B; r \<in> carrier R\<rbrakk> \<Longrightarrow>
(\<forall>j \<le> n. f j \<in> set_mult R A B) \<longrightarrow> r \<cdot>\<^sub>r (nsum R f n) \<in> sum_mult R A B"
apply (cut_tac ring_is_ag)
apply (induct_tac n)
apply (rule impI, simp)
apply (simp add:set_mult_def)
apply ((erule bexE)+, frule sym, thin_tac "x \<cdot>\<^sub>r y = f 0", simp)
apply (frule_tac c = x in subsetD[of A "carrier R"], assumption+)
apply (frule ideal_subset1[of B],
frule_tac c = y in subsetD[of B "carrier R"], assumption,
frule_tac x = r and y = "x \<cdot>\<^sub>r y" in ring_tOp_commute,
simp add:ring_tOp_closed, simp,
subst ring_tOp_assoc, assumption+)
apply (frule_tac x = y and y = r in ring_tOp_commute, assumption+,
simp,
frule_tac x = y and r = r in ideal_ring_multiple [of B], assumption+)
apply (rule times_mem_sum_mult, assumption+)
apply (rule impI)
apply (cut_tac n = n in Nsetn_sub_mem1, simp)
apply (frule_tac f = f and n = n in aGroup.nsum_mem,
frule set_mult_sub [of "A" "B"], simp add:ideal_subset1,
rule allI, rule impI, cut_tac n = n in Nsetn_sub_mem1,
simp add: subsetD,
frule_tac a = "Suc n" in forall_spec, simp)
apply (frule set_mult_sub[of A B], simp add:ideal_subset1,
frule_tac c = "f (Suc n)" in subsetD[of "set_mult R A B" "carrier R"],
assumption)
apply (simp add: ring_distrib1)
apply (rule sum_mult_pOp_closed[of A B], assumption+,
simp add:ideal_subset1, assumption)
apply (frule_tac x = "f (Suc n)" in set_mult_ring_times_closed [of A B _ r],
assumption+, simp, assumption,
frule set_mult_sub_sum_mult[of A B], assumption+,
simp add:subsetD)
done
lemma (in Ring) sum_mult_ring_multiplication:"\<lbrakk>A \<subseteq> carrier R; ideal R B;
r \<in> carrier R; a \<in> sum_mult R A B\<rbrakk> \<Longrightarrow> r \<cdot>\<^sub>r a \<in> sum_mult R A B"
apply (cut_tac ring_is_ag)
apply (frule sum_mult_mem1[of A B a],
simp add:ideal_subset1, assumption)
apply (erule exE, erule bexE, frule sym, thin_tac "\<Sigma>\<^sub>e R f n = a", simp)
apply (subgoal_tac "\<forall>j \<le> n. f j \<in> set_mult R A B")
apply (simp add:sum_mult_ring_multiplicationTr)
apply (simp add:Pi_def)
done
lemma (in Ring) ideal_sum_mult:"\<lbrakk>A \<subseteq> carrier R; A \<noteq> {}; ideal R B\<rbrakk> \<Longrightarrow>
ideal R (sum_mult R A B)"
apply (simp add:ideal_def [of _ "sum_mult R A B"])
apply (cut_tac ring_is_ag)
apply (rule conjI)
apply (rule aGroup.asubg_test, assumption+)
apply (rule subsetI)
apply (frule_tac x = x in sum_mult_mem1[of A B],
simp add:ideal_subset1, assumption,
erule exE, erule bexE, frule sym, thin_tac "\<Sigma>\<^sub>e R f n = x", simp)
apply (rule_tac f = f and n = n in sum_mult_mem[of A B _ _], assumption+)
apply (simp add:ideal_subset1)
apply (simp add:Pi_def)
apply (frule nonempty_ex[of A], erule exE)
apply (frule ideal_zero[of B])
apply (frule_tac a = x and b = \<zero> in times_mem_sum_mult[of A B],
simp add:ideal_subset1, assumption+) apply blast
apply (rule ballI)+
apply (rule_tac a = a and b = "-\<^sub>a b" in sum_mult_pOp_closed[of A B],
assumption, simp add:ideal_subset1, assumption+,
rule_tac x = b in sum_mult_iOp_closed[of A B], assumption+)
apply (rule ballI)+
apply (rule sum_mult_ring_multiplication, assumption+)
done
lemma (in Ring) ideal_inc_set_multTr:"\<lbrakk>A \<subseteq> carrier R; ideal R B; ideal R C;
set_mult R A B \<subseteq> C \<rbrakk> \<Longrightarrow>
\<forall>f \<in> {j. j \<le> (n::nat)} \<rightarrow> set_mult R A B. \<Sigma>\<^sub>e R f n \<in> C"
apply (induct_tac n)
apply (simp add:subsetD)
apply (rule ballI)
apply (
frule_tac f = f and A = "{j. j \<le> Suc n}" and x = "Suc n" and
B = "set_mult R A B"in funcset_mem, simp,
frule_tac c = "f (Suc n)" in subsetD[of "set_mult R A B" "C"],
assumption+, simp)
apply (rule ideal_pOp_closed[of C], assumption+,
cut_tac n = n in Nsetn_sub_mem1,
frule_tac x = f in bspec, simp)
apply (simp add:Pi_def, assumption+)
done
lemma (in Ring) ideal_inc_set_mult:"\<lbrakk>A \<subseteq> carrier R; ideal R B; ideal R C;
set_mult R A B \<subseteq> C \<rbrakk> \<Longrightarrow> sum_mult R A B \<subseteq> C"
apply (rule subsetI)
apply (frule_tac x = x in sum_mult_mem1[of A B],
simp add:ideal_subset1, assumption+)
apply (erule exE, erule bexE, frule sym, thin_tac "\<Sigma>\<^sub>e R f n = x", simp,
thin_tac "x = \<Sigma>\<^sub>e R f n", simp add:subsetD)
apply (simp add:ideal_inc_set_multTr)
done
lemma (in Ring) AB_inc_sum_mult:"\<lbrakk>ideal R A; ideal R B\<rbrakk> \<Longrightarrow>
sum_mult R A B \<subseteq> A \<inter> B"
apply (frule ideal_subset1[of A], frule ideal_subset1[of B])
apply (frule ideal_inc_set_mult [of "A" "B" "A"], assumption+)
apply (rule subsetI,
simp add:set_mult_def, (erule bexE)+, frule sym, thin_tac "xa \<cdot>\<^sub>r y = x",
simp,
frule_tac c = xa in subsetD[of A "carrier R"], assumption+,
frule_tac c = y in subsetD[of B "carrier R"], assumption+,
subst ring_tOp_commute, assumption+,
simp add:ideal_ring_multiple)
apply (frule ideal_inc_set_mult [of "A" "B" "B"], assumption+)
apply (rule subsetI,
simp add:set_mult_def, (erule bexE)+, frule sym, thin_tac "xa \<cdot>\<^sub>r y = x",
simp,
frule_tac c = xa in subsetD[of A "carrier R"], assumption+,
simp add:ideal_ring_multiple)
apply simp
done
lemma (in Ring) sum_mult_is_ideal_prod:"\<lbrakk>ideal R A; ideal R B\<rbrakk> \<Longrightarrow>
sum_mult R A B = A \<diamondsuit>\<^sub>r B"
apply (rule equalityI)
apply (frule ideal_prod_ideal [of "A" "B"], assumption+)
apply (rule ideal_inc_set_mult)
apply (simp add:ideal_subset1)+
apply (rule subsetI)
apply (simp add:set_mult_def ideal_prod_def)
apply (auto del:subsetI)
apply (rule subsetI)
apply (simp add:ideal_prod_def)
apply (frule ideal_subset1[of A],
frule ideal_sum_mult[of A B],
frule ideal_zero[of A], blast, assumption)
apply (frule_tac a = "sum_mult R A B" in forall_spec, simp)
apply (rule subsetI, simp,
thin_tac "\<forall>xa. ideal R xa \<and> {x. \<exists>i\<in>A. \<exists>j\<in>B. x = i \<cdot>\<^sub>r j} \<subseteq> xa \<longrightarrow> x \<in> xa",
(erule bexE)+, simp)
apply (rule times_mem_sum_mult, assumption,
simp add:ideal_subset1, assumption+)
done
lemma (in Ring) ideal_prod_assocTr0:"\<lbrakk>ideal R A; ideal R B; ideal R C; y \<in> C;
z \<in> set_mult R A B\<rbrakk> \<Longrightarrow> z \<cdot>\<^sub>r y \<in> sum_mult R A (B \<diamondsuit>\<^sub>r C)"
apply (simp add:set_mult_def, (erule bexE)+,
frule sym, thin_tac "x \<cdot>\<^sub>r ya = z", simp)
apply (frule_tac h = x in ideal_subset[of A], assumption,
frule_tac h = ya in ideal_subset[of B], assumption,
frule_tac h = y in ideal_subset[of C], assumption,
subst ring_tOp_assoc, assumption+)
apply (frule ideal_subset1[of A],
frule ideal_subset1[of B],
frule ideal_subset1[of C],
frule ideal_prod_ideal[of B C], assumption,
frule ideal_subset1[of "B \<diamondsuit>\<^sub>r C"])
apply (rule times_mem_sum_mult[of A "B \<diamondsuit>\<^sub>r C"], assumption+,
subst sum_mult_is_ideal_prod[of B C, THEN sym], assumption+,
rule times_mem_sum_mult[of B C], assumption+)
done
lemma (in Ring) ideal_prod_assocTr1:"\<lbrakk>ideal R A; ideal R B; ideal R C; y \<in> C\<rbrakk>
\<Longrightarrow> \<forall>f \<in> {j. j\<le>(n::nat)} \<rightarrow> set_mult R A B. (\<Sigma>\<^sub>e R f n) \<cdot>\<^sub>r y \<in> A \<diamondsuit>\<^sub>r (B \<diamondsuit>\<^sub>r C)"
apply (cut_tac ring_is_ag)
apply (frule ideal_prod_ideal[of "B" "C"], assumption+,
subst sum_mult_is_ideal_prod[of A "B \<diamondsuit>\<^sub>r C", THEN sym], assumption+)
apply (induct_tac n)
apply simp
apply (simp add:ideal_prod_assocTr0)
apply (rule ballI,
frule_tac x = f in bspec,
thin_tac "\<forall>f\<in>{j. j \<le> n} \<rightarrow> set_mult R A B.
\<Sigma>\<^sub>e R f n \<cdot>\<^sub>r y \<in> sum_mult R A (B \<diamondsuit>\<^sub>r C)",
rule Pi_I, simp,
frule_tac f = f and A = "{j. j \<le> Suc n}" and B = "set_mult R A B" and
x = x in funcset_mem, simp, assumption)
apply simp
apply (frule ideal_subset1[of A], frule ideal_subset1[of B],
frule set_mult_sub[of A B], assumption,
frule_tac f = f and A = "{j. j \<le> (Suc n)}" in extend_fun[of _ _
"set_mult R A B"
"carrier R"], assumption,
subst ring_distrib2,
simp add:ideal_subset)
apply (rule aGroup.nsum_mem, assumption)
apply (simp add:Pi_def)
apply (simp add:funcset_mem del:Pi_I',
frule_tac f = f and A = "{j. j \<le> Suc n}" and B = "set_mult R A B" and
x = "Suc n" in funcset_mem, simp)
apply (frule ideal_subset1[of A],
frule ideal_zero[of A],
frule ideal_sum_mult[of A "B \<diamondsuit>\<^sub>r C"], blast, assumption)
apply (rule ideal_pOp_closed, assumption+)
apply (simp add:ideal_prod_assocTr0)
done
lemma (in Ring) ideal_quotient_idealTr:"\<lbrakk>ideal R A; ideal R B; ideal R C;
x \<in> carrier R;\<forall>c\<in>C. x \<cdot>\<^sub>r c \<in> ideal_quotient R A B\<rbrakk> \<Longrightarrow>
f\<in>{j. j \<le> n} \<rightarrow> set_mult R B C \<longrightarrow> x \<cdot>\<^sub>r (nsum R f n) \<in> A"
apply (frule ideal_subset1 [of "A"],
frule ideal_subset1 [of "B"])
apply (induct_tac n)
apply (rule impI)
apply (cut_tac n_in_Nsetn[of 0])
apply (frule funcset_mem, assumption+)
apply (thin_tac "f \<in> {j. j \<le> 0} \<rightarrow> set_mult R B C")
apply (simp add:set_mult_def)
apply (erule bexE)+
apply (frule sym, thin_tac "xa \<cdot>\<^sub>r y = f 0", simp)
apply (frule_tac h = xa in ideal_subset[of B], assumption,
frule_tac h = y in ideal_subset[of C], assumption)
apply (frule_tac x = xa and y = y in ring_tOp_commute, assumption+,
simp)
apply (subst ring_tOp_assoc[THEN sym], assumption+)
apply (frule_tac x = y in bspec, assumption,
thin_tac "\<forall>c\<in>C. x \<cdot>\<^sub>r c \<in> A \<dagger>\<^sub>R B")
apply (simp add:ideal_quotient_def)
(****** n ********)
apply (rule impI)
apply (frule func_pre) apply simp
apply (cut_tac ring_is_ag)
apply (frule ideal_subset1[of B], frule ideal_subset1[of C],
frule set_mult_sub[of B C], assumption+)
apply (cut_tac n = n in nsum_memr [of _ "f"],
rule allI, rule impI,
frule_tac x = i in funcset_mem, simp, simp add:subsetD)
apply (frule_tac a = n in forall_spec, simp)
apply (thin_tac "\<forall>l\<le>n. \<Sigma>\<^sub>e R f l \<in> carrier R",
frule_tac f = f and A = "{j. j \<le> Suc n}" and B = "set_mult R B C"
and x = "Suc n" in funcset_mem, simp,
frule_tac c = "f (Suc n)" in subsetD[of "set_mult R B C" " carrier R"],
assumption+)
apply (subst ring_distrib1, assumption+)
apply (rule ideal_pOp_closed[of A], assumption+)
apply (simp add: set_mult_def, (erule bexE)+,
fold set_mult_def,
frule sym, thin_tac "xa \<cdot>\<^sub>r y = f (Suc n)", simp)
apply (frule_tac c = xa in subsetD[of B "carrier R"], assumption+,
frule_tac c = y in subsetD[of C "carrier R"], assumption+,
frule_tac x = xa and y = y in ring_tOp_commute, assumption, simp,
subst ring_tOp_assoc[THEN sym], assumption+)
apply (simp add:ideal_quotient_def)
done
lemma (in Ring) ideal_quotient_ideal:"\<lbrakk>ideal R A; ideal R B; ideal R C\<rbrakk> \<Longrightarrow>
A \<dagger>\<^sub>R B \<dagger>\<^sub>R C = A \<dagger>\<^sub>R B \<diamondsuit>\<^sub>r C"
apply (rule equalityI)
apply (rule subsetI)
apply (simp add:ideal_quotient_def [of _ _ "C"])
apply (erule conjE)
apply (simp add:ideal_quotient_def [of _ _ "B \<diamondsuit>\<^sub>r C"])
apply (rule ballI)
apply (simp add:sum_mult_is_ideal_prod [THEN sym])
apply (simp add:sum_mult_def)
apply (erule exE, erule bexE)
apply (rename_tac x c n f)
apply (frule sym) apply simp
apply (simp add:ideal_quotient_idealTr)
apply (rule subsetI)
apply (simp add:sum_mult_is_ideal_prod [THEN sym])
apply (simp add:ideal_quotient_def)
apply (erule conjE)
apply (rule ballI)
apply (rename_tac x c)
apply (frule ideal_subset [of "C"], assumption+)
apply (simp add:ring_tOp_closed)
apply (rule ballI)
apply (rename_tac x v u)
apply (frule ideal_subset [of "B"], assumption+)
apply (subst ring_tOp_assoc, assumption+)
apply (frule ideal_subset1[of B],
frule ideal_subset1[of C],
frule_tac a = u and b = v in times_mem_sum_mult[of B C], assumption+)
apply (frule_tac x = u and y = v in ring_tOp_commute, assumption,
simp)
done
lemma (in Ring) ideal_prod_assocTr:"\<lbrakk>ideal R A; ideal R B; ideal R C\<rbrakk> \<Longrightarrow>
\<forall>f. (f \<in> {j. j \<le> (n::nat)} \<rightarrow> set_mult R (A \<diamondsuit>\<^sub>r B) C \<longrightarrow>
(\<Sigma>\<^sub>e R f n) \<in> A \<diamondsuit>\<^sub>r (B \<diamondsuit>\<^sub>r C))"
apply (subgoal_tac "\<forall>x\<in>(A \<diamondsuit>\<^sub>r B). \<forall>y\<in>C. x \<cdot>\<^sub>r y \<in> A \<diamondsuit>\<^sub>r (B \<diamondsuit>\<^sub>r C)")
apply (induct_tac n)
apply (rule allI) apply (rule impI)
apply (frule_tac f = f and A = "{j. j \<le> 0}" and B = "set_mult R (A \<diamondsuit>\<^sub>r B) C"
and x = 0 in funcset_mem, simp, simp)
apply (simp add:set_mult_def)
apply ((erule bexE)+, frule sym, thin_tac "x \<cdot>\<^sub>r y = f 0", simp)
apply (rule allI, rule impI)
apply (frule func_pre)
apply (frule_tac a = f in forall_spec, simp,
thin_tac "\<forall>f. f \<in> {j. j \<le> n} \<rightarrow> set_mult R (A \<diamondsuit>\<^sub>r B) C \<longrightarrow>
\<Sigma>\<^sub>e R f n \<in> A \<diamondsuit>\<^sub>r (B \<diamondsuit>\<^sub>r C)",
frule ideal_prod_ideal[of "B" "C"], assumption+,
frule ideal_prod_ideal[of "A" "B \<diamondsuit>\<^sub>r C"], assumption+, simp)
apply (rule ideal_pOp_closed[of "A \<diamondsuit>\<^sub>r (B \<diamondsuit>\<^sub>r C)"], assumption+)
apply (cut_tac n = "Suc n" in n_in_Nsetn,
frule_tac f = f and A = "{j. j \<le> Suc n}" and
B = "set_mult R (A \<diamondsuit>\<^sub>r B) C" and x = "Suc n" in funcset_mem, assumption)
apply (thin_tac "f \<in> {j. j \<le> n} \<rightarrow> set_mult R (A \<diamondsuit>\<^sub>r B) C",
thin_tac "f \<in> {j. j \<le> Suc n} \<rightarrow> set_mult R (A \<diamondsuit>\<^sub>r B) C")
apply (simp add:set_mult_def)
apply ((erule bexE)+,
frule sym, thin_tac "x \<cdot>\<^sub>r y = f (Suc n)", simp)
apply (rule ballI)+
apply (simp add:sum_mult_is_ideal_prod[of A B, THEN sym])
apply (frule ideal_subset1[of A], frule ideal_subset1[of B],
frule_tac x = x in sum_mult_mem1[of A B], assumption+)
apply (erule exE, erule bexE, frule sym, thin_tac "\<Sigma>\<^sub>e R f n = x",
simp)
apply (simp add:ideal_prod_assocTr1)
done
lemma (in Ring) ideal_prod_assoc:"\<lbrakk>ideal R A; ideal R B; ideal R C\<rbrakk> \<Longrightarrow>
(A \<diamondsuit>\<^sub>r B) \<diamondsuit>\<^sub>r C = A \<diamondsuit>\<^sub>r (B \<diamondsuit>\<^sub>r C)"
apply (rule equalityI)
apply (rule subsetI)
apply (frule ideal_prod_ideal[of "A" "B"], assumption+)
apply (frule sum_mult_is_ideal_prod[of "A \<diamondsuit>\<^sub>r B" "C"], assumption+)
apply (frule sym) apply (thin_tac "sum_mult R (A \<diamondsuit>\<^sub>r B) C = (A \<diamondsuit>\<^sub>r B) \<diamondsuit>\<^sub>r C")
apply simp apply (thin_tac "(A \<diamondsuit>\<^sub>r B) \<diamondsuit>\<^sub>r C = sum_mult R (A \<diamondsuit>\<^sub>r B) C")
apply (thin_tac "ideal R (A \<diamondsuit>\<^sub>r B)")
apply (frule ideal_prod_ideal[of "B" "C"], assumption+)
apply (simp add:sum_mult_def)
apply (erule exE, erule bexE)
apply (frule sym, thin_tac "\<Sigma>\<^sub>e R f n = x", simp)
apply (simp add:ideal_prod_assocTr)
apply (rule subsetI)
apply (frule ideal_prod_ideal[of "B" "C"], assumption+)
apply (simp add:ideal_prod_commute [of "A" "B \<diamondsuit>\<^sub>r C"])
apply (frule ideal_prod_ideal[of "A" "B"], assumption+)
apply (simp add:ideal_prod_commute[of "A \<diamondsuit>\<^sub>r B" "C"])
apply (simp add:ideal_prod_commute[of "A" "B"])
apply (simp add:ideal_prod_commute[of "B" "C"])
apply (frule ideal_prod_ideal[of "C" "B"], assumption+)
apply (frule sum_mult_is_ideal_prod[of "C \<diamondsuit>\<^sub>r B" "A"], assumption+)
apply (frule sym) apply (thin_tac "sum_mult R (C \<diamondsuit>\<^sub>r B) A = (C \<diamondsuit>\<^sub>r B) \<diamondsuit>\<^sub>r A")
apply simp apply (thin_tac "(C \<diamondsuit>\<^sub>r B) \<diamondsuit>\<^sub>r A = sum_mult R (C \<diamondsuit>\<^sub>r B) A")
apply (thin_tac "ideal R (C \<diamondsuit>\<^sub>r B)")
apply (frule ideal_prod_ideal[of "B" "A"], assumption+)
apply (simp add:sum_mult_def)
apply (erule exE, erule bexE)
apply (frule sym, thin_tac "\<Sigma>\<^sub>e R f n = x", simp)
apply (simp add:ideal_prod_assocTr)
done
lemma (in Ring) prod_principal_idealTr0:" \<lbrakk>a \<in> carrier R; b \<in> carrier R;
z \<in> set_mult R (R \<diamondsuit>\<^sub>p a) (R \<diamondsuit>\<^sub>p b)\<rbrakk> \<Longrightarrow> z \<in> R \<diamondsuit>\<^sub>p (a \<cdot>\<^sub>r b)"
apply (simp add:set_mult_def, (erule bexE)+,
simp add:Rxa_def, (erule bexE)+, simp)
apply (frule_tac x = r and y = a and z = "ra \<cdot>\<^sub>r b" in ring_tOp_assoc,
assumption+, simp add:ring_tOp_closed, simp)
apply (simp add:ring_tOp_assoc[THEN sym, of a _ b])
apply (frule_tac x = a and y = ra in ring_tOp_commute, assumption+, simp)
apply (simp add:ring_tOp_assoc[of _ a b],
frule_tac x = a and y = b in ring_tOp_closed, assumption)
apply (simp add:ring_tOp_assoc[THEN sym, of _ _ "a \<cdot>\<^sub>r b"],
frule sym, thin_tac "r \<cdot>\<^sub>r ra \<cdot>\<^sub>r (a \<cdot>\<^sub>r b) = z", simp,
frule_tac x = r and y = ra in ring_tOp_closed, assumption+)
apply blast
done
lemma (in Ring) prod_principal_idealTr1:" \<lbrakk>a \<in> carrier R; b \<in> carrier R\<rbrakk> \<Longrightarrow>
\<forall>f \<in> {j. j \<le> (n::nat)} \<rightarrow> set_mult R (R \<diamondsuit>\<^sub>p a) (R \<diamondsuit>\<^sub>p b).
\<Sigma>\<^sub>e R f n \<in> R \<diamondsuit>\<^sub>p (a \<cdot>\<^sub>r b)"
apply (induct_tac n)
apply (rule ballI,
frule_tac f = f in funcset_mem[of _ "{j. j \<le> 0}"
"set_mult R (R \<diamondsuit>\<^sub>p a) (R \<diamondsuit>\<^sub>p b)"], simp)
apply (simp add:prod_principal_idealTr0)
apply (rule ballI,
frule func_pre,
frule_tac x = f in bspec, assumption,
thin_tac "\<forall>f\<in>{j. j \<le> n} \<rightarrow> set_mult R (R \<diamondsuit>\<^sub>p a) (R \<diamondsuit>\<^sub>p b).
\<Sigma>\<^sub>e R f n \<in> R \<diamondsuit>\<^sub>p (a \<cdot>\<^sub>r b)")
apply (frule ring_tOp_closed[of a b], assumption)
apply (frule principal_ideal[of "a \<cdot>\<^sub>r b"], simp,
rule ideal_pOp_closed, assumption+)
apply (cut_tac n = "Suc n" in n_in_Nsetn,
frule_tac f = f and A = "{j. j \<le> Suc n}" and
B = "set_mult R (R \<diamondsuit>\<^sub>p a) (R \<diamondsuit>\<^sub>p b)" in funcset_mem, assumption)
apply (simp add:prod_principal_idealTr0)
done
lemma (in Ring) prod_principal_ideal:"\<lbrakk>a \<in> carrier R; b \<in> carrier R\<rbrakk> \<Longrightarrow>
(Rxa R a) \<diamondsuit>\<^sub>r (Rxa R b) = Rxa R (a \<cdot>\<^sub>r b)"
apply (frule principal_ideal[of "a"],
frule principal_ideal[of "b"])
apply (subst sum_mult_is_ideal_prod[THEN sym, of "Rxa R a" "Rxa R b"],
assumption+)
apply (rule equalityI)
apply (rule subsetI)
apply (simp add:sum_mult_def)
apply (erule exE, erule bexE)
apply (frule sym, thin_tac "\<Sigma>\<^sub>e R f n = x", simp, thin_tac "x = \<Sigma>\<^sub>e R f n")
apply (simp add:prod_principal_idealTr1)
apply (rule subsetI)
apply (simp add:Rxa_def, fold Rxa_def)
apply (erule bexE)
apply (simp add:ring_tOp_assoc[THEN sym])
apply (frule ideal_subset1[of "R \<diamondsuit>\<^sub>p a"],
frule ideal_subset1[of "R \<diamondsuit>\<^sub>p b"])
apply (rule_tac a = "r \<cdot>\<^sub>r a" and b = b in times_mem_sum_mult[of "R \<diamondsuit>\<^sub>p a"
"R \<diamondsuit>\<^sub>p b"], assumption+)
apply (simp add:Rxa_def, blast)
apply (simp add:a_in_principal)
done
lemma (in Ring) principal_ideal_n_pow1:"a \<in> carrier R \<Longrightarrow>
(Rxa R a)\<^bsup>\<diamondsuit>R n\<^esup> = Rxa R (a^\<^bsup>R n \<^esup>)"
apply (cut_tac ring_one)
apply (induct_tac n)
apply simp
apply (cut_tac a_in_principal[of "1\<^sub>r"])
apply (frule principal_ideal[of "1\<^sub>r"])
apply (frule ideal_inc_one, assumption, simp)
apply (simp add:ring_one)
apply simp
apply (frule_tac n = n in npClose[of a],
subst prod_principal_ideal, assumption+)
apply (simp add:ring_tOp_commute)
done
lemma (in Ring) principal_ideal_n_pow:"\<lbrakk>a \<in> carrier R; I = Rxa R a\<rbrakk> \<Longrightarrow>
I \<^bsup>\<diamondsuit>R n\<^esup> = Rxa R (a^\<^bsup>R n\<^esup>)"
apply simp
apply (rule principal_ideal_n_pow1[of "a" "n"], assumption+)
done
text\<open>more about \<open>ideal_n_prod\<close>\<close>
lemma (in Ring) nprod_eqTr:" f \<in> {j. j \<le> (n::nat)} \<rightarrow> carrier R \<and>
g \<in> {j. j \<le> n} \<rightarrow> carrier R \<and> (\<forall>j \<le> n. f j = g j) \<longrightarrow>
nprod R f n = nprod R g n"
apply (induct_tac n)
apply simp
apply (rule impI, (erule conjE)+)
apply (frule func_pre[of f], frule func_pre[of g],
cut_tac n = n in Nsetn_sub_mem1, simp)
done
lemma (in Ring) nprod_eq:"\<lbrakk>\<forall>j \<le> n. f j \<in> carrier R; \<forall>j \<le> n. g j \<in> carrier R;
(\<forall>j \<le> (n::nat). f j = g j)\<rbrakk> \<Longrightarrow> nprod R f n = nprod R g n"
apply (cut_tac nprod_eqTr[of f n g])
apply simp
done
definition
mprod_expR :: "[('b, 'm) Ring_scheme, nat \<Rightarrow> nat, nat \<Rightarrow> 'b, nat] \<Rightarrow> 'b" where
"mprod_expR R e f n = nprod R (\<lambda>j. ((f j)^\<^bsup>R (e j)\<^esup>)) n"
(** Note that e j is a natural number for all j in Nset n **)
lemma (in Ring) mprodR_Suc:"\<lbrakk>e \<in> {j. j \<le> (Suc n)} \<rightarrow> {j. (0::nat) \<le> j};
f \<in> {j. j \<le> (Suc n)} \<rightarrow> carrier R\<rbrakk> \<Longrightarrow>
mprod_expR R e f (Suc n) =
(mprod_expR R e f n) \<cdot>\<^sub>r ((f (Suc n))^\<^bsup>R (e (Suc n))\<^esup>)"
apply (simp add:mprod_expR_def)
done
lemma (in Ring) mprod_expR_memTr:"e \<in> {j. j \<le> n} \<rightarrow> {j. (0::nat) \<le> j} \<and>
f \<in> {j. j \<le> n} \<rightarrow> carrier R \<longrightarrow> mprod_expR R e f n \<in> carrier R"
apply (induct_tac n)
apply (rule impI, (erule conjE)+)
apply (cut_tac n_in_Nsetn[of 0],
simp add: mprod_expR_def)
apply (rule npClose,
simp add:Pi_def)
apply (rule impI, (erule conjE)+)
apply (frule func_pre[of "e"], frule func_pre[of "f"])
apply simp
apply (simp add:mprodR_Suc)
apply (rule ring_tOp_closed, assumption+)
apply (rule npClose, cut_tac n = "Suc n" in n_in_Nsetn)
apply (simp add:Pi_def)
done
lemma (in Ring) mprod_expR_mem:"\<lbrakk> e \<in> {j. j \<le> n} \<rightarrow> {j. (0::nat) \<le> j};
f \<in> {j. j \<le> n} \<rightarrow> carrier R\<rbrakk> \<Longrightarrow> mprod_expR R e f n \<in> carrier R"
apply (simp add:mprod_expR_memTr)
done
lemma (in Ring) prod_n_principal_idealTr:"e \<in> {j. j\<le>n} \<rightarrow> {j. (0::nat)\<le>j} \<and>
f \<in> {j. j\<le>n} \<rightarrow> carrier R \<and> (\<forall>k \<le> n. J k = (Rxa R (f k))\<^bsup>\<diamondsuit>R (e k)\<^esup>) \<longrightarrow>
ideal_n_prod R n J = Rxa R (mprod_expR R e f n)"
apply (induct_tac n)
apply (rule impI) apply (erule conjE)+
apply (simp add:mprod_expR_def)
apply (subgoal_tac "J 0 = R \<diamondsuit>\<^sub>p (f 0) \<^bsup>\<diamondsuit>R (e 0)\<^esup>")
apply simp
apply (rule principal_ideal_n_pow[of "f 0" "R \<diamondsuit>\<^sub>p (f 0)"])
apply (cut_tac n_in_Nsetn[of 0], simp add:Pi_def) apply simp
apply (cut_tac n_in_Nsetn[of 0], simp)
apply (rule impI, (erule conjE)+)
apply (frule func_pre[of "e"], frule func_pre[of "f"])
apply (cut_tac n = n in Nsetn_sub_mem1,
simp add:mprodR_Suc)
apply (cut_tac n = "Suc n" in n_in_Nsetn, simp)
apply (frule_tac A = "{j. j \<le> Suc n}" and x = "Suc n" in funcset_mem[of "f" _ "carrier R"], simp)
apply (frule_tac a = "f (Suc n)" and I = "R \<diamondsuit>\<^sub>p (f (Suc n))" and n = "e (Suc n)" in principal_ideal_n_pow) apply simp
apply (subst prod_principal_ideal[THEN sym])
apply (simp add:mprod_expR_mem)
apply (rule npClose, assumption+) apply simp
done
(************* used in Valuation2.thy *****************)
lemma (in Ring) prod_n_principal_ideal:"\<lbrakk>e \<in> {j. j\<le>n} \<rightarrow> {j. (0::nat)\<le>j};
f \<in> {j. j\<le>n} \<rightarrow> carrier R; \<forall>k\<le> n. J k = (Rxa R (f k))\<^bsup>\<diamondsuit>R (e k)\<^esup>\<rbrakk> \<Longrightarrow>
ideal_n_prod R n J = Rxa R (mprod_expR R e f n)"
apply (simp add:prod_n_principal_idealTr[of e n f J])
done
(*******************************************************)
lemma (in Idomain) a_notin_n_pow1:"\<lbrakk>a \<in> carrier R; \<not> Unit R a; a \<noteq> \<zero>; 0 < n\<rbrakk>
\<Longrightarrow> a \<notin> (Rxa R a) \<^bsup>\<diamondsuit>R (Suc n)\<^esup>"
apply (rule contrapos_pp)
apply (simp del:ipSuc) apply (simp del:ipSuc)
apply (frule principal_ideal[of "a"])
apply (frule principal_ideal_n_pow[of "a" "R \<diamondsuit>\<^sub>p a" "Suc n"])
apply simp apply (simp del:ipSuc)
apply (thin_tac "R \<diamondsuit>\<^sub>p a \<^bsup>\<diamondsuit>R (Suc n)\<^esup> = R \<diamondsuit>\<^sub>p (a^\<^bsup>R n\<^esup> \<cdot>\<^sub>r a)")
apply (thin_tac "ideal R (R \<diamondsuit>\<^sub>p a)")
apply (simp add:Rxa_def)
apply (erule bexE)
apply (frule npClose[of "a" "n"])
apply (simp add:ring_tOp_assoc[THEN sym])
apply (frule ring_l_one[THEN sym, of "a"])
apply (subgoal_tac "1\<^sub>r \<cdot>\<^sub>r a = r \<cdot>\<^sub>r a^\<^bsup>R n\<^esup> \<cdot>\<^sub>r a")
apply (cut_tac b = "r \<cdot>\<^sub>r (a^\<^bsup>R n\<^esup>)" in idom_mult_cancel_r[of "1\<^sub>r" _ "a"])
apply (simp add:ring_one) apply (simp add:ring_tOp_closed)
apply assumption+
apply (thin_tac "1\<^sub>r \<cdot>\<^sub>r a = r \<cdot>\<^sub>r a^\<^bsup>R n\<^esup> \<cdot>\<^sub>r a",
thin_tac "a = 1\<^sub>r \<cdot>\<^sub>r a",
thin_tac "a = r \<cdot>\<^sub>r a^\<^bsup>R n\<^esup> \<cdot>\<^sub>r a")
apply (subgoal_tac "1\<^sub>r = r \<cdot>\<^sub>r (a^\<^bsup>R (Suc (n - Suc 0))\<^esup>)") prefer 2
apply (simp del:ipSuc)
apply (thin_tac "1\<^sub>r = r \<cdot>\<^sub>r a^\<^bsup>R n\<^esup>")
apply (simp del:Suc_pred)
apply (frule npClose[of "a" "n - Suc 0"])
apply (simp add:ring_tOp_assoc[THEN sym])
apply (frule_tac x = r and y = "a^\<^bsup>R (n - Suc 0)\<^esup>" in ring_tOp_closed, assumption)
apply (simp add:ring_tOp_commute[of _ a])
apply (simp add:Unit_def) apply blast
apply simp
done
lemma (in Idomain) a_notin_n_pow2:"\<lbrakk>a \<in> carrier R; \<not> Unit R a; a \<noteq> \<zero>;
0 < n\<rbrakk> \<Longrightarrow> a^\<^bsup>R n\<^esup> \<notin> (Rxa R a) \<^bsup>\<diamondsuit>R (Suc n)\<^esup>"
apply (rule contrapos_pp)
apply (simp del:ipSuc, simp del:ipSuc)
apply (frule principal_ideal[of "a"])
apply (frule principal_ideal_n_pow[of "a" "R \<diamondsuit>\<^sub>p a" "Suc n"])
apply (simp, simp del:ipSuc)
apply (thin_tac "R \<diamondsuit>\<^sub>p a \<^bsup>\<diamondsuit>R (Suc n)\<^esup> = R \<diamondsuit>\<^sub>p (a^\<^bsup>R n\<^esup> \<cdot>\<^sub>r a)")
apply (thin_tac "ideal R (R \<diamondsuit>\<^sub>p a)")
apply (simp add:Rxa_def)
apply (erule bexE)
apply (frule idom_potent_nonzero[of "a" "n"], assumption+)
apply (frule npClose[of "a" "n"])
apply (frule ring_l_one[THEN sym, of "a^\<^bsup>R n\<^esup> "])
apply (subgoal_tac "1\<^sub>r \<cdot>\<^sub>r (a^\<^bsup>R n\<^esup>) = r \<cdot>\<^sub>r ((a^\<^bsup>R n\<^esup>) \<cdot>\<^sub>r a)")
prefer 2 apply simp
apply (thin_tac "a^\<^bsup>R n\<^esup> = 1\<^sub>r \<cdot>\<^sub>r a^\<^bsup>R n\<^esup>",
thin_tac "a^\<^bsup>R n\<^esup> = r \<cdot>\<^sub>r (a^\<^bsup>R n\<^esup> \<cdot>\<^sub>r a)")
apply (simp add:ring_tOp_commute[of "a^\<^bsup>R n\<^esup>" a])
apply (simp add:ring_tOp_assoc[THEN sym])
apply (cut_tac ring_one,
frule_tac b = "r \<cdot>\<^sub>r a" in idom_mult_cancel_r[of "1\<^sub>r" _ "a^\<^bsup>R n\<^esup>"],
simp add:ring_tOp_closed,
assumption+)
apply (simp add:ring_tOp_commute[of _ a])
apply (simp add:Unit_def, blast)
done
lemma (in Idomain) n_pow_not_prime:"\<lbrakk>a \<in> carrier R; a \<noteq> \<zero>; 0 < n\<rbrakk>
\<Longrightarrow> \<not> prime_ideal R ((Rxa R a) \<^bsup>\<diamondsuit>R (Suc n)\<^esup>)"
apply (case_tac "n = 0")
apply simp
apply (case_tac "Unit R a")
apply (simp del:ipSuc add:prime_ideal_def, rule impI)
apply (frule principal_ideal[of "a"])
apply (frule principal_ideal_n_pow[of "a" "R \<diamondsuit>\<^sub>p a" "Suc n"])
apply simp apply (simp del:npow_suc)
apply (simp del:npow_suc add:idom_potent_unit [of "a" "Suc n"])
apply (thin_tac "R \<diamondsuit>\<^sub>p a \<diamondsuit>\<^sub>r R \<diamondsuit>\<^sub>p a \<^bsup>\<diamondsuit>R n\<^esup> = R \<diamondsuit>\<^sub>p (a^\<^bsup>R (Suc n)\<^esup>)")
apply (frule npClose[of "a" "Suc n"])
apply (frule a_in_principal[of "a^\<^bsup>R (Suc n)\<^esup>"])
apply (simp add: ideal_inc_unit)
apply (frule a_notin_n_pow1[of "a" "n"], assumption+)
apply (frule a_notin_n_pow2[of "a" "n"], assumption+)
apply (frule npClose[of "a" "n"])
apply (frule principal_ideal[of "a"])
apply (frule principal_ideal_n_pow[of "a" "R \<diamondsuit>\<^sub>p a" "Suc n"])
apply simp apply (simp del:ipSuc npow_suc)
apply (thin_tac "R \<diamondsuit>\<^sub>p a \<^bsup>\<diamondsuit>R (Suc n)\<^esup> = R \<diamondsuit>\<^sub>p (a^\<^bsup>R (Suc n)\<^esup>)")
apply (subst prime_ideal_def)
apply (simp del:npow_suc) apply (rule impI)
apply (subgoal_tac "(a^\<^bsup>R n\<^esup>) \<cdot>\<^sub>r a \<in> R \<diamondsuit>\<^sub>p (a^\<^bsup>R (Suc n)\<^esup>)")
apply blast
apply (simp add:Rxa_def)
apply (frule ring_tOp_closed[of "a" "a^\<^bsup>R n\<^esup>"], assumption+)
apply (frule ring_l_one[THEN sym, of "a \<cdot>\<^sub>r (a^\<^bsup>R n\<^esup>)"])
apply (cut_tac ring_one)
apply (simp add:ring_tOp_commute[of _ a], blast)
done
lemma (in Idomain) principal_pow_prime_condTr:
"\<lbrakk>a \<in> carrier R; a \<noteq> \<zero>; prime_ideal R ((Rxa R a) \<^bsup>\<diamondsuit>R (Suc n)\<^esup>)\<rbrakk> \<Longrightarrow> n = 0"
apply (rule contrapos_pp, (simp del:ipSuc)+)
apply (frule n_pow_not_prime[of "a" "n"], assumption+)
apply (simp del:ipSuc)
done
lemma (in Idomain) principal_pow_prime_cond:
"\<lbrakk>a \<in> carrier R; a \<noteq> \<zero>; prime_ideal R ((Rxa R a) \<^bsup>\<diamondsuit>R n\<^esup>)\<rbrakk> \<Longrightarrow> n = Suc 0"
apply (case_tac "n = 0")
apply simp
apply (simp add:prime_ideal_def) apply (erule conjE)
apply (cut_tac ring_one, simp)
apply (subgoal_tac "prime_ideal R (R \<diamondsuit>\<^sub>p a \<^bsup>\<diamondsuit>R (Suc (n - Suc 0))\<^esup>)")
apply (frule principal_pow_prime_condTr[of "a" "n - Suc 0"], assumption+)
apply simp apply simp
done
section "Extension and contraction"
locale TwoRings = Ring +
fixes R' (structure)
assumes secondR: "Ring R'"
definition
i_contract :: "['a \<Rightarrow> 'b, ('a, 'm1) Ring_scheme, ('b, 'm2) Ring_scheme,
'b set] \<Rightarrow> 'a set" where
"i_contract f R R' J = invim f (carrier R) J"
definition
i_extension :: "['a \<Rightarrow> 'b, ('a, 'm1) Ring_scheme, ('b, 'm2) Ring_scheme,
'a set] \<Rightarrow> 'b set" where
"i_extension f R R' I = sum_mult R' (f ` I) (carrier R')"
lemma (in TwoRings) i_contract_sub:"\<lbrakk>f \<in> rHom R R'; ideal R' J \<rbrakk> \<Longrightarrow>
(i_contract f R R' J) \<subseteq> carrier R"
by (auto simp add:i_contract_def invim_def)
lemma (in TwoRings) i_contract_ideal:"\<lbrakk>f \<in> rHom R R'; ideal R' J \<rbrakk> \<Longrightarrow>
ideal R (i_contract f R R' J)"
apply (cut_tac Ring,
cut_tac secondR)
apply (rule ideal_condition)
apply (simp add:i_contract_sub)
apply (simp add:i_contract_def invim_def)
apply (cut_tac ring_zero)
apply (cut_tac Ring)
apply (frule rHom_0_0[of R R' f], assumption+,
cut_tac Ring.ideal_zero[of R' J])
apply (frule sym, thin_tac "f \<zero> = \<zero>\<^bsub>R'\<^esub>", simp, blast,
assumption+)
apply (rule ballI)+
apply (simp add:i_contract_def invim_def, (erule conjE)+)
apply (cut_tac ring_is_ag,
frule_tac x = y in aGroup.ag_mOp_closed[of R], assumption)
apply (simp add:aGroup.ag_pOp_closed)
apply (simp add:rHom_add)
apply (frule_tac x = y in rHom_inv_inv[of R R' _ f], assumption+, simp,
thin_tac "f (-\<^sub>a y) = -\<^sub>a\<^bsub>R'\<^esub> (f y)",
frule_tac x = "f y" in Ring.ideal_inv1_closed[of R' J], assumption+,
rule Ring.ideal_pOp_closed[of R'], assumption+)
apply ((rule ballI)+,
simp add:i_contract_def invim_def, erule conjE,
simp add:ring_tOp_closed,
simp add:rHom_tOp)
apply (frule_tac a = r in rHom_mem[of f R R'], assumption,
simp add:Ring.ideal_ring_multiple[of R' J])
done
lemma (in TwoRings) i_contract_mono:"\<lbrakk>f \<in> rHom R R'; ideal R' J1; ideal R' J2;
J1 \<subseteq> J2 \<rbrakk> \<Longrightarrow> i_contract f R R' J1 \<subseteq> i_contract f R R' J2"
apply (rule subsetI)
apply (simp add:i_contract_def invim_def) apply (erule conjE)
apply (rule subsetD, assumption+)
done
lemma (in TwoRings) i_contract_prime:"\<lbrakk>f \<in> rHom R R'; prime_ideal R' P\<rbrakk> \<Longrightarrow>
prime_ideal R (i_contract f R R' P)"
apply (cut_tac Ring,
cut_tac secondR)
apply (simp add:prime_ideal_def, (erule conjE)+)
apply (simp add:i_contract_ideal)
apply (rule conjI)
apply (rule contrapos_pp, simp+)
apply (simp add:i_contract_def invim_def, erule conjE)
apply (simp add:rHom_one)
apply (rule ballI)+
apply (frule_tac a = x in rHom_mem[of "f" "R" "R'"], assumption+,
frule_tac a = y in rHom_mem[of "f" "R" "R'"], assumption+)
apply (rule impI)
apply (simp add:i_contract_def invim_def, erule conjE)
apply (simp add:rHom_tOp)
done
lemma (in TwoRings) i_extension_ideal:"\<lbrakk>f \<in> rHom R R'; ideal R I \<rbrakk> \<Longrightarrow>
ideal R' (i_extension f R R' I)"
apply (cut_tac Ring, cut_tac secondR)
apply (simp add:i_extension_def)
apply (rule Ring.ideal_sum_mult [of "R'" "f ` I" "carrier R'"], assumption+)
apply (rule subsetI)
apply (simp add:image_def)
apply (erule bexE, frule_tac a = xa in rHom_mem[of f R R'],
rule ideal_subset, assumption+, simp)
apply (frule ideal_zero, simp, blast)
apply (simp add:Ring.whole_ideal[of R'])
done
lemma (in TwoRings) i_extension_mono:"\<lbrakk>f \<in> rHom R R'; ideal R I1; ideal R I2;
I1 \<subseteq> I2 \<rbrakk> \<Longrightarrow> (i_extension f R R' I1) \<subseteq> (i_extension f R R' I2)"
apply (rule subsetI)
apply (simp add:i_extension_def)
apply (simp add:sum_mult_def)
apply (erule exE, erule bexE)
apply (cut_tac Ring.set_mult_mono[of R' "f ` I1" "f ` I2" "carrier R'"])
apply (frule_tac f = fa and A = "{j. j \<le> n}" in extend_fun[of _ _
"set_mult R' (f ` I1) (carrier R')" "set_mult R' (f ` I2) (carrier R')"],
assumption+) apply blast
apply (simp add:secondR)
apply (simp add:image_def, rule subsetI, simp, erule bexE,
frule_tac h = xb in ideal_subset[of I1], assumption, simp add:rHom_mem)
apply (simp add:image_def, rule subsetI, simp, erule bexE,
frule_tac h = xb in ideal_subset[of I2], assumption, simp add:rHom_mem)
apply (rule subsetI,
simp add:image_def, erule bexE,
frule_tac c = xb in subsetD[of I1 I2], assumption+, blast)
apply simp
done
lemma (in TwoRings) e_c_inc_self:"\<lbrakk>f \<in> rHom R R'; ideal R I\<rbrakk> \<Longrightarrow>
I \<subseteq> i_contract f R R' (i_extension f R R' I)"
apply (rule subsetI)
apply (simp add:i_contract_def i_extension_def invim_def)
apply (simp add:ideal_subset)
apply (cut_tac secondR,
frule Ring.ring_one [of "R'"])
apply (frule_tac h = x in ideal_subset[of I], assumption,
frule_tac f = f and A = R and R = R' and a = x in rHom_mem, assumption)
apply (frule_tac t = "f x" in Ring.ring_r_one[THEN sym, of R'], assumption)
apply (frule_tac a = "f x" and b = "1\<^sub>r\<^bsub>R'\<^esub>" in Ring.times_mem_sum_mult[of R'
"f ` I" "carrier R'"],
rule subsetI,
simp add:image_def, erule bexE,
frule_tac h = xb in ideal_subset[of I], assumption,
simp add:rHom_mem, simp,
simp add:image_def, blast, assumption+)
apply simp
done
lemma (in TwoRings) c_e_incd_self:"\<lbrakk>f \<in> rHom R R'; ideal R' J \<rbrakk> \<Longrightarrow>
i_extension f R R' (i_contract f R R' J) \<subseteq> J"
apply (rule subsetI)
apply (simp add:i_extension_def)
apply (simp add:sum_mult_def)
apply (erule exE, erule bexE)
apply (cut_tac secondR,
frule_tac n = n and f = fa in Ring.ideal_nsum_closed[of R' J ],
assumption)
apply (rule allI, rule impI) apply (
frule_tac f = fa and A = "{j. j \<le> n}" and
B = "set_mult R' (f ` i_contract f R R' J) (carrier R')" and x = j in
funcset_mem, simp) apply (
thin_tac "fa \<in> {j. j \<le> n} \<rightarrow> set_mult R' (f ` i_contract f R R' J) (carrier R')")
apply (simp add:set_mult_def, (erule bexE)+,
simp add:i_contract_def invim_def, erule conjE)
apply (frule_tac x = "f xa" and r = y in Ring.ideal_ring_multiple1[of R' J],
assumption+, simp)
apply simp
done
lemma (in TwoRings) c_e_c_eq_c:"\<lbrakk>f \<in> rHom R R'; ideal R' J \<rbrakk> \<Longrightarrow>
i_contract f R R' (i_extension f R R' (i_contract f R R' J))
= i_contract f R R' J"
apply (frule i_contract_ideal [of "f" "J"], assumption)
apply (frule e_c_inc_self [of "f" "i_contract f R R' J"], assumption+)
apply (frule c_e_incd_self [of "f" "J"], assumption+)
apply (frule i_contract_mono [of "f"
"i_extension f R R' (i_contract f R R' J)" "J"])
apply (rule i_extension_ideal, assumption+)
apply (rule equalityI, assumption+)
done
lemma (in TwoRings) e_c_e_eq_e:"\<lbrakk>f \<in> rHom R R'; ideal R I \<rbrakk> \<Longrightarrow>
i_extension f R R' (i_contract f R R' (i_extension f R R' I))
= i_extension f R R' I"
apply (frule i_extension_ideal [of "f" "I"], assumption+)
apply (frule c_e_incd_self [of "f" "i_extension f R R' I"], assumption+)
apply (rule equalityI, assumption+)
apply (thin_tac "i_extension f R R' (i_contract f R R' (i_extension f R R' I))
\<subseteq> i_extension f R R' I")
apply (frule e_c_inc_self [of "f" "I"], assumption+)
apply (rule i_extension_mono [of "f" "I"
"i_contract f R R' (i_extension f R R' I)"], assumption+)
apply (rule i_contract_ideal, assumption+)
done
section "Complete system of representatives"
definition
csrp_fn :: "[_, 'a set] \<Rightarrow> 'a set \<Rightarrow> 'a" where
"csrp_fn R I = (\<lambda>x\<in>carrier (R /\<^sub>r I). (if x = I then \<zero>\<^bsub>R\<^esub> else SOME y. y \<in> x))"
definition
csrp :: "[_ , 'a set] \<Rightarrow> 'a set" where
"csrp R I == (csrp_fn R I) ` (carrier (R /\<^sub>r I))"
(** complete system of representatives having 1-1 correspondence with
carrier (R /\<^sub>r I) **)
lemma (in Ring) csrp_mem:"\<lbrakk>ideal R I; a \<in> carrier R\<rbrakk> \<Longrightarrow>
csrp_fn R I (a \<uplus>\<^bsub>R\<^esub> I) \<in> a \<uplus>\<^bsub>R\<^esub> I"
apply (simp add:csrp_fn_def qring_carrier)
apply (case_tac "a \<uplus>\<^bsub>R\<^esub> I = I") apply simp
apply (rule conjI, rule impI)
apply (simp add:ideal_zero)
apply (rule impI)
apply (cut_tac ring_zero)
apply (frule_tac x = \<zero> in bspec, assumption+)
apply (thin_tac "\<forall>a\<in>carrier R. a \<uplus>\<^bsub>R\<^esub> I \<noteq> I")
apply (frule ideal_zero[of "I"])
apply (frule ar_coset_same4[of "I" "\<zero>"], assumption+, simp)
apply simp
apply (rule conjI)
apply (rule impI, rule someI2_ex)
apply (frule a_in_ar_coset[of "I" "a"], assumption+, blast, assumption+)
apply (rule impI)
apply (frule_tac x = a in bspec, assumption+,
thin_tac "\<forall>aa\<in>carrier R. aa \<uplus>\<^bsub>R\<^esub> I \<noteq> a \<uplus>\<^bsub>R\<^esub> I", simp)
done
lemma (in Ring) csrp_same:"\<lbrakk>ideal R I; a \<in> carrier R\<rbrakk> \<Longrightarrow>
csrp_fn R I (a \<uplus>\<^bsub>R\<^esub> I) \<uplus>\<^bsub>R\<^esub> I = a \<uplus>\<^bsub>R\<^esub> I"
apply (frule csrp_mem[of "I" "a"], assumption+)
apply (rule ar_cos_same[of "a" "I" "csrp_fn R I (a \<uplus>\<^bsub>R\<^esub> I)"], assumption+)
done
lemma (in Ring) csrp_mem1:"\<lbrakk>ideal R I; x \<in> carrier (R /\<^sub>r I)\<rbrakk> \<Longrightarrow>
csrp_fn R I x \<in> x"
apply (simp add:qring_carrier, erule bexE, frule sym,
thin_tac "a \<uplus>\<^bsub>R\<^esub> I = x", simp)
apply (simp add:csrp_mem)
done
lemma (in Ring) csrp_fn_mem:"\<lbrakk>ideal R I; x \<in> carrier (R /\<^sub>r I)\<rbrakk> \<Longrightarrow>
(csrp_fn R I x) \<in> carrier R"
apply (simp add:qring_carrier, erule bexE, frule sym,
thin_tac "a \<uplus>\<^bsub>R\<^esub> I = x", simp,
frule_tac a = a in csrp_mem[of "I"], assumption+)
apply (rule_tac a = a and x = "csrp_fn R I (a \<uplus>\<^bsub>R\<^esub> I)" in
ar_coset_subsetD[of "I"], assumption+)
done
lemma (in Ring) csrp_eq_coset:"\<lbrakk>ideal R I; x \<in> carrier (R /\<^sub>r I)\<rbrakk> \<Longrightarrow>
(csrp_fn R I x) \<uplus>\<^bsub>R\<^esub> I = x"
apply (simp add:qring_carrier, erule bexE)
apply (frule sym, thin_tac "a \<uplus>\<^bsub>R\<^esub> I = x", simp)
apply (frule_tac a = a in csrp_mem[of "I"], assumption+)
apply (rule ar_cos_same, assumption+)
done
lemma (in Ring) csrp_nz_nz:"\<lbrakk>ideal R I; x \<in> carrier (R /\<^sub>r I);
x \<noteq> \<zero>\<^bsub>(R /\<^sub>r I)\<^esub>\<rbrakk> \<Longrightarrow> (csrp_fn R I x) \<noteq> \<zero>"
apply (rule contrapos_pp, simp+)
apply (frule csrp_eq_coset[of "I" "x"], assumption+, simp)
apply (simp add:qring_zero[of "I"])
apply (frule ideal_zero[of "I"]) apply (
cut_tac ring_zero)
apply (simp add:Qring_fix1 [of "\<zero>" "I"])
done
lemma (in Ring) csrp_diff_in_vpr:"\<lbrakk>ideal R I; x \<in> carrier R\<rbrakk> \<Longrightarrow>
x \<plusminus> (-\<^sub>a (csrp_fn R I (pj R I x))) \<in> I"
apply (frule csrp_mem[of "I" "x"],
frule csrp_same[of "I" "x"],
simp add:pj_mem, assumption,
frule ar_coset_subsetD[of I x "csrp_fn R I (x \<uplus>\<^bsub>R\<^esub> I)"],
assumption+)
apply (frule belong_ar_coset2[of I x "csrp_fn R I (x \<uplus>\<^bsub>R\<^esub> I)"], assumption+,
frule ideal_inv1_closed[of I "csrp_fn R I (x \<uplus>\<^bsub>R\<^esub> I) \<plusminus> -\<^sub>a x"], assumption+,
cut_tac ring_is_ag,
frule aGroup.ag_mOp_closed[of R x], assumption,
simp add:aGroup.ag_pOp_commute[of R "csrp_fn R I (x \<uplus>\<^bsub>R\<^esub> I)" "-\<^sub>a x"])
apply (simp add:aGroup.ag_p_inv[of R "-\<^sub>a x" "csrp_fn R I (x \<uplus>\<^bsub>R\<^esub> I)"],
simp add:aGroup.ag_inv_inv,
cut_tac Ring, simp add:pj_mem[of R I x])
done
lemma (in Ring) csrp_pj:"\<lbrakk>ideal R I; x \<in> carrier (R /\<^sub>r I)\<rbrakk> \<Longrightarrow>
(pj R I) (csrp_fn R I x) = x"
apply(cut_tac Ring,
frule csrp_fn_mem[of "I" "x"], assumption+,
simp add:pj_mem[of "R" "I" "csrp_fn R I x"],
simp add:csrp_eq_coset)
done
section "Polynomial ring"
text\<open>In this section, we treat a ring of polynomials over a ring S.
Numbers are of type ant\<close>
definition
pol_coeff :: "[('a, 'more) Ring_scheme, (nat \<times> (nat \<Rightarrow> 'a))] \<Rightarrow> bool" where
"pol_coeff S c \<longleftrightarrow> (\<forall>j \<le> (fst c). (snd c) j \<in> carrier S)"
definition
c_max :: "[('a, 'more) Ring_scheme, nat \<times> (nat \<Rightarrow> 'a)] \<Rightarrow> nat" where
"c_max S c = (if {j. j \<le> (fst c) \<and> (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>} = {} then 0 else
n_max {j. j \<le> (fst c) \<and> (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>})"
definition
polyn_expr :: "[('a, 'more) Ring_scheme, 'a, nat, nat \<times> (nat \<Rightarrow> 'a)] \<Rightarrow> 'a" where
"polyn_expr R X k c == nsum R (\<lambda>j. ((snd c) j) \<cdot>\<^sub>r\<^bsub>R\<^esub> (X^\<^bsup>R j\<^esup>)) k"
definition
algfree_cond :: "[('a, 'm) Ring_scheme, ('a, 'm1) Ring_scheme,
'a] \<Rightarrow> bool" where
"algfree_cond R S X \<longleftrightarrow> (\<forall>c. pol_coeff S c \<and> (\<forall>k \<le> (fst c).
(nsum R (\<lambda>j. ((snd c) j) \<cdot>\<^sub>r\<^bsub>R\<^esub> (X^\<^bsup>R j\<^esup>)) k = \<zero>\<^bsub>R\<^esub> \<longrightarrow>
(\<forall>j \<le> k. (snd c) j = \<zero>\<^bsub>S\<^esub>))))"
locale PolynRg = Ring +
fixes S (structure)
fixes X (structure)
assumes X_mem_R:"X \<in> carrier R"
and not_zeroring:"\<not> Zero_ring S"
and subring: "Subring R S"
and algfree: "algfree_cond R S X"
and S_X_generate:"x \<in> carrier R \<Longrightarrow>
\<exists>f. pol_coeff S f \<and> x = polyn_expr R X (fst f) f"
(** a polynomial is an element of a polynomial ring **)
section \<open>Addition and multiplication of \<open>polyn_exprs\<close>\<close>
subsection \<open>Simple properties of a \<open>polyn_ring\<close>\<close>
lemma Subring_subset:"Subring R S \<Longrightarrow> carrier S \<subseteq> carrier R"
by (simp add:Subring_def)
lemma (in Ring) subring_Ring:"Subring R S \<Longrightarrow> Ring S"
by (simp add:Subring_def)
lemma (in Ring) mem_subring_mem_ring:"\<lbrakk>Subring R S; x \<in> carrier S\<rbrakk> \<Longrightarrow>
x \<in> carrier R"
by (simp add:Subring_def, (erule conjE)+, simp add: subsetD)
lemma (in Ring) Subring_pOp_ring_pOp:"\<lbrakk>Subring R S; a \<in> carrier S;
b \<in> carrier S \<rbrakk> \<Longrightarrow> a \<plusminus>\<^bsub>S\<^esub> b = a \<plusminus> b"
apply (simp add:Subring_def, (erule conjE)+)
apply (frule rHom_add[of "ridmap S" S R a b], assumption+)
apply (cut_tac Ring.ring_is_ag[of S],
frule aGroup.ag_pOp_closed[of S a b], assumption+,
simp add:ridmap_def, assumption)
done
lemma (in Ring) Subring_tOp_ring_tOp:"\<lbrakk>Subring R S; a \<in> carrier S;
b \<in> carrier S \<rbrakk> \<Longrightarrow> a \<cdot>\<^sub>r\<^bsub>S\<^esub> b = a \<cdot>\<^sub>r b"
apply (simp add:Subring_def, (erule conjE)+)
apply (frule rHom_tOp[of "S" "R" "a" "b" "ridmap S"], rule Ring_axioms, assumption+)
apply (frule Ring.ring_tOp_closed[of "S" "a" "b"], assumption+,
simp add:ridmap_def)
done
lemma (in Ring) Subring_one_ring_one:"Subring R S \<Longrightarrow> 1\<^sub>r\<^bsub>S\<^esub> = 1\<^sub>r"
apply (simp add:Subring_def, (erule conjE)+)
apply (frule rHom_one[of "S" "R" "ridmap S"], rule Ring_axioms, assumption+)
apply (simp add:ridmap_def, simp add:Ring.ring_one[of S])
done
lemma (in Ring) Subring_zero_ring_zero:"Subring R S \<Longrightarrow> \<zero>\<^bsub>S\<^esub> = \<zero>"
apply (simp add:Subring_def, (erule conjE)+,
frule rHom_0_0[of "S" "R" "ridmap S"], rule Ring_axioms, assumption+,
simp add:ridmap_def, simp add:Ring.ring_zero[of "S"])
done
lemma (in Ring) Subring_minus_ring_minus:"\<lbrakk>Subring R S; x \<in> carrier S\<rbrakk>
\<Longrightarrow> -\<^sub>a\<^bsub>S\<^esub> x = -\<^sub>a x"
apply (simp add:Subring_def, (erule conjE)+, simp add:rHom_def, (erule conjE)+)
apply (cut_tac ring_is_ag, frule Ring.ring_is_ag[of "S"])
apply (frule aHom_inv_inv[of "S" "R" "ridmap S" "x"], assumption+,
frule aGroup.ag_mOp_closed[of "S" "x"], assumption+)
apply (simp add:ridmap_def)
done
lemma (in PolynRg) Subring_pow_ring_pow:"x \<in> carrier S \<Longrightarrow>
x^\<^bsup>S n\<^esup> = x^\<^bsup>R n\<^esup>"
apply (cut_tac subring, frule subring_Ring)
apply (induct_tac n)
apply (simp, simp add:Subring_one_ring_one)
apply (frule_tac n = n in Ring.npClose[of S x], assumption+)
apply (simp add:Subring_tOp_ring_tOp)
done
lemma (in PolynRg) is_Ring: "Ring R" ..
lemma (in PolynRg) polyn_ring_nonzero:"1\<^sub>r \<noteq> \<zero>"
apply (cut_tac Ring, cut_tac subring)
apply (simp add:Subring_zero_ring_zero[THEN sym])
apply (simp add:Subring_one_ring_one[THEN sym])
apply (simp add:not_zeroring)
done
lemma (in PolynRg) polyn_ring_S_nonzero:"1\<^sub>r\<^bsub>S\<^esub> \<noteq> \<zero>\<^bsub>S\<^esub>"
apply (cut_tac subring)
apply (simp add:Subring_zero_ring_zero)
apply (simp add:Subring_one_ring_one)
apply (simp add:polyn_ring_nonzero)
done
lemma (in PolynRg) polyn_ring_X_nonzero:"X \<noteq> \<zero>"
apply (cut_tac algfree,
cut_tac subring)
apply (simp add:algfree_cond_def)
apply (rule contrapos_pp, simp+)
apply (drule_tac x = "Suc 0" in spec)
apply (subgoal_tac "pol_coeff S ((Suc 0),
(\<lambda>j\<in>{l. l \<le> (Suc 0)}. if j = 0 then \<zero>\<^bsub>S\<^esub> else 1\<^sub>r\<^bsub>S\<^esub>))")
apply (drule_tac x = "\<lambda>j\<in>{l. l \<le> (Suc 0)}. if j = 0 then \<zero>\<^bsub>S\<^esub> else 1\<^sub>r\<^bsub>S\<^esub>" in
spec)
apply (erule conjE, simp)
apply (simp only:Nset_1)
apply (drule_tac a = "Suc 0" in forall_spec, simp)
apply simp
apply (cut_tac subring, simp add:Subring_zero_ring_zero,
simp add:Subring_one_ring_one, cut_tac ring_zero, cut_tac ring_one,
simp add:ring_r_one, simp add:ring_times_x_0, cut_tac ring_is_ag,
simp add:aGroup.ag_r_zero,
drule_tac a = "Suc 0" in forall_spec, simp, simp)
apply (cut_tac polyn_ring_S_nonzero, simp add:Subring_zero_ring_zero)
apply (thin_tac "\<forall>b. pol_coeff S (Suc 0, b) \<and>
(\<forall>k\<le>Suc 0. \<Sigma>\<^sub>e R (\<lambda>j. b j \<cdot>\<^sub>r \<zero>^\<^bsup>R j\<^esup>) k = \<zero> \<longrightarrow> (\<forall>j\<le>k. b j = \<zero>\<^bsub>S\<^esub>))",
simp add:pol_coeff_def,
rule allI,
simp add:Subring_def, simp add:Ring.ring_zero,
(rule impI)+,
simp add:Ring.ring_one)
done
subsection "Coefficients of a polynomial"
lemma (in PolynRg) pol_coeff_split:"pol_coeff S f = pol_coeff S (fst f, snd f)"
by simp
lemma (in PolynRg) pol_coeff_cartesian:"pol_coeff S c \<Longrightarrow>
(fst c, snd c) = c"
by simp
lemma (in PolynRg) split_pol_coeff:"\<lbrakk>pol_coeff S c; k \<le> (fst c)\<rbrakk> \<Longrightarrow>
pol_coeff S (k, snd c)"
by (simp add:pol_coeff_def)
lemma (in PolynRg) pol_coeff_pre:"pol_coeff S ((Suc n), f) \<Longrightarrow>
pol_coeff S (n, f)"
apply (simp add:pol_coeff_def)
done
lemma (in PolynRg) pol_coeff_le:"\<lbrakk>pol_coeff S c; n \<le> (fst c)\<rbrakk> \<Longrightarrow>
pol_coeff S (n, (snd c))"
apply (simp add:pol_coeff_def)
done
lemma (in PolynRg) pol_coeff_mem:"\<lbrakk>pol_coeff S c; j \<le> (fst c)\<rbrakk> \<Longrightarrow>
((snd c) j) \<in> carrier S"
by (simp add:pol_coeff_def)
lemma (in PolynRg) pol_coeff_mem_R:"\<lbrakk>pol_coeff S c; j \<le> (fst c)\<rbrakk>
\<Longrightarrow> ((snd c) j) \<in> carrier R"
apply (cut_tac subring, frule subring_Ring)
apply (frule pol_coeff_mem[of c "j"], assumption+,
simp add:mem_subring_mem_ring)
done
lemma (in PolynRg) Slide_pol_coeff:"\<lbrakk>pol_coeff S c; n < (fst c)\<rbrakk> \<Longrightarrow>
pol_coeff S (((fst c) - Suc n), (\<lambda>x. (snd c) (Suc (n + x))))"
apply (simp add: pol_coeff_def)
done
subsection \<open>Addition of \<open>polyn_exprs\<close>\<close>
lemma (in PolynRg) monomial_mem:"pol_coeff S c \<Longrightarrow>
\<forall>j \<le> (fst c). (snd c) j \<cdot>\<^sub>r X^\<^bsup>R j\<^esup> \<in> carrier R"
apply (rule allI, rule impI)
apply (rule ring_tOp_closed)
apply (simp add:pol_coeff_mem_R[of c],
cut_tac X_mem_R, simp add:npClose)
done
lemma (in PolynRg) polyn_mem:"\<lbrakk>pol_coeff S c; k \<le> (fst c)\<rbrakk> \<Longrightarrow>
polyn_expr R X k c \<in> carrier R"
apply (simp add:polyn_expr_def,
cut_tac ring_is_ag)
apply (rule aGroup.nsum_mem[of R k "\<lambda>j. (snd c) j \<cdot>\<^sub>r X^\<^bsup>R j\<^esup>"], assumption+)
apply (simp add:monomial_mem)
done
lemma (in PolynRg) polyn_exprs_eq:"\<lbrakk>pol_coeff S c; pol_coeff S d;
k \<le> (min (fst c) (fst d)); \<forall>j \<le> k. (snd c) j = (snd d) j\<rbrakk> \<Longrightarrow>
polyn_expr R X k c = polyn_expr R X k d"
apply (cut_tac ring_is_ag,
simp add:polyn_expr_def,
cut_tac subring,
cut_tac X_mem_R)
apply (rule aGroup.nsum_eq[of R k "\<lambda>j. (snd c) j \<cdot>\<^sub>r X^\<^bsup>R j\<^esup>"
"\<lambda>j. (snd d) j \<cdot>\<^sub>r X^\<^bsup>R j\<^esup>"], assumption)
apply (simp add:monomial_mem)+
done
lemma (in PolynRg) polyn_expr_restrict:"pol_coeff S (Suc n, f) \<Longrightarrow>
polyn_expr R X n (Suc n, f) = polyn_expr R X n (n, f)"
apply (cut_tac subring, frule subring_Ring,
cut_tac pol_coeff_le[of "(Suc n, f)" n])
apply (cut_tac polyn_exprs_eq[of "(Suc n, f)" "(n, f)" n],
(simp add:pol_coeff_split[THEN sym])+)
done
lemma (in PolynRg) polyn_expr_short:"\<lbrakk>pol_coeff S c; k \<le> (fst c)\<rbrakk> \<Longrightarrow>
polyn_expr R X k c = polyn_expr R X k (k, snd c)"
apply (rule polyn_exprs_eq[of c "(k, snd c)" k], assumption+)
apply (simp add:pol_coeff_def)
apply (simp)
apply simp
done
lemma (in PolynRg) polyn_expr0:"pol_coeff S c \<Longrightarrow>
polyn_expr R X 0 c = (snd c) 0"
apply (simp add:polyn_expr_def)
apply (cut_tac subring,
cut_tac subring_Ring[of S])
apply (frule pol_coeff_mem[of c 0], simp)
apply (frule mem_subring_mem_ring [of S "(snd c) 0"], assumption)
apply (simp add:ring_r_one, assumption)
done
lemma (in PolynRg) polyn_expr_split:"
polyn_expr R X k f = polyn_expr R X k (fst f, snd f)"
by simp
lemma (in PolynRg) polyn_Suc:"Suc n \<le> (fst c) \<Longrightarrow>
polyn_expr R X (Suc n) ((Suc n), (snd c)) =
polyn_expr R X n c \<plusminus> ((snd c) (Suc n)) \<cdot>\<^sub>r (X^\<^bsup>R (Suc n)\<^esup>)"
by (simp add:polyn_expr_def)
lemma (in PolynRg) polyn_Suc_split:"pol_coeff S (Suc n, f) \<Longrightarrow>
polyn_expr R X (Suc n) ((Suc n), f) =
polyn_expr R X n (n, f) \<plusminus> (f (Suc n)) \<cdot>\<^sub>r (X^\<^bsup>R (Suc n)\<^esup>)"
apply (cut_tac polyn_Suc[of n "(Suc n, f)"])
apply (simp del:npow_suc)
apply (subst polyn_expr_short[of "(Suc n, f)" n], assumption+, simp)
apply (simp del:npow_suc)
apply simp
done
lemma (in PolynRg) polyn_n_m:"\<lbrakk>pol_coeff S c; n < m; m \<le> (fst c)\<rbrakk> \<Longrightarrow>
polyn_expr R X m (m, (snd c)) = polyn_expr R X n (n, (snd c)) \<plusminus>
(fSum R (\<lambda>j. ((snd c) j) \<cdot>\<^sub>r (X^\<^bsup>R j\<^esup>)) (Suc n) m)"
apply (simp add:polyn_expr_def, cut_tac ring_is_ag)
apply (rule aGroup.nsum_split1[of "R" m "\<lambda>j. ((snd c) j) \<cdot>\<^sub>r (X^\<^bsup>R j\<^esup>)" n],
assumption+)
apply (rule allI, rule impI)
apply (frule_tac monomial_mem[of c],
frule_tac i = j and j = m and k = "(fst c)" in le_trans, assumption+,
simp+)
done
lemma (in PolynRg) polyn_n_m1:"\<lbrakk>pol_coeff S c; n < m; m \<le> (fst c)\<rbrakk> \<Longrightarrow>
polyn_expr R X m c = polyn_expr R X n c \<plusminus>
(fSum R (\<lambda>j. ((snd c) j) \<cdot>\<^sub>r (X^\<^bsup>R j\<^esup>)) (Suc n) m)"
apply (subst polyn_expr_short[of c n], assumption)
apply (frule_tac x = n and y = m and z = "fst c" in less_le_trans, assumption,
simp add:less_imp_le)
apply (subst polyn_expr_short[of c m], assumption+)
apply (simp add:polyn_n_m)
done
lemma (in PolynRg) polyn_n_m_mem:"\<lbrakk>pol_coeff S c; n < m; m \<le> (fst c)\<rbrakk> \<Longrightarrow>
(fSum R (\<lambda>j. ((snd c) j) \<cdot>\<^sub>r (X^\<^bsup>R j\<^esup>)) (Suc n) m) \<in> carrier R"
apply (simp add:fSum_def)
apply (cut_tac ring_is_ag,
rule_tac n = "m - Suc n" in aGroup.nsum_mem, assumption+)
apply (rule allI, rule impI,
simp del:npow_suc add:cmp_def slide_def)
apply (rule ring_tOp_closed)
apply (simp add:pol_coeff_def)
apply (frule_tac a = "Suc (n + j)" in forall_spec, arith)
apply (cut_tac subring)
apply (simp add:mem_subring_mem_ring)
apply (rule npClose)
apply (cut_tac X_mem_R,
simp del:npow_suc add:npClose)
done
lemma (in PolynRg) polyn_n_ms_eq:"\<lbrakk>pol_coeff S c; pol_coeff S d;
m \<le> min (fst c) (fst d); n < m;
\<forall>j\<in>nset (Suc n) m. (snd c) j = (snd d) j\<rbrakk> \<Longrightarrow>
(fSum R (\<lambda>j. ((snd c) j) \<cdot>\<^sub>r (X^\<^bsup>R j\<^esup>)) (Suc n) m) =
(fSum R (\<lambda>j. ((snd d) j) \<cdot>\<^sub>r (X^\<^bsup>R j\<^esup>)) (Suc n) m)"
apply (cut_tac ring_is_ag)
apply (cut_tac aGroup.fSum_eq1[of R "Suc n" m "\<lambda>j. (snd c) j \<cdot>\<^sub>r X^\<^bsup>R j\<^esup>"
"\<lambda>j. (snd d) j \<cdot>\<^sub>r X^\<^bsup>R j\<^esup>"],
assumption+)
apply (rule Suc_leI, assumption,
simp add:nset_def, simp add:monomial_mem)
apply (frule Suc_leI,
rule ballI, simp add:nset_def)
apply (simp add:monomial_mem)
apply simp
done
lemma (in PolynRg) polyn_addTr:
"(pol_coeff S (n, f)) \<and> (pol_coeff S (n, g)) \<longrightarrow>
(polyn_expr R X n (n, f)) \<plusminus> (polyn_expr R X n (n, g)) =
nsum R (\<lambda>j. ((f j) \<plusminus>\<^bsub>S\<^esub> (g j)) \<cdot>\<^sub>r (X^\<^bsup>R j\<^esup>)) n"
apply (cut_tac subring,
frule subring_Ring[of S])
apply (induct_tac n)
apply (rule impI, simp, erule conjE)
apply (simp add:polyn_expr0)
apply (cut_tac pol_coeff_mem[of "(0, f)" 0], simp,
cut_tac pol_coeff_mem[of "(0, g)" 0], simp,
frule mem_subring_mem_ring[of S "f 0"], assumption+,
frule mem_subring_mem_ring[of S "g 0"], assumption+,
frule Ring.ring_is_ag[of S],
frule aGroup.ag_pOp_closed[of S "f 0" "g 0"], assumption+,
frule mem_subring_mem_ring[of S "f 0 \<plusminus>\<^bsub>S\<^esub> g 0"], assumption+)
apply (simp add:ring_r_one)
apply (simp add:Subring_pOp_ring_pOp[of S "f 0" "g 0"])
apply (simp del:npow_suc)+
apply (rule impI, erule conjE)
apply (frule_tac n = n in pol_coeff_pre[of _ f],
frule_tac n = n in pol_coeff_pre[of _ g], simp del:npow_suc)
apply (cut_tac n = n and c = "(Suc n, f)" in polyn_Suc, simp del:npow_suc,
simp del:npow_suc,
thin_tac "polyn_expr R X (Suc n) (Suc n, f) =
polyn_expr R X n (Suc n, f) \<plusminus> f (Suc n) \<cdot>\<^sub>r X^\<^bsup>R (Suc n)\<^esup>")
apply (cut_tac n = n and c = "(Suc n, g)" in polyn_Suc, simp del:npow_suc,
simp del:npow_suc,
thin_tac "polyn_expr R X (Suc n) (Suc n, g) =
polyn_expr R X n (Suc n, g) \<plusminus> g (Suc n) \<cdot>\<^sub>r X^\<^bsup>R (Suc n)\<^esup>")
apply (cut_tac c = "(Suc n, f)" and k = n in polyn_mem, assumption,
simp del:npow_suc,
cut_tac k = n and c = "(Suc n, g)" in polyn_mem, assumption,
simp del:npow_suc)
apply (frule_tac j = "Suc n" and c = "(Suc n, f)" in pol_coeff_mem_R, simp,
frule_tac j = "Suc n" and c = "(Suc n, g)" in pol_coeff_mem_R, simp,
cut_tac X_mem_R,
frule_tac n = "Suc n" in npClose[of "X"], simp del:npow_suc)
apply (frule_tac x = "f (Suc n)" and y = "X^\<^bsup>R (Suc n)\<^esup>" in ring_tOp_closed,
assumption+,
frule_tac x = "g (Suc n)" and y = "X^\<^bsup>R (Suc n)\<^esup>" in ring_tOp_closed,
assumption+)
apply (cut_tac ring_is_ag,
subst aGroup.pOp_assocTr43, assumption+)
apply (frule_tac x = "f (Suc n) \<cdot>\<^sub>r X^\<^bsup>R (Suc n)\<^esup>" and
y = "polyn_expr R X n (Suc n, g)" in aGroup.ag_pOp_commute[of R],
assumption+, simp del:npow_suc,
thin_tac "f (Suc n) \<cdot>\<^sub>r X^\<^bsup>R (Suc n)\<^esup> \<plusminus> polyn_expr R X n (Suc n, g) =
polyn_expr R X n (Suc n, g) \<plusminus> f (Suc n) \<cdot>\<^sub>r X^\<^bsup>R (Suc n)\<^esup>")
apply (subst aGroup.pOp_assocTr43[THEN sym], assumption+,
simp del:npow_suc add:polyn_expr_restrict)
apply (frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem, simp,
frule_tac c = "(Suc n, g)" and j = "Suc n" in pol_coeff_mem, simp)
apply (subst ring_distrib2[THEN sym], assumption+)
apply (frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem, simp,
frule_tac c = "(Suc n, g)" and j = "Suc n" in pol_coeff_mem, simp)
apply (frule_tac a = "f (Suc n)" and b = "g (Suc n)" in
Subring_pOp_ring_pOp[of S], simp, simp)
apply simp
done
lemma (in PolynRg) polyn_add_n:"\<lbrakk>pol_coeff S (n, f); pol_coeff S (n, g)\<rbrakk> \<Longrightarrow>
(polyn_expr R X n (n, f)) \<plusminus> (polyn_expr R X n (n, g)) =
nsum R (\<lambda>j. ((f j) \<plusminus>\<^bsub>S\<^esub> (g j)) \<cdot>\<^sub>r (X^\<^bsup>R j\<^esup>)) n"
by (simp add:polyn_addTr)
definition
add_cf :: "[('a, 'm) Ring_scheme, nat \<times> (nat \<Rightarrow> 'a), nat \<times> (nat \<Rightarrow> 'a)] \<Rightarrow>
nat \<times> (nat \<Rightarrow> 'a)" where
"add_cf S c d =
(if (fst c) < (fst d) then ((fst d), \<lambda>j. (if j \<le> (fst c)
then (((snd c) j) \<plusminus>\<^bsub>S\<^esub> ((snd d) j)) else ((snd d) j)))
else if (fst c) = (fst d) then ((fst c), \<lambda>j. ((snd c) j \<plusminus>\<^bsub>S\<^esub> (snd d) j))
else ((fst c), \<lambda>j. (if j \<le> (fst d) then
((snd c) j \<plusminus>\<^bsub>S\<^esub> (snd d) j) else ((snd c) j))))"
lemma (in PolynRg) add_cf_pol_coeff:"\<lbrakk>pol_coeff S c; pol_coeff S d\<rbrakk>
\<Longrightarrow> pol_coeff S (add_cf S c d)"
apply (cut_tac subring,
frule subring_Ring[of S], frule Ring.ring_is_ag[of S])
apply (simp add:pol_coeff_def)
apply (rule allI, rule impI)
apply (case_tac "(fst c) < (fst d)", simp add:add_cf_def)
apply (rule impI, rule aGroup.ag_pOp_closed, assumption+, simp+)
apply (drule leI[of "fst c" "fst d"],
drule le_imp_less_or_eq[of "fst d" "fst c"])
apply (erule disjE)
apply (simp add:add_cf_def, rule impI)
apply (frule Ring.ring_is_ag[of S], rule aGroup.ag_pOp_closed, assumption,
simp+)
apply (simp add:add_cf_def)
apply (frule Ring.ring_is_ag[of S], rule aGroup.ag_pOp_closed, assumption,
simp+)
done
lemma (in PolynRg) add_cf_len:"\<lbrakk>pol_coeff S c; pol_coeff S d\<rbrakk>
\<Longrightarrow> fst (add_cf S c d) = (max (fst c) (fst d))"
by (simp add: add_cf_def max.absorb1 max.absorb2)
lemma (in PolynRg) polyn_expr_restrict1:"\<lbrakk>pol_coeff S (n, f);
pol_coeff S (Suc (m + n), g)\<rbrakk> \<Longrightarrow>
polyn_expr R X (m + n) (add_cf S (n, f) (m + n, g)) =
polyn_expr R X (m + n) (m + n, snd (add_cf S (n, f) (Suc (m + n), g)))"
apply (frule pol_coeff_pre[of "m+n" g])
apply (frule add_cf_pol_coeff[of "(n, f)" "(Suc (m + n), g)"], assumption+,
frule add_cf_pol_coeff[of "(n, f)" "(m + n, g)"], assumption+)
apply (rule polyn_exprs_eq[of "add_cf S (n, f) (m + n, g)"
"(m + n, snd (add_cf S (n, f) (Suc (m + n), g)))" "m + n"], assumption+)
apply (rule split_pol_coeff[of "add_cf S (n, f) (Suc (m + n), g)" "m + n"],
assumption, simp add:add_cf_len)
apply (simp add:add_cf_len)
apply (rule allI, rule impI)
apply (simp add:add_cf_def)
done
lemma (in PolynRg) polyn_add_n1:"\<lbrakk>pol_coeff S (n, f); pol_coeff S (n, g)\<rbrakk> \<Longrightarrow>
(polyn_expr R X n (n, f)) \<plusminus> (polyn_expr R X n (n, g)) =
polyn_expr R X n (add_cf S (n, f) (n, g))"
apply (subst polyn_add_n, assumption+)
apply (simp add:polyn_expr_def add_cf_def)
done
lemma (in PolynRg) add_cf_val_hi:"(fst c) < (fst d) \<Longrightarrow>
snd (add_cf S c d) (fst d) = (snd d) (fst d)"
by (simp add:add_cf_def)
lemma (in PolynRg) add_cf_commute:"\<lbrakk>pol_coeff S c; pol_coeff S d\<rbrakk>
\<Longrightarrow> \<forall>j \<le> (max (fst c) (fst d)). snd (add_cf S c d) j =
snd (add_cf S d c) j"
apply (cut_tac subring, frule subring_Ring,
frule Ring.ring_is_ag[of S])
apply (simp add: add_cf_def max.absorb1 max.absorb2)
apply (case_tac "(fst c) = (fst d)", simp add: pol_coeff_def)
apply (rule allI, rule impI,
rule aGroup.ag_pOp_commute[of S], simp+)
apply (case_tac "(fst d) < (fst c)", simp,
rule allI, rule impI,
rule aGroup.ag_pOp_commute, assumption+)
apply (frule_tac x = j and y = "fst d" and z = "fst c" in le_less_trans,
assumption+, frule_tac x = j and y = "fst c" in less_imp_le,
thin_tac "j < fst c", simp add:pol_coeff_mem, simp add:pol_coeff_mem)
apply simp
apply (frule leI[of "fst d" "fst c"],
frule noteq_le_less[of "fst c" "fst d"], assumption,
rule allI, rule impI,
simp)
apply (rule aGroup.ag_pOp_commute, assumption+,
simp add:pol_coeff_mem,
frule_tac x = j and y = "fst c" and z = "fst d" in le_less_trans,
assumption+, frule_tac x = j and y = "fst d" in less_imp_le,
thin_tac "j < fst d", simp add:pol_coeff_mem)
done
lemma (in PolynRg) polyn_addTr1:"pol_coeff S (n, f) \<Longrightarrow>
\<forall>g. pol_coeff S (n + m, g) \<longrightarrow>
(polyn_expr R X n (n, f) \<plusminus> (polyn_expr R X (n + m) ((n + m), g))
= polyn_expr R X (n + m) (add_cf S (n, f) ((n + m), g)))"
apply (cut_tac subring, frule subring_Ring)
apply (induct_tac m)
apply (rule allI, rule impI, simp)
apply (simp add:polyn_add_n1)
apply (simp add:add.commute[of n])
apply (rule allI, rule impI)
apply (frule_tac n = "na + n" and f = g in pol_coeff_pre)
apply (drule_tac a = g in forall_spec, assumption)
apply (cut_tac n = "na + n" and c = "(Suc (na + n), g)" in polyn_Suc,
simp, simp del:npow_suc,
thin_tac "polyn_expr R X (Suc (na + n)) (Suc (na + n), g) =
polyn_expr R X (na + n) (Suc (na + n), g) \<plusminus>
g (Suc (na + n)) \<cdot>\<^sub>r X^\<^bsup>R (Suc (na + n))\<^esup>")
apply (frule_tac c = "(n, f)" and k = n in polyn_mem, simp,
frule_tac c = "(Suc (na + n), g)" and k = "na + n" in polyn_mem, simp,
frule_tac c = "(Suc (na + n), g)" in monomial_mem)
apply (drule_tac a = "Suc (na + n)" in forall_spec, simp del:npow_suc,
cut_tac ring_is_ag,
subst aGroup.ag_pOp_assoc[THEN sym], assumption+, simp del:npow_suc)
apply (simp del:npow_suc add:polyn_expr_restrict)
apply (frule_tac c = "(n, f)" and d = "(Suc (na + n), g)" in
add_cf_pol_coeff, assumption+,
frule_tac c = "(n, f)" and d = "(na + n, g)" in
add_cf_pol_coeff, assumption+)
apply (frule_tac c = "add_cf S (n, f) (Suc (na + n), g)" and
n = "na + n" and m = "Suc (na + n)" in polyn_n_m, simp,
subst add_cf_len, assumption+, simp)
apply (cut_tac k = "Suc (na + n)" and f = "add_cf S (n, f) (Suc (na + n), g)"
in polyn_expr_split)
apply (frule_tac c = "(n, f)" and d = "(Suc (na + n), g)" in
add_cf_len, assumption+, simp del: npow_suc add: max.absorb1 max.absorb2)
apply (thin_tac "polyn_expr R X (Suc (na + n))
(Suc (na + n), snd (add_cf S (n, f) (Suc (na + n), g))) =
polyn_expr R X (na + n)
(na + n, snd (add_cf S (n, f) (Suc (na + n), g))) \<plusminus>
\<Sigma>\<^sub>f R (\<lambda>j. snd (add_cf S (n, f) (Suc (na + n), g)) j \<cdot>\<^sub>r
X^\<^bsup>R j\<^esup>) (Suc (na + n)) (Suc (na + n))",
thin_tac "polyn_expr R X (Suc (na + n)) (add_cf S (n, f) (Suc (na + n),
g)) =
polyn_expr R X (na + n)
(na + n, snd (add_cf S (n, f) (Suc (na + n), g))) \<plusminus>
\<Sigma>\<^sub>f R (\<lambda>j. snd (add_cf S (n, f) (Suc (na + n), g)) j \<cdot>\<^sub>r
X^\<^bsup>R j\<^esup>) (Suc (na + n)) (Suc (na + n))")
apply (simp del:npow_suc add:fSum_def cmp_def slide_def)
apply (cut_tac d = "(Suc (na + n), g)" in add_cf_val_hi[of "(n, f)"],
simp, simp del:npow_suc,
thin_tac "snd (add_cf S (n, f) (Suc (na + n), g)) (Suc (na + n)) =
g (Suc (na + n))")
apply (frule_tac c = "add_cf S (n, f) (Suc (na + n), g)" and k = "na + n" in
polyn_mem, simp,
frule_tac c = "add_cf S (n, f) (na + n, g)" and k = "na + n" in
polyn_mem, simp )
apply (subst add_cf_len, assumption+, simp del:npow_suc)
apply (frule_tac a = "polyn_expr R X (na + n) (add_cf S (n, f) (na + n, g))"
and b = "polyn_expr R X (na + n) (add_cf S (n, f) (Suc (na + n), g))"
and c = "g (Suc (na + n)) \<cdot>\<^sub>r X^\<^bsup>R (Suc (na + n))\<^esup>" in
aGroup.ag_pOp_add_r[of R], assumption+)
apply (rule_tac c = "add_cf S (n, f) (na + n, g)" and
d = "add_cf S (n, f) (Suc (na + n), g)" and k = "na + n" in
polyn_exprs_eq, assumption+, simp,
subst add_cf_len, assumption+)
apply (simp)
apply (rule allI, rule impI,
(subst add_cf_def)+, simp,
frule_tac m = na and g = g in polyn_expr_restrict1[of n f], assumption,
simp del:npow_suc)
done
lemma (in PolynRg) polyn_add:"\<lbrakk>pol_coeff S (n, f); pol_coeff S (m, g)\<rbrakk>
\<Longrightarrow> polyn_expr R X n (n, f) \<plusminus> (polyn_expr R X m (m, g))
= polyn_expr R X (max n m) (add_cf S (n, f) (m, g))"
apply (cut_tac less_linear[of n m])
apply (erule disjE,
frule polyn_addTr1[of n f "m - n"],
drule_tac a = g in forall_spec, simp, simp add: max.absorb1 max.absorb2)
apply (erule disjE,
simp add:polyn_add_n1)
apply (frule polyn_mem[of "(n, f)" n], simp,
frule polyn_mem[of "(m, g)" m], simp)
apply (cut_tac ring_is_ag, simp add:aGroup.ag_pOp_commute)
apply (frule polyn_addTr1[of m g "n - m"],
drule_tac a = f in forall_spec, simp, simp,
frule add_cf_commute[of "(m, g)" "(n, f)"], assumption+,
simp add:max_def,
frule add_cf_pol_coeff[of "(n, f)" "(m, g)"], assumption+,
frule add_cf_pol_coeff[of "(m, g)" "(n, f)"], assumption+)
apply (rule polyn_exprs_eq[of "add_cf S (m, g) (n, f)"
"add_cf S (n, f) (m, g)" n], assumption+)
apply (simp add:add_cf_len, simp)
done
lemma (in PolynRg) polyn_add1:"\<lbrakk>pol_coeff S c; pol_coeff S d\<rbrakk>
\<Longrightarrow> polyn_expr R X (fst c) c \<plusminus> (polyn_expr R X (fst d) d)
= polyn_expr R X (max (fst c) (fst d)) (add_cf S c d)"
apply (cases c)
apply (cases d)
apply (simp add: polyn_add)
done
lemma (in PolynRg) polyn_minus_nsum:"\<lbrakk>pol_coeff S c; k \<le> (fst c)\<rbrakk> \<Longrightarrow>
-\<^sub>a (polyn_expr R X k c) = nsum R (\<lambda>j. ((-\<^sub>a\<^bsub>S\<^esub> ((snd c) j)) \<cdot>\<^sub>r (X^\<^bsup>R j\<^esup>))) k"
apply (cut_tac subring,
frule subring_Ring[of S],
frule Ring.ring_is_ag[of S],
cut_tac ring_is_ag,
cut_tac X_mem_R)
apply (simp add:polyn_expr_def,
subst aGroup.nsum_minus[of R], assumption)
apply (frule monomial_mem[of c], rule allI, rule impI,
frule_tac i = j and j = k and k = "fst c" in le_trans, assumption+,
simp)
apply (rule aGroup.nsum_eq, assumption,
rule allI, rule impI, simp,
rule aGroup.ag_mOp_closed, assumption)
apply (frule monomial_mem[of c],
frule_tac i = j and j = k and k = "fst c" in le_trans, assumption+,
simp)
apply (rule allI, rule impI,
rule ring_tOp_closed)
apply (frule_tac j = j in pol_coeff_mem[of c])
apply (frule_tac i = j and j = k and k = "fst c" in le_trans, assumption+,
simp add:Subring_minus_ring_minus,
frule_tac x = "(snd c) j" in mem_subring_mem_ring[of S], assumption,
simp add:aGroup.ag_mOp_closed,
simp add:npClose)
apply (rule allI, rule impI, simp,
cut_tac j = j in pol_coeff_mem[of c], assumption,
rule_tac i = j and j = k and k = "fst c" in le_trans, assumption+)
apply (simp add:Subring_minus_ring_minus,
frule_tac x = "(snd c) j" in mem_subring_mem_ring[of S], assumption)
apply (subst ring_inv1_1, assumption+)
apply (simp add:npClose, simp)
done
lemma (in PolynRg) minus_pol_coeff:"pol_coeff S c \<Longrightarrow>
pol_coeff S ((fst c), (\<lambda>j. (-\<^sub>a\<^bsub>S\<^esub> ((snd c) j))))"
apply (simp add:pol_coeff_def)
apply (rule allI, rule impI)
apply (cut_tac subring, frule subring_Ring)
apply (frule Ring.ring_is_ag[of "S"])
apply (rule aGroup.ag_mOp_closed, assumption)
apply simp
done
lemma (in PolynRg) polyn_minus:"\<lbrakk>pol_coeff S c; k \<le> (fst c)\<rbrakk> \<Longrightarrow>
-\<^sub>a (polyn_expr R X k c) =
polyn_expr R X k (fst c, (\<lambda>j. (-\<^sub>a\<^bsub>S\<^esub> ((snd c) j))))"
apply (cases c)
apply (subst polyn_minus_nsum)
apply (simp_all add: polyn_expr_def)
done
definition
m_cf :: "[('a, 'm) Ring_scheme, nat \<times> (nat \<Rightarrow> 'a)] \<Rightarrow> nat \<times> (nat \<Rightarrow> 'a)" where
"m_cf S c = (fst c, (\<lambda>j. (-\<^sub>a\<^bsub>S\<^esub> ((snd c) j))))"
lemma (in PolynRg) m_cf_pol_coeff:"pol_coeff S c \<Longrightarrow>
pol_coeff S (m_cf S c)"
by (simp add:m_cf_def, simp add:minus_pol_coeff)
lemma (in PolynRg) m_cf_len:"pol_coeff S c \<Longrightarrow>
fst (m_cf S c) = fst c"
by (simp add:m_cf_def)
lemma (in PolynRg) polyn_minus_m_cf:"\<lbrakk>pol_coeff S c; k \<le> (fst c)\<rbrakk> \<Longrightarrow>
-\<^sub>a (polyn_expr R X k c) =
polyn_expr R X k (m_cf S c)"
by (simp add:m_cf_def polyn_minus)
lemma (in PolynRg) polyn_zero_minus_zero:"\<lbrakk>pol_coeff S c; k \<le> (fst c)\<rbrakk> \<Longrightarrow>
(polyn_expr R X k c = \<zero>) = (polyn_expr R X k (m_cf S c) = \<zero>)"
apply (cut_tac ring_is_ag)
apply (simp add:polyn_minus_m_cf[THEN sym])
apply (rule iffI, simp)
apply (simp add:aGroup.ag_inv_zero)
apply (frule polyn_mem[of c k], assumption)
apply (frule aGroup.ag_inv_inv[of "R" "polyn_expr R X k c"], assumption)
apply (simp add:aGroup.ag_inv_zero)
done
lemma (in PolynRg) coeff_0_pol_0:"\<lbrakk>pol_coeff S c; k \<le> fst c\<rbrakk> \<Longrightarrow>
(\<forall>j\<le> k. (snd c) j = \<zero>\<^bsub>S\<^esub>) = (polyn_expr R X k c = \<zero>)"
apply (rule iffI)
apply (cut_tac ring_is_ag, cut_tac subring,
frule subring_Ring)
apply (simp add:Subring_zero_ring_zero)
apply (simp add:polyn_expr_def,
rule aGroup.nsum_zeroA[of R], assumption)
apply (rule allI, rule impI,
cut_tac X_mem_R)
apply (drule_tac a = j in forall_spec, simp,
frule_tac n = j in npClose[of X], simp)
apply (simp add:ring_times_0_x)
apply (cases c)
using algfree [simplified algfree_cond_def] by (auto simp add: polyn_expr_def)
subsection \<open>Multiplication of \<open>pol_exprs\<close>\<close>
subsection "Multiplication"
definition
ext_cf :: "[('a, 'm) Ring_scheme, nat, nat \<times> (nat \<Rightarrow> 'a)] \<Rightarrow>
nat \<times> (nat \<Rightarrow> 'a)" where
"ext_cf S n c = (n + fst c, \<lambda>i. if n \<le> i then (snd c) (sliden n i) else \<zero>\<^bsub>S\<^esub>)"
(* 0 0 g(0) g(m)
0 n m+n , where (m, g) is a pol_coeff **)
definition
sp_cf :: "[('a, 'm) Ring_scheme, 'a, nat \<times> (nat \<Rightarrow> 'a)] \<Rightarrow> nat \<times> (nat \<Rightarrow> 'a)" where
"sp_cf S a c = (fst c, \<lambda>j. a \<cdot>\<^sub>r\<^bsub>S\<^esub> ((snd c) j))" (* scalar times cf *)
definition
special_cf :: "('a, 'm) Ring_scheme \<Rightarrow> nat \<times> (nat \<Rightarrow> 'a)" ("C\<^sub>0") where
"C\<^sub>0 S = (0, \<lambda>j. 1\<^sub>r\<^bsub>S\<^esub>)"
lemma (in PolynRg) special_cf_pol_coeff:"pol_coeff S (C\<^sub>0 S)"
apply (cut_tac subring, frule subring_Ring)
apply (simp add:pol_coeff_def special_cf_def)
apply (simp add:Ring.ring_one)
done
lemma (in PolynRg) special_cf_len:"fst (C\<^sub>0 S) = 0"
apply (simp add:special_cf_def)
done
lemma (in PolynRg) ext_cf_pol_coeff:"pol_coeff S c \<Longrightarrow>
pol_coeff S (ext_cf S n c)"
apply (simp add: pol_coeff_def ext_cf_def sliden_def)
apply (rule impI)
apply (rule Ring.ring_zero)
apply (rule subring_Ring)
apply (rule subring)
done
lemma (in PolynRg) ext_cf_len:"pol_coeff S c \<Longrightarrow>
fst (ext_cf S m c) = m + fst c"
by (simp add:ext_cf_def)
lemma (in PolynRg) ext_special_cf_len:"fst (ext_cf S m (C\<^sub>0 S)) = m"
apply (cut_tac special_cf_pol_coeff)
apply (simp add:ext_cf_len special_cf_def)
done
lemma (in PolynRg) ext_cf_self:"pol_coeff S c \<Longrightarrow>
\<forall>j \<le> (fst c). snd (ext_cf S 0 c) j = (snd c) j"
apply (rule allI, rule impI, simp add:ext_cf_def sliden_def)
done
lemma (in PolynRg) ext_cf_hi:"pol_coeff S c \<Longrightarrow>
(snd c) (fst c) =
snd (ext_cf S n c) (n + (fst c))"
apply (subst ext_cf_def)
apply (simp add:sliden_def)
done
lemma (in PolynRg) ext_special_cf_hi:"snd (ext_cf S n (C\<^sub>0 S)) n = 1\<^sub>r\<^bsub>S\<^esub>"
apply (cut_tac special_cf_pol_coeff)
apply (cut_tac ext_cf_hi[of "C\<^sub>0 S" n, THEN sym])
apply (simp add:special_cf_def, assumption)
done
lemma (in PolynRg) ext_cf_lo_zero:"\<lbrakk>pol_coeff S c; 0 < n; x \<le> (n - Suc 0)\<rbrakk>
\<Longrightarrow> snd (ext_cf S n c) x = \<zero>\<^bsub>S\<^esub>"
apply (cut_tac Suc_le_mono[THEN sym, of x "n - Suc 0"], simp,
cut_tac x = x and y = "Suc x" and z = n in less_le_trans,
simp, assumption,
simp add:nat_not_le_less[THEN sym, of x n],
thin_tac "x \<le> n - Suc 0")
apply (simp add:ext_cf_def)
done
lemma (in PolynRg) ext_special_cf_lo_zero:"\<lbrakk>0 < n; x \<le> (n - Suc 0)\<rbrakk>
\<Longrightarrow> snd (ext_cf S n (C\<^sub>0 S)) x = \<zero>\<^bsub>S\<^esub>"
by (cut_tac special_cf_pol_coeff,
frule ext_cf_lo_zero[of "C\<^sub>0 S" n], assumption+)
lemma (in PolynRg) sp_cf_pol_coeff:"\<lbrakk>pol_coeff S c; a \<in> carrier S\<rbrakk> \<Longrightarrow>
pol_coeff S (sp_cf S a c)"
apply (cut_tac subring, frule subring_Ring)
apply (simp add:pol_coeff_def sp_cf_def,
rule allI, rule impI,
rule Ring.ring_tOp_closed, assumption+)
apply simp
done
lemma (in PolynRg) sp_cf_len:"\<lbrakk>pol_coeff S c; a \<in> carrier S\<rbrakk> \<Longrightarrow>
fst (sp_cf S a c) = fst c"
by (simp add:sp_cf_def)
lemma (in PolynRg) sp_cf_val:"\<lbrakk>pol_coeff S c; j \<le> (fst c); a \<in> carrier S\<rbrakk> \<Longrightarrow>
snd (sp_cf S a c) j = a \<cdot>\<^sub>r\<^bsub>S\<^esub> ((snd c) j)"
by (simp add:sp_cf_def)
lemma (in PolynRg) polyn_ext_cf_lo_zero:"\<lbrakk>pol_coeff S c; 0 < j\<rbrakk> \<Longrightarrow>
polyn_expr R X (j - Suc 0) (ext_cf S j c) = \<zero>"
apply (simp add:polyn_expr_def, cut_tac ring_is_ag,
rule aGroup.nsum_zeroA, assumption)
apply (rule allI, rule impI)
apply (frule_tac x = ja in ext_cf_lo_zero [of c j], assumption+)
apply (cut_tac X_mem_R, frule_tac n = ja in npClose[of X])
apply (cut_tac subring,
simp add:Subring_zero_ring_zero,
simp add:ring_times_0_x)
done
lemma (in PolynRg) monomial_d:"pol_coeff S c \<Longrightarrow>
polyn_expr R X d (ext_cf S d c) = ((snd c) 0) \<cdot>\<^sub>r X^\<^bsup>R d\<^esup>"
apply (cut_tac ring_is_ag,
cut_tac subring,
cut_tac X_mem_R,
frule subring_Ring[of S])
apply (frule pol_coeff_mem [of c 0], simp)
apply (case_tac "d = 0")
apply simp
apply (simp add:polyn_expr_def ext_cf_def sliden_def)
apply (frule ext_cf_pol_coeff[of c d])
apply (cut_tac polyn_Suc[of "d - Suc 0" "ext_cf S d c"])
apply (simp)
apply (cut_tac polyn_ext_cf_lo_zero[of c d], simp,
thin_tac "polyn_expr R X (d - Suc 0) (ext_cf S d c) = \<zero>")
apply (frule monomial_mem[of "ext_cf S d c"], simp add:ext_cf_len,
drule_tac a = d in forall_spec, simp, simp add:aGroup.ag_l_zero)
apply (subst polyn_expr_short[of "ext_cf S d c" d], assumption,
simp add:ext_cf_len)
apply (simp,
subst ext_cf_def, simp add:sliden_def , assumption+,
simp add:ext_cf_len)
done
lemma (in PolynRg) X_to_d:" X^\<^bsup>R d\<^esup> = polyn_expr R X d (ext_cf S d (C\<^sub>0 S))"
apply (cut_tac special_cf_pol_coeff)
apply (subst monomial_d[of "C\<^sub>0 S" d], assumption+)
apply (subst special_cf_def, simp)
apply (cut_tac subring, frule subring_Ring)
apply (simp add:Subring_one_ring_one)
apply (cut_tac X_mem_R, frule_tac n = d in npClose[of X])
apply (simp add:ring_l_one)
done
lemma (in PolynRg) c_max_ext_special_cf:"c_max S (ext_cf S n (C\<^sub>0 S)) = n"
apply (cut_tac polyn_ring_S_nonzero,
cut_tac subring, frule subring_Ring)
apply (simp add:c_max_def special_cf_def ext_cf_def)
apply (cut_tac n_max[of "{j. (n \<le> j \<longrightarrow> j = n) \<and> n \<le> j}" n])
apply (erule conjE)+ apply simp
apply (rule subsetI, simp, erule conjE, simp)
apply (cut_tac le_refl[of n], blast)
done
lemma (in PolynRg) scalar_times_polynTr:"a \<in> carrier S \<Longrightarrow>
\<forall>f. pol_coeff S (n, f) \<longrightarrow>
a \<cdot>\<^sub>r (polyn_expr R X n (n, f)) = polyn_expr R X n (sp_cf S a (n, f))"
apply (cut_tac subring,
cut_tac X_mem_R,
frule_tac x = a in mem_subring_mem_ring, assumption)
apply (induct_tac n,
rule allI, rule impI, simp add:polyn_expr_def sp_cf_def,
cut_tac n_in_Nsetn[of "0"])
apply (cut_tac subring_Ring,
frule_tac c = "(0, f)" in pol_coeff_mem[of _ "0"], simp)
apply (simp,
frule_tac x = "f 0" in mem_subring_mem_ring, assumption)
apply ( simp add:Subring_tOp_ring_tOp,
frule_tac y = "f 0" in ring_tOp_closed[of a], assumption+,
cut_tac ring_one, simp add:ring_tOp_assoc, assumption)
apply (rule allI, rule impI,
frule subring_Ring,
frule_tac n = n and f = f in pol_coeff_pre,
drule_tac x = f in spec, simp)
apply (cut_tac n = n and c = "(Suc n, f)" in polyn_Suc, simp,
simp del:npow_suc,
thin_tac "polyn_expr R X (Suc n) (Suc n, f) =
polyn_expr R X n (Suc n, f) \<plusminus> f (Suc n) \<cdot>\<^sub>r X^\<^bsup>R (Suc n)\<^esup>")
apply (cut_tac n = n and c = "sp_cf S a (Suc n, f)" in polyn_Suc,
simp add:sp_cf_len)
apply (frule_tac c = "(Suc n, f)" and a = a in sp_cf_len, assumption+,
simp only:fst_conv)
apply (cut_tac k = "Suc n" and f = "sp_cf S a (Suc n, f)" in
polyn_expr_split, simp del:npow_suc,
thin_tac "polyn_expr R X (Suc n) (Suc n, snd (sp_cf S a (Suc n, f))) =
polyn_expr R X n (sp_cf S a (Suc n, f)) \<plusminus>
snd (sp_cf S a (Suc n, f)) (Suc n) \<cdot>\<^sub>r X^\<^bsup>R (Suc n)\<^esup>",
thin_tac "polyn_expr R X (Suc n) (sp_cf S a (Suc n, f)) =
polyn_expr R X n (sp_cf S a (Suc n, f)) \<plusminus>
snd (sp_cf S a (Suc n, f)) (Suc n) \<cdot>\<^sub>r X^\<^bsup>R (Suc n)\<^esup>")
apply (frule_tac c = "(Suc n, f)" and a = a in sp_cf_pol_coeff, assumption)
apply (frule_tac c = "(Suc n, f)" and k = n in polyn_mem,
simp,
frule_tac c = "(Suc n, f)" in monomial_mem,
drule_tac a = "Suc n" in forall_spec, simp,
simp only:snd_conv)
apply (subst ring_distrib1, assumption+,
subst polyn_expr_restrict, assumption+, simp del:npow_suc,
subst sp_cf_val, assumption, simp, assumption,
simp only:snd_conv,
frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem,
simp, simp only:snd_conv,
simp del:npow_suc add:Subring_tOp_ring_tOp,
subst ring_tOp_assoc[THEN sym, of a], assumption+,
simp add:mem_subring_mem_ring, rule npClose, assumption)
apply (cut_tac ring_is_ag,
rule aGroup.ag_pOp_add_r, assumption+,
rule polyn_mem, rule sp_cf_pol_coeff, assumption+,
simp add:sp_cf_len,
rule polyn_mem, assumption, simp add:sp_cf_len,
frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem_R,
simp, simp only:snd_conv,
(rule ring_tOp_closed)+, assumption+, rule npClose, assumption)
apply (rule_tac c = "sp_cf S a (n, f)" and d = "sp_cf S a (Suc n, f)" and
k = n in polyn_exprs_eq, rule sp_cf_pol_coeff, assumption+,
simp add:sp_cf_len)
apply (rule allI, rule impI,
(subst sp_cf_def)+, simp)
done
lemma (in PolynRg) scalar_times_pol_expr:"\<lbrakk>a \<in> carrier S; pol_coeff S c;
n \<le> fst c\<rbrakk> \<Longrightarrow>
a \<cdot>\<^sub>r (polyn_expr R X n c) = polyn_expr R X n (sp_cf S a c)"
apply (cases c) apply (simp only:)
apply (rename_tac m g)
apply (thin_tac "c = (m, g)")
apply (frule_tac c = "(m, g)" and k = n in polyn_expr_short, simp,
simp)
apply (frule scalar_times_polynTr[of a n],
drule_tac x = g in spec)
apply (frule_tac c = "(m, g)" and n = n in pol_coeff_le, simp, simp,
thin_tac "polyn_expr R X n (m, g) = polyn_expr R X n (n, g)",
thin_tac "a \<cdot>\<^sub>r polyn_expr R X n (n, g) =
polyn_expr R X n (sp_cf S a (n, g))")
apply (frule_tac c = "(m, g)" and n = n in pol_coeff_le, simp, simp,
frule_tac c = "(n, g)" and a = a in sp_cf_pol_coeff, assumption,
frule_tac c = "(m, g)" and a = a in sp_cf_pol_coeff, assumption)
apply (rule_tac c = "sp_cf S a (n, g)" and d = "sp_cf S a (m, g)" and
k = n in polyn_exprs_eq, assumption+)
apply (simp add:sp_cf_len)
apply (rule allI, (subst sp_cf_def)+, simp)
done
lemma (in PolynRg) sp_coeff_nonzero:"\<lbrakk>Idomain S; a \<in> carrier S; a \<noteq> \<zero>\<^bsub>S\<^esub>;
pol_coeff S c; (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>; j \<le> (fst c)\<rbrakk> \<Longrightarrow>
snd (sp_cf S a c) j \<noteq> \<zero>\<^bsub>S\<^esub>"
apply (simp add:sp_cf_def)
apply (frule_tac y = "(snd c) j" in Idomain.idom_tOp_nonzeros[of S a],
assumption+,
simp add:pol_coeff_def, simp add:Pi_def, assumption+)
done
lemma (in PolynRg) ext_cf_inductTl:"pol_coeff S (Suc n, f) \<Longrightarrow>
polyn_expr R X (n + j) (ext_cf S j (Suc n, f)) =
polyn_expr R X (n + j) (ext_cf S j (n, f))"
apply (frule pol_coeff_pre[of n f],
frule ext_cf_pol_coeff[of "(Suc n, f)" j],
frule ext_cf_pol_coeff[of "(n, f)" j],
rule polyn_exprs_eq[of "ext_cf S j (Suc n, f)" "ext_cf S j (n, f)"
"n + j"], assumption+)
apply (simp add:ext_cf_len)
apply (rule allI, (subst ext_cf_def)+, simp add:sliden_def)
done
lemma (in PolynRg) low_deg_terms_zeroTr:"
pol_coeff S (n, f) \<longrightarrow>
polyn_expr R X (n + j) (ext_cf S j (n, f)) =
(X^\<^bsup>R j\<^esup>) \<cdot>\<^sub>r (polyn_expr R X n (n, f))"
apply (cut_tac ring_is_ag,
cut_tac X_mem_R, frule npClose[of "X" "j"])
apply (induct_tac n)
apply (rule impI, simp)
apply (case_tac "j = 0", simp add:ext_cf_def sliden_def polyn_expr_def)
apply (frule_tac c = "(0, f)" and j = 0 in pol_coeff_mem_R, simp, simp)
apply (simp add:ring_r_one ring_l_one)
apply (cut_tac polyn_Suc[of "j - Suc 0" "ext_cf S j (0, f)"],
simp del:npow_suc)
apply (frule ext_cf_len[of "(0, f)" j],
cut_tac polyn_expr_split[of j "ext_cf S j (0, f)"], simp,
thin_tac "polyn_expr R X j (ext_cf S j (0, f)) =
polyn_expr R X (j - Suc 0) (ext_cf S j (0, f)) \<plusminus>
snd (ext_cf S j (0, f)) j \<cdot>\<^sub>r X^\<^bsup>R j\<^esup>")
apply (simp add:polyn_ext_cf_lo_zero[of "(0, f)" j],
thin_tac "polyn_expr R X j (j, snd (ext_cf S j (0, f))) =
\<zero> \<plusminus> snd (ext_cf S j (0, f)) j \<cdot>\<^sub>r X^\<^bsup>R j\<^esup>",
frule ext_cf_hi[THEN sym, of "(0, f)" j], simp add:polyn_expr_def)
apply (frule_tac c = "(0, f)" and j = 0 in pol_coeff_mem_R, simp, simp)
apply (subst aGroup.ag_l_zero, assumption, simp add:ring_tOp_closed,
simp add:ring_r_one, subst ring_tOp_commute, assumption+, simp)
apply (simp add:ext_cf_len)
apply (rule impI,
cut_tac subring,
cut_tac subring_Ring[of S],
frule_tac n = n in pol_coeff_pre[of _ "f"])
apply simp
apply (subst polyn_expr_split)
apply (cut_tac n = "n + j" and c = "ext_cf S j (Suc n, f)" in polyn_Suc,
simp add:ext_cf_len)
apply (subst ext_cf_len, assumption+, simp del:npow_suc add:add.commute[of j],
thin_tac "polyn_expr R X (Suc (n + j))
(Suc (n + j), snd (ext_cf S j (Suc n, f))) =
polyn_expr R X (n + j) (ext_cf S j (Suc n, f)) \<plusminus>
snd (ext_cf S j (Suc n, f)) (Suc (n + j)) \<cdot>\<^sub>r X^\<^bsup>R (Suc (n + j))\<^esup>",
subst ext_cf_inductTl, assumption+, simp del:npow_suc,
thin_tac "polyn_expr R X (n + j) (ext_cf S j (n, f)) =
X^\<^bsup>R j\<^esup> \<cdot>\<^sub>r polyn_expr R X n (n, f)")
apply (cut_tac c1 = "(Suc n, f)" and n1 = j in ext_cf_hi[THEN sym],
assumption+,
simp del:npow_suc add:add.commute[of j])
apply (cut_tac n = n and c = "(Suc n, f)" in polyn_Suc, simp,
simp del:npow_suc)
apply (frule_tac c = "(Suc n, f)" and k = n in polyn_mem, simp,
frule_tac c = "(Suc n, f)" and j = "Suc n" in pol_coeff_mem_R, simp,
simp del:npow_suc,
frule_tac x = "f (Suc n)" and y = "X^\<^bsup>R (Suc n)\<^esup>" in ring_tOp_closed,
rule npClose, assumption,
subst ring_distrib1, assumption+)
apply (subst polyn_expr_restrict, assumption+)
apply (rule_tac a = "f (Suc n) \<cdot>\<^sub>r X^\<^bsup>R (Suc (n + j))\<^esup> " and
b = "X^\<^bsup>R j\<^esup> \<cdot>\<^sub>r (f (Suc n) \<cdot>\<^sub>r X^\<^bsup>R (Suc n)\<^esup>)" and
c = "X^\<^bsup>R j\<^esup> \<cdot>\<^sub>r polyn_expr R X n (n, f)" in aGroup.ag_pOp_add_l,
assumption+,
rule ring_tOp_closed, assumption+, rule npClose, assumption,
(rule ring_tOp_closed, assumption+)+,
simp add:polyn_mem,
frule_tac n = "Suc n" in npClose[of X],
subst ring_tOp_assoc[THEN sym], assumption+,
subst ring_tOp_commute[of "X^\<^bsup>R j\<^esup>"], assumption,
simp add:pol_coeff_mem,
subst ring_tOp_assoc, assumption+,
subst npMulDistr[of X], assumption, simp add:add.commute[of j])
apply simp
done
lemma (in PolynRg) low_deg_terms_zero:"pol_coeff S (n, f) \<Longrightarrow>
polyn_expr R X (n + j) (ext_cf S j (n, f)) =
(X^\<^bsup>R j\<^esup>) \<cdot>\<^sub>r (polyn_expr R X n (n, f))"
by (simp add:low_deg_terms_zeroTr)
lemma (in PolynRg) low_deg_terms_zero1:"pol_coeff S c \<Longrightarrow>
polyn_expr R X ((fst c) + j) (ext_cf S j c) =
(X^\<^bsup>R j\<^esup>) \<cdot>\<^sub>r (polyn_expr R X (fst c) c)"
by (cases c) (simp add: low_deg_terms_zeroTr)
lemma (in PolynRg) polyn_expr_tOpTr:"pol_coeff S (n, f) \<Longrightarrow>
\<forall>g. (pol_coeff S (m, g) \<longrightarrow> (\<exists>h. pol_coeff S ((n + m), h) \<and>
h (n + m) = (f n) \<cdot>\<^sub>r\<^bsub>S\<^esub> (g m) \<and>
(polyn_expr R X (n + m) (n + m, h) =
(polyn_expr R X n (n, f)) \<cdot>\<^sub>r (polyn_expr R X m (m, g)))))"
apply (cut_tac subring,
cut_tac X_mem_R,
frule subring_Ring[of S])
apply (induct_tac m)
apply (rule allI, rule impI, simp)
apply (simp add:polyn_expr_def [of R X 0])
apply (frule_tac c = "(0,g)" in pol_coeff_mem[of _ 0], simp, simp,
frule_tac c = "(0,g)" in pol_coeff_mem_R[of _ 0], simp, simp)
apply (simp add:ring_r_one,
frule_tac c = "(n, f)" and k = n in polyn_mem, simp,
simp only:ring_tOp_commute[of "polyn_expr R X n (n, f)"],
subst scalar_times_pol_expr, assumption+, simp)
apply (cut_tac f = "sp_cf S (g 0) (n, f)" in pol_coeff_split)
apply (simp add:sp_cf_len)
apply (cut_tac f = "sp_cf S (g 0) (n, f)" in polyn_expr_split[of n],
simp only:sp_cf_len, simp only:fst_conv,
frule_tac a = "g 0" in sp_cf_pol_coeff[of "(n, f)"], assumption+,
simp,
subgoal_tac "snd (sp_cf S (g 0) (n, f)) n = (f n) \<cdot>\<^sub>r\<^bsub>S\<^esub> (g 0)", blast)
apply (thin_tac "pol_coeff S (n, snd (sp_cf S (g 0) (n, f)))",
thin_tac "polyn_expr R X n (sp_cf S (g 0) (n, f)) =
polyn_expr R X n (n, snd (sp_cf S (g 0) (n, f)))",
thin_tac "pol_coeff S (sp_cf S (g 0) (n, f))")
apply (subst sp_cf_val[of "(n, f)" n], assumption+, simp, assumption, simp,
frule_tac c = "(n,f)" in pol_coeff_mem[of _ n], simp, simp,
simp add:Ring.ring_tOp_commute)
apply (rule allI, rule impI)
apply (frule_tac n = na and f = g in pol_coeff_pre,
drule_tac a = g in forall_spec, assumption+)
apply (erule exE, (erule conjE)+)
apply (cut_tac n = na and c = "(Suc na, g)" in polyn_Suc, (simp del:npow_suc)+,
thin_tac "polyn_expr R X (Suc na) (Suc na, g) =
polyn_expr R X na (Suc na, g) \<plusminus> g (Suc na) \<cdot>\<^sub>r X^\<^bsup>R (Suc na)\<^esup>",
subst polyn_expr_restrict, assumption)
apply (frule_tac c = "(n, f)" and k = n in polyn_mem,simp del:npow_suc,
frule_tac c = "(na, g)" and k = na in polyn_mem, simp del:npow_suc,
frule_tac c = "(Suc na, g)" in monomial_mem, simp del:npow_suc,
drule_tac a = "Suc na" in forall_spec, simp del:npow_suc)
apply (subst ring_distrib1, assumption+)
apply (rotate_tac 8, drule sym,
simp del:npow_suc)
apply (thin_tac "polyn_expr R X n (n, f) \<cdot>\<^sub>r polyn_expr R X na (na, g) =
polyn_expr R X (n + na) (n + na, h)")
apply (frule_tac c = "(Suc na, g)" and j ="Suc na" in pol_coeff_mem_R, simp,
simp del:npow_suc,
frule_tac c = "(Suc na, g)" and j ="Suc na" in pol_coeff_mem, simp,
simp del:npow_suc)
apply (subst ring_tOp_commute, assumption+,
subst ring_tOp_assoc, assumption+, rule npClose, assumption+,
subst low_deg_terms_zero[THEN sym], assumption+)
apply (frule_tac c = "(n, f)" and n = "Suc na" in ext_cf_pol_coeff)
apply (frule_tac c = "ext_cf S (Suc na) (n, f)" and a = "g (Suc na)" in
sp_cf_pol_coeff, assumption)
apply (subst scalar_times_pol_expr, assumption+,
simp add:ext_cf_len,
cut_tac k = "n + Suc na" and
f = "sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))" in
polyn_expr_split,
simp only:sp_cf_len,
thin_tac "polyn_expr R X (n + Suc na)
(sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))) =
polyn_expr R X (n + Suc na)
(fst (ext_cf S (Suc na) (n, f)),
snd (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))))",
simp only:ext_cf_len, simp only:fst_conv,
simp add:add.commute[of _ n])
apply (subst polyn_add, assumption+,
cut_tac f = "sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))" in
pol_coeff_split, simp only:sp_cf_len, simp only:ext_cf_len,
simp add:add.commute[of _ n], simp add: max_def,
frule_tac c = "sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))"
in pol_coeff_cartesian,
simp only:sp_cf_len, simp only:ext_cf_len,
simp add:add.commute[of _ n],
thin_tac "(Suc (n + na),
snd (sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))) =
sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))",
frule_tac c = "(n + na, h)" and
d = "sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))" in
add_cf_pol_coeff, assumption)
apply (cut_tac k = "Suc (n + na)" and f = "add_cf S (n + na, h)
(sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))" in polyn_expr_split,
simp only:mp,
thin_tac "polyn_expr R X (Suc (n + na))
(add_cf S (n + na, h)
(sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))) =
polyn_expr R X (Suc (n + na))
(fst (add_cf S (n + na, h)
(sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))),
snd (add_cf S (n + na, h)
(sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))))",
subst add_cf_len, assumption+,
simp add:sp_cf_len, simp add:ext_cf_len max_def,
cut_tac f = "add_cf S (n + na, h)
(sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f)))" in
pol_coeff_split,
simp only:add_cf_len,
simp only:sp_cf_len, simp add:ext_cf_len, simp add:max_def,
thin_tac "pol_coeff S
(add_cf S (n + na, h)
(sp_cf S (g (Suc na)) (ext_cf S (Suc na) (n, f))))",
subgoal_tac "snd (add_cf S (n + na, h) (sp_cf S (g (Suc na))
(ext_cf S (Suc na) (n, f)))) (Suc (n + na)) = f n \<cdot>\<^sub>r\<^bsub>S\<^esub> g (Suc na)",
simp add:add.commute[of _ n], blast)
apply (subst add_cf_def, simp add:sp_cf_len ext_cf_len,
subst sp_cf_def, simp add:ext_cf_len,
subst ext_cf_def, simp add:sliden_def,
frule pol_coeff_mem[of "(n, f)" n], simp,
simp add:Ring.ring_tOp_commute)
done
lemma (in PolynRg) polyn_expr_tOp:"\<lbrakk>
pol_coeff S (n, f); pol_coeff S (m, g)\<rbrakk> \<Longrightarrow> \<exists>e. pol_coeff S ((n + m), e) \<and>
e (n + m) = (f n) \<cdot>\<^sub>r\<^bsub>S\<^esub> (g m) \<and>
polyn_expr R X (n + m)(n + m, e) =
(polyn_expr R X n (n, f)) \<cdot>\<^sub>r (polyn_expr R X m (m, g))"
by (simp add:polyn_expr_tOpTr)
lemma (in PolynRg) polyn_expr_tOp_c:"\<lbrakk>pol_coeff S c; pol_coeff S d\<rbrakk> \<Longrightarrow>
\<exists>e. pol_coeff S e \<and> (fst e = fst c + fst d) \<and>
(snd e) (fst e) = (snd c (fst c)) \<cdot>\<^sub>r\<^bsub>S\<^esub> (snd d) (fst d) \<and>
polyn_expr R X (fst e) e =
(polyn_expr R X (fst c) c) \<cdot>\<^sub>r (polyn_expr R X (fst d) d)"
by (cases c, cases d) (simp add: polyn_expr_tOpTr)
section "The degree of a polynomial"
lemma (in PolynRg) polyn_degreeTr:"\<lbrakk>pol_coeff S c; k \<le> (fst c)\<rbrakk> \<Longrightarrow>
(polyn_expr R X k c = \<zero> ) = ({j. j \<le> k \<and> (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>} = {})"
apply (subst coeff_0_pol_0[THEN sym, of c k], assumption+)
apply blast
done
lemma (in PolynRg) higher_part_zero:"\<lbrakk>pol_coeff S c; k < fst c;
\<forall>j\<in>nset (Suc k) (fst c). snd c j = \<zero>\<^bsub>S\<^esub>\<rbrakk> \<Longrightarrow>
\<Sigma>\<^sub>f R (\<lambda>j. snd c j \<cdot>\<^sub>r X^\<^bsup>R j\<^esup>) (Suc k) (fst c) = \<zero>"
apply (cut_tac ring_is_ag,
rule aGroup.fSum_zero1[of R k "fst c" "\<lambda>j. snd c j \<cdot>\<^sub>r X^\<^bsup>R j\<^esup>"],
assumption+)
apply (rule ballI,
drule_tac x = j in bspec, assumption, simp)
apply (cut_tac subring,
simp add:Subring_zero_ring_zero,
cut_tac X_mem_R,
frule_tac n = j in npClose[of X],
simp add:ring_times_0_x)
done
lemma (in PolynRg) coeff_nonzero_polyn_nonzero:"\<lbrakk>pol_coeff S c; k \<le> (fst c)\<rbrakk>
\<Longrightarrow> (polyn_expr R X k c \<noteq> \<zero>) = (\<exists>j\<le>k. (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub> )"
by (simp add:coeff_0_pol_0[THEN sym, of c k])
lemma (in PolynRg) pol_expr_unique:"\<lbrakk>p \<in> carrier R; p \<noteq> \<zero>;
pol_coeff S c; p = polyn_expr R X (fst c) c; (snd c) (fst c) \<noteq> \<zero>\<^bsub>S\<^esub>;
pol_coeff S d; p = polyn_expr R X (fst d) d; (snd d) (fst d) \<noteq> \<zero>\<^bsub>S\<^esub>\<rbrakk> \<Longrightarrow>
(fst c) = (fst d) \<and> (\<forall>j \<le> (fst c). (snd c) j = (snd d) j)"
apply (cut_tac ring_is_ag,
cut_tac subring, frule subring_Ring,
frule Ring.ring_is_ag[of S])
apply (frule m_cf_pol_coeff[of d])
apply (frule polyn_minus_m_cf[of d "fst d"], simp)
apply (drule sym, drule sym, simp)
apply (rotate_tac -2, drule sym, drule sym)
apply (frule_tac x = p in aGroup.ag_r_inv1[of R], assumption, simp,
thin_tac "p = polyn_expr R X (fst c) c",
thin_tac "polyn_expr R X (fst d) d = polyn_expr R X (fst c) c",
thin_tac "-\<^sub>a (polyn_expr R X (fst c) c) =
polyn_expr R X (fst d) (m_cf S d)")
apply (frule polyn_add1[of c "m_cf S d"], assumption+, simp add:m_cf_len,
thin_tac "polyn_expr R X (fst c) c \<plusminus> polyn_expr R X (fst d) (m_cf S d)
= polyn_expr R X (max (fst c) (fst d)) (add_cf S c (m_cf S d))",
thin_tac "polyn_expr R X (fst c) c \<noteq>
polyn_expr R X (max (fst c) (fst d)) (add_cf S c (m_cf S d))")
apply (frule_tac add_cf_pol_coeff[of c "m_cf S d"], simp add:m_cf_len)
apply (cut_tac coeff_0_pol_0[THEN sym, of "add_cf S c (m_cf S d)"
"max (fst c) (fst d)"],
drule sym, simp,
thin_tac "polyn_expr R X (max (fst c) (fst d))
(add_cf S c (m_cf S d)) = \<zero>",
thin_tac "pol_coeff S (add_cf S c (m_cf S d))",
thin_tac "pol_coeff S (m_cf S d)")
apply (case_tac "fst c = fst d", simp)
apply (rule allI, rule impI,
drule_tac a = j in forall_spec, assumption)
apply (simp add:add_cf_def m_cf_def m_cf_len)
apply (frule_tac j = j in pol_coeff_mem[of c], simp,
frule_tac j = j in pol_coeff_mem[of d], simp)
apply (subst aGroup.ag_eq_diffzero[of S], assumption+)
apply (simp add:add_cf_def)
apply (case_tac "\<not> (fst c) \<le> (fst d)", simp)
apply (simp add:m_cf_len)
apply (drule_tac a = "fst c" in forall_spec, simp, simp)
apply simp
apply (drule_tac a = "fst d" in forall_spec, simp, simp add:m_cf_len)
apply (case_tac "fst c \<noteq> fst d",
frule noteq_le_less[of "fst c" "fst d"], assumption, simp)
apply (simp add:m_cf_def)
apply (frule pol_coeff_mem[of d "fst d"], simp)
apply (frule Ring.ring_is_ag[of S],
frule aGroup.ag_inv_inv[of S "snd d (fst d)"], assumption)
apply (simp add:aGroup.ag_inv_zero)
apply simp
apply simp
apply (simp add:add_cf_len m_cf_len)
done
lemma (in PolynRg) pol_expr_unique2:"\<lbrakk>pol_coeff S c; pol_coeff S d;
fst c = fst d\<rbrakk> \<Longrightarrow>
(polyn_expr R X (fst c) c = polyn_expr R X (fst d) d ) =
(\<forall>j \<le> (fst c). (snd c) j = (snd d) j)"
apply (cut_tac ring_is_ag,
cut_tac subring, frule subring_Ring,
frule Ring.ring_is_ag[of S])
apply (rule iffI)
apply (frule m_cf_pol_coeff[of d])
apply (frule polyn_mem[of c "fst c"], simp,
frule polyn_mem[of d "fst d"], simp)
apply (frule aGroup.ag_eq_diffzero[of R "polyn_expr R X (fst c) c"
"polyn_expr R X (fst d) d"], assumption+,
simp,
simp only:polyn_minus_m_cf[of d "fst d"],
drule sym, simp)
apply (frule polyn_add1[of c "m_cf S d"], assumption+, simp add:m_cf_len)
apply (thin_tac "polyn_expr R X (fst c) d \<plusminus> polyn_expr R X (fst c)
(m_cf S d) =
polyn_expr R X (fst c) (add_cf S c (m_cf S d))",
thin_tac "polyn_expr R X (fst c) c = polyn_expr R X (fst c) d",
thin_tac "polyn_expr R X (fst c) d \<in> carrier R",
drule sym)
apply (frule_tac add_cf_pol_coeff[of c "m_cf S d"], simp add:m_cf_len)
apply (frule coeff_0_pol_0[THEN sym, of "add_cf S c (m_cf S d)"
"fst c"],
simp add:add_cf_len, simp add:m_cf_len,
thin_tac "\<zero> = polyn_expr R X (fst d) (add_cf S c (m_cf S d))",
thin_tac "pol_coeff S (add_cf S c (m_cf S d))")
apply (simp add:add_cf_def m_cf_def)
apply (rule allI, rule impI)
apply (drule_tac a = j in forall_spec, assumption)
apply (frule_tac j = j in pol_coeff_mem[of c], simp,
frule_tac j = j in pol_coeff_mem[of d], simp)
apply (simp add:aGroup.ag_eq_diffzero[THEN sym])
apply simp
apply (rule polyn_exprs_eq[of c d "fst d"], assumption+)
apply (simp, assumption+)
done
lemma (in PolynRg) pol_expr_unique3:"\<lbrakk>pol_coeff S c; pol_coeff S d;
fst c < fst d\<rbrakk> \<Longrightarrow>
(polyn_expr R X (fst c) c = polyn_expr R X (fst d) d ) =
((\<forall>j \<le> (fst c). (snd c) j = (snd d) j) \<and>
(\<forall>j\<in>nset (Suc (fst c)) (fst d). (snd d) j = \<zero>\<^bsub>S\<^esub>))"
apply (rule iffI)
apply (cut_tac ring_is_ag,
cut_tac subring, frule subring_Ring,
frule Ring.ring_is_ag[of S])
apply (frule m_cf_pol_coeff[of d])
apply (frule polyn_mem[of c "fst c"], simp,
frule polyn_mem[of d "fst d"], simp)
apply (frule aGroup.ag_eq_diffzero[of R "polyn_expr R X (fst c) c"
"polyn_expr R X (fst d) d"], assumption+,
simp,
simp only:polyn_minus_m_cf[of d "fst d"],
drule sym, simp)
apply (frule polyn_add1[of c "m_cf S d"], assumption+, simp add:m_cf_len,
thin_tac "polyn_expr R X (fst c) c \<plusminus> polyn_expr R X (fst d)
(m_cf S d) =
polyn_expr R X (max (fst c) (fst d)) (add_cf S c (m_cf S d))",
thin_tac "polyn_expr R X (fst d) d = polyn_expr R X (fst c) c",
thin_tac "polyn_expr R X (fst c) c \<in> carrier R", drule sym)
apply (frule_tac add_cf_pol_coeff[of c "m_cf S d"], simp add:m_cf_len)
apply (frule coeff_0_pol_0[THEN sym, of "add_cf S c (m_cf S d)"
"max (fst c) (fst d)"],
simp add:add_cf_len m_cf_len, simp,
thin_tac "polyn_expr R X (max (fst c) (fst d))
(add_cf S c (m_cf S d)) = \<zero>",
thin_tac "pol_coeff S (add_cf S c (m_cf S d))")
apply (simp add:add_cf_def m_cf_def max_def)
apply (rule conjI)
apply (rule allI, rule impI,
frule_tac x = j and y = "fst c" and z = "fst d" in le_less_trans,
assumption+,
frule_tac x = j and y = "fst d" in less_imp_le)
apply (drule_tac a = j in forall_spec, simp, simp)
apply (frule_tac j = j in pol_coeff_mem[of c], simp,
frule_tac j = j in pol_coeff_mem[of d], simp)
apply (simp add:aGroup.ag_eq_diffzero[THEN sym])
apply (rule ballI, simp add:nset_def, erule conjE)
apply (cut_tac x = "fst c" and y = "Suc (fst c)" and z = j in
less_le_trans, simp, assumption)
apply (cut_tac m1 = "fst c" and n1 = j in nat_not_le_less[THEN sym], simp)
apply (drule_tac a = j in forall_spec, assumption, simp,
frule_tac j = j in pol_coeff_mem[of d], simp)
apply (frule_tac x = "snd d j" in aGroup.ag_inv_inv[of S], assumption,
simp add:aGroup.ag_inv_inv aGroup.ag_inv_zero)
apply (cut_tac polyn_n_m[of d "fst c" "fst d"])
apply (subst polyn_expr_split[of "fst d" d], simp,
thin_tac "polyn_expr R X (fst d) d =
polyn_expr R X (fst c) (fst c, snd d) \<plusminus>
\<Sigma>\<^sub>f R (\<lambda>j. snd d j \<cdot>\<^sub>r X^\<^bsup>R j\<^esup>) (Suc (fst c)) (fst d)", erule conjE)
apply (subst higher_part_zero[of d "fst c"], assumption+)
apply (frule pol_coeff_le[of d "fst c"], simp add:less_imp_le,
frule polyn_mem[of "(fst c, snd d)" "fst c"], simp,
cut_tac ring_is_ag,
simp add:aGroup.ag_r_zero,
subst polyn_expr_short[THEN sym, of d "fst c"], assumption+,
simp add:less_imp_le)
apply (rule polyn_exprs_eq[of c d "fst c"], assumption+)
apply (simp, assumption+)
apply (simp add:less_imp_le)
done
lemma (in PolynRg) polyn_degree_unique:"\<lbrakk>pol_coeff S c; pol_coeff S d;
polyn_expr R X (fst c) c = polyn_expr R X (fst d) d\<rbrakk> \<Longrightarrow>
c_max S c = c_max S d"
apply (cut_tac ring_is_ag,
cut_tac subring,
frule subring_Ring,
frule Ring.ring_is_ag[of S])
apply (case_tac "polyn_expr R X (fst d) d = \<zero>\<^bsub>R\<^esub>")
apply (cut_tac coeff_0_pol_0[THEN sym, of d "fst d"], simp,
cut_tac coeff_0_pol_0[THEN sym, of c "fst c"], simp)
apply (simp add:c_max_def, assumption, simp, assumption, simp)
apply (frule polyn_mem[of c "fst c"], simp, frule polyn_mem[of d "fst d"],
simp)
apply (frule aGroup.ag_eq_diffzero[of "R" "polyn_expr R X (fst c) c"
"polyn_expr R X (fst d) d"], assumption+)
apply (simp only:polyn_minus_m_cf[of d "fst d"],
frule m_cf_pol_coeff [of d])
apply (frule polyn_add1[of c "m_cf S d"], assumption+,
simp only:m_cf_len)
apply (rotate_tac -1, drule sym, simp,
thin_tac "polyn_expr R X (fst d) d \<plusminus>
polyn_expr R X (fst d) (m_cf S d) = \<zero>",
frule add_cf_pol_coeff[of c "m_cf S d"], assumption+)
apply (cut_tac coeff_0_pol_0[THEN sym, of "add_cf S c (m_cf S d)"
"fst (add_cf S c (m_cf S d))"],
simp add:add_cf_len m_cf_len,
thin_tac "polyn_expr R X (max (fst c) (fst d))
(add_cf S c (m_cf S d)) = \<zero>",
thin_tac "pol_coeff S (add_cf S c (m_cf S d))",
thin_tac "pol_coeff S (m_cf S d)")
apply (frule coeff_nonzero_polyn_nonzero[of d "fst d"], simp, simp)
apply (drule sym, simp)
apply (frule coeff_nonzero_polyn_nonzero[of c "fst c"], simp, simp)
apply (simp add:c_max_def, rule conjI, rule impI, blast,
rule conjI, rule impI, blast)
apply (rule n_max_eq_sets)
apply (rule equalityI)
apply (rule subsetI, simp)
apply (erule conjE)
apply (case_tac "fst c \<le> fst d")
apply (frule_tac i = x in le_trans[of _ "fst c" "fst d"], assumption+, simp,
rule contrapos_pp, simp+, simp add:max_def,
frule_tac i = x in le_trans[of _ "fst c" "fst d"], assumption+,
drule_tac a = x in forall_spec, assumption,
drule le_imp_less_or_eq[of "fst c" "fst d"],
erule disjE, simp add:add_cf_def m_cf_len m_cf_def,
frule_tac j = x in pol_coeff_mem[of c], assumption+,
simp add:aGroup.ag_inv_zero aGroup.ag_r_zero[of S])
apply (simp add:add_cf_def m_cf_len m_cf_def,
rotate_tac -1, drule sym, simp,
frule_tac j = x in pol_coeff_mem[of c], simp,
simp add:aGroup.ag_inv_zero aGroup.ag_r_zero[of S])
apply (simp add:nat_not_le_less) (* \<not> fst c \<le> fst d *)
apply (case_tac "\<not> x \<le> (fst d)", simp,
simp add:nat_not_le_less,
frule_tac x = "fst d" and y = x and z = "fst c" in
less_le_trans, assumption+,
drule_tac x = x in spec, simp add:max_def,
simp add:add_cf_def m_cf_len m_cf_def)
apply (simp,
drule_tac x = x in spec, simp add:max_def,
rule contrapos_pp, simp+,
simp add:add_cf_def m_cf_len m_cf_def,
frule_tac j = x in pol_coeff_mem[of c],
frule_tac x = x and y = "fst d" and z = "fst c" in
le_less_trans, assumption+, simp add:less_imp_le,
simp add:aGroup.ag_inv_zero aGroup.ag_r_zero[of S])
apply (rule subsetI, simp, erule conjE,
case_tac "fst d \<le> fst c",
frule_tac i = x and j = "fst d" and k = "fst c" in le_trans,
assumption+, simp,
drule_tac x = x in spec, simp add:max_def,
rule contrapos_pp, simp+,
simp add:add_cf_def m_cf_len m_cf_def)
apply (case_tac "fst d = fst c", simp, rotate_tac -1, drule sym, simp,
frule_tac j = x in pol_coeff_mem[of d], assumption,
frule_tac x = "snd d x" in aGroup.ag_mOp_closed, assumption+,
simp add:aGroup.ag_l_zero,
frule_tac x = "snd d x" in aGroup.ag_inv_inv[of S],
assumption, simp add:aGroup.ag_inv_zero)
apply (drule noteq_le_less[of "fst d" "fst c"], assumption,
simp,
frule_tac j = x in pol_coeff_mem[of d], assumption,
frule_tac x = "snd d x" in aGroup.ag_mOp_closed, assumption+,
simp add:aGroup.ag_l_zero,
frule_tac x = "snd d x" in aGroup.ag_inv_inv[of S],
assumption, simp add:aGroup.ag_inv_zero)
apply (simp add:nat_not_le_less,
case_tac "\<not> x \<le> fst c", simp,
simp add:nat_not_le_less,
drule_tac x = x in spec, simp add:max_def,
simp add:add_cf_def m_cf_len m_cf_def,
frule_tac j = x in pol_coeff_mem[of d], assumption,
frule_tac x = "snd d x" in aGroup.ag_mOp_closed, assumption+,
simp add:aGroup.ag_l_zero,
frule_tac x = "snd d x" in aGroup.ag_inv_inv[of S],
assumption, simp add:aGroup.ag_inv_zero)
apply (simp,
drule_tac x = x in spec, simp add:max_def,
rule contrapos_pp, simp+,
simp add:add_cf_def m_cf_len m_cf_def,
frule_tac x = x and y = "fst c" and z = "fst d" in le_less_trans,
assumption+, frule_tac x = x and y = "fst d" in less_imp_le,
frule_tac j = x in pol_coeff_mem[of d], assumption,
frule_tac x = "snd d x" in aGroup.ag_mOp_closed, assumption+,
simp add:aGroup.ag_l_zero,
frule_tac x = "snd d x" in aGroup.ag_inv_inv[of S],
assumption, simp add:aGroup.ag_inv_zero)
apply (thin_tac "\<forall>j\<le>max (fst c) (fst d). snd (add_cf S c (m_cf S d)) j = \<zero>\<^bsub>S\<^esub>")
apply (rotate_tac -1, drule sym, simp)
apply (simp add:coeff_0_pol_0[THEN sym, of c "fst c"])
apply blast
apply simp+
done
lemma (in PolynRg) ex_polyn_expr:"p \<in> carrier R \<Longrightarrow>
\<exists>c. pol_coeff S c \<and> p = polyn_expr R X (fst c) c"
apply (cut_tac S_X_generate[of p], blast)
apply assumption
done
lemma (in PolynRg) c_max_eqTr0:"\<lbrakk>pol_coeff S c; k \<le> (fst c);
polyn_expr R X k c = polyn_expr R X (fst c) c; \<exists>j\<le>k. (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>\<rbrakk> \<Longrightarrow>
c_max S (k, snd c) = c_max S c"
apply (simp add:polyn_expr_short[of c k],
frule pol_coeff_le[of c k], assumption+,
rule polyn_degree_unique[of "(k, snd c)" c], assumption+,
simp)
done
definition
cf_sol :: "[('a, 'b) Ring_scheme, ('a, 'b1) Ring_scheme, 'a, 'a,
nat \<times> (nat \<Rightarrow> 'a)] \<Rightarrow> bool" where
"cf_sol R S X p c \<longleftrightarrow> pol_coeff S c \<and> (p = polyn_expr R X (fst c) c)"
definition
deg_n ::"[('a, 'b) Ring_scheme, ('a, 'b1) Ring_scheme, 'a, 'a] \<Rightarrow> nat" where
"deg_n R S X p = c_max S (SOME c. cf_sol R S X p c)"
definition
deg ::"[('a, 'b) Ring_scheme, ('a, 'b1) Ring_scheme, 'a, 'a] \<Rightarrow> ant" where
"deg R S X p = (if p = \<zero>\<^bsub>R\<^esub> then -\<infinity> else (an (deg_n R S X p)))"
lemma (in PolynRg) ex_cf_sol:"p \<in> carrier R \<Longrightarrow>
\<exists>c. cf_sol R S X p c"
apply (unfold cf_sol_def)
apply (frule ex_polyn_expr[of p], (erule exE)+)
apply (cut_tac n = "fst c" in le_refl, blast)
done
lemma (in PolynRg) deg_in_aug_minf:"p \<in> carrier R \<Longrightarrow>
deg R S X p \<in> Z\<^sub>-\<^sub>\<infinity>"
apply (simp add:aug_minf_def deg_def an_def)
done
lemma (in PolynRg) deg_noninf:"p \<in> carrier R \<Longrightarrow>
deg R S X p \<noteq> \<infinity>"
apply (cut_tac deg_in_aug_minf[of p], simp add:deg_def,
simp add:aug_minf_def)
apply (case_tac "p = \<zero>\<^bsub>R\<^esub>", simp+)
done
lemma (in PolynRg) deg_ant_int:"\<lbrakk>p \<in> carrier R; p \<noteq> \<zero>\<rbrakk>
\<Longrightarrow> deg R S X p = ant (int (deg_n R S X p))"
by (simp add:deg_def an_def)
lemma (in PolynRg) deg_an:"\<lbrakk>p \<in> carrier R; p \<noteq> \<zero>\<rbrakk>
\<Longrightarrow> deg R S X p = an (deg_n R S X p)"
by (simp add:deg_def)
lemma (in PolynRg) pol_SOME_1:"p \<in> carrier R \<Longrightarrow>
cf_sol R S X p (SOME f. cf_sol R S X p f)"
apply (frule ex_cf_sol[of p])
apply (rule_tac P = "cf_sol R S X p" in someI_ex, assumption)
done
lemma (in PolynRg) pol_SOME_2:"p \<in> carrier R \<Longrightarrow>
pol_coeff S (SOME c. cf_sol R S X p c) \<and>
p = polyn_expr R X (fst (SOME c. cf_sol R S X p c))
(SOME c. cf_sol R S X p c)"
apply (frule pol_SOME_1[of p])
apply (simp add:cf_sol_def)
done
lemma (in PolynRg) coeff_max_zeroTr:"pol_coeff S c \<Longrightarrow>
\<forall>j. j \<le> (fst c) \<and> (c_max S c) < j \<longrightarrow> (snd c) j = \<zero>\<^bsub>S\<^esub>"
apply (case_tac "\<forall>j \<le> (fst c). (snd c) j = \<zero>\<^bsub>S\<^esub>", rule allI, rule impI,
erule conjE, simp)
apply simp
apply (frule coeff_nonzero_polyn_nonzero[THEN sym, of c "fst c"], simp,
simp)
apply (rule allI, rule impI, erule conjE,
simp add:c_max_def,
simp add:polyn_degreeTr[of c "fst c"])
apply (subgoal_tac "{j. j \<le> (fst c) \<and> (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>} \<subseteq> {j. j \<le> (fst c)}",
frule n_max[of "{j. j \<le> (fst c) \<and> (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>}" "fst c"], blast)
apply (case_tac "\<forall>x\<le>fst c. snd c x = \<zero>\<^bsub>S\<^esub> ", blast, simp)
apply (erule conjE)
apply (rule contrapos_pp, simp+,
thin_tac "\<exists>x\<le>fst c. snd c x \<noteq> \<zero>\<^bsub>S\<^esub>",
thin_tac "{j. j \<le> fst c \<and> snd c j \<noteq> \<zero>\<^bsub>S\<^esub>} \<subseteq> {j. j \<le> fst c}",
thin_tac "snd c (n_max {j. j \<le> fst c \<and> snd c j \<noteq> \<zero>\<^bsub>S\<^esub>}) \<noteq> \<zero>\<^bsub>S\<^esub>",
drule_tac a = j in forall_spec, simp)
apply simp
apply (rule subsetI, simp)
done
lemma (in PolynRg) coeff_max_nonzeroTr:"\<lbrakk>pol_coeff S c;
\<exists>j \<le> (fst c). (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>\<rbrakk> \<Longrightarrow> (snd c) (c_max S c) \<noteq> \<zero>\<^bsub>S\<^esub>"
apply (simp add:c_max_def)
apply (subgoal_tac "{j. j \<le> (fst c) \<and> (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>} \<subseteq> {j. j \<le> (fst c)}",
frule n_max[of "{j. j \<le> (fst c) \<and> (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>}" "fst c"], blast)
apply (erule conjE, simp)
apply (rule subsetI, simp)
done
lemma (in PolynRg) coeff_max_bddTr:"pol_coeff S c \<Longrightarrow> c_max S c \<le> (fst c)"
apply (case_tac "\<forall>j\<le>(fst c). (snd c) j = \<zero>\<^bsub>S\<^esub>", simp add:c_max_def)
apply (simp add:c_max_def,
frule polyn_degreeTr[of c "fst c"], simp, simp,
subgoal_tac "{j. j \<le> (fst c) \<and> (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>} \<subseteq> {j. j \<le> (fst c)}",
frule n_max[of "{j. j \<le> (fst c) \<and> (snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>}" "fst c"],
blast, erule conjE, simp)
apply (rule subsetI, simp)
done
lemma (in PolynRg) pol_coeff_max:"pol_coeff S c \<Longrightarrow>
pol_coeff S ((c_max S c), snd c)"
apply (rule pol_coeff_le[of c "c_max S c"], assumption)
apply (simp add:coeff_max_bddTr)
done
lemma (in PolynRg) polyn_c_max:"pol_coeff S c \<Longrightarrow>
polyn_expr R X (fst c) c = polyn_expr R X (c_max S c) c"
apply (case_tac "(c_max S c) = (fst c)", simp)
apply (frule coeff_max_bddTr[of c],
frule noteq_le_less[of "c_max S c" "fst c"], assumption)
apply (subst polyn_n_m1[of c "c_max S c" "fst c"], assumption+, simp)
apply (frule_tac polyn_mem[of c "c_max S c"], assumption+)
apply (subst higher_part_zero[of c "c_max S c"], assumption+)
apply (frule coeff_max_zeroTr[of c],
rule ballI, simp add:nset_def)
apply (cut_tac ring_is_ag, simp add:aGroup.ag_r_zero)
done
lemma (in PolynRg) pol_deg_eq_c_max:"\<lbrakk>p \<in> carrier R;
pol_coeff S c; p = polyn_expr R X (fst c) c\<rbrakk> \<Longrightarrow>
deg_n R S X p = c_max S c"
apply (cut_tac subring, frule subring_Ring)
apply (frule polyn_c_max[of c])
apply (frule pol_SOME_2[of p])
apply (subst deg_n_def, erule conjE)
apply (rule polyn_degree_unique[of "Eps (cf_sol R S X p)" "c"], simp,
assumption)
apply simp
done
lemma (in PolynRg) pol_deg_le_n:"\<lbrakk>p \<in> carrier R; pol_coeff S c;
p = polyn_expr R X (fst c) c\<rbrakk> \<Longrightarrow> deg_n R S X p \<le> (fst c)"
apply (frule pol_deg_eq_c_max[of p c], assumption+,
frule coeff_max_bddTr[of c])
apply simp
done
lemma (in PolynRg) pol_deg_le_n1:"\<lbrakk>p \<in> carrier R; pol_coeff S c; k \<le> (fst c);
p = polyn_expr R X k c\<rbrakk> \<Longrightarrow> deg_n R S X p \<le> k"
apply (simp add:deg_n_def, drule sym, simp)
apply (frule pol_SOME_2[of p], erule conjE)
apply (frule pol_coeff_le[of c k], assumption)
apply (simp only:polyn_expr_short[of c k])
apply (drule sym)
apply (subst polyn_degree_unique[of "SOME c. cf_sol R S X p c" "(k, snd c)"],
assumption+, simp)
apply (frule coeff_max_bddTr[of "(k, snd c)"], simp)
done
lemma (in PolynRg) pol_len_gt_deg:"\<lbrakk>p \<in> carrier R; pol_coeff S c;
p = polyn_expr R X (fst c) c; deg R S X p < (an j); j \<le> (fst c)\<rbrakk>
\<Longrightarrow> (snd c) j = \<zero>\<^bsub>S\<^esub>"
apply (case_tac "p = \<zero>\<^bsub>R\<^esub>", simp, drule sym)
apply (simp add:coeff_0_pol_0[THEN sym, of c "fst c"])
apply (simp add:deg_def, simp add:aless_natless)
apply (drule sym, simp)
apply (frule coeff_max_zeroTr[of c])
apply (simp add:pol_deg_eq_c_max)
done
lemma (in PolynRg) pol_diff_deg_less:"\<lbrakk>p \<in> carrier R; pol_coeff S c;
p = polyn_expr R X (fst c) c; pol_coeff S d;
fst c = fst d; (snd c) (fst c) = (snd d) (fst d)\<rbrakk> \<Longrightarrow>
p \<plusminus> (-\<^sub>a (polyn_expr R X (fst d) d)) = \<zero> \<or>
deg_n R S X (p \<plusminus> (-\<^sub>a (polyn_expr R X (fst d) d))) < (fst c)"
apply (cut_tac ring_is_ag,
cut_tac subring, frule subring_Ring)
apply (case_tac "p \<plusminus>\<^bsub>R\<^esub> (-\<^sub>a\<^bsub>R\<^esub> (polyn_expr R X (fst d) d)) = \<zero>\<^bsub>R\<^esub>", simp)
apply simp
apply (simp add:polyn_minus_m_cf[of d "fst d"],
frule m_cf_pol_coeff[of d])
apply (cut_tac polyn_add1[of c "m_cf S d"], simp add:m_cf_len,
thin_tac "polyn_expr R X (fst d) c \<plusminus> polyn_expr R X (fst d) (m_cf S d)
= polyn_expr R X (fst d) (add_cf S c (m_cf S d))")
apply (frule add_cf_pol_coeff[of c "m_cf S d"], assumption+)
apply (cut_tac polyn_mem[of "add_cf S c (m_cf S d)" "fst d"],
frule pol_deg_le_n[of "polyn_expr R X (fst d) (add_cf S c (m_cf S d))"
"add_cf S c (m_cf S d)"], assumption+,
simp add:add_cf_len m_cf_len,
simp add:add_cf_len m_cf_len)
apply (rule noteq_le_less[of "deg_n R S X (polyn_expr R X (fst d)
(add_cf S c (m_cf S d)))" "fst d"], assumption)
apply (rule contrapos_pp, simp+)
apply (cut_tac pol_deg_eq_c_max[of "polyn_expr R X (fst d)
(add_cf S c (m_cf S d))" "add_cf S c (m_cf S d)"],
simp,
thin_tac "deg_n R S X (polyn_expr R X (fst d) (add_cf S c (m_cf S d)))
= fst d")
apply (frule coeff_nonzero_polyn_nonzero[of "add_cf S c (m_cf S d)" "fst d"],
simp add:add_cf_len m_cf_len, simp,
thin_tac "polyn_expr R X (fst d) (add_cf S c (m_cf S d)) \<noteq> \<zero>",
frule coeff_max_nonzeroTr[of "add_cf S c (m_cf S d)"],
simp add:add_cf_len m_cf_len,
thin_tac "\<exists>j\<le>fst d. snd (add_cf S c (m_cf S d)) j \<noteq> \<zero>\<^bsub>S\<^esub>",
thin_tac "pol_coeff S (m_cf S d)",
thin_tac "pol_coeff S (add_cf S c (m_cf S d))",
thin_tac "polyn_expr R X (fst d) (add_cf S c (m_cf S d)) \<in>
carrier R", simp,
thin_tac "c_max S (add_cf S c (m_cf S d)) = fst d")
apply (simp add:add_cf_def m_cf_def,
frule pol_coeff_mem[of d "fst d"], simp,
frule Ring.ring_is_ag[of S],
simp add:aGroup.ag_r_inv1, assumption+,
simp add:add_cf_len m_cf_len, assumption,
simp add:add_cf_len m_cf_len, assumption+)
done
lemma (in PolynRg) pol_pre_lt_deg:"\<lbrakk>p \<in> carrier R; pol_coeff S c;
deg_n R S X p \<le> (fst c); (deg_n R S X p) \<noteq> 0;
p = polyn_expr R X (deg_n R S X p) c \<rbrakk> \<Longrightarrow>
(deg_n R S X (polyn_expr R X ((deg_n R S X p) - Suc 0) c)) < (deg_n R S X p)"
apply (frule polyn_expr_short[of c "deg_n R S X p"], assumption)
apply (cut_tac pol_deg_le_n[of "polyn_expr R X (deg_n R S X p - Suc 0) c"
"(deg_n R S X p - Suc 0, snd c)"], simp)
apply (rule polyn_mem[of c "deg_n R S X p - Suc 0"], assumption+,
arith,
rule pol_coeff_le[of c "deg_n R S X p - Suc 0"], assumption,
arith, simp)
apply (subst polyn_expr_short[of c "deg_n R S X p - Suc 0"],
assumption+, arith, simp)
done
lemma (in PolynRg) pol_deg_n:"\<lbrakk>p \<in> carrier R; pol_coeff S c;
n \<le> fst c; p = polyn_expr R X n c; (snd c) n \<noteq> \<zero>\<^bsub>S\<^esub>\<rbrakk> \<Longrightarrow>
deg_n R S X p = n"
apply (simp add:polyn_expr_short[of c n])
apply (frule pol_coeff_le[of c n], assumption+,
cut_tac pol_deg_eq_c_max[of p "(n, snd c)"],
drule sym, simp, simp add:c_max_def)
apply (rule conjI, rule impI, cut_tac le_refl[of n],
thin_tac "deg_n R S X p =
(if \<forall>x\<le>n. snd c x = \<zero>\<^bsub>S\<^esub> then 0
else n_max {j. j \<le> fst (n, snd c) \<and> snd (n, snd c) j \<noteq> \<zero>\<^bsub>S\<^esub>})",
drule_tac a = n in forall_spec, assumption, simp)
apply (rule impI)
apply (cut_tac n_max[of "{j. j \<le> n \<and> snd c j \<noteq> \<zero>\<^bsub>S\<^esub>}" n], erule conjE,
drule_tac x = n in bspec, simp, simp)
apply (rule subsetI, simp, blast,
drule sym, simp, assumption)
apply simp
done
lemma (in PolynRg) pol_expr_deg:"\<lbrakk>p \<in> carrier R; p \<noteq> \<zero>\<rbrakk>
\<Longrightarrow> \<exists>c. pol_coeff S c \<and> deg_n R S X p \<le> (fst c) \<and>
p = polyn_expr R X (deg_n R S X p) c \<and>
(snd c) (deg_n R S X p) \<noteq> \<zero>\<^bsub>S\<^esub>"
apply (cut_tac subring,
frule subring_Ring)
apply (frule ex_polyn_expr[of p], erule exE, erule conjE)
apply (frule_tac c = c in polyn_c_max)
apply (frule_tac c = c in pol_deg_le_n[of p], assumption+)
apply (frule_tac c1 = c and k1 ="fst c" in coeff_0_pol_0[THEN sym], simp)
apply (subgoal_tac "p = polyn_expr R X (deg_n R S X p) c \<and>
snd c (deg_n R S X p) \<noteq> \<zero>\<^bsub>S\<^esub>", blast)
apply (subst pol_deg_eq_c_max, assumption+)+
apply simp
apply (cut_tac c = c in coeff_max_nonzeroTr, simp+)
done
lemma (in PolynRg) deg_n_pos:"p \<in> carrier R \<Longrightarrow> 0 \<le> deg_n R S X p"
by simp
lemma (in PolynRg) pol_expr_deg1:"\<lbrakk>p \<in> carrier R; d = na (deg R S X p)\<rbrakk> \<Longrightarrow>
\<exists>c. (pol_coeff S c \<and> p = polyn_expr R X d c)"
apply (cut_tac subring, frule subring_Ring)
apply (case_tac "p = \<zero>\<^bsub>R\<^esub>",
simp add:deg_def na_minf,
subgoal_tac "pol_coeff S (0, (\<lambda>j. \<zero>\<^bsub>S\<^esub>))",
subgoal_tac "\<zero> = polyn_expr R X d (0, (\<lambda>j. \<zero>\<^bsub>S\<^esub>))", blast,
cut_tac coeff_0_pol_0[of "(d, \<lambda>j. \<zero>\<^bsub>S\<^esub>)" d], simp+,
simp add:pol_coeff_def,
simp add:Ring.ring_zero)
apply (simp add:deg_def na_an,
frule pol_expr_deg[of p], assumption,
erule exE, (erule conjE)+,
unfold split_paired_all, simp, blast)
done
end
diff --git a/thys/Group-Ring-Module/Algebra8.thy b/thys/Group-Ring-Module/Algebra8.thy
--- a/thys/Group-Ring-Module/Algebra8.thy
+++ b/thys/Group-Ring-Module/Algebra8.thy
@@ -1,2823 +1,2824 @@
(** Algebra8
author Hidetsune Kobayashi
Group You Santo
Department of Mathematics
Nihon University
hikoba@math.cst.nihon-u.ac.jp
May 3, 2004.
April 6, 2007 (revised)
chapter 5. Modules
section 4. nsum and Generators(continued)
section 5. existence of homomorphism
section 6. Nakayama's lemma
section 7. direct sum and direct products of modules
**)
theory Algebra8 imports Algebra7 begin
section "nsum and Generators (continued)"
lemma (in Module) unique_expression_last:" \<lbrakk>free_generator R M H;
f \<in> {j. j \<le> Suc n} \<rightarrow> H; s \<in> {j. j \<le> Suc n} \<rightarrow> carrier R;
g \<in> {j. j \<le> Suc n} \<rightarrow> H; t \<in> {j. j \<le> Suc n} \<rightarrow> carrier R;
l_comb R M (Suc n) s f = l_comb R M (Suc n) t g;
inj_on f {j. j \<le> Suc n}; inj_on g {j. j \<le> Suc n};
f (Suc n) = g (Suc n)\<rbrakk> \<Longrightarrow> s (Suc n) = t (Suc n)"
apply (cut_tac sc_Ring, frule Ring.whole_ideal, frule free_generator_sub)
apply (frule l_comb_mem[of "carrier R" H s "Suc n" f], assumption+,
frule l_comb_mem[of "carrier R" H t "Suc n" g], assumption+)
apply (frule ag_r_inv1[of "l_comb R M (Suc n) t g"],
simp only:linear_span_iOp_closedTr2,
thin_tac "l_comb R M (Suc n) t g \<in> carrier M",
frule sym,
thin_tac "l_comb R M (Suc n) s f = l_comb R M (Suc n) t g",
simp,
thin_tac "l_comb R M (Suc n) t g = l_comb R M (Suc n) s f",
subgoal_tac "(\<lambda>x\<in>{j. j \<le> Suc n}. -\<^sub>a\<^bsub>R\<^esub> (t x)) \<in>
{j. j \<le> Suc n} \<rightarrow> carrier R",
simp add:l_comb_Suc,
frule funcset_mem[of s "{j. j \<le> Suc n}" "carrier R" "Suc n"],
simp,
frule funcset_mem[of t "{j. j \<le> Suc n}" "carrier R" "Suc n"],
simp,
frule funcset_mem[of g "{j. j \<le> Suc n}" H "Suc n"],
simp,
frule subsetD[of H "carrier M" "g (Suc n)"], assumption+,
frule func_pre[of s], frule func_pre[of t], frule func_pre[of f],
frule func_pre[of g],
frule func_pre[of "\<lambda>x\<in>{j. j \<le> Suc n}. -\<^sub>a\<^bsub>R\<^esub> (t x)"],
frule l_comb_mem[of "carrier R" H s n f], assumption+,
frule l_comb_mem[of "carrier R" H "\<lambda>x\<in>{j. j \<le> Suc n}. -\<^sub>a\<^bsub>R\<^esub> (t x)"
n g], assumption+,
frule sc_mem[of "s (Suc n)" "g (Suc n)"], assumption+,
cut_tac sc_mem[of "-\<^sub>a\<^bsub>R\<^esub> (t (Suc n))" "g (Suc n)"],
simp add:pOp_assocTr43[of "l_comb R M n s f" "s (Suc n) \<cdot>\<^sub>s g (Suc n)"],
simp add:ag_pOp_commute[of "s (Suc n) \<cdot>\<^sub>s g (Suc n)"],
simp add:pOp_assocTr43[THEN sym, of "l_comb R M n s f"])
apply (frule Ring.ring_is_ag[of R],
frule aGroup.ag_mOp_closed[of R "t (Suc n)"], assumption+,
simp add:sc_l_distr[THEN sym],
simp add:l_comb_add[THEN sym, of "carrier R" H s n f
"\<lambda>x\<in>{j. j \<le> Suc n}. -\<^sub>a\<^bsub>R\<^esub> (t x)" n g])
apply (frule jointfun_hom[of f n H g n H], assumption+,
frule jointfun_hom[of s n "carrier R" "\<lambda>x\<in>{j. j \<le> Suc n}. -\<^sub>a\<^bsub>R\<^esub> (t x)"
n "carrier R"], assumption+, simp)
apply (subgoal_tac "(s (Suc n) \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (t (Suc n))) \<cdot>\<^sub>s g (Suc n) =
l_comb R M 0 (\<lambda>x\<in>{0::nat}. (s (Suc n) \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (t (Suc n))))
(\<lambda>x\<in>{0::nat}. (g (Suc n)))",
simp,
thin_tac "(s (Suc n) \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (t (Suc n))) \<cdot>\<^sub>s g (Suc n) =
l_comb R M 0 (\<lambda>x\<in>{0}. s (Suc n) \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (t (Suc n)))
(\<lambda>x\<in>{0}. g (Suc n))")
apply (subgoal_tac "(\<lambda>x\<in>{0::nat}. (s (Suc n) \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (t (Suc n)))) \<in>
{j. j \<le> (0::nat)} \<rightarrow> carrier R",
subgoal_tac "(\<lambda>x\<in>{0::nat}. g (Suc n)) \<in> {j. j \<le> (0::nat)} \<rightarrow> H")
apply (simp add:l_comb_add[THEN sym] del: Pi_split_insert_domain)
apply (cut_tac unique_expression3_1[of H "jointfun (Suc (n + n))
(jointfun n f n g) 0 (\<lambda>x\<in>{0}. g (Suc n))" "Suc (n + n)"
"jointfun (Suc (n + n))
(jointfun n s n (\<lambda>x\<in>{j. j \<le> Suc n}. -\<^sub>a\<^bsub>R\<^esub> (t x))) 0
(\<lambda>x\<in>{0}. s (Suc n) \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (t (Suc n)))"], simp del: Pi_split_insert_domain,
thin_tac "l_comb R M (Suc (Suc (n + n)))
(jointfun (Suc (n + n)) (jointfun n s n (\<lambda>x\<in>{j. j \<le> Suc n}. -\<^sub>a\<^bsub>R\<^esub> (t x))) 0
(\<lambda>x\<in>{0}. s (Suc n) \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (t (Suc n))))
(jointfun (Suc (n + n)) (jointfun n f n g) 0 (\<lambda>x\<in>{0}. g (Suc n))) =
\<zero>",
thin_tac "(\<lambda>x\<in>{j. j \<le> Suc n}. -\<^sub>a\<^bsub>R\<^esub> (t x)) \<in> {j. j \<le> Suc n} \<rightarrow> carrier R",
thin_tac "l_comb R M n s f \<in> carrier M",
thin_tac "jointfun n f n g \<in> {j. j \<le> Suc (n + n)} \<rightarrow> H",
thin_tac "jointfun n s n (\<lambda>x\<in>{j. j \<le> Suc n}. -\<^sub>a\<^bsub>R\<^esub> (t x))
\<in> {j. j \<le> Suc (n + n)} \<rightarrow> carrier R",
thin_tac "(\<lambda>x\<in>{0}. s (Suc n) \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (t (Suc n))) \<in> {0} \<rightarrow> carrier R")
apply (thin_tac " (-\<^sub>a\<^bsub>R\<^esub> (t (Suc n))) \<cdot>\<^sub>s g (Suc n) \<in> carrier M",
thin_tac "-\<^sub>a\<^bsub>R\<^esub> (t (Suc n)) \<in> carrier R",
thin_tac "l_comb R M n (\<lambda>x\<in>{j. j \<le> Suc n}. -\<^sub>a\<^bsub>R\<^esub> (t x)) g \<in> carrier M")
apply ((erule exE)+, (erule conjE)+, erule exE, (erule conjE)+)
apply (frule_tac s = ta and n = m and m = ga in unique_expression1[of H],
assumption+, simp)
apply (drule_tac x = m in bspec, simp)
apply (simp add:jointfun_def[of "Suc (n + n)"] sliden_def,
thin_tac "\<zero> = l_comb R M m ta ga",
thin_tac "ta m = s (Suc n) \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (t (Suc n))",
frule Ring.ring_is_ag[of R],
rotate_tac -2, frule sym,
subst aGroup.ag_eq_diffzero[of R "s (Suc n)" "t (Suc n)"], assumption+)
apply (rule sym, assumption) apply assumption
apply (rule Pi_I,
case_tac "x \<le> Suc (n + n)", simp add:jointfun_def del: Pi_split_insert_domain,
simp add:Pi_def del: Pi_split_insert_domain,
rule impI, simp add:sliden_def del: Pi_split_insert_domain,
simp add:Pi_def, simp add:jointfun_def sliden_def del: Pi_split_insert_domain)
apply (rule Pi_I,
case_tac "x \<le> Suc (n + n)",
simp add:jointfun_def[of "Suc (n+n)"] del: Pi_split_insert_domain,
simp add:Pi_def,
simp add:jointfun_def[of "Suc (n + n)"] sliden_def del: Pi_split_insert_domain,
rule aGroup.ag_pOp_closed, assumption+,
simp add: Nset_pre1 del: Pi_split_insert_domain,
simp add:im_jointfunTr1[of "Suc (n+n)" "jointfun n f n g" 0
"\<lambda>x\<in>{0}. g (Suc n)"] del: Pi_split_insert_domain,
simp add:im_jointfun[of f n H g n H] del: Pi_split_insert_domain,
(subst jointfun_def[of "Suc (n+n)"])+, simp add:sliden_def del: Pi_split_insert_domain)
apply (rule conjI, frule sym, thin_tac "f (Suc n) = g (Suc n)",
simp add:inj_on_def[of f],
rule contrapos_pp, simp+, simp add:image_def, erule exE,
erule conjE,
drule_tac a = "Suc n" in forall_spec, simp,
drule_tac a = x in forall_spec, simp, simp)
apply (simp add:inj_on_def[of g],
rule contrapos_pp, simp+, simp add:image_def, erule exE,
erule conjE,
drule_tac a = "Suc n" in forall_spec, simp,
drule_tac a = x in forall_spec, simp, simp)
apply (simp)
apply (rule Pi_I, simp,
rule aGroup.ag_pOp_closed, assumption+)
apply (subst l_comb_def, simp,
rule aGroup.ag_mOp_closed, simp add:Ring.ring_is_ag,
simp add:Pi_def, simp add:Pi_def)
apply (rule Pi_I, simp,
frule Ring.ring_is_ag[of R],
rule aGroup.ag_mOp_closed[of R], assumption+,
simp add:Pi_def)
done
lemma (in Module) unique_exprTr7p1:" \<lbrakk>free_generator R M H;
\<forall>f s g t m.
f \<in> {j. j \<le> n} \<rightarrow> H \<and> inj_on f {j. j \<le> n} \<and> s \<in> {j. j \<le> n} \<rightarrow> carrier R \<and>
g \<in> {j. j \<le> m} \<rightarrow> H \<and> inj_on g {j. j \<le> m} \<and> t \<in> {j. j \<le> m} \<rightarrow> carrier R \<and>
l_comb R M n s f = l_comb R M m t g \<and>
(\<forall>j\<le>n. s j \<noteq> \<zero>\<^bsub>R\<^esub>) \<and> (\<forall>k\<le>m. t k \<noteq> \<zero>\<^bsub>R\<^esub>) \<longrightarrow>
n = m \<and>
(\<exists>h. h \<in> {j. j \<le> n} \<rightarrow> {j. j \<le> n} \<and>
(\<forall>l\<le>n. cmp f h l = g l \<and> cmp s h l = t l));
f \<in> {j. j \<le> Suc n} \<rightarrow> H; s \<in> {j. j \<le> Suc n} \<rightarrow> carrier R;
g \<in> {j. j \<le> Suc n} \<rightarrow> H; t \<in> {j. j \<le> Suc n} \<rightarrow> carrier R;
l_comb R M (Suc n) s f = l_comb R M (Suc n) t g; \<forall>j\<le>Suc n. s j \<noteq> \<zero>\<^bsub>R\<^esub>;
\<forall>k\<le>Suc n. t k \<noteq> \<zero>\<^bsub>R\<^esub>; inj_on f {j. j \<le> Suc n}; inj_on g {j. j \<le> Suc n};
f (Suc n) = g (Suc n)\<rbrakk> \<Longrightarrow> \<exists>h. h \<in> {j. j \<le> Suc n} \<rightarrow> {j. j \<le> Suc n} \<and>
(\<forall>l\<le>Suc n. cmp f h l = g l \<and> cmp s h l = t l)"
apply (cut_tac sc_Ring, frule Ring.whole_ideal,
frule Ring.ring_is_ag,
frule free_generator_sub)
apply (frule_tac f = f and n = n and s = s and g = g and t = t in
unique_expression_last[of H], assumption+)
apply (simp add:l_comb_Suc)
apply (frule_tac f = f in func_pre, frule_tac f = s in func_pre,
frule_tac f = g in func_pre, frule_tac f = t in func_pre,
cut_tac n = n in Nsetn_sub,
frule_tac f = f and A = "{j. j \<le> Suc n}" and ?A1.0 = "{j. j \<le> n}" in
restrict_inj, assumption+,
frule_tac f = g and A = "{j. j \<le> Suc n}" and ?A1.0 = "{j. j \<le> n}" in
restrict_inj, assumption+)
apply (frule_tac s = s and m = f and n = n in l_comb_mem[of "carrier R" H],
assumption+,
frule_tac s = t and m = g and n = n in l_comb_mem[of "carrier R" H],
assumption+,
frule_tac f = t and A = "{j. j \<le> Suc n}" and B = "carrier R" and
x = "Suc n" in funcset_mem, simp,
frule_tac f = g and A = "{j. j \<le> Suc n}" and B = H and x = "Suc n"
in funcset_mem, simp,
frule_tac c = "g (Suc n)" in subsetD[of H "carrier M"], assumption+,
frule_tac a = "t (Suc n)" and m = "g (Suc n)" in sc_mem, assumption+)
apply (frule_tac x = "l_comb R M n s f" and y = "t (Suc n) \<cdot>\<^sub>s g (Suc n)" and
z = "-\<^sub>a (t (Suc n) \<cdot>\<^sub>s g (Suc n))" in ag_pOp_assoc, assumption+,
rule ag_mOp_closed, assumption, simp add:ag_r_inv1 ag_r_zero,
frule_tac x = "t (Suc n) \<cdot>\<^sub>s g (Suc n)" in ag_mOp_closed,
simp add:ag_pOp_assoc, simp add:ag_r_inv1 ag_r_zero,
rotate_tac -2, frule sym,
thin_tac "l_comb R M n t g = l_comb R M n s f")
apply (drule_tac x = f in spec,
rotate_tac -1,
drule_tac x = s in spec,
rotate_tac -1,
drule_tac x = g in spec,
rotate_tac -1, drule_tac x = t in spec,
rotate_tac -1, drule_tac x = n in spec)
apply (cut_tac n = n in Nsetn_sub_mem1, simp)
apply (erule exE, erule conjE)
apply (subgoal_tac "(jointfun n h 0 (\<lambda>j. (Suc n))) \<in>
{j. j \<le> Suc n} \<rightarrow> {j. j \<le> Suc n}",
subgoal_tac "\<forall>l\<le>Suc n. cmp f (jointfun n h 0 (\<lambda>j. (Suc n))) l =
g l \<and> cmp s (jointfun n h 0 (\<lambda>j. (Suc n))) l = t l",
thin_tac "\<forall>k\<le>Suc n. t k \<noteq> \<zero>\<^bsub>R\<^esub>", thin_tac "inj_on f {j. j \<le> Suc n}",
thin_tac "inj_on g {j. j \<le> Suc n}", thin_tac "f (Suc n) = g (Suc n)",
thin_tac "s (Suc n) = t (Suc n)",
thin_tac "f \<in> {j. j \<le> n} \<rightarrow> H", thin_tac "s \<in> {j. j \<le> n} \<rightarrow> carrier R",
thin_tac "g \<in> {j. j \<le> n} \<rightarrow> H", thin_tac "t \<in> {j. j \<le> n} \<rightarrow> carrier R",
thin_tac "{j. j \<le> n} \<subseteq> {j. j \<le> Suc n}",
thin_tac "\<forall>l\<le>n. cmp f h l = g l \<and> cmp s h l = t l",
thin_tac "l_comb R M n t g \<in> carrier M", blast)
apply (rule allI, rule impI, case_tac "l \<le> n",
simp add:cmp_def jointfun_def,
simp add:cmp_def, simp add:jointfun_def sliden_def,
frule_tac y = l and x = n in not_le_imp_less, thin_tac "\<not> l \<le> n",
frule_tac m = n and n = l in Suc_leI,
frule_tac m = l and n = "Suc n" in Nat.le_antisym, assumption,
simp)
apply (rule Pi_I,
case_tac "xa \<le> n", simp add:jointfun_def, rule impI,
frule_tac x = x and f = h and A = "{j. j \<le> n}" and
B = "{j. j \<le> n}" in funcset_mem, simp, simp,
simp add:jointfun_def, rule impI,frule_tac x = x and f = h and
A = "{j. j \<le> n}" and B = "{j. j \<le> n}" in funcset_mem, simp, simp)
done
lemma (in Module) unique_expression7:"free_generator R M H \<Longrightarrow>
\<forall>f s g t m. f \<in> {j. j \<le> (n::nat)} \<rightarrow> H \<and> inj_on f {j. j \<le> n} \<and>
s \<in> {j. j \<le> n} \<rightarrow> carrier R \<and>
g \<in> {j. j \<le> (m::nat)} \<rightarrow> H \<and> inj_on g {j. j \<le> m} \<and>
t \<in> {j. j \<le> m} \<rightarrow> carrier R \<and> l_comb R M n s f = l_comb R M m t g \<and>
(\<forall>j \<in> {j. j \<le> n}. s j \<noteq> \<zero>\<^bsub>R\<^esub>) \<and> (\<forall>k \<in> {j. j \<le> m}. t k \<noteq> \<zero>\<^bsub>R\<^esub> ) \<longrightarrow> n = m \<and>
(\<exists>h. h \<in> {j. j \<le> n} \<rightarrow> {j. j \<le> n} \<and> (\<forall>l \<in> {j. j \<le> n}. (cmp f h) l = g l
\<and> (cmp s h) l = t l))"
apply (cut_tac sc_Ring,
frule Ring.ring_is_ag[of R],
frule free_generator_sub[of H])
apply (induct_tac n)
apply (rule allI)+
apply (rule impI, (erule conjE)+)
apply (frule_tac H = H and f = f and n = 0 and s = s and g = g and m = m
and t = t in unique_expression7_1, assumption+, simp del: Pi_split_insert_domain)
apply (frule_tac H = H and f = f and n = 0 and s = s and g = g and m = m
and t = t in unique_expression6, (simp del: Pi_split_insert_domain) +)
apply (simp add:l_comb_def del: Pi_split_insert_domain)
apply (frule_tac f = s in funcset_mem[of _ "{0}" "carrier R" 0], simp del: Pi_split_insert_domain,
frule_tac f = t in funcset_mem[of _ "{0}" "carrier R" 0], simp del: Pi_split_insert_domain,
frule_tac f = g in funcset_mem[of _ "{0}" H 0], simp del: Pi_split_insert_domain,
frule_tac c = "g 0" in subsetD[of "H" "carrier M"], assumption+)
apply (frule_tac a = "s 0" and m = "g 0" in sc_mem, assumption+,
frule_tac a = "t 0" and m = "g 0" in sc_mem, assumption+,
frule_tac a = "s 0 \<cdot>\<^sub>s g 0" and b = "t 0 \<cdot>\<^sub>s g 0" in ag_eq_diffzero)
apply assumption
apply (simp only:sc_minus_am1,
thin_tac "(s 0 \<cdot>\<^sub>s g 0 = t 0 \<cdot>\<^sub>s g 0) = (\<zero> = \<zero>)")
apply (frule_tac x = "t 0" in aGroup.ag_mOp_closed[of R], assumption+)
apply (simp add:sc_l_distr[THEN sym],
frule_tac h = "g 0" and a = "s 0 \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (t 0)" in
free_gen_coeff_zero[of H],
assumption+,
rule aGroup.ag_pOp_closed[of R], assumption+)
apply (simp add:aGroup.ag_eq_diffzero[THEN sym, of R],
simp add:cmp_def,
cut_tac Pi_I[of "{0::nat}" "id" "%_. {0::nat}"],
subgoal_tac "f (id 0) = g 0 \<and> s (id 0) = t 0", blast,
simp add:id_def, simp add:id_def)
(** n = 0 done **)
apply ((rule allI)+, rule impI)
apply (erule conjE)+
apply (frule_tac H = H and f = f and n = "Suc n" and s = s and g = g and
m = m and t = t in unique_expression6, assumption+)
apply (cut_tac k = "Suc n" in finite_Collect_le_nat,
frule_tac A = "{j. j \<le> Suc n}" and f = f in inj_on_iff_eq_card,
simp)
apply (cut_tac k = m in finite_Collect_le_nat,
frule_tac A = "{j. j \<le> m}" and f = g in inj_on_iff_eq_card,
simp,
thin_tac "card (g ` {j. j \<le> m}) = Suc m")
apply (frule sym, thin_tac "Suc n = m", simp)
apply (cut_tac a = "Suc n" and A = "{j. j \<le> Suc n}" and f = g in
mem_in_image2) apply simp apply (
rotate_tac -6) apply ( frule sym,
thin_tac "f ` {j. j \<le> Suc n} = g ` {j. j \<le> Suc n}", simp,
simp add:image_def) apply ( erule exE, erule conjE)
apply (case_tac "x = Suc n", simp, thin_tac "x = Suc n",
rotate_tac -1, frule sym, thin_tac "g (Suc n) = f (Suc n)")
apply (rule_tac n = n and f = f and s = s and g = g and t = t in
unique_exprTr7p1[of H], assumption+)
apply (rotate_tac -2, frule sym, thin_tac "g (Suc n) = f x")
apply (frule Ring.whole_ideal)
apply (frule_tac s = s and n = n and f = f and j = x in
l_comb_transpos1[of "carrier R" H], assumption+,
rule noteq_le_less, assumption+, simp) apply (
thin_tac "l_comb R M (Suc n) s f = l_comb R M (Suc n) t g")
apply (frule_tac f = "cmp f (transpos x (Suc n))" and n = n and
s = "cmp s (transpos x (Suc n))" and g = g and t = t in
unique_exprTr7p1[of H], assumption+)
apply (rule Pi_I,
simp add:cmp_def, rule_tac f = f and A = "{j. j \<le> Suc n}" and B = H
and x = "transpos x (Suc n) xa" in funcset_mem, simp,
simp add:transpos_mem)
apply (rule Pi_I,
simp add:cmp_def, rule_tac f = s and A = "{j. j \<le> Suc n}" and
B = "carrier R" and x = "transpos x (Suc n) xa" in funcset_mem, simp,
simp add:transpos_mem)
apply assumption+
apply (rule allI, rule impI, simp add:cmp_def)
apply (cut_tac i = x and n = "Suc n" and j = "Suc n" and l = j in
transpos_mem, simp+)
apply (cut_tac i = x and n = "Suc n" and j = "Suc n" in transpos_inj,
simp+)
apply (cut_tac i = x and n = "Suc n" and j = "Suc n" in
transpos_hom, simp+)
apply (rule_tac f = "transpos x (Suc n)" and A = "{j. j \<le> Suc n}" and
B = "{j. j \<le> Suc n}" and g = f and C = H in cmp_inj,
assumption+)
apply (simp add:cmp_def transpos_ij_2)
apply (erule exE, erule conjE)
apply (cut_tac i = x and n = "Suc n" and j = "Suc n" in
transpos_hom, simp+)
apply (simp add:cmp_assoc[THEN sym])
apply (frule_tac f = h and A = "{j. j \<le> Suc n}" and B = "{j. j \<le> Suc n}"
and g = "transpos x (Suc n)" and C = "{j. j \<le> Suc n}" in
cmp_fun, assumption+)
apply (thin_tac "\<forall>f s g t m. f \<in> {j. j \<le> n} \<rightarrow> H \<and> inj_on f {j. j \<le> n} \<and>
s \<in> {j. j \<le> n} \<rightarrow> carrier R \<and> g \<in> {j. j \<le> m} \<rightarrow> H \<and> inj_on g {j. j \<le> m} \<and>
t \<in> {j. j \<le> m} \<rightarrow> carrier R \<and> l_comb R M n s f = l_comb R M m t g \<and>
(\<forall>j\<le>n. s j \<noteq> \<zero>\<^bsub>R\<^esub>) \<and> (\<forall>k\<le>m. t k \<noteq> \<zero>\<^bsub>R\<^esub>) \<longrightarrow> n = m \<and>
(\<exists>h. h \<in> {j. j \<le> n} \<rightarrow> {j. j \<le> n} \<and>
(\<forall>l\<le>n. cmp f h l = g l \<and> cmp s h l = t l))",
thin_tac "\<forall>j\<le>Suc n. s j \<noteq> \<zero>\<^bsub>R\<^esub>", thin_tac "\<forall>k\<le>Suc n. t k \<noteq> \<zero>\<^bsub>R\<^esub>",
thin_tac "{y. \<exists>x\<le>Suc n. y = g x} = {y. \<exists>x\<le>Suc n. y = f x}")
apply blast
done
lemma (in Module) gen_mHom_eq:"\<lbrakk>R module N; generator R M H; f \<in> mHom R M N;
g \<in> mHom R M N; \<forall>h\<in>H. f h = g h \<rbrakk> \<Longrightarrow> f = g"
apply (rule mHom_eq[of N], assumption+)
apply (rule ballI)
apply (unfold generator_def, frule conjunct2, fold generator_def)
apply (frule sym, thin_tac "linear_span R M (carrier R) H = carrier M",
simp,
thin_tac "carrier M = linear_span R M (carrier R) H")
apply (simp add:linear_span_def)
apply (case_tac "H = {}", simp, simp add:mHom_0, simp)
apply (erule exE, (erule bexE)+)
apply (unfold generator_def, frule conjunct1, fold generator_def)
apply (cut_tac sc_Ring, frule Ring.whole_ideal)
apply simp
apply (simp add: linmap_im_lincomb [of "carrier R" N])
apply (frule_tac s = s and n = n and f = "cmp f fa" and g = "cmp g fa" in
Module.linear_comb_eq[of N R "f ` H"])
apply (simp add: mHom_def aHom_def cong del: image_cong)
apply (erule conjE)+
apply (simp add: image_sub [of f "carrier M" "carrier N" H] cong del: image_cong)
apply assumption
apply (rule Pi_I, simp add:cmp_def,
simp add:image_def,
frule_tac f = fa and A = "{j. j \<le> n}" and B = H and x = x in
funcset_mem, simp, blast)
apply (rule Pi_I, simp add:cmp_def,
simp add:image_def,
frule_tac f = fa and A = "{j. j \<le> n}" and B = H and x = x in
funcset_mem, simp, blast)
apply (rule ballI, simp add:cmp_def,
frule_tac f = fa and A = "{j. j \<le> n}" and B = H and x = j in
funcset_mem, simp)
apply blast
apply assumption
done
section "Existence of homomorphism"
definition
fgs :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme, 'a set] \<Rightarrow>
'a set" where
(* free generator span, A is a subset of free generator *)
"fgs R M A = linear_span R M (carrier R) A"
definition
fsp :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme,
('b, 'r, 'm2) Module_scheme, 'a \<Rightarrow> 'b, 'a set, 'a set, 'a \<Rightarrow> 'b] \<Rightarrow> bool" where
"fsp R M N f H A g \<longleftrightarrow> g \<in> mHom R (mdl M (fgs R M A)) N \<and> (\<forall>z\<in>A. f z = g z) \<and> A \<subseteq> H"
(* f \<in> H \<rightarrow> N (not a module hom), free span pair condition *)
definition
fsps :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme,
('b, 'r, 'm2) Module_scheme, 'a \<Rightarrow> 'b, 'a set] \<Rightarrow>
(('a set) * ('a \<Rightarrow> 'b)) set" where
"fsps R M N f H = {Z. \<exists>A g. Z = (A, g) \<and> fsp R M N f H A g}"
(* free span pairs *)
definition
od_fm_fun :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme,
('b, 'r, 'm2) Module_scheme, 'a \<Rightarrow> 'b, 'a set] \<Rightarrow>
(('a set) * ('a \<Rightarrow> 'b)) Order" where
"od_fm_fun R M N f H = \<lparr>carrier = fsps R M N f H,
rel = {Y. Y \<in> (fsps R M N f H) \<times> (fsps R M N f H) \<and>
(fst (fst Y)) \<subseteq> (fst (snd Y))} \<rparr>"
(* ordered set free module with function f *)
lemma (in Module) od_fm_fun_carrier:"carrier (od_fm_fun R M N f H) =
fsps R M N f H"
by (simp add:od_fm_fun_def)
lemma (in Module) fgs_submodule:"a \<subseteq> carrier M \<Longrightarrow>
submodule R M (fgs R M a)"
apply (cut_tac sc_Ring, frule Ring.whole_ideal)
apply (subst fgs_def)
apply (rule linear_span_subModule, assumption, assumption)
done
lemma (in Module) fgs_sub_carrier:"a \<subseteq> carrier M \<Longrightarrow> (fgs R M a) \<subseteq> carrier M"
by (frule fgs_submodule[of a],
rule submodule_subset, assumption+)
lemma (in Module) elem_fgs:"\<lbrakk>a \<subseteq> carrier M; x \<in> a\<rbrakk> \<Longrightarrow> x \<in> fgs R M a"
apply (simp add:fgs_def)
apply (frule l_span_cont_H[of a])
apply (simp add:subsetD)
done
lemma (in Module) fst_chain_subset:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H; (a, b) \<in> C\<rbrakk> \<Longrightarrow> a \<subseteq> carrier M"
apply (subgoal_tac "a \<subseteq> H", frule free_generator_sub)
apply (rule subset_trans, assumption+)
apply (frule_tac A = C and B = "fsps R M N f H" and c = "(a, b)" in subsetD,
assumption+)
apply (simp add:fsps_def fsp_def)
done
lemma (in Module) empty_fsp:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N\<rbrakk> \<Longrightarrow> ({}, (\<lambda>x\<in>{\<zero>\<^bsub>M\<^esub>}. \<zero>\<^bsub>N\<^esub>)) \<in> fsps R M N f H"
apply (simp add:fsps_def fsp_def,
simp add:fgs_def linear_span_def,
cut_tac submodule_0,
frule mdl_is_module[of "{\<zero>}"])
apply (rule Module.mHom_test, assumption+)
apply (rule conjI)
apply (rule Pi_I, simp add:mdl_carrier,
rule Module.module_inc_zero[of N R], assumption+)
apply (simp add:mdl_carrier)
apply (simp add:mdl_def, fold mdl_def)
apply (cut_tac ag_inc_zero, simp add:ag_l_zero,
frule Module.module_is_ag[of N R],
frule aGroup.ag_inc_zero[of N], simp add:aGroup.ag_l_zero[of N])
apply (rule ballI, frule_tac a = a in sc_a_0, simp)
apply (simp add:Module.sc_a_0)
done
lemma (in Module) mem_fgs_l_comb:"\<lbrakk>K \<noteq> {}; K \<subseteq> carrier M; x \<in> fgs R M K\<rbrakk> \<Longrightarrow>
\<exists>(n::nat).
\<exists>f \<in> {j. j \<le> (n::nat)} \<rightarrow> K. \<exists>s \<in> {j. j \<le> n} \<rightarrow> carrier R.
x = l_comb R M n s f"
apply (simp add:fgs_def linear_span_def)
done
lemma PairE_lemma: "\<exists>x y. p = (x, y)" by auto
lemma (in Module) fsps_chain_boundTr1:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H;
\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a; \<forall>a b. (a, b) \<in> C \<longrightarrow>
(a, b) \<in> fsps R M N f H; \<exists>x. (\<exists>b. (x, b) \<in> C) \<and> x \<noteq> {}\<rbrakk> \<Longrightarrow>
fa \<in> {j. j \<le> (n::nat)} \<rightarrow> \<Union>{a. \<exists>b. (a, b) \<in> C}
\<longrightarrow> (\<exists>(c, d) \<in> C. fa ` {j. j \<le> n} \<subseteq> c)"
apply (induct_tac n)
apply (rule impI, simp del: Pi_split_insert_domain)
apply ((erule exE)+, erule conjE, erule exE)
apply (frule_tac f = fa and A = "{0}" and B = "\<Union>{a. \<exists>b. (a, b) \<in> C}" and
?A1.0 = "{0}" in image_sub, simp, simp)
apply blast
apply (rule impI)
apply (frule func_pre, simp)
apply (frule_tac f = fa and A = "{j. j \<le> (Suc n)}" and
B = "\<Union>{a. \<exists>b. (a, b) \<in> C}" and x = "Suc n" in funcset_mem, simp)
apply simp
apply ((erule exE)+, erule conjE, erule exE, erule conjE, erule exE)
apply (erule bexE)
apply (drule_tac x = "(xa, ba)" in bspec, assumption,
drule_tac x = xb in bspec, assumption+)
apply (erule disjE, simp)
apply (subgoal_tac "(\<lambda>(c, d). fa ` {j. j \<le> Suc n} \<subseteq> c) xb")
apply auto[1]
apply (subgoal_tac "(\<lambda>(c, d). fa ` {j. j \<le> Suc n} \<subseteq> c) xb")
apply (blast,
cut_tac p = xb in PairE_lemma,
(erule exE)+, simp, rule subsetI, simp add:image_def,
erule exE, erule conjE,
case_tac "xe = Suc n", simp, simp add:subsetD,
frule_tac m = xe and n = "Suc n" in noteq_le_less, assumption,
frule_tac x = xe and n = n in Suc_less_le,
thin_tac "xe \<le> Suc n", thin_tac "xe < Suc n",
cut_tac c = xd and A = "{y. \<exists>x\<le>n. y = fa x}" and B = xc in
subsetD, simp, simp, blast, assumption)
apply (subgoal_tac "(\<lambda>(c, d). fa ` {j. j \<le> Suc n} \<subseteq> c) (xa, ba)")
apply auto[1]
apply (simp, rule subsetI, simp add:image_def, erule exE, erule conjE,
case_tac "xd = Suc n", simp,
frule_tac m = xd and n = "Suc n" in noteq_le_less, assumption,
frule_tac x = xd and n = n in Suc_less_le,
thin_tac "xd \<le> Suc n", thin_tac "xd < Suc n",
cut_tac p = xb in PairE_lemma, (erule exE)+, simp,
cut_tac c = "fa xd" and A = "{y. \<exists>x\<le>n. y = fa x}" and B = xe in
subsetD, assumption, simp, blast)
apply (rule_tac c = "fa xd" and A = xe and B = xa in subsetD, assumption+)
done
lemma (in Module) fsps_chain_boundTr1_1:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H;
\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a;
\<exists>x. (\<exists>b. (x, b) \<in> C) \<and> x \<noteq> {};
fa \<in> {j. j \<le> (n::nat)} \<rightarrow> \<Union>{a. \<exists>b. (a, b) \<in> C}\<rbrakk> \<Longrightarrow>
\<exists>(c, d) \<in> C. fa ` {j. j \<le> n} \<subseteq> c"
apply (subgoal_tac "\<forall>a b. (a, b) \<in> C \<longrightarrow> (a, b) \<in> fsps R M N f H")
apply (simp add:fsps_chain_boundTr1)
apply ((rule allI)+, rule impI, simp add:subsetD)
done
lemma (in Module) fsps_chain_boundTr1_2:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H;
\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a;
\<exists>x. (\<exists>b. (x, b) \<in> C) \<and> x \<noteq> {};
fa \<in> {j. j \<le> (n::nat)} \<rightarrow> \<Union>{a. \<exists>b. (a, b) \<in> C}\<rbrakk> \<Longrightarrow>
\<exists>P \<in> C. fa ` {j. j \<le> n} \<subseteq> fst P"
apply (frule_tac N = N and f = f and C = C and fa = fa and n = n in
fsps_chain_boundTr1_1, assumption+)
apply (erule bexE)
apply (cut_tac p = x in PairE_lemma, (erule exE)+)
apply (subgoal_tac "fa ` {j. j \<le> n} \<subseteq> fst x")
apply blast
apply simp
done
lemma (in Module) eSum_in_SubmoduleTr:"\<lbrakk>H \<subseteq> carrier M; K \<subseteq> H\<rbrakk> \<Longrightarrow>
f \<in> {j. j \<le> (n::nat)} \<rightarrow> K \<and> s \<in> {j. j \<le> n} \<rightarrow> carrier R \<longrightarrow>
l_comb R (mdl M (fgs R M K)) n s f = l_comb R M n s f"
apply (induct_tac n)
apply (rule impI, (erule conjE)+)
apply (simp add:l_comb_def del: Pi_split_insert_domain, simp add:mdl_def del: Pi_split_insert_domain, rule impI)
apply (simp add:fgs_def del: Pi_split_insert_domain, frule subset_trans[of K H "carrier M"], assumption+)
apply (frule funcset_mem[of f "{0}" K 0], simp del: Pi_split_insert_domain,
frule l_span_cont_H[of K], simp add:subsetD)
apply (rule impI, erule conjE)
apply (frule func_pre[of _ _ K], frule func_pre[of _ _ "carrier R"],
simp)
apply (simp add:l_comb_def, simp add:mdl_def)
apply (rule impI, simp add:fgs_def,frule subset_trans[of K H "carrier M"],
assumption+)
apply (frule_tac x = "Suc n" and A = "{j. j \<le> Suc n}" in
funcset_mem[of f _ K], simp,
frule l_span_cont_H[of K], simp add:subsetD)
done
lemma (in Module) eSum_in_Submodule:"\<lbrakk>H \<subseteq> carrier M; K \<subseteq> H;
f \<in> {j. j \<le> (n::nat)} \<rightarrow> K; s \<in> {j. j \<le> n} \<rightarrow> carrier R\<rbrakk> \<Longrightarrow>
l_comb R (mdl M (fgs R M K)) n s f = l_comb R M n s f"
apply (simp add:eSum_in_SubmoduleTr)
done
lemma (in Module) fgs_generator:"H \<subseteq> carrier M \<Longrightarrow>
generator R (mdl M (fgs R M H)) H"
apply (simp add:generator_def, simp add:mdl_def, fold mdl_def)
apply (rule conjI)
apply (simp add:fgs_def)
apply (simp add:l_span_cont_H)
apply (rule equalityI)
apply (rule subsetI, simp add:linear_span_def)
apply (case_tac "H = {}", simp add:mdl_def)
apply (simp add:fgs_def linear_span_def)
apply simp apply (erule exE, (erule bexE)+)
apply (cut_tac f = f and n = n and s = s in eSum_in_Submodule[of H H],
assumption+, simp, assumption+)
apply simp
apply (simp add:fgs_def)
apply (cut_tac sc_Ring, frule Ring.whole_ideal)
apply (simp add:l_comb_mem_linear_span)
apply (rule subsetI)
apply (case_tac "H = {}", simp)
apply (simp add:fgs_def linear_span_def mdl_def)
apply (frule_tac x = x in mem_fgs_l_comb[of H], assumption+)
apply (erule exE, erule bexE, erule bexE)
apply (simp add:eSum_in_Submodule[THEN sym, of H H])
apply (frule fgs_submodule[of H],
frule mdl_is_module[of "fgs R M H"])
apply (rule Module.l_comb_mem_linear_span[of "mdl M (fgs R M H)" R "carrier R" H], assumption+)
apply (cut_tac sc_Ring, rule Ring.whole_ideal, assumption)
apply (simp add:mdl_carrier, simp add:fgs_def l_span_cont_H,
assumption+)
done
lemma (in Module) fgs_mono:"\<lbrakk>free_generator R M H; J \<subseteq> K; K \<subseteq> H\<rbrakk>
\<Longrightarrow> fgs R M J \<subseteq> fgs R M K" (* H1 \<subseteq> H2 , J \<subseteq> K *)
apply (cut_tac sc_Ring)
apply (case_tac "J = {}")
apply (simp add:fgs_def linear_span_def)
apply (case_tac "K = {}") apply simp apply simp
apply (frule nonempty_ex [of K])
apply (erule exE)
apply (subgoal_tac "(\<lambda>j. x) \<in> {j. j \<le> (0::nat)} \<rightarrow> K",
subgoal_tac "(\<lambda>j. \<zero>\<^bsub>R\<^esub>) \<in> {j. j \<le> (0::nat)} \<rightarrow> carrier R",
subgoal_tac "\<zero>\<^bsub>M\<^esub> = l_comb R M 0 (\<lambda>j. \<zero>\<^bsub>R\<^esub>) (\<lambda>j. x)")
apply blast
apply (simp add:l_comb_def)
apply (frule free_generator_sub[of H],
frule subset_trans[of K H "carrier M"], assumption+,
frule_tac c = x in subsetD[of K "carrier M"], assumption+,
simp add:sc_0_m)
apply (rule Pi_I,
frule Ring.ring_is_ag [of "R"],
simp add:aGroup.ag_inc_zero[of R])
apply (rule Pi_I, assumption)
apply (subgoal_tac "K \<noteq> {}") prefer 2 apply (rule contrapos_pp, simp+)
apply (rule subsetI)
apply (simp add:fgs_def, simp add:linear_span_def)
apply (erule exE, (erule bexE)+)
apply (frule_tac f = f and A = "{j. j \<le> n}" and B = J and ?B1.0 = K in
extend_fun, assumption+)
apply blast
done
lemma (in Module) empty_fgs:"fgs R M {} = {\<zero>}"
by(simp add:fgs_def linear_span_def)
lemma (in Module) mem_fsps_snd_mHom:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; (a, b) \<in> fsps R M N f H\<rbrakk> \<Longrightarrow>
b \<in> mHom R (mdl M (fgs R M a)) N"
by (simp add:fsps_def fsp_def)
(*
lemma (in Module) mem_fsps_snd_mHom1:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; (a, b) \<in> fsps R M N f H\<rbrakk> \<Longrightarrow>
b \<in> mHom R (mdl M (fgs R M a)) N"
by (simp add:fsps_def fsp_def) *)
lemma (in Module) mem_fsps_fst_sub:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; (a, b) \<in> fsps R M N f H\<rbrakk> \<Longrightarrow> a \<subseteq> H"
by (simp add:fsps_def fsp_def)
lemma (in Module) mem_fsps_fst_sub1:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; (a, b) \<in> fsps R M N f H\<rbrakk> \<Longrightarrow> a \<subseteq> carrier M"
apply (simp add:fsps_def fsp_def)
apply (frule free_generator_sub[of H],
rule subset_trans[of a H "carrier M"], simp+)
done
lemma (in Module) Order_od_fm_fun:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N\<rbrakk> \<Longrightarrow> Order (od_fm_fun R M N f H)"
apply (rule Order.intro)
apply (simp add:od_fm_fun_carrier)
apply (rule subsetI, simp add:od_fm_fun_def)
apply (simp add:od_fm_fun_def)
apply (simp add:od_fm_fun_def)
apply (cut_tac p = a in PairE_lemma, (erule exE)+,
cut_tac p = b in PairE_lemma, (erule exE)+, simp)
apply (frule_tac a = x and b = y in mem_fsps_snd_mHom[of N H f], assumption+,
frule_tac a = x and b = ya in mem_fsps_snd_mHom[of N H f], assumption+)
apply (frule_tac a = x and b = y in mem_fsps_fst_sub[of N H f], assumption+,
frule free_generator_sub,
frule_tac A = x and B = H and C = "carrier M" in subset_trans,
assumption+)
apply (frule_tac a = x in fgs_submodule,
frule_tac H = "fgs R M x" in mdl_is_module)
apply (thin_tac "a = (x, y)", thin_tac "b = (x, ya)")
apply (simp add:mdl_carrier)
apply (frule_tac H = x in fgs_generator)
apply (frule_tac R = R and M = "mdl M (fgs R M x)" and N = N and H = x
and f = y and g = ya in Module.gen_mHom_eq, assumption+,
simp add:fsps_def fsp_def, assumption)
apply (simp add:od_fm_fun_def)
done
lemma (in Module) fsps_chain_boundTr2_1:"\<lbrakk>R module N;
free_generator R M H; f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H;
(a, b) \<in> C; (aa, ba) \<in> C; x \<in> fgs R M a; x \<in> fgs R M aa; a \<subseteq> aa\<rbrakk>
\<Longrightarrow> b x = ba x"
apply (cut_tac sc_Ring, frule Ring.whole_ideal, frule free_generator_sub)
apply (frule mem_fsps_snd_mHom[of N H f a b], assumption+, simp add:subsetD,
frule mem_fsps_snd_mHom[of N H f aa ba], assumption+, simp add:subsetD)
apply (case_tac "a = {}", simp add:empty_fgs)
apply (cut_tac submodule_0, frule mdl_is_module[of "{\<zero>}"],
frule subsetD[of C "fsps R M N f H" "(aa, ba)"], assumption+)
apply (cut_tac fgs_submodule[of aa],
frule mdl_is_module[of "fgs R M aa"])
apply (frule Module.mHom_0[of "mdl M {\<zero>}" R N b], assumption+,
frule Module.mHom_0[of "mdl M (fgs R M aa)" R N ba], assumption+,
simp add:mdl_def)
apply (frule mem_fsps_fst_sub[of N H f aa ba], assumption+,
rule subset_trans[of aa H "carrier M"], assumption+)
apply (frule nonempty_ex[of a], erule exE,
frule_tac c = xa in subsetD[of a aa], assumption,
frule_tac x = xa in nonempty[of _ aa],
thin_tac "xa \<in> a", thin_tac "xa \<in> aa")
apply (frule mem_fsps_fst_sub[of N H f aa ba], assumption+,
simp add:subsetD)
apply (frule mem_fgs_l_comb[of a x],
frule subset_trans[of a aa H], assumption+,
rule subset_trans[of a H "carrier M"], assumption+)
apply (erule exE, (erule bexE)+)
apply (frule_tac f = fa and A = "{j. j \<le> n}" and B = a and ?B1.0 = aa in
extend_fun, assumption+)
apply (frule subsetD[of C "fsps R M N f H" "(a, b)"], assumption+,
frule subsetD[of C "fsps R M N f H" "(aa, ba)"], assumption+,
thin_tac "C \<subseteq> fsps R M N f H")
apply (cut_tac fgs_submodule[of a],
frule mdl_is_module[of "fgs R M a"],
cut_tac fgs_submodule[of aa],
frule mdl_is_module[of "fgs R M aa"]) (*
apply (simp add:fsps_def fsp_def, (erule conjE)+) *)
apply (frule subset_trans[of a aa H], assumption+)
apply (frule_tac f1 = fa and n1 = n and s1 = s in eSum_in_Submodule[THEN sym,
of H a], assumption+,
frule_tac s = s and n = n and g = fa in Module.linmap_im_lincomb[of
"mdl M (fgs R M a)" R "carrier R" N _ a], assumption+,
simp add:mdl_carrier,
subst fgs_def, rule_tac l_span_cont_H,
frule subset_trans[of a H "carrier M"], assumption+, simp,
thin_tac "b (l_comb R (mdl M (fgs R M a)) n s fa) =
l_comb R N n s (cmp b fa)",
rotate_tac -1, frule sym,
thin_tac "l_comb R M n s fa = l_comb R (mdl M (fgs R M a)) n s fa",
simp,
thin_tac "l_comb R (mdl M (fgs R M a)) n s fa = l_comb R M n s fa")
apply (frule_tac f1 = fa and n1 = n and s1 = s in eSum_in_Submodule[THEN sym,
of H aa], assumption+,
frule_tac s = s and n = n and g = fa in Module.linmap_im_lincomb[of
"mdl M (fgs R M aa)" R "carrier R" N _ aa], assumption+,
simp add:mdl_carrier,
subst fgs_def, rule_tac l_span_cont_H,
frule subset_trans[of a H "carrier M"], assumption+, simp+,
thin_tac "l_comb R M n s fa = l_comb R (mdl M (fgs R M aa)) n s fa",
thin_tac "ba (l_comb R (mdl M (fgs R M aa)) n s fa) =
l_comb R N n s (cmp ba fa)")
apply (simp add:fsps_def fsp_def)
apply (rule_tac s = s and n = n and f = "cmp b fa" and g = "cmp ba fa" in
Module.linear_comb_eq[of N R "f ` H"], assumption+,
simp add:image_sub0, assumption)
apply (rule Pi_I, simp add:cmp_def,
frule_tac x = xa and f = fa and A = "{j. j \<le> n}" and B = a in
funcset_mem, simp,
frule_tac m = "fa xa" in Module.mHom_mem[of "mdl M (fgs R M a)" R N b],
assumption+,
simp add:mdl_carrier fgs_def, rule_tac c = "fa xa" in
subsetD[of a "linear_span R M (carrier R) a"],
rule l_span_cont_H,
frule subset_trans[of a aa H], assumption+,
rule subset_trans[of a H "carrier M"], assumption+,
drule_tac x = "fa xa" in bspec, assumption,
frule_tac c = "fa xa" in subsetD[of a H], assumption+,
rotate_tac -2, frule sym, thin_tac "f (fa xa) = b (fa xa)", simp)
apply (rule Pi_I, simp add:cmp_def,
frule_tac x = xa and f = fa and A = "{j. j \<le> n}" and B = aa in
funcset_mem, simp,
frule_tac c = "fa xa" in subsetD[of aa H], assumption+,
drule_tac x = "fa xa" in bspec, assumption,
rotate_tac -1, frule sym, thin_tac "f (fa xa) = ba (fa xa)",
simp)
apply (rule ballI, simp add:cmp_def,
frule_tac x = j and f = fa and A = "{j. j \<le> n}" and B = a in
funcset_mem, simp)
apply (frule_tac x = "fa j" in bspec, assumption,
thin_tac "\<forall>z\<in>a. f z = b z",
frule_tac c = "fa j" in subsetD[of a aa], assumption+,
frule_tac x = "fa j" in bspec, assumption,
thin_tac "\<forall>z\<in>aa. f z = ba z")
apply (rotate_tac -3, frule sym,
thin_tac "f (fa j) = b (fa j)", frule sym,
thin_tac "f (fa j) = ba (fa j)", simp)
apply (simp add:subset_trans[of aa H "carrier M"])
apply (frule subset_trans[of a aa H], assumption+,
simp add:subset_trans[of a H "carrier M"])
done
lemma (in Module) fsps_chain_boundTr2_2:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H;
\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a; C \<noteq> {}; (a, b) \<in> C;
x \<in> fgs R M a; (a1, b1) \<in> C; x \<in> fgs R M a1\<rbrakk> \<Longrightarrow> b x = b1 x"
apply (frule_tac x = "(a, b)" in bspec, assumption,
thin_tac "\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a",
frule_tac x = "(a1, b1)" in bspec, assumption,
thin_tac "\<forall>ba\<in>C. fst (a, b) \<subseteq> fst ba \<or> fst ba \<subseteq> fst (a, b)", simp)
apply (erule disjE)
apply (rule fsps_chain_boundTr2_1[of N H f C a b a1 b1], assumption+)
apply (frule fsps_chain_boundTr2_1[of N H f C a1 b1 a b], assumption+)
apply (rule sym, assumption)
done
lemma (in Module) fsps_chain_boundTr2:"\<And>x. \<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H;
\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a;
x \<in> (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})); C \<noteq> {}\<rbrakk> \<Longrightarrow>
(THE y. y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and> x \<in> fgs R M a \<and> y = b x)) \<in>
(carrier N) \<and>
(\<exists>a1 b1. (a1, b1) \<in> C \<and> x \<in> fgs R M a1 \<and>
(THE y. y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and> x \<in> fgs R M a \<and> y = b x)) =
b1 x)"
apply (cut_tac sc_Ring, frule Ring.whole_ideal[of R])
apply (case_tac "\<Union>{a. \<exists>b. (a, b) \<in> C} = {}", simp add: empty_fgs,
frule nonempty_ex[of C], erule exE,
cut_tac p = xa in PairE_lemma, (erule exE)+,
frule_tac c = xa in subsetD[of C "fsps R M N f H"], assumption+, simp)
apply (frule_tac a = xb and b = y in mem_fsps_snd_mHom[of N H f],
assumption+, simp add:subsetD)
apply (rename_tac a b a1 b1)
apply (subgoal_tac "\<exists>!y. y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and>
\<zero> \<in> (fgs R M a) \<and> y = b \<zero>)",
subgoal_tac "(THE y. y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and>
\<zero> \<in> fgs R M a \<and> y = b \<zero>)) \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and>
\<zero> \<in> fgs R M a \<and>
(THE y. y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and> \<zero> \<in> fgs R M a \<and> y = b \<zero>)) =
b \<zero>)")
prefer 2
apply (rule theI', simp, simp)
apply (rule ex_ex1I)
apply (cut_tac a = a1 in fgs_submodule)
apply (frule_tac a = a1 and b = b1 in mem_fsps_fst_sub[of N H f],
assumption+, simp add:subsetD, frule free_generator_sub[of H],
rule subset_trans[of _ H "carrier M"], assumption+,
frule_tac H = "fgs R M a1" in mdl_is_module,
frule_tac M = "mdl M (fgs R M a1)" and f = b1 in
Module.mHom_0[of _ R N], assumption+,
frule_tac H = "fgs R M a1" in submodule_inc_0,
frule Module.module_inc_zero[of N],
thin_tac "b1 \<in> mHom R (mdl M (fgs R M a1)) N",
thin_tac "R module mdl M (fgs R M a1)",
simp add:mdl_def)
apply (rotate_tac -3, frule sym, thin_tac "b1 \<zero> = \<zero>\<^bsub>N\<^esub>", blast)
apply ((erule conjE)+, (erule exE)+, (erule conjE)+)
apply (cut_tac a = aa in fgs_submodule,
frule_tac a = aa and b = ba in mem_fsps_fst_sub[of N H f],
assumption+, simp add:subsetD,
frule free_generator_sub, rule subset_trans, assumption+,
simp add:subsetD,
frule_tac H = "fgs R M aa" in mdl_is_module,
frule_tac M = "mdl M (fgs R M aa)" and f = ba in
Module.mHom_0[of _ R N], assumption+,
frule_tac H = "fgs R M aa" in submodule_inc_0,
frule_tac c = "(aa, ba)" in subsetD[of C "fsps R M N f H"],
assumption+, simp,
frule_tac a = aa and b = ba in mem_fsps_snd_mHom[of N H f],
assumption+, simp add:mdl_def)
apply (cut_tac a = ab in fgs_submodule,
frule_tac a = ab and b = bb in mem_fsps_fst_sub[of N H f],
assumption+, simp add:subsetD,
frule free_generator_sub, rule subset_trans[of _ H "carrier M"],
assumption+,
frule_tac H = "fgs R M ab" in mdl_is_module,
frule_tac M = "mdl M (fgs R M ab)" and f = bb in
Module.mHom_0[of _ R N], assumption+)
apply (fold mdl_def,
frule_tac c = "(ab, bb)" in subsetD[of C "fsps R M N f H"],
assumption+,
frule_tac a = ab and b = bb in mem_fsps_snd_mHom[of N H f],
assumption+, simp add:mdl_def)
(*********** case \<Union>{a. \<exists>b. (a, b) \<in> C} = {} done **********)
apply (subgoal_tac "\<exists>!y. y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and>
x \<in> (fgs R M a) \<and> y = b x)",
subgoal_tac "(THE y. y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and>
x \<in> fgs R M a \<and> y = b x)) \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and>
x \<in> fgs R M a \<and>
(THE y. y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and> x \<in> fgs R M a \<and> y = b x)) =
b x)")
prefer 2
apply (rule theI', simp, simp)
apply (rule ex_ex1I)
apply (frule_tac K = "\<Union>{a. \<exists>b. (a, b) \<in> C}" and x = x in mem_fgs_l_comb)
apply (rule subsetI, simp,
thin_tac "\<exists>x. (\<exists>b. (x, b) \<in> C) \<and> x \<noteq> {}",
erule exE, erule conjE, erule exE,
frule_tac a = xb and b = b in mem_fsps_fst_sub[of N H f],
assumption+, simp add:subsetD,
frule free_generator_sub[of H], simp add:subsetD, assumption)
apply (erule exE, (erule bexE)+)
apply (frule_tac fa = fa and n = n in fsps_chain_boundTr1_1[of N H f C],
assumption+)
apply (frule_tac A = "\<Union>{a. \<exists>b. (a, b) \<in> C}" in nonempty_ex,
erule exE, simp, assumption+, erule bexE,
cut_tac p = xa in PairE_lemma, (erule exE)+, simp)
apply (frule_tac a = xb and b = y in mem_fsps_fst_sub[of N H f],
assumption+, simp add:subsetD,
frule free_generator_sub[of H],
frule_tac A = xb in subset_trans[of _ H "carrier M"], assumption+)
apply (frule_tac H = xb and s = s and n = n and f = fa in
l_comb_mem_linear_span[of "carrier R"], assumption+,
rule Pi_I,
frule_tac a = xc and A = "{j. j \<le> n}" and f = fa in
mem_in_image2, simp add:image_def, erule exE, blast,
fold fgs_def)
apply (frule_tac a = xb and b = y in mem_fsps_snd_mHom[of N H f],
assumption+, simp add:subsetD,
frule_tac a = xb in fgs_submodule,
frule_tac H = "fgs R M xb" in mdl_is_module,
frule_tac R = R and M = "mdl M (fgs R M xb)" and N = N and
f = y and m = x in Module.mHom_mem, assumption+,
simp add:mdl_carrier, simp, blast)
apply ((erule conjE)+, (erule exE)+, (erule conjE)+)
apply (frule_tac x = "(a, b)" in bspec, assumption,
rotate_tac -1,
frule_tac x = "(aa, ba)" in bspec, assumption,
thin_tac "\<forall>ba\<in>C. fst (a, b) \<subseteq> fst ba \<or> fst ba \<subseteq> fst (a, b)",
simp)
apply (rule_tac a = a and b = b and ?a1.0 = aa and ?b1.0 = ba and x = x in
fsps_chain_boundTr2_2[of N H f C], assumption+)
done
lemma (in Module) Un_C_submodule:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H; C \<noteq> {};
\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a\<rbrakk> \<Longrightarrow>
submodule R M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))"
apply (cut_tac sc_Ring, frule free_generator_sub[of H],
frule Ring.whole_ideal)
apply (subst fgs_def)
apply (rule linear_span_subModule, assumption+)
apply (rule subsetI, simp)
apply (erule exE, erule conjE, erule exE,
frule_tac c = "(xa, b)" in subsetD[of C "fsps R M N f H"], assumption+,
frule_tac a = xa and b = b in mem_fsps_fst_sub[of N H f], assumption+,
frule free_generator_sub, frule_tac A = xa in
subset_trans[of _ H "carrier M"], assumption+,
simp add:subsetD)
done
lemma (in Module) Un_C_fgs_sub:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H; C \<noteq> {};
\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a\<rbrakk> \<Longrightarrow>
\<Union>{a. \<exists>b. (a, b) \<in> C} \<subseteq> fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})"
apply (simp add:fgs_def)
apply (rule l_span_cont_H)
apply (frule Un_C_submodule[of N H f C], assumption+)
apply (rule subsetI, simp, erule exE, erule conjE, erule exE)
apply (frule_tac a = xa and b = b in mem_fsps_fst_sub[of N H f], assumption+,
simp add:subsetD,
frule free_generator_sub[of H],
frule_tac A = xa and B = H and C = "carrier M" in subset_trans,
assumption+, simp add:subsetD)
done
lemma (in Module) Chain_3_supset:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H; C \<noteq> {};
\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a; (a1, b1) \<in> C; (a2, b2) \<in> C;
(a3, b3) \<in> C\<rbrakk> \<Longrightarrow> \<exists>(g, h)\<in>C. a1 \<subseteq> g \<and> a2 \<subseteq> g \<and> a3 \<subseteq> g"
proof -
assume "(a1, b1) \<in> C"
"(a2, b2) \<in> C"
"(a3, b3) \<in> C"
then have A: "a1 \<in> fst ` C"
"a2 \<in> fst ` C"
"a3 \<in> fst ` C"
by (simp_all add: image_iff) force+
assume "\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a"
with A have "a1 \<subseteq> a2 \<or> a2 \<subseteq> a1"
"a1 \<subseteq> a3 \<or> a3 \<subseteq> a1"
"a2 \<subseteq> a3 \<or> a3 \<subseteq> a2"
by auto
with A have "\<exists>a\<in>fst ` C. a1 \<subseteq> a \<and> a2 \<subseteq> a \<and> a3 \<subseteq> a"
by auto
then obtain a where "a \<in> fst ` C" and C: "a1 \<subseteq> a" "a2 \<subseteq> a" "a3 \<subseteq> a" by blast
then obtain b where "(a, b) \<in> C" by auto
with C show "\<exists>(a, b)\<in>C. a1 \<subseteq> a \<and> a2 \<subseteq> a \<and> a3 \<subseteq> a"
by auto
qed
lemma (in Module) fsps_chain_bound1:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H;
\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a; C \<noteq> {}\<rbrakk> \<Longrightarrow>
(\<lambda>x\<in>(fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})). (THE y. y \<in> carrier N \<and>
(\<exists>a b. (a, b) \<in> C \<and> x \<in> fgs R M a \<and> y = b x))) \<in>
mHom R (mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))) N"
apply (cut_tac sc_Ring, frule Ring.whole_ideal)
apply (frule Un_C_submodule[of N H f C], assumption+)
apply (frule mdl_is_module[of "fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})"])
apply (rule Module.mHom_test[of "mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))" R N
"\<lambda>x\<in>fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}). THE y. y \<in> carrier N \<and>
(\<exists>a b. (a, b) \<in> C \<and> x \<in> fgs R M a \<and> y = b x)"], assumption+)
apply (rule conjI)
apply (rule Pi_I,
frule submodule_subset[of "fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})"],
simp add:mdl_carrier, simp add:fsps_chain_boundTr2)
apply (rule conjI)
apply (simp add:mdl_carrier)
apply (rule conjI)
apply (rule ballI)+
apply (frule Module.module_is_ag[of "mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))"
R])
apply (frule_tac x = m and y = n in aGroup.ag_pOp_closed[of
"mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))"], assumption+)
apply (frule_tac x = m in fsps_chain_boundTr2[of N H f C], assumption+,
simp add:mdl_carrier, assumption, erule conjE, (erule exE)+,
(erule conjE)+,
frule_tac x = n in fsps_chain_boundTr2[of N H f C], assumption+,
simp add:mdl_carrier, assumption, erule conjE, (erule exE)+,
(erule conjE)+,
frule_tac x = "m \<plusminus>\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> n" in
fsps_chain_boundTr2[of N H f C], assumption+,
simp add:mdl_carrier, assumption, erule conjE, (erule exE)+,
(erule conjE)+) apply (simp add:mdl_carrier)
apply (thin_tac "(THE y.
y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and> m \<in> fgs R M a \<and> y = b m)) =
b1 m") apply (
thin_tac "(THE y.
y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and> n \<in> fgs R M a \<and> y = b n)) =
b1a n") apply (
thin_tac "(THE y.
y \<in> carrier N \<and>
(\<exists>a b. (a, b) \<in> C \<and>
m \<plusminus>\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> n \<in> fgs R M a \<and>
y = b (m \<plusminus>\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> n))) =
b1b (m \<plusminus>\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> n)")
apply (thin_tac "m \<plusminus>\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> n
\<in> fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})",
thin_tac "b1b (m \<plusminus>\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> n) \<in> carrier N",
simp add:mdl_def, fold mdl_def)
apply (cut_tac ?a1.0 = a1 and ?b1.0 = b1 and ?a2.0 = a1a and ?b2.0 = b1a and
?a3.0 = a1b and ?b3.0 = b1b in Chain_3_supset[of N H f C], assumption+)
apply (erule bexE)
apply (cut_tac p = x in PairE_lemma, (erule exE)+, simp, (erule conjE)+)
apply (frule_tac J = a1b and K = xa in fgs_mono[of H], assumption+,
rule_tac a = xa and b = y in mem_fsps_fst_sub[of N H f], assumption+,
simp add:subsetD,
frule_tac J = a1 and K = xa in fgs_mono[of H], assumption+,
rule_tac a = xa and b = y in mem_fsps_fst_sub[of N H f], assumption+,
simp add:subsetD,
frule_tac J = a1a and K = xa in fgs_mono[of H], assumption+,
rule_tac a = xa and b = y in mem_fsps_fst_sub[of N H f], assumption+,
simp add:subsetD)
apply (frule_tac c = m and A = "fgs R M a1" and B = "fgs R M xa" in subsetD,
assumption+,
frule_tac c = n and A = "fgs R M a1a" and B = "fgs R M xa" in subsetD,
assumption+,
frule_tac c = "m \<plusminus> n" and A = "fgs R M a1b" and B = "fgs R M xa" in
subsetD, assumption+)
apply (frule_tac a = a1 and b = b1 and aa = xa and ba = y in
fsps_chain_boundTr2_1[of N H f C], assumption+, simp,
frule_tac a = a1a and b = b1a and aa = xa and ba = y in
fsps_chain_boundTr2_1[of N H f C], assumption+, simp,
frule_tac a = a1b and b = b1b and aa = xa and ba = y in
fsps_chain_boundTr2_1[of N H f C], assumption+, simp)
apply (thin_tac "submodule R M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))",
thin_tac "R module mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))",
thin_tac "m \<in> fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})",
thin_tac "n \<in> fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})",
thin_tac "aGroup (mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})))",
thin_tac "y m \<in> carrier N", thin_tac "(a1, b1) \<in> C",
thin_tac "m \<in> fgs R M a1", thin_tac "y n \<in> carrier N",
thin_tac "(a1a, b1a) \<in> C", thin_tac "n \<in> fgs R M a1a",
thin_tac "(a1b, b1b) \<in> C", thin_tac "m \<plusminus> n \<in> fgs R M a1b",
thin_tac "a1 \<subseteq> xa", thin_tac "a1a \<subseteq> xa",
thin_tac "a1b \<subseteq> xa", thin_tac "fgs R M a1 \<subseteq> fgs R M xa",
thin_tac "fgs R M a1a \<subseteq> fgs R M xa",
thin_tac "b1b (m \<plusminus> n) = y (m \<plusminus> n)",
thin_tac "b1 m = y m", thin_tac "b1a n = y n")
apply (cut_tac a = xa in fgs_submodule,
frule_tac c = "(xa, y)" in subsetD[of C "fsps R M N f H"],
assumption+,
simp add:mem_fsps_fst_sub1)
apply (frule_tac H = "fgs R M xa" in mdl_is_module)
apply (frule_tac a = xa and b = y in mem_fsps_snd_mHom[of N H f],
assumption+, simp add:subsetD)
apply (frule_tac R = R and M = "mdl M (fgs R M xa)" and N = N and f = y and
m = m and n = n in Module.mHom_add, assumption+,
simp add:mdl_carrier, simp add:mdl_carrier, simp add:mdl_def)
(************** mHom_add ***********************)
apply (rule ballI)+
apply (frule_tac a = a and m = m in Module.sc_mem[of
"mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))" R], assumption+)
apply (frule_tac x = m in fsps_chain_boundTr2[of N H f C], assumption+,
simp add:mdl_carrier, assumption, erule conjE, (erule exE)+,
(erule conjE)+,
frule_tac x = "a \<cdot>\<^sub>s\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> m" in
fsps_chain_boundTr2[of N H f C], assumption+,
simp add:mdl_carrier, assumption, erule conjE, (erule exE)+,
(erule conjE)+)
apply (simp add:mdl_carrier)
apply (thin_tac "(THE y.
y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and> m \<in> fgs R M a \<and> y = b m)) =
b1 m",
thin_tac "b1a (a \<cdot>\<^sub>s\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> m) \<in> carrier N",
thin_tac "(THE y.
y \<in> carrier N \<and>
(\<exists>aa b.
(aa, b) \<in> C \<and>
a \<cdot>\<^sub>s\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> m \<in> fgs R M aa \<and>
y = b (a \<cdot>\<^sub>s\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> m))) =
b1a (a \<cdot>\<^sub>s\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> m)",
thin_tac "a \<cdot>\<^sub>s\<^bsub>mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))\<^esub> m
\<in> fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})",
thin_tac "submodule R M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))",
thin_tac "R module mdl M (fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}))")
apply (simp add:mdl_def,
thin_tac "m \<in> fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})",
thin_tac "b1 m \<in> carrier N")
apply (frule_tac x = "(a1, b1)" in bspec, assumption,
rotate_tac -1,
drule_tac x = "(a1a, b1a)" in bspec, assumption,
simp)
apply (erule disjE,
frule_tac J = a1 and K = a1a in fgs_mono[of H], assumption+,
rule_tac a = a1a and b = b1a in mem_fsps_fst_sub[of N H f],
assumption+, simp add:subsetD,
frule_tac c = m and A = "fgs R M a1" and B = "fgs R M a1a" in subsetD,
assumption+)
apply (cut_tac a = a1a in fgs_submodule,
frule_tac c = "(a1a, b1a)" in subsetD[of C "fsps R M N f H"],
assumption+, simp add:mem_fsps_fst_sub1,
frule_tac H = "fgs R M a1a" in mdl_is_module,
frule_tac a = a1a and b = b1a in mem_fsps_snd_mHom[of N H f],
assumption+, simp add:subsetD)
apply (frule_tac R = R and M = "mdl M (fgs R M a1a)" and N = N and m = m
and f = b1a and a = a in Module.mHom_lin, assumption+,
simp add:mdl_carrier, assumption+, simp add:mdl_def,
fold mdl_def)
apply (frule_tac a = a1 and b = b1 and ?a1.0 = a1a and ?b1.0 = b1a and x = m
in fsps_chain_boundTr2_2[of N H f C], assumption+, simp)
apply (frule_tac J = a1a and K = a1 in fgs_mono[of H], assumption+,
rule_tac a = a1 and b = b1 in mem_fsps_fst_sub[of N H f],
assumption+, simp add:subsetD,
frule_tac c = "a \<cdot>\<^sub>s m" and A = "fgs R M a1a" and B = "fgs R M a1" in
subsetD, assumption+)
apply (frule_tac a = a1a and b = b1a and ?a1.0 = a1 and ?b1.0 = b1 and
x = "a \<cdot>\<^sub>s m" in fsps_chain_boundTr2_2[of N H f C], assumption+, simp)
apply (cut_tac a = a1 in fgs_submodule,
frule_tac c = "(a1, b1)" in subsetD[of C "fsps R M N f H"],
assumption+, simp add:mem_fsps_fst_sub1,
frule_tac H = "fgs R M a1" in mdl_is_module,
frule_tac a = a1 and b = b1 in mem_fsps_snd_mHom[of N H f],
assumption+, simp add:subsetD)
apply (frule_tac R = R and M = "mdl M (fgs R M a1)" and N = N and m = m
and f = b1 and a = a in Module.mHom_lin, assumption+,
simp add:mdl_carrier, assumption+, simp add:mdl_def)
done
lemma (in Module) fsps_chain_bound2:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; C \<subseteq> fsps R M N f H; C \<noteq> {};
\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a\<rbrakk> \<Longrightarrow>
\<forall>y\<in>(\<Union>{a. \<exists>b. (a, b) \<in> C}). (\<lambda>x\<in>(fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})).
(THE y. y \<in> carrier N \<and> (\<exists>a b. (a, b) \<in> C \<and> x \<in> fgs R M a \<and> y = b x))) y =
f y"
apply (rule ballI)
apply (subgoal_tac "y \<in> fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})")
prefer 2
apply simp
apply (erule exE, erule conjE, erule exE)
apply (frule Un_C_fgs_sub[of N H f C], assumption+)
apply (rule_tac c = y in subsetD[of "\<Union>{a. \<exists>b. (a, b) \<in> C}"
"fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})"], assumption+)
apply (simp, blast)
apply simp
apply (frule_tac x = y in fsps_chain_boundTr2[of N H f C], assumption+)
apply (erule conjE, (erule exE)+, erule conjE, erule exE, (erule conjE)+)
apply (simp,
thin_tac "(THE ya. ya \<in> carrier N \<and>
(\<exists>a b. (a, b) \<in> C \<and> y \<in> fgs R M a \<and> ya = b y)) = b1 y")
apply (frule_tac c = "(x, b)" in subsetD[of C "fsps R M N f H"],
assumption+,
simp add:fsps_def fsp_def, (erule conjE)+,
thin_tac "\<forall>z\<in>x. f z = b z")
apply (frule_tac x = y and a = x and b = b and ?a1.0 = a1 and ?b1.0 = b1 in
fsps_chain_boundTr2_2[of N H f C], assumption+,
simp add:fsps_def fsp_def, assumption+, subst fgs_def,
rule_tac c = y and A = x and B = "linear_span R M (carrier R) x"
in subsetD,
rule l_span_cont_H,
frule free_generator_sub[of H], rule subset_trans[of _ H "carrier M"],
assumption+, simp)
done
lemma (in Module) od_fm_fun_Chain:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; Algebra2.Chain (od_fm_fun R M N f H) C; C \<noteq> {}\<rbrakk> \<Longrightarrow>
\<forall>a\<in>C. \<forall>b\<in>C. fst a \<subseteq> fst b \<or> fst b \<subseteq> fst a"
apply (frule Order_od_fm_fun[of N H f], assumption+)
apply (rule ballI)+
apply (simp add:Algebra2.Chain_def, erule conjE, simp add:Torder_def, erule conjE,
simp add:Torder_axioms_def)
apply (simp add:Order.Iod_carrier)
apply (auto simp add: Iod_def od_fm_fun_def ole_def)
apply blast
done
lemma (in Module) od_fm_fun_inPr0:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; Algebra2.Chain (od_fm_fun R M N f H) C; C \<noteq> {};
\<exists>b. (y, b) \<in> C; z \<in> y\<rbrakk> \<Longrightarrow> z \<in> fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})"
apply (frule Un_C_fgs_sub[of N H f C], assumption+)
apply (simp add:Algebra2.Chain_def od_fm_fun_def, assumption)
apply (simp add:od_fm_fun_Chain)
apply (rule_tac subsetD[of "\<Union>{a. \<exists>b. (a, b) \<in> C}"
"fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C})" z], assumption+)
apply (simp, blast)
done
lemma (in Module) od_fm_fun_indPr1:" \<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; Algebra2.Chain (od_fm_fun R M N f H) C; C \<noteq> {}\<rbrakk> \<Longrightarrow>
(\<Union>{a. \<exists>b. (a, b) \<in> C},
\<lambda>x \<in> fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}). THE y. y \<in> carrier N \<and>
(\<exists>a b. (a, b) \<in> C \<and> x \<in> fgs R M a \<and> y = b x)) \<in> fsps R M N f H"
apply (simp add:fsps_def fsp_def, rule conjI)
apply (rule fsps_chain_bound1[of N H f C], assumption+)
apply (simp add:od_fm_fun_def Algebra2.Chain_def)
apply (simp add:od_fm_fun_Chain, assumption)
apply (rule conjI)
apply (rule allI, rule impI)
apply (rule ballI)
apply (simp add:od_fm_fun_inPr0[of N H f C])
apply (frule fsps_chain_bound2[of N H f C], assumption+,
simp add:Algebra2.Chain_def od_fm_fun_def, assumption)
apply (simp add:od_fm_fun_Chain)
apply (drule_tac x = z in bspec)
apply (simp, blast)
apply (simp add:od_fm_fun_inPr0)
apply (rule subsetI, simp, erule exE, erule conjE, erule exE)
apply (frule Order_od_fm_fun[of N H f], assumption+,
frule Order.Chain_sub[of "od_fm_fun R M N f H" C], assumption,
thin_tac "Algebra2.Chain (od_fm_fun R M N f H) C",
thin_tac "Order (od_fm_fun R M N f H)")
apply (simp add:od_fm_fun_def)
apply (frule_tac a = xa and b = b in mem_fsps_fst_sub[of N H f], assumption+,
simp add:subsetD, simp add:subsetD)
done
lemma (in Module) od_fm_fun_indPr2:" \<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; Algebra2.Chain (od_fm_fun R M N f H) C; C \<noteq> {}\<rbrakk> \<Longrightarrow>
ub\<^bsub>od_fm_fun R M N f H\<^esub> C (\<Union>{a. \<exists>b. (a, b) \<in> C},
\<lambda>x \<in> fgs R M (\<Union>{a. \<exists>b. (a, b) \<in> C}). THE y. y \<in> carrier N \<and>
(\<exists>a b. (a, b) \<in> C \<and> x \<in> fgs R M a \<and> y = b x))"
apply (frule Order_od_fm_fun[of N H f], assumption+,
simp add:upper_bound_def)
apply (simp add:od_fm_fun_def, fold od_fm_fun_def,
frule od_fm_fun_indPr1[of N H f C], assumption+, simp)
apply (rule ballI,
subst od_fm_fun_def, subst ole_def, simp)
apply (rule conjI)
apply (frule Order.Chain_sub[of "od_fm_fun R M N f H" C], assumption,
simp add:od_fm_fun_carrier, simp add:subsetD)
apply (unfold split_paired_all,
rule subsetI, simp, blast)
done
lemma (in Module) od_fm_fun_inductive:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N\<rbrakk> \<Longrightarrow> inductive_set (od_fm_fun R M N f H)"
apply (frule Order_od_fm_fun[of N H f], assumption+)
apply (simp add:inductive_set_def)
apply (rule allI, rule impI)
apply (case_tac "C \<noteq> {}")
apply (frule_tac C = C in od_fm_fun_indPr2[of N H f], assumption+)
apply blast
apply simp
apply (simp add:upper_bound_def)
apply (subst od_fm_fun_def, simp)
apply (frule empty_fsp[of N H f], assumption+, blast)
done
lemma (in Module) sSum_eq:"\<lbrakk>R module N; free_generator R M H; H1 \<subseteq> H;
h \<in> H - H1\<rbrakk> \<Longrightarrow> (fgs R M H1) \<minusplus> (fgs R M {h}) = fgs R M (H1 \<union> {h})"
apply (cut_tac sc_Ring, frule Ring.whole_ideal)
apply (frule free_generator_sub)
apply (frule subset_trans[of H1 H "carrier M"], assumption+,
frule fgs_sub_carrier[of H1])
apply (cut_tac subset_trans[of "{h}" "H - H1" "carrier M"],
frule fgs_sub_carrier[of "{h}"])
apply (simp add:set_sum[of "fgs R M H1" "fgs R M {h}"])
apply (rule equalityI)
apply (rule subsetI, simp, (erule bexE)+)
apply (case_tac "H1 = {}", simp add:empty_fgs)
apply (frule_tac c = k in subsetD[of "fgs R M {h}" "carrier M"], assumption+)
apply (simp add:ag_l_zero)
apply (cut_tac fgs_mono[of H H1 "insert h H1"],
frule_tac c = ha in subsetD[of "fgs R M H1" "fgs R M (insert h H1)"],
assumption+,
cut_tac fgs_mono[of H "{h}" "insert h H1"],
frule_tac c = k in subsetD[of "fgs R M {h}" "fgs R M (insert h H1)"],
assumption+)
apply (simp add:fgs_def,
cut_tac sc_Ring, frule Ring.whole_ideal,
rule linear_span_pOp_closed, assumption+,
rule subsetI, simp, erule disjE, simp, simp add:subsetD, assumption+,
rule subsetI, simp, rule subsetI, simp, erule disjE, simp,
simp add:subsetD, assumption+, rule subsetI, simp,
rule subsetI, simp, erule disjE, simp, simp add:subsetD)
apply (rule subsetI)
apply (cut_tac x = x in mem_fgs_l_comb[of "insert h H1"],
simp,
rule subsetI, simp, erule disjE, simp, simp add:subsetD, assumption+)
apply (erule exE, (erule bexE)+)
apply (case_tac "f ` {j. j \<le> n} \<subseteq> H1")
apply (frule_tac s = s and n = n and f = f in
l_comb_mem_linear_span[of "carrier R" H1], assumption+)
apply (rule Pi_I)
apply (frule_tac f = f and A = "{j. j \<le> n}" and B = "insert h H1" and
a = xa in mem_in_image, assumption+, simp add:subsetD)
apply (fold fgs_def)
apply (frule fgs_sub_carrier[of H1])
apply (frule_tac c = "l_comb R M n s f" in subsetD[of "fgs R M H1"
"carrier M"], assumption+)
apply (frule_tac t = "l_comb R M n s f" in ag_r_zero[THEN sym], simp)
apply (cut_tac fgs_submodule[of "{h}"],
frule submodule_inc_0[of "fgs R M {h}"], blast,
rule subsetI, simp)
apply (frule_tac f = f and A = "{j. j \<le> n}" and B = "insert h H1" in
image_sub0)
apply (frule_tac Y = "f ` {j. j \<le> n}" and a = h and X = H1 in sub_inserted1,
assumption+, erule conjE,
thin_tac "\<not> f ` {j. j \<le> n} \<subseteq> H1",
thin_tac "f ` {j. j \<le> n} \<subseteq> insert h H1", erule conjE)
apply (cut_tac H = "insert h H1" and f = f and n = n and s = s in
unique_expression2)
apply (rule subsetI, simp, erule disjE, simp, simp add:subsetD)
apply assumption+
apply ((erule exE)+, (erule conjE)+, simp)
apply (simp add:bij_to_def, erule conjE,
simp add:surj_to_def)
apply (rotate_tac -2, frule sym, thin_tac "g ` {j. j \<le> m} = f ` {j. j \<le> n}",
simp, thin_tac "f ` {j. j \<le> n} = g ` {j. j \<le> m}", simp add:image_def,
erule exE,
thin_tac "f \<in> {j. j \<le> n} \<rightarrow> insert h H1",
thin_tac "s \<in> {j. j \<le> n} \<rightarrow> carrier R",
thin_tac "l_comb R M n s f = l_comb R M m t g", erule conjE)
apply (case_tac "m = 0", simp,
frule_tac a = H1 in fgs_submodule,
frule submodule_inc_0[of "fgs R M H1"],
cut_tac a = "insert (g 0) H1" in fgs_submodule,
rule subsetI, simp, erule disjE, simp, simp add:subsetD,
frule_tac H = "fgs R M (insert (g 0) H1)" in submodule_subset,
frule_tac c = "l_comb R M 0 t g" and A = "fgs R M (insert (g 0) H1)"
and B = "carrier M" in subsetD, assumption+,
frule_tac t = "l_comb R M 0 t g" in ag_l_zero[THEN sym],
subgoal_tac "l_comb R M 0 t g \<in> fgs R M {g 0}", blast)
apply (subst fgs_def,
cut_tac s = t and n = m and f = g in
l_comb_mem_linear_span[of "carrier R" "{h}"], assumption+,
rule subsetI, simp, simp, simp,
simp, simp)
(** case n \<noteq> 0 **)
apply (cut_tac f = g and n = "m - Suc 0" and s = t and l = xa in
unique_expression3[of "insert h H1"],
rule subsetI, simp, erule disjE, simp, simp add:subsetD)
apply simp+
apply (rule contrapos_pp, simp+, simp add:image_def)
apply (erule bexE, simp add:inj_on_def)
apply (drule_tac a = xa in forall_spec, simp)
apply (drule_tac a = xb in forall_spec,
erule conjE, assumption, simp)
apply ((erule exE)+, (erule conjE)+, simp)
apply (thin_tac "g \<in> {j. j \<le> m} \<rightarrow> insert (g xa) H1",
thin_tac "t \<in> {j. j \<le> m} \<rightarrow> carrier R",
thin_tac "l_comb R M m t g = l_comb R M ma ta ga",
thin_tac "inj_on g {j. j \<le> m}",
thin_tac "g xa \<in> carrier M",
thin_tac "fgs R M {g xa} \<subseteq> carrier M")
apply (rotate_tac 13, frule sym, thin_tac "h = g xa", simp)
apply (thin_tac "g xa = h", thin_tac "0 < m", thin_tac "xa \<le> m",
thin_tac "ta ma = t xa", simp)
apply (rename_tac x g m t)
apply (case_tac "m = 0", simp,
frule_tac a = H1 in fgs_submodule,
frule submodule_inc_0[of "fgs R M H1"],
cut_tac a = "insert h H1" in fgs_submodule,
rule subsetI, simp, erule disjE, simp, simp add:subsetD,
frule subset_trans[of H1 H "carrier M"], assumption+,
simp add:subsetD,
frule_tac H = "fgs R M (insert h H1)" in submodule_subset,
frule_tac c = "l_comb R M 0 t g" and A = "fgs R M (insert h H1)"
and B = "carrier M" in subsetD, assumption+,
frule_tac t = "l_comb R M 0 t g" in ag_l_zero[THEN sym],
subgoal_tac "l_comb R M 0 t g \<in> fgs R M {h}", blast)
apply (subst fgs_def,
cut_tac s = t and n = 0 and f = g in
l_comb_mem_linear_span[of "carrier R" "{h}"],
assumption, rule subsetI, simp, simp add:subsetD,
simp, simp, assumption)
apply (subgoal_tac "l_comb R M m t g = l_comb R M (Suc (m - Suc 0)) t g",
simp del:Suc_pred,
frule insert_sub[of H1 H h], assumption+,
frule subset_trans[of "insert h H1" H "carrier M"], assumption+,
subst l_comb_Suc[of "insert h H1" "carrier R"], assumption+,
(subst Suc_pred, assumption,
thin_tac "l_comb R M m t g = l_comb R M (Suc (m - Suc 0)) t g",
simp)+)
apply (subgoal_tac "l_comb R M (m - Suc 0) t g \<in> fgs R M H1",
subgoal_tac "t m \<cdot>\<^sub>s h \<in> fgs R M {h}", blast,
subst fgs_def,
rule_tac h = h and a = "t m" in
sc_linear_span[of "carrier R" "{h}"], assumption+,
rule subsetI, simp, simp add:Pi_def, simp,
cut_tac a = "{h}" in fgs_submodule, rule subsetI, simp,
cut_tac s = t and n = "m - Suc 0" and f = g in
l_comb_mem_linear_span[of "carrier R" H1], assumption,
simp add:subset_trans[of H1 H "carrier M"],
rule func_pre, simp,
rule Pi_I, simp,
frule_tac f = g and A = "{k. k \<le> m}" and B = "insert h H1"
and x = xa in funcset_mem, simp, simp, simp add:inj_on_def)
apply (erule disjE)
apply (drule_tac a = m in forall_spec, simp,
drule_tac a = xa in forall_spec, simp, simp, simp)
apply (cut_tac n1 = xa and m1 = "xa - Suc 0" in Suc_le_mono[THEN sym],
simp, simp add:fgs_def, simp)
apply (rule subsetI, simp)
apply (rule subsetI, simp add:subsetD)
done
definition
monofun :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme,
('b, 'r, 'm2) Module_scheme, 'a \<Rightarrow> 'b, 'a set, 'a] \<Rightarrow> ('a \<Rightarrow> 'b)" where
"monofun R M N f H h = (\<lambda>x\<in>fgs R M {h}.
(THE y. (\<exists>r\<in>carrier R. x = r \<cdot>\<^sub>s\<^bsub>M\<^esub> h \<and> y = r \<cdot>\<^sub>s\<^bsub>N\<^esub> (f h))))"
lemma (in Module) fgs_single_span:"\<lbrakk>h \<in> carrier M; x \<in> (fgs R M {h})\<rbrakk> \<Longrightarrow>
\<exists>a\<in>carrier R. x = a \<cdot>\<^sub>s h"
apply (simp add:fgs_def linear_span_def)
apply (erule exE, (erule bexE)+)
apply (cut_tac sc_Ring, frule Ring.whole_ideal)
apply (frule_tac A = "carrier R" and z = h and h = f and n = n and t = s in
single_span, assumption+)
apply (erule bexE, simp add:l_comb_def del: Pi_split_insert_domain)
apply (frule_tac f = sa and A = "{0}" and B = "carrier R" and
x = 0 in funcset_mem, simp, blast)
done
lemma (in Module) monofun_mHomTr:"\<lbrakk>h \<in> H; free_generator R M H;
a \<in> carrier R; r \<in> carrier R; a \<cdot>\<^sub>s h = r \<cdot>\<^sub>s h\<rbrakk> \<Longrightarrow> a = r"
apply (cut_tac sc_Ring, frule free_generator_sub,
frule subsetD[of H "carrier M" h], assumption+)
apply (frule_tac a = a and m = h in sc_mem, assumption+,
frule_tac a = r and m = h in sc_mem, assumption+)
apply (frule ag_eq_diffzero[of "a \<cdot>\<^sub>s h" "r \<cdot>\<^sub>s h"], assumption+,
simp only:sc_minus_am1, thin_tac "(a \<cdot>\<^sub>s h = r \<cdot>\<^sub>s h) = (\<zero> = \<zero>)")
apply (frule Ring.ring_is_ag[of R],
frule aGroup.ag_mOp_closed[of R r], assumption+)
apply (simp add:sc_l_distr[THEN sym])
apply (frule free_gen_coeff_zero[of H h "a \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> r"], assumption+)
apply (rule aGroup.ag_pOp_closed, assumption+)
apply (simp add:aGroup.ag_eq_diffzero)
done
lemma (in Module) monofun_mhomTr1:"\<lbrakk>R module N; h \<in> H; free_generator R M H;
f \<in> H \<rightarrow> carrier N; a \<in> carrier R\<rbrakk> \<Longrightarrow>
monofun R M N f H h (a \<cdot>\<^sub>s h) = a \<cdot>\<^sub>s\<^bsub>N\<^esub> (f h)"
apply (frule free_generator_sub)
apply (simp add:monofun_def)
apply (cut_tac a = "{h}" and x = h in elem_fgs, rule subsetI,
simp add:subsetD, simp,
cut_tac fgs_submodule[of "{h}"])
apply (frule submodule_sc_closed[of "fgs R M {h}" a h], assumption+,
simp)
apply (subgoal_tac "\<exists>!y. \<exists>r\<in>carrier R. a \<cdot>\<^sub>s h = r \<cdot>\<^sub>s h \<and> y = r \<cdot>\<^sub>s\<^bsub>N\<^esub> f h",
subgoal_tac "\<exists>r\<in>carrier R. a \<cdot>\<^sub>s h = r \<cdot>\<^sub>s h \<and>
(THE y. \<exists>s\<in>carrier R. a \<cdot>\<^sub>s h = s \<cdot>\<^sub>s h \<and> y = s \<cdot>\<^sub>s\<^bsub>N\<^esub> (f h)) =
r \<cdot>\<^sub>s\<^bsub>N\<^esub> (f h)")
prefer 2 apply (rule theI', simp)
apply (thin_tac "\<exists>!y. \<exists>r\<in>carrier R. a \<cdot>\<^sub>s h = r \<cdot>\<^sub>s h \<and> y = r \<cdot>\<^sub>s\<^bsub>N\<^esub> f h")
apply (erule bexE, erule conjE, simp,
thin_tac "(THE y. \<exists>s\<in>carrier R. r \<cdot>\<^sub>s h = s \<cdot>\<^sub>s h \<and> y = s \<cdot>\<^sub>s\<^bsub>N\<^esub> f h) = r \<cdot>\<^sub>s\<^bsub>N\<^esub> f h")
apply (frule_tac r = r in monofun_mHomTr[of h H a], assumption+, simp)
apply (rule ex_ex1I)
apply blast
apply ((erule bexE)+, (erule conjE)+)
apply (frule_tac r = r in monofun_mHomTr[of h H a], assumption+,
frule_tac r = ra in monofun_mHomTr[of h H a], assumption+)
apply (thin_tac "a \<cdot>\<^sub>s h = r \<cdot>\<^sub>s h", thin_tac "a \<cdot>\<^sub>s h = ra \<cdot>\<^sub>s h",
rotate_tac -2, frule sym, thin_tac "a = r", frule sym,
thin_tac "a = ra", simp)
apply (rule subsetI, simp add:subsetD)
done
lemma (in Module) monofun_mem:"\<lbrakk>R module N; h \<in> H; free_generator R M H;
x \<in> fgs R M {h}; f \<in> H \<rightarrow> carrier N\<rbrakk> \<Longrightarrow>
monofun R M N f H h x \<in> carrier N"
apply (frule free_generator_sub,
frule subsetD[of H "carrier M" h], assumption+,
frule fgs_single_span[of h x], assumption+)
apply (erule bexE, simp add:monofun_mhomTr1,
frule funcset_mem[of f H "carrier N" h], assumption)
apply (simp add:Module.sc_mem[of N R _ "f h"])
done
lemma (in Module) monofun_eq_f:"\<lbrakk>R module N; h \<in> H; free_generator R M H;
f \<in> H \<rightarrow> carrier N\<rbrakk> \<Longrightarrow> monofun R M N f H h h = f h"
apply (cut_tac sc_Ring)
apply (frule_tac N = N and h = h and f = f and a = "1\<^sub>r\<^bsub>R\<^esub>" in monofun_mhomTr1,
assumption+)
apply (simp add:Ring.ring_one)
apply (frule free_generator_sub, frule subsetD[of H "carrier M" h],
assumption+,
frule funcset_mem[of f H "carrier N" h], assumption)
apply (simp add:sprod_one, simp add:Module.sprod_one)
done
lemma (in Module) sSum_unique:"\<lbrakk>R module N; free_generator R M H; H1 \<subseteq> H;
h \<in> H - H1; x1 \<in>(fgs R M H1); x2 \<in> (fgs R M H1);
y1 \<in> (fgs R M {h}); y2 \<in> (fgs R M {h}); x1 \<plusminus> y1 = x2 \<plusminus> y2\<rbrakk> \<Longrightarrow>
x1 = x2 \<and> y1 = y2"
apply (cut_tac sc_Ring, frule Ring.whole_ideal [of "R"],
frule free_generator_sub)
apply (frule_tac subset_trans[of H1 H "carrier M"], assumption+)
apply (frule subsetD[of H "carrier M" h], simp)
apply (frule_tac fgs_single_span[of h y1], assumption+,
frule_tac fgs_single_span[of h y2], assumption+)
apply (erule bexE, erule bexE)
apply (frule_tac a = a and m = h in sc_mem, assumption+,
frule_tac a = aa and m = h in sc_mem, assumption+)
apply (frule subset_trans[of H1 H "carrier M"], assumption+)
apply (frule_tac A = "carrier R" and H = H1 in lin_span_sub_carrier,
assumption+)
apply (fold fgs_def)
apply (frule subsetD[of "fgs R M H1" "carrier M" x1], assumption+,
frule subsetD[of "fgs R M H1" "carrier M" x2], assumption+)
apply simp
apply (frule ag_mOp_closed[of x2])
apply (frule_tac z = "a \<cdot>\<^sub>s h" in ag_pOp_assoc[of "-\<^sub>a x2" "x1"], assumption+,
simp only:ag_pOp_assoc[THEN sym], simp add:ag_l_inv1,
simp add:ag_l_zero,
frule ag_pOp_closed[of "-\<^sub>a x2" x1], assumption+,
frule_tac x = "a \<cdot>\<^sub>s h" in ag_mOp_closed)
apply (frule_tac y = "a \<cdot>\<^sub>s h" and z = "-\<^sub>a (a \<cdot>\<^sub>s h)" in
ag_pOp_assoc[of "-\<^sub>a x2 \<plusminus> x1"], assumption+,
simp add:ag_r_inv1 ag_r_zero, thin_tac "-\<^sub>a x2 \<plusminus> x1 \<plusminus> a \<cdot>\<^sub>s h = aa \<cdot>\<^sub>s h")
apply (frule fgs_submodule[of H1])
apply (frule submodule_mOp_closed[of "fgs R M H1" x2], assumption,
frule submodule_pOp_closed[of "fgs R M H1" "-\<^sub>a x2" "x1"], assumption+)
apply (case_tac "H1 = {}", simp add:empty_fgs)
apply (simp add:ag_eq_diffzero[THEN sym])
apply (frule mem_fgs_l_comb[of H1 "-\<^sub>a x2 \<plusminus> x1"], assumption+,
erule exE, (erule bexE)+)
apply (case_tac "-\<^sub>a\<^bsub>M\<^esub> x2 \<plusminus>\<^bsub>M\<^esub> x1 = \<zero>\<^bsub>M\<^esub>", simp, rotate_tac -2, frule sym,
thin_tac "\<zero> = l_comb R M n s f", simp,
simp add:ag_pOp_commute[of "-\<^sub>a x2" " x1"])
apply (simp add:ag_eq_diffzero[of x1 x2])
apply (simp add:ag_eq_diffzero[THEN sym])
apply simp
apply (frule_tac f = f and n = n and s = s in unique_expression2[of H1],
assumption+, (erule exE)+, (erule conjE)+, simp,
simp add:bij_to_def, erule conjE,
thin_tac "f \<in> {j. j \<le> n} \<rightarrow> H1",
thin_tac "s \<in> {j. j \<le> n} \<rightarrow> carrier R",
thin_tac "surj_to g {j. j \<le> m} (f ` {j. j \<le> n})",
thin_tac "l_comb R M n s f = l_comb R M m t g", simp)
apply (frule_tac f = g and A = "{j. j \<le> m}" and B = H1 and ?B1.0 = H in
extend_fun, assumption+)
apply (frule_tac f = g and n = m and s = t in unique_expression4[of H],
simp, (erule exE)+, (erule conjE)+, erule exE, (erule conjE)+)
apply (frule_tac f = g and A = "{j. j \<le> m}" and B = H1 in image_sub0,
frule_tac A = "ga ` {k. k \<le> ma}" and B = "g ` {j. j \<le> m}" and
C = H1 in subset_trans, assumption+, simp,
thin_tac "g \<in> {j. j \<le> m} \<rightarrow> H1",
thin_tac "t \<in> {j. j \<le> m} \<rightarrow> carrier R",
thin_tac "ga ` {k. k \<le> ma} \<subseteq> g ` {k. k \<le> m}",
thin_tac "l_comb R M m t g = l_comb R M ma ta ga",
thin_tac "g ` {j. j \<le> m} \<subseteq> H1",
thin_tac "inj_on g {j. j \<le> m}",
thin_tac "g \<in> {j. j \<le> m} \<rightarrow> H", simp)
apply (rename_tac a aa m g t)
apply (simp add:sc_minus_am1,
cut_tac sc_Ring, frule Ring.ring_is_ag,
frule_tac x = a in aGroup.ag_mOp_closed[of R], assumption)
apply (simp add:sc_l_distr[THEN sym])
apply (case_tac "aa \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> a = \<zero>\<^bsub>R\<^esub>", simp add:sc_0_m,
frule_tac x = aa and y = "-\<^sub>a\<^bsub>R\<^esub> a" in aGroup.ag_pOp_closed, assumption+)
apply (subgoal_tac "(\<lambda>x\<in>{0::nat}. (aa \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> a)) \<in>
{j. j \<le> (0::nat)} \<rightarrow> carrier R",
subgoal_tac "(\<lambda>x\<in>{0::nat}. h) \<in>
{j. j \<le> (0::nat)} \<rightarrow> H")
apply (frule_tac f = "\<lambda>x\<in>{0::nat}. h" and n = 0 and
s = "\<lambda>x\<in>{0::nat}. (aa \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> a)" and g = g and m = m and t = t in
unique_expression5[of H])
apply (simp)
apply (simp add:inj_on_def)
apply (simp, assumption+)
apply (simp add:l_comb_def[of _ _ 0], rule ballI, simp)
apply simp
apply (frule_tac A = "(\<lambda>x\<in>{0}. h) ` {j. j \<le> 0}" and B = "g ` {j. j \<le> m}"
and C = H1 in subset_trans, assumption+,
thin_tac "(\<lambda>x\<in>{0}. h) \<in> {j. j \<le> 0} \<rightarrow> H",
thin_tac "(\<lambda>x\<in>{0}. aa \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> a) \<in> {j. j \<le> 0} \<rightarrow> carrier R",
thin_tac "(\<lambda>x\<in>{0}. h) ` {j. j \<le> 0} \<subseteq> g ` {j. j \<le> m}",
thin_tac "g ` {k. k \<le> m} \<subseteq> H1")
apply (simp, simp, simp)
done
lemma (in Module) ex_extensionTr:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; H1 \<subseteq> H; h \<in> H; h \<notin> H1;
g \<in> mHom R (mdl M (fgs R M H1)) N;
x \<in> fgs R M H1 \<minusplus> (fgs R M {h})\<rbrakk> \<Longrightarrow>
\<exists>k1\<in> fgs R M H1. \<exists>k2\<in>fgs R M {h}. x = k1 \<plusminus> k2 \<and>
(THE y. \<exists>h1\<in>fgs R M H1. \<exists>h2\<in>fgs R M {h}. x = h1 \<plusminus> h2 \<and> y = g h1 \<plusminus>\<^bsub>N\<^esub>
(monofun R M N f H h h2)) = g k1 \<plusminus>\<^bsub>N\<^esub> (monofun R M N f H h k2)"
apply (frule free_generator_sub)
apply (frule sSum_eq[THEN sym, of N H H1 h], assumption+, simp,
simp)
apply (subgoal_tac "\<exists>!y. \<exists>h1\<in>fgs R M H1. \<exists>h2\<in>fgs R M {h}. x = h1 \<plusminus> h2 \<and>
y = g h1 \<plusminus>\<^bsub>N\<^esub> (monofun R M N f H h h2)",
subgoal_tac "\<exists>h1\<in>fgs R M H1. \<exists>h2\<in>fgs R M {h}. x = h1 \<plusminus> h2 \<and>
(THE y. \<exists>h1\<in>fgs R M H1. \<exists>h2\<in>fgs R M {h}. x = h1 \<plusminus> h2 \<and>
y = g h1 \<plusminus>\<^bsub>N\<^esub> (monofun R M N f H h h2)) = g h1 \<plusminus>\<^bsub>N\<^esub> (monofun R M N f H h h2)")
prefer 2 apply (rule theI') apply simp
apply (thin_tac " \<exists>!y. \<exists>h1\<in>fgs R M H1.
\<exists>h2\<in>fgs R M {h}. x = h1 \<plusminus> h2 \<and> y = g h1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h h2")
apply blast
apply (rule ex_ex1I)
apply (thin_tac "fgs R M (insert h H1) = fgs R M H1 \<minusplus> (fgs R M {h})")
apply (cut_tac fgs_sub_carrier[of H1], cut_tac fgs_sub_carrier[of "{h}"])
apply (simp add:set_sum)
apply (erule bexE)+
apply (frule subsetD[of H "carrier M" h], assumption+)
apply (frule_tac x = k in monofun_mem[of N h H _ f], assumption+)
apply blast
apply (rule subsetI, simp, simp add:subsetD,
rule subset_trans[of H1 H "carrier M"], assumption+)
apply ((erule bexE)+, (erule conjE)+, simp)
apply (frule_tac ?x1.0 = h1a and ?x2.0 = h1 and ?y1.0 = h2a and ?y2.0 = h2
in sSum_unique[of N H H1 h], assumption+, simp, assumption+)
apply ((erule conjE)+, simp)
done
lemma (in Module) monofun_add:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; h \<in> H; x \<in> fgs R M {h}; y \<in> fgs R M {h}\<rbrakk> \<Longrightarrow>
monofun R M N f H h (x \<plusminus> y) =
monofun R M N f H h x \<plusminus>\<^bsub>N\<^esub> (monofun R M N f H h y)"
apply (frule free_generator_sub,
cut_tac sc_Ring, frule Ring.ring_is_ag,
frule subsetD[of H "carrier M" h], assumption+,
frule fgs_single_span[of h x], assumption+,
frule fgs_single_span[of h y], assumption+, (erule bexE)+, simp,
simp add:sc_l_distr[THEN sym])
apply (frule_tac x = a and y = aa in aGroup.ag_pOp_closed[of R], assumption+)
apply (simp add:monofun_mhomTr1,
frule funcset_mem[of f H "carrier N" h], assumption+)
apply (simp add:Module.sc_l_distr)
done
lemma (in Module) monofun_sprod:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; h \<in> H; x \<in> fgs R M {h}; a \<in> carrier R\<rbrakk> \<Longrightarrow>
monofun R M N f H h ( a \<cdot>\<^sub>s x) = a \<cdot>\<^sub>s\<^bsub>N\<^esub> (monofun R M N f H h x)"
apply (frule free_generator_sub,
cut_tac sc_Ring,
frule subsetD[of H "carrier M" h], assumption+,
frule fgs_single_span[of h x], assumption+,
erule bexE, simp add:sc_assoc[THEN sym])
apply (frule_tac x = a and y = aa in Ring.ring_tOp_closed, assumption+)
apply (simp add:monofun_mhomTr1,
frule funcset_mem[of f H "carrier N" h], assumption+)
apply (simp add:Module.sc_assoc)
done
lemma (in Module) monofun_0:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; h \<in> H\<rbrakk> \<Longrightarrow> monofun R M N f H h \<zero> = \<zero>\<^bsub>N\<^esub>"
apply (cut_tac fgs_submodule[of "{h}"],
frule submodule_inc_0[of "fgs R M {h}"])
apply (cut_tac sc_Ring, frule Ring.ring_zero,
frule monofun_sprod[of N H f h \<zero> "\<zero>\<^bsub>R\<^esub>"], assumption+,
cut_tac ag_inc_zero)
apply (simp add:sc_0_m,
frule monofun_mem[of N h H \<zero> f], assumption+,
simp add:Module.sc_0_m)
apply (frule free_generator_sub, rule subsetI, simp add:subsetD)
done
lemma (in Module) ex_extension:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N; H1 \<subseteq> H; h \<in> H - H1; (H1, g) \<in> fsps R M N f H\<rbrakk> \<Longrightarrow>
\<exists>k. ((H1 \<union> {h}), k) \<in> fsps R M N f H"
apply (frule free_generator_sub,
cut_tac sc_Ring,
frule subset_trans[of H1 H "carrier M"], assumption+,
frule subsetD[of H "carrier M" h], simp+,
frule singleton_sub[of h "carrier M"])
apply (frule fgs_submodule[of "{h}"],
frule fgs_submodule[of H1],
frule submodule_subset[of "fgs R M {h}"],
frule submodule_subset[of "fgs R M H1"],
frule insert_sub[of H1 "carrier M" h], assumption,
frule fgs_submodule[of "insert h H1"])
apply (simp add:fsps_def fsp_def)
apply (erule conjE)+
apply (subgoal_tac "(\<lambda>x \<in> (fgs R M (H1 \<union> {h})).
(THE y. (\<exists>h1\<in>(fgs R M H1). \<exists>h2\<in>(fgs R M {h}). x = h1 \<plusminus> h2 \<and>
y = g h1 \<plusminus>\<^bsub>N\<^esub> (monofun R M N f H h h2)))) \<in>
mHom R (mdl M (fgs R M (insert h H1))) N \<and>
f h = (\<lambda>x \<in> (fgs R M (H1 \<union> {h})).
(THE y. (\<exists>h1\<in>(fgs R M H1). \<exists>h2\<in>(fgs R M {h}). x = h1 \<plusminus> h2 \<and>
y = g h1 \<plusminus>\<^bsub>N\<^esub> (monofun R M N f H h h2)))) h \<and>
(\<forall>z\<in>H1. g z = (\<lambda>x \<in> (fgs R M (H1 \<union> {h})).
(THE y. (\<exists>h1\<in>(fgs R M H1). \<exists>h2\<in>(fgs R M {h}). x = h1 \<plusminus> h2 \<and>
y = g h1 \<plusminus>\<^bsub>N\<^esub> (monofun R M N f H h h2)))) z)")
apply blast
apply (rule conjI)
apply (rule Module.mHom_test)
apply (rule mdl_is_module, assumption+)
apply (rule conjI,
rule Pi_I, simp add:mdl_carrier)
apply (frule sSum_eq[THEN sym, of N H H1 h], assumption+, simp, simp)
apply (frule_tac h = h and H = H and ?H1.0 = H1 and g = g and x = x and N = N
in ex_extensionTr, assumption+,
(erule bexE)+, erule conjE, rotate_tac -1,
frule sym,
thin_tac "(THE y. \<exists>h1\<in>fgs R M H1. \<exists>h2\<in>fgs R M {h}.
x = h1 \<plusminus> h2 \<and> y = g h1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h h2) =
g k1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h k2")
apply (frule mdl_is_module[of "fgs R M H1"],
frule_tac f = g and m = k1 in Module.mHom_mem[of "mdl M (fgs R M H1)"
R N], assumption+, simp add:mdl_carrier)
apply (frule_tac x = k2 in monofun_mem[of N h H _ f], assumption+)
apply (frule Module.module_is_ag[of N],
frule_tac x = "g k1" and y = "monofun R M N f H h k2 " in
aGroup.ag_pOp_closed[of N], assumption+, simp)
apply (simp add:mdl_carrier)
apply (rule conjI, (rule ballI)+)
apply (frule mdl_is_module[of "fgs R M (insert h H1)"],
frule Module.module_is_ag[of "mdl M (fgs R M (insert h H1))"],
frule_tac x = m and y = n in aGroup.ag_pOp_closed[of
"mdl M (fgs R M (insert h H1))"],
(simp add:mdl_carrier)+)
apply (frule sSum_eq[THEN sym, of N H H1 h], assumption+, simp, simp)
apply (frule_tac h = h and H = H and ?H1.0 = H1 and g = g and N = N and
x = "m \<plusminus>\<^bsub>mdl M (fgs R M H1 \<minusplus> fgs R M {h})\<^esub> n" in ex_extensionTr, assumption+,
(erule bexE)+, (erule conjE)+, simp,
thin_tac "(THE y. \<exists>h1\<in>fgs R M H1. \<exists>h2\<in>fgs R M {h}.
k1 \<plusminus> k2 = h1 \<plusminus> h2 \<and> y = g h1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h h2) =
g k1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h k2")
apply (frule_tac h = h and H = H and ?H1.0 = H1 and g = g and N = N and
x = m in ex_extensionTr, assumption+, (erule bexE)+, (erule conjE)+,
simp,
thin_tac "(THE y. \<exists>h1\<in>fgs R M H1. \<exists>h2\<in>fgs R M {h}.
k1a \<plusminus> k2a = h1 \<plusminus> h2 \<and> y = g h1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h h2) =
g k1a \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h k2a")
apply (frule_tac h = h and H = H and ?H1.0 = H1 and g = g and N = N and
x = n in ex_extensionTr, assumption+, (erule bexE)+, (erule conjE)+,
simp,
thin_tac "(THE y. \<exists>h1\<in>fgs R M H1. \<exists>h2\<in>fgs R M {h}.
k1b \<plusminus> k2b = h1 \<plusminus> h2 \<and> y = g h1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h h2) =
g k1b \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h k2b")
apply (simp add:mdl_def[of M "fgs R M H1 \<minusplus> fgs R M {h}"],
fold mdl_def)
apply (frule_tac c = k1a in subsetD[of "fgs R M H1" "carrier M"],
assumption+,
frule_tac c = k1b in subsetD[of "fgs R M H1" "carrier M"],
assumption+) apply (
frule_tac c = k2a in subsetD[of "fgs R M {h}" "carrier M"],
assumption+,
frule_tac c = k2b in subsetD[of "fgs R M {h}" "carrier M"],
assumption+)
apply (simp add:pOp_assocTr41[THEN sym], simp add:pOp_assocTr42,
frule_tac x = k2a and y = k1b in ag_pOp_commute, assumption+, simp,
thin_tac "k2a \<plusminus> k1b = k1b \<plusminus> k2a", simp add:pOp_assocTr42[THEN sym],
simp add:pOp_assocTr41)
apply(frule_tac h = k1a and k = k1b in submodule_pOp_closed[of "fgs R M H1"],
assumption+)
apply(frule_tac h = k2a and k = k2b in
submodule_pOp_closed[of "fgs R M {h}"], assumption+)
apply (frule_tac ?x1.0 = "k1a \<plusminus> k1b" and ?x2.0 = k1 and ?y1.0 = "k2a \<plusminus> k2b"
and ?y2.0 = k2 in sSum_unique[of N H H1 h], assumption+, simp,
assumption+, erule conjE, rotate_tac -2, frule sym,
thin_tac "k1a \<plusminus> k1b = k1", frule sym, thin_tac "k2a \<plusminus> k2b = k2",
thin_tac "k1a \<plusminus> k1b \<plusminus> (k2a \<plusminus> k2b) = k1 \<plusminus> k2", simp,
frule mdl_is_module[of "fgs R M H1"])
apply (frule_tac m = k1a and n = k1b in Module.mHom_add[of
"mdl M (fgs R M H1)" R N g], assumption+,
simp add:mdl_carrier, simp add:mdl_carrier,
simp add:mdl_def, fold mdl_def)
apply (simp add: monofun_add)
apply (frule_tac m = k1a in Module.mHom_mem[of "mdl M (fgs R M H1)" R N g],
assumption+, simp add:mdl_carrier,
frule_tac m = k1b in Module.mHom_mem[of "mdl M (fgs R M H1)" R N g],
assumption+, simp add:mdl_carrier,
frule_tac x = k2a in monofun_mem[of N h H _ f], assumption+,
frule_tac x = k2b in monofun_mem[of N h H _ f], assumption+)
apply (frule Module.module_is_ag[of N],
simp add:aGroup.pOp_assocTr41[of N, THEN sym],
simp add:aGroup.pOp_assocTr42[of N],
simp add:aGroup.ag_pOp_commute[of N])
apply (rule ballI)+
apply (frule mdl_is_module[of "fgs R M (insert h H1)"])
apply (frule_tac a = a and m = m in Module.sc_mem[of
"mdl M (fgs R M (insert h H1))" R], assumption+,
simp add:mdl_carrier, simp add:mdl_carrier)
apply (frule sSum_eq[THEN sym, of N H H1 h], assumption+, simp)
apply (frule_tac h = h and H = H and ?H1.0 = H1 and g = g and N = N and
x = "a \<cdot>\<^sub>s\<^bsub>mdl M (fgs R M (insert h H1))\<^esub> m" in ex_extensionTr, assumption+,
simp)
apply ((erule bexE)+, erule conjE, simp,
thin_tac "(THE y. \<exists>h1\<in>fgs R M H1. \<exists>h2\<in>fgs R M {h}.
k1 \<plusminus> k2 = h1 \<plusminus> h2 \<and> y = g h1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h h2) =
g k1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h k2")
apply (simp add:mdl_def, fold mdl_def)
apply (frule_tac h = h and H = H and ?H1.0 = H1 and g = g and N = N and
x = m in ex_extensionTr, assumption+,
(erule bexE)+, erule conjE, simp,
thin_tac "(THE y. \<exists>h1\<in>fgs R M H1. \<exists>h2\<in>fgs R M {h}.
k1a \<plusminus> k2a = h1 \<plusminus> h2 \<and> y = g h1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h h2) =
g k1a \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h k2a",
frule_tac c = k1a in subsetD[of "fgs R M H1" "carrier M"], assumption+,
frule_tac c = k1 in subsetD[of "fgs R M H1" "carrier M"], assumption+,
frule_tac c = k2a in subsetD[of "fgs R M {h}" "carrier M"], assumption+,
frule_tac c = k2 in subsetD[of "fgs R M {h}" "carrier M"], assumption+) apply (simp add:sc_r_distr,
frule_tac a = a and h = k1a in
submodule_sc_closed[of "fgs R M H1"], assumption+,
frule_tac a = a and h = k2a in
submodule_sc_closed[of "fgs R M {h}"], assumption+)
apply (frule_tac ?x1.0 = k1 and ?x2.0 = "a \<cdot>\<^sub>s k1a" and ?y1.0 = k2 and
?y2.0 = "a \<cdot>\<^sub>s k2a" in sSum_unique[of N H H1 h], assumption+,
simp, assumption+, rule sym, assumption,
thin_tac "a \<cdot>\<^sub>s k1a \<plusminus> a \<cdot>\<^sub>s k2a = k1 \<plusminus> k2", simp)
apply (frule subset_trans[of H1 H "carrier M"], assumption+,
frule mdl_is_module[of "fgs R M H1"])
apply (frule_tac a = a and m = k1a in Module.mHom_lin[of "mdl M (fgs R M H1)"
R N _ g], assumption+, simp add:mdl_carrier, assumption+)
apply (simp add:mdl_def, fold mdl_def)
apply (simp add:monofun_sprod)
apply (frule_tac m = k1a in Module.mHom_mem[of "mdl M (fgs R M H1)" R N g],
assumption+, simp add:mdl_carrier,
frule_tac x = k2a in monofun_mem[of N h H _ f], assumption+,
simp add:Module.sc_r_distr[of N])
apply (rule conjI) apply simp
apply (cut_tac insert_sub[of H1 "carrier M" h],
frule elem_fgs[of "insert h H1" h], simp, simp)
apply (frule_tac h = h and H = H and ?H1.0 = H1 and g = g and N = N and
x = h in ex_extensionTr, assumption+)
apply (frule sSum_eq[THEN sym, of N H H1 h], assumption+, simp, simp)
apply ((erule bexE)+, (erule conjE), simp,
thin_tac "(THE y.
\<exists>h1\<in>fgs R M H1.
\<exists>h2\<in>fgs R M {k1 \<plusminus> k2}.
k1 \<plusminus> k2 = h1 \<plusminus> h2 \<and>
y = g h1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H (k1 \<plusminus> k2) h2) =
g k1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H (k1 \<plusminus> k2) k2")
apply (rotate_tac -1, frule sym, thin_tac "h = k1 \<plusminus> k2", simp)
apply (cut_tac elem_fgs[of "{h}" h],
frule ag_l_zero[THEN sym, of h])
apply (frule submodule_inc_0[of "fgs R M H1"])
apply (frule_tac ?x1.0 = k1 and ?x2.0 = \<zero> and ?y1.0 = k2 and ?y2.0 = h in
sSum_unique[of N H H1 h], assumption+, simp, assumption+,
simp, simp, simp add: monofun_eq_f)
apply (cut_tac Module.mHom_0[of "mdl M (fgs R M H1)" R N g],
simp add:mdl_def, fold mdl_def)
apply (frule funcset_mem[of f H "carrier N" h], assumption+,
frule Module.module_is_ag[of N], simp add:aGroup.l_zero,
rule mdl_is_module, assumption+,
rule subsetI, simp, simp, assumption+)
apply (rule ballI, simp,
cut_tac x = z in elem_fgs[of "insert h H1"],
rule subsetI, simp, erule disjE, simp, simp add:subsetD,
rule_tac c = z in subsetD[of "H1" "insert h H1"],
rule subsetI, simp, assumption, simp)
apply (frule_tac h = h and H = H and ?H1.0 = H1 and g = g and N = N and
x = z in ex_extensionTr, assumption+)
apply (frule sSum_eq[THEN sym, of N H H1 h], assumption+, simp, simp)
apply ((erule bexE)+, (erule conjE), simp)
apply (thin_tac "(THE y. \<exists>h1\<in>fgs R M H1. \<exists>h2\<in>fgs R M {h}.
k1 \<plusminus> k2 = h1 \<plusminus> h2 \<and> y = g h1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h h2) =
g k1 \<plusminus>\<^bsub>N\<^esub> monofun R M N f H h k2")
apply (rotate_tac -1, frule sym, thin_tac "z = k1 \<plusminus> k2", simp)
apply (frule_tac x = z in elem_fgs[of H1], assumption,
frule_tac c = z in subsetD[of H1 "carrier M"], assumption+,
frule_tac t = z in ag_r_zero[THEN sym])
apply (frule_tac ?x1.0 = k1 and ?x2.0 = z and ?y1.0 = k2 and ?y2.0 = \<zero> in
sSum_unique[of N H H1 h], assumption+, simp, assumption+,
simp add:submodule_inc_0, simp, simp)
apply (simp add: monofun_0)
apply (frule mdl_is_module[of "fgs R M H1"],
frule_tac m = z in Module.mHom_mem[of "mdl M (fgs R M H1)" R N g],
assumption+, simp add:mdl_carrier)
apply (frule Module.module_is_ag[of N],
simp add:aGroup.ag_r_zero)
done
lemma (in Module) mHom_mHom:"\<lbrakk>R module N; g \<in> mHom R (mdl M (carrier M)) N\<rbrakk>
\<Longrightarrow> g \<in> mHom R M N"
apply (rule mHom_test, assumption)
apply (rule conjI)
apply (simp add:mHom_def aHom_def mdl_def)
apply (simp add:mHom_def aHom_def mdl_def)
done
lemma (in Module) exist_extension_mhom:"\<lbrakk>R module N; free_generator R M H;
f \<in> H \<rightarrow> carrier N\<rbrakk> \<Longrightarrow> \<exists>g\<in>mHom R M N. \<forall>x\<in>H. g x = f x"
+supply [[simproc del: defined_all]]
apply (frule Order_od_fm_fun[of N H f], assumption+)
apply (frule_tac N = N and H = H and f = f in od_fm_fun_inductive,
assumption+)
apply (frule_tac D = "od_fm_fun R M N f H" in Order.g_Zorn_lemma3,
assumption+)
apply (erule bexE)
apply (rule contrapos_pp, simp+)
apply (simp add:maximal_element_def)
apply (simp add:od_fm_fun_carrier)
apply (cut_tac p = m in PairE_lemma, (erule exE)+, simp)
apply (frule_tac a = x and b = y in mem_fsps_fst_sub[of N H f], assumption+)
apply (case_tac "x \<noteq> H", frule not_sym)
apply (frule_tac A = H and B = x in diff_nonempty, assumption,
frule_tac A = "H - x" in nonempty_ex, erule exE,
rename_tac m H1 g h)
apply (frule_tac ?H1.0 = H1 and h = h and g = g in ex_extension[of N H f],
assumption+, erule exE)
apply (drule_tac x = "(H1 \<union> {h}, k)" in bspec, assumption)
apply (simp add:od_fm_fun_def ole_def, fold od_fm_fun_def,
simp add:insert_inc1, (erule conjE)+,
cut_tac a = h and A = H1 in insert_inc2,
rotate_tac -3, frule sym, thin_tac "H1 = insert h H1", simp)
apply simp
apply (frule_tac a = H and b = y in mem_fsps_snd_mHom[of N H f], assumption+)
apply (simp add:fgs_def free_generator_def generator_def)
apply (frule_tac g = y in mHom_mHom[of N], assumption+,
(erule conjE)+, thin_tac "\<forall>n s. s \<in> {j. j \<le> n} \<rightarrow> carrier R \<and>
(\<exists>f. f \<in> {j. j \<le> n} \<rightarrow> H \<and>
inj_on f {j. j \<le> n} \<and> l_comb R M n s f = \<zero>) \<longrightarrow>
s \<in> {j. j \<le> n} \<rightarrow> {\<zero>\<^bsub>R\<^esub>}")
apply (drule_tac x = y in bspec, assumption+,
simp add:fsps_def fsp_def)
done
section "Nakayama lemma"
definition
Lcg :: "[('r, 'm) Ring_scheme, ('a, 'r, 'm1) Module_scheme, nat] \<Rightarrow> bool" where
"Lcg R M j \<longleftrightarrow> (\<exists>H. finite_generator R M H \<and> j = card H)"
(* Least cardinality of generator *)
lemma (in Module) NAKTr1:"M fgover R \<Longrightarrow>
\<exists>H. finite_generator R M H \<and> (LEAST j.
\<exists>L. finite_generator R M L \<and> j = card L) = card H"
apply (simp add:fGOver_def)
apply (erule exE)
apply (rule LeastI)
apply (rule_tac x = H in exI) apply simp
done
lemma (in Module) NAKTr2:"\<lbrakk>Lcg R M j; k < (LEAST j. Lcg R M j)\<rbrakk> \<Longrightarrow>
\<not> Lcg R M k"
apply (rule not_less_Least, assumption+)
done
lemma (in Module) NAKTr3:"\<lbrakk>M fgover R; H \<subseteq> carrier M; finite H;
card H < (LEAST j. \<exists>L. finite_generator R M L \<and> j = card L)\<rbrakk> \<Longrightarrow>
\<not> finite_generator R M H"
apply (frule NAKTr1, erule exE, erule conjE)
apply (cut_tac j = "card Ha" and k = "card H" in NAKTr2)
apply (simp add:Lcg_def, blast)
apply (simp add:Lcg_def)
apply (simp add:Lcg_def)
apply (rule contrapos_pp, simp+)
apply (drule_tac a = H in forall_spec, assumption)
apply simp
done
lemma (in Module) finite_gen_over_ideal:"\<lbrakk>ideal R A; h \<in> {j. j \<le> (n::nat)} \<rightarrow>
carrier M; generator R M (h ` {j. j \<le> n}); A \<odot>\<^bsub>R\<^esub> M = carrier M;
m \<in> carrier M \<rbrakk> \<Longrightarrow> \<exists>s\<in>{j. j \<le> n} \<rightarrow> A. m = l_comb R M n s h"
apply (frule l_span_closed3[of "A" "h ` {j. j \<le> n}"], assumption+)
apply (rotate_tac -1, frule sym,
thin_tac "linear_span R M A (h ` {j. j \<le> n}) = carrier M")
apply (frule eq_set_inc[of m "carrier M"
"linear_span R M A (h ` {j. j \<le> n})"], assumption+)
apply (thin_tac "carrier M = linear_span R M A (h ` {j. j \<le> n})")
apply (simp add:linear_span_def)
apply (cut_tac Nset_nonempty[of n], simp)
apply (case_tac "\<forall>x. \<not> x \<le> n", simp, simp) (** simp no good **)
apply (erule exE, (erule bexE)+)
apply (simp add:finite_lin_span)
done
lemma (in Module) NAKTr4:"\<lbrakk>ideal R A; h \<in> {j. j \<le> (k::nat)} \<rightarrow> carrier M;
0 < k; h ` {j. j \<le> k} \<subseteq> carrier M; s \<in> {j. j \<le> k} \<rightarrow> A;
h k = \<Sigma>\<^sub>e M (\<lambda>j. s j \<cdot>\<^sub>s (h j)) (k - Suc 0) \<plusminus> (s k \<cdot>\<^sub>s (h k))\<rbrakk> \<Longrightarrow>
(1\<^sub>r\<^bsub>R\<^esub> \<plusminus>\<^bsub>R\<^esub> (-\<^sub>a\<^bsub>R\<^esub> (s k))) \<cdot>\<^sub>s (h k) = \<Sigma>\<^sub>e M (\<lambda>j. s j \<cdot>\<^sub>s (h j)) (k - Suc 0)"
apply (cut_tac sc_Ring)
apply (frule funcset_mem[of h "{j. j \<le> k}" "carrier M" k], simp)
apply (frule_tac funcset_mem[of s "{j. j \<le> k}" A k], simp,
frule_tac h = "s k" in Ring.ideal_subset[of R A], assumption+)
apply (cut_tac n = "k - Suc 0" and f = "\<lambda>j. s j \<cdot>\<^sub>s h j" in nsum_mem)
apply (rule allI, rule impI, rule sc_mem,
frule_tac x = j in funcset_mem[of s "{j. j \<le> k}" A ],
simp)
apply (simp add:Ring.ideal_subset)
apply (cut_tac i = j and j = "k - Suc 0" and k = k in
le_trans, assumption+, simp,
rule_tac x = j in funcset_mem[of h "{j. j \<le> k}" "carrier M"],
assumption, simp)
apply (frule sc_mem[of "s k" "h k"], assumption+,
frule ag_mOp_closed[of "s k \<cdot>\<^sub>s h k"])
apply (frule_tac ag_pOp_add_r[of "h k"
"\<Sigma>\<^sub>e M (\<lambda>j. s j \<cdot>\<^sub>s h j) (k - Suc 0) \<plusminus> s k \<cdot>\<^sub>s h k" "-\<^sub>a (s k \<cdot>\<^sub>s h k)"],
rule ag_pOp_closed, assumption+,
thin_tac "h k = \<Sigma>\<^sub>e M (\<lambda>j. s j \<cdot>\<^sub>s h j) (k - Suc 0) \<plusminus> s k \<cdot>\<^sub>s h k",
simp add:ag_pOp_assoc ag_r_inv1 ag_r_zero,
simp add:sc_minus_am1)
apply (frule Ring.ring_is_ag,
frule aGroup.ag_mOp_closed[of R "s k"], assumption+)
(** I don't know how to change s k to 1\<^sub>r\<^bsub>R\<^esub> \<cdot>\<^sub>s s k **)
apply (subgoal_tac "(1\<^sub>r\<^bsub>R\<^esub> \<plusminus>\<^bsub>R\<^esub> (-\<^sub>a\<^bsub>R\<^esub> (s k))) \<cdot>\<^sub>s h k = h k \<plusminus> (-\<^sub>a\<^bsub>R\<^esub> (s k)) \<cdot>\<^sub>s h k",
simp,
thin_tac "h k \<plusminus> (-\<^sub>a\<^bsub>R\<^esub> (s k)) \<cdot>\<^sub>s h k = \<Sigma>\<^sub>e M (\<lambda>j. s j \<cdot>\<^sub>s h j) (k - Suc 0)")
apply (cut_tac Ring.ring_one[of R],
simp add:sc_l_distr sprod_one, assumption)
done
lemma (in Module) NAKTr5:"\<lbrakk>\<not> zeroring R; ideal R A; A \<subseteq> J_rad R;
A \<odot>\<^bsub>R\<^esub> M = carrier M; finite_generator R M H; card H = Suc k; 0 < k\<rbrakk> \<Longrightarrow>
\<exists>h \<in> {j. j \<le> k} \<rightarrow> carrier M. H = h ` {j. j \<le> k} \<and>
h k \<in> linear_span R M A (h ` {j. j \<le> (k - Suc 0)})"
apply (cut_tac sc_Ring)
apply (insert Nset2_finiteTr [of "k"])
apply (subgoal_tac "\<exists>h. h \<in> {j. j \<le> k} \<rightarrow> H \<and> surj_to h {j. j \<le> k} H")
prefer 2
apply (subgoal_tac "finite H")
apply blast apply (simp add:finite_generator_def)
apply (thin_tac " \<forall>A. finite A \<and> card A = Suc k \<longrightarrow>
(\<exists>f. f \<in> {j. j \<le> k} \<rightarrow> A \<and> surj_to f {j. j \<le> k} A)")
apply (erule exE)
apply (simp add:surj_to_def, erule conjE) apply (rotate_tac -1)
apply (drule sym)
apply (frule_tac f = h and A = "{j. j \<le> k}" and B = "H" and x = k in
funcset_mem)
apply simp
apply (simp add:finite_generator_def,
simp add:generator_def, erule conjE)
apply (cut_tac A = "h ` {j. j \<le> k}" and B = "carrier M" and c = "h k" in
subsetD, assumption, simp)
apply (frule_tac h = h and n = k and m = "h k" in finite_gen_over_ideal [of
"A"])
apply (fastforce simp add:Pi_def image_def)
apply (simp add:generator_def)
apply assumption+
apply (erule bexE)
apply (frule_tac f = s and A = "{j. j \<le> k}" and B = A and x = k in
funcset_mem, simp, simp add:subsetD)
apply (subgoal_tac "l_comb R M k s h = l_comb R M (Suc (k - Suc 0)) s h",
simp del:Suc_pred,
thin_tac "l_comb R M k s h = l_comb R M (Suc (k - Suc 0)) s h")
apply (cut_tac s = s and n = "k - Suc 0" and f = h and H = "h ` {j. j \<le> k}"
in l_comb_Suc[of _ A], assumption+,
frule Ring.ideal_subset1[of R A], assumption,
simp add:extend_fun, simp)
apply (rotate_tac -3, frule sym,
thin_tac "h k = l_comb R M (Suc (k - Suc 0)) s h", simp)
apply (frule_tac h = h and k = k and s = s in NAKTr4[of A])
apply (fastforce simp add:Pi_def image_def)
apply(assumption+, rule sym, simp add:l_comb_def)
apply (cut_tac H = "h ` {j. j \<le> (k - Suc 0)}" and s = s and n = "k - Suc 0"
and f = h in l_comb_mem_linear_span[of A], assumption+)
apply (rule_tac A = "h ` {j. j \<le> k - Suc 0}" and
B = "h ` {j. j \<le> k}" and C = "carrier M" in subset_trans)
apply (rule subsetI, simp add:image_def, erule exE, erule conjE,
cut_tac i = xa and j = "k - Suc 0" and k = k in le_trans,
assumption, subst Suc_le_mono[THEN sym], simp, blast, assumption)
apply (rule Pi_I, simp)
apply (frule_tac x = x in le_pre_le[of _ k], simp add:Pi_def)
apply (rule Pi_I, simp)
apply (unfold l_comb_def)
apply (rotate_tac -2, frule sym,
thin_tac "(1\<^sub>r\<^bsub>R\<^esub> \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (s k)) \<cdot>\<^sub>s h k = \<Sigma>\<^sub>e M (\<lambda>j. s j \<cdot>\<^sub>s h j) (k - Suc 0)",
thin_tac "\<Sigma>\<^sub>e M (\<lambda>j. s j \<cdot>\<^sub>s h j) (k - Suc 0) \<plusminus> s k \<cdot>\<^sub>s h k = h k",
simp,
thin_tac "\<Sigma>\<^sub>e M (\<lambda>j. s j \<cdot>\<^sub>s h j) (k - Suc 0) = (1\<^sub>r\<^bsub>R\<^esub> \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> (s k)) \<cdot>\<^sub>s h k")
apply (frule_tac x = "s k" in Ring.J_rad_unit [of R], assumption,
frule_tac f = s and A = "{j. j \<le> k}" and B = A and x = k in
funcset_mem, simp, simp add:subsetD)
apply (frule Ring.ring_one[of R],
frule_tac a = "1\<^sub>r\<^bsub>R\<^esub>" in forall_spec, assumption)
apply (frule_tac h = "s k" in Ring.ideal_subset[of R A], assumption+,
frule Ring.ring_is_ag[of R],
frule_tac x = "s k" in aGroup.ag_mOp_closed[of R], assumption,
thin_tac "\<forall>y. y \<in> carrier R \<longrightarrow> Unit R (1\<^sub>r\<^bsub>R\<^esub> \<plusminus>\<^bsub>R\<^esub> (-\<^sub>a\<^bsub>R\<^esub> (s k)) \<cdot>\<^sub>r\<^bsub>R\<^esub> y)",
simp add:Ring.ring_r_one)
apply (simp add:Unit_def, erule conjE, erule bexE)
apply (simp add:Ring.ring_tOp_commute)
apply (frule_tac H = "h ` {j. j \<le> k - Suc 0}" and r = b and
x = "(1\<^sub>r\<^bsub>R\<^esub> \<plusminus>\<^bsub>R\<^esub> (-\<^sub>a\<^bsub>R\<^esub> (s k))) \<cdot>\<^sub>s h k" in linear_span_sc_closed[of A])
apply (rule_tac A = "h ` {j. j \<le> k - Suc 0}" and B = "h ` {j. j \<le> k}"
and C = "carrier M" in subset_trans,
rule subsetI, simp add:image_def, erule exE, erule conjE)
apply (frule_tac x = xa and n = k in le_pre_le, blast, assumption+)
apply (simp add:sc_assoc[THEN sym] sprod_one)
apply(simp add:Pi_def)
apply blast
apply simp
done
lemma (in Module) NAK:"\<lbrakk>\<not> zeroring R; M fgover R; ideal R A; A \<subseteq> J_rad R;
A \<odot>\<^bsub>R\<^esub> M = carrier M \<rbrakk> \<Longrightarrow> carrier M = {\<zero>}"
apply (cut_tac sc_Ring)
apply (frule NAKTr1, erule exE, erule conjE)
apply (simp add:finite_generator_def, frule conjunct1, frule conjunct2,
fold finite_generator_def)
apply (case_tac "H = {}", simp add:generator_def linear_span_def,
frule_tac A = H in nonempty_card_pos1, assumption+)
apply (case_tac "card H = Suc 0", simp,
frule_tac A = H in nonempty_ex, erule exE, rename_tac H h,
frule_tac t = H and a1 = h in card1_tr0[THEN sym], assumption+, simp)
apply (frule_tac H = "{h}" in generator_sub_carrier,
frule sym, thin_tac "A \<odot>\<^bsub>R\<^esub> M = carrier M", simp,
thin_tac "carrier M = A \<odot>\<^bsub>R\<^esub> M", simp add:smodule_ideal_coeff_def,
simp add:generator_def)
apply (erule conjE, rotate_tac -1, frule sym,
thin_tac "linear_span R M (carrier R) {h} = carrier M", simp)
apply (rotate_tac -1, frule sym,
frule_tac a = h and A = "linear_span R M (carrier R) {h}" and
B = "carrier M" in eq_set_inc, assumption+,
frule_tac a = h and A = "carrier M" in singleton_sub)
apply (frule_tac H = "{h}" in l_spanA_l_span[of A], assumption+,
thin_tac "carrier M = linear_span R M (carrier R) {h}",
thin_tac "linear_span R M (carrier R) {h} = carrier M", simp,
thin_tac "h \<in> linear_span R M (carrier R) {h}",
thin_tac "linear_span R M A (linear_span R M (carrier R) {h}) =
linear_span R M A {h}",
thin_tac "(LEAST j. \<exists>L. finite L \<and>
L \<subseteq> linear_span R M (carrier R) {h} \<and>
linear_span R M (carrier R) L =
linear_span R M (carrier R) {h} \<and>
j = card L) = Suc 0")
apply (frule_tac h = h and x = h in mem_single_l_span1[of A],
assumption+, erule bexE,
frule_tac h = a in Ring.ideal_subset[of R A], assumption+,
frule_tac a = a and m = h in sc_mem, assumption+,
frule_tac x = "a \<cdot>\<^sub>s h" in ag_mOp_closed)
apply (frule_tac a = h and b = "a \<cdot>\<^sub>s h" and c = "-\<^sub>a (a \<cdot>\<^sub>s h)" in
ag_pOp_add_r, assumption+,
thin_tac "h = a \<cdot>\<^sub>s h", simp add:ag_r_inv1,
simp add:sc_minus_am1,
frule Ring.ring_is_ag[of R],
frule_tac x = a in aGroup.ag_mOp_closed, assumption)
apply (subgoal_tac "h \<plusminus> (-\<^sub>a\<^bsub>R\<^esub> a) \<cdot>\<^sub>s h = 1\<^sub>r\<^bsub>R\<^esub> \<cdot>\<^sub>s h \<plusminus> (-\<^sub>a\<^bsub>R\<^esub> a) \<cdot>\<^sub>s h", simp)
apply (frule Ring.ring_one[of R],
simp add:sc_l_distr[THEN sym],
frule_tac x = a in Ring.J_rad_unit [of R], assumption,
simp add:subsetD,
drule_tac a = "1\<^sub>r\<^bsub>R\<^esub>" in forall_spec, assumption,
simp add:Ring.ring_r_one,
simp add:Unit_def, erule conjE, erule bexE,
simp add:Ring.ring_tOp_commute,
frule_tac a = b and b = "1\<^sub>r\<^bsub>R\<^esub> \<plusminus>\<^bsub>R\<^esub> -\<^sub>a\<^bsub>R\<^esub> a" and m = h in sc_assoc,
rule aGroup.ag_pOp_closed, assumption+, simp add:sprod_one sc_a_0,
simp add:l_span_zero, simp add:sprod_one)
apply (rotate_tac -1, frule not_sym,
frule_tac m = "Suc 0" and n = "card H" in noteq_le_less, assumption,
thin_tac "Suc 0 \<le> card H", thin_tac "card H \<noteq> Suc 0")
apply (frule_tac a = "Suc 0" and b = "card H" and c = "Suc 0" in
diff_less_mono, simp,
thin_tac "Suc 0 < card H")
apply (frule_tac H = H and k = "card H - Suc 0" in NAKTr5[of A],
assumption+, simp, simp)
apply (erule bexE)
apply (frule_tac H = H and ?H1.0 = "h ` {j. j \<le> card H - Suc 0 - Suc 0}" in
generator_generator,
frule_tac f = h and A = "{j. j \<le> card H - Suc 0}" and B = "carrier M"
and ?A1.0 = "{j. j \<le> card H - Suc 0 - Suc 0}" in image_sub,
rule subsetI, simp+)
apply (erule conjE)
apply (cut_tac f = h and n = "card H - Suc (Suc 0)" in image_Nset_Suc)
apply (frule_tac m = "Suc 0" and n = "card H" in Suc_leI)
apply (simp add:Suc_Suc_Tr)
apply (frule_tac f = h and A = "{j. j \<le> card H - Suc 0}" and
B = "carrier M" and ?A1.0 = "{j. j \<le> card H - Suc 0 - Suc 0}" in
image_sub,
rule subsetI, simp)
apply (frule_tac H = "h ` {j. j \<le> card H - Suc 0 - Suc 0}" in
lin_span_coeff_mono[of A], simp, simp)
apply (frule_tac c = "h (card H - Suc 0)" and A =
"linear_span R M A (h ` {j. j \<le> card H - Suc (Suc 0)})" and
B = "linear_span R M (carrier R)
(h ` {j. j \<le> card H - Suc (Suc 0)})" in subsetD, assumption)
apply (rule_tac A = H and B = "insert (h (card H - Suc 0))
(h ` {j. j \<le> card H - Suc (Suc 0)})" and
C = "linear_span R M (carrier R)
(h ` {j. j \<le> card H - Suc (Suc 0)})" in subset_trans)
apply simp
apply (rule_tac A = "h ` {j. j \<le> card H - Suc (Suc 0)}" and
B = "linear_span R M (carrier R)
(h ` {j. j \<le> card H - Suc (Suc 0)})" and a = "h (card H - Suc 0)"
in insert_sub)
apply (rule subsetI)
apply (rule_tac H = "h ` {j. j \<le> card H - Suc (Suc 0)}" and h = x in
h_in_linear_span, assumption+)
apply (frule_tac H = "h ` {j. j \<le> card H - Suc 0 - Suc 0}" in NAKTr3)
apply (simp add:generator_sub_carrier)
apply (rule finite_imageI, simp)
apply (thin_tac "H = h ` {j. j \<le> card H - Suc 0} \<and>
h (card H - Suc 0)
\<in> linear_span R M A (h ` {j. j \<le> card H - Suc 0 - Suc 0})")
apply (cut_tac k = "card H - Suc 0 - Suc 0" in finite_Collect_le_nat,
frule_tac A = "{j. j \<le> card H - Suc 0 - Suc 0}" and f = h in
card_image_le,
simp)
apply (frule_tac m = "Suc 0" and n = "card H" in Suc_leI,
simp add:Suc_Suc_Tr, simp add:finite_generator_def) (*
apply (rule_tac i = "card (h ` {j. j \<le> card H - Suc (Suc 0)})" and
j = "card H - Suc 0" and k = "card H" in Nat.le_less_trans, assumption+)
apply simp*)
apply (cut_tac k = "card H - Suc 0 - Suc 0" in finite_Collect_le_nat,
frule_tac F = "{j. j \<le> card H - Suc 0 - Suc 0}" and h = h in
finite_imageI, simp add:finite_generator_def)
done
lemma (in Module) fg_qmodule:"\<lbrakk> M fgover R; submodule R M N\<rbrakk> \<Longrightarrow>
(M /\<^sub>m N) fgover R"
apply (frule mpj_mHom[of N])
apply (frule mpj_surjec [of N])
apply (rule surjec_finitely_gen [of "M /\<^sub>m N" "mpj M N"])
apply (simp add:qmodule_module)
apply assumption+
done
lemma (in Module) NAK1:"\<lbrakk>\<not> zeroring R; M fgover R; submodule R M N;
ideal R A; A \<subseteq> J_rad R; carrier M = A \<odot>\<^bsub>R\<^esub> M \<minusplus> N \<rbrakk> \<Longrightarrow> carrier M = N"
apply (subgoal_tac "A \<odot>\<^bsub>R\<^esub> (M /\<^sub>m N) = carrier (M /\<^sub>m N)")
apply (frule fg_qmodule [of N], assumption+)
apply (frule qmodule_module [of N])
apply (frule Module.NAK [of "M /\<^sub>m N" R "A"], assumption+,
thin_tac "R module M /\<^sub>m N", thin_tac "M /\<^sub>m N fgover R",
thin_tac "A \<odot>\<^bsub>R\<^esub> M /\<^sub>m N = carrier (M /\<^sub>m N)",
thin_tac "carrier M = A \<odot>\<^bsub>R\<^esub> M \<minusplus> N")
apply (simp add:qmodule_def)
apply (rule equalityI)
apply (rule subsetI)
apply (frule_tac m = x in set_mr_cos_mem[of N], assumption)
apply simp
apply (frule_tac m = x in m_in_mr_coset[of N], assumption, simp)
apply (simp add:submodule_subset)
apply (subst smodule_ideal_coeff_def)
apply (frule qmodule_module [of N])
apply (frule Module.lin_span_sub_carrier[of "M /\<^sub>m N" R A "carrier (M /\<^sub>m N)"],
assumption, simp)
apply (rule equalityI, assumption+)
apply (thin_tac "linear_span R (M /\<^sub>m N) A (carrier (M /\<^sub>m N)) \<subseteq>
carrier (M /\<^sub>m N)")
apply (rule subsetI)
apply (simp add:qmodule_carr)
apply (frule_tac x = x in mem_set_mr_cos[of N], assumption, erule bexE)
apply (frule smodule_ideal_coeff_is_Submodule[of A],
frule submodule_subset,
frule submodule_subset[of "A \<odot>\<^bsub>R\<^esub> M"])
apply (frule_tac x = m in mem_set_sum[of "A \<odot>\<^bsub>R\<^esub> M" "N"], assumption+,
simp, (erule bexE)+)
apply simp
apply (frule submodule_asubg[of N])
apply (frule sym, thin_tac "carrier M = A \<odot>\<^bsub>R\<^esub> M \<minusplus> N", simp)
apply (frule_tac c = h in subsetD[of "A \<odot>\<^bsub>R\<^esub> M" "carrier M"], assumption+,
frule_tac c = k in subsetD[of N "carrier M"], assumption+,
frule_tac x = h and y = k in ag_pOp_commute, assumption+, simp)
apply (simp add:arcos_fixed[THEN sym, of N], thin_tac "h \<plusminus> k = k \<plusminus> h",
thin_tac "m = k \<plusminus> h",
thin_tac "submodule R M (A \<odot>\<^bsub>R\<^esub> M)",
thin_tac "A \<odot>\<^bsub>R\<^esub> M \<subseteq> carrier M",
thin_tac "A \<odot>\<^bsub>R\<^esub> M \<minusplus> N = carrier M")
apply (simp add:smodule_ideal_coeff_def)
apply (frule_tac h = h in linmap_im_linspan1 [of A "M /\<^sub>m N" "mpj M N"
"carrier M"], assumption+,
rule mpj_mHom[of N], assumption, simp, assumption)
apply (frule_tac m = h in elem_mpj[of _ N], assumption,
frule_tac mpj_surjec[of N],
simp add:surjec_def surj_to_def qmodule_carr)
done
section "Direct sum and direct products of modules"
definition
prodM_sprod :: "[('r, 'm) Ring_scheme, 'i set,
'i \<Rightarrow> ('a, 'r, 'm1) Module_scheme] \<Rightarrow> 'r \<Rightarrow> ('i \<Rightarrow> 'a) \<Rightarrow> ('i \<Rightarrow> 'a)" where
"prodM_sprod R I A = (\<lambda>a\<in>carrier R. \<lambda>g\<in>carr_prodag I A.
(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>(A j)\<^esub> (g j)))"
definition
prodM :: "[('r, 'm) Ring_scheme, 'i set, 'i \<Rightarrow> ('a, 'r, 'm1) Module_scheme] \<Rightarrow>
\<lparr> carrier:: ('i \<Rightarrow> 'a) set,
pop::['i \<Rightarrow> 'a, 'i \<Rightarrow> 'a] \<Rightarrow> ('i \<Rightarrow> 'a),
mop:: ('i \<Rightarrow> 'a) \<Rightarrow> ('i \<Rightarrow> 'a), zero::('i \<Rightarrow> 'a),
sprod :: ['r, 'i \<Rightarrow> 'a] \<Rightarrow> ('i \<Rightarrow> 'a) \<rparr>" where
"prodM R I A = \<lparr>carrier = carr_prodag I A,
pop = prod_pOp I A, mop = prod_mOp I A,
zero = prod_zero I A, sprod = prodM_sprod R I A \<rparr>"
definition
mProject :: "[('r, 'm) Ring_scheme, 'i set,
'i \<Rightarrow> ('a, 'r, 'more) Module_scheme, 'i] \<Rightarrow> ('i \<Rightarrow> 'a) \<Rightarrow> 'a" where
"mProject R I A j = (\<lambda>f\<in>carr_prodag I A. f j)"
abbreviation
PRODMODULES ("(3m\<Pi>\<^bsub>_ _\<^esub> _)" [72,72,73]72) where
"m\<Pi>\<^bsub>R I\<^esub> A == prodM R I A"
lemma (in Ring) prodM_carr:"\<lbrakk>\<forall>i\<in>I. (R module (M i))\<rbrakk> \<Longrightarrow>
carrier (prodM R I M) = carr_prodag I M"
apply (simp add:prodM_def)
done
lemma (in Ring) prodM_mem_eq:"\<lbrakk>\<forall>i\<in>I. (R module (M i));
m1 \<in> carrier (prodM R I M); m2 \<in> carrier (prodM R I M);
\<forall>i\<in>I. m1 i = m2 i \<rbrakk> \<Longrightarrow> m1 = m2"
apply (simp add:prodM_carr [of I M])
apply (rule carr_prodag_mem_eq [of I M])
apply (rule ballI, drule_tac x = k in bspec, assumption,
simp add:Module.module_is_ag)
apply assumption+
done
lemma (in Ring) prodM_sprod_mem:"\<lbrakk>\<forall>i\<in>I. (R module (M i)); a \<in> carrier R;
m \<in> carr_prodag I M\<rbrakk> \<Longrightarrow> prodM_sprod R I M a m \<in> carr_prodag I M"
apply (simp add:prodM_sprod_def carr_prodag_def)
apply (erule conjE)+
apply (rule conjI)
apply (rule Pi_I)
apply (simp add:Un_carrier_def)
apply (drule_tac x = x in bspec, assumption,
drule_tac x = x in bspec, assumption+)
apply (frule_tac R = R and M = "M x" and a = a and m = "m x" in
Module.sc_mem, assumption+, blast)
apply (rule ballI,
drule_tac x = i in bspec, assumption,
drule_tac x = i in bspec, assumption+)
apply (frule_tac M = "M i" and a = a and m = "m i" in
Module.sc_mem [of _ "R"], assumption+)
done
lemma (in Ring) prodM_sprod_val:"\<lbrakk>\<forall>i\<in>I. (R module (M i)); a \<in> carrier R;
m \<in> carr_prodag I M; j \<in> I\<rbrakk> \<Longrightarrow> (prodM_sprod R I M a m) j = a \<cdot>\<^sub>s\<^bsub>(M j)\<^esub> (m j)"
apply (simp add:prodM_sprod_def)
done
lemma (in Ring) prodM_Module:"\<forall>i\<in>I. (R module (M i)) \<Longrightarrow>
R module (prodM R I M)"
apply (rule Module.intro)
apply (cut_tac prodag_aGroup[of I M])
apply (simp add:prodM_def prodag_def aGroup_def)
apply (erule conjE)+
apply (rule allI, rule impI)
apply (rotate_tac -2,
frule_tac a = a in forall_spec, assumption)
apply (thin_tac "\<forall>a. a \<in> carr_prodag I M \<longrightarrow>
prod_pOp I M (prod_zero I M) a = a",
thin_tac "\<forall>i\<in>I. R module M i",
thin_tac "\<forall>a. a \<in> carr_prodag I M \<longrightarrow> (\<forall>b. b \<in> carr_prodag I M \<longrightarrow>
(\<forall>c. c \<in> carr_prodag I M \<longrightarrow> prod_pOp I M (prod_pOp I M a b) c =
prod_pOp I M a (prod_pOp I M b c)))",
thin_tac "\<forall>a. a \<in> carr_prodag I M \<longrightarrow>
prod_pOp I M (prod_mOp I M a) a = prod_zero I M",
frule_tac a = "prod_zero I M" in forall_spec, assumption,
thin_tac "\<forall>a. a \<in> carr_prodag I M \<longrightarrow> (\<forall>b. b \<in> carr_prodag I M \<longrightarrow>
prod_pOp I M a b = prod_pOp I M b a)",
frule_tac a = a in forall_spec, assumption,
thin_tac "\<forall>b. b \<in> carr_prodag I M \<longrightarrow> prod_pOp I M (prod_zero I M) b =
prod_pOp I M b (prod_zero I M)", simp)
apply (rule ballI, drule_tac x = k in bspec, assumption,
simp add:Module.module_is_ag)
apply (rule Module_axioms.intro)
apply (cut_tac Ring, assumption)
apply (simp add:prodM_carr, simp add:prodM_def prodM_sprod_mem)
apply (simp add:prodM_carr, simp add:prodM_def)
apply (frule_tac a = a and m = m in prodM_sprod_mem[of I M], assumption+,
frule_tac a = b and m = m in prodM_sprod_mem[of I M], assumption+,
(cut_tac ring_is_ag,
frule_tac x = a and y = b in aGroup.ag_pOp_closed[of R], assumption+),
frule_tac a = "a \<plusminus> b" and m = m in prodM_sprod_mem[of I M],
assumption+)
apply (rule prodM_mem_eq[of I M], assumption, simp add:prodM_carr)
apply (simp add:prodM_carr,
rule_tac X = "prodM_sprod R I M a m" and Y = "prodM_sprod R I M b m"
in prod_pOp_mem[of I M],
(rule ballI, drule_tac x = k in bspec, assumption,
simp add:Module.module_is_ag),
assumption+)
apply (rule ballI,
simp add:prodM_sprod_def prod_pOp_def,
frule_tac x = i in bspec, assumption,
thin_tac "\<forall>i\<in>I. R module M i",
thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_prodag I M",
thin_tac "(\<lambda>j\<in>I. b \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_prodag I M",
thin_tac "(\<lambda>j\<in>I. (a \<plusminus> b) \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_prodag I M",
simp add:carr_prodag_def, (erule conjE)+,
drule_tac x = i in bspec, assumption+)
apply (cut_tac Ring, simp add:Module.sc_l_distr)
apply (simp add:prodM_carr)
apply (simp add:prodM_def)
apply (cut_tac X = m and Y = n in prod_pOp_mem[of I M],
(rule ballI, frule_tac x = k in bspec, assumption,
thin_tac "\<forall>i\<in>I. R module M i", simp add:Module.module_is_ag),
assumption+)
apply (frule_tac a = a and m = "prod_pOp I M m n" in prodM_sprod_mem[of I M],
assumption+)
apply (rule prodM_mem_eq[of I M], assumption, simp add:prodM_carr)
apply (simp add:prodM_carr,
frule_tac a = a and m = m in prodM_sprod_mem[of I M], assumption+,
frule_tac a = a and m = n in prodM_sprod_mem[of I M], assumption+,
rule_tac X = "prodM_sprod R I M a m" and Y = "prodM_sprod R I M a n"
in prod_pOp_mem[of I M],
(rule ballI, drule_tac x = k in bspec, assumption,
simp add:Module.module_is_ag),
assumption+)
apply (rule ballI)
apply (frule_tac a = a and m = m in prodM_sprod_mem[of I M], assumption+,
frule_tac a = a and m = n in prodM_sprod_mem[of I M], assumption+,
simp add:prod_pOp_def prodM_sprod_def,
drule_tac x = i in bspec, assumption,
thin_tac "(\<lambda>x\<in>I. m x \<plusminus>\<^bsub>M x\<^esub> n x) \<in> carr_prodag I M",
thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>M j\<^esub> (if j \<in> I then m j \<plusminus>\<^bsub>M j\<^esub> n j else undefined))
\<in> carr_prodag I M",
thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_prodag I M",
thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>M j\<^esub> n j) \<in> carr_prodag I M")
apply (simp add:carr_prodag_def, (erule conjE)+,
drule_tac x = i in bspec, assumption,
drule_tac x = i in bspec, assumption)
apply (simp add:Module.sc_r_distr)
apply (simp add:prodM_carr,
frule_tac a = b and m = m in prodM_sprod_mem[of I M], assumption+,
frule_tac a = a and m = "prodM_sprod R I M b m" in prodM_sprod_mem[of
I M], assumption+)
apply (frule_tac x = a and y = b in ring_tOp_closed, assumption+,
frule_tac a = "a \<cdot>\<^sub>r b" and m = m in prodM_sprod_mem[of I M],
assumption+)
apply (simp add:prodM_def,
rule prodM_mem_eq, assumption+, simp add:prodM_carr,
simp add:prodM_carr)
apply (rule ballI, simp add:prodM_sprod_def,
drule_tac x = i in bspec, assumption,
thin_tac "(\<lambda>j\<in>I. b \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_prodag I M",
thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>M j\<^esub> (if j \<in> I then b \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j else undefined))
\<in> carr_prodag I M",
thin_tac "(\<lambda>j\<in>I. (a \<cdot>\<^sub>r b) \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_prodag I M",
simp add:carr_prodag_def, (erule conjE)+,
drule_tac x = i in bspec, assumption)
apply (simp add:Module.sc_assoc)
apply (simp add:prodM_carr, cut_tac ring_one,
frule_tac a = "1\<^sub>r" and m = m in prodM_sprod_mem[of I M], assumption+,
rule prodM_mem_eq, assumption+, simp add:prodM_carr,
simp add:prodM_def, simp add:prodM_def)
apply (rule ballI, simp add:prodM_def,
thin_tac "prodM_sprod R I M 1\<^sub>r m \<in> carr_prodag I M",
simp add:prodM_sprod_def,
drule_tac x = i in bspec, assumption,
simp add:carr_prodag_def, (erule conjE)+,
drule_tac x = i in bspec, assumption,
simp add:Module.sprod_one)
done
definition
dsumM :: "[('r, 'm) Ring_scheme, 'i set, 'i \<Rightarrow> ('a, 'r, 'more) Module_scheme]
\<Rightarrow> \<lparr> carrier:: ('i \<Rightarrow> 'a) set,
pop::['i \<Rightarrow> 'a, 'i \<Rightarrow> 'a] \<Rightarrow> ('i \<Rightarrow> 'a),
mop:: ('i \<Rightarrow> 'a) \<Rightarrow> ('i \<Rightarrow> 'a),
zero::('i \<Rightarrow> 'a),
sprod :: ['r, 'i \<Rightarrow> 'a] \<Rightarrow> ('i \<Rightarrow> 'a) \<rparr>" where
"dsumM R I A = \<lparr> carrier = carr_dsumag I A,
pop = prod_pOp I A, mop = prod_mOp I A,
zero = prod_zero I A, sprod = prodM_sprod R I A\<rparr>"
abbreviation
DSUMMOD ("(3\<^bsub>_\<^esub>\<Sigma>\<^sub>d\<^bsub>_\<^esub> _)" [72,72,73]72) where
"\<^bsub>R\<^esub>\<Sigma>\<^sub>d\<^bsub>I\<^esub> A == dsumM R I A"
lemma (in Ring) dsumM_carr:"carrier (dsumM R I M) = carr_dsumag I M"
by (simp add:dsumM_def)
lemma (in Ring) dsum_sprod_mem:"\<lbrakk>\<forall>i\<in>I. R module M i; a \<in> carrier R;
b \<in> carr_dsumag I M\<rbrakk> \<Longrightarrow> prodM_sprod R I M a b \<in> carr_dsumag I M"
apply (simp add:carr_dsumag_def)
apply (simp add:finiteHom_def, erule conjE, erule exE, (erule conjE)+,
frule_tac a = a and m = b in prodM_sprod_mem[of I M], assumption+)
apply (subgoal_tac "\<forall>j\<in>I - H. prodM_sprod R I M a b j = \<zero>\<^bsub>M j\<^esub>", blast)
apply (rule ballI,
drule_tac x = j in bspec, assumption,
drule_tac x = j in bspec, simp)
apply (subst prodM_sprod_def, simp)
apply (simp add:Module.sc_a_0)
done
lemma (in Ring) carr_dsum_prod:"carr_dsumag I M \<subseteq> carr_prodag I M"
apply (rule subsetI) apply (simp add:carr_dsumag_def finiteHom_def)
done
lemma (in Ring) carr_dsum_prod1:"
\<forall>x. x \<in> carr_dsumag I M \<longrightarrow> x \<in> carr_prodag I M"
apply (rule allI) apply (rule impI)
apply (cut_tac carr_dsum_prod [of I M])
apply (simp add:subsetD)
done
lemma (in Ring) carr_dsumM_mem_eq:"\<lbrakk>\<forall>i\<in>I. R module M i; x \<in> carr_dsumag I M;
y \<in> carr_dsumag I M; \<forall>j\<in>I. x j = y j\<rbrakk> \<Longrightarrow> x = y"
apply (cut_tac carr_dsum_prod [of I M])
apply (frule subsetD [of "carr_dsumag I M" "carr_prodag I M" "x"], assumption+,
frule subsetD [of "carr_dsumag I M" "carr_prodag I M" "y"], assumption+)
apply (subgoal_tac "\<forall>i\<in>I. aGroup (M i)")
apply (rule carr_prodag_mem_eq [of "I" "M"], assumption+)
apply (rule ballI)
apply (drule_tac x = i in bspec, assumption,
simp add:Module.module_is_ag)
done
lemma (in Ring) dsumM_Module:"\<forall>i\<in>I. R module (M i) \<Longrightarrow> R module (\<^bsub>R\<^esub>\<Sigma>\<^sub>d\<^bsub>I\<^esub> M)"
apply (rule Module.intro)
apply (cut_tac prodag_aGroup[of I M])
apply (cut_tac carr_dsum_prod1[of I M])
apply (simp add:dsumM_def prodag_def aGroup_def,
cut_tac dsum_pOp_func[of I M], simp,
cut_tac dsum_iOp_func[of I M], simp,
cut_tac dsum_zero_func[of I M], simp)
apply (erule conjE)+
apply (rule allI, rule impI,
thin_tac "\<forall>a. a \<in> carr_prodag I M \<longrightarrow>
(\<forall>b. b \<in> carr_prodag I M \<longrightarrow>
(\<forall>c. c \<in> carr_prodag I M \<longrightarrow>
prod_pOp I M (prod_pOp I M a b) c =
prod_pOp I M a (prod_pOp I M b c)))",
thin_tac "\<forall>i\<in>I. R module M i")
apply (frule_tac a = a in forall_spec, assumption,
drule_tac a = "prod_zero I M" in forall_spec, assumption)
apply (drule_tac a = a in forall_spec, assumption,
rotate_tac -1,
drule_tac a = "prod_zero I M" in forall_spec, assumption, simp)
apply (rule ballI, drule_tac x = k in bspec, assumption,
simp add:Module.module_is_ag)
apply (rule ballI, drule_tac x = k in bspec, assumption,
(simp add:Module.module_is_ag)+)
apply (rule ballI, drule_tac x = k in bspec, assumption+,
simp add:Module.module_is_ag)+
apply (rule Module_axioms.intro)
apply (cut_tac Ring, simp del: Pi_split_insert_domain,
simp add:dsumM_carr dsumM_def del: Pi_split_insert_domain,
simp add:dsum_sprod_mem del: Pi_split_insert_domain)
(* sc_l_distr *)
apply (cut_tac ring_is_ag,
frule_tac x = a and y = b in aGroup.ag_pOp_closed, assumption+,
simp add:dsumM_carr del: Pi_split_insert_domain,
frule_tac a = a and b = m in dsum_sprod_mem[of I M], assumption+,
frule_tac a = b and b = m in dsum_sprod_mem[of I M], assumption+,
frule_tac a = "a \<plusminus> b" and b = m in dsum_sprod_mem[of I M], assumption+,
simp add:dsumM_def del: Pi_split_insert_domain,
cut_tac X = "prodM_sprod R I M a m" and Y = "prodM_sprod R I M b m" in
prod_pOp_mem[of I M],
(rule ballI, frule_tac x = k in bspec, assumption,
thin_tac "\<forall>i\<in>I. R module M i", simp add:Module.module_is_ag del: Pi_split_insert_domain),
(cut_tac carr_dsum_prod[of I M], simp add:subsetD del: Pi_split_insert_domain)+)
apply(cut_tac carr_dsum_prod1[of I M])
apply(rule prodM_mem_eq)
apply(assumption+)
apply((simp add:prodM_carr del: Pi_split_insert_domain)+)
apply(rule ballI)
apply(simp add:prodM_sprod_def prod_pOp_def del: Pi_split_insert_domain PiE_restrict)
apply(thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_dsumag I M",
thin_tac "(\<lambda>j\<in>I. b \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_dsumag I M",
thin_tac "(\<lambda>j\<in>I. (a \<plusminus> b) \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_dsumag I M",
thin_tac "(\<lambda>x\<in>I. (if x \<in> I then a \<cdot>\<^sub>s\<^bsub>M x\<^esub> m x else undefined) \<plusminus>\<^bsub>M x\<^esub>
(if x \<in> I then b \<cdot>\<^sub>s\<^bsub>M x\<^esub> m x else undefined))
\<in> carr_prodag I M")
apply (frule_tac a = m in forall_spec, assumption,
thin_tac "\<forall>x. x \<in> carr_dsumag I M \<longrightarrow> x \<in> carr_prodag I M",
frule_tac x = i in bspec, assumption,
thin_tac "\<forall>i\<in>I. R module M i",
thin_tac "m \<in> carr_dsumag I M",
simp add:carr_prodag_def, (erule conjE)+,
frule_tac x = i in bspec, assumption,
thin_tac "\<forall>i\<in>I. m i \<in> carrier (M i)")
apply (simp add:Module.sc_l_distr)
(** sc_r_distr **)
apply (simp add:dsumM_carr,
frule_tac a = a and b = m in dsum_sprod_mem[of I M], assumption+,
frule_tac a = a and b = n in dsum_sprod_mem[of I M], assumption+)
apply (simp add:dsumM_def, cut_tac carr_dsum_prod1[of I M],
cut_tac X = m and Y = n in prod_pOp_mem[of I M],
(rule ballI, frule_tac x = k in bspec, assumption,
thin_tac "\<forall>i\<in>I. R module M i", simp add:Module.module_is_ag), simp+,
frule_tac a = a and m = "prod_pOp I M m n" in prodM_sprod_mem[of I M],
assumption, simp)
apply (cut_tac X = "prodM_sprod R I M a m" and Y = "prodM_sprod R I M a n" in
prod_pOp_mem[of I M],
(rule ballI, frule_tac x = k in bspec, assumption,
thin_tac "\<forall>i\<in>I. R module M i", simp add:Module.module_is_ag), simp+,
rule_tac ?m1.0 = "prodM_sprod R I M a (prod_pOp I M m n)" and
?m2.0 = "prod_pOp I M (prodM_sprod R I M a m) (prodM_sprod R I M a n)"
in prodM_mem_eq[of I M], assumption, (simp add:prodM_carr)+)
apply (rule ballI,
frule_tac x = i in bspec, assumption,
thin_tac "prodM_sprod R I M a m \<in> carr_dsumag I M",
thin_tac "prodM_sprod R I M a n \<in> carr_dsumag I M")
apply (rotate_tac 1,
frule_tac a = m in forall_spec, assumption,
frule_tac a = n in forall_spec, assumption,
thin_tac "\<forall>x. x \<in> carr_dsumag I M \<longrightarrow> x \<in> carr_prodag I M",
thin_tac "m \<in> carr_dsumag I M",
thin_tac "n \<in> carr_dsumag I M")
apply (frule_tac a = a and m = m in prodM_sprod_mem[of I M], assumption+,
frule_tac a = a and m = n in prodM_sprod_mem[of I M], assumption+)
apply (simp add:prodM_sprod_def prod_pOp_def del: PiE_restrict,
thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>M j\<^esub> (if j \<in> I then m j \<plusminus>\<^bsub>M j\<^esub> n j else undefined))
\<in> carr_prodag I M",
thin_tac "(\<lambda>x\<in>I. (if x \<in> I then a \<cdot>\<^sub>s\<^bsub>M x\<^esub> m x else undefined) \<plusminus>\<^bsub>M x\<^esub>
(if x \<in> I then a \<cdot>\<^sub>s\<^bsub>M x\<^esub> n x else undefined))
\<in> carr_prodag I M",
thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_prodag I M",
thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>M j\<^esub> n j) \<in> carr_prodag I M",
thin_tac "(\<lambda>x\<in>I. m x \<plusminus>\<^bsub>M x\<^esub> n x) \<in> carr_prodag I M")
apply (simp add:carr_prodag_def, (erule conjE)+,
frule_tac x = i in bspec, assumption,
thin_tac "\<forall>i\<in>I. R module M i",
frule_tac x = i in bspec, assumption,
thin_tac "\<forall>i\<in>I. m i \<in> carrier (M i)",
frule_tac x = i in bspec, assumption,
simp add:Module.sc_r_distr)
apply (frule_tac x = a and y = b in ring_tOp_closed, assumption+,
simp add:dsumM_carr,
frule_tac a = b and b = m in dsum_sprod_mem[of I M], assumption+,
frule_tac a = a and b = "prodM_sprod R I M b m" in dsum_sprod_mem[of I M],
assumption+,
frule_tac a = "a \<cdot>\<^sub>r b" and b = m in dsum_sprod_mem[of I M], assumption+)
apply (cut_tac carr_dsum_prod1[of I M])
apply (simp add:dsumM_def,
rule_tac ?m1.0 = "prodM_sprod R I M (a \<cdot>\<^sub>r b) m" and
?m2.0 = "prodM_sprod R I M a (prodM_sprod R I M b m)"
in prodM_mem_eq[of I M], assumption, (simp add:prodM_carr)+)
apply (rule ballI, simp add:prodM_sprod_def,
frule_tac x = i in bspec, assumption,
thin_tac "\<forall>i\<in>I. R module M i",
thin_tac "(\<lambda>j\<in>I. b \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_dsumag I M",
thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>M j\<^esub> (if j \<in> I then b \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j else undefined))
\<in> carr_dsumag I M",
thin_tac "(\<lambda>j\<in>I. (a \<cdot>\<^sub>r b) \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_dsumag I M",
frule_tac a = m in forall_spec, assumption,
thin_tac "\<forall>x. x \<in> carr_dsumag I M \<longrightarrow> x \<in> carr_prodag I M",
thin_tac "m \<in> carr_dsumag I M",
simp add:carr_prodag_def, (erule conjE)+,
frule_tac x = i in bspec, assumption,
thin_tac "\<forall>i\<in>I. m i \<in> carrier (M i)",
simp add:Module.sc_assoc)
apply (simp add:dsumM_carr)
apply (simp add:dsumM_def)
apply (cut_tac ring_one,
cut_tac carr_dsum_prod1[of I M],
frule_tac a = "1\<^sub>r" and b = m in dsum_sprod_mem[of I M], assumption+,
rule_tac ?m1.0 = "prodM_sprod R I M 1\<^sub>r m" and
?m2.0 = m in prodM_mem_eq[of I M], assumption, (simp add:prodM_carr)+,
frule_tac a = m in forall_spec, assumption,
frule_tac a = "prodM_sprod R I M 1\<^sub>r m" in forall_spec, assumption,
thin_tac "\<forall>x. x \<in> carr_dsumag I M \<longrightarrow> x \<in> carr_prodag I M")
apply (rule ballI,
frule_tac x = i in bspec, assumption,
thin_tac "\<forall>i\<in>I. R module M i")
apply (simp add:prodM_sprod_def)
apply (thin_tac "(\<lambda>j\<in>I. 1\<^sub>r \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_dsumag I M",
thin_tac "(\<lambda>j\<in>I. 1\<^sub>r \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_prodag I M")
apply (simp add:carr_prodag_def, (erule conjE)+,
frule_tac x = i in bspec, assumption,
simp add:Module.sprod_one)
done
definition
ringModule :: "('r, 'b) Ring_scheme \<Rightarrow> ('r, 'r) Module"
("(M\<^bsub>_\<^esub>)" [998]999) where
"M\<^bsub>R\<^esub> = \<lparr>carrier = carrier R, pop = pop R, mop = mop R,
zero = zero R, sprod = tp R\<rparr>"
lemma (in Ring) ringModule_Module:"R module M\<^bsub>R\<^esub>"
apply (rule Module.intro)
apply (simp add:aGroup_def, simp add:ringModule_def,
simp add:pop_closed, simp add:pop_aassoc,
simp add:pop_commute, simp add:mop_closed,
simp add:l_m, simp add:ex_zero,
simp add:l_zero)
apply (rule Module_axioms.intro)
apply (cut_tac Ring, assumption,
simp add:ringModule_def, simp add:ring_tOp_closed,
simp add:ringModule_def, simp add:ring_distrib2,
simp add:ringModule_def, simp add:ring_distrib1,
simp add:ringModule_def, simp add:ring_tOp_assoc,
simp add:ringModule_def, simp add:rg_l_unit)
done
definition
dsumMhom :: "['i set, 'i \<Rightarrow> ('a, 'r, 'm) Module_scheme,
'i \<Rightarrow> ('b, 'r, 'm1) Module_scheme, 'i \<Rightarrow> ('a \<Rightarrow> 'b)] \<Rightarrow> ('i \<Rightarrow> 'a) \<Rightarrow>
('i \<Rightarrow> 'b)" where
"dsumMhom I A B S = (\<lambda>f\<in>carr_dsumag I A. (\<lambda>k\<in>I. (S k) (f k)))"
(** I \<longrightarrow> A
\ |
\ | S
B **)
lemma (in Ring) dsumMhom_mem:"\<lbrakk>\<forall>i\<in>I. R module M i; \<forall>i\<in>I. R module N i;
\<forall>i\<in>I. S i \<in> mHom R (M i) (N i); x \<in> carr_dsumag I M\<rbrakk>
\<Longrightarrow> dsumMhom I M N S x \<in> carr_dsumag I N"
apply (subst carr_dsumag_def, simp, simp add:finiteHom_def)
apply (rule conjI)
apply (simp add:dsumMhom_def)
apply (cut_tac carr_dsum_prod1[of I M],
drule_tac a = x in forall_spec, assumption) apply (
subst carr_prodag_def, simp)
apply (rule conjI,
rule Pi_I,
thin_tac "x \<in> carr_dsumag I M", simp,
simp add:Un_carrier_def,
drule_tac x = xa in bspec, assumption,
drule_tac x = xa in bspec, assumption,
drule_tac x = xa in bspec, assumption,
simp add:carr_prodag_def, (erule conjE)+,
drule_tac x = xa in bspec, assumption)
apply (frule_tac R = R and M = "M xa" and N = "N xa" and f = "S xa" and
m = "x xa" in Module.mHom_mem, assumption+, blast)
apply (rule ballI,
drule_tac x = i in bspec, assumption,
drule_tac x = i in bspec, assumption,
drule_tac x = i in bspec, assumption,
simp add:carr_prodag_def, (erule conjE)+,
drule_tac x = i in bspec, assumption)
apply (simp add:Module.mHom_mem)
apply (cut_tac carr_dsum_prod1[of I M],
drule_tac a = x in forall_spec, assumption)
apply (simp add:dsumMhom_def)
apply (simp add:carr_prodag_def, (erule conjE)+)
apply (simp add:carr_dsumag_def, simp add:finiteHom_def)
apply (erule conjE, erule exE)
apply (subgoal_tac "\<forall>j\<in>I - H. S j (x j) = \<zero>\<^bsub>N j\<^esub>", blast)
apply (rule ballI,
drule_tac x = j in bspec, simp,
drule_tac x = j in bspec, simp,
drule_tac x = j in bspec, simp,
simp add:carr_prodag_def, (erule conjE)+,
drule_tac x = j in bspec, simp)
apply (simp add:Module.mHom_0)
done
lemma (in Ring) dsumMhom_mHom:"\<lbrakk>\<forall>i\<in>I. (R module (M i));
\<forall>i\<in>I. (R module (N i)); \<forall>i\<in>I. ((S i) \<in> mHom R (M i) (N i))\<rbrakk> \<Longrightarrow>
dsumMhom I M N S \<in> mHom R (dsumM R I M) (dsumM R I N)"
apply (subst mHom_def, simp)
apply (rule conjI)
apply (simp add:aHom_def,
rule conjI,
simp add:dsumM_def dsumMhom_mem)
apply (rule conjI,
simp add:dsumMhom_def extensional_def dsumM_def)
apply (rule ballI)+
apply (simp add:dsumM_def)
apply (subgoal_tac "\<forall>i\<in>I. aGroup (M i)",
frule_tac X = a and Y = b in dsum_pOp_mem [of I M], assumption+,
frule_tac x = "prod_pOp I M a b" in dsumMhom_mem [of I M N S],
assumption+,
frule_tac x = a in dsumMhom_mem [of I M N S], assumption+,
frule_tac x = b in dsumMhom_mem [of I M N S], assumption+)
apply (subgoal_tac "\<forall>i\<in>I. aGroup (N i)")
apply (frule_tac X = "dsumMhom I M N S a" and Y = "dsumMhom I M N S b" in
dsum_pOp_mem [of I N], assumption+)
apply (rule carr_dsumM_mem_eq [of I N], assumption+, rule ballI)
apply (cut_tac carr_dsum_prod1 [of I M], cut_tac carr_dsum_prod1 [of I N])
apply (simp add:prod_pOp_def)
apply (simp add:dsumMhom_def)
apply (rule_tac R = R and M = "M j" and N = "N j" and m = "a j" and n = "b j"
in Module.mHom_add)
apply simp+
apply (simp add:carr_dsumag_def carr_prodag_def finiteHom_def,
simp add:carr_dsumag_def carr_prodag_def finiteHom_def,
rule ballI, rotate_tac 1,
drule_tac x = i in bspec, assumption,
simp add:Module.module_is_ag,
rule ballI, drule_tac x = i in bspec, assumption,
simp add:Module.module_is_ag)
apply (rule ballI)+
apply (simp add:dsumM_def)
apply (frule_tac a = a and b = m in dsum_sprod_mem [of I M], assumption+)
apply (frule_tac x = "prodM_sprod R I M a m" in
dsumMhom_mem [of I M N S], assumption+)
apply (frule_tac x = m in dsumMhom_mem [of I M N S], assumption+)
apply (frule_tac a = a and b = "dsumMhom I M N S m" in
dsum_sprod_mem [of I N], assumption+)
apply (rule carr_dsumM_mem_eq [of I N], assumption+)
apply (rule ballI)
apply (cut_tac carr_dsum_prod1 [of I])
apply (drule_tac a = m in forall_spec, assumption)
apply (cut_tac carr_dsum_prod1 [of I N],
frule_tac a = "dsumMhom I M N S m" in forall_spec,
assumption)
apply (simp add:dsumMhom_def prodM_sprod_def)
apply (thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>M j\<^esub> m j) \<in> carr_dsumag I M",
thin_tac "(\<lambda>k\<in>I. S k (if k \<in> I then a \<cdot>\<^sub>s\<^bsub>M k\<^esub> m k else undefined))
\<in> carr_dsumag I N",
thin_tac "(\<lambda>k\<in>I. S k (m k)) \<in> carr_dsumag I N",
thin_tac "(\<lambda>j\<in>I. a \<cdot>\<^sub>s\<^bsub>N j\<^esub> (if j \<in> I then S j (m j) else undefined))
\<in> carr_dsumag I N",
thin_tac "\<forall>x. x \<in> carr_dsumag I N \<longrightarrow> x \<in> carr_prodag I N")
apply (drule_tac x = j in bspec, assumption,
drule_tac x = j in bspec, assumption,
drule_tac x = j in bspec, assumption,
simp add:carr_prodag_def, (erule conjE)+,
drule_tac x = j in bspec, assumption)
apply (simp add:Module.mHom_lin)
done
end
diff --git a/thys/HOL-CSP/CSP.thy b/thys/HOL-CSP/CSP.thy
--- a/thys/HOL-CSP/CSP.thy
+++ b/thys/HOL-CSP/CSP.thy
@@ -1,4079 +1,4079 @@
(*<*)
\<comment>\<open> ********************************************************************
* Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL
* Version : 2.0
*
* Author : Burkhart Wolff, Safouan Taha, Lina Ye.
* (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff)
*
* This file : A Combined CSP Theory
*
* Copyright (c) 2009 Université Paris-Sud, France
*
* All rights reserved.
*
* Redistribution and use in source and binary forms, with or without
* modification, are permitted provided that the following conditions are
* met:
*
* * Redistributions of source code must retain the above copyright
* notice, this list of conditions and the following disclaimer.
*
* * Redistributions in binary form must reproduce the above
* copyright notice, this list of conditions and the following
* disclaimer in the documentation and/or other materials provided
* with the distribution.
*
* * Neither the name of the copyright holders nor the names of its
* contributors may be used to endorse or promote products derived
* from this software without specific prior written permission.
*
* THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
* "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
* LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
* A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT
* OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL,
* SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT
* LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE,
* DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY
* THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT
* (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE
* OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
******************************************************************************\<close>
(*>*)
chapter\<open> The "Laws" of CSP \<close>
theory CSP
imports Bot Skip Stop Mprefix Det Ndet Seq Hide Sync Mndet "HOL-Eisbach.Eisbach"
begin
method exI for y::'a = (rule_tac exI[where x = y])
section\<open> General Laws\<close>
lemma skip_Neq_stop: "SKIP \<noteq> STOP"
by (auto simp: Process_eq_spec F_SKIP F_STOP D_SKIP D_STOP Un_def)
lemma bot_less1[simp]: "\<bottom> \<le> (X::'a process)"
by (simp add: le_approx_implies_le_ref)
lemma bot_less2[simp]: "Bot \<le> (X::'a process)"
by simp
section\<open> Deterministic Choice Operator Laws \<close>
lemma mono_det_FD_onside[simp]: " P \<le> P' \<Longrightarrow> (P \<box> S) \<le> (P' \<box> S)"
unfolding le_ref_def F_det D_det using F_subset_imp_T_subset by blast
lemma mono_det_FD[simp]: " \<lbrakk>P \<le> P'; S \<le> S'\<rbrakk> \<Longrightarrow> (P \<box> S) \<le> (P' \<box> S')"
by (metis Det_commute dual_order.trans mono_det_FD_onside)
lemma mono_det_ref: " \<lbrakk>P \<sqsubseteq> P'; S \<sqsubseteq> S'\<rbrakk> \<Longrightarrow> (P \<box> S) \<sqsubseteq> (P' \<box> S')"
using below_trans mono_Det mono_Det_sym by blast
lemma det_bot : "(P \<box> \<bottom>) = \<bottom>"
by (auto simp add:Process_eq_spec D_det F_det is_processT2 D_imp_front_tickFree F_UU D_UU)
lemma det_STOP: "(P \<box> STOP) = P"
by (auto simp add: Process_eq_spec D_det F_det D_STOP F_STOP T_STOP
Un_def Sigma_def is_processT8 is_processT6_S2)
lemma det_id: "(P \<box> P) = P"
by (auto simp: Process_eq_spec D_det F_det Un_def Sigma_def is_processT8 is_processT6_S2)
lemma det_assoc: "((M \<box> P) \<box> Q) = (M \<box> (P \<box> Q))"
by (auto simp add: Process_eq_spec D_det F_det Un_def Sigma_def T_det)
section\<open> NonDeterministic Choice Operator Laws \<close>
lemma mono_ndet_FD[simp]: " \<lbrakk>P \<le> P'; S \<le> S'\<rbrakk> \<Longrightarrow> (P \<sqinter> S) \<le> (P' \<sqinter> S')"
by (auto simp: le_ref_def F_ndet D_ndet)
lemma mono_ndet_FD_left[simp]: "P \<le> Q \<Longrightarrow> (P \<sqinter> S) \<le> Q"
by (metis D_ndet F_ndet dual_order.trans le_ref_def le_sup_iff order_refl)
lemma mono_ndet_FD_right[simp]: "P \<le> Q \<Longrightarrow> (S \<sqinter> P) \<le> Q"
by (metis D_ndet F_ndet dual_order.trans le_ref_def le_sup_iff order_refl)
lemma mono_ndet_ref: " \<lbrakk>P \<sqsubseteq> P'; S \<sqsubseteq> S'\<rbrakk> \<Longrightarrow> (P \<sqinter> S) \<sqsubseteq> (P' \<sqinter> S')"
using below_trans mono_Ndet mono_Ndet_sym by blast
lemma ndet_bot: "(P \<sqinter> \<bottom>) = \<bottom>"
by (auto simp: Process_eq_spec D_ndet F_ndet is_processT2 D_imp_front_tickFree F_UU D_UU)
lemma ndet_id: "(P \<sqinter> P) = P"
by (simp_all add: F_ndet D_ndet Process_eq_spec)
lemma non_det_assoc: "((M \<sqinter> P) \<sqinter> Q) = (M \<sqinter> (P \<sqinter> Q))"
by (simp_all add: F_ndet D_ndet Process_eq_spec Un_assoc)
subsection\<open> Multi-Operators laws \<close>
lemma det_distrib: "(M \<box> (P \<sqinter> Q)) = ((M \<box> P) \<sqinter> (M \<box> Q))"
by (auto simp add: Process_eq_spec F_det D_det F_ndet D_ndet Un_def T_ndet)
lemma non_det_distrib: "(M \<sqinter> (P \<box> Q)) = ((M \<sqinter> P) \<box> (M \<sqinter> Q))"
by (auto simp add: Process_eq_spec F_det D_det F_ndet
D_ndet Un_def T_ndet is_processT8 is_processT6_S2)
lemma "(P \<sqinter> Q) \<le> (P \<box> Q)"
apply(auto simp add:le_ref_def D_ndet D_det F_ndet F_det T_ndet T_det Ra_def min_elems_def)
using is_processT6_S2 NF_ND by blast+
section\<open> Sequence Operator Laws \<close>
subsection\<open>Preliminaries\<close>
definition F_minus_D_seq where
"F_minus_D_seq P Q \<equiv> {(t, X). (t, X \<union> {tick}) \<in> F P \<and> tickFree t} \<union>
{(t, X). \<exists>t1 t2. t = t1 @ t2 \<and> t1 @ [tick] \<in> T P \<and> (t2, X) \<in> F Q}"
lemma F_minus_D_seq_opt: "(a,b) \<in> F(P `;` Q) = (a \<in> D(P `;` Q) \<or> (a,b) \<in> F_minus_D_seq P Q)"
by (auto simp add:F_seq D_seq F_minus_D_seq_def)
lemma Process_eq_spec_optimized_seq :
"((P `;` Q) = (U `;` S)) = (D (P `;` Q) = D (U `;` S) \<and>
F_minus_D_seq P Q \<subseteq> F (U `;` S) \<and>
F_minus_D_seq U S \<subseteq> F (P `;` Q))"
unfolding Process_eq_spec_optimized[of "(P `;` Q)" "(U `;` S)"]
by (auto simp:F_minus_D_seq_opt)
subsection\<open>Laws\<close>
lemma mono_seq_FD[simp]: " \<lbrakk>P \<le> P'; S \<le> S'\<rbrakk> \<Longrightarrow> (P `;` S) \<le> (P' `;` S')"
apply (auto simp: le_ref_def F_seq D_seq)
by (metis F_subset_imp_T_subset subsetCE)+
lemma mono_seq_ref: " \<lbrakk>P \<sqsubseteq> P'; S \<sqsubseteq> S'\<rbrakk> \<Longrightarrow> (P `;` S) \<sqsubseteq> (P' `;` S')"
using below_trans mono_Seq mono_Seq_sym by blast
lemma Bot_SEQ: "(\<bottom> `;` P) = \<bottom>"
apply(auto simp add: Process_eq_spec D_seq F_seq front_tickFree_append F_UU D_UU T_UU)
using front_tickFree_append front_tickFree_implies_tickFree is_processT2 apply blast
using D_imp_front_tickFree front_tickFree_append front_tickFree_implies_tickFree apply blast
using front_tickFree_charn tickFree_Nil apply blast
using D_imp_front_tickFree front_tickFree_append front_tickFree_implies_tickFree apply blast
using front_tickFree_Nil tickFree_Nil by blast
lemma SEQ_SKIP: "(P `;` SKIP) = P"
apply (auto simp add: Process_eq_spec F_seq D_seq F_SKIP D_SKIP is_processT7 is_processT8_S T_F_spec is_processT6_S1)
apply (meson insertI2 is_processT4 subsetI)
apply (meson append_T_imp_tickFree is_processT5_S7 non_tickFree_tick not_Cons_self2 tickFree_append)
using T_F_spec insert_absorb is_processT5_S2 apply fastforce
apply (metis F_T is_processT nonTickFree_n_frontTickFree)
by (metis append_Nil2 front_tickFree_mono is_processT nonTickFree_n_frontTickFree not_Cons_self)
lemma SKIP_SEQ: "(SKIP `;` P) = P"
by (auto simp add: Process_eq_spec D_seq T_SKIP F_seq F_SKIP D_SKIP is_processT8_Pair)
lemma STOP_SEQ: "(STOP `;` P) = STOP"
by (auto simp: Process_eq_spec F_seq D_seq F_STOP D_STOP T_STOP Un_def)
subsection\<open> Multi-Operators laws \<close>
lemma SEQ_Ndet_distrR: "((P \<sqinter> Q) `;` S) = ((P `;` S) \<sqinter> (Q `;` S))"
by (auto simp add: Process_eq_spec D_seq D_ndet T_ndet Un_def F_seq T_seq F_ndet)
lemma SEQ_Ndet_distrL: "(P `;` (Q \<sqinter> S)) = ((P `;` Q) \<sqinter> (P `;` S))"
by (auto simp add: Process_eq_spec D_seq D_ndet T_ndet Un_def F_seq F_ndet)
lemma SEQ_assoc_D: "D(P `;` (Q `;` S)) = D((P `;` Q) `;` S)"
proof(safe, goal_cases)
case (1 x)
then show ?case
apply(auto simp add:D_seq T_seq)
using front_tickFree_Nil apply blast
apply (metis append.assoc append_single_T_imp_tickFree tickFree_append)
by (metis append.assoc)
next
case (2 x)
then show ?case
proof(auto simp add:D_seq T_seq, goal_cases)
case (1 t2 t1a t2a)
then show ?case using front_tickFree_append by fastforce
next
case (2 t1 t2 t1a t2a)
then obtain t2b where "t2a = t2b@[tick]"
by (metis T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append)
with 2 show ?case by auto
next
case (3 t1 t2 t1a t2a)
then show ?case by (metis front_tickFree_implies_tickFree process_charn)
next
case (4 t1 t2 t1a t2a)
then obtain t2b where "t2a = t2b@[tick]"
by (metis D_T T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append)
with 4 show ?case
by (metis D_imp_front_tickFree append.assoc butlast_snoc front_tickFree_implies_tickFree process_charn)
next
case (5 t1 t2 t1a t2a)
then show ?case by (metis front_tickFree_implies_tickFree process_charn)
next
case (6 t1 t2 t1a t2a)
then obtain t2b where "t2a = t2b@[tick]"
by (metis D_T T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append)
with 6 show ?case
by (metis D_imp_front_tickFree append.assoc butlast_snoc front_tickFree_implies_tickFree process_charn)
qed
qed
lemma SEQ_assoc: "(P `;` (Q `;` S)) = ((P `;` Q) `;` S)"
proof (auto simp: Process_eq_spec_optimized_seq SEQ_assoc_D, goal_cases)
case (1 a b)
then show ?case
proof(auto simp add:F_minus_D_seq_def SEQ_assoc_D F_minus_D_seq_opt append_single_T_imp_tickFree del:conjI,
goal_cases 11 12)
case (11 t1 t2)
then show ?case by (metis (mono_tags, lifting) D_seq SEQ_assoc_D UnCI mem_Collect_eq)
next
case (12 t1 t1a t2a)
hence "(t1 @ t1a) @ [tick] \<in> T (P `;` Q)" by (auto simp:T_seq)
then show ?case
using 12(2)[rule_format, of "t1@t1a" t2a] 12(4,5,6) by simp
qed
next
case (2 a b)
then show ?case
proof(auto simp add:F_minus_D_seq_def SEQ_assoc_D F_minus_D_seq_opt append_single_T_imp_tickFree T_seq del:conjI,
goal_cases 21 22 23 24 25 26)
case 21
then show ?case
by (metis (mono_tags, lifting) D_seq UnCI append_Nil2 front_tickFree_Nil mem_Collect_eq)
next
case (22 t1 t2 t1a t2a)
then obtain t2b where "t2a = t2b@[tick]"
by (metis T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append)
with 22 show ?case using append.assoc butlast_snoc by auto
next
case (23 t1 t2 t1a t2a)
hence "t1 \<in> D (P `;` (Q `;` S))"
by (simp add:D_seq) (metis append_Nil2 front_tickFree_implies_tickFree process_charn)
with 23 SEQ_assoc_D show ?case by (metis front_tickFree_implies_tickFree process_charn)
next
case (24 t1 t2 t1a t2a)
then obtain t2b where "t2a = t2b@[tick]"
by (metis D_T T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append)
with 24 show ?case by (metis D_T append.assoc butlast_snoc)
next
case (25 t1 t2 t1a t2a)
hence "t1 \<in> D (P `;` (Q `;` S))"
by (simp add:D_seq) (metis append_Nil2 front_tickFree_implies_tickFree process_charn)
with 25 SEQ_assoc_D show ?case by (metis front_tickFree_implies_tickFree process_charn)
next
case (26 t1 t2 t1a t2a)
then obtain t2b where "t2a = t2b@[tick]"
by (metis D_T T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append)
with 26 show ?case by (metis D_T append.assoc butlast_snoc)
qed
qed
section\<open> The Multi-Prefix Operator Laws \<close>
lemma mono_mprefix_FD[simp]: "\<forall>x \<in> A. P x \<le> P' x \<Longrightarrow> Mprefix A P \<le> Mprefix A P'"
by (auto simp: le_ref_def F_Mprefix D_Mprefix) blast+
lemmas mono_mprefix_ref = mono_Mprefix0
lemma Mprefix_STOP: "(Mprefix {} P) = STOP"
by (auto simp:Process_eq_spec F_Mprefix D_Mprefix D_STOP F_STOP)
subsection\<open> Multi-Operators laws \<close>
lemma mprefix_Un_distr: "(Mprefix (A \<union> B) P) = ((Mprefix A P) \<box> (Mprefix B P))"
apply (unfold Process_eq_spec, rule conjI)
apply (auto, (elim disjE conjE | simp_all add: F_det F_Mprefix Un_def image_def)+, auto)
by (auto simp add: D_det D_Mprefix Un_def)
lemma mprefix_seq: "((Mprefix A P) `;` Q) = (Mprefix A (\<lambda>x. (P x) `;` Q))"
proof (unfold Process_eq_spec, rule conjI, rule subset_antisym, goal_cases)
case 1
then show ?case
apply(unfold split_def F_seq D_seq F_Mprefix T_Mprefix D_Mprefix)
apply(rule subsetI, simp_all, elim disjE conjE exE)
apply(rule disjI1, simp, meson tickFree_tl)
apply (rule disjI2, rule conjI, simp) apply auto[1]
apply (auto simp add:hd_append)[1]
using tickFree_tl apply fastforce
by (auto simp add: hd_append)[1]
next
case 2
then show ?case
apply(unfold split_def F_seq D_seq F_Mprefix T_Mprefix D_Mprefix)
apply(rule subsetI, simp_all, elim disjE conjE exE)
apply(rule disjI1, simp, blast)
apply(rule disjI1, metis event.simps(3) list.exhaust_sel tickFree_Cons)
proof(goal_cases)
case (1 x a t1 t2)
then show ?case
apply(rule_tac disjI2, rule_tac disjI1, rule_tac x="(ev a)#t1" in exI)
using hd_Cons_tl image_iff by fastforce
next
case (2 x a t1 t2)
then show ?case
apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI1, rule_tac x="(ev a)#t1" in exI)
using hd_Cons_tl image_iff by fastforce
next
case (3 x a t1 t2)
then show ?case
apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac x="(ev a)#t1" in exI)
using hd_Cons_tl image_iff by fastforce
qed
next
case 3
then show ?case
apply (auto simp add: D_Mprefix D_seq T_Mprefix)
using tickFree_tl apply blast
apply (metis event.distinct(1) hd_append image_iff list.sel(1))
apply (metis event.distinct(1) hd_append list.sel(1) tl_append2)
apply (metis (no_types, hide_lams) append_Cons event.distinct(1) image_eqI list.distinct(1)
list.exhaust_sel list.sel(1) list.sel(3) tickFree_Cons)
by (metis append_Cons list.exhaust_sel list.sel(1) list.sel(3))
qed
subsection\<open> Derivative Operators laws \<close>
lemma mprefix_singl: "(Mprefix {a} P) = (a \<rightarrow> (P a))"
by (simp add:write0_def Mprefix_def, rule arg_cong[of _ _ "\<lambda>x. Abs_process x"]) fastforce
lemma mono_read_FD: "(\<And>x. P x \<le> Q x) \<Longrightarrow> (c `?` x \<rightarrow> (P x)) \<le> (c `?` x \<rightarrow> (Q x))"
by (simp add: read_def)
lemma mono_write_FD: "(P \<le> Q) \<Longrightarrow> (c `!` x \<rightarrow> P) \<le> (c `!` x \<rightarrow> Q)"
by (simp add: write_def)
lemma mono_write0_FD: "P \<le> Q \<Longrightarrow> (a \<rightarrow> P) \<le> (a \<rightarrow> Q)"
by (simp add: write0_def)
lemma mono_read_ref: "(\<And>x. P x \<sqsubseteq> Q x) \<Longrightarrow> (c `?` x \<rightarrow> (P x)) \<sqsubseteq> (c `?` x \<rightarrow> (Q x))"
by (simp add: mono_Mprefix0 read_def)
lemma mono_write_ref: "(P \<sqsubseteq> Q) \<Longrightarrow> (c `!` x \<rightarrow> P) \<sqsubseteq> (c `!` x \<rightarrow> Q)"
by (simp add: mono_Mprefix0 write_def)
lemma mono_write0_ref: "P \<sqsubseteq> Q \<Longrightarrow> (a \<rightarrow> P) \<sqsubseteq> (a \<rightarrow> Q)"
by (simp add: mono_Mprefix0 write0_def)
lemma write0_non_det: "(a \<rightarrow> (P \<sqinter> Q)) = ((a \<rightarrow> P) \<sqinter> (a \<rightarrow> Q))"
by (auto simp: Process_eq_spec write0_def D_ndet F_ndet F_Mprefix D_Mprefix Un_def)
lemma write0_det_non_det: "((a \<rightarrow> P) \<box> (a \<rightarrow> Q)) = ((a \<rightarrow> P) \<sqinter> (a \<rightarrow> Q))"
by (auto simp: Process_eq_spec write0_def D_ndet F_ndet F_det D_det) (simp add: F_Mprefix)+
section\<open> The Hiding Operator Laws \<close>
subsection\<open> Preliminaries \<close>
lemma elemDIselemHD: "t \<in> D P \<Longrightarrow> trace_hide t (ev ` A) \<in> D (P \\ A)"
proof (cases "tickFree t")
case True
assume "t \<in> D P"
with True show ?thesis by (simp add:D_hiding, rule_tac x=t in exI, rule_tac x="[]" in exI, simp)
next
case False
assume a:"t \<in> D P"
with False obtain t' where "t = t'@[tick]" using D_imp_front_tickFree nonTickFree_n_frontTickFree by blast
with a show ?thesis apply (auto simp add:D_hiding, rule_tac x=t' in exI, rule_tac x="[tick]" in exI)
by (meson front_tickFree_implies_tickFree front_tickFree_single is_processT)
qed
lemma length_strict_mono: "strict_mono (f::nat \<Rightarrow> 'a list) \<Longrightarrow> length (f i) \<ge> i + length (f 0)"
apply(induct i, simp)
by (metis dual_order.trans lessI less_length_mono not_less not_less_eq_eq plus_nat.simps(2) strict_mono_def)
lemma mono_trace_hide: "a \<le> b \<Longrightarrow> trace_hide a (ev ` A) \<le> trace_hide b (ev ` A)"
by (metis filter_append le_list_def)
lemma mono_constant:
assumes "mono (f::nat \<Rightarrow> 'a event list)" and "\<forall>i. f i \<le> a"
shows "\<exists>i. \<forall>j\<ge>i. f j = f i"
proof -
from assms(2) have "\<forall>i. length (f i) \<le> length a"
by (simp add: le_length_mono)
hence aa:"finite {length (f i)|i. True}"
using finite_nat_set_iff_bounded_le by auto
define lm where l2:"lm = Max {length (f i)|i. True}"
with aa obtain im where "length (f im) = lm" using Max_in by fastforce
with l2 assms(1) show ?thesis
apply(rule_tac x=im in exI, intro impI allI)
by (metis (mono_tags, lifting) Max_ge aa antisym le_length_mono le_neq_trans less_length_mono
mem_Collect_eq mono_def order_less_irrefl)
qed
lemma elemTIselemHT: "t \<in> T P \<Longrightarrow> trace_hide t (ev ` A) \<in> T (P \\ A)"
proof (cases "tickFree t")
case True
assume a:"t \<in> T P"
with True show ?thesis
proof (cases "(\<exists>ta. trace_hide t (ev ` A) = trace_hide ta (ev ` A) \<and> (ta, ev ` A) \<in> F P)")
case True
thus ?thesis by (simp add:T_hiding)
next
case False
with a False inf_hidden[of A t P] obtain f where "isInfHiddenRun f P A \<and> t \<in> range f" by auto
with True show ?thesis
by (simp add:T_hiding, rule_tac disjI2, rule_tac x=t in exI, rule_tac x="[]" in exI, auto)
qed
next
case False
assume a:"t \<in> T P"
with False obtain t' where tt:"t = t'@[tick]" using T_nonTickFree_imp_decomp by auto
with a show ?thesis
proof (cases "(\<exists>ta. trace_hide t (ev ` A) = trace_hide ta (ev ` A) \<and> (ta, ev ` A) \<in> F P)")
case True
thus ?thesis by (simp add:T_hiding)
next
case False
assume "t \<in> T P"
with False inf_hidden[of A t P] obtain f where "isInfHiddenRun f P A \<and> t \<in> range f" by auto
then show ?thesis
apply (simp add:T_hiding, rule_tac disjI2, rule_tac x=t' in exI, rule_tac x="[tick]" in exI, auto)
apply (metis append_T_imp_tickFree list.distinct(1) tt)
using tt apply force
by (metis False append_T_imp_tickFree is_processT5_S7 non_tickFree_tick not_Cons_self2 tickFree_append tt)
qed
qed
lemma hide_un_D1: "D (P \\ (A \<union> B)) \<subseteq> D ((P \\ A) \\ B)"
proof (simp add:conj_commute D_hiding, intro conjI subset_antisym subsetI, simp_all, goal_cases)
case (1 x)
then obtain t u where B1:"x = trace_hide t (ev ` (A \<union> B)) @ u" and B2:"tickFree t" and B3:"front_tickFree u" and
B4:"(t \<in> D P \<or> (\<exists>(f:: nat \<Rightarrow> 'a event list). isInfHiddenRun f P (A \<union> B) \<and> t \<in> range f))" by auto
thus ?case
apply(erule_tac disjE)
apply(rule_tac x="trace_hide t (ev ` A)" in exI, rule_tac x=u in exI)
apply(simp add:hiding_tickFree tickFree_def)
apply(rule disjI1, rule_tac x=t in exI, rule_tac x="[]" in exI, simp)
proof(goal_cases)
case 1
then obtain f n where fc:"isInfHiddenRun f P (A \<union> B) \<and> t = f n" by auto
define ff where "ff = (\<lambda>i. take (i + length (f 0)) (f i))"
with fc have ffc:"isInfHiddenRun ff P (A \<union> B) \<and> t \<in> range ff"
proof (auto, goal_cases)
case 1
{ fix x
from length_strict_mono[of f "Suc x ", OF 1(2)]
have a:"take (x + length (f 0)) (f (Suc x)) < take ((Suc x) + length (f 0)) (f (Suc x))"
by (simp add: take_Suc_conv_app_nth)
from 1(2)[THEN strict_monoD, of x "Suc x", simplified]
obtain t where "f (Suc x) = (f x @ t)" by (metis le_list_def less_imp_le)
with length_strict_mono[of f x, OF 1(2)]
have "take (x + length (f 0)) (f x) = take (x + length (f 0)) (f (Suc x))" by simp
with a have "take (x + length (f 0)) (f x) < take (Suc x + length (f 0)) (f (Suc x))" by simp
}
thus ?case by (meson lift_Suc_mono_less strict_mono_def)
next
case (2 i)
have "take (i + length (f 0)) (f i) \<le> f i"
using append_take_drop_id le_list_def by blast
also have "\<forall>x y. x \<le> y \<and> y \<in> T P \<longrightarrow> x \<in> T P" using is_processT3_ST_pref by blast
ultimately show ?case using "2"(3) by blast
next
case (3 i)
hence "(f i) \<ge> (f 0)" using strict_mono_less_eq by blast
hence "take (i + length (f 0)) (f i) \<ge> (f 0)"
by (metis add.commute append_eq_conv_conj le_list_def take_add)
hence a:"[x\<leftarrow>take (i + length (f 0)) (f i) . x \<notin> ev ` A \<and> x \<notin> ev ` B] \<ge> [x\<leftarrow>f 0 . x \<notin> ev ` A \<and> x \<notin> ev ` B]"
by (metis (no_types, lifting) filter_append le_list_def)
have "take (i + length (f 0)) (f i) \<le> f i"
using append_take_drop_id le_list_def by blast
hence "[x\<leftarrow>take (i + length (f 0)) (f i) . x \<notin> ev ` A \<and> x \<notin> ev ` B] \<le> [x\<leftarrow>f i . x \<notin> ev ` A \<and> x \<notin> ev ` B]"
by (metis (no_types, lifting) filter_append le_list_def)
with a 3(4) show ?case by (metis (no_types, lifting) dual_order.antisym)
next
case 4
have "f (length (f n) - length (f 0)) \<ge> f n"
by (simp add: "4"(2) add_le_imp_le_diff length_strict_mono strict_mono_less_eq)
hence "f n = (\<lambda>i. take (i + length (f 0)) (f i)) (length (f n) - length (f 0))"
by (metis "4"(2) add.commute append_eq_conv_conj diff_is_0_eq'
le_add_diff_inverse le_list_def le_zero_eq nat_le_linear strict_mono_less_eq)
then show ?case by blast
qed
thus ?case proof(cases "\<exists>m. (\<forall>i>m. last (ff i) \<in> (ev ` A))")
case True
then obtain m where mc:"\<forall>i>m. last (ff i) \<in> (ev ` A)" by blast
hence mc2:"tickFree (ff m)"
by (metis (no_types, lifting) B2 event.distinct(1) ffc
image_iff mem_Collect_eq set_filter tickFree_def)
with mc mc2 1 ffc show ?thesis
apply(rule_tac x="trace_hide (ff m) (ev ` A)" in exI, rule_tac x=u in exI, simp, intro conjI)
apply (metis (no_types, lifting) mem_Collect_eq set_filter tickFree_def)
apply (metis (no_types, lifting) rangeE)
apply(rule disjI1, rule_tac x="ff m" in exI, rule_tac x="[]" in exI, intro conjI, simp_all)
apply(rule disjI2, rule_tac x="\<lambda>i. ff(i + m)" in exI, intro conjI)
apply (metis (no_types, lifting) add.commute add.right_neutral rangeI)
apply (simp add: strict_mono_def)
apply blast
proof(rule allI, goal_cases)
case (1 i)
from ffc ff_def True have "\<exists>t. (ff (i + m)) = (ff m) @ t \<and> set t \<subseteq> (ev ` A)"
proof(induct i)
case 0
then show ?case by fastforce
next
case (Suc i)
then obtain tt where tc:"(ff (i + m)) = (ff m) @ tt \<and> set tt \<subseteq> (ev ` A)" by blast
from ffc ff_def length_strict_mono[of ff] have lc:"length (ff (Suc i + m)) = Suc (length (ff (i + m)))"
by (metis (no_types, lifting) add_Suc fc length_strict_mono length_take min.absorb2)
from True obtain l where lc2:"l = last (ff (Suc i + m)) \<and> l \<in> (ev ` A)"
by (meson less_add_same_cancel2 mc zero_less_Suc)
from ffc obtain r where rc:"ff (Suc i + m) = ff (i + m)@r"
by (metis add.commute add_Suc_right le_list_def lessI less_imp_le strict_mono_less_eq)
with lc have "length r = 1" by (metis Suc_eq_plus1 add_left_cancel length_append)
with rc lc2 have "r = [l]"
by (metis (no_types, lifting) Nil_is_append_conv Suc_eq_plus1 append_butlast_last_id
append_eq_append_conv append_eq_append_conv2 length_Cons length_append)
with Suc lc2 tc rc show ?case by (rule_tac x="tt@[l]" in exI, auto)
qed
then show ?case using filter_empty_conv by fastforce
qed
next
case False
{ fix i
assume as:"(i::nat) > 0"
with ffc obtain tt where ttc:"ff i = ff 0 @ tt \<and> set tt \<subseteq> ev ` (A \<union> B)"
unfolding isInfHiddenRun_1 by blast
with ffc as have "tt \<noteq> []" using strict_mono_eq by fastforce
with ttc have "last (ff i) \<in> ev ` (A \<union> B)" by auto
}
hence as2:"\<forall>i. \<exists>j>i. last(ff j) \<in> ((ev ` B) - (ev ` A))"
by (metis DiffI False UnE gr_implies_not_zero gr_zeroI image_Un)
define ffb where "ffb = rec_nat t (\<lambda>i t. (let j = SOME j. ff j > t \<and> last(ff j) \<in> ((ev ` B) - (ev ` A)) in ff j))"
with as2 have ffbff:"\<And>n. ffb n \<in> range ff"
by (metis (no_types, lifting) ffc old.nat.exhaust old.nat.simps(6) old.nat.simps(7) rangeI)
from 1 ffb_def show ?thesis
apply(rule_tac x="trace_hide t (ev ` A)" in exI, rule_tac x=u in exI, simp, intro conjI)
apply (meson filter_is_subset set_rev_mp tickFree_def)
proof(rule disjI2, rule_tac x="\<lambda>i. trace_hide (ffb i) (ev ` A)" in exI, intro conjI, goal_cases)
case 1
then show ?case by (metis (no_types, lifting) old.nat.simps(6) rangeI)
next
case 2
{ fix n
have a0:"(ffb (Suc n)) = ff (SOME j. ff j > ffb n \<and> last(ff j) \<in> ((ev ` B) - (ev ` A)))" by (simp add: ffb_def)
from ffbff obtain i where a1:"ffb n = ff i" by blast
with as2 have "\<exists>j. ff j > ffb n \<and> last(ff j) \<in> ((ev ` B) - (ev ` A))"
by (metis ffc strict_mono_def)
with a0 a1 have a:"(ffb (Suc n)) > (ffb n) \<and> last (ffb (Suc n)) \<notin> (ev ` A)"
by (metis (no_types, lifting) Diff_iff tfl_some)
then obtain r where "ffb (Suc n) = (ffb n)@r \<and> last r \<notin> (ev ` A)"
by (metis append_self_conv last_append le_list_def less_list_def)
hence "trace_hide (ffb (Suc n)) (ev ` A) > trace_hide (ffb n) (ev ` A)"
by (metis (no_types, lifting) a append_self_conv filter_append filter_empty_conv
last_in_set le_list_def less_list_def)
}
then show ?case by (metis (mono_tags, lifting) lift_Suc_mono_less_iff strict_monoI)
next
case 3
then show ?case by (metis (mono_tags) elemTIselemHT ffbff ffc rangeE)
next
case 4
from ffbff ffc show ?case by (metis rangeE trace_hide_union)
qed
qed
qed
qed
lemma hide_un_D2: "finite A \<Longrightarrow> D ((P \\ A) \\ B) \<subseteq> D (P \\ (A \<union> B))"
proof (simp add:conj_commute D_hiding, intro conjI subset_antisym subsetI, simp_all, goal_cases)
case (1 x)
then obtain t u where B1:"x = trace_hide t (ev ` B) @ u" and B2:"tickFree t" and B3:"front_tickFree u" and
B4:"(t \<in> D (P \\ A) \<or> (\<exists>(f:: nat \<Rightarrow> 'a event list). isInfHiddenRun f (P \\ A) B \<and> t \<in> range f))"
by (simp add:D_hiding) blast
thus ?case
proof(erule_tac disjE, auto simp add:D_hiding, goal_cases)
case (1 ta ua)
then show ?case
by (rule_tac x=ta in exI, rule_tac x="trace_hide ua (ev ` B) @ u" in exI, auto simp add: front_tickFree_append tickFree_def)
next
case (2 ua f xa)
then show ?case
apply(rule_tac x="f xa" in exI, rule_tac x="trace_hide ua (ev ` B) @ u" in exI, intro conjI disjI2)
apply(auto simp add: front_tickFree_append tickFree_def)
by (rule_tac x="f" in exI) (metis (no_types) filter_filter rangeI)
next
case (3 f xx)
note "3a" = 3
then show ?case
proof(cases "\<exists>i. f i \<in> D (P \\ A)")
case True
with 3 show ?thesis
proof(auto simp add:D_hiding, goal_cases)
case (1 i ta ua)
then show ?case
apply (rule_tac x=ta in exI, rule_tac x="trace_hide ua (ev ` B) @ u" in exI, intro conjI)
apply (metis (full_types) front_tickFree_append hiding_tickFree tickFree_append)
apply(subgoal_tac "trace_hide (f xx) (ev ` B) = trace_hide (f i) (ev ` B)", auto)
apply (metis (full_types) filter_append filter_filter)
by (metis (full_types) filter_append filter_filter)
next
case (2 i ua fa xa)
hence "trace_hide (f xx) (ev ` B) = trace_hide (f i) (ev ` B)" by metis
with 2 show ?case
apply (rule_tac x="fa xa" in exI, rule_tac x="trace_hide ua (ev ` B) @ u" in exI, intro conjI)
apply (metis (full_types) front_tickFree_append hiding_tickFree tickFree_append)
apply (simp_all)
apply(rule_tac disjI2, rule_tac x=fa in exI, auto)
by (metis (no_types) filter_filter)
qed
next
case False
with 3 have Falsebis:"\<forall>i. (f i \<in> T (P \\ A) \<and> f i \<notin> D (P \\ A))" by blast
with T_hiding[of P A] D_hiding[of P A] have "\<forall>i. (f i \<in> {trace_hide t (ev ` A) |t. (t, ev ` A) \<in> F P})"
by (metis (no_types, lifting) UnE)
hence ff0:"\<forall>i. (\<exists>t. f i = trace_hide t (ev ` A) \<and> t \<in> T P)" using F_T by fastforce
define ff where ff1:"ff = (\<lambda>i. SOME t. f i = trace_hide t (ev ` A) \<and> t \<in> T P)"
hence "inj ff" unfolding inj_def by (metis (mono_tags, lifting) "3"(4) ff0 strict_mono_eq tfl_some)
hence ff2:"infinite (range ff)" using finite_imageD by blast
{ fix tt i
assume "tt \<in> range ff"
then obtain k where "ff k = tt" using finite_nat_set_iff_bounded_le by blast
hence kk0:"f k = trace_hide tt (ev ` A) \<and> tt \<in> T P" using ff1 by (metis (mono_tags, lifting) ff0 someI_ex)
hence "set (take i tt) \<subseteq> set (f i) \<union> (ev ` A)"
proof(cases "k \<le> i")
case True
hence "set (f k) \<subseteq> set (f i)"
by (metis "3"(4) le_list_def set_append strict_mono_less_eq sup.cobounded1)
moreover from kk0 have "set (take i tt) \<subseteq> set (f k) \<union> (ev ` A)" using in_set_takeD by fastforce
ultimately show ?thesis by blast
next
case False
have a:"length (f i) \<ge> i" by (meson "3"(4) dual_order.trans le_add1 length_strict_mono)
have b:"f i \<le> f k" using "3"(4) False nat_le_linear strict_mono_less_eq by blast
with a have c:"take i (f k) \<le> (f i)"
by (metis append_eq_conv_conj le_add_diff_inverse le_list_def take_add)
from kk0[THEN conjunct1] have c1:"f k = (trace_hide (take i tt) (ev ` A))@(trace_hide (drop i tt) (ev ` A))"
by (metis append_take_drop_id filter_append)
have "length (trace_hide (take i tt) (ev ` A)) \<le> i"
by (metis length_filter_le length_take min.absorb2 nat_le_linear order.trans take_all)
with c1 have "take i (f k) \<ge> (trace_hide (take i tt) (ev ` A))" by (simp add: le_list_def)
with c obtain t where d:"f i = (trace_hide (take i tt) (ev ` A))@t"
by (metis append.assoc le_list_def)
then show ?thesis using in_set_takeD by fastforce
qed
} note ee=this
{ fix i
have "finite {(take i t)|t. t \<in> range ff}"
proof(induct i, simp)
case (Suc i)
have ff:"{take (Suc i) t|t. t \<in> range ff} \<subseteq> {(take i t)|t. t \<in> range ff} \<union>
(\<Union>e\<in>(set (f (Suc i)) \<union> (ev ` A)). {(take i t)@[e]|t. t \<in> range ff})" (is "?AA \<subseteq> ?BB")
proof
fix t
assume "t \<in> ?AA"
then obtain t' where hh:"t' \<in> range ff \<and> t = take (Suc i) t'"
using finite_nat_set_iff_bounded_le by blast
with ee[of t'] show "t \<in> ?BB"
proof(cases "length t' > i")
case True
hence ii:"take (Suc i) t' = take i t' @ [t'!i]" by (simp add: take_Suc_conv_app_nth)
with ee[of t' "Suc i"] have "t'!i \<in> set (f (Suc i)) \<union> (ev ` A)" by (simp add: hh)
with ii hh show ?thesis by blast
next
case False
hence "take (Suc i) t' = take i t'" by fastforce
with hh show ?thesis by auto
qed
qed
{ fix e
have "{x @ [e] |x. \<exists>t. x = take i t \<and> t \<in> range ff} = {take i t' @ [e] |t'. t' \<in> range ff}"
by auto
hence "finite({(take i t') @ [e] |t'. t'\<in>range ff})"
using finite_image_set[of _ "\<lambda>t. t@[e]", OF Suc] by auto
} note gg=this
have "finite(set (f (Suc i)) \<union> (ev ` A))" using 1(1) by simp
with ff gg Suc show ?case by (metis (no_types, lifting) finite_UN finite_Un finite_subset)
qed
} note ff=this
hence "\<forall>i. {take i t |t. t \<in> range ff} = {t. \<exists>t'. t = take i (ff t')}" by auto
with KoenigLemma[of "range ff"] ff ff2
obtain f' where gg:"strict_mono (f':: nat \<Rightarrow> 'a event list) \<and>
range f' \<subseteq> {t. \<exists>t'\<in>range ff. t \<le> t'}" by auto
{ fix i
from gg obtain n where aa:"f' i \<le> ff n" by blast
have "\<exists>t. f n = f 0 @ t" by (metis "3a"(4) le0 le_list_def strict_mono_less_eq)
with mono_trace_hide[OF aa, of A] ff0 ff1 have "\<exists>t. trace_hide (f' i) (ev ` A) \<le> f 0 @ t"
by (metis (mono_tags, lifting) someI_ex)
} note zz=this
{
define ff' where "ff' = (\<lambda>i. trace_hide (f' i) (ev ` A))"
with gg have "mono ff'"
by (rule_tac monoI, simp add: mono_trace_hide strict_mono_less_eq)
assume aa:"\<forall>i. trace_hide (f' i) (ev ` A) \<le> f 0"
with aa mono_constant have "\<exists>i. \<forall>j\<ge>i. ff' j = ff' i" using \<open>mono ff'\<close> ff'_def by blast
then obtain m where bb:"\<forall>j\<ge>m. ff' j = ff' m" by blast
have "ff' m \<in> D (P \\ A)"
proof(simp add:D_hiding, rule_tac x="f' m" in exI, rule_tac x="[]" in exI, intro conjI, simp, goal_cases)
case 1
from gg have "f' m < f' (Suc m)" by (meson lessI strict_monoD)
moreover from gg obtain k where "f' (Suc m) \<le> ff k" by blast
moreover have "ff k \<in> T P" by (metis (mono_tags, lifting) ff0 ff1 someI_ex)
ultimately show ?case
by (metis NF_NT append_Nil2 front_tickFree_mono is_processT le_list_def less_list_def)
next
case 2
then show ?case unfolding ff'_def by simp
next
case 3
then show ?case
proof(rule disjI2, rule_tac x="\<lambda>i. f' (m + i)" in exI, simp_all, intro conjI allI, goal_cases)
case 1
show ?case using gg[THEN conjunct1]
by (rule_tac strict_monoI, simp add: strict_monoD)
next
case (2 i)
from gg obtain k where "f' (m + i) \<le> ff k" by blast
moreover from ff0 ff1 have "ff k \<in> T P" by (metis (mono_tags, lifting) someI_ex)
ultimately have "f' (m + i) \<in> T P" using is_processT3_ST_pref by blast
then show ?case by (simp add: add.commute)
next
case (3 i)
then show ?case using bb[THEN spec, of "m+i", simplified] unfolding ff'_def by assumption
next
case 4
then show ?case by (metis Nat.add_0_right rangeI)
qed
qed
with gg False have "False"
by (metis (no_types, lifting) Falsebis aa append_Nil2 ff'_def front_tickFree_mono is_processT is_processT2_TR le_list_def)
}
with zz obtain m where hh:"trace_hide (f' m) (ev ` A) \<ge> f 0"
unfolding le_list_def by (metis append_eq_append_conv2)
from ff0 ff1 gg show ?thesis
proof(auto simp add:T_hiding, rule_tac x="f' m" in exI, rule_tac x=u in exI, intro conjI, simp_all add:3(3), goal_cases)
case 1
hence "f' m < f' (Suc m)" by (meson lessI strict_monoD)
moreover from gg obtain k where "f' (Suc m) \<le> ff k" by blast
moreover have "ff k \<in> T P" by (metis (mono_tags, lifting) ff0 ff1 someI_ex)
ultimately show ?case
by (metis NF_NT append_Nil2 front_tickFree_mono is_processT le_list_def less_list_def)
next
case 2
from gg obtain k where "f' m \<le> ff k" by blast
with ff0 ff1 mono_trace_hide[of "f' m"] have "trace_hide (f' m) (ev ` A) \<le> f k"
by (metis (mono_tags, lifting) someI_ex)
with mono_trace_hide[OF this, of B] mono_trace_hide[OF hh, of B] 3(6)[THEN spec, of k] 3(6)
show ?case by (metis (full_types) dual_order.antisym filter_filter)
next
case 3 show ?case
proof(rule disjI2, rule_tac x="\<lambda>i. f' (m + i)" in exI, simp_all, intro conjI allI, goal_cases)
case 1
then show ?case by (metis Nat.add_0_right rangeI)
next
case 2 with 3(4) show ?case
by (rule_tac strict_monoI, simp add: strict_monoD)
next
case (3 i)
from gg obtain k where "f' (m + i) \<le> ff k" by blast
moreover from ff0 ff1 have "ff k \<in> T P" by (metis (mono_tags, lifting) someI_ex)
ultimately have "f' (m + i) \<in> T P" using is_processT3_ST_pref by blast
then show ?case by (simp add: add.commute)
next
case (4 i)
from gg obtain k where "f' (m + i) \<le> ff k" by blast
with ff0 ff1 mono_trace_hide[of "f' (m + i)"] have ll:"trace_hide (f' (m + i)) (ev ` A) \<le> f k"
by (metis (mono_tags, lifting) someI_ex)
{ fix a b c assume "(a::'a event list) \<le> b" and "b \<le> c" and "c \<le> a" hence "b = c" by auto}
note jj=this
from jj[OF mono_trace_hide[OF hh, of B],
OF mono_trace_hide[THEN mono_trace_hide, of "f' m" "f' (m + i)" B A,
OF gg[THEN conjunct1, THEN strict_mono_mono,
THEN monoD, of "m" "m+i", simplified]]]
mono_trace_hide[OF ll, of B]
show ?case unfolding "3a"(6) [THEN spec, of k] by simp
qed
qed
qed
qed
qed
lemma hide_un_D: "finite A \<Longrightarrow> D ((P \\ A) \\ B) = D (P \\ (A \<union> B))"
using hide_un_D1 hide_un_D2 by blast
subsection\<open> Laws \<close>
lemma mono_hide_FD[simp]: "P \<le> Q \<Longrightarrow> P \\ A \<le> Q \\ A"
apply (auto simp: le_ref_def F_hiding D_hiding)
using F_subset_imp_T_subset by blast+
lemmas mono_hide_ref = mono_hiding
lemma hide_un: "finite A \<Longrightarrow> P \\ (A \<union> B) = (P \\ A) \\ B"
proof (simp add:Process_eq_spec, intro conjI, unfold F_hiding, goal_cases)
case 1
thus ?case (is "{(s, X). ?A s X} \<union> {(s, X). ?B s} = {(s, X). (\<exists>t. (?C1 s t \<and> (t, X \<union> ev ` B) \<in> ?C2 \<union> ?C3))} \<union> {(s, X). ?D s}")
proof(unfold F_hiding set_eq_subset Un_subset_iff, intro conjI, goal_cases)
case 1
then show ?case
by (auto, metis (no_types) filter_filter image_Un sup_commute sup_left_commute)
next
case 2
then show ?case
by (rule_tac Un_mono[of "{}", simplified], insert hide_un_D[of A P B], simp add:D_hiding)
next
case 3
have "{(s, X). (\<exists>t. (?C1 s t \<and> (t, X \<union> ev ` B) \<in> ?C2))} \<subseteq> {(s, X). ?A s X}"
by (auto, metis (no_types) filter_filter image_Un sup_commute sup_left_commute)
moreover have "{(s, X). (\<exists>t. (?C1 s t \<and> (t, X \<union> ev ` B) \<in> ?C3))} \<subseteq> {(s, X). ?B s}"
proof(auto,goal_cases)
case (1 ta u)
then show ?case using hiding_fronttickFree by blast
next
case (2 u f x)
then show ?case
apply(rule_tac x="f x" in exI, rule_tac x="trace_hide u (ev ` B)" in exI, auto)
using hiding_fronttickFree apply blast
apply(erule_tac x=f in allE) by (metis (no_types) filter_filter rangeI)
qed
moreover have "{(s, X). (\<exists>t. (?C1 s t \<and> (t, X \<union> ev ` B) \<in> ?C2 \<union> ?C3))} =
{(s, X). (\<exists>t. (?C1 s t \<and> (t, X \<union> ev ` B) \<in> ?C2 ))} \<union>
{(s, X). (\<exists>t. (?C1 s t \<and> (t, X \<union> ev ` B) \<in> ?C3))}" by blast
ultimately show ?case by (metis (no_types, lifting) Un_mono)
next
case 4
then show ?case
by (rule_tac Un_mono[of "{}", simplified], insert hide_un_D[of A P B], simp add:D_hiding)
qed
next
case 2
then show ?case by (simp add: hide_un_D)
qed
lemma hide_set_bot: "(\<bottom> \\ A) = \<bottom>"
apply(auto simp add:Process_eq_spec D_hiding F_hiding F_UU D_UU)
using hiding_fronttickFree apply blast
using front_tickFree_append hiding_tickFree apply blast
using front_tickFree_append hiding_tickFree apply blast
apply (metis (full_types) append_Nil filter.simps(1) tickFree_Nil tickFree_implies_front_tickFree)
using front_tickFree_append hiding_tickFree apply blast
using front_tickFree_append hiding_tickFree apply blast
using tickFree_Nil by fastforce
lemma hide_set_STOP: "(STOP \\ A) = STOP"
apply(auto simp add:Process_eq_spec D_hiding F_hiding F_STOP D_STOP T_STOP)
by (metis (full_types) lessI less_irrefl strict_mono_eq) +
lemma hide_set_SKIP: "(SKIP \\ A) = SKIP"
apply(auto simp add:Process_eq_spec D_hiding F_hiding F_SKIP D_SKIP T_SKIP split:if_splits)
apply (metis filter.simps(1) non_tickFree_tick)
apply (metis (full_types) hiding_tickFree n_not_Suc_n non_tickFree_tick strict_mono_eq)
apply (metis (full_types) hiding_tickFree n_not_Suc_n non_tickFree_tick strict_mono_eq)
apply (metis event.distinct(1) filter.simps(1) imageE)
apply (metis event.distinct(1) filter.simps(1) filter.simps(2) imageE)
by (metis (full_types) hiding_tickFree n_not_Suc_n non_tickFree_tick strict_mono_eq)
lemma hide_set_empty: "(P \\ {}) = P"
apply(auto simp add:Process_eq_spec D_hiding F_hiding is_processT7 is_processT8_S strict_mono_eq)
by (metis append_Nil2 front_tickFree_implies_tickFree front_tickFree_single is_processT
nonTickFree_n_frontTickFree)
subsection\<open> Multi-Operators laws \<close>
lemma hide_ndet: "(P \<sqinter> Q) \\ A = ((P \\ A) \<sqinter> (Q \\ A))"
proof(auto simp add:Process_eq_spec D_hiding F_hiding,
simp_all add: F_ndet T_ndet D_ndet D_hiding F_hiding, goal_cases)
case (1 b t)
then show ?case by blast
next
case (2 b t u)
then show ?case by blast
next
case (3 b u f x)
from 3(4) have A:"infinite ({i. f i \<in> T P}) \<or> infinite ({i. f i \<in> T Q})"
using finite_nat_set_iff_bounded by auto
hence "(\<forall>i. f i \<in> T P) \<or> (\<forall>i. f i \<in> T Q)"
by (metis (no_types, lifting) 3(3) finite_nat_set_iff_bounded
is_processT3_ST_pref mem_Collect_eq not_less strict_mono_less_eq)
with A show ?case by (metis (no_types, lifting) 3(1,2,3,5) rangeI)
next
case (4 a b)
then show ?case by blast
next
case (5 t u)
then show ?case by blast
next
case (6 u f x)
from 6(4) have A:"infinite ({i. f i \<in> T P}) \<or> infinite ({i. f i \<in> T Q})"
using finite_nat_set_iff_bounded by auto
hence "(\<forall>i. f i \<in> T P) \<or> (\<forall>i. f i \<in> T Q)"
by (metis (no_types, lifting) 6(3) finite_nat_set_iff_bounded
is_processT3_ST_pref mem_Collect_eq not_less strict_mono_less_eq)
with A show ?case by (metis (no_types, lifting) 6(1,2,3,5) rangeI)
next
case (7 x)
then show ?case by blast
qed
lemma hide_mprefix_distr: "(B \<inter> A) = {} \<Longrightarrow> ((Mprefix A P) \\ B) = (Mprefix A (\<lambda>x. ((P x) \\ B)))"
proof (auto simp add: Process_eq_spec,
simp_all add: F_Mprefix T_Mprefix D_Mprefix D_hiding F_hiding,
goal_cases)
case (1 x b) then show ?case
proof(elim exE disjE conjE, goal_cases)
case (1 t)
then show ?case by (simp add: inf_sup_distrib2)
next
case (2 t a)
then show ?case
by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2)
imageE list.sel(1) list.sel(3) list.collapse neq_Nil_conv)
next
case (3 t u a)
hence "hd t \<notin> ev ` B" by force
with 3 have"x = hd t # trace_hide (tl t) (ev ` B) @ u"
by (metis append_Cons filter.simps(2) list.exhaust_sel)
with 3 show ?case by (metis list.distinct(1) list.sel(1) list.sel(3) tickFree_tl)
next
case (4 t u f)
then obtain k where kk:"t = f k" by blast
from 4 obtain a where "f 1 \<noteq> [] \<and> ev a = hd (f 1)" and aa:"a \<in> A"
by (metis less_numeral_extra(1) nil_le not_less strict_mono_less_eq)
with 4(1) 4(6)[THEN spec, of 0] 4(7)[THEN spec, of 1] have "f 0 \<noteq> [] \<and> hd (f 0) = ev a"
apply auto
apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject
filter.simps(2) hd_Cons_tl imageE list.distinct(1))
apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2)
hd_Cons_tl imageE list.distinct(1))
by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2)
hd_Cons_tl imageE list.sel(1))
with 4(1, 7) aa have nf: "\<forall>i. f i \<noteq> [] \<and> hd (f i) = ev a"
by (metis (mono_tags, hide_lams) "4"(5) append_Cons le_list_def le_simps(2) list.distinct(1)
list.exhaust_sel list.sel(1) neq0_conv old.nat.distinct(1) strict_mono_less_eq)
with 4(5) have sm:"strict_mono (tl \<circ> f)" by (simp add: less_tail strict_mono_def)
with 4 aa kk nf show ?case
apply(rule_tac disjI2, intro conjI)
apply (metis (no_types, lifting) Nil_is_append_conv disjoint_iff_not_equal event.inject
filter.simps(2) hd_Cons_tl imageE list.distinct(1))
apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2)
hd_Cons_tl hd_append2 image_iff list.distinct(1) list.sel(1))
apply(rule_tac x=a in exI, intro conjI disjI2)
apply (metis disjoint_iff_not_equal event.inject filter.simps(2)
hd_Cons_tl hd_append2 image_iff list.distinct(1) list.sel(1))
apply(rule_tac x="tl t" in exI, rule_tac x="u" in exI, intro conjI, simp_all)
apply (metis tickFree_tl)
apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2)
hd_Cons_tl imageE list.distinct(1) list.sel(3) tl_append2)
apply(subst disj_commute, rule_tac disjCI)
apply(rule_tac x="tl \<circ> f" in exI, intro conjI)
apply auto
apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl list.sel(3))
done
qed
next
case (2 x b)
then show ?case proof(elim exE disjE conjE, goal_cases)
case 1 then show ?case
apply(rule_tac disjI1, rule_tac x="[]" in exI)
by (simp add: disjoint_iff_not_equal inf_sup_distrib2)
next
case (2 a t) then show ?case
apply(rule_tac disjI1, rule_tac x="(ev a)#t" in exI)
using list.collapse by fastforce
next
case (3 a t u)
then show ?case
apply(rule_tac disjI2, rule_tac x="(ev a)#t" in exI, rule_tac x=u in exI)
using list.collapse by fastforce
next
case (4 a t u f)
then show ?case
apply(rule_tac disjI2)
apply(rule_tac x="(ev a) # t" in exI, rule_tac x="u" in exI, intro conjI, simp)
apply auto[1]
using hd_Cons_tl apply fastforce
apply(rule_tac disjI2, rule_tac x="\<lambda>i. (ev a) # (f i)" in exI)
by (auto simp add: less_cons strict_mono_def)
qed
next
case (3 x) then show ?case
proof(elim exE disjE conjE, goal_cases)
case (1 t u a)
hence aa: "hd (trace_hide t (ev ` B)) = ev a \<and> trace_hide t (ev ` B) \<noteq> []"
by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl
imageE list.distinct(1) list.sel(1))
with 1 have "hd x = ev a \<and> x \<noteq> []" by simp
with 1 show ?case
apply(intro conjI, simp_all, rule_tac x=a in exI, simp)
apply(rule_tac x="tl t" in exI, rule_tac x="u" in exI, intro conjI, simp_all)
using tickFree_tl apply blast
using aa by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2)
hd_Cons_tl imageE list.sel(3) tl_append2)
next
case (2 t u f)
then obtain k where kk:"t = f k" by blast
from 2 obtain a where "f 1 \<noteq> [] \<and> ev a = hd (f 1)" and aa:"a \<in> A"
by (metis less_numeral_extra(1) nil_le not_less strict_mono_less_eq)
with 2(1) 2(6)[THEN spec, of 0] 2(7)[THEN spec, of 1] have "f 0 \<noteq> [] \<and> hd (f 0) = ev a"
apply auto
apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject
filter.simps(2) hd_Cons_tl imageE list.distinct(1))
apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2)
hd_Cons_tl imageE list.distinct(1))
by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2)
hd_Cons_tl imageE list.sel(1))
with 2(1, 7) aa have nf: "\<forall>i. f i \<noteq> [] \<and> hd (f i) = ev a"
by (metis (mono_tags, hide_lams) 2(5) append_Cons le_list_def le_simps(2) list.distinct(1)
list.exhaust_sel list.sel(1) neq0_conv old.nat.distinct(1) strict_mono_less_eq)
with 2(5) have sm:"strict_mono (tl \<circ> f)" by (simp add: less_tail strict_mono_def)
from 2(1,4) nf aa kk have x1:"x \<noteq> []"
by (metis Nil_is_append_conv disjoint_iff_not_equal event.inject filter.simps(2)
hd_Cons_tl imageE list.distinct(1))
with 2(1,4) nf aa kk have x2: "hd (trace_hide t (ev ` B)) = ev a \<and> trace_hide t (ev ` B) \<noteq> []"
by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl
imageE list.distinct(1) list.sel(1))
with 2(1,4) nf aa kk x1 have x3:"hd x = ev a" by simp
with 2 aa kk nf sm x1 show ?case
apply(intro conjI, simp_all)
apply(rule_tac x="tl t" in exI, rule_tac x="u" in exI, intro conjI, simp_all)
apply (metis tickFree_tl)
apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2)
hd_Cons_tl imageE list.distinct(1) list.sel(3) tl_append2)
apply(subst disj_commute, rule_tac disjCI)
apply(rule_tac x="tl \<circ> f" in exI, intro conjI)
apply auto
apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl list.sel(3))
by (metis (no_types, lifting) filter.simps(2) hd_Cons_tl list.sel(3))
qed
next
case (4 x) then show ?case
proof(elim exE disjE conjE, goal_cases)
case (1 a t u)
then show ?case
apply(rule_tac x="(ev a)#t" in exI, rule_tac x=u in exI)
using list.collapse by fastforce
next
case (2 a t u f)
then show ?case
apply(rule_tac x="(ev a) # t" in exI, rule_tac x="u" in exI, intro conjI, simp_all)
apply auto[1]
using hd_Cons_tl apply fastforce
apply(rule_tac disjI2, rule_tac x="\<lambda>i. (ev a) # (f i)" in exI)
by (auto simp add: less_cons strict_mono_def)
qed
qed
lemma no_hide_read: "(\<forall>y. c y \<notin> B) \<Longrightarrow> ((c `?` x \<rightarrow> (P x)) \\ B) = (c `?` x \<rightarrow> ((P x) \\ B))"
by (simp add: read_def o_def, subst hide_mprefix_distr, auto)
lemma no_hide_write0: "a \<notin> B \<Longrightarrow> ((a \<rightarrow> P) \\ B) = (a \<rightarrow> (P \\ B))"
by (simp add: hide_mprefix_distr write0_def)
lemma hide_write0: "a \<in> B \<Longrightarrow> ((a \<rightarrow> P) \\ B) = (P \\ B)"
proof (auto simp add: write0_def Process_eq_spec,
simp_all add: F_Mprefix T_Mprefix D_Mprefix D_hiding F_hiding,
goal_cases)
case (1 x b)
then show ?case
apply(elim exE disjE conjE)
apply (metis filter.simps(2) hd_Cons_tl image_eqI)
apply (metis (no_types, lifting) filter.simps(2) image_eqI list.sel(1)
list.sel(3) neq_Nil_conv tickFree_tl)
proof(goal_cases)
case (1 t u f)
have fS: "strict_mono (f \<circ> Suc)" by (metis "1"(5) Suc_mono comp_def strict_mono_def)
from 1 have aa:"\<forall>i. f (Suc i) \<noteq> []"
by (metis (full_types) less_Suc_eq_le less_irrefl nil_le strict_mono_less_eq)
with fS have sm:"strict_mono (tl \<circ> f \<circ> Suc)" by (simp add: less_tail strict_mono_def)
with 1 aa show ?case
apply(subst disj_commute, rule_tac disjCI, simp)
apply(rule_tac x="tl (f 1)" in exI, rule_tac x="u" in exI, intro conjI, simp_all)
apply (metis hiding_tickFree imageE tickFree_tl)
apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl image_eqI rangeE)
apply(subst disj_commute, rule_tac disjCI)
apply(rule_tac x="tl \<circ> f \<circ> Suc" in exI, intro conjI)
apply auto
apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl list.sel(3))
done
qed
next
case (2 aa b)
then show ?case
apply(elim exE disjE conjE)
apply (metis (no_types, lifting) filter.simps(2) image_eqI list.distinct(1)
list.sel(1) list.sel(3))
proof(goal_cases)
case (1 t u)
then show ?case by (rule_tac disjI2, rule_tac x="(ev a)#t" in exI, auto)
next
case (2 t u f)
then show ?case
apply(rule_tac disjI2)
apply(rule_tac x="(ev a) # t" in exI, rule_tac x="u" in exI, intro conjI, simp_all)
apply(rule_tac disjI2)
apply(rule_tac x="\<lambda>i. (ev a) # (f i)" in exI, intro conjI)
by (auto simp add: less_cons strict_mono_def)
qed
next
case (3 x)
then show ?case
apply(elim exE disjE conjE)
apply(rule_tac x="tl t" in exI, rule_tac x="u" in exI, intro conjI, simp_all)
using tickFree_tl apply blast
apply (metis filter.simps(2) hd_Cons_tl image_eqI)
proof(goal_cases)
case (1 t u f)
have fS: "strict_mono (f \<circ> Suc)" by (metis "1"(5) Suc_mono comp_def strict_mono_def)
from 1 have aa:"\<forall>i. f (Suc i) \<noteq> []"
by (metis (full_types) less_Suc_eq_le less_irrefl nil_le strict_mono_less_eq)
with fS have sm:"strict_mono (tl \<circ> f \<circ> Suc)" by (simp add: less_tail strict_mono_def)
with 1 aa show ?case
apply(rule_tac x="tl (f 1)" in exI, rule_tac x="u" in exI, intro conjI, simp_all)
apply (metis hiding_tickFree imageE tickFree_tl)
apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl image_eqI rangeE)
apply(subst disj_commute, rule_tac disjCI)
apply(rule_tac x="tl \<circ> f \<circ> Suc" in exI, intro conjI)
apply auto
apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl list.sel(3))
done
qed
next
case (4 x)
then show ?case
apply(elim exE disjE conjE)
proof(goal_cases)
case (1 t u)
then show ?case by (rule_tac x="(ev a)#t" in exI, auto)
next
case (2 t u f)
then show ?case
apply(rule_tac x="(ev a) # t" in exI, rule_tac x="u" in exI, intro conjI, simp_all)
apply(rule_tac disjI2)
apply(rule_tac x="\<lambda>i. (ev a) # (f i)" in exI, intro conjI)
by (auto simp add: less_cons strict_mono_def)
qed
qed
lemma no_hide_write: "(\<forall>y. c y \<notin> B) \<Longrightarrow> ((c `!` a \<rightarrow> P) \\ B) = (c `!` a \<rightarrow> (P \\ B))"
by(simp add: write_def, subst hide_mprefix_distr, auto)
lemma hide_write: "(c a) \<in> B ==> ((c `!` a \<rightarrow> P) \\ B) = (P \\ B)"
by (simp add: write_def hide_write0 mprefix_singl)
section\<open> The Sync Operator Laws \<close>
subsection\<open> Preliminaries \<close>
lemma syncTlEmpty:"a setinterleaves (([], u), A) \<Longrightarrow> tl a setinterleaves (([], tl u), A)"
by (case_tac u, simp, case_tac a, simp_all split:if_splits)
lemma syncHd_Tl: "a setinterleaves ((t, u), A) \<and> hd t \<in> A \<and> hd u \<notin> A \<Longrightarrow> hd a = hd u \<and> tl a setinterleaves ((t, tl u), A)"
by (case_tac u) (case_tac t, auto split:if_splits)+
lemma syncHdAddEmpty: "(tl a) setinterleaves (([], u), A) \<and> hd a \<notin> A \<and> a \<noteq> [] \<Longrightarrow> a setinterleaves (([], hd a # u), A)"
using hd_Cons_tl by fastforce
lemma syncHdAdd: "(tl a) setinterleaves ((t, u), A) \<and> hd a \<notin> A \<and> hd t \<in> A \<and> a \<noteq> [] \<Longrightarrow> a setinterleaves ((t, hd a # u), A)"
by (case_tac a, simp, case_tac t, auto)
lemmas syncHdAdd1 = syncHdAdd[of "a#r", simplified] for a r
lemma syncSameHdTl:"a setinterleaves ((t, u), A) \<and> hd t \<in> A \<and> hd u \<in> A \<Longrightarrow> hd t = hd u \<and> hd a = hd t \<and> (tl a) setinterleaves ((tl t, tl u), A)"
by (case_tac u) (case_tac t, auto split:if_splits)+
lemma synSingleHeadAdd: "a setinterleaves ((t, u), A) \<and> h \<notin> A\<Longrightarrow> (h#a) setinterleaves ((h#t, u), A)"
by (case_tac u, auto split:if_splits)
lemma TickLeftSync:"tick \<in> A \<and> front_tickFree t \<and> s setinterleaves (([tick], t), A) \<Longrightarrow> s = t \<and> last t = tick"
proof(induct t arbitrary: s)
case Nil
then show ?case by (simp add: Nil.prems)
next
case (Cons a t)
then show ?case
apply (auto split:if_splits)
using emptyLeftProperty apply blast
apply (metis last_ConsR last_snoc nonTickFree_n_frontTickFree tickFree_Cons)
by (metis append_Cons append_Nil front_tickFree_mono)+
qed
lemma EmptyLeftSync:"s setinterleaves (([], t), A) \<Longrightarrow> s = t \<and> set t \<inter> A = {}"
by (meson Int_emptyI emptyLeftNonSync emptyLeftProperty)
lemma tick_T_F:"t@[tick] \<in> T P \<Longrightarrow> (t@[tick], X) \<in> F P"
using append_T_imp_tickFree is_processT5_S7 by force
lemma event_set: "(e::'a event) \<in> insert tick (range ev)"
by (metis event.exhaust insert_iff rangeI)
lemma mprefix_Par_Int_distr1_D: " A \<subseteq> C \<Longrightarrow> B \<subseteq> C \<Longrightarrow> D (Mprefix A P \<lbrakk>C\<rbrakk> Mprefix B Q) = D ( \<box> x \<in> A \<inter> B \<rightarrow> (P x \<lbrakk>C\<rbrakk> Q x))"
apply(auto,simp_all add:D_sync F_sync F_Mprefix T_Mprefix D_Mprefix)
apply(elim exE disjE conjE)
apply (metis (no_types, lifting) Sync.sym empty_iff image_iff insertI2 list.exhaust_sel setinterleaving.simps(2) subsetCE)
proof(goal_cases)
case (1 x t u r v a aa)
from 1(1,2,6,8,11,13) have aa1: "hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by blast
from 1(5,6,7,8,11,13) aa1 have aa2: " hd x \<in> ev ` (A \<inter> B)"
by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff syncSameHdTl)
then show ?case
using 1(3,4,5,6,7,9,10,13,14) by (metis (no_types, lifting) Nil_is_append_conv aa1 empty_setinterleaving event.inject hd_append2 syncSameHdTl tickFree_tl tl_append2)
next
case (2 x t u r v a)
then show ?case
by (metis (no_types, lifting) Sync.si_empty3 equals0D imageE image_eqI insertI2 list.exhaust_sel subsetCE)
next
case (3 x t u r v a aa)
from 3(1,2,6,8,11,13) have aa1: "hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by blast
from 3(5,6,7,8,11,13) aa1 have aa2: " hd x \<in> ev ` (A \<inter> B)"
by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff syncSameHdTl)
then show ?case
using 3(3,4,5,6,7,9,10,13,14) by (metis (no_types, lifting) Nil_is_append_conv aa1 empty_setinterleaving event.inject hd_append2 syncSameHdTl tickFree_tl tl_append2)
next
case (4 x t u r v a)
then show ?case
by (metis (no_types, lifting) Sync.si_empty3 equals0D imageE image_eqI insertI2 list.exhaust_sel subsetCE)
next
case (5 x t u r v a aa)
from 5(1,2,6,8,11,13) have aa1: "hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by blast
from 5(5,6,7,8,11,13) aa1 have aa2: " hd x \<in> ev ` (A \<inter> B)"
by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff syncSameHdTl)
then show ?case
using 5(3,4,5,6,7,9,10,13,14) by (metis aa1 append_Nil2 empty_setinterleaving event.inject syncSameHdTl)
next
case (6 x t u r v a)
then show ?case
by (metis (no_types, lifting) Sync.si_empty3 equals0D imageE image_eqI insertI2 list.exhaust_sel subsetCE)
next
case (7 x t u r v a aa)
from 7(1,2,6,8,11,13) have aa1: "hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by blast
from 7(5,6,7,8,11,13) aa1 have aa2: " hd x \<in> ev ` (A \<inter> B)"
by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff syncSameHdTl)
then show ?case
using 7(3,4,5,6,7,9,10,13,14) by (metis aa1 append_Nil2 empty_setinterleaving event.inject syncSameHdTl)
next
case (8 x)
then show ?case
apply(elim exE disjE conjE)
proof(goal_cases)
case (1 a t u r v)
obtain r1 t1 u1 where aa0: "r1=hd x#r\<and>t1=hd x#t\<and>u1=hd x#u"
by auto
from 1(3,4,5,7,8,9) have aa1: "tickFree r1\<and>x = r1 @ v"
by (metis Cons_eq_appendI aa0 event.distinct(1) list.exhaust_sel tickFree_Cons)
from 1(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>t1 \<noteq> []"
using aa0 subsetCE by auto
from 1(4,5,10,11) aa0 have aa3: "(tl t1) \<in> D (P a)\<and>(tl u1) \<in> T (Q a)\<and> ev a = hd t1\<and>ev a = hd u1\<and> hd t1 \<in> ev ` A\<and> hd u1 \<in> ev ` B"
by auto
then show ?case
using "1"(6) aa1 aa2 by fastforce
next
case (2 a t u r v)
obtain r1 t1 u1 where aa0: "r1=hd x#r\<and>t1=hd x#t\<and>u1=hd x#u"
by auto
from 2(3,4,5,7,8,9) have aa1: "tickFree r1\<and>x = r1 @ v"
by (metis Cons_eq_appendI aa0 event.distinct(1) list.exhaust_sel tickFree_Cons)
from 2(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>t1 \<noteq> []"
using aa0 subsetCE by auto
from 2(4,5,10,11) aa0 have aa3: "(tl t1) \<in> D (Q a)\<and>(tl u1) \<in> T (P a)\<and> ev a = hd t1\<and>ev a = hd u1\<and> hd t1 \<in> ev ` A\<and> hd u1 \<in> ev ` B"
by auto
then show ?case
using "2"(6) aa1 aa2 by fastforce
next
case (3 a t u r v)
obtain r1 t1 u1 where aa0: "r1=hd x#r\<and>t1=hd x#t\<and>u1=hd x#u"
by auto
from 3(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>t1 \<noteq> []"
using aa0 subsetCE by auto
from 3(3,4,5,7,8,10,11) aa0 have aa3: "(tl t1) \<in> D (P a)\<and>(tl u1) \<in> T (Q a)\<and> ev a = hd t1\<and>ev a = hd u1\<and> hd t1 \<in> ev ` A\<and> hd u1 \<in> ev ` B\<and>x = r1 @ v"
by (metis (no_types, lifting) Int_lower1 Int_lower2 append_Nil2 imageE image_eqI list.exhaust_sel list.sel(1) list.sel(3) subsetCE)
then show ?case
by (metis (no_types, lifting) "3"(6) "3"(7) aa2 event.inject imageE)
next
case (4 a t u r v)
obtain r1 t1 u1 where aa0: "r1=hd x#r\<and>t1=hd x#t\<and>u1=hd x#u"
by auto
from 4(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>t1 \<noteq> []"
using aa0 subsetCE by auto
from 4(3,4,5,7,8,10,11) aa0 have aa3: "(tl t1) \<in> D (Q a)\<and>(tl u1) \<in> T (P a)\<and> ev a = hd t1\<and>ev a = hd u1\<and> hd t1 \<in> ev ` A\<and> hd u1 \<in> ev ` B\<and>x = r1 @ v"
by (metis (no_types, lifting) Int_lower1 Int_lower2 append_Nil2 imageE image_eqI list.exhaust_sel list.sel(1) list.sel(3) subsetCE)
then show ?case
by (metis (no_types, lifting) "4"(6) "4"(7) aa2 event.inject imageE)
qed
qed
lemma hide_interleave: "A \<inter> C = {} \<Longrightarrow> r setinterleaves ((t, u), C) \<Longrightarrow> (trace_hide r A) setinterleaves ((trace_hide t A, trace_hide u A), C)"
proof(induct r arbitrary:t u)
case Nil
then show ?case
using EmptyLeftSync empty_setinterleaving by fastforce
next
case (Cons a r)
then show ?case
apply(cases t) using EmptyLeftSync apply fastforce
apply(cases u) apply fastforce
- proof(simp split:if_splits, goal_cases)
+ using [[simproc del: defined_all]] proof (simp split:if_splits, goal_cases)
case (1 a list aa lista)
then show ?case by fastforce
next
case (2 a list aa lista)
then show ?case
apply(erule_tac disjE)
using Cons(1)[of list u]
apply (metis (no_types, lifting) filter.simps(2) synSingleHeadAdd)
using Cons(1)[of t lista]
by (metis (no_types, lifting) Sync.sym filter.simps(2) synSingleHeadAdd)
next
case (3 a list aa lista)
then show ?case by fastforce
qed
qed
lemma non_sync_interleaving: "(set t \<union> set u) \<inter> C = {} \<Longrightarrow> setinterleaving (t, C, u) \<noteq> {}" (* FINALLY NON-USED*)
proof(induct t arbitrary:u)
case Nil
then show ?case
by (metis Un_empty_left disjoint_iff_not_equal emptyLeftSelf empty_iff empty_set)
next
case (Cons a t)
then show ?case
proof(induct u)
case Nil
then show ?case using Sync.sym by auto
next
case (Cons a u)
then show ?case by auto
qed
qed
lemma interleave_hide: "A \<inter> C = {} \<Longrightarrow> ra setinterleaves ((trace_hide t A, trace_hide u A), C) \<Longrightarrow> \<exists>r. ra = trace_hide r A \<and> r setinterleaves ((t, u), C)"
proof(induct "length t + length u" arbitrary:ra t u rule:nat_less_induct)
case Ind:1
then show ?case
proof (cases t)
case Nilt: Nil
from Ind(2) have a:"set (trace_hide u A) \<inter> C = {} \<Longrightarrow> set u \<inter> C = {}" by auto
hence b: "u setinterleaves ((t, u), C)"
by (metis (no_types, lifting) Ind(3) EmptyLeftSync Nilt disjoint_iff_not_equal emptyLeftSelf filter.simps(1))
then show ?thesis
apply(rule_tac x=u in exI)
using EmptyLeftSync[of "ra" C "trace_hide u A"] a b Cons Nilt Ind(3) by auto
next
case Const: (Cons ta tlist)
then show ?thesis
proof(cases u)
case Nilu: Nil
from Ind(2) have a:"set (trace_hide t A) \<inter> C = {} \<Longrightarrow> set t \<inter> C = {}" by auto
hence b: "t setinterleaves ((t, u), C)"
by (metis Ind(3) Nilu EmptyLeftSync disjoint_iff_not_equal emptyLeftSelf filter.simps(1) Sync.sym)
then show ?thesis
apply(rule_tac x=t in exI)
using EmptyLeftSync[of "ra" C "trace_hide t A"] a b Ind Nilu by (metis Sync.sym filter.simps(1))
next
case Consu: (Cons ua ulist)
with Const Ind(2,3) show ?thesis
proof(auto split:if_splits simp del:setinterleaving.simps, goal_cases)
case 1
then show ?case
proof (auto, goal_cases)
case (1 raa)
with Ind(1)[THEN spec, of "length tlist + length u", simplified Const, simplified,
THEN spec, THEN spec, of tlist u, simplified Ind, simplified,
THEN spec, of raa] obtain r where
"raa = trace_hide r A \<and> r setinterleaves ((tlist, u), C)" by auto
then show ?case using "1"(4) Consu by force
next
case (2 raa)
with Ind(1)[THEN spec, of "length t + length ulist", simplified Consu, simplified,
THEN spec, THEN spec, of t ulist, simplified Ind, simplified,
THEN spec, of raa] obtain r where
"raa = trace_hide r A \<and> r setinterleaves ((t, ulist), C)" by auto
then show ?case using "2"(2) "2"(4) by force
next
case (3 raa)
with Ind(1)[THEN spec, of "length tlist + length ulist", simplified Const Consu, simplified,
THEN spec, THEN spec, of tlist ulist, simplified Ind, simplified,
THEN spec, of raa] obtain r where
"raa = trace_hide r A \<and> r setinterleaves ((tlist, ulist), C)" by auto
then show ?case using "3"(4) by force
next
case (4 raa)
with Ind(1)[THEN spec, of "length t + length ulist", simplified Consu, simplified,
THEN spec, THEN spec, of t ulist, simplified Ind, simplified,
THEN spec, of raa] obtain r where
"raa = trace_hide r A \<and> r setinterleaves ((t, ulist), C)" by auto
then show ?case using "4"(5) Const by force
next
case (5 raa)
with Ind(1)[THEN spec, of "length tlist + length u", simplified Const, simplified,
THEN spec, THEN spec, of tlist u, simplified Ind, simplified,
THEN spec, of raa] obtain r where
"raa = trace_hide r A \<and> r setinterleaves ((tlist, u), C)" by auto
then show ?case using "5"(4) Consu by force
next
case (6 raa)
with Ind(1)[THEN spec, of "length t + length ulist", simplified Consu, simplified,
THEN spec, THEN spec, of t ulist, simplified Ind, simplified,
THEN spec, of raa] obtain r where
"raa = trace_hide r A \<and> r setinterleaves ((t, ulist), C)" by auto
then show ?case using "6"(5) Const by force
next
case (7 raa)
with Ind(1)[THEN spec, of "length tlist + length u", simplified Const, simplified,
THEN spec, THEN spec, of tlist u, simplified Ind, simplified,
THEN spec, of raa] obtain r where
"raa = trace_hide r A \<and> r setinterleaves ((tlist, u), C)" by auto
then show ?case using "7"(4) Consu by force
qed
next
case 2
with Ind(1)[THEN spec, of "length t + length ulist", simplified Consu, simplified,
THEN spec, THEN spec, of t ulist, simplified Ind, simplified,
THEN spec, of ra] obtain r where
"ra = trace_hide r A \<and> r setinterleaves ((t, ulist), C)" by auto
then show ?case
apply(rule_tac x="ua#r" in exI)
using "2"(5) Const Ind.prems(1) by auto
next
case 3
with Ind(1)[THEN spec, of "length tlist + length u", simplified Const, simplified,
THEN spec, THEN spec, of tlist u, simplified Ind, simplified,
THEN spec, of ra] obtain r where
"ra = trace_hide r A \<and> r setinterleaves ((tlist, u), C)" by auto
then show ?case
apply(rule_tac x="ta#r" in exI)
using "3"(4) Consu Ind.prems(1) by auto
next
case 4
with Ind(1)[THEN spec, of "length tlist + length ulist", simplified Const Consu, simplified,
THEN spec, THEN spec, of tlist ulist, simplified Ind, simplified,
THEN spec, of ra] obtain r where
"ra = trace_hide r A \<and> r setinterleaves ((tlist, ulist), C)" by auto
then show ?case
apply(rule_tac x="ta#ua#r" in exI)
using "4"(4,5) Consu Const Ind.prems(1) apply auto
using synSingleHeadAdd apply blast
by (metis Sync.sym synSingleHeadAdd)
qed
qed
qed
qed
lemma interleave_size: "s setinterleaves ((t,u), C) \<Longrightarrow> length s = length t + length u - length(filter (\<lambda>x. x\<in>C) t)"
proof(induct s arbitrary:t u)
case Nil
then show ?case using EmptyLeftSync empty_setinterleaving by fastforce
next
case (Cons a list)
then show ?case
apply(cases t) using emptyLeftProperty apply fastforce
apply(cases u)
apply (metis (no_types, lifting) Sync.sym add_diff_cancel_right' emptyLeftNonSync emptyLeftProperty filter_empty_conv)
proof(auto split:if_splits, goal_cases 11 12 13 14 15)
case (11 tlist ulist)
then show ?case
by (metis (no_types, lifting) Suc_diff_le le_add1 length_filter_le order_trans)
next
case (12 ta tlist ulist)
with 12(3)[rule_format, of "ta#tlist" ulist] show ?case
by simp (metis Suc_diff_le le_add1 length_filter_le order_trans)
next
case (13 tlist ua ulist)
with 13(3)[rule_format, of tlist "ua#ulist"] show ?case
by simp (metis Suc_diff_le le_less_trans length_filter_le less_SucI less_Suc_eq_le less_add_Suc1)
next
case (14 ta tlist ulist)
with 14(3)[rule_format, of "ta#tlist" ulist] show ?case
by simp (metis Suc_diff_le le_less_trans length_filter_le less_SucI less_Suc_eq_le less_add_Suc1)
next
case (15 tlist ua ulist)
with 15(3)[rule_format, of tlist "ua#ulist"] show ?case
by simp (metis Suc_diff_le le_less_trans length_filter_le less_SucI less_Suc_eq_le less_add_Suc1)
qed
qed
lemma interleave_eq_size: "s setinterleaves ((t,u), C) \<Longrightarrow> s' setinterleaves ((t,u), C) \<Longrightarrow> length s = length s'"
by (simp add: interleave_size)
lemma interleave_set: "s setinterleaves ((t,u), C) \<Longrightarrow> set t \<union> set u \<subseteq> set s"
proof(induct s arbitrary: t u)
case Nil
then show ?case using EmptyLeftSync empty_setinterleaving by blast
next
case (Cons a s)
then show ?case
apply(cases t) using emptyLeftProperty apply fastforce
apply(cases u) apply (metis Sync.sym Un_empty_right emptyLeftProperty empty_set eq_refl)
by (auto simp add: subset_iff split:if_splits)
qed
lemma interleave_order: "s setinterleaves ((t1@t2,u), C) \<Longrightarrow> set(t2) \<subseteq> set(drop (length t1) s)"
proof(induct s arbitrary: t1 t2 u)
case Nil
then show ?case using empty_setinterleaving by auto
next
case (Cons a s)
then show ?case
apply(cases t1) using append_self_conv2 interleave_set apply fastforce
apply(cases u) apply (metis EmptyLeftSync Sync.sym append_eq_conv_conj order_refl)
proof (auto simp add: subset_iff split:if_splits, goal_cases)
case (1 a list lista t)
then show ?case using Cons(1)[of "a#list" "t2" "lista", simplified, OF 1(6)]
by (meson Suc_n_not_le_n contra_subsetD nat_le_linear set_drop_subset_set_drop)
next
case (2 a list lista t)
then show ?case using Cons(1)[of "a#list" "t2" "lista", simplified, OF 2(7)]
by (meson Suc_n_not_le_n contra_subsetD nat_le_linear set_drop_subset_set_drop)
qed
qed
lemma interleave_append: "s setinterleaves ((t1@t2,u), C) \<Longrightarrow> \<exists> u1 u2 s1 s2. u = u1@u2 \<and> s = s1@s2 \<and> s1 setinterleaves ((t1,u1), C) \<and> s2 setinterleaves ((t2,u2), C)"
proof(induct s arbitrary: t1 t2 u)
case Nil
then show ?case using empty_setinterleaving by fastforce
next
case (Cons a s)
then show ?case
apply(cases t1) using append_self_conv2 interleave_set apply fastforce
apply(cases u) apply(exI "[]", exI "[]") apply auto[1]
apply (metis (no_types, hide_lams) Nil_is_append_conv append_Cons)
proof (auto simp add: subset_iff split:if_splits, goal_cases)
case (1 list lista)
with Cons(1)[of "list" "t2" "lista", simplified, OF 1(5)] show ?case
proof(elim exE conjE, goal_cases)
case (1 u1 u2 s1 s2)
then show ?case
by (exI "a#u1", exI "u2", simp) (metis append_Cons)
qed
next
case (2 aa list lista)
with Cons(1)[of "aa#list" "t2" "lista", simplified, OF 2(6)] show ?case
proof(elim exE conjE, goal_cases)
case (1 u1 u2 s1 s2)
then show ?case
by (exI "a#u1", exI "u2", simp) (metis append_Cons)
qed
next
case (3 list aa lista)
with Cons(1)[of "list" "t2" "aa#lista", simplified, OF 3(6)] show ?case
proof(elim exE conjE, goal_cases)
case (1 u1 u2 s1 s2)
then show ?case
apply (exI "u1", exI "u2", simp) by (metis append_Cons synSingleHeadAdd)
qed
next
case (4 aa list lista)
with Cons(1)[of "aa#list" "t2" "lista", simplified, OF 4(6)] show ?case
proof(elim exE conjE, goal_cases)
case (1 u1 u2 s1 s2)
then show ?case
by (exI "a#u1", exI "u2", simp) (metis append_Cons)
qed
next
case (5 list aa lista)
with Cons(1)[of "list" "t2" "aa#lista", simplified, OF 5(6)] show ?case
proof(elim exE conjE, goal_cases)
case (1 u1 u2 s1 s2)
then show ?case
apply (exI "u1", exI "u2", simp) by (metis append_Cons synSingleHeadAdd)
qed
qed
qed
lemma interleave_append_sym: "s setinterleaves ((t,u1@u2), C) \<Longrightarrow> \<exists> t1 t2 s1 s2. t = t1@t2 \<and> s = s1@s2 \<and> s1 setinterleaves ((t1,u1), C) \<and> s2 setinterleaves ((t2,u2), C)"
by (metis (no_types) Sync.sym interleave_append)
lemma interleave_append_tail: "s setinterleaves ((t1,u), C) \<Longrightarrow> (set t2) \<inter> C = {} \<Longrightarrow> (s@t2) setinterleaves ((t1@t2,u), C)"
proof(induct s arbitrary: t1 t2 u)
case Nil
then show ?case by (metis Set.set_insert Sync.sym append_Nil disjoint_insert(2) emptyLeftSelf empty_setinterleaving)
next
case (Cons a s)
then show ?case
apply(cases u) using EmptyLeftSync Sync.sym apply fastforce
apply(cases t1, cases t2) apply simp apply fastforce
- proof(auto, goal_cases)
+ using [[simproc del: defined_all]] proof(auto, goal_cases)
case (1 list lista)
with 1(1)[OF 1(7) 1(2)] show ?case by simp
next
case (2 list aa lista)
with 2(1)[OF 2(8) 2(2)] show ?case by simp
next
case (3 list aa lista)
with 3(1)[OF 3(9) 3(2)] show ?case by simp
next
case (4 list aa lista)
with 4(1)[OF 4(9) 4(2)] show ?case by simp
qed
qed
lemma interleave_append_tail_sym: "s setinterleaves ((t,u1), C) \<Longrightarrow> (set u2) \<inter> C = {} \<Longrightarrow> (s@u2) setinterleaves ((t,u1@u2), C)"
by (metis (no_types) Sync.sym interleave_append_tail)
lemma interleave_assoc_1: "tu setinterleaves ((t, u), A) \<Longrightarrow> tuv setinterleaves ((tu, v), A) \<Longrightarrow> \<exists>uv. uv setinterleaves ((u, v), A) \<and> tuv setinterleaves ((t, uv), A)"
proof(induct tuv arbitrary: t tu u v)
case Nil
then show ?case using EmptyLeftSync empty_setinterleaving by blast
next
case Cons_tuv:(Cons a tuv)
then show ?case
proof(cases t)
case Nil_t:Nil
with Cons_tuv(2) have *:"tu = u" using EmptyLeftSync by blast
show ?thesis
proof(cases u)
case Nil_u:Nil
with Nil_t Cons_tuv show ?thesis using EmptyLeftSync by fastforce
next
case Cons_u:(Cons ua ulist)
with Nil_t Cons_tuv(2) have "ua \<notin> A" by force
show ?thesis
proof(cases v)
case Nil_v:Nil
with * Nil_t Cons_tuv(2,3) Cons_u show ?thesis using Sync.sym by blast
next
case Cons_v:(Cons va vlist)
with * Nil_t Cons_tuv(2,3) Cons_u show ?thesis apply(exI "a#tuv", auto split:if_splits)
using Cons_tuv.hyps Cons_tuv.prems(1) emptyLeftProperty by blast+
qed
qed
next
case Cons_t:(Cons ta tlist)
show ?thesis
proof(cases u)
case Nil_u:Nil
with Cons_t Cons_tuv show ?thesis
by (metis Sync.sym emptyLeftProperty emptyLeftSelf empty_set equals0D ftf_syn21)
next
case Cons_u:(Cons ua ulist)
show ?thesis
proof(cases v)
case Nil_v:Nil
with Cons_tuv(3) have "a # tuv = tu" by (simp add: Nil_v EmptyLeftSync Sync.sym)
with Nil_v Cons_u Cons_t Cons_tuv show ?thesis apply(exI u, auto split:if_splits)
apply (metis Cons_t Nil_v Sync.sym emptyLeftNonSync list.set_intros(1))
using Cons_tuv(1)[of tuv tlist u] apply(metis (no_types, lifting) Sync.sym emptyLeftNonSync emptyLeftSelf list.sel(3) syncTlEmpty)
by (metis Cons_t Sync.sym emptyLeftProperty) fastforce+
next
case Cons_v:(Cons va vlist)
with Cons_t Cons_u Cons_tuv(2,3) show ?thesis
proof(auto split:if_splits, goal_cases)
case (1 tulist)
from Cons_tuv(1)[OF 1(5) 1(10)] obtain uvlist
where "uvlist setinterleaves ((ulist, vlist), A) \<and> tuv setinterleaves ((tlist, uvlist), A)" by blast
with 1 show ?case by(exI "va#uvlist", simp)
next
case (2 u)
then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym synSingleHeadAdd)
next
case (3 u)
then show ?case by (metis Cons_tuv.hyps Sync.sym synSingleHeadAdd)
next
case (4 u)
then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym synSingleHeadAdd)
next
case (5 u)
then show ?case by (metis Cons_tuv.hyps Sync.sym synSingleHeadAdd)
next
case (6 u)
then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym synSingleHeadAdd)
next
case (7 u)
then show ?case by (metis Cons_tuv.hyps Sync.sym synSingleHeadAdd)
next
case (8 tulist)
from Cons_tuv(1)[OF 8(5) 8(9)] obtain uvlist
where "uvlist setinterleaves ((va#ulist, va#vlist), A) \<and> tuv setinterleaves ((tlist, uvlist), A)" by blast
with 8 show ?case using synSingleHeadAdd by (exI "uvlist", simp) blast
next
case (9 u)
then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym synSingleHeadAdd)
next
case (10 tulist)
from Cons_tuv(1)[OF 10(6) 10(10)] obtain uvlist
where "uvlist setinterleaves ((ua#ulist, va#vlist), A) \<and> tuv setinterleaves ((tlist, uvlist), A)" by blast
with 10 show ?case using synSingleHeadAdd by (exI "uvlist", simp) blast
next
case (11 u)
then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym synSingleHeadAdd)
next
case (12 u)
then show ?case using Cons_tuv.hyps by fastforce
next
case (13 u)
then show ?case by (metis Cons_tuv.hyps Sync.sym synSingleHeadAdd)
next
case (14 u)
then show ?case by (metis Cons_tuv.hyps Sync.sym synSingleHeadAdd)
next
case (15 u)
then show ?case by (metis Cons_tuv.hyps Sync.sym synSingleHeadAdd)
next
case (16 u)
then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym synSingleHeadAdd)
next
case (17 u)
then show ?case by (metis Cons_tuv.hyps Sync.sym synSingleHeadAdd)
next
case (18 u)
then show ?case using Cons_tuv.hyps by fastforce
next
case (19 u)
then show ?case using Cons_tuv.hyps by fastforce
next
case (20 u)
then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym synSingleHeadAdd)
next
case (21 u)
then show ?case using Cons_tuv.hyps by fastforce
qed
qed
qed
qed
qed
lemma interleave_assoc_2:
assumes *:"uv setinterleaves ((u, v), A)" and
**:"tuv setinterleaves ((t, uv), A)"
shows "\<exists>tu. tu setinterleaves ((t, u), A) \<and> tuv setinterleaves ((tu, v), A)"
using "*" "**" Sync.sym interleave_assoc_1 by blast
subsection\<open> Laws \<close>
lemma sync_commute: "(P \<lbrakk> A \<rbrakk> Q) = (Q \<lbrakk> A \<rbrakk> P)"
by (simp add: Process_eq_spec mono_D_syn mono_F_syn)
lemma mono_sync_FD_oneside[simp]: "P \<le> P' \<Longrightarrow> (P \<lbrakk> A \<rbrakk> Q) \<le> (P' \<lbrakk> A \<rbrakk> Q)"
apply(auto simp:le_ref_def F_sync D_sync)
using F_subset_imp_T_subset front_tickFree_Nil by blast+
lemma mono_sync_FD[simp]: "\<lbrakk>P \<le> P'; Q \<le> Q'\<rbrakk> \<Longrightarrow> (P \<lbrakk> A \<rbrakk> Q) \<le> (P' \<lbrakk> A \<rbrakk> Q')"
using mono_sync_FD_oneside[of P P' A Q] mono_sync_FD_oneside[of Q Q' A P]
by (simp add: order_trans sync_commute)
lemma mono_sync_ref: "\<lbrakk>P \<sqsubseteq> P'; Q \<sqsubseteq> Q'\<rbrakk> \<Longrightarrow> (P \<lbrakk> A \<rbrakk> Q) \<sqsubseteq> (P' \<lbrakk> A \<rbrakk> Q')"
using mono_Sync_sym[of Q Q' P A] mono_Sync[of P P' A Q'] below_trans by blast
lemma par_Int_bot: "(P \<lbrakk> A \<rbrakk> \<bottom>) = \<bottom>"
apply(auto simp add:Process_eq_spec, simp_all add:D_sync F_sync F_UU D_UU)
apply (metis D_imp_front_tickFree append_Nil2 front_tickFree_append ftf_syn is_processT is_processT2_TR)
apply (metis Nil_elem_T Sync.si_empty1 append_Nil front_tickFree_Nil insertI1 non_tickFree_implies_nonMt)
apply (metis D_imp_front_tickFree append_Nil2 front_tickFree_append ftf_syn is_processT2_TR)
by (metis Nil_elem_T Sync.si_empty1 append_Nil front_tickFree_Nil insertI1 non_tickFree_implies_nonMt)
lemma par_Int_skip: "(SKIP \<lbrakk> A \<rbrakk> SKIP) = SKIP"
apply(auto simp add:Process_eq_spec, simp_all add:D_sync F_sync F_SKIP D_SKIP)
apply(elim exE conjE disjE, auto)
apply (metis Sync.si_empty1 inf.idem insertI1 sup_commute sup_inf_absorb)
apply(exI "[tick]", exI "[tick]", simp)
by blast
lemma par_Int_skip_stop: "(SKIP \<lbrakk> A \<rbrakk> STOP) = STOP"
proof(auto simp add:Process_eq_spec, simp_all add:D_sync F_sync F_SKIP D_SKIP F_STOP D_STOP, goal_cases)
case (1 a b)
then show ?case by (elim exE conjE disjE, auto)
next
case (2 a b)
then show ?case
apply(rule_tac x="[]" in exI, simp, rule_tac x="b-{tick}" in exI, simp, rule_tac x="b" in exI)
by blast
qed
subsection\<open> Multi-Operators laws \<close>
lemma mprefix_Par_Int_distr: "
\<lbrakk>B \<inter> C = {}; A \<subseteq> C\<rbrakk> \<Longrightarrow> (\<box> x \<in> A \<rightarrow> P x \<lbrakk> C \<rbrakk> \<box> x \<in> B \<rightarrow> Q x) = \<box> x \<in> B \<rightarrow> (\<box> x \<in> A \<rightarrow> P x \<lbrakk> C \<rbrakk> Q x)"
apply(auto simp add: Process_eq_spec, simp_all add: F_Mprefix T_Mprefix D_Mprefix D_sync F_sync)
apply(elim exE disjE conjE)
apply auto[1]
proof(goal_cases)
case (1 a b t u X Y aa)
then show ?case
by (metis emptyLeftProperty syncTlEmpty)
next
case (2 a b t u X Y aa)
then show ?case
by (metis (no_types, lifting) Sync.sym empty_iff image_iff list.exhaust_sel setinterleaving.simps(2) subsetCE subset_insertI)
next
case (3 a b t u X Y aa ab)
from 3(1,2,7,8) have a1: "hd t \<in>insert tick(ev ` C)\<and>hd u \<notin>insert tick(ev ` C)"
using insertI2 by auto
then show ?case
by (metis "3"(10) "3"(11) "3"(12) "3"(3) "3"(4) "3"(5) "3"(7) "3"(8) "3"(9) empty_setinterleaving syncHd_Tl)
next
case (4 a b t u r v aa)
then show ?case
by (metis (no_types, lifting) Sync.si_empty3 equals0D imageE list.exhaust_sel rev_image_eqI subset_eq subset_insertI)
next
case (5 a b t u r v aa ab)
from 5(5,6,7) have a1: " r \<noteq> []\<and>a\<noteq> []\<and>hd t \<in>insert tick (ev ` C)\<and>hd u \<notin>insert tick (ev ` C)"
using "5"(1) "5"(11) "5"(13) "5"(2) "5"(8) disjoint_iff_not_equal empty_setinterleaving insert_iff set_mp by fastforce
from 5(5,6) a1 have a3: "hd u=hd a\<and>tl r setinterleaves ((t, tl u), insert tick (ev ` C))"
by (metis hd_append2 syncHd_Tl)
then show ?case
using 5(3,4,5,7,8,9,10,11,13,14) by (metis a1 a3 image_eqI tickFree_tl tl_append2)
next
case (6 a b t u r v aa)
from 6(5,6,9) have a1:"hd a=hd t\<and>tl r setinterleaves ((tl t, u), insert tick (ev ` C))"
using "6"(7) Sync.sym emptyLeftProperty syncTlEmpty by fastforce
then show ?case
using 6(3,4,5,6,7,8,9,10,11) by (metis EmptyLeftSync Nil_is_append_conv Sync.sym a1 tickFree_tl tl_append2)
next
case (7 a b t u r v aa ab)
from 7(2,1,8,11,13) have a2: "hd t\<notin> insert tick(ev ` C)\<and>hd u\<in>insert tick(ev ` C)"
using image_eqI by auto
from 7(5,6) a2 have a3: "hd a=hd t\<and>hd r=hd t\<and>a \<noteq> [] \<and>tl a =tl r @ v \<and> tl r setinterleaves ((tl t, u), insert tick (ev ` C))"
by (metis Nil_is_append_conv Sync.sym empty_setinterleaving hd_append2 syncHd_Tl tl_append2)
then show ?case
using 7(3,4,8,9,10,11,13,14) by (metis a3 list.sel(2) tickFree_tl)
next
case (8 a b t u r v aa)
then show ?case
by (metis (no_types, lifting) Sync.sym empty_iff image_iff insertI2 list.exhaust_sel setinterleaving.simps(2) subsetCE)
next
case (9 a b t u r v aa ab)
from 9(2,8,1,11,13) have a2: "hd t\<in> insert tick(ev ` C)\<and>hd u\<notin> insert tick(ev ` C)"
by auto
then show ?case
by (metis "9"(10) "9"(11) "9"(13) "9"(14) "9"(3) "9"(4) "9"(5) "9"(6) "9"(7) "9"(8) "9"(9) append_Nil2 empty_setinterleaving image_eqI syncHd_Tl)
next
case (10 a b t u r v aa)
then show ?case
by (metis (no_types, lifting) EmptyLeftSync Sync.sym append_Nil2 syncTlEmpty)
next
case (11 a b t u r v aa ab)
from 11(2,8,1,11,13) have a2: "hd u\<in> insert tick(ev ` C)\<and>hd t\<notin> insert tick(ev ` C)"
by auto
then show ?case
by (metis "11"(10) "11"(11) "11"(13) "11"(14) "11"(3) "11"(4) "11"(5) "11"(6) "11"(8) "11"(9) Sync.sym append_Nil2 empty_setinterleaving syncHd_Tl)
next
case (12 a b)
then show ?case
apply(elim exE disjE conjE)
proof(goal_cases)
case 1
obtain X where aa1: "(X::'a event set)=b-ev ` A" by auto
from aa1 have aa2: "X \<inter> ev ` A = {}\<and> b \<inter> ev ` B = {}"
by (simp add: "1"(4) Int_commute)
from 1(1,2,3,4) aa1 have a12: "b = (X \<union> b) \<inter> insert tick (ev ` C) \<union> X \<inter> b"
by auto
from 1(3) have a21: "a setinterleaves (([], []), insert tick (ev ` C))"
by auto
then show ?case
using 1(1,2,3,4) a12 a21 by (metis aa2)
next
case (2 aa t u X Y)
from 2(1,3,4,7,9) have aa1: "a setinterleaves ((t, hd a#u), insert tick (ev ` C))"
by (metis (no_types, lifting) disjoint_iff_not_equal event.distinct(1) event.inject imageE insertE syncHdAddEmpty)
then show ?case
using 2(4,5,6,7,8,10) by (metis list.distinct(1) list.sel(1) list.sel(3))
next
case (3 aa t u X Y ab)
from 3(1,2,4) have aa0: "hd a\<notin> insert tick (ev ` C)\<and>hd t\<in> insert tick (ev ` C)"
using "3"(10) by auto
then show ?case
using 3(3,4,5,6,7,8,9,10,11,12) by (metis list.distinct(1) list.sel(1) list.sel(3) syncHdAdd)
next
case (4 aa t u r v ab)
from 4(1,2,4,7,11) have aa1:"tickFree (hd a#r)\<and>hd a\<notin>insert tick (ev ` C)\<and>hd t\<in>insert tick (ev ` C)"
by auto
from 4(4,5,9,10,12) aa1 have aa2: "(hd a#r) setinterleaves ((t,(hd a#u)), insert tick (ev ` C))\<and>aa \<in> B \<and>hd a# u \<noteq> [] \<and> hd(hd a# u) = ev aa \<and> u \<in> T (Q aa)"
by (metis (no_types, lifting) event.inject imageE list.distinct(1) list.sel(1) syncHdAdd1)
then show ?case
using 4(3,6,8,10,11,13,14) by (metis (no_types, lifting) aa1 aa2 append_Cons list.exhaust_sel list.sel(3))
next
case (5 aa t u r v)
from 5(1,4,7,9,11) have aa2:"(hd a#r) setinterleaves (((hd a#t),u), insert tick (ev ` C))\<and>tickFree (hd a#r)"
using "5"(11) "5"(7) Sync.sym by auto
then show ?case
using 5(3,4,5,6,8,10,11) by (metis append_Cons list.exhaust_sel list.sel(1) list.sel(3) list.simps(3))
next
case (6 aa t u r v ab)
from 6(1,4,7) have aa1:"tickFree (hd a#r)\<and>hd a\<notin>insert tick (ev ` C)\<and>(hd a#r) setinterleaves (((hd a#t),u), insert tick (ev ` C))"
by (metis (no_types, lifting) "6"(9) disjoint_iff_not_equal event.inject event.simps(3) imageE insertE synSingleHeadAdd tickFree_Cons)
then show ?case
using 6(3,4,5,6,8,10,11,13,14) by (metis aa1 append_Cons list.distinct(1) list.exhaust_sel list.sel(1) list.sel(3))
next
case (7 aa t u r v ab)
from 7(1,4,5,12) have aa1: "aa \<in> B \<and>hd a# u \<noteq> [] \<and> hd(hd a# u) = ev aa\<and>u \<in> T (Q aa)\<and>hd a \<notin>insert tick (ev ` C)"
using insertE by auto
from 7(1,2,4,9,10,11) aa1 have aa2: "(hd a#r) setinterleaves ((t,(hd a#u)), insert tick (ev ` C))"
by (metis (no_types, lifting) imageE rev_image_eqI subset_iff subset_insertI syncHdAdd1)
then show ?case
using 7(3,6,7,8,10,11,13,14) by (metis aa1 append.right_neutral list.exhaust_sel list.sel(3))
next
case (8 aa t u r v)
from 8(1,4,9) have aa1:"(hd a#r) setinterleaves (((hd a#t),u), insert tick (ev ` C))"
using "8"(11) by auto
then show ?case
using 8(3,4,5,6,7,8,10,11) by (metis EmptyLeftSync Sync.sym append.right_neutral list.exhaust_sel list.sel(3))
next
case (9 aa t u r v ab)
from 9(1,2,4,5,10,11,13) have aa2: "aa \<in> B \<and>hd a# t \<noteq> [] \<and> hd(hd a# t) = ev aa \<and> t \<in> D (Q aa)\<and>hd a\<notin> insert tick (ev ` C)\<and>hd u\<in> insert tick (ev ` C)"
by auto
then show ?case
using 9(3,4,5,6,7,8,9,11,13,14) by (metis (no_types, lifting) Sync.sym append.right_neutral list.sel(3) syncHdAdd)
qed
next
case (13 x)
then show ?case
apply(elim exE disjE conjE)
apply (metis (no_types, lifting) Sync.sym empty_iff image_iff insertI2 list.exhaust_sel setinterleaving.simps(2) subsetCE)
proof(goal_cases)
case (1 t u r v a aa)
from 1(1,2,8,11,13) have aa0: "hd t\<in>insert tick (ev ` C)\<and>hd u\<notin>insert tick (ev ` C)"
by auto
from 1(5,6,7) have aa1: "x \<noteq> []\<and>hd r=hd u\<and>(tl r) setinterleaves ((t,tl u), insert tick (ev ` C))"
by (metis Nil_is_append_conv aa0 empty_setinterleaving syncHd_Tl)
then show ?case
using 1(3,4,5,7,8,9,10,11,13,14) by (metis (no_types, lifting) empty_setinterleaving hd_append2 image_eqI list.sel(2) tickFree_tl tl_append2)
next
case (2 t u r v a)
from 2(5,6,7,9) have aa1: "x \<noteq> []\<and>hd r=hd t"
by (metis Nil_is_append_conv Sync.sym emptyLeftProperty)
then show ?case
using 2(3,4,5,6,7,8,9,10,11) by (metis (no_types, lifting) Sync.sym emptyLeftProperty hd_append2 syncTlEmpty tickFree_tl tl_append2)
next
case (3 t u r v a aa)
from 3(1,2,8,11,13) have aa0: "hd t\<notin>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by auto
from 3(5,6,7) have aa1: "x \<noteq> []\<and>hd r=hd t\<and>hd u\<in>insert tick (ev ` C)\<and>(tl r) setinterleaves ((tl t,u), insert tick (ev ` C))"
by (metis Nil_is_append_conv Sync.sym aa0 empty_setinterleaving syncHd_Tl)
then show ?case
using 3(3,4,5,6,7,8,9,10,11,13,14) by (metis (no_types, lifting) empty_setinterleaving hd_append2 tickFree_tl tl_append2)
next
case (4 t u r v a)
then show ?case
by (metis (no_types, lifting) Sync.sym equals0D image_iff list.exhaust_sel setinterleaving.simps(2) subsetCE subset_insertI)
next
case (5 t u r v a aa)
from 5(1,2,8,11,13) have aa0: "hd t\<in>insert tick (ev ` C)\<and>hd u\<notin>insert tick (ev ` C)"
by auto
from 5(5,6,7) have aa1: "x \<noteq> []\<and>(tl r) setinterleaves ((t,tl u), insert tick (ev ` C))"
using aa0 empty_setinterleaving syncHd_Tl[THEN conjunct2] by fastforce
then show ?case
using 5(3,4,5,6,7,8,9,10,11,13,14) by (metis (no_types, lifting) aa0 image_iff self_append_conv syncHd_Tl)
next
case (6 t u r v a)
then show ?case
by (metis Sync.sym append_Nil2 emptyLeftProperty syncTlEmpty)
next
case (7 t u r v a aa)
from 7(1,2,8,11,13) have aa0: "hd t\<notin>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by auto
then show ?case
by (metis "7"(10) "7"(11) "7"(13) "7"(14) "7"(3) "7"(4) "7"(5) "7"(6) "7"(8) "7"(9) Sync.sym append_Nil2 empty_setinterleaving syncHd_Tl)
qed
next
case (14 x)
then show ?case
apply(elim exE disjE conjE)
proof(goal_cases)
case (1 a t u r v aa)
from 1(1,4,7) have aa1: "tickFree (hd x#r)\<and>hd x\<notin>insert tick (ev ` C)"
by auto
from 1(1,2,4,5,9,10,11,12) aa1 have aa2: "(hd x#r) setinterleaves ((t, hd x#u), insert tick (ev ` C))\<and>hd(hd x#u) =ev a\<and>u\<in>T(Q a)\<and>a\<in>B\<and>hd x#u\<noteq>[]\<and>tickFree (hd x#r)\<and>hd x\<notin>insert tick (ev ` C)"
by (metis (no_types, lifting) event.inject imageE image_eqI insertI2 list.sel(1) list.simps(3) subsetCE syncHdAdd1)
then show ?case
using 1(3,6,8,10,11,13,14) by (metis (no_types, lifting) aa1 append_Cons list.exhaust_sel list.sel(3))
next
case (2 a t u r v)
from 2(1,4,7,9,11) have aa1: "tickFree (hd x#r)\<and>(hd x#r) setinterleaves ((hd x#t, u), insert tick (ev ` C))"
using Sync.sym by auto
then show ?case
using 2(3,4,5,6,8,10,11) by (metis append_Cons list.distinct(1) list.exhaust_sel list.sel(1) list.sel(3))
next
case (3 a t u r v aa)
from 3(1,4,7) have aa1: "tickFree (hd x#r)\<and>hd x\<notin> insert tick (ev ` C)"
by auto
from 3(1,2,4,9,11,12,13) aa1 have aa2: "(hd x#r) setinterleaves ((hd x#t, u), insert tick (ev ` C))\<and>tickFree (hd x#r)\<and>hd x\<notin> insert tick (ev ` C)"
by (metis (no_types, lifting) Sync.sym image_subset_iff set_rev_mp subset_insertI syncHdAdd1)
then show ?case
using 3(3,4,5,6,8,10,11,13,14) by (metis aa1 append_Cons list.distinct(1) list.exhaust_sel list.sel(1) list.sel(3))
next
case (4 a t u r v aa)
from 4(1,4) have aa0: "hd x\<notin>insert tick (ev ` C)"
by auto
from 4(1,2,4,5,9,10,11,12) aa0 have aa1: "(hd x#r) setinterleaves ((t, hd x#u), insert tick (ev ` C))\<and>hd(hd x#u) =ev a\<and>u\<in>T(Q a)\<and>a\<in>B\<and>hd x#u\<noteq>[]"
by (metis (no_types, lifting) event.inject imageE image_eqI insertI2 list.sel(1) list.simps(3) subsetCE syncHdAdd1)
then show ?case
using 4(3,6,7,8,10,11,13,14) by (metis list.exhaust_sel list.sel(3) self_append_conv)
next
case (5 a t u r v)
from 5(9) have aa1: "(hd x#r) setinterleaves ((hd x#t, u), insert tick (ev ` C))"
using "14"(1) "5"(11) "5"(4) by auto
then show ?case
using 5(3,4,5,6,7,8,10,11) by (metis append.right_neutral list.distinct(1) list.exhaust_sel list.sel(1) list.sel(3))
next
case (6 a t u r v aa)
from 6(1,4) have aa0: "hd x\<notin>insert tick (ev ` C)\<and>(hd x#r) setinterleaves ((hd x#t, u), insert tick (ev ` C))"
by (metis (no_types, lifting) "6"(9) disjoint_iff_not_equal event.inject event.simps(3) imageE insertE synSingleHeadAdd)
then show ?case
using 6(3,4,5,6,7,8,10,11,13,14) by (metis append_Nil2 list.distinct(1) list.exhaust_sel list.sel(1) list.sel(3))
qed
qed
lemma mprefix_Par_Int:
"\<lbrakk>B \<inter> C = {}; A \<inter> C = {}\<rbrakk> \<Longrightarrow>
(Mprefix A P \<lbrakk> C \<rbrakk> Mprefix B Q) = (\<box> x \<in> A \<rightarrow> (P x \<lbrakk> C \<rbrakk> Mprefix B Q) \<box> (\<box> y \<in> B \<rightarrow> (Mprefix A P \<lbrakk> C \<rbrakk> Q y)))"
apply(auto simp add:Process_eq_spec, simp_all add:D_sync F_sync F_Mprefix T_Mprefix D_Mprefix F_det D_det)
apply(elim exE disjE conjE)
proof(goal_cases)
case (1 a b t u X Y)
then show ?case using IntE Un_iff by auto
next
case (2 a b t u X Y aa)
then show ?case
by (metis emptyLeftProperty syncTlEmpty)
next
case (3 a b t u X Y aa)
then show ?case
by (metis (no_types, lifting) Sync.sym emptyLeftProperty syncTlEmpty)
next
case (4 a b t u X Y aa ab)
from 4(1,2,7,8) have aa: "hd t\<notin>insert tick (ev ` C)\<and>hd u\<notin>insert tick (ev ` C)" by auto
from 4(3,6,7,8) obtain t1 t2 u1 u2 where aa0: "t=t1#t2\<and>u=u1#u2" using list.exhaust_sel by blast
from 4(3,4,6) aa aa0 have aa1: "(tl a) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd a=hd t\<or>(tl a) setinterleaves ((t,tl u), insert tick (ev ` C))\<and>hd a=hd u" by auto
then show ?case
by (metis "4"(10) "4"(11) "4"(12) "4"(3) "4"(4) "4"(5) "4"(6) "4"(7) "4"(8) "4"(9) empty_setinterleaving)
next
case (5 a b t u r v aa)
then show ?case
apply(rule_tac disjI2, rule_tac disjI1, intro conjI, simp)
using empty_setinterleaving apply blast
apply(rule_tac disjI2, rule_tac disjI1, intro conjI, simp)
using empty_setinterleaving apply blast
using Sync.sym emptyLeftProperty apply fastforce
by (metis (no_types, lifting) Sync.sym emptyLeftProperty hd_append2 syncTlEmpty tickFree_tl tl_append2)
next
case (6 a b t u r v aa ab)
from 6(1,2,7,8,11,13) have aa: "hd t\<notin>insert tick (ev ` C)\<and>hd u\<notin>insert tick (ev ` C)"
by auto
from 6(7,8,12,13) obtain t1 t2 u1 u2 where aa0: "t=t1#t2\<and>u=u1#u2"
using list.exhaust_sel by blast
from 6(6) aa aa0 have aa1: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<or>(tl r) setinterleaves ((t,tl u), insert tick (ev ` C))\<and>hd r=hd u"
by auto
from 6(4,5,6,7,8,9,10,12,13,14) aa0 have aa2: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<Longrightarrow>tl t \<in> D (P aa)\<and>tl u \<in> T (Q ab)\<and> a \<noteq> [] \<and>hd a \<in> ev ` A\<and>hd a= ev aa\<and>tl a =(tl r) @ v\<and>tickFree (tl r)\<and> hd u = ev ab"
by (metis Nil_is_append_conv empty_setinterleaving hd_append2 tickFree_tl tl_append2)
from 6(4,5,7,8,10,11,12,13,14) aa0 have aa3: "(tl r) setinterleaves ((t, tl u), insert tick (ev ` C))\<and>hd r=hd u\<Longrightarrow> tl t \<in> D (P aa)\<and>tl u \<in> T (Q ab)\<and> a \<noteq> [] \<and>hd a \<in> ev ` B\<and>tl a =(tl r) @ v\<and>tickFree (tl r)\<and>hd a=ev ab\<and> hd u = ev ab"
by (metis (no_types, lifting) Nil_is_append_conv empty_setinterleaving hd_append2 image_eqI list.sel(2) tickFree_tl tl_append2)
then show ?case
by (metis "6"(11) "6"(12) "6"(3) "6"(7) "6"(8) "6"(9) aa1 aa2)
next
case (7 a b t u r v aa)
then show ?case
apply(rule_tac disjI2, rule_tac disjI1, intro conjI, simp)
using empty_setinterleaving apply blast
apply(rule_tac disjI2, rule_tac disjI2,rule_tac disjI2, intro conjI,simp)
using empty_setinterleaving apply blast
using Sync.sym emptyLeftProperty apply fastforce
by (metis (no_types, lifting) Sync.sym emptyLeftProperty hd_append2 syncTlEmpty tickFree_tl tl_append2)
next
case (8 a b t u r v aa ab)
from 8(1,2,7,8,11,13) have aa: "hd t\<notin>insert tick (ev ` C)\<and>hd u\<notin>insert tick (ev ` C) "
by auto
from 8(7,8,12,13) obtain t1 t2 u1 u2 where aa0: "t=t1#t2\<and>u=u1#u2"
using list.exhaust_sel by blast
from 8(6) aa aa0 have aa1: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<or>(tl r) setinterleaves ((t,tl u), insert tick (ev ` C))\<and>hd r=hd u"
by auto
from 8(4,5,6,7,8,9,10,12,13,14) aa0 have aa2: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<Longrightarrow>tl t \<in> D (Q aa)\<and>tl u \<in> T (P ab)\<and> a \<noteq> [] \<and>hd a \<in> ev ` B\<and>hd a= ev aa\<and>tl a =(tl r) @ v\<and>tickFree (tl r)\<and> hd u = ev ab"
by (metis Nil_is_append_conv empty_setinterleaving hd_append2 tickFree_tl tl_append2)
from 8(4,5,7,8,10,11,12,13,14) aa0 have aa3: "(tl r) setinterleaves ((t, tl u), insert tick (ev ` C))\<and>hd r=hd u\<Longrightarrow> tl t \<in> D (Q aa)\<and>tl u \<in> T (P ab)\<and> a \<noteq> [] \<and>hd a \<in> ev ` A\<and>tl a =(tl r) @ v\<and>tickFree (tl r)\<and>hd a=ev ab\<and> hd u = ev ab"
by (metis (no_types, lifting) Nil_is_append_conv empty_setinterleaving hd_append2 image_eqI list.sel(2) tickFree_tl tl_append2)
then show ?case
by (metis "8"(11) "8"(12) "8"(3) "8"(7) "8"(8) "8"(9) aa1 aa2)
next
case (9 a b t u r v aa)
then show ?case
by (metis (no_types, lifting) Sync.sym append_Nil2 emptyLeftProperty syncTlEmpty)
next
case (10 a b t u r v aa ab)
from 10(1,2,7,8,11,13) have aa: "hd t\<notin>insert tick (ev ` C)\<and>hd u\<notin>insert tick (ev ` C)"
by auto
from 10(7,8,12,13) obtain t1 t2 u1 u2 where aa0: "t=t1#t2\<and>u=u1#u2"
using list.exhaust_sel by blast
from 10(6) aa aa0 have aa1: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<or>(tl r) setinterleaves ((t,tl u), insert tick (ev ` C))\<and>hd r=hd u"
by auto
then show ?case
by (metis (no_types, lifting) "10"(10) "10"(11) "10"(13) "10"(14) "10"(3) "10"(4) "10"(5) "10"(6) "10"(7) "10"(8) "10"(9) append_self_conv empty_setinterleaving image_eqI)
next
case (11 a b t u r v aa)
then show ?case
by (metis (no_types, lifting) Sync.sym append_Nil2 emptyLeftProperty syncTlEmpty)
next
case (12 a b t u r v aa ab)
from 12(1,2,7,8,11,13) have aa: "hd t\<notin>insert tick (ev ` C)\<and>hd u\<notin>insert tick (ev ` C)"
by auto
from 12(7,8,12,13) obtain t1 t2 u1 u2 where aa0: "t=t1#t2\<and>u=u1#u2"
using list.exhaust_sel by blast
from 12(6) aa aa0 have aa1: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<or>(tl r) setinterleaves ((t,tl u), insert tick (ev ` C))\<and>hd r=hd u"
by auto
then show ?case
by (metis (no_types, lifting) "12"(10) "12"(11) "12"(13) "12"(14) "12"(3) "12"(4) "12"(5) "12"(6) "12"(7) "12"(8) "12"(9) append_Nil2 empty_setinterleaving image_eqI)
next
case (13 a b)
then show ?case
apply(elim disjE conjE, simp_all)
proof(goal_cases)
case 1
then show ?case
by(rule_tac disjI1, rule_tac x="[]" in exI, rule_tac x="[]" in exI, rule_tac x="b- ev ` A" in exI, simp, intro conjI, auto)
next
case 2
then show ?case
apply(elim exE disjE conjE)
proof(goal_cases)
case (1 aa t u X Y)
then show ?case
apply(rule_tac disjI1, rule_tac x="hd a#t" in exI, rule_tac x="u" in exI, rule_tac x="X" in exI, intro conjI, simp_all)
apply (blast, intro conjI)
apply auto[1]
apply auto[1]
by (metis list.exhaust_sel)
next
case (2 aa t u X Y aaa)
from 2(2,4) have aa0: "hd a\<notin>insert tick (ev ` C)"
by auto
from 2(1, 2, 7) aa0 have aa1: "a setinterleaves ((hd a#t, u), insert tick (ev ` C))"
by (metis list.exhaust_sel synSingleHeadAdd)
then show ?case
using 2(2,5,6,8,9,10,11,12) by (metis list.distinct(1) list.sel(1) list.sel(3))
next
case (3 aa t u r v)
then show ?case
apply(rule_tac disjI2, rule_tac x="hd a#t" in exI, rule_tac x="u" in exI,rule_tac x="hd a#r" in exI,rule_tac x="v" in exI, simp)
by (metis disjoint_iff_not_equal event.distinct(1) event.inject imageE list.exhaust_sel)
next
case (4 aa t u r v aaa)
then show ?case
apply(rule_tac disjI2, rule_tac x="hd a#t" in exI, rule_tac x="u" in exI,rule_tac x="hd a#r" in exI,rule_tac x="v" in exI, intro conjI, auto)
by (metis list.exhaust_sel, simp add: disjoint_iff_not_equal image_iff synSingleHeadAdd)
next
case (5 aa t u r v aaa)
from 5(2,4,7) have aa0: "hd a\<notin>insert tick (ev ` C)\<and>tickFree (hd a#r)"
by auto
from 5(1,2, 7,8,9) aa0 have aa1: "(hd a#r) setinterleaves ((t, hd a#u), insert tick (ev ` C))\<and> a = (hd a#r) @ v"
by (metis Sync.sym append_Cons list.exhaust_sel synSingleHeadAdd)
then show ?case
using 5(2,5,6,10,11,12,13,14) aa0 event.inject list.sel(1) by fastforce
next
case (6 aa t u r v)
then show ?case
apply(rule_tac disjI2, rule_tac x="hd a#t" in exI, rule_tac x="u" in exI,rule_tac x="hd a#r" in exI,rule_tac x="v" in exI, simp)
by (metis disjoint_iff_not_equal event.distinct(1) event.inject imageE list.exhaust_sel)
next
case (7 aa t u r v aaa)
then show ?case
apply(rule_tac disjI2, rule_tac x="hd a#t" in exI, rule_tac x="u" in exI,rule_tac x="hd a#r" in exI,rule_tac x="v" in exI, intro conjI, auto)
by (metis list.exhaust_sel, simp add: disjoint_iff_not_equal image_iff synSingleHeadAdd)
next
case (8 aa t u r v aaa)
from 8(2,4,7) have aa0: "hd a\<notin>insert tick (ev ` C) "
by auto
from 8(1,2, 7,8,9) aa0 have aa1: "(hd a#r) setinterleaves ((t, hd a#u), insert tick (ev ` C))\<and> a = (hd a#r) @ v"
by (metis Sync.sym append_Cons list.exhaust_sel synSingleHeadAdd)
then show ?case
using 8(2,5,7,10,11,12,13,14) by fastforce
qed
next
case 3
then show ?case
apply(elim exE disjE conjE)
proof(goal_cases)
case (1 aa t u X Y)
then show ?case
apply(rule_tac disjI1, rule_tac x="t" in exI, rule_tac x="hd a#u" in exI, rule_tac x="X" in exI, simp_all, intro conjI, simp_all)
apply auto[1]
apply auto[1]
by (metis list.exhaust_sel)
next
case (2 aa t u X Y aaa)
from 2(2,3) have aa0: "hd a\<notin>insert tick (ev ` C)"
by auto
from 2(1, 2, 8) aa0 have aa1: "a setinterleaves ((t, hd a#u), insert tick (ev ` C))"
by (metis Sync.sym list.exhaust_sel synSingleHeadAdd)
then show ?case
using 2(2,5,6,7,9,10,11,12) by (metis list.distinct(1) list.sel(1) list.sel(3))
next
case (3 aa t u r v aaa)
from 3(2,3, 7) have aa0: "hd a\<notin>insert tick (ev ` C)\<and>tickFree (hd a#r)"
by auto
from 3(1, 2, 8, 9) aa0 have aa1: "(hd a#r) setinterleaves ((t, hd a#u), insert tick (ev ` C))\<and> a =(hd a# r) @ v"
by (metis Sync.sym append_Cons list.exhaust_sel synSingleHeadAdd)
then show ?case
using 3(2,5,6,10,11,12,13,14) aa0 by fastforce
next
case (4 aa t u r v)
then show ?case
apply(rule_tac disjI2, rule_tac x="hd a#t" in exI, rule_tac x="u" in exI,rule_tac x="hd a#r" in exI,rule_tac x="v" in exI, simp, intro conjI, simp_all)
apply auto[1]
apply auto[1]
by (metis list.exhaust_sel)
next
case (5 aa t u r v aaa)
from 5(2,3,7) have aa0: "hd a\<notin>insert tick (ev ` C)\<and>tickFree (hd a#r)"
by auto
from 5(1,2, 7,8,9) aa0 have aa1: "(hd a#r) setinterleaves ((hd a#t, u), insert tick (ev ` C))\<and> a = (hd a#r) @ v"
by (metis Sync.sym append_Cons list.exhaust_sel synSingleHeadAdd)
then show ?case
using 5(2,5,6,10,11,13,14) aa0 by fastforce
next
case (6 aa t u r v aaa)
from 6(2,3,7) have aa0: "hd a\<notin>insert tick (ev ` C)"
by auto
from 6(1,2, 7,8,9) aa0 have aa1: "(hd a#r) setinterleaves ((t, hd a#u), insert tick (ev ` C))\<and> a = (hd a#r) @ v"
by (metis Sync.sym append_Cons list.exhaust_sel synSingleHeadAdd)
then show ?case
using 6(2,5,7,10,11,12,13,14) by fastforce
next
case (7 aa t u r v)
then show ?case
apply(rule_tac disjI2, rule_tac x="hd a#t" in exI, rule_tac x="u" in exI,rule_tac x="hd a#r" in exI,rule_tac x="v" in exI, simp)
by (metis (no_types, lifting) disjoint_iff_not_equal event.distinct(1) event.inject imageE list.exhaust_sel)
next
case (8 aa t u r v aaa)
from 8(2,3,7) have aa0: "hd a\<notin>insert tick (ev ` C) "
by auto
from 8(1,2, 7,8,9) aa0 have aa1: "(hd a#r) setinterleaves ((hd a#t, u), insert tick (ev ` C))\<and> a = (hd a#r) @ v"
by (metis Sync.sym append_Cons list.exhaust_sel synSingleHeadAdd)
then show ?case
using 8(2,5,7,10,11,13,14) by fastforce
qed
qed
next
case (14 x)
then show ?case
apply(elim exE disjE conjE)
proof(goal_cases)
case (1 t u r v a)
then show ?case
apply(rule_tac disjI1, intro conjI, simp_all)
using empty_setinterleaving apply blast
using Sync.sym emptyLeftProperty apply fastforce
by (metis (no_types, lifting) Sync.sym emptyLeftProperty hd_append2 syncTlEmpty tickFree_tl tl_append2)
next
case (2 t u r v a aa)
from 2(1,2,7,8,11,13) have aa: "hd t\<notin>insert tick (ev ` C)\<and>hd u\<notin>insert tick (ev ` C)"
by auto
from 2(7,8,12,13) obtain t1 t2 u1 u2 where aa0: "t=t1#t2\<and>u=u1#u2"
using list.exhaust_sel by blast
from 2(6) aa aa0 have aa1: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<or>(tl r) setinterleaves ((t,tl u), insert tick (ev ` C))\<and>hd r=hd u"
by auto
from 2(4,5,6,7,8,9,10,12,13,14) aa0 have aa2: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<Longrightarrow>tl t \<in> D (P a)\<and>tl u \<in> T (Q aa)\<and> x \<noteq> [] \<and>hd x \<in> ev ` A\<and>hd x= ev a\<and>tl x =(tl r) @ v\<and>tickFree (tl r)\<and> hd u = ev aa"
by (metis Nil_is_append_conv empty_setinterleaving hd_append2 tickFree_tl tl_append2)
from 2(4,5,7,8,10,11,12,13,14) aa0 have aa3: "(tl r) setinterleaves ((t, tl u), insert tick (ev ` C))\<and>hd r=hd u\<Longrightarrow> tl t \<in> D (P a)\<and>tl u \<in> T (Q aa)\<and> x \<noteq> [] \<and>hd x \<in> ev ` B\<and>tl x =(tl r) @ v\<and>tickFree (tl r)\<and>hd x=ev aa\<and> hd u = ev aa"
by (metis (no_types, lifting) Nil_is_append_conv empty_setinterleaving hd_append2 image_eqI list.sel(2) tickFree_tl tl_append2)
then show ?case
using aa1 aa2 aa3 2(3,7,8,9,11,12) by (metis)
next
case (3 t u r v a)
then show ?case
apply(rule_tac disjI2, intro conjI, simp_all)
using empty_setinterleaving apply blast
using Sync.sym emptyLeftProperty apply fastforce
by (metis (no_types, lifting) Sync.sym emptyLeftProperty hd_append2 syncTlEmpty tickFree_tl tl_append2)
next
case (4 t u r v a aa)
from 4(1,2,7,8,11,13) have aa: "hd t\<notin>insert tick (ev ` C)\<and>hd u\<notin>insert tick (ev ` C)"
by auto
from 4(7,8,12,13) obtain t1 t2 u1 u2 where aa0: "t=t1#t2\<and>u=u1#u2"
using list.exhaust_sel by blast
from 4(6) aa aa0 have aa1: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<or>(tl r) setinterleaves ((t,tl u), insert tick (ev ` C))\<and>hd r=hd u"
by auto
from 4(4,5,6,7,8,9,10,12,13,14) aa0 have aa2: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<Longrightarrow>tl t \<in> D (Q a)\<and>tl u \<in> T (P aa)\<and> x \<noteq> [] \<and>hd x \<in> ev ` B\<and>hd x= ev a\<and>tl x =(tl r) @ v\<and>tickFree (tl r)\<and> hd u = ev aa"
by (metis Nil_is_append_conv empty_setinterleaving hd_append2 tickFree_tl tl_append2)
from 4(4,5,7,8,10,11,12,13,14) aa0 have aa3: "(tl r) setinterleaves ((t, tl u), insert tick (ev ` C))\<and>hd r=hd u\<Longrightarrow> tl t \<in> D (Q a)\<and>tl u \<in> T (P aa)\<and> x \<noteq> [] \<and>hd x \<in> ev ` A\<and>tl x =(tl r) @ v\<and>tickFree (tl r)\<and>hd x=ev aa\<and> hd u = ev aa"
by (metis (no_types, lifting) Nil_is_append_conv empty_setinterleaving hd_append2 image_eqI list.sel(2) tickFree_tl tl_append2)
then show ?case
by (metis "4"(11) "4"(12) "4"(3) "4"(7) "4"(8) "4"(9) aa1 aa2)
next
case (5 t u r v a)
then show ?case
by (metis EmptyLeftSync Sync.sym append_Nil2 syncTlEmpty)
next
case (6 t u r v a aa)
from 6(1,2,7,8,11,13) have aa: "hd t\<notin>insert tick (ev ` C)\<and>hd u\<notin>insert tick (ev ` C)"
by auto
from 6(7,8,12,13) obtain t1 t2 u1 u2 where aa0: "t=t1#t2\<and>u=u1#u2"
using list.exhaust_sel by blast
from 6(6) aa aa0 have aa1: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<or>(tl r) setinterleaves ((t,tl u), insert tick (ev ` C))\<and>hd r=hd u"
by auto
then show ?case
by (metis "6"(10) "6"(11) "6"(13) "6"(14) "6"(3) "6"(4) "6"(5) "6"(6) "6"(7) "6"(8) "6"(9) append_Nil2 empty_setinterleaving image_eqI)
next
case (7 t u r v a)
then show ?case
by (metis EmptyLeftSync Sync.sym append_Nil2 syncTlEmpty)
next
case (8 t u r v a aa)
from 8(1,2,7,8,11,13) have aa: "hd t\<notin>insert tick (ev ` C)\<and>hd u\<notin>insert tick (ev ` C)"
by auto
from 8(7,8,12,13) obtain t1 t2 u1 u2 where aa0: "t=t1#t2\<and>u=u1#u2"
using list.exhaust_sel by blast
from 8(6) aa aa0 have aa1: "(tl r) setinterleaves ((tl t, u), insert tick (ev ` C))\<and>hd r=hd t\<or>(tl r) setinterleaves ((t,tl u), insert tick (ev ` C))\<and>hd r=hd u"
by auto
then show ?case
by (metis "8"(10) "8"(11) "8"(13) "8"(14) "8"(3) "8"(4) "8"(5) "8"(6) "8"(7) "8"(8) "8"(9) append_Nil2 empty_setinterleaving image_eqI)
qed
next
case (15 x)
then show ?case
apply(elim exE disjE conjE )
proof(goal_cases)
case (1 a t u r v)
then show ?case
apply(rule_tac x="hd x#t" in exI, rule_tac x="u" in exI, rule_tac x="hd x#r" in exI, rule_tac x="v" in exI, simp)
by (metis disjoint_iff_not_equal event.distinct(1) event.inject imageE list.exhaust_sel)
next
case (2 a t u r v aa)
then show ?case
apply(rule_tac x="hd x#t" in exI, rule_tac x="u" in exI, rule_tac x="hd x#r" in exI, rule_tac x="v" in exI, simp, auto)
by (metis list.exhaust_sel, simp add: disjoint_iff_not_equal image_iff synSingleHeadAdd)
next
case (3 a t u r v aa)
from 3(2,4) have aa0: "hd x\<notin>insert tick (ev ` C)"
by auto
from 3(3,4,8,9) aa0 have aa: "(hd x#r) setinterleaves ((t, hd x#u), insert tick (ev ` C))\<and> x =(hd x#r) @ v "
by (metis Sync.sym append_Cons list.exhaust_sel synSingleHeadAdd)
then show ?case
using 3(4,5,6,7,10,11,12,13,14) aa0 event.inject by fastforce
next
case (4 a t u r v)
then show ?case
apply(rule_tac x="hd x#t" in exI, rule_tac x="u" in exI, rule_tac x="hd x#r" in exI, rule_tac x="v" in exI, simp)
by (metis disjoint_iff_not_equal event.distinct(1) event.inject imageE list.exhaust_sel)
next
case (5 a t u r v aa)
then show ?case
apply(rule_tac x="hd x#t" in exI, rule_tac x="u" in exI, rule_tac x="hd x#r" in exI, rule_tac x="v" in exI, simp, auto)
by (metis list.exhaust_sel, simp add: disjoint_iff_not_equal image_iff synSingleHeadAdd)
next
case (6 a t u r v aa)
from 6(2,4) have aa0: "hd x\<notin>insert tick (ev ` C)"
by auto
from 6(3,4,8,9) aa0 have aa: "(hd x#r) setinterleaves ((t, hd x#u), insert tick (ev ` C))\<and> x =(hd x#r) @ v "
by (metis Sync.sym append_Cons list.exhaust_sel synSingleHeadAdd)
then show ?case
using 6(4,5,6,7,10,11,12,13,14) by (metis (no_types, lifting) event.inject imageE list.sel(1) list.sel(3))
next
case (7 a t u r v aa)
from 7(1,4) have aa0: "hd x\<notin>insert tick (ev ` C)"
by auto
from 7(3,4,8,9) aa0 have aa: "(hd x#r) setinterleaves ((t, hd x#u), insert tick (ev ` C))\<and> x =(hd x#r) @ v "
by (metis Sync.sym append_Cons list.exhaust_sel synSingleHeadAdd)
then show ?case
using 7(4,5,6,7,10,11,12,13,14) aa0 by fastforce
next
case (8 a t u r v)
then show ?case
apply(rule_tac x="hd x#t" in exI, rule_tac x="u" in exI, rule_tac x="hd x#r" in exI, rule_tac x="v" in exI, simp, auto)
by (metis list.exhaust_sel)
next
case (9 a t u r v aa)
then show ?case
apply(rule_tac x="hd x#t" in exI, rule_tac x="u" in exI, rule_tac x="hd x#r" in exI, rule_tac x="v" in exI, simp, auto)
by (metis list.exhaust_sel, simp add: disjoint_iff_not_equal image_iff synSingleHeadAdd)
next
case (10 a t u r v aa)
from 10(1,4) have aa0: "hd x\<notin>insert tick (ev ` C)"
by auto
from 10(3,4,8,9) aa0 have aa: "(hd x#r) setinterleaves ((t, hd x#u), insert tick (ev ` C))\<and> x =(hd x#r) @ v "
by (metis Sync.sym append_Cons list.exhaust_sel synSingleHeadAdd)
then show ?case
using 10(4,5,6,7,10,11,12,13,14) by (metis (no_types, lifting) event.inject imageE list.sel(1) list.sel(3))
next
case (11 a t u r v)
then show ?case
apply(rule_tac x="hd x#t" in exI, rule_tac x="u" in exI, rule_tac x="hd x#r" in exI, rule_tac x="v" in exI, simp, auto)
by (metis list.exhaust_sel)
next
case (12 a t u r v aa)
then show ?case
apply(rule_tac x="hd x#t" in exI, rule_tac x="u" in exI, rule_tac x="hd x#r" in exI, rule_tac x="v" in exI, simp, auto)
by (metis list.exhaust_sel, simp add: disjoint_iff_not_equal image_iff synSingleHeadAdd)
qed
qed
lemma mprefix_Par_distr:
" \<lbrakk>A \<subseteq> C; B \<subseteq> C\<rbrakk> \<Longrightarrow> (Mprefix A P \<lbrakk> C \<rbrakk> Mprefix B Q) = \<box> x \<in> A \<inter> B \<rightarrow> (P x \<lbrakk> C \<rbrakk> Q x)"
apply(auto simp add:Process_eq_spec, simp_all add:D_sync F_sync F_Mprefix T_Mprefix D_Mprefix)
apply(elim exE disjE conjE)
apply auto[1]
apply (metis (no_types, lifting) empty_iff image_iff list.exhaust_sel setinterleaving.simps(2) subsetCE subset_insertI)
apply (metis (no_types, lifting) Sync.sym empty_iff image_iff insertI2 list.exhaust_sel setinterleaving.simps(2) subsetCE)
proof(goal_cases)
case (1 a b t u X Y aa ab)
from 1(1,2,3,4,7,8) have a0:"a \<noteq> []\<and>hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
using empty_setinterleaving by blast
from 1(4,7,8) a0 have a1: "hd a \<in> ev ` (A \<inter> B)"
using syncSameHdTl by fastforce
then show ?case
using 1(4,5,9,10,11,12) by (metis a0 event.inject syncSameHdTl syncSameHdTl)
next
case (2 a b t u r v aa)
then show ?case
by (metis (no_types, lifting) Sync.si_empty3 contra_subsetD equals0D imageE list.exhaust_sel rev_image_eqI subset_insertI)
next
case (3 a b t u r v aa ab)
from 3(1,2,3,4,5,6,7,8,11,13) have a0:"a \<noteq> []\<and>hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by (metis (no_types, lifting) Nil_is_append_conv empty_setinterleaving image_iff insertI2 subsetCE)
from 3(5,6,7) a0 have a2: "hd r=hd t\<and>hd r=hd u\<and>hd r=hd a\<and>hd a \<in> ev ` (A \<inter> B)"
by (metis (mono_tags, lifting) "3"(11) "3"(13) "3"(8) IntI empty_setinterleaving event.inject hd_append2 imageE image_eqI syncSameHdTl)
then show ?case
by (metis (no_types, lifting) "3"(10) "3"(13) "3"(14) "3"(3) "3"(4) "3"(5) "3"(6) "3"(7) "3"(9) a0 empty_setinterleaving event.inject syncSameHdTl tickFree_tl tl_append2)
next
case (4 a b t u r v aa)
then show ?case
by (metis (no_types, lifting) Sync.si_empty3 contra_subsetD equals0D imageE list.exhaust_sel rev_image_eqI subset_insertI)
next
case (5 a b t u r v aa ab)
from 5(1,2,3,4,5,6,7,8,11,13) have a0:"a \<noteq> []\<and>hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by (metis (no_types, lifting) Nil_is_append_conv empty_setinterleaving image_iff insertI2 subsetCE)
from 5(5,6,7) a0 have a2: "hd r=hd t\<and>hd r=hd u\<and>hd r=hd a\<and>hd a \<in> ev ` (A \<inter> B)"
by (metis (mono_tags, lifting) "5"(11) "5"(13) "5"(8) IntI empty_setinterleaving event.inject hd_append2 imageE image_eqI syncSameHdTl)
then show ?case
by (metis (no_types, lifting) "5"(10) "5"(13) "5"(14) "5"(3) "5"(4) "5"(5) "5"(6) "5"(7) "5"(9) a0 empty_setinterleaving event.inject syncSameHdTl tickFree_tl tl_append2)
next
case (6 a b t u r v aa)
then show ?case
by (metis (no_types, lifting) Sync.si_empty3 contra_subsetD equals0D imageE image_eqI list.exhaust_sel subset_insertI)
next
case (7 a b t u r v aa ab)
from 7(1,2,3,4,5,6,7,8,11,13) have a0:"a \<noteq> []\<and>hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by (metis (no_types, lifting) Nil_is_append_conv empty_setinterleaving image_iff insertI2 subsetCE)
from 7(5,6,7) a0 have a2: "hd r=hd t\<and>hd r=hd u\<and>hd r=hd a\<and>hd a \<in> ev ` (A \<inter> B)"
by (metis "7"(11) "7"(13) "7"(4) "7"(8) IntI append_Nil2 event.inject imageE image_eqI syncSameHdTl)
then show ?case
by (metis "7"(10) "7"(13) "7"(14) "7"(3) "7"(4) "7"(5) "7"(6) "7"(9) a0 event.inject self_append_conv syncSameHdTl)
next
case (8 a b t u r v aa)
then show ?case
by (metis (no_types, lifting) Sync.sym equals0D imageE image_eqI insertI2 list.exhaust_sel setinterleaving.simps(2) subsetCE)
next
case (9 a b t u r v aa ab)
from 9(1,2,3,4,5,6,7,8,11,13) have a0:"a \<noteq> []\<and>hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by (metis (no_types, lifting) Nil_is_append_conv empty_setinterleaving image_iff insertI2 subsetCE)
from 9(5,6,7) a0 have a2: "hd r=hd t\<and>hd r=hd u\<and>hd r=hd a\<and>hd a \<in> ev ` (A \<inter> B)"
by (metis "9"(11) "9"(13) "9"(4) "9"(8) IntI append_Nil2 event.inject imageE image_eqI syncSameHdTl)
then show ?case
by (metis "9"(10) "9"(13) "9"(14) "9"(3) "9"(4) "9"(5) "9"(6) "9"(9) a0 event.inject self_append_conv syncSameHdTl)
next
case (10 a b)
then show ?case
apply(elim exE disjE conjE)
proof(goal_cases)
case 1
obtain X Y where aa0: "(X::'a event set)=b-ev ` A\<and>(Y::'a event set)=b-ev ` B"
by auto
from 1(1,2,4) aa0 have aa1: "X\<union>Y=b\<and>X\<inter>Y=b-(ev ` (A \<union> B))\<and>(A \<union> B)\<subseteq>C"
by auto
from aa0 aa1 have aa2: "b = (X \<union> Y) \<inter> insert tick (ev ` C) \<union> X \<inter> Y"
by auto
then show ?case
by (metis "1"(3) Diff_disjoint Int_commute Sync.si_empty1 aa0 insert_iff)
next
case (2 aa t u X Y)
from 2(5) obtain t1 u1 where aa0: "t1=(hd a)#t\<and>u1=(hd a)#u" by auto
from 2(2,4) aa0 have aa1: "hd t1\<in>insert tick (ev ` C)\<and>hd u1\<in>insert tick (ev ` C)\<and>hd u1=hd t1"
using image_eqI by auto
from 2(1,3,4,8) aa0 aa1 have aa2: "a setinterleaves ((t1, u1), insert tick (ev ` C))"
using list.exhaust_sel by auto
from 2(4,5) aa0 have aa3: "t1\<noteq>[]\<and>u1\<noteq>[]\<and> hd t1 \<in> ev ` A\<and> hd u1 \<in> ev ` B"
by auto
then show ?case
using 2(5,6,7,9) by (metis aa1 aa0 aa2 list.sel(3) syncSameHdTl)
next
case (3 aa t u r v)
from 3(3,5,8) obtain r1 t1 u1 where aa0: "r1=ev aa#r\<and>t1=ev aa#t\<and>u1=ev aa#u\<and>a=r1@v"
using hd_Cons_tl by fastforce
from 3(1,2,4,5,7,9) aa0 have aa1: "tickFree r1\<and>ev aa\<in>insert tick (ev ` C)"
by (metis IntD2 event.inject event.simps(3) imageE image_eqI insertI2 subsetCE tickFree_Cons)
from 3(3,5,8,9) aa0 aa1 have aa2:"r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>hd t1 \<in> ev ` A\<and>tl t1 \<in> D (P aa)\<and>hd u1 \<in> ev ` B \<and> tl u1 \<in> T (Q aa)"
using "3"(10) "3"(11) "3"(4) imageE image_eqI by force
then show ?case
by (metis (no_types, lifting) "3"(6) aa0 aa1 event.inject imageE list.sel(1) list.simps(3))
next
case (4 aa t u r v)
from 4(3,5,8) obtain r1 t1 u1 where aa0: "r1=ev aa#r\<and>t1=ev aa#t\<and>u1=ev aa#u\<and>a=r1@v"
using hd_Cons_tl by fastforce
from 4(1,2,4,5,7,9) aa0 have aa1: "tickFree r1\<and>ev aa\<in>insert tick (ev ` C)"
using event.distinct(1) by fastforce
from 4(3,5,8,9) aa0 aa1 have aa2:"r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>hd t1 \<in> ev ` A\<and>tl t1 \<in> D (Q aa)\<and>hd u1 \<in> ev ` B \<and> tl u1 \<in> T (P aa)"
using "4"(10) "4"(11) "4"(4) by force
then show ?case
by (metis (no_types, lifting) "4"(6) aa0 aa1 event.inject imageE list.sel(1) list.simps(3))
next
case (5 aa t u r v)
from 5(3,5,8) obtain r1 t1 u1 where aa0: "r1=ev aa#r\<and>t1=ev aa#t\<and>u1=ev aa#u"
using "5"(7) by auto
from 5(3,5,8,9) aa0 have aa: "a=r1@v"
using list.exhaust_sel by force
from 5(2,3,4,5,8,9) aa0 have aa2:"r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>hd t1 \<in> ev ` A\<and>tl t1 \<in> D (P aa)\<and>hd u1 \<in> ev ` B \<and> tl u1 \<in> T (Q aa)"
using "5"(10) "5"(11) list.sel(1) list.sel(3) subsetCE by auto
then show ?case
by (metis (no_types, lifting) "5"(6) "5"(7) aa aa0 aa2 event.inject imageE list.distinct(1) list.sel(1))
next
case (6 aa t u r v)
from 6(3,5,8) obtain r1 t1 u1 where aa0: "r1=ev aa#r\<and>t1=ev aa#t\<and>u1=ev aa#u" by simp
from 6(3,5,8,9) aa0 have aa: "a=r1@v"
using list.exhaust_sel by force
from 6(2,3,4,5,8,9) aa0 have aa2:"r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>hd t1 \<in> ev ` A\<and>tl t1 \<in> D (Q aa)\<and>hd u1 \<in> ev ` B \<and> tl u1 \<in> T (P aa)"
using "6"(10) "6"(11) IntD2 Int_lower1 by auto
then show ?case
by (metis (no_types, lifting) "6"(6) "6"(7) aa aa0 aa2 event.inject imageE list.distinct(1) list.sel(1))
qed
next
case (11 x)
then show ?case
apply(elim exE disjE conjE)
apply (metis (no_types, lifting) Sync.sym empty_iff image_iff insertI2 list.exhaust_sel setinterleaving.simps(2) subsetCE)
proof(goal_cases)
case (1 t u r v a aa)
from 1(1,2,6,8,11,13) have aa1: "hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by blast
from 1(5,6,7,8,11,13) aa1 have aa2: " hd x \<in> ev ` (A \<inter> B)"
by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff syncSameHdTl)
then show ?case
using 1(3,4,5,6,7,9,10,13,14) by (metis (no_types, lifting) Nil_is_append_conv aa1 empty_setinterleaving event.inject hd_append2 syncSameHdTl tickFree_tl tl_append2)
next
case (2 t u r v a)
then show ?case
by (metis (no_types, lifting) Sync.si_empty3 equals0D imageE image_eqI insertI2 list.exhaust_sel subsetCE)
next
case (3 t u r v a aa)
from 3(1,2,6,8,11,13) have aa1: "hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by blast
from 3(5,6,7,8,11,13) aa1 have aa2: " hd x \<in> ev ` (A \<inter> B)"
by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff syncSameHdTl)
then show ?case
using 3(3,4,5,6,7,9,10,13,14) by (metis (no_types, lifting) Nil_is_append_conv aa1 empty_setinterleaving event.inject hd_append2 syncSameHdTl tickFree_tl tl_append2)
next
case (4 t u r v a)
then show ?case
by (metis (no_types, lifting) Sync.si_empty3 equals0D imageE image_eqI insertI2 list.exhaust_sel subsetCE)
next
case (5 t u r v a aa)
from 5(1,2,6,8,11,13) have aa1: "hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by blast
from 5(5,6,7,8,11,13) aa1 have aa2: " hd x \<in> ev ` (A \<inter> B)"
by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff syncSameHdTl)
then show ?case
using 5(3,4,5,6,7,9,10,13,14) by (metis aa1 append_Nil2 empty_setinterleaving event.inject syncSameHdTl)
next
case (6 t u r v a)
then show ?case
by (metis (no_types, lifting) Sync.si_empty3 equals0D imageE image_eqI insertI2 list.exhaust_sel subsetCE)
next
case (7 t u r v a aa)
from 7(1,2,6,8,11,13) have aa1: "hd t\<in>insert tick (ev ` C)\<and>hd u\<in>insert tick (ev ` C)"
by blast
from 7(5,6,7,8,11,13) aa1 have aa2: " hd x \<in> ev ` (A \<inter> B)"
by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff syncSameHdTl)
then show ?case
using 7(3,4,5,6,7,9,10,13,14) by (metis aa1 append_Nil2 empty_setinterleaving event.inject syncSameHdTl syncSameHdTl)
qed
next
case (12 x)
then show ?case
apply(elim exE disjE conjE)
proof(goal_cases)
case (1 a t u r v)
obtain r1 t1 u1 where aa0: "r1=hd x#r\<and>t1=hd x#t\<and>u1=hd x#u\<and>tickFree r1\<and>x = r1 @ v"
using "1"(3) "1"(5) "1"(7) "1"(8) hd_Cons_tl by fastforce
from 1(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>t1 \<noteq> []"
using aa0 by fastforce
from 1(4,5,10,11) aa0 have aa3: "(tl t1) \<in> D (P a)\<and>(tl u1) \<in> T (Q a)\<and> ev a = hd t1\<and>ev a = hd u1\<and> hd t1 \<in> ev ` A\<and> hd u1 \<in> ev ` B"
by force
then show ?case
by (metis (no_types, lifting) "1"(6) aa0 aa2 event.inject imageE)
next
case (2 a t u r v)
obtain r1 t1 u1 where aa0: "r1=hd x#r\<and>t1=hd x#t\<and>u1=hd x#u"
by auto
from 2(3,4,5,7,8,9) have aa1: "tickFree r1\<and>x = r1 @ v"
by (metis Cons_eq_appendI aa0 event.distinct(1) list.exhaust_sel tickFree_Cons)
from 2(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>t1 \<noteq> []"
using aa0 subsetCE by auto
from 2(4,5,10,11) aa0 have aa3: "(tl t1) \<in> D (Q a)\<and>(tl u1) \<in> T (P a)\<and> ev a = hd t1\<and>ev a = hd u1\<and> hd t1 \<in> ev ` A\<and> hd u1 \<in> ev ` B"
by auto
then show ?case
using "2"(6) aa1 aa2 by fastforce
next
case (3 a t u r v)
obtain r1 t1 u1 where aa0: "r1=hd x#r\<and>t1=hd x#t\<and>u1=hd x#u"
by auto
from 3(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>t1 \<noteq> []"
using aa0 subsetCE by auto
from 3(3,4,5,7,8,10,11) aa0 have aa3: "(tl t1) \<in> D (P a)\<and>(tl u1) \<in> T (Q a)\<and> ev a = hd t1\<and>ev a = hd u1\<and> hd t1 \<in> ev ` A\<and> hd u1 \<in> ev ` B\<and>x = r1 @ v"
by (metis (no_types, lifting) Int_lower1 Int_lower2 append_Nil2 imageE image_eqI list.exhaust_sel list.sel(1) list.sel(3) subsetCE)
then show ?case
by (metis (no_types, lifting) "3"(6) "3"(7) aa2 event.inject imageE)
next
case (4 a t u r v)
obtain r1 t1 u1 where aa0: "r1=hd x#r\<and>t1=hd x#t\<and>u1=hd x#u"
by auto
from 4(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` C))\<and>t1 \<noteq> []"
using aa0 subsetCE by auto
from 4(3,4,5,7,8,10,11) aa0 have aa3: "(tl t1) \<in> D (Q a)\<and>(tl u1) \<in> T (P a)\<and> ev a = hd t1\<and>ev a = hd u1\<and> hd t1 \<in> ev ` A\<and> hd u1 \<in> ev ` B\<and>x = r1 @ v"
by (metis (no_types, lifting) Int_lower1 Int_lower2 append_Nil2 imageE image_eqI list.exhaust_sel list.sel(1) list.sel(3) subsetCE)
then show ?case
by (metis (no_types, lifting) "4"(6) "4"(7) aa2 event.inject imageE)
qed
qed
lemma mprefix_Par_Int_skip: "(Mprefix A P \<lbrakk> B \<rbrakk> SKIP) = (\<box> x \<in> A - B \<rightarrow> (P x \<lbrakk> B \<rbrakk> SKIP))"
proof (auto simp add:Process_eq_spec_optimized, goal_cases)
case (1 x)
then show ?case
proof(simp_all add:D_sync D_Mprefix D_SKIP T_SKIP, elim exE conjE, goal_cases)
case (1 t u r v a)
then show ?case
apply(intro conjI)
using empty_setinterleaving apply blast
apply(elim disjE)
apply (metis (no_types, lifting) DiffI Sync.sym emptyLeftProperty empty_iff hd_append2 image_diff_subset insertI2 list.exhaust_sel setinterleaving.simps(2) subsetCE)
apply (metis D_imp_front_tickFree Sync.sym TickLeftSync empty_iff empty_setinterleaving event.distinct(1) ftf_syn21 insertI1 list.expand list.sel(1) list.set_intros(1) syncHd_Tl syncSameHdTl tickFree_def)
using Sync.sym emptyLeftNonSync emptyLeftProperty hd_in_set apply fastforce
apply (metis (no_types, lifting) DiffI Sync.sym append_Nil2 event.distinct(1) image_diff_subset insertI1 insertI2 list.sel(1) subsetCE syncHd_Tl syncSameHdTl)
apply(rule_tac x="a" in exI, intro conjI)
apply (metis Sync.sym emptyLeftProperty empty_setinterleaving hd_append2 insertI1 list.sel(1) syncHd_Tl syncSameHdTl)
apply(rule_tac x="tl t" in exI, rule_tac x="u" in exI, rule_tac x="tl r" in exI, rule_tac x="v" in exI)
by (metis (no_types, lifting) Sync.sym empty_setinterleaving event.distinct(1) insertI1 list.sel(1) syncHd_Tl syncSameHdTl syncTlEmpty tickFree_tl tl_append2)
qed
next
case (2 x)
then show ?case
proof(simp_all add:D_sync D_Mprefix D_SKIP T_SKIP, elim exE conjE, goal_cases)
case (1 a t u r v)
then show ?case
apply(rule_tac x="hd x#t" in exI, rule_tac x="u" in exI, rule_tac x="hd x#r" in exI,rule_tac x="v" in exI, auto)
using hd_Cons_tl list.exhaust_sel by force+
qed
next
case (3 a b)
then show ?case
proof(simp add:F_sync F_Mprefix F_SKIP, elim exE disjE, simp_all, goal_cases)
case (1 t u X)
{ fix X Y A B
assume "b = (X \<union> Y) \<inter> insert tick (ev ` B) \<union> X \<inter> Y" and "X \<inter> ev ` A = {}" and "tick \<notin> Y"
hence "b \<inter> ev ` (A - B) = {}" by (rule_tac equals0I, auto)
} note l1 = this
from 1 show ?case
apply(elim exE conjE disjE)
using l1 apply simp
apply(rule disjI2, simp)
apply(intro disjI2 conjI)
using empty_setinterleaving apply blast
apply (metis (no_types, hide_lams) DiffI Sync.sym UnI2 contra_subsetD emptyLeftNonSync
emptyLeftProperty hd_in_set image_diff_subset insert_is_Un)
apply (metis Sync.sym Un_commute emptyLeftProperty list.sel(1) list.sel(3) neq_Nil_conv syncTlEmpty)
apply(intro disjI2 conjI)
using empty_setinterleaving apply blast
apply (metis (no_types, hide_lams) DiffI Sync.sym event.distinct(1) event.inject
imageE image_eqI insertI1 insertI2 list.sel(1) syncHd_Tl syncSameHdTl)
apply (metis Sync.sym event.distinct(1) insertI1 list.sel(1) syncHd_Tl syncSameHdTl)
done
next
case (2 t u r v)
from 2(2) have "a \<in> D (Mprefix A P \<lbrakk>B\<rbrakk> SKIP)"
apply(simp add:D_sync) using T_SKIP by blast
with 2(3) have "a \<in> D (\<box> x \<in> A - B \<rightarrow> (P x \<lbrakk> B \<rbrakk> SKIP))" by simp
with 2(2) show ?case
by (simp add:D_sync D_Mprefix, rule_tac disjI2, metis)
qed
next
case (4 a b)
then show ?case
proof(simp add: F_sync F_Mprefix F_SKIP, elim exE disjE, simp_all, goal_cases)
case 1
then show ?case
apply(rule_tac disjI1, rule_tac x="[]" in exI, simp, rule_tac x="[]" in exI, simp)
apply(rule_tac x="b - (ev ` B)" in exI, rule_tac conjI)
by (blast, rule_tac x="b - {tick}" in exI, blast)
next
case 2
then show ?case
proof (elim conjE exE, elim disjE exE, simp_all, goal_cases)
case (1 aa t u X)
then show ?case
apply(erule_tac conjE, erule_tac exE, simp_all)
apply(rule_tac disjI1, rule_tac x="(ev aa)#t" in exI, rule_tac x=u in exI, rule_tac x=X in exI, intro conjI, simp_all)
apply auto[1]
by (metis (no_types, lifting) DiffE event.distinct(1) event.inject imageE insertE list.exhaust_sel synSingleHeadAdd)
next
case (2 aa t u r v)
from 2(2) have "a \<in> D (\<box> x \<in> A - B \<rightarrow> (P x \<lbrakk> B \<rbrakk> SKIP))"
apply(simp add:D_Mprefix D_sync) by (metis "2"(1) "2"(3) "2"(4))
with 2(5) have "a \<in> D (Mprefix A P \<lbrakk>B\<rbrakk> SKIP)" by simp
with 2(2) show ?case by (simp add:D_sync D_Mprefix)
qed
qed
qed
lemma par_int_ndet_distrib: "((P \<sqinter> Q) \<lbrakk> A \<rbrakk> M) = ((P \<lbrakk> A \<rbrakk> M) \<sqinter> (Q \<lbrakk> A \<rbrakk> M))"
apply(auto simp:Process_eq_spec, simp_all add: D_sync F_sync D_ndet F_ndet T_ndet)
by blast+
lemma prefix_Par_Int_skip1: "a \<notin> A \<Longrightarrow> (a \<rightarrow> P \<lbrakk>A\<rbrakk> SKIP) = (a \<rightarrow> (P \<lbrakk>A\<rbrakk> SKIP))"
by (metis (no_types, lifting) write0_def empty_Diff insert_Diff_if mprefix_Par_Int_skip mprefix_singl)
lemma prefix_Par_Int_skip2: "a \<in> A \<Longrightarrow> (a \<rightarrow> P \<lbrakk>A\<rbrakk> SKIP) = STOP"
by (simp add: Mprefix_STOP mprefix_Par_Int_skip write0_def)
lemma prefix_Par_Int_skip: "(a \<rightarrow> P \<lbrakk>A\<rbrakk> SKIP) = (if a \<in> A then STOP else (a \<rightarrow> (P \<lbrakk>A \<rbrakk> SKIP)))"
by (auto simp add: prefix_Par_Int_skip2 prefix_Par_Int_skip1)
lemma prefix_par_Int1: "\<lbrakk>a \<in> A; b \<in> A; a \<noteq> b\<rbrakk> \<Longrightarrow> (a \<rightarrow> P \<lbrakk>A\<rbrakk> (b \<rightarrow> Q)) = STOP"
using mprefix_Par_distr[of "{a}" A "{b}" "\<lambda>x. P" "\<lambda>x. Q", simplified]
by (auto, simp add: Mprefix_STOP mprefix_singl)
lemma prefix_par_Int2: "a \<in> A \<Longrightarrow> ((a \<rightarrow> P) \<lbrakk>A\<rbrakk> (a \<rightarrow> Q)) = (a \<rightarrow> (P \<lbrakk> A \<rbrakk> Q))"
by (simp add:write0_def mprefix_Par_distr[of "{a}" A "{a}" "\<lambda>x. P" "\<lambda>x. Q", simplified])
lemma prefix_Par_Int3: "\<lbrakk>a \<in> C; b \<notin> C\<rbrakk> \<Longrightarrow> (a \<rightarrow> P \<lbrakk>C\<rbrakk> (b \<rightarrow> Q)) = (b \<rightarrow> ((a \<rightarrow> P) \<lbrakk>C\<rbrakk> Q))"
using mprefix_Par_Int_distr[of "{b}" C "{a}" "\<lambda>x. P" "\<lambda>x. Q", simplified]
by (auto simp add:write0_def sync_commute)
lemma hide_sync_D1: "finite A \<Longrightarrow> A \<inter> C = {} \<Longrightarrow> D ((P \<lbrakk>C\<rbrakk> Q) \\ A) \<subseteq> D ((P \\ A) \<lbrakk>C\<rbrakk> (Q \\ A))"
proof (simp only:D_sync T_sync D_hiding T_hiding, intro subsetI CollectI,
elim CollectE exE conjE, elim exE CollectE disjE, goal_cases)
case (1 x t u ta ua r v)
then show ?case (is "\<exists>t ua ra va. ?P t ua ra va")
apply (subgoal_tac "?P (trace_hide ta (ev ` A)) (trace_hide ua (ev ` A)) (trace_hide r (ev ` A)) ((trace_hide v (ev ` A))@u)")
apply (metis (no_types, lifting))
apply(simp_all add:front_tickFree_append hiding_tickFree tickFree_def disjoint_iff_not_equal
hide_interleave[of "ev ` A" "insert tick (ev ` C)" r ta ua])
apply(elim conjE, erule_tac disjE, rule_tac disjI1, rule_tac conjI, case_tac "tickFree ta")
apply(meson front_tickFree_charn self_append_conv tickFree_def)
apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_syn21 ftf_syn32 in_set_conv_decomp_last
insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def)
apply(subst disj_commute, rule_tac disjCI, simp)
subgoal proof(goal_cases)
case 1
hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7)
from 1(11) 1(10) inf_hidden[of A ua Q] obtain ff where "isInfHiddenRun ff Q A \<and> ua \<in> range ff" by auto
with 1(2,3,4,7) a show ?case
apply(rule_tac x=ua in exI, rule_tac x="[]" in exI)
using front_tickFree_Nil tickFree_def by blast
qed
apply(rule_tac disjI2, rule_tac conjI, case_tac "tickFree ta")
apply(meson front_tickFree_charn self_append_conv tickFree_def)
apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_syn21 ftf_syn32 in_set_conv_decomp_last
insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def)
apply(subst disj_commute, rule_tac disjCI, simp)
proof(goal_cases)
case 1
hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7)
from 1(10) 1(11) inf_hidden[of A ua P] obtain ff where "isInfHiddenRun ff P A \<and> ua \<in> range ff" by auto
with 1(2,3,4,7) a show ?case
apply(rule_tac x=ua in exI, rule_tac x="[]" in exI)
using front_tickFree_Nil tickFree_def by blast
qed
next
case (2 x t u f)
then show ?case (is "\<exists>t ua ra va. ?P t ua ra va")
proof(elim UnE exE conjE rangeE CollectE, goal_cases 21)
case (21 xa)
note aa = 21(7)[rule_format, of xa]
with 21 show ?case
proof(elim UnE exE conjE rangeE CollectE, goal_cases 211 212)
case (211 taa uaa)
then show ?case (is "\<exists>t ua ra va. ?P t ua ra va")
proof (cases "\<forall>i. f i \<in> {s. \<exists>t u. t \<in> T P \<and> u \<in> T Q \<and> s setinterleaves ((t, u), ev ` C \<union> {tick})}")
case True
define ftu where "ftu \<equiv> \<lambda>i. (SOME (t,u). t \<in> T P \<and> u \<in> T Q \<and> (f i) setinterleaves ((t, u), ev ` C \<union> {tick}))"
define ft fu where "ft \<equiv> fst \<circ> ftu" and "fu \<equiv> snd \<circ> ftu"
have "inj ftu"
proof
fix x y
assume ftuxy: "ftu x = ftu y"
obtain t u where tu:"(t, u) = ftu x" by (metis surj_pair)
have "\<exists>t u. t \<in> T P \<and> u \<in> T Q \<and> f x setinterleaves ((t, u), ev ` C \<union> {tick})"
using True[rule_format, of x] by simp
with tu have a:"(f x) setinterleaves ((t, u), ev ` C \<union> {tick})"
unfolding ftu_def by (metis (mono_tags, lifting) exE_some old.prod.case tu)
have "\<exists>t u. t \<in> T P \<and> u \<in> T Q \<and> f y setinterleaves ((t, u), ev ` C \<union> {tick})"
using True[rule_format, of y] by simp
with ftuxy tu have b:"(f y) setinterleaves ((t, u), ev ` C \<union> {tick})"
unfolding ftu_def by (metis (mono_tags, lifting) exE_some old.prod.case tu)
from interleave_eq_size[OF a b] have "length (f x) = length (f y)" by assumption
with 211(6) show "x = y"
by (metis add.left_neutral less_length_mono linorder_neqE_nat not_add_less2 strict_mono_def)
qed
hence inf_ftu: "infinite (range ftu)" using finite_imageD by blast
have "range ftu \<subseteq> range ft \<times> range fu"
by (clarify, metis comp_apply fst_conv ft_def fu_def rangeI snd_conv)
with inf_ftu have inf_ft_fu: "infinite (range ft) \<or> infinite (range fu)"
by (meson finite_SigmaI infinite_super)
have a:"isInfHiddenRun f (P \<lbrakk>C\<rbrakk> Q) A"
using "2"(6) T_sync by blast
{ fix i
from a obtain t where "f i = f 0 @ t \<and> set t \<subseteq> (ev ` A)"
unfolding isInfHiddenRun_1 by blast
hence "set (f i) \<subseteq> set (f 0) \<union> (ev ` A)" by auto
} note b = this
{ fix x
obtain t u where tu:"(t, u) = ftu x" by (metis surj_pair)
have "\<exists>t u. t \<in> T P \<and> u \<in> T Q \<and> f x setinterleaves ((t, u), ev ` C \<union> {tick})"
using True[rule_format, of x] by simp
with tu have a:"t \<in> T P \<and> u \<in> T Q \<and> (f x) setinterleaves ((t, u), ev ` C \<union> {tick})"
unfolding ftu_def by (metis (mono_tags, lifting) exE_some old.prod.case tu)
hence "ft x \<in> T P \<and> fu x \<in> T Q \<and> (f x) setinterleaves ((ft x, fu x), ev ` C \<union> {tick})"
by (metis comp_apply fst_conv ft_def fu_def snd_conv tu)
hence "ft x \<in> T P \<and> fu x \<in> T Q \<and> set (ft x) \<union> set (fu x) \<subseteq> set (f x)
\<and> (f x) setinterleaves ((ft x, fu x), ev ` C \<union> {tick})" using interleave_set by blast
} note ft_fu_f = this
with b have c:"ft i \<in> T P \<and> fu i \<in> T Q \<and> set (ft i) \<union> set (fu i) \<subseteq> set (f 0) \<union> ev ` A" for i by blast
with inf_ft_fu show ?thesis
proof(elim disjE, goal_cases 2111 2112)
case 2111
have "\<exists>t'\<in>range ft. t = take i t' \<longrightarrow> (set t \<subseteq> set (f 0) \<union> ev ` A \<and> length t \<le> i)" for i
by simp (metis "211"(9) b)
hence "{t. \<exists>t'\<in>range ft. t = take i t'} \<subseteq> {t. set t \<subseteq> set (f 0) \<union> ev ` A \<and> length t \<le> i} " for i
by auto (meson UnE c in_set_takeD subsetCE sup.boundedE)
with 2(1) finite_lists_length_le[of "set (f 0) \<union> ev ` A"] have "finite {t. \<exists>t'\<in>range ft. t = take i t'}" for i
by (meson List.finite_set finite_Un finite_imageI finite_subset)
from KoenigLemma[of "range ft", OF 2111(2), rule_format, OF this] obtain ftf::"nat \<Rightarrow> 'a event list" where
d:"strict_mono ftf \<and> range ftf \<subseteq> {t. \<exists>t'\<in>range ft. t \<le> t'}" by blast
with c d[THEN conjunct2] have e:"range ftf \<subseteq> T P" using is_processT3_ST_pref by blast
define ftfs where f:"ftfs = ftf \<circ> (\<lambda>n. n + length (f 0))"
from e d[THEN conjunct1] strict_mono_def have f1:"range ftfs \<subseteq> T P" and f2:"strict_mono ftfs"
by (auto simp add: strict_mono_def f)
{ fix i
have t0:"length (ftfs 0) \<ge> length (f 0)" by (metis d[THEN conjunct1] add_leE comp_def f length_strict_mono)
obtain tt where t1:"ftfs i = (ftfs 0) @ tt" by (metis f2 le0 le_list_def strict_mono_less_eq)
from d[THEN conjunct2] f obtain j where t2:"ftfs i \<le> ft j" by simp blast
obtain tt1 where t3: "ft j = (ftfs 0) @ tt @ tt1" by (metis append.assoc le_list_def t1 t2)
with t0 interleave_order[of "f j" "ftfs 0" "tt@tt1" "ev ` C \<union> {tick}" "fu j"] ft_fu_f
have t4:"set tt \<subseteq> set (drop (length (f 0)) (f j))"
by (metis Un_subset_iff set_append set_drop_subset_set_drop subset_Un_eq)
with 21 isInfHiddenRun_1[of f _ A] have t5: "set (drop (length (f 0)) (f j)) \<subseteq> ev ` A"
by (metis (full_types) a append_eq_conv_conj)
with t4 have "set tt \<subseteq> ev ` A" by simp
with t1 have "trace_hide (ftfs i) (ev ` A) = trace_hide (ftfs 0) (ev ` A)"
by (simp add: subset_eq)
} note f3 = this
from d[THEN conjunct2] f obtain i where f4:"ftfs 0 \<le> ft i" by simp blast
show ?case
apply(rule_tac x="trace_hide (ft i) (ev ` A)" in exI, rule_tac x="trace_hide (fu i) (ev ` A)" in exI)
apply(rule_tac x="trace_hide (f i) (ev ` A)" in exI, rule_tac x="u" in exI)
proof (intro conjI, goal_cases 21111 21112 21113 21114 21115)
case 21111
then show ?case using 2 by (simp add: "211"(3))
next
case 21112
then show ?case by (metis "211"(4) "211"(9) a hiding_tickFree)
next
case 21113
then show ?case by (metis "211"(5) "211"(8) "211"(9))
next
case 21114
then show ?case
apply(rule hide_interleave)
using "211"(2) apply blast
using ft_fu_f by simp
next
case 21115
from f4 obtain u where f5:"ft i = (ftfs 0) @ u" by (metis le_list_def)
have "tickFree (f i)" by (metis "211"(4) "211"(8) "211"(9) hiding_tickFree)
with ft_fu_f have f6:"tickFree (ft i)" by (meson subsetCE sup.boundedE tickFree_def)
show ?case
apply (rule disjI1, rule conjI, simp)
apply(rule_tac x="ftfs 0" in exI, rule_tac x="trace_hide u (ev ` A)" in exI, intro conjI)
apply (metis f5 front_tickFree_Nil front_tickFree_mono ft_fu_f hiding_fronttickFree is_processT2_TR)
apply (metis f5 f6 tickFree_append)
apply (simp add: f5, rule disjI2, rule_tac x=ftfs in exI)
using f1 f2 f3 apply blast
using elemTIselemHT[of "fu i" Q A, simplified T_hiding] ft_fu_f by blast
qed
next
case 2112
have "\<exists>t'\<in>range fu. t = take i t' \<longrightarrow> (set t \<subseteq> set (f 0) \<union> ev ` A \<and> length t \<le> i)" for i
by simp (metis "211"(9) b)
hence "{t. \<exists>t'\<in>range fu. t = take i t'} \<subseteq> {t. set t \<subseteq> set (f 0) \<union> ev ` A \<and> length t \<le> i} " for i
by auto (meson UnE c in_set_takeD subsetCE sup.boundedE)
with 2(1) finite_lists_length_le[of "set (f 0) \<union> ev ` A"] have "finite {t. \<exists>t'\<in>range fu. t = take i t'}" for i
by (meson List.finite_set finite_Un finite_imageI finite_subset)
from KoenigLemma[of "range fu", OF 2112(2), rule_format, OF this] obtain fuf::"nat \<Rightarrow> 'a event list" where
d:"strict_mono fuf \<and> range fuf \<subseteq> {t. \<exists>t'\<in>range fu. t \<le> t'}" by blast
with c d[THEN conjunct2] have e:"range fuf \<subseteq> T Q" using is_processT3_ST_pref by blast
define fufs where f:"fufs = fuf \<circ> (\<lambda>n. n + length (f 0))"
from e d[THEN conjunct1] strict_mono_def have f1:"range fufs \<subseteq> T Q" and f2:"strict_mono fufs"
by (auto simp add: strict_mono_def f)
{ fix i
have t0:"length (fufs 0) \<ge> length (f 0)" by (metis d[THEN conjunct1] add_leE comp_def f length_strict_mono)
obtain tt where t1:"fufs i = (fufs 0) @ tt" by (metis f2 le0 le_list_def strict_mono_less_eq)
from d[THEN conjunct2] f obtain j where t2:"fufs i \<le> fu j" by simp blast
obtain tt1 where t3: "fu j = (fufs 0) @ tt @ tt1" by (metis append.assoc le_list_def t1 t2)
with t0 interleave_order[of "f j" "fufs 0" "tt@tt1" "ev ` C \<union> {tick}" "ft j"] ft_fu_f
have t4:"set tt \<subseteq> set (drop (length (f 0)) (f j))"
by (metis (no_types, hide_lams) Sync.sym Un_subset_iff order_trans set_append set_drop_subset_set_drop)
with 21 isInfHiddenRun_1[of f _ A] have t5: "set (drop (length (f 0)) (f j)) \<subseteq> ev ` A"
by (metis (full_types) a append_eq_conv_conj)
with t4 have "set tt \<subseteq> ev ` A" by simp
with t1 have "trace_hide (fufs i) (ev ` A) = trace_hide (fufs 0) (ev ` A)"
by (simp add: subset_eq)
} note f3 = this
from d[THEN conjunct2] f obtain i where f4:"fufs 0 \<le> fu i" by simp blast
show ?case
apply(rule_tac x="trace_hide (fu i) (ev ` A)" in exI, rule_tac x="trace_hide (ft i) (ev ` A)" in exI)
apply(rule_tac x="trace_hide (f i) (ev ` A)" in exI, rule_tac x="u" in exI)
proof (intro conjI, goal_cases 21111 21112 21113 21114 21115)
case 21111
then show ?case using 2 by (simp add: "211"(3))
next
case 21112
then show ?case by (metis "211"(4) "211"(9) a hiding_tickFree)
next
case 21113
then show ?case by (metis "211"(5) "211"(8) "211"(9))
next
case 21114
then show ?case
apply(rule hide_interleave)
using "211"(2) apply blast
using ft_fu_f using Sync.sym by blast
next
case 21115
from f4 obtain u where f5:"fu i = (fufs 0) @ u" by (metis le_list_def)
have "tickFree (f i)" by (metis "211"(4) "211"(8) "211"(9) hiding_tickFree)
with ft_fu_f have f6:"tickFree (fu i)" by (meson subsetCE sup.boundedE tickFree_def)
show ?case
apply (rule disjI2, rule conjI, simp)
apply(rule_tac x="fufs 0" in exI, rule_tac x="trace_hide u (ev ` A)" in exI, intro conjI)
apply (metis f5 front_tickFree_Nil front_tickFree_mono ft_fu_f hiding_fronttickFree is_processT2_TR)
apply (metis f5 f6 tickFree_append)
apply (simp add: f5, rule disjI2, rule_tac x=fufs in exI)
using f1 f2 f3 apply blast
using elemTIselemHT[of "ft i" P A, simplified T_hiding] ft_fu_f by blast
qed
qed
next
case False
then obtain xaa where "f xaa \<notin> {s. \<exists>t u. t \<in> T P \<and> u \<in> T Q \<and> s setinterleaves ((t, u), ev ` C \<union> {tick})}" by blast
with 211(7)[rule_format, of xaa] obtain ta ua ra va where bb:"front_tickFree va \<and>
(tickFree ra \<or> va = []) \<and> f xaa = ra @ va \<and> ra setinterleaves ((ta, ua), ev ` C \<union> {tick}) \<and>
(ta \<in> D P \<and> ua \<in> T Q \<or> ta \<in> D Q \<and> ua \<in> T P)" by blast
from 211 have cc:"x = trace_hide (f xaa) (ev ` A) @ u" by (metis (no_types, lifting))
from bb 211(9) "211"(8) "21"(4) have "tickFree ra \<and> tickFree va" by (metis hiding_tickFree tickFree_append)
with cc bb 211 show ?thesis
apply(subgoal_tac "?P (trace_hide ta (ev ` A)) (trace_hide ua (ev ` A)) (trace_hide ra (ev ` A)) ((trace_hide va (ev ` A))@u)")
apply (metis (no_types, lifting))
apply(simp_all add:front_tickFree_append hiding_tickFree tickFree_def disjoint_iff_not_equal
hide_interleave[of "ev ` A" "insert tick (ev ` C)" ra ta ua])
apply(elim conjE disjE, rule_tac disjI1, rule_tac conjI, case_tac "tickFree ta")
apply(meson front_tickFree_charn self_append_conv tickFree_def)
apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_syn21 ftf_syn32 in_set_conv_decomp_last
insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def)
apply(subst disj_commute, rule_tac disjCI, simp)
subgoal proof(goal_cases 2111)
case 2111
hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7)
from 2111(20) 2111(21) inf_hidden[of A ua Q] obtain ff where "isInfHiddenRun ff Q A \<and> ua \<in> range ff" by auto
with 2111(2,3,4,7) a show ?case
apply(rule_tac x=ua in exI, rule_tac x="[]" in exI)
using front_tickFree_Nil tickFree_def by blast
qed
apply(rule_tac disjI2, rule_tac conjI, case_tac "tickFree ta")
apply(meson front_tickFree_charn self_append_conv tickFree_def)
apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_syn21 ftf_syn32 in_set_conv_decomp_last
insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def)
apply(subst disj_commute, rule_tac disjCI, simp)
proof(goal_cases 2111)
case 2111
hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7)
from 2111(20) 2111(21) inf_hidden[of A ua P] obtain ff where "isInfHiddenRun ff P A \<and> ua \<in> range ff" by auto
with 2111(2,3,4,7) a show ?case
apply(rule_tac x=ua in exI, rule_tac x="[]" in exI)
using front_tickFree_Nil tickFree_def by blast
qed
qed
next
case (212 ta ua r v)
then show ?case (is "\<exists>t ua ra va. ?P t ua ra va")
apply (subgoal_tac "?P (trace_hide ta (ev ` A)) (trace_hide ua (ev ` A)) (trace_hide r (ev ` A)) ((trace_hide v (ev ` A))@u)")
apply (metis (no_types, lifting))
apply(simp_all add:front_tickFree_append hiding_tickFree tickFree_def disjoint_iff_not_equal
hide_interleave[of "ev ` A" "insert tick (ev ` C)" r ta ua])
apply(erule_tac disjE, rule_tac disjI1, rule_tac conjI, case_tac "tickFree ta")
apply(meson front_tickFree_charn self_append_conv tickFree_def)
apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_syn21 ftf_syn32 in_set_conv_decomp_last
insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def)
apply(subst disj_commute, rule_tac disjCI, simp)
subgoal proof(goal_cases 2121)
case 2121
hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7)
from 2121(14) 2121(13) inf_hidden[of A ua Q] obtain ff where "isInfHiddenRun ff Q A \<and> ua \<in> range ff" by auto
with 2121(2,3,4,7) a show ?case
apply(rule_tac x=ua in exI, rule_tac x="[]" in exI)
using front_tickFree_Nil tickFree_def by blast
qed
apply(rule_tac disjI2, rule_tac conjI, case_tac "tickFree ta")
apply(meson front_tickFree_charn self_append_conv tickFree_def)
apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_syn21 ftf_syn32 in_set_conv_decomp_last
insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def)
apply(subst disj_commute, rule_tac disjCI, simp)
proof(goal_cases 2121)
case 2121
hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7)
from 2121(14) 2121(13) inf_hidden[of A ua P] obtain ff where "isInfHiddenRun ff P A \<and> ua \<in> range ff" by auto
with 2121(2,3,4,7) a show ?case
apply(rule_tac x=ua in exI, rule_tac x="[]" in exI)
using front_tickFree_Nil tickFree_def by blast
qed
qed
qed
qed
lemma hide_sync_D2:
assumes *:"A \<inter> C = {}"
shows "D ((P \\ A) \<lbrakk>C\<rbrakk> (Q \\ A)) \<subseteq> D ((P \<lbrakk>C\<rbrakk> Q) \\ A)"
proof -
{ fix P Q
have "{s. \<exists>t u r v. front_tickFree v \<and> (tickFree r \<or> v = []) \<and>
s = r @ v \<and> r setinterleaves ((t, u), insert tick (ev ` C)) \<and>
(t \<in> D (P \\ A) \<and> u \<in> T (Q \\ A))} \<subseteq> D ((P \<lbrakk>C\<rbrakk> Q) \\ A)"
proof(simp add:D_hiding D_sync, intro subsetI CollectI, elim exE conjE CollectE, goal_cases a)
case (a x t u r v ta1 ta2)
from a(1,3-9) show ?case
proof(erule_tac disjE, goal_cases)
case 1
with interleave_append[of r "trace_hide ta1 (ev ` A)" "ta2" "insert tick (ev ` C)" u] obtain u1 u2 r1 r2
where a1:"u = u1 @ u2" and a2:"r = r1 @ r2" and
a3:"r1 setinterleaves ((trace_hide ta1 (ev ` A), u1), insert tick (ev ` C))" and
a4:"r2 setinterleaves ((ta2, u2), insert tick (ev ` C))" by blast
with 1 show ?case
proof(auto simp only:T_hiding, goal_cases 11 12 13)
case (11 ua)
from trace_hide_append[OF 11(12)] obtain ua1 where a5:"u1 = trace_hide ua1 (ev ` A) \<and> ua1 \<le> ua \<and> ua1 \<in> T Q"
using "11"(13) F_T is_processT3_ST_pref le_list_def by blast
from interleave_hide[OF _ 11(10)[simplified a5]] obtain ra1
where a6:"r1 = trace_hide ra1 (ev ` A) \<and> ra1 setinterleaves ((ta1, ua1), insert tick (ev ` C))"
using * by blast
from a(2) 11 a5 a6 show ?case
apply(exI ra1, exI "r2@v", intro conjI, simp_all (asm_lr))
apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_syn is_processT2_TR)
apply (metis equals0D ftf_syn1 ftf_syn21 insertI1 tickFree_def)
using front_tickFree_Nil by blast
next
case (12 ua1 ua2)
then show ?case
proof(cases "u1 \<le> trace_hide ua1 (ev ` A)")
case True
then obtain uu1 where "u1@uu1 = trace_hide ua1 (ev ` A)" using le_list_def by blast
with trace_hide_append[OF this] obtain uaa1 where a5:"u1 = trace_hide uaa1 (ev ` A) \<and> uaa1 \<le> ua1 \<and> uaa1 \<in> T Q"
using "12"(15) NT_ND is_processT3_ST_pref le_list_def by blast
from interleave_hide[OF _ 12(10)[simplified a5]] obtain rr1
where a6:"r1 = trace_hide rr1 (ev ` A) \<and> rr1 setinterleaves ((ta1, uaa1), insert tick (ev ` C))"
using * by blast
from a(2) 12 a5 a6 show ?thesis
apply(exI rr1, exI "r2@v", intro conjI, simp_all (asm_lr))
apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_syn is_processT2_TR)
apply (metis equals0D ftf_syn1 ftf_syn21 insertI1 tickFree_def)
using front_tickFree_Nil by blast
next
case False
with 12(14) obtain uaa1 where "u1 = trace_hide ua1 (ev ` A)@uaa1"
by (metis append_eq_append_conv2 le_list_def)
with a3 interleave_append_sym[of r1 "trace_hide ta1 (ev ` A)" "insert tick (ev ` C)" "trace_hide ua1 (ev ` A)" uaa1]
obtain taa1 taa2 ra1 ra2
where aa1:"trace_hide ta1 (ev ` A) = taa1 @ taa2" and aa2:"r1 = ra1 @ ra2" and
aa3:"ra1 setinterleaves ((taa1, trace_hide ua1 (ev ` A)), insert tick (ev ` C))" and
aa4:"ra2 setinterleaves ((taa2,uaa1), insert tick (ev ` C))" by blast
with trace_hide_append[OF aa1[symmetric]] obtain taaa1 where a5:"taa1 = trace_hide taaa1 (ev ` A) \<and> taaa1 \<le> ta1 \<and> taaa1 \<in> T P"
using "12"(7) D_T is_processT3_ST_pref le_list_def by blast
with interleave_hide[OF _ aa3[simplified a5]] obtain rr1
where a6:"ra1 = trace_hide rr1 (ev ` A) \<and> rr1 setinterleaves ((taaa1, ua1), insert tick (ev ` C))"
using * by blast
from a(2) 12 a5 show ?thesis
apply(exI rr1, exI "ra2@r2@v", intro conjI, simp_all (asm_lr))
apply (metis aa2 front_tickFree_append front_tickFree_mono ftf_syn hiding_tickFree self_append_conv tickFree_append)
apply (metis a6 ftf_syn1 le_list_def tickFree_append tickFree_def)
using a6 aa2 apply blast
using Sync.sym a6 front_tickFree_Nil by blast
qed
next
case (13 ua fa xa)
with isInfHiddenRun_1[of fa Q] have a0:"\<forall>i. \<exists>t. fa i = fa 0 @ t \<and> set t \<subseteq> (ev ` A)" by simp
define tlf where "tlf \<equiv> \<lambda>i. drop (length (fa 0)) (fa i)"
have tlf_def2:"(fa x) = (fa 0) @ (tlf x)" for x by (metis a0 append_eq_conv_conj tlf_def)
{ fix x y
assume *:"(x::nat) < y"
hence "fa x = fa 0 @ tlf x \<and> fa y = fa 0 @ tlf y" by (metis tlf_def2)
hence "tlf x < tlf y"
using strict_monoD[OF 13(16), of x y] by (simp add: A "*" le_list_def)
} note tlf_strict_mono = this
have tlf_range:"set (tlf i) \<subseteq> (ev ` A)" for i
by (metis a0 append_eq_conv_conj tlf_def)
from 13 show ?case
proof(cases "u1 \<le> trace_hide (fa xa) (ev ` A)")
case True
then obtain uu1 where "u1@uu1 = trace_hide (fa xa) (ev ` A)" using le_list_def by blast
with trace_hide_append[OF this] obtain uaa1 where a5:"u1 = trace_hide uaa1 (ev ` A) \<and> uaa1 \<le> (fa xa) \<and> uaa1 \<in> T Q"
by (metis "13"(17) is_processT3_ST le_list_def)
from interleave_hide[OF _ 13(10)[simplified a5]] obtain rr1
where a6:"r1 = trace_hide rr1 (ev ` A) \<and> rr1 setinterleaves ((ta1, uaa1), insert tick (ev ` C))"
using * by blast
from a(2) 13 a5 a6 show ?thesis
apply(exI rr1, exI "r2@v", intro conjI, simp_all (asm_lr))
apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_syn is_processT2_TR)
apply (metis equals0D ftf_syn1 ftf_syn21 insertI1 tickFree_def)
using front_tickFree_Nil by blast
next
case False
with 13(14) obtain uaa1 where "u1 = trace_hide (fa 0) (ev ` A)@uaa1"
by (metis "13"(18) append_eq_append_conv2 le_list_def)
with a3 interleave_append_sym[of r1 "trace_hide ta1 (ev ` A)" "insert tick (ev ` C)" "trace_hide (fa 0) (ev ` A)" uaa1]
obtain taa1 taa2 ra1 ra2
where aa1:"trace_hide ta1 (ev ` A) = taa1 @ taa2" and aa2:"r1 = ra1 @ ra2" and
aa3:"ra1 setinterleaves ((taa1, trace_hide (fa 0) (ev ` A)), insert tick (ev ` C))" and
aa4:"ra2 setinterleaves ((taa2,uaa1), insert tick (ev ` C))" by blast
with trace_hide_append[OF aa1[symmetric]] obtain taaa1 where a5:"taa1 = trace_hide taaa1 (ev ` A) \<and> taaa1 \<le> ta1 \<and> taaa1 \<in> T P"
using "13"(7) D_T is_processT3_ST_pref le_list_def by blast
with interleave_hide[OF _ aa3[simplified a5]] obtain rr1
where a6:"ra1 = trace_hide rr1 (ev ` A) \<and> rr1 setinterleaves ((taaa1, (fa 0)), insert tick (ev ` C))"
using * by blast
from a(2) 13 a0 a5 a6 aa1 aa2 aa3 show ?thesis
apply(exI rr1, exI "ra2@r2@v", intro conjI, simp_all (asm_lr))
apply (metis front_tickFree_append front_tickFree_mono ftf_syn hiding_tickFree self_append_conv tickFree_append)
apply (metis a6 ftf_syn1 le_list_def tickFree_append tickFree_def)
proof (rule_tac disjI2, exI "\<lambda>i. rr1@(tlf i)", intro conjI, goal_cases 131 132 133 134)
case 131
then show ?case apply(rule_tac strict_monoI, drule_tac tlf_strict_mono, rule_tac less_append) by assumption
next
case 132
have "(rr1@tlf i) setinterleaves ((taaa1, fa i),insert tick (ev ` C))" for i
apply(subst tlf_def2, rule interleave_append_tail_sym[OF a6[THEN conjunct2], of "tlf i"])
using * tlf_range by blast
then show ?case
unfolding T_sync using a5 13(17) by auto
next
case 133
from tlf_range have "trace_hide (y @ tlf i) (ev ` A) = trace_hide y (ev ` A)" for i y
apply(induct y) using filter_empty_conv apply fastforce by simp
then show ?case by simp
next
case 134
have "tlf 0 = []"
by (simp add: tlf_def)
then show ?case by simp
qed
qed
qed
next
case 2
from 2(8) obtain nta::nat and ff::"nat \<Rightarrow> 'a event list"
where "strict_mono ff" and ff:"(\<forall>i. ff i \<in> T P) \<and> ta1 = ff nta \<and>
(\<forall>i. trace_hide (ff i) (ev ` A) = trace_hide (ff 0) (ev ` A))" by blast
define f where "f \<equiv> (\<lambda>i. ff (i + nta))"
with ff have f1:"strict_mono f" and f2:"(\<forall>i. f i \<in> T P)" and f3:"ta1 = f 0" and
f4:"(\<forall>i. trace_hide (f i) (ev ` A) = trace_hide ta1 (ev ` A))"
apply auto apply(rule strict_monoI, rule \<open>strict_mono ff\<close>[THEN strict_monoD])
by simp metis
with isInfHiddenRun_1[of f P] have a0:"\<forall>i. \<exists>t. f i = f 0 @ t \<and> set t \<subseteq> (ev ` A)" by simp
define tlf where "tlf \<equiv> \<lambda>i. drop (length (f 0)) (f i)"
have tlf_def2:"(f x) = ta1 @ (tlf x)" for x by (metis a0 f3 append_eq_conv_conj tlf_def)
{ fix x y
assume *:"(x::nat) < y"
hence "f x = f 0 @ tlf x \<and> f y = f 0 @ tlf y" by (metis f3 tlf_def2)
hence "tlf x < tlf y"
using strict_monoD[OF f1, of x y] by (simp add: A "*" le_list_def)
} note tlf_strict_mono = this
have tlf_range:"set (tlf i) \<subseteq> (ev ` A)" for i
by (metis a0 append_eq_conv_conj tlf_def)
with 2 interleave_append[of r "trace_hide ta1 (ev ` A)" "ta2" "insert tick (ev ` C)" u] obtain u1 u2 r1 r2
where a1:"u = u1 @ u2" and a2:"r = r1 @ r2" and
a3:"r1 setinterleaves ((trace_hide ta1 (ev ` A), u1), insert tick (ev ` C))" and
a4:"r2 setinterleaves ((ta2, u2), insert tick (ev ` C))" by blast
with 2(1-6) f1 f2 f3 f4 show ?case
proof(auto simp only:T_hiding, goal_cases 21 22 23)
case (21 ua)
from trace_hide_append[OF 21(14)] obtain ua1 where a5:"u1 = trace_hide ua1 (ev ` A) \<and> ua1 \<le> ua \<and> ua1 \<in> T Q"
using "21"(15) F_T is_processT3_ST_pref le_list_def by blast
from interleave_hide[OF _ a3[simplified a5]] obtain ra1
where a6:"r1 = trace_hide ra1 (ev ` A) \<and> ra1 setinterleaves ((ta1, ua1), insert tick (ev ` C))"
using * by blast
from a(2) 21 a5 a6 show ?case
apply(exI ra1, exI "r2@v", intro conjI, simp_all (asm_lr))
apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_syn is_processT2_TR)
apply (metis equals0D ftf_syn1 ftf_syn21 insertI1 tickFree_def)
proof (rule_tac disjI2, exI "\<lambda>i. ra1@(tlf i)", intro conjI, goal_cases 211 212 213 214)
case 211
then show ?case apply(rule_tac strict_monoI, drule_tac tlf_strict_mono, rule_tac less_append) by assumption
next
case 212
have "(ra1@tlf i) setinterleaves ((f i, ua1),insert tick (ev ` C))" for i
apply(subst tlf_def2, rule interleave_append_tail[OF a6[THEN conjunct2], of "tlf i"])
using * tlf_range by blast
then show ?case
unfolding T_sync using a5 21(15) f2 by auto
next
case 213
from tlf_range have "trace_hide (y @ tlf i) (ev ` A) = trace_hide y (ev ` A)" for i y
apply(induct y) using filter_empty_conv apply fastforce by simp
then show ?case by simp
next
case 214
have "tlf 0 = []"
by (simp add: tlf_def)
then show ?case by simp
qed
next
case (22 ua1 ua2)
then show ?case
proof(cases "u1 \<le> trace_hide ua1 (ev ` A)")
case True
then obtain uu1 where "u1@uu1 = trace_hide ua1 (ev ` A)" using le_list_def by blast
from trace_hide_append[OF this] obtain uaa1 where a5:"u1 = trace_hide uaa1 (ev ` A) \<and> uaa1 \<le> ua1 \<and> uaa1 \<in> T Q"
using "22"(17) D_T is_processT3_ST le_list_def by blast
from interleave_hide[OF _ a3[simplified a5]] obtain ra1
where a6:"r1 = trace_hide ra1 (ev ` A) \<and> ra1 setinterleaves ((ta1, uaa1), insert tick (ev ` C))"
using * by blast
from a(2) 22 a5 a6 show ?thesis
apply(exI ra1, exI "r2@v", intro conjI, simp_all (asm_lr))
apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_syn is_processT2_TR)
apply (metis equals0D ftf_syn1 ftf_syn21 insertI1 tickFree_def)
proof (rule_tac disjI2, exI "\<lambda>i. ra1@(tlf i)", intro conjI, goal_cases 211 212 213 214)
case 211
then show ?case apply(rule_tac strict_monoI, drule_tac tlf_strict_mono, rule_tac less_append) by assumption
next
case 212
have "(ra1@tlf i) setinterleaves ((f i, uaa1),insert tick (ev ` C))" for i
apply(subst tlf_def2, rule interleave_append_tail[OF a6[THEN conjunct2], of "tlf i"])
using * tlf_range by blast
then show ?case
unfolding T_sync using a5 22(15) f2 by auto
next
case 213
from tlf_range have "trace_hide (y @ tlf i) (ev ` A) = trace_hide y (ev ` A)" for i y
apply(induct y) using filter_empty_conv apply fastforce by simp
then show ?case by simp
next
case 214
have "tlf 0 = []"
by (simp add: tlf_def)
then show ?case by simp
qed
next
case False
with 22(16) obtain uaa1 where "u1 = trace_hide ua1 (ev ` A)@uaa1"
by (metis append_eq_append_conv2 le_list_def)
with a3 interleave_append_sym[of r1 "trace_hide ta1 (ev ` A)" "insert tick (ev ` C)" "trace_hide ua1 (ev ` A)" uaa1]
obtain taa1 taa2 ra1 ra2
where aa1:"trace_hide ta1 (ev ` A) = taa1 @ taa2" and aa2:"r1 = ra1 @ ra2" and
aa3:"ra1 setinterleaves ((taa1, trace_hide ua1 (ev ` A)), insert tick (ev ` C))" and
aa4:"ra2 setinterleaves ((taa2,uaa1), insert tick (ev ` C))" by blast
with trace_hide_append[OF aa1[symmetric]] obtain taaa1 where a5:"taa1 = trace_hide taaa1 (ev ` A) \<and> taaa1 \<le> ta1 \<and> taaa1 \<in> T P"
using f2 f3 is_processT3_ST_pref le_list_def by metis
with interleave_hide[OF _ aa3[simplified a5]] obtain rr1
where a6:"ra1 = trace_hide rr1 (ev ` A) \<and> rr1 setinterleaves ((taaa1, ua1), insert tick (ev ` C))"
using * by blast
from a(2) 22 a5 show ?thesis
apply(exI rr1, exI "ra2@r2@v", intro conjI, simp_all (asm_lr))
apply (metis a(5) a(8) aa2 front_tickFree_append front_tickFree_mono ftf_syn hiding_tickFree
is_processT2_TR self_append_conv tickFree_append)
apply (metis a6 ftf_syn1 le_list_def tickFree_append tickFree_def)
using a6 aa2 apply blast
using Sync.sym a6[THEN conjunct2] front_tickFree_Nil by (metis append_Nil2)
qed
next
case (23 ua fa xa)
with isInfHiddenRun_1[of fa Q] have a0:"\<forall>i. \<exists>t. fa i = fa 0 @ t \<and> set t \<subseteq> (ev ` A)" by simp
define tlfa where "tlfa \<equiv> \<lambda>i. drop (length (fa 0)) (fa i)"
have tlfa_def2:"(fa x) = (fa 0) @ (tlfa x)" for x by (metis a0 append_eq_conv_conj tlfa_def)
{ fix x y
assume *:"(x::nat) < y"
hence "fa x = fa 0 @ tlfa x \<and> fa y = fa 0 @ tlfa y" by (metis tlfa_def2)
hence "tlfa x < tlfa y"
using strict_monoD[OF 23(18), of x y] by (simp add: A "*" le_list_def)
} note tlfa_strict_mono = this
have tlfa_range:"set (tlfa i) \<subseteq> (ev ` A)" for i
by (metis a0 append_eq_conv_conj tlfa_def)
from 23 show ?case
proof(cases "u1 \<le> trace_hide (fa xa) (ev ` A)")
case True
then obtain uu1 where "u1@uu1 = trace_hide (fa xa) (ev ` A)" using le_list_def by blast
with trace_hide_append[OF this] obtain uaa1 where a5:"u1 = trace_hide uaa1 (ev ` A) \<and> uaa1 \<le> (fa xa) \<and> uaa1 \<in> T Q"
by (metis "23"(19) is_processT3_ST le_list_def)
from interleave_hide[OF _ a3[simplified a5]] obtain rr1
where a6:"r1 = trace_hide rr1 (ev ` A) \<and> rr1 setinterleaves ((ta1, uaa1), insert tick (ev ` C))"
using * by blast
from a(2) 23 a5 a6 show ?thesis
apply(exI rr1, exI "r2@v", intro conjI, simp_all (asm_lr))
apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_syn is_processT2_TR)
apply (metis equals0D ftf_syn1 ftf_syn21 insertI1 tickFree_def)
proof (rule_tac disjI2, exI "\<lambda>i. rr1@(tlf i)", intro conjI, goal_cases 211 212 213 214)
case 211
then show ?case apply(rule_tac strict_monoI, drule_tac tlf_strict_mono, rule_tac less_append) by assumption
next
case 212
have "(rr1@tlf i) setinterleaves ((f i, uaa1),insert tick (ev ` C))" for i
apply(subst tlf_def2, rule interleave_append_tail[OF a6[THEN conjunct2], of "tlf i"])
using * tlf_range by blast
then show ?case
unfolding T_sync using a5 23(17) f2 by auto
next
case 213
from tlf_range have "trace_hide (y @ tlf i) (ev ` A) = trace_hide y (ev ` A)" for i y
apply(induct y) using filter_empty_conv apply fastforce by simp
then show ?case by simp
next
case 214
have "tlf 0 = []"
by (simp add: tlf_def)
then show ?case by simp
qed
next
case False
with 23(16) obtain uaa1 where "u1 = trace_hide (fa 0) (ev ` A)@uaa1"
by (metis "23"(20) append_eq_append_conv2 le_list_def)
with a3 interleave_append_sym[of r1 "trace_hide ta1 (ev ` A)" "insert tick (ev ` C)" "trace_hide (fa 0) (ev ` A)" uaa1]
obtain taa1 taa2 ra1 ra2
where aa1:"trace_hide ta1 (ev ` A) = taa1 @ taa2" and aa2:"r1 = ra1 @ ra2" and
aa3:"ra1 setinterleaves ((taa1, trace_hide (fa 0) (ev ` A)), insert tick (ev ` C))" and
aa4:"ra2 setinterleaves ((taa2,uaa1), insert tick (ev ` C))" by blast
with trace_hide_append[OF aa1[symmetric]] obtain taaa1 where a5:"taa1 = trace_hide taaa1 (ev ` A) \<and> taaa1 \<le> ta1 \<and> taaa1 \<in> T P"
by (metis f2 f3 is_processT3_ST le_list_def)
with interleave_hide[OF _ aa3[simplified a5]] obtain rr1
where a6:"ra1 = trace_hide rr1 (ev ` A) \<and> rr1 setinterleaves ((taaa1, (fa 0)), insert tick (ev ` C))"
using * by blast
from a(2) 23 a0 a5 a6 aa1 aa2 aa3 show ?thesis
apply(exI rr1, exI "ra2@r2@v", intro conjI, simp_all (asm_lr))
apply (metis a(5) a(8) front_tickFree_append front_tickFree_mono ftf_syn hiding_tickFree is_processT2_TR self_append_conv)
apply (metis a6 ftf_syn1 le_list_def tickFree_append tickFree_def)
proof (rule_tac disjI2, exI "\<lambda>i. rr1@(tlfa i)", intro conjI, goal_cases 131 132 133 134)
case 131
then show ?case apply(rule_tac strict_monoI, drule_tac tlfa_strict_mono, rule_tac less_append) by assumption
next
case 132
have "(rr1@tlfa i) setinterleaves ((taaa1, fa i),insert tick (ev ` C))" for i
apply(subst tlfa_def2, rule interleave_append_tail_sym[OF a6[THEN conjunct2], of "tlfa i"])
using * tlfa_range by blast
then show ?case
unfolding T_sync using a5 23(19) by auto
next
case 133
from tlfa_range have "trace_hide (y @ tlfa i) (ev ` A) = trace_hide y (ev ` A)" for i y
apply(induct y) using filter_empty_conv apply fastforce by simp
then show ?case by simp
next
case 134
have "tlfa 0 = []"
by (simp add: tlfa_def)
then show ?case by simp
qed
qed
qed
qed
qed
}
from this[of P Q] this[of Q P] sync_commute[of P C Q] show ?thesis
by (simp add:D_sync T_sync) blast
qed
lemma hide_sync: "finite A \<Longrightarrow> A \<inter> C = {} \<Longrightarrow> ((P \<lbrakk>C\<rbrakk> Q) \\ A) = ((P \\ A) \<lbrakk>C\<rbrakk> (Q \\ A))"
proof(auto simp add: Process_eq_spec_optimized subset_antisym[OF hide_sync_D1 hide_sync_D2],
simp_all add: F_sync F_hiding,
goal_cases)
case (1 a b)
then show ?case
proof(elim disjE, goal_cases 11 12)
case 11
then show ?case
proof(elim exE conjE, elim disjE, goal_cases 111 112)
case (111 c)
from 111(7) obtain t u X Y where
a: "(t, X) \<in> F P" and
b: "(u, Y) \<in> F Q" and
c: "c setinterleaves ((t, u), insert tick (ev ` C))" and
d: "b \<union> ev ` A = (X \<union> Y) \<inter> insert tick (ev ` C) \<union> X \<inter> Y" by auto
from 1(2) d have e:"ev ` A \<subseteq> X \<inter> Y" by blast
from a b c d e 111(2,3) show ?case
apply(rule_tac disjI1)
apply(rule_tac x="trace_hide t (ev ` A)" in exI, rule_tac x="trace_hide u (ev ` A)" in exI)
apply(rule_tac x="X - ((ev ` A) - b)" in exI, intro conjI disjI1, rule_tac x="t" in exI)
apply (metis Int_subset_iff Un_Diff_Int Un_least is_processT4 sup_ge1)
apply(rule_tac x="Y - ((ev ` A) - b)" in exI, intro conjI disjI1, rule_tac x="u" in exI)
apply (metis Int_subset_iff Un_Diff_Int Un_least is_processT4 sup_ge1)
using hide_interleave[of "ev ` A" "insert tick (ev ` C)" c t u] by blast+
next
case (112 t)
from 112(7) have a:"front_tickFree t"
by (metis (mono_tags, hide_lams) D_imp_front_tickFree append_Nil2
front_tickFree_append ftf_syn is_processT2_TR)
from 112 have "t \<in> D (P \<lbrakk>C\<rbrakk> Q)" by (simp add:D_sync)
with 112(1,2,3) have "a \<in> D ((P \<lbrakk>C\<rbrakk> Q) \\ A)"
apply (simp add:D_hiding)
proof(case_tac "tickFree t", goal_cases 1121 1122)
case 1121
then show ?case
by (rule_tac x=t in exI, rule_tac x="[]" in exI, simp)
next
case 1122
with a obtain t' where "t = t'@[tick]" using nonTickFree_n_frontTickFree by blast
with a show ?case
apply (rule_tac x="t'" in exI, rule_tac x="[tick]" in exI, auto)
using front_tickFree_mono apply blast
using 1122(4) is_processT9 by blast
qed
then show ?case
apply(rule_tac disjI2)
using hide_sync_D1[of A C P Q, OF 112(1,2)] D_sync[of "P \\ A" C "Q \\ A"] by auto
qed
next
case 12
hence "a \<in> D ((P \<lbrakk>C\<rbrakk> Q) \\ A)" by (simp add:D_hiding)
then show ?case
apply(rule_tac disjI2)
using hide_sync_D1[of A C P Q, OF 12(1,2)] D_sync[of "P \\ A" C "Q \\ A"] by auto
qed
next
case (2 a b)
then show ?case
proof(elim disjE, goal_cases 21 22)
case 21
then show ?case
proof(elim exE conjE, elim disjE, goal_cases 211 212 213 214)
case (211 t u X Y)
then obtain ta ua where t211: "t = trace_hide ta (ev ` A) \<and> (ta, X \<union> ev ` A) \<in> F P" and
u211: "u = trace_hide ua (ev ` A) \<and> (ua, Y \<union> ev ` A) \<in> F Q" by blast
from interleave_hide[of "(ev ` A)" "insert tick (ev ` C)" a ta ua] "211"(1,2,3) t211 u211 obtain aa
where a211: "a = trace_hide aa (ev ` A) \<and> aa setinterleaves ((ta, ua), insert tick (ev ` C))" by blast
with 211(1,2,3,4) t211 u211 show ?case
apply(rule_tac disjI1, rule_tac x="aa" in exI, simp)
apply(rule_tac disjI1, rule_tac x="ta" in exI, rule_tac x="ua" in exI)
apply(rule_tac x="X \<union> ev ` A" in exI, rule conjI) using is_processT4 apply blast
apply(rule_tac x="Y \<union> ev ` A" in exI, rule conjI) using is_processT4 by blast+
next
case (212 t u X Y)
hence "u \<in> D (Q \\ A)" and "t \<in> T (P \\ A)"
apply (simp add:D_hiding)
using "212"(8) F_T elemTIselemHT by blast
with 212(3,8) have "a \<in> D (((P \\ A) \<lbrakk>C\<rbrakk> (Q \\ A)))"
apply (simp add:D_sync)
using Sync.sym front_tickFree_charn by blast
then show ?case
apply (rule_tac disjI2)
using hide_sync_D2[of A C P Q, OF 2(2)] D_hiding[of "(P \<lbrakk>C\<rbrakk> Q)", simplified] by auto
next
case (213 t u X Y)
hence "u \<in> T (Q \\ A)" and "t \<in> D (P \\ A)"
by (simp_all add:D_hiding, metis F_T elemTIselemHT)
with 213(3,8) have "a \<in> D (((P \\ A) \<lbrakk>C\<rbrakk> (Q \\ A)))"
apply (simp add:D_sync)
using Sync.sym front_tickFree_charn by blast
then show ?case
apply (rule_tac disjI2)
using hide_sync_D2[of A C P Q, OF 2(2)] D_hiding[of "(P \<lbrakk>C\<rbrakk> Q)", simplified] by auto
next
case (214 t u X Y)
hence "u \<in> T (Q \\ A)" and "t \<in> D (P \\ A)"
by (simp_all add:D_hiding T_hiding)
with 214(3,8) have "a \<in> D (((P \\ A) \<lbrakk>C\<rbrakk> (Q \\ A)))"
apply (simp add:D_sync)
using Sync.sym front_tickFree_charn by blast
then show ?case
apply (rule_tac disjI2)
using hide_sync_D2[of A C P Q, OF 2(2)] D_hiding[of "(P \<lbrakk>C\<rbrakk> Q)", simplified] by auto
qed
next
case 22
hence "a \<in> D (((P \\ A) \<lbrakk>C\<rbrakk> (Q \\ A)))" by (simp add:D_sync)
then show ?case
apply (rule_tac disjI2)
using hide_sync_D2[of A C P Q, OF 2(2)] D_hiding[of "(P \<lbrakk>C\<rbrakk> Q)", simplified] by auto
qed
qed
lemma sync_assoc_oneside_D: "D (P \<lbrakk>C\<rbrakk> (Q \<lbrakk>C\<rbrakk> S)) \<subseteq> D (P \<lbrakk>C\<rbrakk> Q \<lbrakk>C\<rbrakk> S)"
proof(auto simp add:D_sync T_sync del:disjE, goal_cases a)
case (a tt uu rr vv)
from a(3,4) show ?case
proof(auto, goal_cases)
case (1 t u)
with interleave_assoc_2[OF 1(5) 1(1)] obtain tu
where *:"tu setinterleaves ((tt, t), insert tick (ev ` C))" and **:"rr setinterleaves ((tu, u), insert tick (ev ` C))"
by blast
with a(1,2) 1 show ?case by (metis append.right_neutral front_tickFree_charn)
next
case (2 t u u1 u2)
with interleave_append_sym [OF 2(1)] obtain t1 t2 r1 r2
where a1:"tt = t1 @ t2" and a2:"rr = r1 @ r2" and
a3:"r1 setinterleaves ((t1, u1), insert tick (ev ` C))" and
a4:"r2 setinterleaves ((t2, u2), insert tick (ev ` C))" by blast
with interleave_assoc_2[OF 2(5) a3] obtain tu
where *:"tu setinterleaves ((t1, t), insert tick (ev ` C))" and **:"r1 setinterleaves ((tu, u), insert tick (ev ` C))"
by blast
with a1 a2 a3 a(1,2) 2 show ?case
apply(exI tu, exI u, exI r1, exI "r2@vv", intro conjI, simp_all)
apply (metis D_imp_front_tickFree append_Nil2 front_tickFree_append front_tickFree_mono ftf_syn)
using D_imp_front_tickFree front_tickFree_append front_tickFree_mono ftf_syn apply blast
by (metis NT_ND Sync.sym append_Nil2 front_tickFree_Nil is_processT3_ST)
next
case (3 v u u1 u2)
with interleave_append_sym [OF 3(1)] obtain t1 t2 r1 r2
where a1:"tt = t1 @ t2" and a2:"rr = r1 @ r2" and
a3:"r1 setinterleaves ((t1, u1), insert tick (ev ` C))" and
a4:"r2 setinterleaves ((t2, u2), insert tick (ev ` C))" by blast
with interleave_assoc_2[OF _ a3, of u v] obtain tu
where *:"tu setinterleaves ((t1, u), insert tick (ev ` C))" and **:"r1 setinterleaves ((tu, v), insert tick (ev ` C))"
using Sync.sym 3(5) by blast
with a1 a2 a3 a(1,2) 3 show ?case
apply(exI v, exI tu, exI r1, exI "r2@vv", intro conjI, simp_all)
apply (metis D_imp_front_tickFree append_Nil2 front_tickFree_append front_tickFree_mono ftf_syn)
using D_imp_front_tickFree front_tickFree_append front_tickFree_mono ftf_syn apply blast
using Sync.sym apply blast by (meson D_T is_processT3_ST)
next
case (4 t u)
with interleave_assoc_2[OF 4(3) 4(1)] obtain tu
where *:"tu setinterleaves ((tt, t), insert tick (ev ` C))" and **:"rr setinterleaves ((tu, u), insert tick (ev ` C))"
by blast
with a(1,2) 4 show ?case
apply(exI tu, exI u, exI rr, exI vv, simp) using NT_ND front_tickFree_Nil by blast
next
case (5 t u)
with interleave_assoc_2[OF _ 5(1), of u t] obtain tu
where *:"tu setinterleaves ((tt, u), insert tick (ev ` C))" and **:"rr setinterleaves ((tu, t), insert tick (ev ` C))"
using Sync.sym by blast
with a(1,2) 5 show ?case
apply(exI tu, exI t, exI rr, exI vv, simp) using NT_ND front_tickFree_Nil by blast
next
case (6 t u u1 u2)
with interleave_append [OF 6(1)] obtain t1 t2 r1 r2
where a1:"uu = t1 @ t2" and a2:"rr = r1 @ r2" and
a3:"r1 setinterleaves ((t1, u1), insert tick (ev ` C))" and
a4:"r2 setinterleaves ((t2, u2), insert tick (ev ` C))" using Sync.sym by blast
with interleave_assoc_2[OF 6(5) a3] obtain tu
where *:"tu setinterleaves ((t1, t), insert tick (ev ` C))" and **:"r1 setinterleaves ((tu, u), insert tick (ev ` C))"
by blast
with a1 a2 a3 a(1,2) 6 show ?case
apply(exI tu, exI u, exI r1, exI "r2@vv", intro conjI, simp_all)
apply (metis append_Nil2 front_tickFree_append front_tickFree_mono ftf_syn is_processT2_TR)
apply (meson front_tickFree_append front_tickFree_mono ftf_syn is_processT2_TR)
by (metis Sync.sym append_Nil2 front_tickFree_Nil is_processT3_ST)
next
case (7 v u t1 t2)
with interleave_append [OF 7(1)] obtain u1 u2 r1 r2
where a1:"uu = u1 @ u2" and a2:"rr = r1 @ r2" and
a3:"r1 setinterleaves ((t1, u1), insert tick (ev ` C))" and
a4:"r2 setinterleaves ((t2, u2), insert tick (ev ` C))" by blast
with interleave_assoc_2[of t1 u _ v r1 u1] obtain tu
where *:"tu setinterleaves ((u1, u), insert tick (ev ` C))" and **:"r1 setinterleaves ((tu, v), insert tick (ev ` C))"
using 7(5) Sync.sym by blast
with a1 a2 a3 a(1,2) 7 show ?case
apply(exI v, exI tu, exI r1, exI "r2@vv", intro conjI, simp_all)
apply (metis append_Nil2 front_tickFree_append front_tickFree_mono ftf_syn is_processT2_TR)
apply (meson front_tickFree_append front_tickFree_mono ftf_syn is_processT2_TR)
using Sync.sym apply blast by (meson D_T is_processT3_ST)
next
case (8 t u)
with interleave_assoc_2[OF 8(3), of rr uu] obtain tu
where *:"tu setinterleaves ((uu, t), insert tick (ev ` C))" and **:"rr setinterleaves ((tu, u), insert tick (ev ` C))"
using Sync.sym by blast
with a(1,2) 8 show ?case
apply(exI tu, exI u, exI rr, exI vv, simp) using Sync.sym front_tickFree_Nil by blast
next
case (9 t u)
with interleave_assoc_1[OF 9(3), of rr uu] obtain uv
where *:"uv setinterleaves ((u, uu), insert tick (ev ` C))" and **:"rr setinterleaves ((t, uv), insert tick (ev ` C))"
by blast
with a(1,2) 9 show ?case
apply(exI t, exI uv, exI rr, exI vv, simp) using Sync.sym front_tickFree_Nil by blast
qed
qed
lemma sync_assoc_oneside: "((P \<lbrakk>C\<rbrakk> Q) \<lbrakk>C\<rbrakk> S) \<le> (P \<lbrakk>C\<rbrakk> (Q \<lbrakk>C\<rbrakk> S))"
proof(unfold le_ref_def, rule conjI, goal_cases D F)
case D
then show ?case using sync_assoc_oneside_D by assumption
next
case F
then show ?case
unfolding F_sync proof(rule Un_least, goal_cases FF DD)
case (FF)
then show ?case
proof(rule subsetI, simp add:D_sync T_sync split_paired_all, elim exE conjE disjE, goal_cases)
case (1 r Xtuv t uv Xt Xuv u v Xu Xv)
with interleave_assoc_2[OF 1(6), OF 1(2)] obtain tu
where *:"tu setinterleaves ((t, u), insert tick (ev ` C))"
and **:"r setinterleaves ((tu, v), insert tick (ev ` C))" by blast
with 1(1,3,4,5,7) show ?case
apply(rule_tac disjI1, exI tu, exI v, exI "(Xt \<union> Xu) \<inter> insert tick (ev ` C) \<union> Xt \<inter> Xu", intro conjI disjI1, simp_all)
by blast+
next
case (2 r Xtuv t uv Xt Xuv u v uv1 uv2)
with interleave_append_sym [of r t _ uv1 uv2] obtain t1 t2 r1 r2
where a1:"t = t1 @ t2" and a2:"r = r1 @ r2" and
a3:"r1 setinterleaves ((t1, uv1), insert tick (ev ` C))" and
a4:"r2 setinterleaves ((t2, uv2), insert tick (ev ` C))" by metis
with interleave_assoc_2[OF 2(6), OF a3] obtain tu
where *:"tu setinterleaves ((t1, u), insert tick (ev ` C))"
and **:"r1 setinterleaves ((tu, v), insert tick (ev ` C))" by blast
from 2(1) a1 obtain Xt1 where ***:"(t1, Xt1) \<in> F P" using is_processT3_SR by blast
from 2 a1 a2 a3 a4 * ** *** show ?case
apply(rule_tac disjI2, exI tu, exI v, exI r1, exI r2, intro conjI disjI1, simp_all)
apply (metis front_tickFree_charn front_tickFree_mono ftf_syn is_processT2)
apply (metis equals0D ftf_syn1 ftf_syn21 insertI1 tickFree_def)
using F_T Sync.sym front_tickFree_Nil by blast
next
case (3 r Xtuv t uv Xt Xuv v u uv1 uv2)
with interleave_append_sym [of r t _ uv1 uv2] obtain t1 t2 r1 r2
where a1:"t = t1 @ t2" and a2:"r = r1 @ r2" and
a3:"r1 setinterleaves ((t1, uv1), insert tick (ev ` C))" and
a4:"r2 setinterleaves ((t2, uv2), insert tick (ev ` C))" by metis
with interleave_assoc_2[OF _ a3, of u v] obtain tu
where *:"tu setinterleaves ((t1, u), insert tick (ev ` C))"
and **:"r1 setinterleaves ((tu, v), insert tick (ev ` C))" using "3"(6) Sync.sym by blast
from 3(1) a1 obtain Xt1 where ***:"(t1, Xt1) \<in> F P" using is_processT3_SR by blast
from 3 a1 a2 a3 a4 * ** *** show ?case
apply(rule_tac disjI2, exI v, exI tu, exI r1, exI r2, intro conjI, simp_all)
apply (metis front_tickFree_charn front_tickFree_mono ftf_syn is_processT2)
apply (metis equals0D ftf_syn1 ftf_syn21 insertI1 tickFree_def)
using Sync.sym F_T by blast+
next
case (4 r Xtuv t uv Xt Xuv u v uv1 uv2)
with interleave_assoc_2[OF 4(6), of r t] obtain tu
where *:"tu setinterleaves ((t, u), insert tick (ev ` C))"
and **:"r setinterleaves ((tu, v), insert tick (ev ` C))" by auto
from 4 * ** show ?case
apply(rule_tac disjI2, exI tu, exI v, exI r, exI "[]", intro conjI, simp_all)
by (metis "4"(4) F_T Sync.sym append.right_neutral)
next
case (5 r Xtuv t uv Xt Xuv v u uv1 uv2)
from 5(2,7,5,6) interleave_assoc_2[of uv u _ v r t] obtain tu
where *:"tu setinterleaves ((t, u), insert tick (ev ` C))"
and **:"r setinterleaves ((tu, v), insert tick (ev ` C))" by (metis Sync.sym append_Nil2)
from 5 * ** show ?case
apply(rule_tac disjI2, exI v, exI tu, exI r, exI "[]", intro conjI, simp_all)
using Sync.sym F_T by blast+
qed
next
case (DD)
then show ?case
using sync_assoc_oneside_D[of P C Q S] by (simp add:D_sync Collect_mono_iff sup.coboundedI2)
qed
qed
lemma sync_assoc: "((P \<lbrakk>C\<rbrakk> Q) \<lbrakk>C\<rbrakk> S) = (P \<lbrakk>C\<rbrakk> (Q \<lbrakk>C\<rbrakk> S))"
by (metis antisym_conv sync_assoc_oneside sync_commute)
subsection\<open> Derivative Operators laws \<close>
lemma mono_Inter_ref: " \<lbrakk>P \<sqsubseteq> P'; Q \<sqsubseteq> Q'\<rbrakk> \<Longrightarrow> (P ||| Q) \<sqsubseteq> (P'||| Q')"
by (rule mono_sync_ref)
lemma mono_Par_ref: "\<lbrakk>P \<sqsubseteq> P'; Q \<sqsubseteq> Q'\<rbrakk> \<Longrightarrow> (P || Q) \<sqsubseteq> (P' || Q')"
by (rule mono_sync_ref)
lemma Inter_commute: "(P ||| Q) = (Q ||| P)"
by(rule sync_commute)
lemma par_comm: "(P || Q) =(Q || P)"
by(rule sync_commute)
lemma par_id:"(P || P) = P"
oops
lemma Inter_skip1: "(P ||| SKIP) = P"
proof(auto simp:Process_eq_spec D_sync F_sync F_SKIP D_SKIP T_SKIP, goal_cases)
case (1 a t X Y)
then have aa:"(X \<union> Y) \<inter> {tick} \<union> X \<inter> Y \<subseteq> X" by (simp add: Int_Un_distrib2)
have "front_tickFree t" using "1"(1) is_processT2 by blast
with 1 EmptyLeftSync[of a "{tick}" t] have bb:"a=t"
using Sync.sym by blast
from 1(1) aa bb show ?case using is_processT4 by blast
next
case (2 a t X Y)
have "front_tickFree t" using 2(1) is_processT2 by blast
with 2 TickLeftSync[of "{tick}" t a] have bb:"a=t \<and> last t=tick"
using Sync.sym by blast
with 2 have "a \<in> T P \<and> last a = tick" using F_T by blast
with tick_T_F[of "butlast a" P "(X \<union> Y) \<inter> {tick} \<union> X \<inter> Y"] show ?case
by (metis "2"(2) EmptyLeftSync bb self_append_conv2 snoc_eq_iff_butlast)
next
case (3 b t r v)
with 3 EmptyLeftSync[of r "{tick}" t] have bb:"r=t"
using Sync.sym by blast
with 3 show ?case using is_processT by blast
next
case (4 b t r v)
with TickLeftSync[of "{tick}" t r, simplified] have bb:"r=t"
using D_imp_front_tickFree Sync.sym by blast
with 4 show ?case using is_processT by blast
next
case (5 b t r)
with EmptyLeftSync[of r "{tick}" t] have bb:"r=t"
using Sync.sym by blast
with 5 show ?case using is_processT by blast
next
case (6 b t r)
with TickLeftSync[of "{tick}" t r, simplified] have bb:"r=t"
using D_imp_front_tickFree Sync.sym by blast
with 6 show ?case using is_processT by blast
next
case (7 a b)
then show ?case
apply(cases "tickFree a")
apply(rule_tac x=a in exI, rule_tac x="[]" in exI, rule_tac x=b in exI, simp)
apply(rule_tac x="b - {tick}" in exI, simp, intro conjI)
using emptyLeftSelf[of a "{tick}"] Sync.sym tickFree_def apply fastforce
apply blast
apply(rule_tac x=a in exI, rule_tac x="[tick]" in exI, rule_tac x=b in exI, simp, rule conjI)
subgoal proof -
assume a1: "\<not> tickFree a"
assume a2: "(a, b) \<in> F P"
then obtain ees :: "'a event list \<Rightarrow> 'a event list" where
f3: "a = ees a @ [tick]"
using a1 by (meson is_processT2 nonTickFree_n_frontTickFree)
then have "ees a setinterleaves (([], ees a), {tick})"
using a2 by (metis (no_types) DiffD2 Diff_insert_absorb emptyLeftSelf front_tickFree_implies_tickFree is_processT2 tickFree_def)
then show "a setinterleaves ((a, [tick]), {tick})"
using f3 by (metis (no_types) Sync.sym addSyncEmpty insertI1)
qed
apply(rule_tac x="b - {tick}" in exI) by blast
next
case (8 t r v)
with EmptyLeftSync[of r "{tick}" t] have bb:"r=t"
using Sync.sym by blast
with 8 show ?case using is_processT by blast
next
case (9 t r v)
with TickLeftSync[of "{tick}" t r, simplified] have bb:"r=t"
using D_imp_front_tickFree Sync.sym by blast
with 9 show ?case using is_processT by blast
next
case (10 t r)
with EmptyLeftSync[of r "{tick}" t] have bb:"r=t"
using Sync.sym by blast
with 10 show ?case by simp
next
case (11 t r)
with TickLeftSync[of "{tick}" t r, simplified] have bb:"r=t"
using D_imp_front_tickFree Sync.sym by blast
with 11 show ?case by simp
next
case (12 x)
then show ?case
proof(cases "tickFree x")
case True
with 12 show ?thesis
apply(rule_tac x=x in exI, rule_tac x="[]" in exI, rule_tac x=x in exI, rule_tac x="[]" in exI, simp)
by (metis Sync.sym emptyLeftSelf singletonD tickFree_def)
next
case False
with 12 show ?thesis
apply(rule_tac x="butlast x" in exI, rule_tac x="[]" in exI, rule_tac x="butlast x" in exI, rule_tac x="[tick]" in exI, simp, intro conjI)
apply (metis D_imp_front_tickFree append_butlast_last_id front_tickFree_mono list.distinct(1) tickFree_Nil)
using NT_ND T_nonTickFree_imp_decomp apply fastforce
apply (metis NT_ND Sync.sym append_T_imp_tickFree append_butlast_last_id emptyLeftSelf
list.distinct(1) singletonD tickFree_Nil tickFree_def)
using D_imp_front_tickFree is_processT9 nonTickFree_n_frontTickFree by fastforce
qed
qed
lemma Inter_stop_seq_stop: "(P ||| STOP) = (P `;` STOP)"
proof (auto simp add:Process_eq_spec D_sync F_sync F_STOP D_STOP T_STOP F_seq D_seq, goal_cases)
case (1 a t X Y)
have "insert tick ((X \<union> Y) \<inter> {tick} \<union> X \<inter> Y) \<subseteq> X \<union> {tick}" by blast
with 1 show ?case
by (metis (mono_tags, lifting) EmptyLeftSync Sync.sym is_processT is_processT5_S2 no_Trace_implies_no_Failure)
next
case (2 a t X)
then show ?case using EmptyLeftSync Sync.sym tickFree_def by fastforce
next
case (3 b t r v)
then show ?case using EmptyLeftSync Sync.sym tickFree_def by blast
next
case (4 t r v)
then show ?case using EmptyLeftSync Sync.sym by blast
next
case (5 b t r)
then show ?case by (metis EmptyLeftSync Sync.sym is_processT)
next
case (6 t r)
then show ?case using EmptyLeftSync Sync.sym tickFree_def by fastforce
next
case (7 a b)
then show ?case
apply(rule_tac x=a in exI, rule_tac x="b" in exI, intro conjI)
apply (meson insert_iff is_processT subsetI)
using emptyLeftSelf[of a "{tick}"] Sync.sym tickFree_def apply fastforce
by blast
next
case (8 a b)
then show ?case
apply(rule_tac x=a in exI, rule_tac x="b-{tick}" in exI, intro conjI)
apply (meson NF_NT is_processT6)
using emptyLeftSelf[of a "{tick}"] Sync.sym tickFree_def
apply (metis append_T_imp_tickFree list.distinct(1) singletonD)
by blast
next
case (9 b t1 t2)
hence "t1 setinterleaves ((t1, []), {tick})"
using emptyLeftSelf[of t1 "{tick}"] Sync.sym tickFree_def by fastforce
with 9(1)[THEN spec, THEN spec, of t1 t1, simplified] have False
using "9"(2) "9"(3) "9"(4) by blast
then show ?case ..
next
case (10 t r v)
then show ?case
apply(rule_tac x=r in exI, rule_tac x="v" in exI, auto)
using EmptyLeftSync Sync.sym by blast
next
case (11 t r)
then show ?case
apply(rule_tac x=r in exI, rule_tac x="[]" in exI, simp)
using EmptyLeftSync Sync.sym tickFree_def by fastforce
next
case (12 t1 t2)
then show ?case
apply(rule_tac x=t1 in exI, rule_tac x=t1 in exI, rule_tac x=t2 in exI, simp)
by (metis Sync.sym emptyLeftSelf singletonD tickFree_def)
qed
lemma par_ndet_distrib2: "((P \<sqinter> Q) || M) = ((P || M) \<sqinter> (Q || M))"
by (rule par_int_ndet_distrib)
lemma par_stop: "P \<noteq> \<bottom> \<Longrightarrow> (P || STOP) = STOP"
proof(auto simp add:Process_eq_spec,
simp_all add:D_sync F_sync F_STOP D_STOP T_STOP is_processT2 D_imp_front_tickFree F_UU D_UU,
goal_cases)
case (1 a b aa ba)
then show ?case
proof(elim exE conjE disjE, goal_cases)
case (1 t X Y)
then show ?case
proof(cases t)
case Nil then show ?thesis using "1"(4) EmptyLeftSync by blast
next
case (Cons a list)
with 1 show ?thesis by (simp split:if_splits, cases a, simp_all)
qed
next
case (2 t r v)
hence "[] \<in> D P"
proof(cases t)
case Nil with 2 show ?thesis by simp
next
case (Cons a list)
with 2 show ?thesis by (simp split:if_splits, cases a, simp_all)
qed
with 2(1,2) show ?case using NF_ND is_processT7 by fastforce
next
case (3 t r v)
hence "[] \<in> D P"
proof(cases t)
case Nil with 3 show ?thesis by simp
next
case (Cons a list)
with 3 show ?thesis by (simp split:if_splits, cases a, simp_all)
qed
with 3(1,2) show ?case using NF_ND is_processT7 by fastforce
qed
next
case (2 a b aa ba)
then show ?case
apply(rule_tac disjI1, rule_tac x="[]" in exI, rule_tac x="{}" in exI, simp, rule_tac conjI)
using is_processT apply blast
apply(rule_tac x=ba in exI, rule_tac set_eqI, rule_tac iffI)
using event_set apply blast
by fast
next
case (3 a b x)
then show ?case
proof(elim exE conjE disjE, goal_cases)
case (1 t r v)
hence "[] \<in> D P"
proof(cases t)
case Nil with 1 show ?thesis by simp
next
case (Cons a list)
with 1 show ?thesis by (simp split:if_splits, cases a, simp_all)
qed
with 3(1,2) show ?case using NF_ND is_processT7 by fastforce
next
case (2 t r v)
hence "[] \<in> D P"
proof(cases t)
case Nil with 2 show ?thesis by simp
next
case (Cons a list)
with 2 show ?thesis by (simp split:if_splits, cases a, simp_all)
qed
with 3(1,2) show ?case using NF_ND is_processT7 by fastforce
qed
next
case (4 x a b)
then show ?case
proof(elim exE conjE disjE, goal_cases)
case (1 t X Y)
then show ?case
proof(cases t)
case Nil then show ?thesis using "1"(4) EmptyLeftSync by blast
next
case (Cons a list)
with 1 show ?thesis by (simp split:if_splits, cases a, simp_all)
qed
next
case (2 t r v)
hence "[] \<in> D P"
proof(cases t)
case Nil with 2 show ?thesis by simp
next
case (Cons a list)
with 2 show ?thesis by (simp split:if_splits, cases a, simp_all)
qed
with 2(1,2) show ?case using NF_ND is_processT7 by fastforce
next
case (3 t r v)
hence "[] \<in> D P"
proof(cases t)
case Nil with 3 show ?thesis by simp
next
case (Cons a list)
with 3 show ?thesis by (simp split:if_splits, cases a, simp_all)
qed
with 3(1,2) show ?case using NF_ND is_processT7 by fastforce
qed
next
case (5 x a b)
then show ?case
apply(rule_tac disjI1, rule_tac x="[]" in exI, rule_tac x="{}" in exI, simp, rule_tac conjI)
using is_processT apply blast
apply(rule_tac x=b in exI, rule_tac set_eqI, rule_tac iffI)
using event_set apply blast
by fast
next
case (6 x xa)
then show ?case
proof(elim exE conjE disjE, goal_cases)
case (1 t r v)
hence "[] \<in> D P"
proof(cases t)
case Nil with 1 show ?thesis by simp
next
case (Cons a list)
with 1 show ?thesis by (simp split:if_splits, cases a, simp_all)
qed
with 6(1,2) show ?case using NF_ND is_processT7 by fastforce
next
case (2 t r v)
hence "[] \<in> D P"
proof(cases t)
case Nil with 2 show ?thesis by simp
next
case (Cons a list)
with 2 show ?thesis by (simp split:if_splits, cases a, simp_all)
qed
with 6(1,2) show ?case using NF_ND is_processT7 by fastforce
qed
qed
lemma par_assoc: "((P || Q) || S) = (P || (Q || S))"
by (rule sync_assoc)
lemma prefix_par: "((a \<rightarrow> P) || (a \<rightarrow> Q)) = (a \<rightarrow> (P || Q))"
by (simp add: prefix_par_Int2)
lemma prefix_Inter: "((a \<rightarrow> P) ||| (b \<rightarrow> Q)) = ((a \<rightarrow> (P ||| (b \<rightarrow> Q))) \<box> (b \<rightarrow> ((a \<rightarrow> P) ||| Q)))"
using mprefix_Par_Int[of "{b}" "{}" "{a}" "\<lambda>x. P" "\<lambda>x. Q", simplified]
unfolding write0_def by assumption
section\<open>Multiple Non Deterministic Operator Laws\<close>
lemma Mprefix_refines_Mndet: "(\<sqinter> x \<in> A \<rightarrow> P x) \<le> (\<box> x \<in> A \<rightarrow> P x)"
proof(cases "A = {}")
case True
then show ?thesis by (simp add: Mprefix_STOP)
next
case False
then show ?thesis
by (auto simp add:le_ref_def D_Mprefix write0_def F_Mprefix F_mndet D_mndet)
qed
lemma mndet_refine_FD: "a\<in>A \<Longrightarrow> (\<sqinter> x \<in> A \<rightarrow> (P x)) \<le> (a \<rightarrow> (P a))"
by(subgoal_tac "A \<noteq> {}", auto simp:le_ref_def D_mndet F_mndet)
lemma mndet_subset_FD:"A \<noteq> {} \<Longrightarrow> (\<sqinter>xa\<in> A \<union> B \<rightarrow> P) \<le> (\<sqinter>xa\<in> A \<rightarrow> P)"
by (simp add: D_mndet F_mndet le_ref_def)
lemma mono_mndet_FD[simp]: "\<forall>x \<in> A. P x \<le> P' x \<Longrightarrow> (\<sqinter> x \<in> A \<rightarrow> (P x)) \<le> (\<sqinter> x \<in> A \<rightarrow> (P' x))"
apply(cases "A \<noteq> {}", auto simp:le_ref_def D_mndet F_mndet)
by(meson contra_subsetD le_ref_def mono_write0_FD)+
lemma mndet_FD_subset : "A \<noteq> {} \<Longrightarrow> A \<subseteq> B \<Longrightarrow> ((\<sqinter> x \<in> B \<rightarrow> P x) \<le> (\<sqinter> x \<in> A \<rightarrow> P x))"
using mndet_distrib [of "B - A" "A"]
by (metis bot.extremum_uniqueI mndet_distrib mono_ndet_FD_right order_refl sup.absorb_iff1)
section \<open> Infra-structure for Communication Primitives \<close>
lemma read_read_sync:
assumes contained: "(\<And>y. c y \<in> C)"
shows "((c `?` x \<rightarrow> P x)\<lbrakk> C \<rbrakk>(c `?` x \<rightarrow> Q x)) = (c `?` x \<rightarrow> ((P x) \<lbrakk> C \<rbrakk> (Q x)))"
proof -
have A: "range c \<subseteq> C" by(insert contained, auto)
show ?thesis
by(auto simp: read_def o_def Set.Int_absorb2 mprefix_Par_distr A)
qed
lemma read_read_nonsync:
" \<lbrakk>\<And>y. c y \<notin> C; \<And>y. d y \<in> C\<rbrakk> \<Longrightarrow>
((c `?` x \<rightarrow> (P x)) \<lbrakk>C\<rbrakk> (d `?` x \<rightarrow>(Q x))) = (c `?` x \<rightarrow> ((P x) \<lbrakk>C\<rbrakk> (d `?` x \<rightarrow>(Q x))))"
apply (auto simp add: read_def o_def, goal_cases)
apply(insert mprefix_Par_Int_distr[of "range c" C "range d" "\<lambda>x. (Q (inv d x))" "\<lambda>x. (P (inv c x))"])
apply(subst (1 2) sync_commute)
by auto
lemma write_read_sync:
assumes contained: "\<And>y. c y \<in> C"
assumes is_construct: "inj c"
shows "((c `!` a \<rightarrow> P) \<lbrakk> C \<rbrakk> (c `?` x \<rightarrow> Q x)) = (c `!` a \<rightarrow> (P \<lbrakk> C \<rbrakk> (Q a)))"
proof -
have A : "range c \<subseteq> C" by(insert contained, auto)
have B : " {c a} \<inter> range c \<inter> C = {c a}" by(insert contained, auto)
show ?thesis
apply(simp add: read_def write_def o_def, subst mprefix_Par_distr)
by (auto simp: A B contained is_construct mprefix_Par_distr Hilbert_Choice.inv_f_f mprefix_singl)
qed
lemma write_read_nonsync:
"\<lbrakk>d a \<notin> C; \<And>y. c y \<in> C\<rbrakk> \<Longrightarrow> ((d `!` a \<rightarrow> P)\<lbrakk>C\<rbrakk>(c `?` x \<rightarrow> Q x)) = (d `!` a \<rightarrow> (P\<lbrakk>C\<rbrakk>(c `?` x \<rightarrow> Q x)))"
apply (auto simp add: read_def o_def write_def, goal_cases)
apply(insert mprefix_Par_Int_distr[of "{d a}" C "range c" "\<lambda>x. (Q (inv c x))" "\<lambda>x. P"])
apply(subst (1 2) sync_commute)
by auto
lemma write0_read_nonsync:
"\<lbrakk>d \<in> C; \<And>y. c y \<notin> C\<rbrakk> \<Longrightarrow> ((d \<rightarrow> P)\<lbrakk>C\<rbrakk>(c `?` x \<rightarrow> Q x)) = (c `?` x \<rightarrow> ((d \<rightarrow> P)\<lbrakk>C\<rbrakk>Q x))"
apply(simp add: read_def o_def)
apply(subst mprefix_singl[symmetric])
apply(subst mprefix_singl[symmetric])
apply(subst mprefix_Par_Int_distr)
by auto
lemma write0_write_nonsync:
"\<lbrakk>d a \<notin> C; c \<in> C\<rbrakk> \<Longrightarrow> ((c \<rightarrow> Q)\<lbrakk>C\<rbrakk>(d `!` a \<rightarrow> P)) = (d `!` a \<rightarrow> ((c \<rightarrow> Q)\<lbrakk>C\<rbrakk>P))"
apply(simp add: write_def )
apply(subst mprefix_singl[symmetric])
apply(subst mprefix_singl[symmetric])
by (auto simp: mprefix_Par_Int_distr)
(* ************************************************************************* *)
(* *)
(* Setup for rewriting *)
(* *)
(* ************************************************************************* *)
lemmas sync_rules = prefix_par_Int2
read_read_sync read_read_nonsync
write_read_sync write_read_nonsync
write0_read_nonsync
write0_write_nonsync
lemmas hide_rules = no_hide_read no_hide_write hide_write hide_write0
lemmas mono_rules = mono_read_ref mono_write_ref mono_write0_ref
end
diff --git a/thys/Incompleteness/Coding.thy b/thys/Incompleteness/Coding.thy
--- a/thys/Incompleteness/Coding.thy
+++ b/thys/Incompleteness/Coding.thy
@@ -1,859 +1,860 @@
chapter\<open>De Bruijn Syntax, Quotations, Codes, V-Codes\<close>
theory Coding
imports SyntaxN
begin
declare fresh_Nil [iff]
section \<open>de Bruijn Indices (locally-nameless version)\<close>
nominal_datatype dbtm = DBZero | DBVar name | DBInd nat | DBEats dbtm dbtm
nominal_datatype dbfm =
DBMem dbtm dbtm
| DBEq dbtm dbtm
| DBDisj dbfm dbfm
| DBNeg dbfm
| DBEx dbfm
declare dbtm.supp [simp]
declare dbfm.supp [simp]
fun lookup :: "name list \<Rightarrow> nat \<Rightarrow> name \<Rightarrow> dbtm"
where
"lookup [] n x = DBVar x"
| "lookup (y # ys) n x = (if x = y then DBInd n else (lookup ys (Suc n) x))"
lemma fresh_imp_notin_env: "atom name \<sharp> e \<Longrightarrow> name \<notin> set e"
by (metis List.finite_set fresh_finite_set_at_base fresh_set)
lemma lookup_notin: "x \<notin> set e \<Longrightarrow> lookup e n x = DBVar x"
by (induct e arbitrary: n) auto
lemma lookup_in:
"x \<in> set e \<Longrightarrow> \<exists>k. lookup e n x = DBInd k \<and> n \<le> k \<and> k < n + length e"
apply (induct e arbitrary: n)
apply (auto intro: Suc_leD)
apply (metis Suc_leD add_Suc_right add_Suc_shift)
done
lemma lookup_fresh: "x \<sharp> lookup e n y \<longleftrightarrow> y \<in> set e \<or> x \<noteq> atom y"
by (induct arbitrary: n rule: lookup.induct) (auto simp: pure_fresh fresh_at_base)
lemma lookup_eqvt[eqvt]: "(p \<bullet> lookup xs n x) = lookup (p \<bullet> xs) (p \<bullet> n) (p \<bullet> x)"
by (induct xs arbitrary: n) (simp_all add: permute_pure)
lemma lookup_inject [iff]: "(lookup e n x = lookup e n y) \<longleftrightarrow> x = y"
apply (induct e n x arbitrary: y rule: lookup.induct, force, simp)
by (metis Suc_n_not_le_n dbtm.distinct(7) dbtm.eq_iff(3) lookup_in lookup_notin)
nominal_function trans_tm :: "name list \<Rightarrow> tm \<Rightarrow> dbtm"
where
"trans_tm e Zero = DBZero"
| "trans_tm e (Var k) = lookup e 0 k"
| "trans_tm e (Eats t u) = DBEats (trans_tm e t) (trans_tm e u)"
by (auto simp: eqvt_def trans_tm_graph_aux_def) (metis tm.strong_exhaust)
nominal_termination (eqvt)
by lexicographic_order
lemma fresh_trans_tm_iff [simp]: "i \<sharp> trans_tm e t \<longleftrightarrow> i \<sharp> t \<or> i \<in> atom ` set e"
by (induct t rule: tm.induct, auto simp: lookup_fresh fresh_at_base)
lemma trans_tm_forget: "atom i \<sharp> t \<Longrightarrow> trans_tm [i] t = trans_tm [] t"
by (induct t rule: tm.induct, auto simp: fresh_Pair)
nominal_function (invariant "\<lambda>(xs, _) y. atom ` set xs \<sharp>* y")
trans_fm :: "name list \<Rightarrow> fm \<Rightarrow> dbfm"
where
"trans_fm e (Mem t u) = DBMem (trans_tm e t) (trans_tm e u)"
| "trans_fm e (Eq t u) = DBEq (trans_tm e t) (trans_tm e u)"
| "trans_fm e (Disj A B) = DBDisj (trans_fm e A) (trans_fm e B)"
| "trans_fm e (Neg A) = DBNeg (trans_fm e A)"
| "atom k \<sharp> e \<Longrightarrow> trans_fm e (Ex k A) = DBEx (trans_fm (k#e) A)"
+supply [[simproc del: defined_all]]
apply(simp add: eqvt_def trans_fm_graph_aux_def)
apply(erule trans_fm_graph.induct)
using [[simproc del: alpha_lst]]
apply(auto simp: fresh_star_def)
apply(rule_tac y=b and c=a in fm.strong_exhaust)
apply(auto simp: fresh_star_def)
apply(erule_tac c=ea in Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: fresh_star_Pair perm_supp_eq)
apply (simp add: fresh_star_def)
done
nominal_termination (eqvt)
by lexicographic_order
lemma fresh_trans_fm [simp]: "i \<sharp> trans_fm e A \<longleftrightarrow> i \<sharp> A \<or> i \<in> atom ` set e"
by (nominal_induct A avoiding: e rule: fm.strong_induct, auto simp: fresh_at_base)
abbreviation DBConj :: "dbfm \<Rightarrow> dbfm \<Rightarrow> dbfm"
where "DBConj t u \<equiv> DBNeg (DBDisj (DBNeg t) (DBNeg u))"
lemma trans_fm_Conj [simp]: "trans_fm e (Conj A B) = DBConj (trans_fm e A) (trans_fm e B)"
by (simp add: Conj_def)
lemma trans_tm_inject [iff]: "(trans_tm e t = trans_tm e u) \<longleftrightarrow> t = u"
proof (induct t arbitrary: e u rule: tm.induct)
case Zero show ?case
apply (cases u rule: tm.exhaust, auto)
apply (metis dbtm.distinct(1) dbtm.distinct(3) lookup_in lookup_notin)
done
next
case (Var i) show ?case
apply (cases u rule: tm.exhaust, auto)
apply (metis dbtm.distinct(1) dbtm.distinct(3) lookup_in lookup_notin)
apply (metis dbtm.distinct(10) dbtm.distinct(11) lookup_in lookup_notin)
done
next
case (Eats tm1 tm2) thus ?case
apply (cases u rule: tm.exhaust, auto)
apply (metis dbtm.distinct(12) dbtm.distinct(9) lookup_in lookup_notin)
done
qed
lemma trans_fm_inject [iff]: "(trans_fm e A = trans_fm e B) \<longleftrightarrow> A = B"
proof (nominal_induct A avoiding: e B rule: fm.strong_induct)
case (Mem tm1 tm2) thus ?case
by (rule fm.strong_exhaust [where y=B and c=e]) (auto simp: fresh_star_def)
next
case (Eq tm1 tm2) thus ?case
by (rule fm.strong_exhaust [where y=B and c=e]) (auto simp: fresh_star_def)
next
case (Disj fm1 fm2) show ?case
by (rule fm.strong_exhaust [where y=B and c=e]) (auto simp: Disj fresh_star_def)
next
case (Neg fm) show ?case
by (rule fm.strong_exhaust [where y=B and c=e]) (auto simp: Neg fresh_star_def)
next
case (Ex name fm)
thus ?case using [[simproc del: alpha_lst]]
proof (cases rule: fm.strong_exhaust [where y=B and c="(e, name)"], simp_all add: fresh_star_def)
fix name'::name and fm'::fm
assume name': "atom name' \<sharp> (e, name)"
assume "atom name \<sharp> fm' \<or> name = name'"
thus "(trans_fm (name # e) fm = trans_fm (name' # e) fm') = ([[atom name]]lst. fm = [[atom name']]lst. fm')"
(is "?lhs = ?rhs")
proof (rule disjE)
assume "name = name'"
thus "?lhs = ?rhs"
by (metis fresh_Pair fresh_at_base(2) name')
next
assume name: "atom name \<sharp> fm'"
have eq1: "(name \<leftrightarrow> name') \<bullet> trans_fm (name' # e) fm' = trans_fm (name' # e) fm'"
by (simp add: flip_fresh_fresh name)
have eq2: "(name \<leftrightarrow> name') \<bullet> ([[atom name']]lst. fm') = [[atom name']]lst. fm'"
by (rule flip_fresh_fresh) (auto simp: Abs_fresh_iff name)
show "?lhs = ?rhs" using name' eq1 eq2 Ex(1) Ex(3) [of "name#e" "(name \<leftrightarrow> name') \<bullet> fm'"]
by (simp add: flip_fresh_fresh) (metis Abs1_eq(3))
qed
qed
qed
lemma trans_fm_perm:
assumes c: "atom c \<sharp> (i,j,A,B)"
and t: "trans_fm [i] A = trans_fm [j] B"
shows "(i \<leftrightarrow> c) \<bullet> A = (j \<leftrightarrow> c) \<bullet> B"
proof -
have c_fresh1: "atom c \<sharp> trans_fm [i] A"
using c by (auto simp: supp_Pair)
moreover
have i_fresh: "atom i \<sharp> trans_fm [i] A"
by auto
moreover
have c_fresh2: "atom c \<sharp> trans_fm [j] B"
using c by (auto simp: supp_Pair)
moreover
have j_fresh: "atom j \<sharp> trans_fm [j] B"
by auto
ultimately have "((i \<leftrightarrow> c) \<bullet> (trans_fm [i] A)) = ((j \<leftrightarrow> c) \<bullet> trans_fm [j] B)"
by (simp only: flip_fresh_fresh t)
then have "trans_fm [c] ((i \<leftrightarrow> c) \<bullet> A) = trans_fm [c] ((j \<leftrightarrow> c) \<bullet> B)"
by simp
then show "(i \<leftrightarrow> c) \<bullet> A = (j \<leftrightarrow> c) \<bullet> B" by simp
qed
section\<open>Characterising the Well-Formed de Bruijn Formulas\<close>
subsection\<open>Well-Formed Terms\<close>
inductive wf_dbtm :: "dbtm \<Rightarrow> bool"
where
Zero: "wf_dbtm DBZero"
| Var: "wf_dbtm (DBVar name)"
| Eats: "wf_dbtm t1 \<Longrightarrow> wf_dbtm t2 \<Longrightarrow> wf_dbtm (DBEats t1 t2)"
equivariance wf_dbtm
inductive_cases Zero_wf_dbtm [elim!]: "wf_dbtm DBZero"
inductive_cases Var_wf_dbtm [elim!]: "wf_dbtm (DBVar name)"
inductive_cases Ind_wf_dbtm [elim!]: "wf_dbtm (DBInd i)"
inductive_cases Eats_wf_dbtm [elim!]: "wf_dbtm (DBEats t1 t2)"
declare wf_dbtm.intros [intro]
lemma wf_dbtm_imp_is_tm:
assumes "wf_dbtm x"
shows "\<exists>t::tm. x = trans_tm [] t"
using assms
proof (induct rule: wf_dbtm.induct)
case Zero thus ?case
by (metis trans_tm.simps(1))
next
case (Var i) thus ?case
by (metis lookup.simps(1) trans_tm.simps(2))
next
case (Eats dt1 dt2) thus ?case
by (metis trans_tm.simps(3))
qed
lemma wf_dbtm_trans_tm: "wf_dbtm (trans_tm [] t)"
by (induct t rule: tm.induct) auto
theorem wf_dbtm_iff_is_tm: "wf_dbtm x \<longleftrightarrow> (\<exists>t::tm. x = trans_tm [] t)"
by (metis wf_dbtm_imp_is_tm wf_dbtm_trans_tm)
nominal_function abst_dbtm :: "name \<Rightarrow> nat \<Rightarrow> dbtm \<Rightarrow> dbtm"
where
"abst_dbtm name i DBZero = DBZero"
| "abst_dbtm name i (DBVar name') = (if name = name' then DBInd i else DBVar name')"
| "abst_dbtm name i (DBInd j) = DBInd j"
| "abst_dbtm name i (DBEats t1 t2) = DBEats (abst_dbtm name i t1) (abst_dbtm name i t2)"
apply (simp add: eqvt_def abst_dbtm_graph_aux_def, auto)
apply (metis dbtm.exhaust)
done
nominal_termination (eqvt)
by lexicographic_order
nominal_function subst_dbtm :: "dbtm \<Rightarrow> name \<Rightarrow> dbtm \<Rightarrow> dbtm"
where
"subst_dbtm u i DBZero = DBZero"
| "subst_dbtm u i (DBVar name) = (if i = name then u else DBVar name)"
| "subst_dbtm u i (DBInd j) = DBInd j"
| "subst_dbtm u i (DBEats t1 t2) = DBEats (subst_dbtm u i t1) (subst_dbtm u i t2)"
by (auto simp: eqvt_def subst_dbtm_graph_aux_def) (metis dbtm.exhaust)
nominal_termination (eqvt)
by lexicographic_order
lemma fresh_iff_non_subst_dbtm: "subst_dbtm DBZero i t = t \<longleftrightarrow> atom i \<sharp> t"
by (induct t rule: dbtm.induct) (auto simp: pure_fresh fresh_at_base(2))
lemma lookup_append: "lookup (e @ [i]) n j = abst_dbtm i (length e + n) (lookup e n j)"
by (induct e arbitrary: n) (auto simp: fresh_Cons)
lemma trans_tm_abs: "trans_tm (e@[name]) t = abst_dbtm name (length e) (trans_tm e t)"
by (induct t rule: tm.induct) (auto simp: lookup_notin lookup_append)
subsection\<open>Well-Formed Formulas\<close>
nominal_function abst_dbfm :: "name \<Rightarrow> nat \<Rightarrow> dbfm \<Rightarrow> dbfm"
where
"abst_dbfm name i (DBMem t1 t2) = DBMem (abst_dbtm name i t1) (abst_dbtm name i t2)"
| "abst_dbfm name i (DBEq t1 t2) = DBEq (abst_dbtm name i t1) (abst_dbtm name i t2)"
| "abst_dbfm name i (DBDisj A1 A2) = DBDisj (abst_dbfm name i A1) (abst_dbfm name i A2)"
| "abst_dbfm name i (DBNeg A) = DBNeg (abst_dbfm name i A)"
| "abst_dbfm name i (DBEx A) = DBEx (abst_dbfm name (i+1) A)"
apply (simp add: eqvt_def abst_dbfm_graph_aux_def, auto)
apply (metis dbfm.exhaust)
done
nominal_termination (eqvt)
by lexicographic_order
nominal_function subst_dbfm :: "dbtm \<Rightarrow> name \<Rightarrow> dbfm \<Rightarrow> dbfm"
where
"subst_dbfm u i (DBMem t1 t2) = DBMem (subst_dbtm u i t1) (subst_dbtm u i t2)"
| "subst_dbfm u i (DBEq t1 t2) = DBEq (subst_dbtm u i t1) (subst_dbtm u i t2)"
| "subst_dbfm u i (DBDisj A1 A2) = DBDisj (subst_dbfm u i A1) (subst_dbfm u i A2)"
| "subst_dbfm u i (DBNeg A) = DBNeg (subst_dbfm u i A)"
| "subst_dbfm u i (DBEx A) = DBEx (subst_dbfm u i A)"
by (auto simp: eqvt_def subst_dbfm_graph_aux_def) (metis dbfm.exhaust)
nominal_termination (eqvt)
by lexicographic_order
lemma fresh_iff_non_subst_dbfm: "subst_dbfm DBZero i t = t \<longleftrightarrow> atom i \<sharp> t"
by (induct t rule: dbfm.induct) (auto simp: fresh_iff_non_subst_dbtm)
section\<open>Well formed terms and formulas (de Bruijn representation)\<close>
inductive wf_dbfm :: "dbfm \<Rightarrow> bool"
where
Mem: "wf_dbtm t1 \<Longrightarrow> wf_dbtm t2 \<Longrightarrow> wf_dbfm (DBMem t1 t2)"
| Eq: "wf_dbtm t1 \<Longrightarrow> wf_dbtm t2 \<Longrightarrow> wf_dbfm (DBEq t1 t2)"
| Disj: "wf_dbfm A1 \<Longrightarrow> wf_dbfm A2 \<Longrightarrow> wf_dbfm (DBDisj A1 A2)"
| Neg: "wf_dbfm A \<Longrightarrow> wf_dbfm (DBNeg A)"
| Ex: "wf_dbfm A \<Longrightarrow> wf_dbfm (DBEx (abst_dbfm name 0 A))"
equivariance wf_dbfm
lemma atom_fresh_abst_dbtm [simp]: "atom i \<sharp> abst_dbtm i n t"
by (induct t rule: dbtm.induct) (auto simp: pure_fresh)
lemma atom_fresh_abst_dbfm [simp]: "atom i \<sharp> abst_dbfm i n A"
by (nominal_induct A arbitrary: n rule: dbfm.strong_induct) auto
text\<open>Setting up strong induction: "avoiding" for name. Necessary to allow some proofs to go through\<close>
nominal_inductive wf_dbfm
avoids Ex: name
by (auto simp: fresh_star_def)
inductive_cases Mem_wf_dbfm [elim!]: "wf_dbfm (DBMem t1 t2)"
inductive_cases Eq_wf_dbfm [elim!]: "wf_dbfm (DBEq t1 t2)"
inductive_cases Disj_wf_dbfm [elim!]: "wf_dbfm (DBDisj A1 A2)"
inductive_cases Neg_wf_dbfm [elim!]: "wf_dbfm (DBNeg A)"
inductive_cases Ex_wf_dbfm [elim!]: "wf_dbfm (DBEx z)"
declare wf_dbfm.intros [intro]
lemma trans_fm_abs: "trans_fm (e@[name]) A = abst_dbfm name (length e) (trans_fm e A)"
apply (nominal_induct A avoiding: name e rule: fm.strong_induct)
apply (auto simp: trans_tm_abs fresh_Cons fresh_append)
apply (metis One_nat_def Suc_eq_plus1 append_Cons list.size(4))
done
lemma abst_trans_fm: "abst_dbfm name 0 (trans_fm [] A) = trans_fm [name] A"
by (metis append_Nil list.size(3) trans_fm_abs)
lemma abst_trans_fm2: "i \<noteq> j \<Longrightarrow> abst_dbfm i (Suc 0) (trans_fm [j] A) = trans_fm [j,i] A"
using trans_fm_abs [where e="[j]" and name=i]
by auto
lemma wf_dbfm_imp_is_fm:
assumes "wf_dbfm x" shows "\<exists>A::fm. x = trans_fm [] A"
using assms
proof (induct rule: wf_dbfm.induct)
case (Mem t1 t2) thus ?case
by (metis trans_fm.simps(1) wf_dbtm_imp_is_tm)
next
case (Eq t1 t2) thus ?case
by (metis trans_fm.simps(2) wf_dbtm_imp_is_tm)
next
case (Disj fm1 fm2) thus ?case
by (metis trans_fm.simps(3))
next
case (Neg fm) thus ?case
by (metis trans_fm.simps(4))
next
case (Ex fm name) thus ?case
apply auto
apply (rule_tac x="Ex name A" in exI)
apply (auto simp: abst_trans_fm)
done
qed
lemma wf_dbfm_trans_fm: "wf_dbfm (trans_fm [] A)"
apply (nominal_induct A rule: fm.strong_induct)
apply (auto simp: wf_dbtm_trans_tm abst_trans_fm)
apply (metis abst_trans_fm wf_dbfm.Ex)
done
lemma wf_dbfm_iff_is_fm: "wf_dbfm x \<longleftrightarrow> (\<exists>A::fm. x = trans_fm [] A)"
by (metis wf_dbfm_imp_is_fm wf_dbfm_trans_fm)
lemma dbtm_abst_ignore [simp]:
"abst_dbtm name i (abst_dbtm name j t) = abst_dbtm name j t"
by (induct t rule: dbtm.induct) auto
lemma abst_dbtm_fresh_ignore [simp]: "atom name \<sharp> u \<Longrightarrow> abst_dbtm name j u = u"
by (induct u rule: dbtm.induct) auto
lemma dbtm_subst_ignore [simp]:
"subst_dbtm u name (abst_dbtm name j t) = abst_dbtm name j t"
by (induct t rule: dbtm.induct) auto
lemma dbtm_abst_swap_subst:
"name \<noteq> name' \<Longrightarrow> atom name' \<sharp> u \<Longrightarrow>
subst_dbtm u name (abst_dbtm name' j t) = abst_dbtm name' j (subst_dbtm u name t)"
by (induct t rule: dbtm.induct) auto
lemma dbfm_abst_swap_subst:
"name \<noteq> name' \<Longrightarrow> atom name' \<sharp> u \<Longrightarrow>
subst_dbfm u name (abst_dbfm name' j A) = abst_dbfm name' j (subst_dbfm u name A)"
by (induct A arbitrary: j rule: dbfm.induct) (auto simp: dbtm_abst_swap_subst)
lemma subst_trans_commute [simp]:
"atom i \<sharp> e \<Longrightarrow> subst_dbtm (trans_tm e u) i (trans_tm e t) = trans_tm e (subst i u t)"
apply (induct t rule: tm.induct)
apply (auto simp: lookup_notin fresh_imp_notin_env)
apply (metis abst_dbtm_fresh_ignore dbtm_subst_ignore lookup_fresh lookup_notin subst_dbtm.simps(2))
done
lemma subst_fm_trans_commute [simp]:
"subst_dbfm (trans_tm [] u) name (trans_fm [] A) = trans_fm [] (A (name::= u))"
apply (nominal_induct A avoiding: name u rule: fm.strong_induct)
apply (auto simp: lookup_notin abst_trans_fm [symmetric])
apply (metis dbfm_abst_swap_subst fresh_at_base(2) fresh_trans_tm_iff)
done
lemma subst_fm_trans_commute_eq:
"du = trans_tm [] u \<Longrightarrow> subst_dbfm du i (trans_fm [] A) = trans_fm [] (A(i::=u))"
by (metis subst_fm_trans_commute)
section\<open>Quotations\<close>
fun htuple :: "nat \<Rightarrow> hf" where
"htuple 0 = \<langle>0,0\<rangle>"
| "htuple (Suc k) = \<langle>0, htuple k\<rangle>"
fun HTuple :: "nat \<Rightarrow> tm" where
"HTuple 0 = HPair Zero Zero"
| "HTuple (Suc k) = HPair Zero (HTuple k)"
lemma eval_tm_HTuple [simp]: "\<lbrakk>HTuple n\<rbrakk>e = htuple n"
by (induct n) auto
lemma fresh_HTuple [simp]: "x \<sharp> HTuple n"
by (induct n) auto
lemma HTuple_eqvt[eqvt]: "(p \<bullet> HTuple n) = HTuple (p \<bullet> n)"
by (induct n, auto simp: HPair_eqvt permute_pure)
lemma htuple_nonzero [simp]: "htuple k \<noteq> 0"
by (induct k) auto
lemma htuple_inject [iff]: "htuple i = htuple j \<longleftrightarrow> i=j"
proof (induct i arbitrary: j)
case 0 show ?case
by (cases j) auto
next
case (Suc i) show ?case
by (cases j) (auto simp: Suc)
qed
subsection \<open>Quotations of de Bruijn terms\<close>
definition nat_of_name :: "name \<Rightarrow> nat"
where "nat_of_name x = nat_of (atom x)"
lemma nat_of_name_inject [simp]: "nat_of_name n1 = nat_of_name n2 \<longleftrightarrow> n1 = n2"
by (metis nat_of_name_def atom_components_eq_iff atom_eq_iff sort_of_atom_eq)
definition name_of_nat :: "nat \<Rightarrow> name"
where "name_of_nat n \<equiv> Abs_name (Atom (Sort ''SyntaxN.name'' []) n)"
lemma nat_of_name_Abs_eq [simp]: "nat_of_name (Abs_name (Atom (Sort ''SyntaxN.name'' []) n)) = n"
by (auto simp: nat_of_name_def atom_name_def Abs_name_inverse)
lemma nat_of_name_name_eq [simp]: "nat_of_name (name_of_nat n) = n"
by (simp add: name_of_nat_def)
lemma name_of_nat_nat_of_name [simp]: "name_of_nat (nat_of_name i) = i"
by (metis nat_of_name_inject nat_of_name_name_eq)
lemma HPair_neq_ORD_OF [simp]: "HPair x y \<noteq> ORD_OF i"
by (metis Not_Ord_hpair Ord_ord_of eval_tm_HPair eval_tm_ORD_OF)
text\<open>Infinite support, so we cannot use nominal primrec.\<close>
function quot_dbtm :: "dbtm \<Rightarrow> tm"
where
"quot_dbtm DBZero = Zero"
| "quot_dbtm (DBVar name) = ORD_OF (Suc (nat_of_name name))"
| "quot_dbtm (DBInd k) = HPair (HTuple 6) (ORD_OF k)"
| "quot_dbtm (DBEats t u) = HPair (HTuple 1) (HPair (quot_dbtm t) (quot_dbtm u))"
by (rule dbtm.exhaust) auto
termination
by lexicographic_order
lemma quot_dbtm_inject_lemma [simp]: "\<lbrakk>quot_dbtm t\<rbrakk>e = \<lbrakk>quot_dbtm u\<rbrakk>e \<longleftrightarrow> t=u"
proof (induct t arbitrary: u rule: dbtm.induct)
case DBZero show ?case
by (induct u rule: dbtm.induct) auto
next
case (DBVar name) show ?case
by (induct u rule: dbtm.induct) (auto simp: hpair_neq_Ord')
next
case (DBInd k) show ?case
by (induct u rule: dbtm.induct) (auto simp: hpair_neq_Ord hpair_neq_Ord')
next
case (DBEats t1 t2) thus ?case
by (induct u rule: dbtm.induct) (simp_all add: hpair_neq_Ord)
qed
lemma quot_dbtm_inject [iff]: "quot_dbtm t = quot_dbtm u \<longleftrightarrow> t=u"
by (metis quot_dbtm_inject_lemma)
subsection \<open>Quotations of de Bruijn formulas\<close>
text\<open>Infinite support, so we cannot use nominal primrec.\<close>
function quot_dbfm :: "dbfm \<Rightarrow> tm"
where
"quot_dbfm (DBMem t u) = HPair (HTuple 0) (HPair (quot_dbtm t) (quot_dbtm u))"
| "quot_dbfm (DBEq t u) = HPair (HTuple 2) (HPair (quot_dbtm t) (quot_dbtm u))"
| "quot_dbfm (DBDisj A B) = HPair (HTuple 3) (HPair (quot_dbfm A) (quot_dbfm B))"
| "quot_dbfm (DBNeg A) = HPair (HTuple 4) (quot_dbfm A)"
| "quot_dbfm (DBEx A) = HPair (HTuple 5) (quot_dbfm A)"
by (rule_tac y=x in dbfm.exhaust, auto)
termination
by lexicographic_order
lemma htuple_minus_1: "n > 0 \<Longrightarrow> htuple n = \<langle>0, htuple (n - 1)\<rangle>"
by (metis Suc_diff_1 htuple.simps(2))
lemma HTuple_minus_1: "n > 0 \<Longrightarrow> HTuple n = HPair Zero (HTuple (n - 1))"
by (metis Suc_diff_1 HTuple.simps(2))
lemmas HTS = HTuple_minus_1 HTuple.simps \<comment> \<open>for freeness reasoning on codes\<close>
lemma quot_dbfm_inject_lemma [simp]: "\<lbrakk>quot_dbfm A\<rbrakk>e = \<lbrakk>quot_dbfm B\<rbrakk>e \<longleftrightarrow> A=B"
proof (induct A arbitrary: B rule: dbfm.induct)
case (DBMem t u) show ?case
by (induct B rule: dbfm.induct) (simp_all add: htuple_minus_1)
next
case (DBEq t u) show ?case
by (induct B rule: dbfm.induct) (auto simp: htuple_minus_1)
next
case (DBDisj A B') thus ?case
by (induct B rule: dbfm.induct) (simp_all add: htuple_minus_1)
next
case (DBNeg A) thus ?case
by (induct B rule: dbfm.induct) (simp_all add: htuple_minus_1)
next
case (DBEx A) thus ?case
by (induct B rule: dbfm.induct) (simp_all add: htuple_minus_1)
qed
class quot =
fixes quot :: "'a \<Rightarrow> tm" ("\<guillemotleft>_\<guillemotright>")
instantiation tm :: quot
begin
definition quot_tm :: "tm \<Rightarrow> tm"
where "quot_tm t = quot_dbtm (trans_tm [] t)"
instance ..
end
lemma quot_dbtm_fresh [simp]: "s \<sharp> (quot_dbtm t)"
by (induct t rule: dbtm.induct) auto
lemma quot_tm_fresh [simp]: fixes t::tm shows "s \<sharp> \<guillemotleft>t\<guillemotright>"
by (simp add: quot_tm_def)
lemma quot_Zero [simp]: "\<guillemotleft>Zero\<guillemotright> = Zero"
by (simp add: quot_tm_def)
lemma quot_Var: "\<guillemotleft>Var x\<guillemotright> = SUCC (ORD_OF (nat_of_name x))"
by (simp add: quot_tm_def)
lemma quot_Eats: "\<guillemotleft>Eats x y\<guillemotright> = HPair (HTuple 1) (HPair \<guillemotleft>x\<guillemotright> \<guillemotleft>y\<guillemotright>)"
by (simp add: quot_tm_def)
text\<open>irrelevance of the environment for quotations, because they are ground terms\<close>
lemma eval_quot_dbtm_ignore:
"\<lbrakk>quot_dbtm t\<rbrakk>e = \<lbrakk>quot_dbtm t\<rbrakk>e'"
by (induct t rule: dbtm.induct) auto
lemma eval_quot_dbfm_ignore:
"\<lbrakk>quot_dbfm A\<rbrakk>e = \<lbrakk>quot_dbfm A\<rbrakk>e'"
by (induct A rule: dbfm.induct) (auto intro: eval_quot_dbtm_ignore)
instantiation fm :: quot
begin
definition quot_fm :: "fm \<Rightarrow> tm"
where "quot_fm A = quot_dbfm (trans_fm [] A)"
instance ..
end
lemma quot_dbfm_fresh [simp]: "s \<sharp> (quot_dbfm A)"
by (induct A rule: dbfm.induct) auto
lemma quot_fm_fresh [simp]: fixes A::fm shows "s \<sharp> \<guillemotleft>A\<guillemotright>"
by (simp add: quot_fm_def)
lemma quot_fm_permute [simp]: fixes A:: fm shows "p \<bullet> \<guillemotleft>A\<guillemotright> = \<guillemotleft>A\<guillemotright>"
by (metis fresh_star_def perm_supp_eq quot_fm_fresh)
lemma quot_Mem: "\<guillemotleft>x IN y\<guillemotright> = HPair (HTuple 0) (HPair (\<guillemotleft>x\<guillemotright>) (\<guillemotleft>y\<guillemotright>))"
by (simp add: quot_fm_def quot_tm_def)
lemma quot_Eq: "\<guillemotleft>x EQ y\<guillemotright> = HPair (HTuple 2) (HPair (\<guillemotleft>x\<guillemotright>) (\<guillemotleft>y\<guillemotright>))"
by (simp add: quot_fm_def quot_tm_def)
lemma quot_Disj: "\<guillemotleft>A OR B\<guillemotright> = HPair (HTuple 3) (HPair (\<guillemotleft>A\<guillemotright>) (\<guillemotleft>B\<guillemotright>))"
by (simp add: quot_fm_def)
lemma quot_Neg: "\<guillemotleft>Neg A\<guillemotright> = HPair (HTuple 4) (\<guillemotleft>A\<guillemotright>)"
by (simp add: quot_fm_def)
lemma quot_Ex: "\<guillemotleft>Ex i A\<guillemotright> = HPair (HTuple 5) (quot_dbfm (trans_fm [i] A))"
by (simp add: quot_fm_def)
lemma eval_quot_fm_ignore: fixes A:: fm shows "\<lbrakk>\<guillemotleft>A\<guillemotright>\<rbrakk>e = \<lbrakk>\<guillemotleft>A\<guillemotright>\<rbrakk>e'"
by (metis eval_quot_dbfm_ignore quot_fm_def)
lemmas quot_simps = quot_Var quot_Eats quot_Eq quot_Mem quot_Disj quot_Neg quot_Ex
section\<open>Definitions Involving Coding\<close>
definition q_Var :: "name \<Rightarrow> hf"
where "q_Var i \<equiv> succ (ord_of (nat_of_name i))"
definition q_Ind :: "hf \<Rightarrow> hf"
where "q_Ind k \<equiv> \<langle>htuple 6, k\<rangle>"
abbreviation Q_Eats :: "tm \<Rightarrow> tm \<Rightarrow> tm"
where "Q_Eats t u \<equiv> HPair (HTuple (Suc 0)) (HPair t u)"
definition q_Eats :: "hf \<Rightarrow> hf \<Rightarrow> hf"
where "q_Eats x y \<equiv> \<langle>htuple 1, x, y\<rangle>"
abbreviation Q_Succ :: "tm \<Rightarrow> tm"
where "Q_Succ t \<equiv> Q_Eats t t"
definition q_Succ :: "hf \<Rightarrow> hf"
where "q_Succ x \<equiv> q_Eats x x"
lemma quot_Succ: "\<guillemotleft>SUCC x\<guillemotright> = Q_Succ \<guillemotleft>x\<guillemotright>"
by (auto simp: SUCC_def quot_Eats)
abbreviation Q_HPair :: "tm \<Rightarrow> tm \<Rightarrow> tm"
where "Q_HPair t u \<equiv>
Q_Eats (Q_Eats Zero (Q_Eats (Q_Eats Zero u) t))
(Q_Eats (Q_Eats Zero t) t)"
definition q_HPair :: "hf \<Rightarrow> hf \<Rightarrow> hf"
where "q_HPair x y \<equiv>
q_Eats (q_Eats 0 (q_Eats (q_Eats 0 y) x))
(q_Eats (q_Eats 0 x) x)"
abbreviation Q_Mem :: "tm \<Rightarrow> tm \<Rightarrow> tm"
where "Q_Mem t u \<equiv> HPair (HTuple 0) (HPair t u)"
definition q_Mem :: "hf \<Rightarrow> hf \<Rightarrow> hf"
where "q_Mem x y \<equiv> \<langle>htuple 0, x, y\<rangle>"
abbreviation Q_Eq :: "tm \<Rightarrow> tm \<Rightarrow> tm"
where "Q_Eq t u \<equiv> HPair (HTuple 2) (HPair t u)"
definition q_Eq :: "hf \<Rightarrow> hf \<Rightarrow> hf"
where "q_Eq x y \<equiv> \<langle>htuple 2, x, y\<rangle>"
abbreviation Q_Disj :: "tm \<Rightarrow> tm \<Rightarrow> tm"
where "Q_Disj t u \<equiv> HPair (HTuple 3) (HPair t u)"
definition q_Disj :: "hf \<Rightarrow> hf \<Rightarrow> hf"
where "q_Disj x y \<equiv> \<langle>htuple 3, x, y\<rangle>"
abbreviation Q_Neg :: "tm \<Rightarrow> tm"
where "Q_Neg t \<equiv> HPair (HTuple 4) t"
definition q_Neg :: "hf \<Rightarrow> hf"
where "q_Neg x \<equiv> \<langle>htuple 4, x\<rangle>"
abbreviation Q_Conj :: "tm \<Rightarrow> tm \<Rightarrow> tm"
where "Q_Conj t u \<equiv> Q_Neg (Q_Disj (Q_Neg t) (Q_Neg u))"
definition q_Conj :: "hf \<Rightarrow> hf \<Rightarrow> hf"
where "q_Conj t u \<equiv> q_Neg (q_Disj (q_Neg t) (q_Neg u))"
abbreviation Q_Imp :: "tm \<Rightarrow> tm \<Rightarrow> tm"
where "Q_Imp t u \<equiv> Q_Disj (Q_Neg t) u"
definition q_Imp :: "hf \<Rightarrow> hf \<Rightarrow> hf"
where "q_Imp t u \<equiv> q_Disj (q_Neg t) u"
abbreviation Q_Ex :: "tm \<Rightarrow> tm"
where "Q_Ex t \<equiv> HPair (HTuple 5) t"
definition q_Ex :: "hf \<Rightarrow> hf"
where "q_Ex x \<equiv> \<langle>htuple 5, x\<rangle>"
abbreviation Q_All :: "tm \<Rightarrow> tm"
where "Q_All t \<equiv> Q_Neg (Q_Ex (Q_Neg t))"
definition q_All :: "hf \<Rightarrow> hf"
where "q_All x \<equiv> q_Neg (q_Ex (q_Neg x))"
lemmas q_defs = q_Var_def q_Ind_def q_Eats_def q_HPair_def q_Eq_def q_Mem_def
q_Disj_def q_Neg_def q_Conj_def q_Imp_def q_Ex_def q_All_def
lemma q_Eats_iff [iff]: "q_Eats x y = q_Eats x' y' \<longleftrightarrow> x=x' \<and> y=y'"
by (metis hpair_iff q_Eats_def)
lemma quot_subst_eq: "\<guillemotleft>A(i::=t)\<guillemotright> = quot_dbfm (subst_dbfm (trans_tm [] t) i (trans_fm [] A))"
by (metis quot_fm_def subst_fm_trans_commute)
lemma Q_Succ_cong: "H \<turnstile> x EQ x' \<Longrightarrow> H \<turnstile> Q_Succ x EQ Q_Succ x'"
by (metis HPair_cong Refl)
section\<open>Quotations are Injective\<close>
subsection\<open>Terms\<close>
lemma eval_tm_inject [simp]: fixes t::tm shows "\<lbrakk>\<guillemotleft>t\<guillemotright>\<rbrakk> e = \<lbrakk>\<guillemotleft>u\<guillemotright>\<rbrakk> e \<longleftrightarrow> t=u"
proof (induct t arbitrary: u rule: tm.induct)
case Zero thus ?case
by (cases u rule: tm.exhaust) (auto simp: quot_Var quot_Eats)
next
case (Var i) thus ?case
apply (cases u rule: tm.exhaust, auto)
apply (auto simp: quot_Var quot_Eats)
done
next
case (Eats t1 t2) thus ?case
apply (cases u rule: tm.exhaust, auto)
apply (auto simp: quot_Eats quot_Var)
done
qed
subsection\<open>Formulas\<close>
lemma eval_fm_inject [simp]: fixes A::fm shows "\<lbrakk>\<guillemotleft>A\<guillemotright>\<rbrakk> e = \<lbrakk>\<guillemotleft>B\<guillemotright>\<rbrakk> e \<longleftrightarrow> A=B"
proof (nominal_induct B arbitrary: A rule: fm.strong_induct)
case (Mem tm1 tm2) thus ?case
by (cases A rule: fm.exhaust, auto simp: quot_simps htuple_minus_1)
next
case (Eq tm1 tm2) thus ?case
by (cases A rule: fm.exhaust, auto simp: quot_simps htuple_minus_1)
next
case (Neg \<alpha>) thus ?case
by (cases A rule: fm.exhaust, auto simp: quot_simps htuple_minus_1)
next
case (Disj fm1 fm2)
thus ?case
by (cases A rule: fm.exhaust, auto simp: quot_simps htuple_minus_1)
next
case (Ex i \<alpha>)
thus ?case
apply (induct A arbitrary: i rule: fm.induct)
apply (auto simp: trans_fm_perm quot_simps htuple_minus_1 Abs1_eq_iff_all)
by (metis (no_types) Abs1_eq_iff_all(3) dbfm.eq_iff(5) fm.eq_iff(5) fresh_Nil trans_fm.simps(5))
qed
subsection\<open>The set \<open>\<Gamma>\<close> of Definition 1.1, constant terms used for coding\<close>
inductive coding_tm :: "tm \<Rightarrow> bool"
where
Ord: "\<exists>i. x = ORD_OF i \<Longrightarrow> coding_tm x"
| HPair: "coding_tm x \<Longrightarrow> coding_tm y \<Longrightarrow> coding_tm (HPair x y)"
declare coding_tm.intros [intro]
lemma coding_tm_Zero [intro]: "coding_tm Zero"
by (metis ORD_OF.simps(1) Ord)
lemma coding_tm_HTuple [intro]: "coding_tm (HTuple k)"
by (induct k, auto)
inductive_simps coding_tm_HPair [simp]: "coding_tm (HPair x y)"
lemma quot_dbtm_coding [simp]: "coding_tm (quot_dbtm t)"
apply (induct t rule: dbtm.induct, auto)
apply (metis ORD_OF.simps(2) Ord)
done
lemma quot_dbfm_coding [simp]: "coding_tm (quot_dbfm fm)"
by (induct fm rule: dbfm.induct, auto)
lemma quot_fm_coding: fixes A::fm shows "coding_tm \<guillemotleft>A\<guillemotright>"
by (metis quot_dbfm_coding quot_fm_def)
inductive coding_hf :: "hf \<Rightarrow> bool"
where
Ord: "\<exists>i. x = ord_of i \<Longrightarrow> coding_hf x"
| HPair: "coding_hf x \<Longrightarrow> coding_hf y \<Longrightarrow> coding_hf (\<langle>x,y\<rangle>)"
declare coding_hf.intros [intro]
lemma coding_hf_0 [intro]: "coding_hf 0"
by (metis coding_hf.Ord ord_of.simps(1))
inductive_simps coding_hf_hpair [simp]: "coding_hf (\<langle>x,y\<rangle>)"
lemma coding_tm_hf [simp]: "coding_tm t \<Longrightarrow> coding_hf \<lbrakk>t\<rbrakk>e"
by (induct t rule: coding_tm.induct) auto
section \<open>V-Coding for terms and formulas, for the Second Theorem\<close>
text\<open>Infinite support, so we cannot use nominal primrec.\<close>
function vquot_dbtm :: "name set \<Rightarrow> dbtm \<Rightarrow> tm"
where
"vquot_dbtm V DBZero = Zero"
| "vquot_dbtm V (DBVar name) = (if name \<in> V then Var name
else ORD_OF (Suc (nat_of_name name)))"
| "vquot_dbtm V (DBInd k) = HPair (HTuple 6) (ORD_OF k)"
| "vquot_dbtm V (DBEats t u) = HPair (HTuple 1) (HPair (vquot_dbtm V t) (vquot_dbtm V u))"
by (auto, rule_tac y=b in dbtm.exhaust, auto)
termination
by lexicographic_order
lemma fresh_vquot_dbtm [simp]: "i \<sharp> vquot_dbtm V tm \<longleftrightarrow> i \<sharp> tm \<or> i \<notin> atom ` V"
by (induct tm rule: dbtm.induct) (auto simp: fresh_at_base pure_fresh)
text\<open>Infinite support, so we cannot use nominal primrec.\<close>
function vquot_dbfm :: "name set \<Rightarrow> dbfm \<Rightarrow> tm"
where
"vquot_dbfm V (DBMem t u) = HPair (HTuple 0) (HPair (vquot_dbtm V t) (vquot_dbtm V u))"
| "vquot_dbfm V (DBEq t u) = HPair (HTuple 2) (HPair (vquot_dbtm V t) (vquot_dbtm V u))"
| "vquot_dbfm V (DBDisj A B) = HPair (HTuple 3) (HPair (vquot_dbfm V A) (vquot_dbfm V B))"
| "vquot_dbfm V (DBNeg A) = HPair (HTuple 4) (vquot_dbfm V A)"
| "vquot_dbfm V (DBEx A) = HPair (HTuple 5) (vquot_dbfm V A)"
by (auto, rule_tac y=b in dbfm.exhaust, auto)
termination
by lexicographic_order
lemma fresh_vquot_dbfm [simp]: "i \<sharp> vquot_dbfm V fm \<longleftrightarrow> i \<sharp> fm \<or> i \<notin> atom ` V"
by (induct fm rule: dbfm.induct) (auto simp: HPair_def HTuple_minus_1)
class vquot =
fixes vquot :: "'a \<Rightarrow> name set \<Rightarrow> tm" ("\<lfloor>_\<rfloor>_" [0,1000]1000)
instantiation tm :: vquot
begin
definition vquot_tm :: "tm \<Rightarrow> name set \<Rightarrow> tm"
where "vquot_tm t V = vquot_dbtm V (trans_tm [] t)"
instance ..
end
lemma vquot_dbtm_empty [simp]: "vquot_dbtm {} t = quot_dbtm t"
by (induct t rule: dbtm.induct) auto
lemma vquot_tm_empty [simp]: fixes t::tm shows "\<lfloor>t\<rfloor>{} = \<guillemotleft>t\<guillemotright>"
by (simp add: vquot_tm_def quot_tm_def)
lemma vquot_dbtm_eq: "atom ` V \<inter> supp t = atom ` W \<inter> supp t \<Longrightarrow> vquot_dbtm V t = vquot_dbtm W t"
by (induct t rule: dbtm.induct) (auto simp: image_iff, blast+)
instantiation fm :: vquot
begin
definition vquot_fm :: "fm \<Rightarrow> name set \<Rightarrow> tm"
where "vquot_fm A V = vquot_dbfm V (trans_fm [] A)"
instance ..
end
lemma vquot_fm_fresh [simp]: fixes A::fm shows "i \<sharp> \<lfloor>A\<rfloor>V \<longleftrightarrow> i \<sharp> A \<or> i \<notin> atom ` V"
by (simp add: vquot_fm_def)
lemma vquot_dbfm_empty [simp]: "vquot_dbfm {} A = quot_dbfm A"
by (induct A rule: dbfm.induct) auto
lemma vquot_fm_empty [simp]: fixes A::fm shows "\<lfloor>A\<rfloor>{} = \<guillemotleft>A\<guillemotright>"
by (simp add: vquot_fm_def quot_fm_def)
lemma vquot_dbfm_eq: "atom ` V \<inter> supp A = atom ` W \<inter> supp A \<Longrightarrow> vquot_dbfm V A = vquot_dbfm W A"
by (induct A rule: dbfm.induct) (auto simp: intro!: vquot_dbtm_eq, blast+)
lemma vquot_fm_insert:
fixes A::fm shows "atom i \<notin> supp A \<Longrightarrow> \<lfloor>A\<rfloor>(insert i V) = \<lfloor>A\<rfloor>V"
by (auto simp: vquot_fm_def supp_conv_fresh intro: vquot_dbfm_eq)
declare HTuple.simps [simp del]
end
diff --git a/thys/Incompleteness/SyntaxN.thy b/thys/Incompleteness/SyntaxN.thy
--- a/thys/Incompleteness/SyntaxN.thy
+++ b/thys/Incompleteness/SyntaxN.thy
@@ -1,1593 +1,1594 @@
chapter\<open>Syntax of Terms and Formulas using Nominal Logic\<close>
theory SyntaxN
imports Nominal2.Nominal2 HereditarilyFinite.OrdArith
begin
section\<open>Terms and Formulas\<close>
subsection\<open>Hf is a pure permutation type\<close>
instantiation hf :: pt
begin
definition "p \<bullet> (s::hf) = s"
instance
by standard (simp_all add: permute_hf_def)
end
instance hf :: pure
proof qed (rule permute_hf_def)
atom_decl name
declare fresh_set_empty [simp]
lemma supp_name [simp]: fixes i::name shows "supp i = {atom i}"
by (rule supp_at_base)
subsection\<open>The datatypes\<close>
nominal_datatype tm = Zero | Var name | Eats tm tm
nominal_datatype fm =
Mem tm tm (infixr "IN" 150)
| Eq tm tm (infixr "EQ" 150)
| Disj fm fm (infixr "OR" 130)
| Neg fm
| Ex x::name f::fm binds x in f
text \<open>Mem, Eq are atomic formulas; Disj, Neg, Ex are non-atomic\<close>
declare tm.supp [simp] fm.supp [simp]
subsection\<open>Substitution\<close>
nominal_function subst :: "name \<Rightarrow> tm \<Rightarrow> tm \<Rightarrow> tm"
where
"subst i x Zero = Zero"
| "subst i x (Var k) = (if i=k then x else Var k)"
| "subst i x (Eats t u) = Eats (subst i x t) (subst i x u)"
by (auto simp: eqvt_def subst_graph_aux_def) (metis tm.strong_exhaust)
nominal_termination (eqvt)
by lexicographic_order
lemma fresh_subst_if [simp]:
"j \<sharp> subst i x t \<longleftrightarrow> (atom i \<sharp> t \<and> j \<sharp> t) \<or> (j \<sharp> x \<and> (j \<sharp> t \<or> j = atom i))"
by (induct t rule: tm.induct) (auto simp: fresh_at_base)
lemma forget_subst_tm [simp]: "atom a \<sharp> tm \<Longrightarrow> subst a x tm = tm"
by (induct tm rule: tm.induct) (simp_all add: fresh_at_base)
lemma subst_tm_id [simp]: "subst a (Var a) tm = tm"
by (induct tm rule: tm.induct) simp_all
lemma subst_tm_commute [simp]:
"atom j \<sharp> tm \<Longrightarrow> subst j u (subst i t tm) = subst i (subst j u t) tm"
by (induct tm rule: tm.induct) (auto simp: fresh_Pair)
lemma subst_tm_commute2 [simp]:
"atom j \<sharp> t \<Longrightarrow> atom i \<sharp> u \<Longrightarrow> i \<noteq> j \<Longrightarrow> subst j u (subst i t tm) = subst i t (subst j u tm)"
by (induct tm rule: tm.induct) auto
lemma repeat_subst_tm [simp]: "subst i u (subst i t tm) = subst i (subst i u t) tm"
by (induct tm rule: tm.induct) auto
nominal_function subst_fm :: "fm \<Rightarrow> name \<Rightarrow> tm \<Rightarrow> fm" ("_'(_::=_')" [1000, 0, 0] 200)
where
Mem: "(Mem t u)(i::=x) = Mem (subst i x t) (subst i x u)"
| Eq: "(Eq t u)(i::=x) = Eq (subst i x t) (subst i x u)"
| Disj: "(Disj A B)(i::=x) = Disj (A(i::=x)) (B(i::=x))"
| Neg: "(Neg A)(i::=x) = Neg (A(i::=x))"
| Ex: "atom j \<sharp> (i, x) \<Longrightarrow> (Ex j A)(i::=x) = Ex j (A(i::=x))"
apply (simp add: eqvt_def subst_fm_graph_aux_def)
apply auto [16]
apply (rule_tac y=a and c="(aa, b)" in fm.strong_exhaust)
apply (auto simp: eqvt_at_def fresh_star_def fresh_Pair fresh_at_base)
apply (metis flip_at_base_simps(3) flip_fresh_fresh)
done
nominal_termination (eqvt)
by lexicographic_order
lemma size_subst_fm [simp]: "size (A(i::=x)) = size A"
by (nominal_induct A avoiding: i x rule: fm.strong_induct) auto
lemma forget_subst_fm [simp]: "atom a \<sharp> A \<Longrightarrow> A(a::=x) = A"
by (nominal_induct A avoiding: a x rule: fm.strong_induct) (auto simp: fresh_at_base)
lemma subst_fm_id [simp]: "A(a::=Var a) = A"
by (nominal_induct A avoiding: a rule: fm.strong_induct) (auto simp: fresh_at_base)
lemma fresh_subst_fm_if [simp]:
"j \<sharp> (A(i::=x)) \<longleftrightarrow> (atom i \<sharp> A \<and> j \<sharp> A) \<or> (j \<sharp> x \<and> (j \<sharp> A \<or> j = atom i))"
by (nominal_induct A avoiding: i x rule: fm.strong_induct) (auto simp: fresh_at_base)
lemma subst_fm_commute [simp]:
"atom j \<sharp> A \<Longrightarrow> (A(i::=t))(j::=u) = A(i ::= subst j u t)"
by (nominal_induct A avoiding: i j t u rule: fm.strong_induct) (auto simp: fresh_at_base)
lemma repeat_subst_fm [simp]: "(A(i::=t))(i::=u) = A(i ::= subst i u t)"
by (nominal_induct A avoiding: i t u rule: fm.strong_induct) auto
lemma subst_fm_Ex_with_renaming:
"atom i' \<sharp> (A, i, j, t) \<Longrightarrow> (Ex i A)(j ::= t) = Ex i' (((i \<leftrightarrow> i') \<bullet> A)(j ::= t))"
by (rule subst [of "Ex i' ((i \<leftrightarrow> i') \<bullet> A)" "Ex i A"])
(auto simp: Abs1_eq_iff flip_def swap_commute)
text \<open>the simplifier cannot apply the rule above, because
it introduces a new variable at the right hand side.\<close>
simproc_setup subst_fm_renaming ("(Ex i A)(j ::= t)") = \<open>fn _ => fn ctxt => fn ctrm =>
let
val _ $ (_ $ i $ A) $ j $ t = Thm.term_of ctrm
val atoms = Simplifier.prems_of ctxt
|> map_filter (fn thm => case Thm.prop_of thm of
_ $ (Const (@{const_name fresh}, _) $ atm $ _) => SOME (atm) | _ => NONE)
|> distinct ((=))
fun get_thm atm =
let
val goal = HOLogic.mk_Trueprop (mk_fresh atm (HOLogic.mk_tuple [A, i, j, t]))
in
SOME ((Goal.prove ctxt [] [] goal (K (asm_full_simp_tac ctxt 1)))
RS @{thm subst_fm_Ex_with_renaming} RS eq_reflection)
handle ERROR _ => NONE
end
in
get_first get_thm atoms
end
\<close>
subsection\<open>Semantics\<close>
definition e0 :: "(name, hf) finfun" \<comment> \<open>the null environment\<close>
where "e0 \<equiv> finfun_const 0"
nominal_function eval_tm :: "(name, hf) finfun \<Rightarrow> tm \<Rightarrow> hf"
where
"eval_tm e Zero = 0"
| "eval_tm e (Var k) = finfun_apply e k"
| "eval_tm e (Eats t u) = eval_tm e t \<triangleleft> eval_tm e u"
by (auto simp: eqvt_def eval_tm_graph_aux_def) (metis tm.strong_exhaust)
nominal_termination (eqvt)
by lexicographic_order
syntax
"_EvalTm" :: "tm \<Rightarrow> (name, hf) finfun \<Rightarrow> hf" ("\<lbrakk>_\<rbrakk>_" [0,1000]1000)
translations
"\<lbrakk>tm\<rbrakk>e" == "CONST eval_tm e tm"
nominal_function eval_fm :: "(name, hf) finfun \<Rightarrow> fm \<Rightarrow> bool"
where
"eval_fm e (t IN u) \<longleftrightarrow> \<lbrakk>t\<rbrakk>e \<^bold>\<in> \<lbrakk>u\<rbrakk>e"
| "eval_fm e (t EQ u) \<longleftrightarrow> \<lbrakk>t\<rbrakk>e = \<lbrakk>u\<rbrakk>e"
| "eval_fm e (A OR B) \<longleftrightarrow> eval_fm e A \<or> eval_fm e B"
| "eval_fm e (Neg A) \<longleftrightarrow> (~ eval_fm e A)"
| "atom k \<sharp> e \<Longrightarrow> eval_fm e (Ex k A) \<longleftrightarrow> (\<exists>x. eval_fm (finfun_update e k x) A)"
+supply [[simproc del: defined_all]]
apply(simp add: eqvt_def eval_fm_graph_aux_def)
apply(auto del: iffI)[16]
apply(rule_tac y=b and c="(a)" in fm.strong_exhaust)
apply(auto simp: fresh_star_def)[5]
using [[simproc del: alpha_lst]] apply clarsimp
apply(erule_tac c="(ea)" in Abs_lst1_fcb2')
apply(rule pure_fresh)
apply(simp add: fresh_star_def)
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq)
done
nominal_termination (eqvt)
by lexicographic_order
lemma eval_tm_rename:
assumes "atom k' \<sharp> t"
shows "\<lbrakk>t\<rbrakk>(finfun_update e k x) = \<lbrakk>(k' \<leftrightarrow> k) \<bullet> t\<rbrakk>(finfun_update e k' x)"
using assms
by (induct t rule: tm.induct) (auto simp: permute_flip_at)
lemma eval_fm_rename:
assumes "atom k' \<sharp> A"
shows "eval_fm (finfun_update e k x) A = eval_fm (finfun_update e k' x) ((k' \<leftrightarrow> k) \<bullet> A)"
using assms
apply (nominal_induct A avoiding: e k k' x rule: fm.strong_induct)
apply (simp_all add: eval_tm_rename[symmetric], metis)
apply (simp add: fresh_finfun_update fresh_at_base finfun_update_twist)
done
lemma better_ex_eval_fm[simp]:
"eval_fm e (Ex k A) \<longleftrightarrow> (\<exists>x. eval_fm (finfun_update e k x) A)"
proof -
obtain k'::name where k': "atom k' \<sharp> (k, e, A)"
by (rule obtain_fresh)
then have eq: "Ex k' ((k' \<leftrightarrow> k) \<bullet> A) = Ex k A"
by (simp add: Abs1_eq_iff flip_def)
have "eval_fm e (Ex k' ((k' \<leftrightarrow> k) \<bullet> A)) = (\<exists>x. eval_fm (finfun_update e k' x) ((k' \<leftrightarrow> k) \<bullet> A))"
using k' by simp
also have "... = (\<exists>x. eval_fm (finfun_update e k x) A)"
by (metis eval_fm_rename k' fresh_Pair)
finally show ?thesis
by (metis eq)
qed
lemma forget_eval_tm [simp]: "atom i \<sharp> t \<Longrightarrow> \<lbrakk>t\<rbrakk>(finfun_update e i x) = \<lbrakk>t\<rbrakk>e"
by (induct t rule: tm.induct) (simp_all add: fresh_at_base)
lemma forget_eval_fm [simp]:
"atom k \<sharp> A \<Longrightarrow> eval_fm (finfun_update e k x) A = eval_fm e A"
by (nominal_induct A avoiding: k e rule: fm.strong_induct)
(simp_all add: fresh_at_base finfun_update_twist)
lemma eval_subst_tm: "\<lbrakk>subst i t u\<rbrakk>e = \<lbrakk>u\<rbrakk>(finfun_update e i \<lbrakk>t\<rbrakk>e)"
by (induct u rule: tm.induct) (auto)
lemma eval_subst_fm: "eval_fm e (fm(i::= t)) = eval_fm (finfun_update e i \<lbrakk>t\<rbrakk>e) fm"
by (nominal_induct fm avoiding: i t e rule: fm.strong_induct)
(simp_all add: eval_subst_tm finfun_update_twist fresh_at_base)
subsection\<open>Derived syntax\<close>
subsubsection\<open>Ordered pairs\<close>
definition HPair :: "tm \<Rightarrow> tm \<Rightarrow> tm"
where "HPair a b = Eats (Eats Zero (Eats (Eats Zero b) a)) (Eats (Eats Zero a) a)"
lemma HPair_eqvt [eqvt]: "(p \<bullet> HPair a b) = HPair (p \<bullet> a) (p \<bullet> b)"
by (auto simp: HPair_def)
lemma fresh_HPair [simp]: "x \<sharp> HPair a b \<longleftrightarrow> (x \<sharp> a \<and> x \<sharp> b)"
by (auto simp: HPair_def)
lemma HPair_injective_iff [iff]: "HPair a b = HPair a' b' \<longleftrightarrow> (a = a' \<and> b = b')"
by (auto simp: HPair_def)
lemma subst_tm_HPair [simp]: "subst i x (HPair a b) = HPair (subst i x a) (subst i x b)"
by (auto simp: HPair_def)
lemma eval_tm_HPair [simp]: "\<lbrakk>HPair a b\<rbrakk>e = hpair \<lbrakk>a\<rbrakk>e \<lbrakk>b\<rbrakk>e"
by (auto simp: HPair_def hpair_def)
subsubsection\<open>Ordinals\<close>
definition
SUCC :: "tm \<Rightarrow> tm" where
"SUCC x \<equiv> Eats x x"
fun ORD_OF :: "nat \<Rightarrow> tm"
where
"ORD_OF 0 = Zero"
| "ORD_OF (Suc k) = SUCC (ORD_OF k)"
lemma eval_tm_SUCC [simp]: "\<lbrakk>SUCC t\<rbrakk>e = succ \<lbrakk>t\<rbrakk>e"
by (simp add: SUCC_def succ_def)
lemma SUCC_fresh_iff [simp]: "a \<sharp> SUCC t \<longleftrightarrow> a \<sharp> t"
by (simp add: SUCC_def)
lemma SUCC_eqvt [eqvt]: "(p \<bullet> SUCC a) = SUCC (p \<bullet> a)"
by (simp add: SUCC_def)
lemma SUCC_subst [simp]: "subst i t (SUCC k) = SUCC (subst i t k)"
by (simp add: SUCC_def)
lemma eval_tm_ORD_OF [simp]: "\<lbrakk>ORD_OF n\<rbrakk>e = ord_of n"
by (induct n) auto
lemma ORD_OF_fresh [simp]: "a \<sharp> ORD_OF n"
by (induct n) (auto simp: SUCC_def)
lemma ORD_OF_eqvt [eqvt]: "(p \<bullet> ORD_OF n) = ORD_OF (p \<bullet> n)"
by (induct n) (auto simp: permute_pure SUCC_eqvt)
subsection\<open>Derived logical connectives\<close>
abbreviation Imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" (infixr "IMP" 125)
where "Imp A B \<equiv> Disj (Neg A) B"
abbreviation All :: "name \<Rightarrow> fm \<Rightarrow> fm"
where "All i A \<equiv> Neg (Ex i (Neg A))"
abbreviation All2 :: "name \<Rightarrow> tm \<Rightarrow> fm \<Rightarrow> fm" \<comment> \<open>bounded universal quantifier, for Sigma formulas\<close>
where "All2 i t A \<equiv> All i ((Var i IN t) IMP A)"
subsubsection\<open>Conjunction\<close>
definition Conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" (infixr "AND" 135)
where "Conj A B \<equiv> Neg (Disj (Neg A) (Neg B))"
lemma Conj_eqvt [eqvt]: "p \<bullet> (A AND B) = (p \<bullet> A) AND (p \<bullet> B)"
by (simp add: Conj_def)
lemma fresh_Conj [simp]: "a \<sharp> A AND B \<longleftrightarrow> (a \<sharp> A \<and> a \<sharp> B)"
by (auto simp: Conj_def)
lemma supp_Conj [simp]: "supp (A AND B) = supp A \<union> supp B"
by (auto simp: Conj_def)
lemma size_Conj [simp]: "size (A AND B) = size A + size B + 4"
by (simp add: Conj_def)
lemma Conj_injective_iff [iff]: "(A AND B) = (A' AND B') \<longleftrightarrow> (A = A' \<and> B = B')"
by (auto simp: Conj_def)
lemma subst_fm_Conj [simp]: "(A AND B)(i::=x) = (A(i::=x)) AND (B(i::=x))"
by (auto simp: Conj_def)
lemma eval_fm_Conj [simp]: "eval_fm e (Conj A B) \<longleftrightarrow> (eval_fm e A \<and> eval_fm e B)"
by (auto simp: Conj_def)
subsubsection\<open>If and only if\<close>
definition Iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" (infixr "IFF" 125)
where "Iff A B = Conj (Imp A B) (Imp B A)"
lemma Iff_eqvt [eqvt]: "p \<bullet> (A IFF B) = (p \<bullet> A) IFF (p \<bullet> B)"
by (simp add: Iff_def)
lemma fresh_Iff [simp]: "a \<sharp> A IFF B \<longleftrightarrow> (a \<sharp> A \<and> a \<sharp> B)"
by (auto simp: Conj_def Iff_def)
lemma size_Iff [simp]: "size (A IFF B) = 2*(size A + size B) + 8"
by (simp add: Iff_def)
lemma Iff_injective_iff [iff]: "(A IFF B) = (A' IFF B') \<longleftrightarrow> (A = A' \<and> B = B')"
by (auto simp: Iff_def)
lemma subst_fm_Iff [simp]: "(A IFF B)(i::=x) = (A(i::=x)) IFF (B(i::=x))"
by (auto simp: Iff_def)
lemma eval_fm_Iff [simp]: "eval_fm e (Iff A B) \<longleftrightarrow> (eval_fm e A \<longleftrightarrow> eval_fm e B)"
by (auto simp: Iff_def)
section\<open>Axioms and Theorems\<close>
subsection\<open>Logical axioms\<close>
inductive_set boolean_axioms :: "fm set"
where
Ident: "A IMP A \<in> boolean_axioms"
| DisjI1: "A IMP (A OR B) \<in> boolean_axioms"
| DisjCont: "(A OR A) IMP A \<in> boolean_axioms"
| DisjAssoc: "(A OR (B OR C)) IMP ((A OR B) OR C) \<in> boolean_axioms"
| DisjConj: "(C OR A) IMP (((Neg C) OR B) IMP (A OR B)) \<in> boolean_axioms"
lemma boolean_axioms_hold: "A \<in> boolean_axioms \<Longrightarrow> eval_fm e A"
by (induct rule: boolean_axioms.induct, auto)
inductive_set special_axioms :: "fm set" where
I: "A(i::=x) IMP (Ex i A) \<in> special_axioms"
lemma special_axioms_hold: "A \<in> special_axioms \<Longrightarrow> eval_fm e A"
by (induct rule: special_axioms.induct, auto) (metis eval_subst_fm)
inductive_set induction_axioms :: "fm set" where
ind:
"atom (j::name) \<sharp> (i,A)
\<Longrightarrow> A(i::=Zero) IMP ((All i (All j (A IMP (A(i::= Var j) IMP A(i::= Eats(Var i)(Var j))))))
IMP (All i A))
\<in> induction_axioms"
lemma twist_forget_eval_fm [simp]:
"atom j \<sharp> (i, A)
\<Longrightarrow> eval_fm (finfun_update (finfun_update (finfun_update e i x) j y) i z) A =
eval_fm (finfun_update e i z) A"
by (metis finfun_update_twice finfun_update_twist forget_eval_fm fresh_Pair)
lemma induction_axioms_hold: "A \<in> induction_axioms \<Longrightarrow> eval_fm e A"
by (induction rule: induction_axioms.induct) (auto simp: eval_subst_fm intro: hf_induct_ax)
subsection \<open>Concrete variables\<close>
declare Abs_name_inject[simp]
abbreviation
"X0 \<equiv> Abs_name (Atom (Sort ''SyntaxN.name'' []) 0)"
abbreviation
"X1 \<equiv> Abs_name (Atom (Sort ''SyntaxN.name'' []) (Suc 0))"
\<comment> \<open>We prefer @{term "Suc 0"} because simplification will transform 1 to that form anyway.\<close>
abbreviation
"X2 \<equiv> Abs_name (Atom (Sort ''SyntaxN.name'' []) 2)"
abbreviation
"X3 \<equiv> Abs_name (Atom (Sort ''SyntaxN.name'' []) 3)"
abbreviation
"X4 \<equiv> Abs_name (Atom (Sort ''SyntaxN.name'' []) 4)"
subsection\<open>The HF axioms\<close>
definition HF1 :: fm where \<comment> \<open>the axiom @{term"z=0 \<longleftrightarrow> (\<forall>x. \<not> x \<^bold>\<in> z)"}\<close>
"HF1 = (Var X0 EQ Zero) IFF (All X1 (Neg (Var X1 IN Var X0)))"
lemma HF1_holds: "eval_fm e HF1"
by (auto simp: HF1_def)
definition HF2 :: fm where \<comment> \<open>the axiom @{term"z = x \<triangleleft> y \<longleftrightarrow> (\<forall>u. u \<^bold>\<in> z \<longleftrightarrow> u \<^bold>\<in> x | u=y)"}\<close>
"HF2 \<equiv> Var X0 EQ Eats (Var X1) (Var X2) IFF
All X3 (Var X3 IN Var X0 IFF Var X3 IN Var X1 OR Var X3 EQ Var X2)"
lemma HF2_holds: "eval_fm e HF2"
by (auto simp: HF2_def)
definition HF_axioms where "HF_axioms = {HF1, HF2}"
lemma HF_axioms_hold: "A \<in> HF_axioms \<Longrightarrow> eval_fm e A"
by (auto simp: HF_axioms_def HF1_holds HF2_holds)
subsection\<open>Equality axioms\<close>
definition refl_ax :: fm where
"refl_ax = Var X1 EQ Var X1"
lemma refl_ax_holds: "eval_fm e refl_ax"
by (auto simp: refl_ax_def)
definition eq_cong_ax :: fm where
"eq_cong_ax = ((Var X1 EQ Var X2) AND (Var X3 EQ Var X4)) IMP
((Var X1 EQ Var X3) IMP (Var X2 EQ Var X4))"
lemma eq_cong_ax_holds: "eval_fm e eq_cong_ax"
by (auto simp: Conj_def eq_cong_ax_def)
definition mem_cong_ax :: fm where
"mem_cong_ax = ((Var X1 EQ Var X2) AND (Var X3 EQ Var X4)) IMP
((Var X1 IN Var X3) IMP (Var X2 IN Var X4))"
lemma mem_cong_ax_holds: "eval_fm e mem_cong_ax"
by (auto simp: Conj_def mem_cong_ax_def)
definition eats_cong_ax :: fm where
"eats_cong_ax = ((Var X1 EQ Var X2) AND (Var X3 EQ Var X4)) IMP
((Eats (Var X1) (Var X3)) EQ (Eats (Var X2) (Var X4)))"
lemma eats_cong_ax_holds: "eval_fm e eats_cong_ax"
by (auto simp: Conj_def eats_cong_ax_def)
definition equality_axioms :: "fm set" where
"equality_axioms = {refl_ax, eq_cong_ax, mem_cong_ax, eats_cong_ax}"
lemma equality_axioms_hold: "A \<in> equality_axioms \<Longrightarrow> eval_fm e A"
by (auto simp: equality_axioms_def refl_ax_holds eq_cong_ax_holds mem_cong_ax_holds eats_cong_ax_holds)
subsection\<open>The proof system\<close>
text\<open>This arbitrary additional axiom generalises the statements of the incompleteness theorems
and other results to any formal system stronger than the HF theory. The additional axiom
could be the conjunction of any finite number of assertions. Any more general extension
must be a form that can be formalised for the proof predicate.\<close>
consts extra_axiom :: fm
specification (extra_axiom)
extra_axiom_holds: "eval_fm e extra_axiom"
by (rule exI [where x = "Zero IN Eats Zero Zero"], auto)
inductive hfthm :: "fm set \<Rightarrow> fm \<Rightarrow> bool" (infixl "\<turnstile>" 55)
where
Hyp: "A \<in> H \<Longrightarrow> H \<turnstile> A"
| Extra: "H \<turnstile> extra_axiom"
| Bool: "A \<in> boolean_axioms \<Longrightarrow> H \<turnstile> A"
| Eq: "A \<in> equality_axioms \<Longrightarrow> H \<turnstile> A"
| Spec: "A \<in> special_axioms \<Longrightarrow> H \<turnstile> A"
| HF: "A \<in> HF_axioms \<Longrightarrow> H \<turnstile> A"
| Ind: "A \<in> induction_axioms \<Longrightarrow> H \<turnstile> A"
| MP: "H \<turnstile> A IMP B \<Longrightarrow> H' \<turnstile> A \<Longrightarrow> H \<union> H' \<turnstile> B"
| Exists: "H \<turnstile> A IMP B \<Longrightarrow> atom i \<sharp> B \<Longrightarrow> \<forall>C \<in> H. atom i \<sharp> C \<Longrightarrow> H \<turnstile> (Ex i A) IMP B"
text\<open>Soundness theorem!\<close>
theorem hfthm_sound: assumes "H \<turnstile> A" shows "(\<forall>B\<in>H. eval_fm e B) \<Longrightarrow> eval_fm e A"
using assms
proof (induct arbitrary: e)
case (Hyp A H) thus ?case
by auto
next
case (Extra H) thus ?case
by (metis extra_axiom_holds)
next
case (Bool A H) thus ?case
by (metis boolean_axioms_hold)
next
case (Eq A H) thus ?case
by (metis equality_axioms_hold)
next
case (Spec A H) thus ?case
by (metis special_axioms_hold)
next
case (HF A H) thus ?case
by (metis HF_axioms_hold)
next
case (Ind A H) thus ?case
by (metis induction_axioms_hold)
next
case (MP H A B H') thus ?case
by auto
next
case (Exists H A B i e) thus ?case
by auto (metis forget_eval_fm)
qed
subsection\<open>Derived rules of inference\<close>
lemma contraction: "insert A (insert A H) \<turnstile> B \<Longrightarrow> insert A H \<turnstile> B"
by (metis insert_absorb2)
lemma thin_Un: "H \<turnstile> A \<Longrightarrow> H \<union> H' \<turnstile> A"
by (metis Bool MP boolean_axioms.Ident sup_commute)
lemma thin: "H \<turnstile> A \<Longrightarrow> H \<subseteq> H' \<Longrightarrow> H' \<turnstile> A"
by (metis Un_absorb1 thin_Un)
lemma thin0: "{} \<turnstile> A \<Longrightarrow> H \<turnstile> A"
by (metis sup_bot_left thin_Un)
lemma thin1: "H \<turnstile> B \<Longrightarrow> insert A H \<turnstile> B"
by (metis subset_insertI thin)
lemma thin2: "insert A1 H \<turnstile> B \<Longrightarrow> insert A1 (insert A2 H) \<turnstile> B"
by (blast intro: thin)
lemma thin3: "insert A1 (insert A2 H) \<turnstile> B \<Longrightarrow> insert A1 (insert A2 (insert A3 H)) \<turnstile> B"
by (blast intro: thin)
lemma thin4:
"insert A1 (insert A2 (insert A3 H)) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 H))) \<turnstile> B"
by (blast intro: thin)
lemma rotate2: "insert A2 (insert A1 H) \<turnstile> B \<Longrightarrow> insert A1 (insert A2 H) \<turnstile> B"
by (blast intro: thin)
lemma rotate3: "insert A3 (insert A1 (insert A2 H)) \<turnstile> B \<Longrightarrow> insert A1 (insert A2 (insert A3 H)) \<turnstile> B"
by (blast intro: thin)
lemma rotate4:
"insert A4 (insert A1 (insert A2 (insert A3 H))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 H))) \<turnstile> B"
by (blast intro: thin)
lemma rotate5:
"insert A5 (insert A1 (insert A2 (insert A3 (insert A4 H)))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 (insert A5 H)))) \<turnstile> B"
by (blast intro: thin)
lemma rotate6:
"insert A6 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 H))))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 H))))) \<turnstile> B"
by (blast intro: thin)
lemma rotate7:
"insert A7 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 H)))))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 H)))))) \<turnstile> B"
by (blast intro: thin)
lemma rotate8:
"insert A8 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 H))))))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 H))))))) \<turnstile> B"
by (blast intro: thin)
lemma rotate9:
"insert A9 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 H)))))))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 H)))))))) \<turnstile> B"
by (blast intro: thin)
lemma rotate10:
"insert A10 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 H))))))))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 H))))))))) \<turnstile> B"
by (blast intro: thin)
lemma rotate11:
"insert A11 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 H)))))))))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 H)))))))))) \<turnstile> B"
by (blast intro: thin)
lemma rotate12:
"insert A12 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 H))))))))))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 H))))))))))) \<turnstile> B"
by (blast intro: thin)
lemma rotate13:
"insert A13 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 H)))))))))))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 (insert A13 H)))))))))))) \<turnstile> B"
by (blast intro: thin)
lemma rotate14:
"insert A14 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 (insert A13 H))))))))))))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 (insert A13 (insert A14 H))))))))))))) \<turnstile> B"
by (blast intro: thin)
lemma rotate15:
"insert A15 (insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 (insert A13 (insert A14 H)))))))))))))) \<turnstile> B
\<Longrightarrow> insert A1 (insert A2 (insert A3 (insert A4 (insert A5 (insert A6 (insert A7 (insert A8 (insert A9 (insert A10 (insert A11 (insert A12 (insert A13 (insert A14 (insert A15 H)))))))))))))) \<turnstile> B"
by (blast intro: thin)
lemma MP_same: "H \<turnstile> A IMP B \<Longrightarrow> H \<turnstile> A \<Longrightarrow> H \<turnstile> B"
by (metis MP Un_absorb)
lemma MP_thin: "HA \<turnstile> A IMP B \<Longrightarrow> HB \<turnstile> A \<Longrightarrow> HA \<union> HB \<subseteq> H \<Longrightarrow> H \<turnstile> B"
by (metis MP_same le_sup_iff thin)
lemma MP_null: "{} \<turnstile> A IMP B \<Longrightarrow> H \<turnstile> A \<Longrightarrow> H \<turnstile> B"
by (metis MP_same thin0)
lemma Disj_commute: "H \<turnstile> B OR A \<Longrightarrow> H \<turnstile> A OR B"
using DisjConj [of B A B] Ident [of B]
by (metis Bool MP_same)
lemma S: assumes "H \<turnstile> A IMP (B IMP C)" "H' \<turnstile> A IMP B" shows "H \<union> H' \<turnstile> A IMP C"
proof -
have "H' \<union> H \<turnstile> (Neg A) OR (C OR (Neg A))"
by (metis Bool MP MP_same boolean_axioms.DisjConj Disj_commute DisjAssoc assms)
thus ?thesis
by (metis Bool Disj_commute Un_commute MP_same DisjAssoc DisjCont DisjI1)
qed
lemma Assume: "insert A H \<turnstile> A"
by (metis Hyp insertI1)
lemmas AssumeH = Assume Assume [THEN rotate2] Assume [THEN rotate3] Assume [THEN rotate4] Assume [THEN rotate5]
Assume [THEN rotate6] Assume [THEN rotate7] Assume [THEN rotate8] Assume [THEN rotate9] Assume [THEN rotate10]
Assume [THEN rotate11] Assume [THEN rotate12]
declare AssumeH [intro!]
lemma Imp_triv_I: "H \<turnstile> B \<Longrightarrow> H \<turnstile> A IMP B"
by (metis Bool Disj_commute MP_same boolean_axioms.DisjI1)
lemma DisjAssoc1: "H \<turnstile> A OR (B OR C) \<Longrightarrow> H \<turnstile> (A OR B) OR C"
by (metis Bool MP_same boolean_axioms.DisjAssoc)
lemma DisjAssoc2: "H \<turnstile> (A OR B) OR C \<Longrightarrow> H \<turnstile> A OR (B OR C)"
by (metis DisjAssoc1 Disj_commute)
lemma Disj_commute_Imp: "H \<turnstile> (B OR A) IMP (A OR B)"
using DisjConj [of B A B] Ident [of B]
by (metis Bool DisjAssoc2 Disj_commute MP_same)
lemma Disj_Semicong_1: "H \<turnstile> A OR C \<Longrightarrow> H \<turnstile> A IMP B \<Longrightarrow> H \<turnstile> B OR C"
using DisjConj [of A C B]
by (metis Bool Disj_commute MP_same)
lemma Imp_Imp_commute: "H \<turnstile> B IMP (A IMP C) \<Longrightarrow> H \<turnstile> A IMP (B IMP C)"
by (metis DisjAssoc1 DisjAssoc2 Disj_Semicong_1 Disj_commute_Imp)
subsection\<open>The Deduction Theorem\<close>
lemma deduction_Diff: assumes "H \<turnstile> B" shows "H - {C} \<turnstile> C IMP B"
using assms
proof (induct)
case (Hyp A H) thus ?case
by (metis Bool Imp_triv_I boolean_axioms.Ident hfthm.Hyp member_remove remove_def)
next
case (Extra H) thus ?case
by (metis Imp_triv_I hfthm.Extra)
next
case (Bool A H) thus ?case
by (metis Imp_triv_I hfthm.Bool)
next
case (Eq A H) thus ?case
by (metis Imp_triv_I hfthm.Eq)
next
case (Spec A H) thus ?case
by (metis Imp_triv_I hfthm.Spec)
next
case (HF A H) thus ?case
by (metis Imp_triv_I hfthm.HF)
next
case (Ind A H) thus ?case
by (metis Imp_triv_I hfthm.Ind)
next
case (MP H A B H')
hence "(H-{C}) \<union> (H'-{C}) \<turnstile> Imp C B"
by (simp add: S)
thus ?case
by (metis Un_Diff)
next
case (Exists H A B i) show ?case
proof (cases "C \<in> H")
case True
hence "atom i \<sharp> C" using Exists by auto
moreover have "H - {C} \<turnstile> A IMP C IMP B" using Exists
by (metis Imp_Imp_commute)
ultimately have "H - {C} \<turnstile> (Ex i A) IMP C IMP B" using Exists
by (metis fm.fresh(3) fm.fresh(4) hfthm.Exists member_remove remove_def)
thus ?thesis
by (metis Imp_Imp_commute)
next
case False
hence "H - {C} = H" by auto
thus ?thesis using Exists
by (metis Imp_triv_I hfthm.Exists)
qed
qed
theorem Imp_I [intro!]: "insert A H \<turnstile> B \<Longrightarrow> H \<turnstile> A IMP B"
by (metis Diff_insert_absorb Imp_triv_I deduction_Diff insert_absorb)
lemma anti_deduction: "H \<turnstile> A IMP B \<Longrightarrow> insert A H \<turnstile> B"
by (metis Assume MP_same thin1)
subsection\<open>Cut rules\<close>
lemma cut: "H \<turnstile> A \<Longrightarrow> insert A H' \<turnstile> B \<Longrightarrow> H \<union> H' \<turnstile> B"
by (metis MP Un_commute Imp_I)
lemma cut_same: "H \<turnstile> A \<Longrightarrow> insert A H \<turnstile> B \<Longrightarrow> H \<turnstile> B"
by (metis Un_absorb cut)
lemma cut_thin: "HA \<turnstile> A \<Longrightarrow> insert A HB \<turnstile> B \<Longrightarrow> HA \<union> HB \<subseteq> H \<Longrightarrow> H \<turnstile> B"
by (metis thin cut)
lemma cut0: "{} \<turnstile> A \<Longrightarrow> insert A H \<turnstile> B \<Longrightarrow> H \<turnstile> B"
by (metis cut_same thin0)
lemma cut1: "{A} \<turnstile> B \<Longrightarrow> H \<turnstile> A \<Longrightarrow> H \<turnstile> B"
by (metis cut sup_bot_right)
lemma rcut1: "{A} \<turnstile> B \<Longrightarrow> insert B H \<turnstile> C \<Longrightarrow> insert A H \<turnstile> C"
by (metis Assume cut1 cut_same rotate2 thin1)
lemma cut2: "\<lbrakk>{A,B} \<turnstile> C; H \<turnstile> A; H \<turnstile> B\<rbrakk> \<Longrightarrow> H \<turnstile> C"
by (metis Un_empty_right Un_insert_right cut cut_same)
lemma rcut2: "{A,B} \<turnstile> C \<Longrightarrow> insert C H \<turnstile> D \<Longrightarrow> H \<turnstile> B \<Longrightarrow> insert A H \<turnstile> D"
by (metis Assume cut2 cut_same insert_commute thin1)
lemma cut3: "\<lbrakk>{A,B,C} \<turnstile> D; H \<turnstile> A; H \<turnstile> B; H \<turnstile> C\<rbrakk> \<Longrightarrow> H \<turnstile> D"
by (metis MP_same cut2 Imp_I)
lemma cut4: "\<lbrakk>{A,B,C,D} \<turnstile> E; H \<turnstile> A; H \<turnstile> B; H \<turnstile> C; H \<turnstile> D\<rbrakk> \<Longrightarrow> H \<turnstile> E"
by (metis MP_same cut3 [of B C D] Imp_I)
section\<open>Miscellaneous logical rules\<close>
lemma Disj_I1: "H \<turnstile> A \<Longrightarrow> H \<turnstile> A OR B"
by (metis Bool MP_same boolean_axioms.DisjI1)
lemma Disj_I2: "H \<turnstile> B \<Longrightarrow> H \<turnstile> A OR B"
by (metis Disj_commute Disj_I1)
lemma Peirce: "H \<turnstile> (Neg A) IMP A \<Longrightarrow> H \<turnstile> A"
using DisjConj [of "Neg A" A A] DisjCont [of A]
by (metis Bool MP_same boolean_axioms.Ident)
lemma Contra: "insert (Neg A) H \<turnstile> A \<Longrightarrow> H \<turnstile> A"
by (metis Peirce Imp_I)
lemma Imp_Neg_I: "H \<turnstile> A IMP B \<Longrightarrow> H \<turnstile> A IMP (Neg B) \<Longrightarrow> H \<turnstile> Neg A"
by (metis DisjConj [of B "Neg A" "Neg A"] DisjCont Bool Disj_commute MP_same)
lemma NegNeg_I: "H \<turnstile> A \<Longrightarrow> H \<turnstile> Neg (Neg A)"
using DisjConj [of "Neg (Neg A)" "Neg A" "Neg (Neg A)"]
by (metis Bool Ident MP_same)
lemma NegNeg_D: "H \<turnstile> Neg (Neg A) \<Longrightarrow> H \<turnstile> A"
by (metis Disj_I1 Peirce)
lemma Neg_D: "H \<turnstile> Neg A \<Longrightarrow> H \<turnstile> A \<Longrightarrow> H \<turnstile> B"
by (metis Imp_Neg_I Imp_triv_I NegNeg_D)
lemma Disj_Neg_1: "H \<turnstile> A OR B \<Longrightarrow> H \<turnstile> Neg B \<Longrightarrow> H \<turnstile> A"
by (metis Disj_I1 Disj_Semicong_1 Disj_commute Peirce)
lemma Disj_Neg_2: "H \<turnstile> A OR B \<Longrightarrow> H \<turnstile> Neg A \<Longrightarrow> H \<turnstile> B"
by (metis Disj_Neg_1 Disj_commute)
lemma Neg_Disj_I: "H \<turnstile> Neg A \<Longrightarrow> H \<turnstile> Neg B \<Longrightarrow> H \<turnstile> Neg (A OR B)"
by (metis Bool Disj_Neg_1 MP_same boolean_axioms.Ident DisjAssoc)
lemma Conj_I [intro!]: "H \<turnstile> A \<Longrightarrow> H \<turnstile> B \<Longrightarrow> H \<turnstile> A AND B"
by (metis Conj_def NegNeg_I Neg_Disj_I)
lemma Conj_E1: "H \<turnstile> A AND B \<Longrightarrow> H \<turnstile> A"
by (metis Conj_def Bool Disj_Neg_1 NegNeg_D boolean_axioms.DisjI1)
lemma Conj_E2: "H \<turnstile> A AND B \<Longrightarrow> H \<turnstile> B"
by (metis Conj_def Bool Disj_I2 Disj_Neg_2 MP_same DisjAssoc Ident)
lemma Conj_commute: "H \<turnstile> B AND A \<Longrightarrow> H \<turnstile> A AND B"
by (metis Conj_E1 Conj_E2 Conj_I)
lemma Conj_E: assumes "insert A (insert B H) \<turnstile> C" shows "insert (A AND B) H \<turnstile> C"
apply (rule cut_same [where A=A], metis Conj_E1 Hyp insertI1)
by (metis (full_types) AssumeH(2) Conj_E2 assms cut_same [where A=B] insert_commute thin2)
lemmas Conj_EH = Conj_E Conj_E [THEN rotate2] Conj_E [THEN rotate3] Conj_E [THEN rotate4] Conj_E [THEN rotate5]
Conj_E [THEN rotate6] Conj_E [THEN rotate7] Conj_E [THEN rotate8] Conj_E [THEN rotate9] Conj_E [THEN rotate10]
declare Conj_EH [intro!]
lemma Neg_I0: assumes "(\<And>B. atom i \<sharp> B \<Longrightarrow> insert A H \<turnstile> B)" shows "H \<turnstile> Neg A"
by (rule Imp_Neg_I [where B = "Zero IN Zero"]) (auto simp: assms)
lemma Neg_mono: "insert A H \<turnstile> B \<Longrightarrow> insert (Neg B) H \<turnstile> Neg A"
by (rule Neg_I0) (metis Hyp Neg_D insert_commute insertI1 thin1)
lemma Conj_mono: "insert A H \<turnstile> B \<Longrightarrow> insert C H \<turnstile> D \<Longrightarrow> insert (A AND C) H \<turnstile> B AND D"
by (metis Conj_E1 Conj_E2 Conj_I Hyp Un_absorb2 cut insertI1 subset_insertI)
lemma Disj_mono:
assumes "insert A H \<turnstile> B" "insert C H \<turnstile> D" shows "insert (A OR C) H \<turnstile> B OR D"
proof -
{ fix A B C H
have "insert (A OR C) H \<turnstile> (A IMP B) IMP C OR B"
by (metis Bool Hyp MP_same boolean_axioms.DisjConj insertI1)
hence "insert A H \<turnstile> B \<Longrightarrow> insert (A OR C) H \<turnstile> C OR B"
by (metis MP_same Un_absorb Un_insert_right Imp_I thin_Un)
}
thus ?thesis
by (metis cut_same assms thin2)
qed
lemma Disj_E:
assumes A: "insert A H \<turnstile> C" and B: "insert B H \<turnstile> C" shows "insert (A OR B) H \<turnstile> C"
by (metis A B Disj_mono NegNeg_I Peirce)
lemmas Disj_EH = Disj_E Disj_E [THEN rotate2] Disj_E [THEN rotate3] Disj_E [THEN rotate4] Disj_E [THEN rotate5]
Disj_E [THEN rotate6] Disj_E [THEN rotate7] Disj_E [THEN rotate8] Disj_E [THEN rotate9] Disj_E [THEN rotate10]
declare Disj_EH [intro!]
lemma Contra': "insert A H \<turnstile> Neg A \<Longrightarrow> H \<turnstile> Neg A"
by (metis Contra Neg_mono)
lemma NegNeg_E [intro!]: "insert A H \<turnstile> B \<Longrightarrow> insert (Neg (Neg A)) H \<turnstile> B"
by (metis NegNeg_D Neg_mono)
declare NegNeg_E [THEN rotate2, intro!]
declare NegNeg_E [THEN rotate3, intro!]
declare NegNeg_E [THEN rotate4, intro!]
declare NegNeg_E [THEN rotate5, intro!]
declare NegNeg_E [THEN rotate6, intro!]
declare NegNeg_E [THEN rotate7, intro!]
declare NegNeg_E [THEN rotate8, intro!]
lemma Imp_E:
assumes A: "H \<turnstile> A" and B: "insert B H \<turnstile> C" shows "insert (A IMP B) H \<turnstile> C"
proof -
have "insert (A IMP B) H \<turnstile> B"
by (metis Hyp A thin1 MP_same insertI1)
thus ?thesis
by (metis cut [where B=C] Un_insert_right sup_commute sup_idem B)
qed
lemma Imp_cut:
assumes "insert C H \<turnstile> A IMP B" "{A} \<turnstile> C"
shows "H \<turnstile> A IMP B"
by (metis Contra Disj_I1 Neg_mono assms rcut1)
lemma Iff_I [intro!]: "insert A H \<turnstile> B \<Longrightarrow> insert B H \<turnstile> A \<Longrightarrow> H \<turnstile> A IFF B"
by (metis Iff_def Conj_I Imp_I)
lemma Iff_MP_same: "H \<turnstile> A IFF B \<Longrightarrow> H \<turnstile> A \<Longrightarrow> H \<turnstile> B"
by (metis Iff_def Conj_E1 MP_same)
lemma Iff_MP2_same: "H \<turnstile> A IFF B \<Longrightarrow> H \<turnstile> B \<Longrightarrow> H \<turnstile> A"
by (metis Iff_def Conj_E2 MP_same)
lemma Iff_refl [intro!]: "H \<turnstile> A IFF A"
by (metis Hyp Iff_I insertI1)
lemma Iff_sym: "H \<turnstile> A IFF B \<Longrightarrow> H \<turnstile> B IFF A"
by (metis Iff_def Conj_commute)
lemma Iff_trans: "H \<turnstile> A IFF B \<Longrightarrow> H \<turnstile> B IFF C \<Longrightarrow> H \<turnstile> A IFF C"
unfolding Iff_def
by (metis Conj_E1 Conj_E2 Conj_I Disj_Semicong_1 Disj_commute)
lemma Iff_E:
"insert A (insert B H) \<turnstile> C \<Longrightarrow> insert (Neg A) (insert (Neg B) H) \<turnstile> C \<Longrightarrow> insert (A IFF B) H \<turnstile> C"
apply (auto simp: Iff_def insert_commute)
apply (metis Disj_I1 Hyp anti_deduction insertCI)
apply (metis Assume Disj_I1 anti_deduction)
done
lemma Iff_E1:
assumes A: "H \<turnstile> A" and B: "insert B H \<turnstile> C" shows "insert (A IFF B) H \<turnstile> C"
by (metis Iff_def A B Conj_E Imp_E insert_commute thin1)
lemma Iff_E2:
assumes A: "H \<turnstile> A" and B: "insert B H \<turnstile> C" shows "insert (B IFF A) H \<turnstile> C"
by (metis Iff_def A B Bool Conj_E2 Conj_mono Imp_E boolean_axioms.Ident)
lemma Iff_MP_left: "H \<turnstile> A IFF B \<Longrightarrow> insert A H \<turnstile> C \<Longrightarrow> insert B H \<turnstile> C"
by (metis Hyp Iff_E2 cut_same insertI1 insert_commute thin1)
lemma Iff_MP_left': "H \<turnstile> A IFF B \<Longrightarrow> insert B H \<turnstile> C \<Longrightarrow> insert A H \<turnstile> C"
by (metis Iff_MP_left Iff_sym)
lemma Swap: "insert (Neg B) H \<turnstile> A \<Longrightarrow> insert (Neg A) H \<turnstile> B"
by (metis NegNeg_D Neg_mono)
lemma Cases: "insert A H \<turnstile> B \<Longrightarrow> insert (Neg A) H \<turnstile> B \<Longrightarrow> H \<turnstile> B"
by (metis Contra Neg_D Neg_mono)
lemma Neg_Conj_E: "H \<turnstile> B \<Longrightarrow> insert (Neg A) H \<turnstile> C \<Longrightarrow> insert (Neg (A AND B)) H \<turnstile> C"
by (metis Conj_I Swap thin1)
lemma Disj_CI: "insert (Neg B) H \<turnstile> A \<Longrightarrow> H \<turnstile> A OR B"
by (metis Contra Disj_I1 Disj_I2 Swap)
lemma Disj_3I: "insert (Neg A) (insert (Neg C) H) \<turnstile> B \<Longrightarrow> H \<turnstile> A OR B OR C"
by (metis Disj_CI Disj_commute insert_commute)
lemma Contrapos1: "H \<turnstile> A IMP B \<Longrightarrow> H \<turnstile> Neg B IMP Neg A"
by (metis Bool MP_same boolean_axioms.DisjConj boolean_axioms.Ident)
lemma Contrapos2: "H \<turnstile> (Neg B) IMP (Neg A) \<Longrightarrow> H \<turnstile> A IMP B"
by (metis Bool MP_same boolean_axioms.DisjConj boolean_axioms.Ident)
lemma ContraAssumeN [intro]: "B \<in> H \<Longrightarrow> insert (Neg B) H \<turnstile> A"
by (metis Hyp Swap thin1)
lemma ContraAssume: "Neg B \<in> H \<Longrightarrow> insert B H \<turnstile> A"
by (metis Disj_I1 Hyp anti_deduction)
lemma ContraProve: "H \<turnstile> B \<Longrightarrow> insert (Neg B) H \<turnstile> A"
by (metis Swap thin1)
lemma Disj_IE1: "insert B H \<turnstile> C \<Longrightarrow> insert (A OR B) H \<turnstile> A OR C"
by (metis Assume Disj_mono)
lemmas Disj_IE1H = Disj_IE1 Disj_IE1 [THEN rotate2] Disj_IE1 [THEN rotate3] Disj_IE1 [THEN rotate4] Disj_IE1 [THEN rotate5]
Disj_IE1 [THEN rotate6] Disj_IE1 [THEN rotate7] Disj_IE1 [THEN rotate8]
declare Disj_IE1H [intro!]
subsection\<open>Quantifier reasoning\<close>
lemma Ex_I: "H \<turnstile> A(i::=x) \<Longrightarrow> H \<turnstile> Ex i A"
by (metis MP_same Spec special_axioms.intros)
lemma Ex_E:
assumes "insert A H \<turnstile> B" "atom i \<sharp> B" "\<forall>C \<in> H. atom i \<sharp> C"
shows "insert (Ex i A) H \<turnstile> B"
by (metis Exists Imp_I anti_deduction assms)
lemma Ex_E_with_renaming:
assumes "insert ((i \<leftrightarrow> i') \<bullet> A) H \<turnstile> B" "atom i' \<sharp> (A,i,B)" "\<forall>C \<in> H. atom i' \<sharp> C"
shows "insert (Ex i A) H \<turnstile> B"
proof -
have "Ex i A = Ex i' ((i \<leftrightarrow> i') \<bullet> A)" using assms
apply (auto simp: Abs1_eq_iff fresh_Pair)
apply (metis flip_at_simps(2) fresh_at_base_permute_iff)+
done
thus ?thesis
by (metis Ex_E assms fresh_Pair)
qed
lemmas Ex_EH = Ex_E Ex_E [THEN rotate2] Ex_E [THEN rotate3] Ex_E [THEN rotate4] Ex_E [THEN rotate5]
Ex_E [THEN rotate6] Ex_E [THEN rotate7] Ex_E [THEN rotate8] Ex_E [THEN rotate9] Ex_E [THEN rotate10]
declare Ex_EH [intro!]
lemma Ex_mono: "insert A H \<turnstile> B \<Longrightarrow> \<forall>C \<in> H. atom i \<sharp> C \<Longrightarrow> insert (Ex i A) H \<turnstile> (Ex i B)"
by (auto simp add: intro: Ex_I [where x="Var i"])
lemma All_I [intro!]: "H \<turnstile> A \<Longrightarrow> \<forall>C \<in> H. atom i \<sharp> C \<Longrightarrow> H \<turnstile> All i A"
by (auto intro: ContraProve Neg_I0)
lemma All_D: "H \<turnstile> All i A \<Longrightarrow> H \<turnstile> A(i::=x)"
by (metis Assume Ex_I NegNeg_D Neg_mono SyntaxN.Neg cut_same)
lemma All_E: "insert (A(i::=x)) H \<turnstile> B \<Longrightarrow> insert (All i A) H \<turnstile> B"
by (metis Ex_I NegNeg_D Neg_mono SyntaxN.Neg)
lemma All_E': "H \<turnstile> All i A \<Longrightarrow> insert (A(i::=x)) H \<turnstile> B \<Longrightarrow> H \<turnstile> B"
by (metis All_D cut_same)
lemma All2_E: "\<lbrakk>atom i \<sharp> t; H \<turnstile> x IN t; insert (A(i::=x)) H \<turnstile> B\<rbrakk> \<Longrightarrow> insert (All2 i t A) H \<turnstile> B"
apply (rule All_E [where x=x], auto)
by (metis Swap thin1)
lemma All2_E': "\<lbrakk>H \<turnstile> All2 i t A; H \<turnstile> x IN t; insert (A(i::=x)) H \<turnstile> B; atom i \<sharp> t\<rbrakk> \<Longrightarrow> H \<turnstile> B"
by (metis All2_E cut_same)
subsection\<open>Congruence rules\<close>
lemma Neg_cong: "H \<turnstile> A IFF A' \<Longrightarrow> H \<turnstile> Neg A IFF Neg A'"
by (metis Iff_def Conj_E1 Conj_E2 Conj_I Contrapos1)
lemma Disj_cong: "H \<turnstile> A IFF A' \<Longrightarrow> H \<turnstile> B IFF B' \<Longrightarrow> H \<turnstile> A OR B IFF A' OR B'"
by (metis Conj_E1 Conj_E2 Disj_mono Iff_I Iff_def anti_deduction)
lemma Conj_cong: "H \<turnstile> A IFF A' \<Longrightarrow> H \<turnstile> B IFF B' \<Longrightarrow> H \<turnstile> A AND B IFF A' AND B'"
by (metis Conj_def Disj_cong Neg_cong)
lemma Imp_cong: "H \<turnstile> A IFF A' \<Longrightarrow> H \<turnstile> B IFF B' \<Longrightarrow> H \<turnstile> (A IMP B) IFF (A' IMP B')"
by (metis Disj_cong Neg_cong)
lemma Iff_cong: "H \<turnstile> A IFF A' \<Longrightarrow> H \<turnstile> B IFF B' \<Longrightarrow> H \<turnstile> (A IFF B) IFF (A' IFF B')"
by (metis Iff_def Conj_cong Imp_cong)
lemma Ex_cong: "H \<turnstile> A IFF A' \<Longrightarrow> \<forall>C \<in> H. atom i \<sharp> C \<Longrightarrow> H \<turnstile> (Ex i A) IFF (Ex i A')"
apply (rule Iff_I)
apply (metis Ex_mono Hyp Iff_MP_same Un_absorb Un_insert_right insertI1 thin_Un)
apply (metis Ex_mono Hyp Iff_MP2_same Un_absorb Un_insert_right insertI1 thin_Un)
done
lemma All_cong: "H \<turnstile> A IFF A' \<Longrightarrow> \<forall>C \<in> H. atom i \<sharp> C \<Longrightarrow> H \<turnstile> (All i A) IFF (All i A')"
by (metis Ex_cong Neg_cong)
lemma Subst: "H \<turnstile> A \<Longrightarrow> \<forall>B \<in> H. atom i \<sharp> B \<Longrightarrow> H \<turnstile> A (i::=x)"
by (metis All_D All_I)
section\<open>Equality reasoning\<close>
subsection\<open>The congruence property for @{term Eq}, and other basic properties of equality\<close>
lemma Eq_cong1: "{} \<turnstile> (t EQ t' AND u EQ u') IMP (t EQ u IMP t' EQ u')"
proof -
obtain v2::name and v3::name and v4::name
where v2: "atom v2 \<sharp> (t,X1,X3,X4)"
and v3: "atom v3 \<sharp> (t,t',X1,v2,X4)"
and v4: "atom v4 \<sharp> (t,t',u,X1,v2,v3)"
by (metis obtain_fresh)
have "{} \<turnstile> (Var X1 EQ Var X2 AND Var X3 EQ Var X4) IMP (Var X1 EQ Var X3 IMP Var X2 EQ Var X4)"
by (rule Eq) (simp add: eq_cong_ax_def equality_axioms_def)
hence "{} \<turnstile> (Var X1 EQ Var X2 AND Var X3 EQ Var X4) IMP (Var X1 EQ Var X3 IMP Var X2 EQ Var X4)"
by (drule_tac i=X1 and x="Var X1" in Subst) simp_all
hence "{} \<turnstile> (Var X1 EQ Var v2 AND Var X3 EQ Var X4) IMP (Var X1 EQ Var X3 IMP Var v2 EQ Var X4)"
by (drule_tac i=X2 and x="Var v2" in Subst) simp_all
hence "{} \<turnstile> (Var X1 EQ Var v2 AND Var v3 EQ Var X4) IMP (Var X1 EQ Var v3 IMP Var v2 EQ Var X4)"
using v2
by (drule_tac i=X3 and x="Var v3" in Subst) simp_all
hence "{} \<turnstile> (Var X1 EQ Var v2 AND Var v3 EQ Var v4) IMP (Var X1 EQ Var v3 IMP Var v2 EQ Var v4)"
using v2 v3
by (drule_tac i=X4 and x="Var v4" in Subst) simp_all
hence "{} \<turnstile> (t EQ Var v2 AND Var v3 EQ Var v4) IMP (t EQ Var v3 IMP Var v2 EQ Var v4)"
using v2 v3 v4
by (drule_tac i=X1 and x=t in Subst) simp_all
hence "{} \<turnstile> (t EQ t' AND Var v3 EQ Var v4) IMP (t EQ Var v3 IMP t' EQ Var v4)"
using v2 v3 v4
by (drule_tac i=v2 and x="t'" in Subst) simp_all
hence "{} \<turnstile> (t EQ t' AND u EQ Var v4) IMP (t EQ u IMP t' EQ Var v4)"
using v3 v4
by (drule_tac i=v3 and x=u in Subst) simp_all
thus ?thesis
using v4
by (drule_tac i=v4 and x="u'" in Subst) simp_all
qed
lemma Refl [iff]: "H \<turnstile> t EQ t"
proof -
have "{} \<turnstile> Var X1 EQ Var X1"
by (rule Eq) (simp add: equality_axioms_def refl_ax_def)
hence "{} \<turnstile> t EQ t"
by (drule_tac i=X1 and x=t in Subst) simp_all
thus ?thesis
by (metis empty_subsetI thin)
qed
text\<open>Apparently necessary in order to prove the congruence property.\<close>
lemma Sym: assumes "H \<turnstile> t EQ u" shows "H \<turnstile> u EQ t"
proof -
have "{} \<turnstile> (t EQ u AND t EQ t) IMP (t EQ t IMP u EQ t)"
by (rule Eq_cong1)
moreover have "{t EQ u} \<turnstile> t EQ u AND t EQ t"
by (metis Assume Conj_I Refl)
ultimately have "{t EQ u} \<turnstile> u EQ t"
by (metis MP_same MP Refl sup_bot_left)
thus "H \<turnstile> u EQ t" by (metis assms cut1)
qed
lemma Sym_L: "insert (t EQ u) H \<turnstile> A \<Longrightarrow> insert (u EQ t) H \<turnstile> A"
by (metis Assume Sym Un_empty_left Un_insert_left cut)
lemma Trans: assumes "H \<turnstile> x EQ y" "H \<turnstile> y EQ z" shows "H \<turnstile> x EQ z"
proof -
have "\<And>H. H \<turnstile> (x EQ x AND y EQ z) IMP (x EQ y IMP x EQ z)"
by (metis Eq_cong1 bot_least thin)
moreover have "{x EQ y, y EQ z} \<turnstile> x EQ x AND y EQ z"
by (metis Assume Conj_I Refl thin1)
ultimately have "{x EQ y, y EQ z} \<turnstile> x EQ z"
by (metis Hyp MP_same insertI1)
thus ?thesis
by (metis assms cut2)
qed
lemma Eq_cong:
assumes "H \<turnstile> t EQ t'" "H \<turnstile> u EQ u'" shows "H \<turnstile> t EQ u IFF t' EQ u'"
proof -
{ fix t t' u u'
assume "H \<turnstile> t EQ t'" "H \<turnstile> u EQ u'"
moreover have "{t EQ t', u EQ u'} \<turnstile> t EQ u IMP t' EQ u'" using Eq_cong1
by (metis Assume Conj_I MP_null insert_commute)
ultimately have "H \<turnstile> t EQ u IMP t' EQ u'"
by (metis cut2)
}
thus ?thesis
by (metis Iff_def Conj_I assms Sym)
qed
lemma Eq_Trans_E: "H \<turnstile> x EQ u \<Longrightarrow> insert (t EQ u) H \<turnstile> A \<Longrightarrow> insert (x EQ t) H \<turnstile> A"
by (metis Assume Sym_L Trans cut_same thin1 thin2)
subsection\<open>The congruence property for @{term Mem}\<close>
lemma Mem_cong1: "{} \<turnstile> (t EQ t' AND u EQ u') IMP (t IN u IMP t' IN u')"
proof -
obtain v2::name and v3::name and v4::name
where v2: "atom v2 \<sharp> (t,X1,X3,X4)"
and v3: "atom v3 \<sharp> (t,t',X1,v2,X4)"
and v4: "atom v4 \<sharp> (t,t',u,X1,v2,v3)"
by (metis obtain_fresh)
have "{} \<turnstile> (Var X1 EQ Var X2 AND Var X3 EQ Var X4) IMP (Var X1 IN Var X3 IMP Var X2 IN Var X4)"
by (metis mem_cong_ax_def equality_axioms_def insert_iff Eq)
hence "{} \<turnstile> (Var X1 EQ Var v2 AND Var X3 EQ Var X4) IMP (Var X1 IN Var X3 IMP Var v2 IN Var X4)"
by (drule_tac i=X2 and x="Var v2" in Subst) simp_all
hence "{} \<turnstile> (Var X1 EQ Var v2 AND Var v3 EQ Var X4) IMP (Var X1 IN Var v3 IMP Var v2 IN Var X4)"
using v2
by (drule_tac i=X3 and x="Var v3" in Subst) simp_all
hence "{} \<turnstile> (Var X1 EQ Var v2 AND Var v3 EQ Var v4) IMP (Var X1 IN Var v3 IMP Var v2 IN Var v4)"
using v2 v3
by (drule_tac i=X4 and x="Var v4" in Subst) simp_all
hence "{} \<turnstile> (t EQ Var v2 AND Var v3 EQ Var v4) IMP (t IN Var v3 IMP Var v2 IN Var v4)"
using v2 v3 v4
by (drule_tac i=X1 and x=t in Subst) simp_all
hence "{} \<turnstile> (t EQ t' AND Var v3 EQ Var v4) IMP (t IN Var v3 IMP t' IN Var v4)"
using v2 v3 v4
by (drule_tac i=v2 and x=t' in Subst) simp_all
hence "{} \<turnstile> (t EQ t' AND u EQ Var v4) IMP (t IN u IMP t' IN Var v4)"
using v3 v4
by (drule_tac i=v3 and x=u in Subst) simp_all
thus ?thesis
using v4
by (drule_tac i=v4 and x=u' in Subst) simp_all
qed
lemma Mem_cong:
assumes "H \<turnstile> t EQ t'" "H \<turnstile> u EQ u'" shows "H \<turnstile> t IN u IFF t' IN u'"
proof -
{ fix t t' u u'
have cong: "{t EQ t', u EQ u'} \<turnstile> t IN u IMP t' IN u'"
by (metis AssumeH(2) Conj_I MP_null Mem_cong1 insert_commute)
}
thus ?thesis
by (metis Iff_def Conj_I cut2 assms Sym)
qed
subsection\<open>The congruence properties for @{term Eats} and @{term HPair}\<close>
lemma Eats_cong1: "{} \<turnstile> (t EQ t' AND u EQ u') IMP (Eats t u EQ Eats t' u')"
proof -
obtain v2::name and v3::name and v4::name
where v2: "atom v2 \<sharp> (t,X1,X3,X4)"
and v3: "atom v3 \<sharp> (t,t',X1,v2,X4)"
and v4: "atom v4 \<sharp> (t,t',u,X1,v2,v3)"
by (metis obtain_fresh)
have "{} \<turnstile> (Var X1 EQ Var X2 AND Var X3 EQ Var X4) IMP (Eats (Var X1) (Var X3) EQ Eats (Var X2) (Var X4))"
by (metis eats_cong_ax_def equality_axioms_def insert_iff Eq)
hence "{} \<turnstile> (Var X1 EQ Var v2 AND Var X3 EQ Var X4) IMP (Eats (Var X1) (Var X3) EQ Eats (Var v2) (Var X4))"
by (drule_tac i=X2 and x="Var v2" in Subst) simp_all
hence "{} \<turnstile> (Var X1 EQ Var v2 AND Var v3 EQ Var X4) IMP (Eats (Var X1) (Var v3) EQ Eats (Var v2) (Var X4))"
using v2
by (drule_tac i=X3 and x="Var v3" in Subst) simp_all
hence "{} \<turnstile> (Var X1 EQ Var v2 AND Var v3 EQ Var v4) IMP (Eats (Var X1) (Var v3) EQ Eats (Var v2) (Var v4))"
using v2 v3
by (drule_tac i=X4 and x="Var v4" in Subst) simp_all
hence "{} \<turnstile> (t EQ Var v2 AND Var v3 EQ Var v4) IMP (Eats t (Var v3) EQ Eats (Var v2) (Var v4))"
using v2 v3 v4
by (drule_tac i=X1 and x=t in Subst) simp_all
hence "{} \<turnstile> (t EQ t' AND Var v3 EQ Var v4) IMP (Eats t (Var v3) EQ Eats t' (Var v4))"
using v2 v3 v4
by (drule_tac i=v2 and x=t' in Subst) simp_all
hence "{} \<turnstile> (t EQ t' AND u EQ Var v4) IMP (Eats t u EQ Eats t' (Var v4))"
using v3 v4
by (drule_tac i=v3 and x=u in Subst) simp_all
thus ?thesis
using v4
by (drule_tac i=v4 and x=u' in Subst) simp_all
qed
lemma Eats_cong: "\<lbrakk>H \<turnstile> t EQ t'; H \<turnstile> u EQ u'\<rbrakk> \<Longrightarrow> H \<turnstile> Eats t u EQ Eats t' u'"
by (metis Conj_I anti_deduction Eats_cong1 cut1)
lemma HPair_cong: "\<lbrakk>H \<turnstile> t EQ t'; H \<turnstile> u EQ u'\<rbrakk> \<Longrightarrow> H \<turnstile> HPair t u EQ HPair t' u'"
by (metis HPair_def Eats_cong Refl)
lemma SUCC_cong: "H \<turnstile> t EQ t' \<Longrightarrow> H \<turnstile> SUCC t EQ SUCC t'"
by (metis Eats_cong SUCC_def)
subsection\<open>Substitution for Equalities\<close>
lemma Eq_subst_tm_Iff: "{t EQ u} \<turnstile> subst i t tm EQ subst i u tm"
by (induct tm rule: tm.induct) (auto simp: Eats_cong)
lemma Eq_subst_fm_Iff: "insert (t EQ u) H \<turnstile> A(i::=t) IFF A(i::=u)"
proof -
have "{ t EQ u } \<turnstile> A(i::=t) IFF A(i::=u)"
by (nominal_induct A avoiding: i t u rule: fm.strong_induct)
(auto simp: Disj_cong Neg_cong Ex_cong Mem_cong Eq_cong Eq_subst_tm_Iff)
thus ?thesis
by (metis Assume cut1)
qed
lemma Var_Eq_subst_Iff: "insert (Var i EQ t) H \<turnstile> A(i::=t) IFF A"
by (metis Eq_subst_fm_Iff Iff_sym subst_fm_id)
lemma Var_Eq_imp_subst_Iff: "H \<turnstile> Var i EQ t \<Longrightarrow> H \<turnstile> A(i::=t) IFF A"
by (metis Var_Eq_subst_Iff cut_same)
subsection\<open>Congruence Rules for Predicates\<close>
lemma P1_cong:
fixes tms :: "tm list"
assumes "\<And>i t x. atom i \<sharp> tms \<Longrightarrow> (P t)(i::=x) = P (subst i x t)" and "H \<turnstile> x EQ x'"
shows "H \<turnstile> P x IFF P x'"
proof -
obtain i::name where i: "atom i \<sharp> tms"
by (metis obtain_fresh)
have "insert (x EQ x') H \<turnstile> (P (Var i))(i::=x) IFF (P(Var i))(i::=x')"
by (rule Eq_subst_fm_Iff)
thus ?thesis using assms i
by (metis cut_same subst.simps(2))
qed
lemma P2_cong:
fixes tms :: "tm list"
assumes sub: "\<And>i t u x. atom i \<sharp> tms \<Longrightarrow> (P t u)(i::=x) = P (subst i x t) (subst i x u)"
and eq: "H \<turnstile> x EQ x'" "H \<turnstile> y EQ y'"
shows "H \<turnstile> P x y IFF P x' y'"
proof -
have yy': "{ y EQ y' } \<turnstile> P x' y IFF P x' y'"
by (rule P1_cong [where tms="[y,x']@tms"]) (auto simp: fresh_Cons sub)
have "{ x EQ x' } \<turnstile> P x y IFF P x' y"
by (rule P1_cong [where tms="[y,x']@tms"]) (auto simp: fresh_Cons sub)
hence "{x EQ x', y EQ y'} \<turnstile> P x y IFF P x' y'"
by (metis Assume Iff_trans cut1 rotate2 yy')
thus ?thesis
by (metis cut2 eq)
qed
lemma P3_cong:
fixes tms :: "tm list"
assumes sub: "\<And>i t u v x. atom i \<sharp> tms \<Longrightarrow>
(P t u v)(i::=x) = P (subst i x t) (subst i x u) (subst i x v)"
and eq: "H \<turnstile> x EQ x'" "H \<turnstile> y EQ y'" "H \<turnstile> z EQ z'"
shows "H \<turnstile> P x y z IFF P x' y' z'"
proof -
obtain i::name where i: "atom i \<sharp> (z,z',y,y',x,x')"
by (metis obtain_fresh)
have tl: "{ y EQ y', z EQ z' } \<turnstile> P x' y z IFF P x' y' z'"
by (rule P2_cong [where tms="[z,z',y,y',x,x']@tms"]) (auto simp: fresh_Cons sub)
have hd: "{ x EQ x' } \<turnstile> P x y z IFF P x' y z"
by (rule P1_cong [where tms="[z,y,x']@tms"]) (auto simp: fresh_Cons sub)
have "{x EQ x', y EQ y', z EQ z'} \<turnstile> P x y z IFF P x' y' z'"
by (metis Assume thin1 hd [THEN cut1] tl Iff_trans)
thus ?thesis
by (rule cut3) (rule eq)+
qed
lemma P4_cong:
fixes tms :: "tm list"
assumes sub: "\<And>i t1 t2 t3 t4 x. atom i \<sharp> tms \<Longrightarrow>
(P t1 t2 t3 t4)(i::=x) = P (subst i x t1) (subst i x t2) (subst i x t3) (subst i x t4)"
and eq: "H \<turnstile> x1 EQ x1'" "H \<turnstile> x2 EQ x2'" "H \<turnstile> x3 EQ x3'" "H \<turnstile> x4 EQ x4'"
shows "H \<turnstile> P x1 x2 x3 x4 IFF P x1' x2' x3' x4'"
proof -
obtain i::name where i: "atom i \<sharp> (x4,x4',x3,x3',x2,x2',x1,x1')"
by (metis obtain_fresh)
have tl: "{ x2 EQ x2', x3 EQ x3', x4 EQ x4' } \<turnstile> P x1' x2 x3 x4 IFF P x1' x2' x3' x4'"
by (rule P3_cong [where tms="[x4,x4',x3,x3',x2,x2',x1,x1']@tms"]) (auto simp: fresh_Cons sub)
have hd: "{ x1 EQ x1' } \<turnstile> P x1 x2 x3 x4 IFF P x1' x2 x3 x4"
by (auto simp: fresh_Cons sub intro!: P1_cong [where tms="[x4,x3,x2,x1']@tms"])
have "{x1 EQ x1', x2 EQ x2', x3 EQ x3', x4 EQ x4'} \<turnstile> P x1 x2 x3 x4 IFF P x1' x2' x3' x4'"
by (metis Assume thin1 hd [THEN cut1] tl Iff_trans)
thus ?thesis
by (rule cut4) (rule eq)+
qed
section\<open>Zero and Falsity\<close>
lemma Mem_Zero_iff:
assumes "atom i \<sharp> t" shows "H \<turnstile> (t EQ Zero) IFF (All i (Neg ((Var i) IN t)))"
proof -
obtain i'::name where i': "atom i' \<sharp> (t, X0, X1, i)"
by (rule obtain_fresh)
have "{} \<turnstile> ((Var X0) EQ Zero) IFF (All X1 (Neg ((Var X1) IN (Var X0))))"
by (simp add: HF HF_axioms_def HF1_def)
then have "{} \<turnstile> (((Var X0) EQ Zero) IFF (All X1 (Neg ((Var X1) IN (Var X0)))))(X0 ::= t)"
by (rule Subst) simp
hence "{} \<turnstile> (t EQ Zero) IFF (All i' (Neg ((Var i') IN t)))" using i'
by simp
also have "... = (FRESH i'. (t EQ Zero) IFF (All i' (Neg ((Var i') IN t))))"
using i' by simp
also have "... = (t EQ Zero) IFF (All i (Neg ((Var i) IN t)))"
using assms by simp
finally show ?thesis
by (metis empty_subsetI thin)
qed
lemma Mem_Zero_E [intro!]: "insert (x IN Zero) H \<turnstile> A"
proof -
obtain i::name where "atom i \<sharp> Zero"
by (rule obtain_fresh)
hence "{} \<turnstile> All i (Neg ((Var i) IN Zero))"
by (metis Mem_Zero_iff Iff_MP_same Refl)
hence "{} \<turnstile> Neg (x IN Zero)"
by (drule_tac x=x in All_D) simp
thus ?thesis
by (metis Contrapos2 Hyp Imp_triv_I MP_same empty_subsetI insertI1 thin)
qed
declare Mem_Zero_E [THEN rotate2, intro!]
declare Mem_Zero_E [THEN rotate3, intro!]
declare Mem_Zero_E [THEN rotate4, intro!]
declare Mem_Zero_E [THEN rotate5, intro!]
declare Mem_Zero_E [THEN rotate6, intro!]
declare Mem_Zero_E [THEN rotate7, intro!]
declare Mem_Zero_E [THEN rotate8, intro!]
subsection\<open>The Formula @{term Fls}; Consistency of the Calculus\<close>
definition Fls where "Fls \<equiv> Zero IN Zero"
lemma Fls_eqvt [eqvt]: "(p \<bullet> Fls) = Fls"
by (simp add: Fls_def)
lemma Fls_fresh [simp]: "a \<sharp> Fls"
by (simp add: Fls_def)
lemma Neg_I [intro!]: "insert A H \<turnstile> Fls \<Longrightarrow> H \<turnstile> Neg A"
unfolding Fls_def
by (rule Neg_I0) (metis Mem_Zero_E cut_same)
lemma Neg_E [intro!]: "H \<turnstile> A \<Longrightarrow> insert (Neg A) H \<turnstile> Fls"
by (rule ContraProve)
declare Neg_E [THEN rotate2, intro!]
declare Neg_E [THEN rotate3, intro!]
declare Neg_E [THEN rotate4, intro!]
declare Neg_E [THEN rotate5, intro!]
declare Neg_E [THEN rotate6, intro!]
declare Neg_E [THEN rotate7, intro!]
declare Neg_E [THEN rotate8, intro!]
text\<open>We need these because Neg (A IMP B) doesn't have to be syntactically a conjunction.\<close>
lemma Neg_Imp_I [intro!]: "H \<turnstile> A \<Longrightarrow> insert B H \<turnstile> Fls \<Longrightarrow> H \<turnstile> Neg (A IMP B)"
by (metis NegNeg_I Neg_Disj_I Neg_I)
lemma Neg_Imp_E [intro!]: "insert (Neg B) (insert A H) \<turnstile> C \<Longrightarrow> insert (Neg (A IMP B)) H \<turnstile> C"
apply (rule cut_same [where A=A])
apply (metis Assume Disj_I1 NegNeg_D Neg_mono)
apply (metis Swap Imp_I rotate2 thin1)
done
declare Neg_Imp_E [THEN rotate2, intro!]
declare Neg_Imp_E [THEN rotate3, intro!]
declare Neg_Imp_E [THEN rotate4, intro!]
declare Neg_Imp_E [THEN rotate5, intro!]
declare Neg_Imp_E [THEN rotate6, intro!]
declare Neg_Imp_E [THEN rotate7, intro!]
declare Neg_Imp_E [THEN rotate8, intro!]
lemma Fls_E [intro!]: "insert Fls H \<turnstile> A"
by (metis Mem_Zero_E Fls_def)
declare Fls_E [THEN rotate2, intro!]
declare Fls_E [THEN rotate3, intro!]
declare Fls_E [THEN rotate4, intro!]
declare Fls_E [THEN rotate5, intro!]
declare Fls_E [THEN rotate6, intro!]
declare Fls_E [THEN rotate7, intro!]
declare Fls_E [THEN rotate8, intro!]
lemma truth_provable: "H \<turnstile> (Neg Fls)"
by (metis Fls_E Neg_I)
lemma ExFalso: "H \<turnstile> Fls \<Longrightarrow> H \<turnstile> A"
by (metis Neg_D truth_provable)
text\<open>Thanks to Andrei Popescu for pointing out that consistency was provable here.\<close>
proposition consistent: "\<not> {} \<turnstile> Fls"
by (meson empty_iff eval_fm.simps(4) hfthm_sound truth_provable)
subsection\<open>More properties of @{term Zero}\<close>
lemma Eq_Zero_D:
assumes "H \<turnstile> t EQ Zero" "H \<turnstile> u IN t" shows "H \<turnstile> A"
proof -
obtain i::name where i: "atom i \<sharp> t"
by (rule obtain_fresh)
with assms have an: "H \<turnstile> (All i (Neg ((Var i) IN t)))"
by (metis Iff_MP_same Mem_Zero_iff)
have "H \<turnstile> Neg (u IN t)" using All_D [OF an, of u] i
by simp
thus ?thesis using assms
by (metis Neg_D)
qed
lemma Eq_Zero_thm:
assumes "atom i \<sharp> t" shows "{All i (Neg ((Var i) IN t))} \<turnstile> t EQ Zero"
by (metis Assume Iff_MP2_same Mem_Zero_iff assms)
lemma Eq_Zero_I:
assumes insi: "insert ((Var i) IN t) H \<turnstile> Fls" and i1: "atom i \<sharp> t" and i2: "\<forall>B \<in> H. atom i \<sharp> B"
shows "H \<turnstile> t EQ Zero"
proof -
have "H \<turnstile> All i (Neg ((Var i) IN t))"
by (metis All_I Neg_I i2 insi)
thus ?thesis
by (metis cut_same cut [OF Eq_Zero_thm [OF i1] Hyp] insertCI insert_is_Un)
qed
subsection\<open>Basic properties of @{term Eats}\<close>
lemma Eq_Eats_iff:
assumes "atom i \<sharp> (z,t,u)"
shows "H \<turnstile> (z EQ Eats t u) IFF (All i (Var i IN z IFF Var i IN t OR Var i EQ u))"
proof -
obtain v1::name and v2::name and i'::name
where v1: "atom v1 \<sharp> (z,X0,X2,X3)"
and v2: "atom v2 \<sharp> (t,z,X0,v1,X3)"
and i': "atom i' \<sharp> (t,u,z,X0,v1,v2,X3)"
by (metis obtain_fresh)
have "{} \<turnstile> ((Var X0) EQ (Eats (Var X1) (Var X2))) IFF
(All X3 (Var X3 IN Var X0 IFF Var X3 IN Var X1 OR Var X3 EQ Var X2))"
by (simp add: HF HF_axioms_def HF2_def)
hence "{} \<turnstile> ((Var X0) EQ (Eats (Var X1) (Var X2))) IFF
(All X3 (Var X3 IN Var X0 IFF Var X3 IN Var X1 OR Var X3 EQ Var X2))"
by (drule_tac i=X0 and x="Var X0" in Subst) simp_all
hence "{} \<turnstile> ((Var X0) EQ (Eats (Var v1) (Var X2))) IFF
(All X3 (Var X3 IN Var X0 IFF Var X3 IN Var v1 OR Var X3 EQ Var X2))"
using v1 by (drule_tac i=X1 and x="Var v1" in Subst) simp_all
hence "{} \<turnstile> ((Var X0) EQ (Eats (Var v1) (Var v2))) IFF
(All X3 (Var X3 IN Var X0 IFF Var X3 IN Var v1 OR Var X3 EQ Var v2))"
using v1 v2 by (drule_tac i=X2 and x="Var v2" in Subst) simp_all
hence "{} \<turnstile> (((Var X0) EQ (Eats (Var v1) (Var v2))) IFF
(All X3 (Var X3 IN Var X0 IFF Var X3 IN Var v1 OR Var X3 EQ Var v2)))(X0 ::= z)"
by (rule Subst) simp
hence "{} \<turnstile> ((z EQ (Eats (Var v1) (Var v2))) IFF
(All i' (Var i' IN z IFF Var i' IN Var v1 OR Var i' EQ Var v2)))"
using v1 v2 i' by (simp add: Conj_def Iff_def)
hence "{} \<turnstile> (z EQ (Eats t (Var v2))) IFF
(All i' (Var i' IN z IFF Var i' IN t OR Var i' EQ Var v2))"
using v1 v2 i' by (drule_tac i=v1 and x=t in Subst) simp_all
hence "{} \<turnstile> (z EQ Eats t u) IFF
(All i' (Var i' IN z IFF Var i' IN t OR Var i' EQ u))"
using v1 v2 i' by (drule_tac i=v2 and x=u in Subst) simp_all
also have "... = (FRESH i'. (z EQ Eats t u) IFF (All i' (Var i' IN z IFF Var i' IN t OR Var i' EQ u)))"
using i' by simp
also have "... = (z EQ Eats t u) IFF (All i (Var i IN z IFF Var i IN t OR Var i EQ u))"
using assms i' by simp
finally show ?thesis
by (rule thin0)
qed
lemma Eq_Eats_I:
"H \<turnstile> All i (Var i IN z IFF Var i IN t OR Var i EQ u) \<Longrightarrow> atom i \<sharp> (z,t,u) \<Longrightarrow> H \<turnstile> z EQ Eats t u"
by (metis Iff_MP2_same Eq_Eats_iff)
lemma Mem_Eats_Iff:
"H \<turnstile> x IN (Eats t u) IFF x IN t OR x EQ u"
proof -
obtain i::name where "atom i \<sharp> (Eats t u, t, u)"
by (rule obtain_fresh)
thus ?thesis
using Iff_MP_same [OF Eq_Eats_iff, THEN All_D]
by auto
qed
lemma Mem_Eats_I1: "H \<turnstile> u IN t \<Longrightarrow> H \<turnstile> u IN Eats t z"
by (metis Disj_I1 Iff_MP2_same Mem_Eats_Iff)
lemma Mem_Eats_I2: "H \<turnstile> u EQ z \<Longrightarrow> H \<turnstile> u IN Eats t z"
by (metis Disj_I2 Iff_MP2_same Mem_Eats_Iff)
lemma Mem_Eats_E:
assumes A: "insert (u IN t) H \<turnstile> C" and B: "insert (u EQ z) H \<turnstile> C"
shows "insert (u IN Eats t z) H \<turnstile> C"
by (rule Mem_Eats_Iff [of _ u t z, THEN Iff_MP_left']) (metis A B Disj_E)
lemmas Mem_Eats_EH = Mem_Eats_E Mem_Eats_E [THEN rotate2] Mem_Eats_E [THEN rotate3] Mem_Eats_E [THEN rotate4] Mem_Eats_E [THEN rotate5]
Mem_Eats_E [THEN rotate6] Mem_Eats_E [THEN rotate7] Mem_Eats_E [THEN rotate8]
declare Mem_Eats_EH [intro!]
lemma Mem_SUCC_I1: "H \<turnstile> u IN t \<Longrightarrow> H \<turnstile> u IN SUCC t"
by (metis Mem_Eats_I1 SUCC_def)
lemma Mem_SUCC_I2: "H \<turnstile> u EQ t \<Longrightarrow> H \<turnstile> u IN SUCC t"
by (metis Mem_Eats_I2 SUCC_def)
lemma Mem_SUCC_Refl [simp]: "H \<turnstile> k IN SUCC k"
by (metis Mem_SUCC_I2 Refl)
lemma Mem_SUCC_E:
assumes "insert (u IN t) H \<turnstile> C" "insert (u EQ t) H \<turnstile> C" shows "insert (u IN SUCC t) H \<turnstile> C"
by (metis assms Mem_Eats_E SUCC_def)
lemmas Mem_SUCC_EH = Mem_SUCC_E Mem_SUCC_E [THEN rotate2] Mem_SUCC_E [THEN rotate3] Mem_SUCC_E [THEN rotate4] Mem_SUCC_E [THEN rotate5]
Mem_SUCC_E [THEN rotate6] Mem_SUCC_E [THEN rotate7] Mem_SUCC_E [THEN rotate8]
lemma Eats_EQ_Zero_E: "insert (Eats t u EQ Zero) H \<turnstile> A"
by (metis Assume Eq_Zero_D Mem_Eats_I2 Refl)
lemmas Eats_EQ_Zero_EH = Eats_EQ_Zero_E Eats_EQ_Zero_E [THEN rotate2] Eats_EQ_Zero_E [THEN rotate3] Eats_EQ_Zero_E [THEN rotate4] Eats_EQ_Zero_E [THEN rotate5]
Eats_EQ_Zero_E [THEN rotate6] Eats_EQ_Zero_E [THEN rotate7] Eats_EQ_Zero_E [THEN rotate8]
declare Eats_EQ_Zero_EH [intro!]
lemma Eats_EQ_Zero_E2: "insert (Zero EQ Eats t u) H \<turnstile> A"
by (metis Eats_EQ_Zero_E Sym_L)
lemmas Eats_EQ_Zero_E2H = Eats_EQ_Zero_E2 Eats_EQ_Zero_E2 [THEN rotate2] Eats_EQ_Zero_E2 [THEN rotate3] Eats_EQ_Zero_E2 [THEN rotate4] Eats_EQ_Zero_E2 [THEN rotate5]
Eats_EQ_Zero_E2 [THEN rotate6] Eats_EQ_Zero_E2 [THEN rotate7] Eats_EQ_Zero_E2 [THEN rotate8]
declare Eats_EQ_Zero_E2H [intro!]
section\<open>Bounded Quantification involving @{term Eats}\<close>
lemma All2_cong: "H \<turnstile> t EQ t' \<Longrightarrow> H \<turnstile> A IFF A' \<Longrightarrow> \<forall>C \<in> H. atom i \<sharp> C \<Longrightarrow> H \<turnstile> (All2 i t A) IFF (All2 i t' A')"
by (metis All_cong Imp_cong Mem_cong Refl)
lemma All2_Zero_E [intro!]: "H \<turnstile> B \<Longrightarrow> insert (All2 i Zero A) H \<turnstile> B"
by (rule thin1)
lemma All2_Eats_I_D:
"atom i \<sharp> (t,u) \<Longrightarrow> { All2 i t A, A(i::=u)} \<turnstile> (All2 i (Eats t u) A)"
apply (auto, auto intro!: Ex_I [where x="Var i"])
apply (metis Assume thin1 Var_Eq_subst_Iff [THEN Iff_MP_same])
done
lemma All2_Eats_I:
"\<lbrakk>atom i \<sharp> (t,u); H \<turnstile> All2 i t A; H \<turnstile> A(i::=u)\<rbrakk> \<Longrightarrow> H \<turnstile> (All2 i (Eats t u) A)"
by (rule cut2 [OF All2_Eats_I_D], auto)
lemma All2_Eats_E1:
"\<lbrakk>atom i \<sharp> (t,u); \<forall>C \<in> H. atom i \<sharp> C\<rbrakk> \<Longrightarrow> insert (All2 i (Eats t u) A) H \<turnstile> All2 i t A"
by auto (metis Assume Ex_I Imp_E Mem_Eats_I1 Neg_mono subst_fm_id)
lemma All2_Eats_E2:
"\<lbrakk>atom i \<sharp> (t,u); \<forall>C \<in> H. atom i \<sharp> C\<rbrakk> \<Longrightarrow> insert (All2 i (Eats t u) A) H \<turnstile> A(i::=u)"
by (rule All_E [where x=u]) (auto intro: ContraProve Mem_Eats_I2)
lemma All2_Eats_E:
assumes i: "atom i \<sharp> (t,u)"
and B: "insert (All2 i t A) (insert (A(i::=u)) H) \<turnstile> B"
shows "insert (All2 i (Eats t u) A) H \<turnstile> B"
using i
apply (rule cut_thin [OF All2_Eats_E2, where HB = "insert (All2 i (Eats t u) A) H"], auto)
apply (rule cut_thin [OF All2_Eats_E1 B], auto)
done
lemma All2_SUCC_I:
"atom i \<sharp> t \<Longrightarrow> H \<turnstile> All2 i t A \<Longrightarrow> H \<turnstile> A(i::=t) \<Longrightarrow> H \<turnstile> (All2 i (SUCC t) A)"
by (simp add: SUCC_def All2_Eats_I)
lemma All2_SUCC_E:
assumes "atom i \<sharp> t"
and "insert (All2 i t A) (insert (A(i::=t)) H) \<turnstile> B"
shows "insert (All2 i (SUCC t) A) H \<turnstile> B"
by (simp add: SUCC_def All2_Eats_E assms)
lemma All2_SUCC_E':
assumes "H \<turnstile> u EQ SUCC t"
and "atom i \<sharp> t" "\<forall>C \<in> H. atom i \<sharp> C"
and "insert (All2 i t A) (insert (A(i::=t)) H) \<turnstile> B"
shows "insert (All2 i u A) H \<turnstile> B"
by (metis All2_SUCC_E Iff_MP_left' Iff_refl All2_cong assms)
section\<open>Induction\<close>
lemma Ind:
assumes j: "atom (j::name) \<sharp> (i,A)"
and prems: "H \<turnstile> A(i::=Zero)" "H \<turnstile> All i (All j (A IMP (A(i::= Var j) IMP A(i::= Eats(Var i)(Var j)))))"
shows "H \<turnstile> A"
proof -
have "{A(i::=Zero), All i (All j (A IMP (A(i::= Var j) IMP A(i::= Eats(Var i)(Var j)))))} \<turnstile> All i A"
by (metis j hfthm.Ind ind anti_deduction insert_commute)
hence "H \<turnstile> (All i A)"
by (metis cut2 prems)
thus ?thesis
by (metis All_E' Assume subst_fm_id)
qed
end
diff --git a/thys/Iptables_Semantics/Semantics_Goto.thy b/thys/Iptables_Semantics/Semantics_Goto.thy
--- a/thys/Iptables_Semantics/Semantics_Goto.thy
+++ b/thys/Iptables_Semantics/Semantics_Goto.thy
@@ -1,1170 +1,1171 @@
theory Semantics_Goto
imports Main Firewall_Common "Common/List_Misc" "HOL-Library.LaTeXsugar"
begin
section\<open>Big Step Semantics with Goto\<close>
text\<open>
We extend the iptables semantics to support the goto action.
A goto directly continues processing at the start of the called chain.
It does not change the call stack.
In contrast to calls, goto does not return.
Consequently, everything behind a matching goto cannot be reached.
\<close>
text\<open>
This theory is structured as follows.
Fist, the goto semantics are introduced.
Then, we show that those semantics are deterministic.
Finally, we present two methods to remove gotos.
The first unfolds goto.
The second replaces gotos with calls.
Finally, since the goto rules makes all proofs quite ugly, we never mention the goto semantics again.
As we have shown, we can get rid of the gotos easily, thus, we stick to the nicer iptables semantics without goto.
\<close>
context
begin
subsection\<open>Semantics\<close>
private type_synonym 'a ruleset = "string \<rightharpoonup> 'a rule list"
private type_synonym ('a, 'p) matcher = "'a \<Rightarrow> 'p \<Rightarrow> bool"
qualified fun matches :: "('a, 'p) matcher \<Rightarrow> 'a match_expr \<Rightarrow> 'p \<Rightarrow> bool" where
"matches \<gamma> (MatchAnd e1 e2) p \<longleftrightarrow> matches \<gamma> e1 p \<and> matches \<gamma> e2 p" |
"matches \<gamma> (MatchNot me) p \<longleftrightarrow> \<not> matches \<gamma> me p" |
"matches \<gamma> (Match e) p \<longleftrightarrow> \<gamma> e p" |
"matches _ MatchAny _ \<longleftrightarrow> True"
(*
main:
call foo
deny-all
foo:
goto bar
bar:
[nothing]
The call returns, even if a goto is executed in the called chains. The deny-all will be executed!
Chain OUTPUT (policy ACCEPT 98 packets, 34936 bytes)
pkts bytes target prot opt in out source destination
1 84 all -- * * 0.0.0.0/0 127.42.0.1
1 84 foo all -- * * 0.0.0.0/0 127.42.0.1
1 84 all -- * * 0.0.0.0/0 127.42.0.1
Chain bar (1 references)
pkts bytes target prot opt in out source destination
Chain foo (1 references)
pkts bytes target prot opt in out source destination
1 84 bar all -- * * 0.0.0.0/0 0.0.0.0/0 [goto]
*)
qualified fun no_matching_Goto :: "('a, 'p) matcher \<Rightarrow> 'p \<Rightarrow> 'a rule list \<Rightarrow> bool" where
"no_matching_Goto _ _ [] \<longleftrightarrow> True" |
"no_matching_Goto \<gamma> p ((Rule m (Goto _))#rs) \<longleftrightarrow> \<not> matches \<gamma> m p \<and> no_matching_Goto \<gamma> p rs" |
"no_matching_Goto \<gamma> p (_#rs) \<longleftrightarrow> no_matching_Goto \<gamma> p rs"
inductive iptables_goto_bigstep :: "'a ruleset \<Rightarrow> ('a, 'p) matcher \<Rightarrow> 'p \<Rightarrow> 'a rule list \<Rightarrow> state \<Rightarrow> state \<Rightarrow> bool"
("_,_,_\<turnstile>\<^sub>g \<langle>_, _\<rangle> \<Rightarrow> _" [60,60,60,20,98,98] 89)
for \<Gamma> and \<gamma> and p where
skip: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[], t\<rangle> \<Rightarrow> t" |
accept: "matches \<gamma> m p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m Accept], Undecided\<rangle> \<Rightarrow> Decision FinalAllow" |
drop: "matches \<gamma> m p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m Drop], Undecided\<rangle> \<Rightarrow> Decision FinalDeny" |
reject: "matches \<gamma> m p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m Reject], Undecided\<rangle> \<Rightarrow> Decision FinalDeny" |
log: "matches \<gamma> m p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m Log], Undecided\<rangle> \<Rightarrow> Undecided" |
empty: "matches \<gamma> m p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m Empty], Undecided\<rangle> \<Rightarrow> Undecided" |
nomatch: "\<not> matches \<gamma> m p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m a], Undecided\<rangle> \<Rightarrow> Undecided" |
decision: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Decision X\<rangle> \<Rightarrow> Decision X" |
seq: "\<lbrakk>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1, Undecided\<rangle> \<Rightarrow> t; \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2, t\<rangle> \<Rightarrow> t'; no_matching_Goto \<gamma> p rs\<^sub>1\<rbrakk> \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1@rs\<^sub>2, Undecided\<rangle> \<Rightarrow> t'" |
call_return: "\<lbrakk> matches \<gamma> m p; \<Gamma> chain = Some (rs\<^sub>1@[Rule m' Return]@rs\<^sub>2);
matches \<gamma> m' p; \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1, Undecided\<rangle> \<Rightarrow> Undecided;
no_matching_Goto \<gamma> p rs\<^sub>1\<rbrakk> \<Longrightarrow>
\<comment> \<open>we do not support a goto in the first part if you want to return\<close>
\<comment> \<open>probably unhandled case:\<close>
\<comment> \<open>\<^verbatim>\<open>main:\<close>\<close>
\<comment> \<open>\<^verbatim>\<open> call foo\<close>\<close>
\<comment> \<open>\<^verbatim>\<open>foo:\<close>\<close>
\<comment> \<open>\<^verbatim>\<open> goto bar\<close>\<close>
\<comment> \<open>\<^verbatim>\<open>bar:\<close>\<close>
\<comment> \<open> Return //returns to \<^verbatim>\<open>call foo\<close>\<close>
\<comment> \<open>But this would be a really awkward ruleset!\<close>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Call chain)], Undecided\<rangle> \<Rightarrow> Undecided" |
call_result: "\<lbrakk> matches \<gamma> m p; \<Gamma> chain = Some rs; \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t \<rbrakk> \<Longrightarrow>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Call chain)], Undecided\<rangle> \<Rightarrow> t" | \<comment> \<open>goto handling here seems okay\<close>
goto_decision: "\<lbrakk> matches \<gamma> m p; \<Gamma> chain = Some rs; \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> Decision X \<rbrakk> \<Longrightarrow>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>(Rule m (Goto chain))#rest, Undecided\<rangle> \<Rightarrow> Decision X" |
goto_no_decision: "\<lbrakk> matches \<gamma> m p; \<Gamma> chain = Some rs; \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> Undecided \<rbrakk> \<Longrightarrow>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>(Rule m (Goto chain))#rest, Undecided\<rangle> \<Rightarrow> Undecided"
text\<open>
The semantic rules again in pretty format:
\begin{center}
@{thm[mode=Axiom] skip [no_vars]}\\[1ex]
@{thm[mode=Rule] accept [no_vars]}\\[1ex]
@{thm[mode=Rule] drop [no_vars]}\\[1ex]
@{thm[mode=Rule] reject [no_vars]}\\[1ex]
@{thm[mode=Rule] log [no_vars]}\\[1ex]
@{thm[mode=Rule] empty [no_vars]}\\[1ex]
@{thm[mode=Rule] nomatch [no_vars]}\\[1ex]
@{thm[mode=Rule] decision [no_vars]}\\[1ex]
@{thm[mode=Rule] seq [no_vars]} \\[1ex]
@{thm[mode=Rule] call_return [no_vars]}\\[1ex]
@{thm[mode=Rule] call_result [no_vars]}\\[1ex]
@{thm[mode=Rule] goto_decision [no_vars]}\\[1ex]
@{thm[mode=Rule] goto_no_decision [no_vars]}
\end{center}
\<close>
private lemma deny:
"matches \<gamma> m p \<Longrightarrow> a = Drop \<or> a = Reject \<Longrightarrow> iptables_goto_bigstep \<Gamma> \<gamma> p [Rule m a] Undecided (Decision FinalDeny)"
by (auto intro: drop reject)
private lemma iptables_goto_bigstep_induct
[case_names
Skip Allow Deny Log Nomatch Decision Seq Call_return Call_result Goto_Decision Goto_no_Decision,
induct pred: iptables_goto_bigstep]:
"\<lbrakk> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs,s\<rangle> \<Rightarrow> t;
\<And>t. P [] t t;
\<And>m a. matches \<gamma> m p \<Longrightarrow> a = Accept \<Longrightarrow> P [Rule m a] Undecided (Decision FinalAllow);
\<And>m a. matches \<gamma> m p \<Longrightarrow> a = Drop \<or> a = Reject \<Longrightarrow> P [Rule m a] Undecided (Decision FinalDeny);
\<And>m a. matches \<gamma> m p \<Longrightarrow> a = Log \<or> a = Empty \<Longrightarrow> P [Rule m a] Undecided Undecided;
\<And>m a. \<not> matches \<gamma> m p \<Longrightarrow> P [Rule m a] Undecided Undecided;
\<And>rs X. P rs (Decision X) (Decision X);
\<And>rs rs\<^sub>1 rs\<^sub>2 t t'. rs = rs\<^sub>1 @ rs\<^sub>2 \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1,Undecided\<rangle> \<Rightarrow> t \<Longrightarrow> P rs\<^sub>1 Undecided t \<Longrightarrow>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2,t\<rangle> \<Rightarrow> t' \<Longrightarrow> P rs\<^sub>2 t t' \<Longrightarrow> no_matching_Goto \<gamma> p rs\<^sub>1 \<Longrightarrow>
P rs Undecided t';
\<And>m a chain rs\<^sub>1 m' rs\<^sub>2. matches \<gamma> m p \<Longrightarrow> a = Call chain \<Longrightarrow>
\<Gamma> chain = Some (rs\<^sub>1 @ [Rule m' Return] @ rs\<^sub>2) \<Longrightarrow>
matches \<gamma> m' p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1,Undecided\<rangle> \<Rightarrow> Undecided \<Longrightarrow>
no_matching_Goto \<gamma> p rs\<^sub>1 \<Longrightarrow> P rs\<^sub>1 Undecided Undecided \<Longrightarrow>
P [Rule m a] Undecided Undecided;
\<And>m a chain rs t. matches \<gamma> m p \<Longrightarrow> a = Call chain \<Longrightarrow> \<Gamma> chain = Some rs \<Longrightarrow>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs,Undecided\<rangle> \<Rightarrow> t \<Longrightarrow> P rs Undecided t \<Longrightarrow> P [Rule m a] Undecided t;
\<And>m a chain rs rest X. matches \<gamma> m p \<Longrightarrow> a = Goto chain \<Longrightarrow> \<Gamma> chain = Some rs \<Longrightarrow>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs,Undecided\<rangle> \<Rightarrow> (Decision X) \<Longrightarrow> P rs Undecided (Decision X) \<Longrightarrow>
P (Rule m a#rest) Undecided (Decision X);
\<And>m a chain rs rest. matches \<gamma> m p \<Longrightarrow> a = Goto chain \<Longrightarrow> \<Gamma> chain = Some rs \<Longrightarrow>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs,Undecided\<rangle> \<Rightarrow> Undecided \<Longrightarrow> P rs Undecided Undecided \<Longrightarrow>
P (Rule m a#rest) Undecided Undecided\<rbrakk> \<Longrightarrow>
P rs s t"
by (induction rule: iptables_goto_bigstep.induct) auto
subsubsection\<open>Forward reasoning\<close>
private lemma decisionD: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r, s\<rangle> \<Rightarrow> t \<Longrightarrow> s = Decision X \<Longrightarrow> t = Decision X"
by (induction rule: iptables_goto_bigstep_induct) auto
private lemma iptables_goto_bigstep_to_undecided: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, s\<rangle> \<Rightarrow> Undecided \<Longrightarrow> s = Undecided"
by (metis decisionD state.exhaust)
private lemma iptables_goto_bigstep_to_decision: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Decision Y\<rangle> \<Rightarrow> Decision X \<Longrightarrow> Y = X"
by (metis decisionD state.inject)
private lemma skipD: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r, s\<rangle> \<Rightarrow> t \<Longrightarrow> r = [] \<Longrightarrow> s = t"
by (induction rule: iptables_goto_bigstep.induct) auto
private lemma gotoD: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r, s\<rangle> \<Rightarrow> t \<Longrightarrow> r = [Rule m (Goto chain)] \<Longrightarrow> s = Undecided \<Longrightarrow> matches \<gamma> m p \<Longrightarrow>
\<exists> rs. \<Gamma> chain = Some rs \<and> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs,s\<rangle> \<Rightarrow> t"
by (induction rule: iptables_goto_bigstep.induct) (auto dest: skipD elim: list_app_singletonE)
private lemma not_no_matching_Goto_singleton_cases: "\<not> no_matching_Goto \<gamma> p [Rule m a] \<longleftrightarrow> (\<exists> chain. a = (Goto chain)) \<and> matches \<gamma> m p"
by(case_tac a) (simp_all)
private lemma no_matching_Goto_Cons: "no_matching_Goto \<gamma> p [r] \<Longrightarrow> no_matching_Goto \<gamma> p rs \<Longrightarrow> no_matching_Goto \<gamma> p (r#rs)"
by(cases r)(rename_tac m a, case_tac a, simp_all)
private lemma no_matching_Goto_head: "no_matching_Goto \<gamma> p (r#rs) \<Longrightarrow> no_matching_Goto \<gamma> p [r]"
by(cases r)(rename_tac m a, case_tac a, simp_all)
private lemma no_matching_Goto_tail: "no_matching_Goto \<gamma> p (r#rs) \<Longrightarrow> no_matching_Goto \<gamma> p rs"
by(cases r)(rename_tac m a, case_tac a, simp_all)
private lemma not_no_matching_Goto_cases:
assumes "\<not> no_matching_Goto \<gamma> p rs" "rs \<noteq> []"
shows "\<exists>rs1 m chain rs2. rs = rs1@(Rule m (Goto chain))#rs2 \<and> no_matching_Goto \<gamma> p rs1 \<and> matches \<gamma> m p"
using assms
proof(induction rs)
case Nil thus ?case by simp
next
case (Cons r rs)
note Cons_outer=this
from Cons have "\<not> no_matching_Goto \<gamma> p (r # rs)" by simp
show ?case
proof(cases rs)
case Nil
obtain m a where "r = Rule m a" by (cases r) simp
with \<open>\<not> no_matching_Goto \<gamma> p (r # rs)\<close> Nil not_no_matching_Goto_singleton_cases have "(\<exists> chain. a = (Goto chain)) \<and> matches \<gamma> m p" by metis
from this obtain chain where "a = (Goto chain)" and "matches \<gamma> m p" by blast
have "r # rs = [] @ Rule m (Goto chain) # []" "no_matching_Goto \<gamma> p []" "matches \<gamma> m p"
by (simp_all add: \<open>a = Goto chain\<close> \<open>r = Rule m a\<close> Nil \<open>matches \<gamma> m p\<close>)
thus ?thesis by blast
next
case(Cons r' rs')
with Cons_outer have "r # rs = r # r' # rs'" by simp
show ?thesis
proof(cases"no_matching_Goto \<gamma> p [r]")
case True
with \<open>\<not> no_matching_Goto \<gamma> p (r # rs)\<close> have "\<not> no_matching_Goto \<gamma> p rs" by (meson no_matching_Goto_Cons)
have "rs \<noteq> []" using Cons by simp
from Cons_outer(1)[OF \<open>\<not> no_matching_Goto \<gamma> p rs\<close> \<open>rs \<noteq> []\<close>]
obtain rs1 m chain rs2 where "rs = rs1 @ Rule m (Goto chain) # rs2" "no_matching_Goto \<gamma> p rs1" "matches \<gamma> m p" by blast
with \<open>r # rs = r # r' # rs'\<close> \<open>no_matching_Goto \<gamma> p [r]\<close> no_matching_Goto_Cons
have "r # rs = r # rs1 @ Rule m (Goto chain) # rs2 \<and> no_matching_Goto \<gamma> p (r#rs1) \<and> matches \<gamma> m p" by fast
thus ?thesis
apply(rule_tac x="r#rs1" in exI)
by auto
next
case False
obtain m a where "r = Rule m a" by (cases r) simp
with False not_no_matching_Goto_singleton_cases have "(\<exists> chain. a = (Goto chain)) \<and> matches \<gamma> m p" by metis
from this obtain chain where "a = (Goto chain)" and "matches \<gamma> m p" by blast
have "r # rs = [] @ Rule m (Goto chain) # rs" "no_matching_Goto \<gamma> p []" "matches \<gamma> m p"
by (simp_all add: \<open>a = Goto chain\<close> \<open>r = Rule m a\<close> \<open>matches \<gamma> m p\<close>)
thus ?thesis by blast
qed
qed
qed
private lemma seq_cons_Goto_Undecided:
assumes "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)], Undecided\<rangle> \<Rightarrow> Undecided"
and "\<not> matches \<gamma> m p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> Undecided"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>Rule m (Goto chain) # rs, Undecided\<rangle> \<Rightarrow> Undecided"
proof(cases "matches \<gamma> m p")
case True
from True assms have "\<exists>rs. \<Gamma> chain = Some rs \<and> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> Undecided" by(auto dest: gotoD)
with True show ?thesis using goto_no_decision by fast
next
case False
with assms have " \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)] @ rs, Undecided\<rangle> \<Rightarrow> Undecided" by(auto dest: seq)
with False show ?thesis by simp
qed
private lemma seq_cons_Goto_t:
"\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)], Undecided\<rangle> \<Rightarrow> t \<Longrightarrow> matches \<gamma> m p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>Rule m (Goto chain) # rs, Undecided\<rangle> \<Rightarrow> t"
apply(frule gotoD)
apply(simp_all)
apply(clarify)
apply(cases t)
apply(auto intro: iptables_goto_bigstep.intros)
done
private lemma no_matching_Goto_append: "no_matching_Goto \<gamma> p (rs1@rs2) \<longleftrightarrow> no_matching_Goto \<gamma> p rs1 \<and> no_matching_Goto \<gamma> p rs2"
by(induction \<gamma> p rs1 rule: no_matching_Goto.induct) (simp_all)
private lemma no_matching_Goto_append1: "no_matching_Goto \<gamma> p (rs1@rs2) \<Longrightarrow> no_matching_Goto \<gamma> p rs1"
using no_matching_Goto_append by fast
private lemma no_matching_Goto_append2: "no_matching_Goto \<gamma> p (rs1@rs2) \<Longrightarrow> no_matching_Goto \<gamma> p rs2"
using no_matching_Goto_append by fast
private lemma seq_cons:
assumes "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[r],Undecided\<rangle> \<Rightarrow> t" and "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs,t\<rangle> \<Rightarrow> t'" and "no_matching_Goto \<gamma> p [r]"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r#rs, Undecided\<rangle> \<Rightarrow> t'"
proof -
from assms have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[r] @ rs, Undecided\<rangle> \<Rightarrow> t'" by (rule seq)
thus ?thesis by simp
qed
context
notes skipD[dest] list_app_singletonE[elim]
begin
lemma acceptD: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r, s\<rangle> \<Rightarrow> t \<Longrightarrow> r = [Rule m Accept] \<Longrightarrow> matches \<gamma> m p \<Longrightarrow> s = Undecided \<Longrightarrow> t = Decision FinalAllow"
by (induction rule: iptables_goto_bigstep.induct) auto
lemma dropD: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r, s\<rangle> \<Rightarrow> t \<Longrightarrow> r = [Rule m Drop] \<Longrightarrow> matches \<gamma> m p \<Longrightarrow> s = Undecided \<Longrightarrow> t = Decision FinalDeny"
by (induction rule: iptables_goto_bigstep.induct) auto
lemma rejectD: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r, s\<rangle> \<Rightarrow> t \<Longrightarrow> r = [Rule m Reject] \<Longrightarrow> matches \<gamma> m p \<Longrightarrow> s = Undecided \<Longrightarrow> t = Decision FinalDeny"
by (induction rule: iptables_goto_bigstep.induct) auto
lemma logD: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r, s\<rangle> \<Rightarrow> t \<Longrightarrow> r = [Rule m Log] \<Longrightarrow> matches \<gamma> m p \<Longrightarrow> s = Undecided \<Longrightarrow> t = Undecided"
by (induction rule: iptables_goto_bigstep.induct) auto
lemma emptyD: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r, s\<rangle> \<Rightarrow> t \<Longrightarrow> r = [Rule m Empty] \<Longrightarrow> matches \<gamma> m p \<Longrightarrow> s = Undecided \<Longrightarrow> t = Undecided"
by (induction rule: iptables_goto_bigstep.induct) auto
lemma nomatchD: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r, s\<rangle> \<Rightarrow> t \<Longrightarrow> r = [Rule m a] \<Longrightarrow> s = Undecided \<Longrightarrow> \<not> matches \<gamma> m p \<Longrightarrow> t = Undecided"
by (induction rule: iptables_goto_bigstep.induct) auto
lemma callD:
assumes "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r, s\<rangle> \<Rightarrow> t" "r = [Rule m (Call chain)]" "s = Undecided" "matches \<gamma> m p" "\<Gamma> chain = Some rs"
obtains "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs,s\<rangle> \<Rightarrow> t"
| rs\<^sub>1 rs\<^sub>2 m' where "rs = rs\<^sub>1 @ Rule m' Return # rs\<^sub>2" "matches \<gamma> m' p" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1,s\<rangle> \<Rightarrow> Undecided" "no_matching_Goto \<gamma> p rs\<^sub>1" "t = Undecided"
using assms
proof (induction r s t arbitrary: rs rule: iptables_goto_bigstep.induct)
case (seq rs\<^sub>1)
thus ?case by (cases rs\<^sub>1) auto
qed auto
end
private lemmas iptables_goto_bigstepD = skipD acceptD dropD rejectD logD emptyD nomatchD decisionD callD gotoD
private lemma seq':
assumes "rs = rs\<^sub>1 @ rs\<^sub>2" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1,s\<rangle> \<Rightarrow> t" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2,t\<rangle> \<Rightarrow> t'" and "no_matching_Goto \<gamma> p rs\<^sub>1"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs,s\<rangle> \<Rightarrow> t'"
using assms by (cases s) (auto intro: seq decision dest: decisionD)
private lemma seq'_cons: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[r],s\<rangle> \<Rightarrow> t \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs,t\<rangle> \<Rightarrow> t' \<Longrightarrow> no_matching_Goto \<gamma> p [r] \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r#rs, s\<rangle> \<Rightarrow> t'"
by (metis decision decisionD state.exhaust seq_cons)
private lemma no_matching_Goto_take: "no_matching_Goto \<gamma> p rs \<Longrightarrow> no_matching_Goto \<gamma> p (take n rs)"
apply(induction n arbitrary: rs)
apply(simp_all)
apply(rename_tac r rs)
apply(case_tac rs)
apply(simp_all)
apply(rename_tac r' rs')
apply(case_tac r')
apply(simp)
apply(rename_tac m a)
by(case_tac a) (simp_all)
private lemma seq_split:
assumes "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, s\<rangle> \<Rightarrow> t" "rs = rs\<^sub>1@rs\<^sub>2"
obtains (no_matching_Goto) t' where "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1,s\<rangle> \<Rightarrow> t'" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2,t'\<rangle> \<Rightarrow> t" "no_matching_Goto \<gamma> p rs\<^sub>1"
| (matching_Goto) "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1,s\<rangle> \<Rightarrow> t" "\<not> no_matching_Goto \<gamma> p rs\<^sub>1"
proof -
have "(\<exists>t'. \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1,s\<rangle> \<Rightarrow> t' \<and> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2,t'\<rangle> \<Rightarrow> t \<and> no_matching_Goto \<gamma> p rs\<^sub>1) \<or> (\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1,s\<rangle> \<Rightarrow> t \<and> \<not> no_matching_Goto \<gamma> p rs\<^sub>1)"
using assms
proof (induction rs s t arbitrary: rs\<^sub>1 rs\<^sub>2 rule: iptables_goto_bigstep_induct)
case Skip thus ?case by (auto intro: iptables_goto_bigstep.intros simp add: accept)
next
case Allow thus ?case by (cases rs\<^sub>1) (auto intro: iptables_goto_bigstep.intros simp add: accept)
next
case Deny thus ?case by (cases rs\<^sub>1) (auto intro: iptables_goto_bigstep.intros simp add: deny)
next
case Log thus ?case by (cases rs\<^sub>1) (auto intro: iptables_goto_bigstep.intros simp add: log empty)
next
- case Nomatch thus ?case by (cases rs\<^sub>1) (auto intro: iptables_goto_bigstep.intros simp add: not_no_matching_Goto_singleton_cases)
+ case Nomatch thus ?case by (cases rs\<^sub>1)
+ (auto intro: iptables_goto_bigstep.intros simp add: not_no_matching_Goto_singleton_cases, meson nomatch not_no_matching_Goto_singleton_cases skip)
next
case Decision thus ?case by (auto intro: iptables_goto_bigstep.intros)
next
case (Seq rs rsa rsb t t')
hence rs: "rsa @ rsb = rs\<^sub>1 @ rs\<^sub>2" by simp
note List.append_eq_append_conv_if[simp]
from rs show ?case
proof (cases rule: list_app_eq_cases)
case longer
with Seq have t1: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>take (length rsa) rs\<^sub>1, Undecided\<rangle> \<Rightarrow> t"
by simp
from Seq.IH(2)[OF longer(2)] have IH:
"(\<exists>t'a. \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>drop (length rsa) rs\<^sub>1, t\<rangle> \<Rightarrow> t'a \<and> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2, t'a\<rangle> \<Rightarrow> t' \<and> no_matching_Goto \<gamma> p (drop (length rsa) rs\<^sub>1)) \<or>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>drop (length rsa) rs\<^sub>1, t\<rangle> \<Rightarrow> t' \<and> \<not> no_matching_Goto \<gamma> p (drop (length rsa) rs\<^sub>1)" (is "?IH_no_Goto \<or> ?IH_Goto") by simp
thus ?thesis
proof(rule disjE)
assume IH: ?IH_no_Goto
from IH obtain t2
where t2a: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>drop (length rsa) rs\<^sub>1,t\<rangle> \<Rightarrow> t2"
and rs_part2: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2,t2\<rangle> \<Rightarrow> t'"
and "no_matching_Goto \<gamma> p (drop (length rsa) rs\<^sub>1)"
by blast
with t1 rs_part2 have rs_part1: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>take (length rsa) rs\<^sub>1 @ drop (length rsa) rs\<^sub>1, Undecided\<rangle> \<Rightarrow> t2"
using Seq.hyps(4) longer(1) seq by fastforce
have "no_matching_Goto \<gamma> p (take (length rsa) rs\<^sub>1 @ drop (length rsa) rs\<^sub>1)"
using Seq.hyps(4) \<open>no_matching_Goto \<gamma> p (drop (length rsa) rs\<^sub>1)\<close> longer(1)
no_matching_Goto_append by fastforce
with Seq rs_part1 rs_part2 show ?thesis by auto
next
assume ?IH_Goto
thus ?thesis by (metis Seq.hyps(2) Seq.hyps(4) append_take_drop_id longer(1) no_matching_Goto_append2 seq')
qed
next
case shorter
from shorter rs have rsa': "rsa = rs\<^sub>1 @ take (length rsa - length rs\<^sub>1) rs\<^sub>2"
by (metis append_eq_conv_conj length_drop)
from shorter rs have rsb': "rsb = drop (length rsa - length rs\<^sub>1) rs\<^sub>2"
by (metis append_eq_conv_conj length_drop)
from Seq.hyps(4) rsa' no_matching_Goto_append2 have
no_matching_Goto_rs2: "no_matching_Goto \<gamma> p (take (length rsa - length rs\<^sub>1) rs\<^sub>2)" by metis
from rsb' Seq.hyps have t2: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>drop (length rsa - length rs\<^sub>1) rs\<^sub>2,t\<rangle> \<Rightarrow> t'"
by blast
from Seq.IH(1)[OF rsa'] have IH:
"(\<exists>t'. \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1, Undecided\<rangle> \<Rightarrow> t' \<and> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>take (length rsa - length rs\<^sub>1) rs\<^sub>2, t'\<rangle> \<Rightarrow> t \<and> no_matching_Goto \<gamma> p rs\<^sub>1) \<or>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1, Undecided\<rangle> \<Rightarrow> t \<and> \<not> no_matching_Goto \<gamma> p rs\<^sub>1" (is "?IH_no_Goto \<or> ?IH_Goto") by simp
thus ?thesis
proof(rule disjE)
assume IH: ?IH_no_Goto
from IH obtain t1
where t1a: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1,Undecided\<rangle> \<Rightarrow> t1"
and t1b: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>take (length rsa - length rs\<^sub>1) rs\<^sub>2,t1\<rangle> \<Rightarrow> t"
and "no_matching_Goto \<gamma> p rs\<^sub>1"
by blast
from no_matching_Goto_rs2 t2 seq' t1b have rs2: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2,t1\<rangle> \<Rightarrow> t'"
by fastforce
from t1a rs2 \<open>no_matching_Goto \<gamma> p rs\<^sub>1\<close> show ?thesis by fast
next
assume ?IH_Goto
thus ?thesis by (metis Seq.hyps(4) no_matching_Goto_append1 rsa')
qed
qed
next
case Call_return
hence "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1, Undecided\<rangle> \<Rightarrow> Undecided" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2, Undecided\<rangle> \<Rightarrow> Undecided"
by (case_tac [!] rs\<^sub>1) (auto intro: iptables_goto_bigstep.skip iptables_goto_bigstep.call_return)
thus ?case by fast
next
case (Call_result _ _ _ _ t)
show ?case
proof (cases rs\<^sub>1)
case Nil
with Call_result have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1, Undecided\<rangle> \<Rightarrow> Undecided" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2, Undecided\<rangle> \<Rightarrow> t"
by (auto intro: iptables_goto_bigstep.intros)
thus ?thesis using local.Nil by auto
next
case Cons
with Call_result have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1, Undecided\<rangle> \<Rightarrow> t" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2, t\<rangle> \<Rightarrow> t"
by (auto intro: iptables_goto_bigstep.intros)
thus ?thesis by fast
qed
next
case (Goto_Decision m a chain rs rest X)
thus ?case
proof (cases rs\<^sub>1)
case Nil
with Goto_Decision have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1, Undecided\<rangle> \<Rightarrow> Undecided" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2, Undecided\<rangle> \<Rightarrow> Decision X"
by (auto intro: iptables_goto_bigstep.intros)
thus ?thesis using local.Nil by auto
next
case Cons
with Goto_Decision have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1, Undecided\<rangle> \<Rightarrow> Decision X" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2, Decision X\<rangle> \<Rightarrow> Decision X"
by (auto intro: iptables_goto_bigstep.intros)
thus ?thesis by fast
qed
next
case (Goto_no_Decision m a chain rs rest rs\<^sub>1)
from Goto_no_Decision have rs1rs2: "Rule m (Goto chain) # rest = rs\<^sub>1 @ rs\<^sub>2" by simp
from goto_no_decision[OF Goto_no_Decision(1)] Goto_no_Decision(3) Goto_no_Decision(4)
have x: "\<And>rest. \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>Rule m (Goto chain) # rest, Undecided\<rangle> \<Rightarrow> Undecided" by simp
show ?case
proof (cases rs\<^sub>1)
case Nil
with Goto_no_Decision have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1, Undecided\<rangle> \<Rightarrow> Undecided" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2, Undecided\<rangle> \<Rightarrow> Undecided"
by (auto intro: iptables_goto_bigstep.intros)
thus ?thesis by fast
next
case (Cons rs\<^sub>1a rs\<^sub>1s)
with rs1rs2 have "rs\<^sub>1 = Rule m (Goto chain) # (take (length rs\<^sub>1s) rest)" by simp
from Cons rs1rs2 have"rs\<^sub>2 = drop (length rs\<^sub>1s) rest" by simp
from Cons Goto_no_Decision have 1: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1, Undecided\<rangle> \<Rightarrow> Undecided"
using x by auto[1]
have 2: "\<not> no_matching_Goto \<gamma> p rs\<^sub>1"
by (simp add: Goto_no_Decision.hyps(1) \<open>rs\<^sub>1 = Rule m (Goto chain) # take (length rs\<^sub>1s) rest\<close>)
from 1 2 show ?thesis by fast
qed
qed
thus ?thesis using matching_Goto no_matching_Goto by blast
qed
private lemma seqE:
assumes "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1@rs\<^sub>2, s\<rangle> \<Rightarrow> t"
obtains (no_matching_Goto) ti where "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1,s\<rangle> \<Rightarrow> ti" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>2,ti\<rangle> \<Rightarrow> t" "no_matching_Goto \<gamma> p rs\<^sub>1"
| (matching_Goto) "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs\<^sub>1,s\<rangle> \<Rightarrow> t" "\<not> no_matching_Goto \<gamma> p rs\<^sub>1"
using assms by (force elim: seq_split)
private lemma seqE_cons:
assumes "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r#rs, s\<rangle> \<Rightarrow> t"
obtains (no_matching_Goto) ti where "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[r],s\<rangle> \<Rightarrow> ti" "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs,ti\<rangle> \<Rightarrow> t" "no_matching_Goto \<gamma> p [r]"
| (matching_Goto) "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[r],s\<rangle> \<Rightarrow> t" "\<not> no_matching_Goto \<gamma> p [r]"
using assms by (metis append_Cons append_Nil seqE)
private lemma seqE_cons_Undecided:
assumes "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r#rs, Undecided\<rangle> \<Rightarrow> t"
obtains (no_matching_Goto) ti where "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[r],Undecided\<rangle> \<Rightarrow> ti" and "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs,ti\<rangle> \<Rightarrow> t" and "no_matching_Goto \<gamma> p [r]"
| (matching_Goto) m chain rs' where "r = Rule m (Goto chain)" and "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)],Undecided\<rangle> \<Rightarrow> t" and "matches \<gamma> m p" "\<Gamma> chain = Some rs'"
using assms
proof(cases rule: seqE_cons)
case no_matching_Goto thus ?thesis using local.that by simp
next
case matching_Goto
from this(2) not_no_matching_Goto_singleton_cases[of \<gamma> p "(get_match r)" "(get_action r)", simplified] have
"((\<exists>chain. (get_action r) = Goto chain) \<and> matches \<gamma> (get_match r) p)" by simp
from this obtain chain m where r: "r = Rule m (Goto chain)" "matches \<gamma> m p" by(cases r) auto
from matching_Goto r have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)],Undecided\<rangle> \<Rightarrow> t" by simp
from gotoD[OF matching_Goto(1)] r \<open>matches \<gamma> m p\<close> obtain rs' where "\<Gamma> chain = Some rs'" by blast
from local.that
show ?thesis using \<open>\<Gamma> chain = Some rs'\<close> \<open>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)], Undecided\<rangle> \<Rightarrow> t\<close> r(1) r(2) by blast
qed
private lemma nomatch':
assumes "\<And>r. r \<in> set rs \<Longrightarrow> \<not> matches \<gamma> (get_match r) p"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, s\<rangle> \<Rightarrow> s"
proof(cases s)
case Undecided
have "\<forall>r\<in>set rs. \<not> matches \<gamma> (get_match r) p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> Undecided"
proof(induction rs)
case Nil
thus ?case by (fast intro: skip)
next
case (Cons r rs)
hence "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[r], Undecided\<rangle> \<Rightarrow> Undecided"
by (cases r) (auto intro: nomatch)
with Cons show ?case
by (metis list.set_intros(1) list.set_intros(2) not_no_matching_Goto_singleton_cases rule.collapse seq'_cons)
qed
with assms Undecided show ?thesis by simp
qed (blast intro: decision)
private lemma no_free_return: assumes "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m Return], Undecided\<rangle> \<Rightarrow> t" and "matches \<gamma> m p" shows "False"
proof -
{ fix a s
have no_free_return_hlp: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>a,s\<rangle> \<Rightarrow> t \<Longrightarrow> matches \<gamma> m p \<Longrightarrow> s = Undecided \<Longrightarrow> a = [Rule m Return] \<Longrightarrow> False"
proof (induction rule: iptables_goto_bigstep.induct)
case (seq rs\<^sub>1)
thus ?case
by (cases rs\<^sub>1) (auto dest: skipD)
qed simp_all
} with assms show ?thesis by blast
qed
subsection\<open>Determinism\<close>
private lemma iptables_goto_bigstep_Undecided_Undecided_deterministic:
"\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> Undecided \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t \<Longrightarrow> t = Undecided"
proof(induction rs Undecided Undecided arbitrary: t rule: iptables_goto_bigstep_induct)
case Skip thus ?case by(fastforce dest: skipD logD emptyD nomatchD decisionD)
next
case Log thus ?case by(fastforce dest: skipD logD emptyD nomatchD decisionD)
next
case Nomatch thus ?case by(fastforce dest: skipD logD emptyD nomatchD decisionD)
next
case Seq thus ?case by (metis iptables_goto_bigstep_to_undecided seqE)
next
case (Call_return m a chain rs\<^sub>1 m' rs\<^sub>2)
from Call_return have " \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Call chain)], Undecided\<rangle> \<Rightarrow> Undecided"
apply(frule_tac rs\<^sub>1=rs\<^sub>1 and m'=m' and chain=chain in call_return)
by(simp_all)
with Call_return show ?case
apply simp
apply (metis callD no_free_return seqE seqE_cons)
done
next
case Call_result thus ?case by (meson callD)
next
case Goto_no_Decision thus ?case by (metis gotoD no_matching_Goto.simps(2) option.sel seqE_cons)
qed
private lemma iptables_goto_bigstep_Undecided_deterministic:
"\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t' \<Longrightarrow> t' = t"
proof(induction rs Undecided t arbitrary: t' rule: iptables_goto_bigstep_induct)
case Skip thus ?case by(fastforce dest: skipD logD emptyD nomatchD decisionD)
next
case Allow thus ?case by (auto dest: iptables_goto_bigstepD)
next
case Deny thus ?case by (auto dest: iptables_goto_bigstepD)
next
case Log thus ?case by (auto dest: iptables_goto_bigstepD)
next
case Nomatch thus ?case by (auto dest: iptables_goto_bigstepD)
next
case Seq thus ?case by (metis decisionD seqE state.exhaust)
next
case Call_return thus ?case by (meson call_return iptables_goto_bigstep_Undecided_Undecided_deterministic)
next
case Call_result thus ?case by (metis callD call_result iptables_goto_bigstep_Undecided_Undecided_deterministic)
next
case Goto_Decision thus ?case by (metis gotoD no_matching_Goto.simps(2) option.sel seqE_cons)
next
case Goto_no_Decision thus ?case by (meson goto_no_decision iptables_goto_bigstep_Undecided_Undecided_deterministic)
qed
qualified theorem iptables_goto_bigstep_deterministic: assumes "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, s\<rangle> \<Rightarrow> t" and "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, s\<rangle> \<Rightarrow> t'" shows "t = t'"
using assms
apply(cases s)
apply(simp add: iptables_goto_bigstep_Undecided_deterministic)
by(auto dest: decisionD)
subsection\<open>Matching\<close>
private lemma matches_rule_and_simp_help:
assumes "matches \<gamma> m p"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule (MatchAnd m m') a'], s\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m' a'], s\<rangle> \<Rightarrow> t" (is "?l \<longleftrightarrow>?r")
proof
assume ?l thus ?r
by (induction "[Rule (MatchAnd m m') a']" s t rule: iptables_goto_bigstep_induct)
(auto intro: iptables_goto_bigstep.intros simp: assms Cons_eq_append_conv dest: skipD)
next
assume ?r thus ?l
by (induction "[Rule m' a']" s t rule: iptables_goto_bigstep_induct)
(auto intro: iptables_goto_bigstep.intros simp: assms Cons_eq_append_conv dest: skipD)
qed
private lemma matches_MatchNot_simp:
assumes "matches \<gamma> m p"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule (MatchNot m) a], Undecided\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[], Undecided\<rangle> \<Rightarrow> t" (is "?l \<longleftrightarrow> ?r")
proof
assume ?l thus ?r
by (induction "[Rule (MatchNot m) a]" "Undecided" t rule: iptables_goto_bigstep_induct)
(auto intro: iptables_goto_bigstep.intros simp: assms Cons_eq_append_conv dest: skipD)
next
assume ?r
hence "t = Undecided"
by (metis skipD)
with assms show ?l
by (fastforce intro: nomatch)
qed
private lemma matches_MatchNotAnd_simp:
assumes "matches \<gamma> m p"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule (MatchAnd (MatchNot m) m') a], Undecided\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[], Undecided\<rangle> \<Rightarrow> t" (is "?l \<longleftrightarrow> ?r")
proof
assume ?l thus ?r
by (induction "[Rule (MatchAnd (MatchNot m) m') a]" "Undecided" t rule: iptables_goto_bigstep_induct)
(auto intro: iptables_goto_bigstep.intros simp add: assms Cons_eq_append_conv dest: skipD)
next
assume ?r
hence "t = Undecided"
by (metis skipD)
with assms show ?l
by (fastforce intro: nomatch)
qed
private lemma matches_rule_and_simp:
assumes "matches \<gamma> m p"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule (MatchAnd m m') a'], s\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m' a'], s\<rangle> \<Rightarrow> t"
proof (cases s)
case Undecided
with assms show ?thesis
by (simp add: matches_rule_and_simp_help)
next
case Decision
thus ?thesis by (metis decision decisionD)
qed
qualified definition add_match :: "'a match_expr \<Rightarrow> 'a rule list \<Rightarrow> 'a rule list" where
"add_match m rs = map (\<lambda>r. case r of Rule m' a' \<Rightarrow> Rule (MatchAnd m m') a') rs"
private lemma add_match_split: "add_match m (rs1@rs2) = add_match m rs1 @ add_match m rs2"
unfolding add_match_def
by (fact map_append)
private lemma add_match_split_fst: "add_match m (Rule m' a' # rs) = Rule (MatchAnd m m') a' # add_match m rs"
unfolding add_match_def
by simp
private lemma matches_add_match_no_matching_Goto_simp: "matches \<gamma> m p \<Longrightarrow> no_matching_Goto \<gamma> p (add_match m rs) \<Longrightarrow> no_matching_Goto \<gamma> p rs"
apply(induction rs)
apply(simp_all)
apply(rename_tac r rs)
apply(case_tac r)
apply(simp add: add_match_split_fst no_matching_Goto_tail)
apply(drule no_matching_Goto_head)
apply(rename_tac m' a')
apply(case_tac a')
apply simp_all
done
private lemma matches_add_match_no_matching_Goto_simp2: "matches \<gamma> m p \<Longrightarrow> no_matching_Goto \<gamma> p rs \<Longrightarrow> no_matching_Goto \<gamma> p (add_match m rs)"
apply(induction rs)
apply(simp add: add_match_def)
apply(rename_tac r rs)
apply(case_tac r)
apply(simp add: add_match_split_fst no_matching_Goto_tail)
apply(rename_tac m' a')
apply(case_tac a')
apply simp_all
done
private lemma matches_add_match_MatchNot_no_matching_Goto_simp: "\<not> matches \<gamma> m p \<Longrightarrow> no_matching_Goto \<gamma> p (add_match m rs)"
apply(induction rs)
apply(simp add: add_match_def)
apply(rename_tac r rs)
apply(case_tac r)
apply(simp add: add_match_split_fst no_matching_Goto_tail)
apply(rename_tac m' a')
apply(case_tac a')
apply simp_all
done
private lemma not_matches_add_match_simp:
assumes "\<not> matches \<gamma> m p"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match m rs, Undecided\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[], Undecided\<rangle> \<Rightarrow> t"
proof(induction rs)
case Nil thus ?case unfolding add_match_def by simp
next
case (Cons r rs)
obtain m' a where r: "r = Rule m' a" by(cases r, simp)
let ?lhs="\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>Rule (MatchAnd m m') a # add_match m rs, Undecided\<rangle> \<Rightarrow> t"
let ?rhs="\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[], Undecided\<rangle> \<Rightarrow> t"
{ assume ?lhs
from \<open>?lhs\<close> Cons have ?rhs
proof(cases \<Gamma> \<gamma> p "Rule (MatchAnd m m') a" "add_match m rs" t rule: seqE_cons_Undecided)
case (no_matching_Goto ti)
hence "ti = Undecided" by (simp add: assms nomatchD)
with no_matching_Goto Cons show ?thesis by simp
next
case (matching_Goto) with Cons assms show ?thesis by force
qed
} note 1=this
{ assume ?rhs
hence "t = Undecided" using skipD by metis
with Cons.IH \<open>?rhs\<close> have ?lhs
by (meson assms matches.simps(1) nomatch not_no_matching_Goto_singleton_cases seq_cons)
} with 1 show ?case by(auto simp add: r add_match_split_fst)
qed
private lemma matches_add_match_MatchNot_simp:
assumes m: "matches \<gamma> m p"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match (MatchNot m) rs, s\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[], s\<rangle> \<Rightarrow> t" (is "?l s \<longleftrightarrow> ?r s")
proof (cases s)
case Undecided
have "?l Undecided \<longleftrightarrow> ?r Undecided"
proof
assume "?l Undecided" with m show "?r Undecided"
proof (induction rs)
case Nil
thus ?case
unfolding add_match_def by simp
next
case (Cons r rs)
thus ?case
by (cases r) (metis matches_MatchNotAnd_simp skipD seqE_cons add_match_split_fst)
qed
next
assume "?r Undecided" with m show "?l Undecided"
proof (induction rs)
case Nil
thus ?case
unfolding add_match_def by simp
next
case (Cons r rs)
hence "t = Undecided" using skipD by metis
with Cons show ?case
apply (cases r)
apply(simp add: add_match_split_fst)
by (metis matches.simps(1) matches.simps(2) matches_MatchNotAnd_simp not_no_matching_Goto_singleton_cases seq_cons)
qed
qed
with Undecided show ?thesis by fast
next
case (Decision d)
thus ?thesis
by(metis decision decisionD)
qed
private lemma just_show_all_bigstep_semantics_equalities_with_start_Undecided:
"\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs1, Undecided\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs2, Undecided\<rangle> \<Rightarrow> t \<Longrightarrow>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs1, s\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs2, s\<rangle> \<Rightarrow> t"
apply(cases s)
apply(simp)
apply(simp)
using decision decisionD by fastforce
private lemma matches_add_match_simp_helper:
assumes m: "matches \<gamma> m p"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match m rs, Undecided\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t" (is "?l \<longleftrightarrow> ?r")
proof
assume ?l with m show ?r
proof (induction rs)
case Nil
thus ?case
unfolding add_match_def by simp
next
case (Cons r rs)
obtain m' a where r: "r = Rule m' a" by(cases r, simp)
from Cons have " \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>Rule (MatchAnd m m') a # add_match m rs, Undecided\<rangle> \<Rightarrow> t"
by(simp add: r add_match_split_fst)
from this Cons have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>Rule m' a # rs, Undecided\<rangle> \<Rightarrow> t"
proof(cases rule: seqE_cons_Undecided)
case (no_matching_Goto ti)
from no_matching_Goto(3) Cons.prems(1) not_no_matching_Goto_singleton_cases
have "no_matching_Goto \<gamma> p [Rule m' a]" by (metis matches.simps(1))
with no_matching_Goto Cons show ?thesis
apply(simp add: matches_rule_and_simp)
apply(cases ti)
apply (simp add: seq'_cons)
by (metis decision decisionD seq'_cons)
next
case (matching_Goto) with Cons show ?thesis
apply(clarify)
apply(simp add: matches_rule_and_simp_help)
by (simp add: seq_cons_Goto_t)
qed
thus ?case by(simp add: r)
qed
next
assume ?r with m show ?l
proof (induction rs)
case Nil
thus ?case
unfolding add_match_def by simp
next
case (Cons r rs)
obtain m' a where r: "r = Rule m' a" by(cases r, simp)
from Cons have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>Rule m' a # rs, Undecided\<rangle> \<Rightarrow> t" by(simp add: r)
from this have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>Rule (MatchAnd m m') a # add_match m rs, Undecided\<rangle> \<Rightarrow> t"
proof(cases \<Gamma> \<gamma> p "Rule m' a" rs t rule: seqE_cons_Undecided)
case (no_matching_Goto ti)
from no_matching_Goto Cons.prems matches_rule_and_simp[symmetric] have
"\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule (MatchAnd m m') a], Undecided\<rangle> \<Rightarrow> ti" by fast
with Cons.prems Cons.IH no_matching_Goto show ?thesis
apply(cases ti)
apply (metis matches.simps(1) not_no_matching_Goto_singleton_cases seq_cons)
apply (metis decision decisionD matches.simps(1) not_no_matching_Goto_singleton_cases seq_cons)
done
next
case (matching_Goto) with Cons show ?thesis
by (simp add: matches_rule_and_simp_help seq_cons_Goto_t)
qed
thus ?case by(simp add: r add_match_split_fst)
qed
qed
private lemma matches_add_match_simp:
"matches \<gamma> m p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match m rs, s\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, s\<rangle> \<Rightarrow> t"
apply(rule just_show_all_bigstep_semantics_equalities_with_start_Undecided)
by(simp add: matches_add_match_simp_helper)
private lemma not_matches_add_matchNot_simp:
"\<not> matches \<gamma> m p \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match (MatchNot m) rs, s\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, s\<rangle> \<Rightarrow> t"
by (simp add: matches_add_match_simp)
subsection\<open>Goto Unfolding\<close>
private lemma unfold_Goto_Undecided:
assumes chain_defined: "\<Gamma> chain = Some rs" and no_matching_Goto_rs: "no_matching_Goto \<gamma> p rs"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>(Rule m (Goto chain))#rest, Undecided\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match m rs @ add_match (MatchNot m) rest, Undecided\<rangle> \<Rightarrow> t"
(is "?l \<longleftrightarrow> ?r")
proof
assume ?l
thus ?r
proof(cases rule: seqE_cons_Undecided)
case (no_matching_Goto ti)
from no_matching_Goto have "\<not> matches \<gamma> m p" by simp
with no_matching_Goto have ti: "ti = Undecided" using nomatchD by metis
from \<open>\<not> matches \<gamma> m p\<close> have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match m rs, Undecided\<rangle> \<Rightarrow> Undecided"
using not_matches_add_match_simp skip by fast
from \<open>\<not> matches \<gamma> m p\<close> matches_add_match_MatchNot_no_matching_Goto_simp have "no_matching_Goto \<gamma> p (add_match m rs)" by force
from no_matching_Goto ti have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rest, Undecided\<rangle> \<Rightarrow> t" by simp
with not_matches_add_matchNot_simp[OF \<open>\<not> matches \<gamma> m p\<close>] have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match (MatchNot m) rest, Undecided\<rangle> \<Rightarrow> t" by simp
show ?thesis
by (meson \<open>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match (MatchNot m) rest, Undecided\<rangle> \<Rightarrow> t\<close> \<open>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match m rs, Undecided\<rangle> \<Rightarrow> Undecided\<close> \<open>no_matching_Goto \<gamma> p (add_match m rs)\<close> seq)
next
case (matching_Goto m chain rs')
from matching_Goto gotoD assms have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t" by fastforce
hence 1: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match m rs, Undecided\<rangle> \<Rightarrow> t" by (simp add: matches_add_match_simp matching_Goto(3))
have 2: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match (MatchNot m) rest, t\<rangle> \<Rightarrow> t" by (simp add: matches_add_match_MatchNot_simp matching_Goto(3) skip)
from no_matching_Goto_rs matches_add_match_no_matching_Goto_simp2 matching_Goto have 3: "no_matching_Goto \<gamma> p (add_match m rs)" by fast
from 1 2 3 show ?thesis using matching_Goto(1) seq by fastforce
qed
next
assume ?r
thus ?l
proof(cases "matches \<gamma> m p")
case True
have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t"
by (metis True \<open>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match m rs @ add_match (MatchNot m) rest, Undecided\<rangle> \<Rightarrow> t\<close>
matches_add_match_MatchNot_simp matches_add_match_simp_helper self_append_conv seq' seqE)
show ?l
apply(cases t)
using goto_no_decision[OF True] chain_defined apply (metis \<open>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t\<close>)
using goto_decision[OF True, of \<Gamma> chain rs _ rest] chain_defined apply (metis \<open>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t\<close>)
done
next
case False
with \<open>?r\<close> have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match (MatchNot m) rest, Undecided\<rangle> \<Rightarrow> t"
by (metis matches_add_match_MatchNot_no_matching_Goto_simp not_matches_add_match_simp seqE skipD)
with False have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rest, Undecided\<rangle> \<Rightarrow> t" by (meson not_matches_add_matchNot_simp)
show ?l by (meson False \<open>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rest, Undecided\<rangle> \<Rightarrow> t\<close> nomatch not_no_matching_Goto_singleton_cases seq_cons)
qed
qed
(*
This theorem allows us to unfold the deepest goto in a ruleset.
This can be iterated to get to the higher-level gotos.
*)
qualified theorem unfold_Goto:
assumes chain_defined: "\<Gamma> chain = Some rs" and no_matching_Goto_rs: "no_matching_Goto \<gamma> p rs"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>(Rule m (Goto chain))#rest, s\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>add_match m rs @ add_match (MatchNot m) rest, s\<rangle> \<Rightarrow> t"
apply(rule just_show_all_bigstep_semantics_equalities_with_start_Undecided)
using assms unfold_Goto_Undecided by fast
text\<open>A chain that will definitely come to a direct decision\<close>
qualified fun terminal_chain :: "'a rule list \<Rightarrow> bool" where
"terminal_chain [] = False" |
"terminal_chain [Rule MatchAny Accept] = True" |
"terminal_chain [Rule MatchAny Drop] = True" |
"terminal_chain [Rule MatchAny Reject] = True" |
"terminal_chain ((Rule _ (Goto _))#rs) = False" |
"terminal_chain ((Rule _ (Call _))#rs) = False" |
"terminal_chain ((Rule _ Return)#rs) = False" |
"terminal_chain ((Rule _ Unknown)#rs) = False" |
"terminal_chain (_#rs) = terminal_chain rs"
private lemma terminal_chain_no_matching_Goto: "terminal_chain rs \<Longrightarrow> no_matching_Goto \<gamma> p rs"
by(induction rs rule: terminal_chain.induct) simp_all
text\<open>A terminal chain means (if the semantics are actually defined) that the chain will
ultimately yield a final filtering decision, for all packets.\<close>
qualified lemma "terminal_chain rs \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t \<Longrightarrow> \<exists>X. t = Decision X"
apply(induction rs)
apply(simp)
apply(rename_tac r rs)
apply(case_tac r)
apply(rename_tac m a)
apply(simp)
apply(frule_tac \<gamma>=\<gamma> and p=p in terminal_chain_no_matching_Goto)
apply(case_tac a)
apply(simp_all)
apply(erule seqE_cons, simp_all,
metis iptables_goto_bigstepD matches.elims terminal_chain.simps terminal_chain.simps terminal_chain.simps)+
done
private lemma replace_Goto_with_Call_in_terminal_chain_Undecided:
assumes chain_defined: "\<Gamma> chain = Some rs" and terminal_chain: "terminal_chain rs"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)], Undecided\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Call chain)], Undecided\<rangle> \<Rightarrow> t"
(is "?l \<longleftrightarrow> ?r")
proof
assume ?l
thus ?r
proof(cases rule: seqE_cons_Undecided)
case (no_matching_Goto ti)
from no_matching_Goto have "\<not> matches \<gamma> m p" by simp
with nomatch have 1: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)], Undecided\<rangle> \<Rightarrow> Undecided" by fast
from \<open>\<not> matches \<gamma> m p\<close> nomatch have 2: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Call chain)], Undecided\<rangle> \<Rightarrow> Undecided" by fast
from 1 2 show ?thesis
using \<open>?l\<close> iptables_goto_bigstep_Undecided_Undecided_deterministic by fastforce
next
case (matching_Goto m chain rs')
from matching_Goto gotoD assms have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t" by fastforce
from call_result[OF \<open>matches \<gamma> m p\<close> chain_defined \<open>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t\<close>] show ?thesis
by (metis matching_Goto(1) rule.sel(1))
qed
next
assume ?r
thus ?l
proof(cases "matches \<gamma> m p")
case True
{fix rs1::"'a rule list" and m' and rs2
have "terminal_chain (rs1 @ Rule m' Return # rs2) \<Longrightarrow> False"
apply(induction rs1)
apply(simp_all)
apply(rename_tac r' rs')
apply(case_tac r')
apply(rename_tac m a)
apply(simp_all)
apply(case_tac a)
apply(simp_all)
apply (metis append_is_Nil_conv hd_Cons_tl terminal_chain.simps)+
done
} note no_return=this
have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t"
apply(rule callD[OF \<open>?r\<close> _ _ True chain_defined])
apply(simp_all)
using no_return terminal_chain by blast
show ?l
apply(cases t)
using goto_no_decision[OF True] chain_defined apply (metis \<open>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t\<close>)
using goto_decision[OF True, of \<Gamma> chain rs _ "[]"] chain_defined apply (metis \<open>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t\<close>)
done
next
case False
show ?l using False \<open>\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Call chain)], Undecided\<rangle> \<Rightarrow> t\<close> nomatch nomatchD by fastforce
qed
qed
qualified theorem replace_Goto_with_Call_in_terminal_chain:
assumes chain_defined: "\<Gamma> chain = Some rs" and terminal_chain: "terminal_chain rs"
shows "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)], s\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Call chain)], s\<rangle> \<Rightarrow> t"
apply(rule just_show_all_bigstep_semantics_equalities_with_start_Undecided)
using assms replace_Goto_with_Call_in_terminal_chain_Undecided by fast
qualified fun rewrite_Goto_chain_safe :: "(string \<rightharpoonup> 'a rule list) \<Rightarrow> 'a rule list \<Rightarrow> ('a rule list) option" where
"rewrite_Goto_chain_safe _ [] = Some []" |
"rewrite_Goto_chain_safe \<Gamma> ((Rule m (Goto chain))#rs) =
(case (\<Gamma> chain) of None \<Rightarrow> None
| Some rs' \<Rightarrow> (if
\<not> terminal_chain rs'
then
None
else
map_option (\<lambda>rs. Rule m (Call chain) # rs) (rewrite_Goto_chain_safe \<Gamma> rs)
)
)" |
"rewrite_Goto_chain_safe \<Gamma> (r#rs) = map_option (\<lambda>rs. r # rs) (rewrite_Goto_chain_safe \<Gamma> rs)"
private fun rewrite_Goto_safe_internal
:: "(string \<times> 'a rule list) list \<Rightarrow> (string \<times> 'a rule list) list \<Rightarrow> (string \<times> 'a rule list) list option" where
"rewrite_Goto_safe_internal _ [] = Some []" |
"rewrite_Goto_safe_internal \<Gamma> ((chain_name, rs)#cs) =
(case rewrite_Goto_chain_safe (map_of \<Gamma>) rs of
None \<Rightarrow> None
| Some rs' \<Rightarrow> map_option (\<lambda>rst. (chain_name, rs')#rst) (rewrite_Goto_safe_internal \<Gamma> cs)
)"
qualified fun rewrite_Goto_safe :: "(string \<times> 'a rule list) list \<Rightarrow> (string \<times> 'a rule list) list option" where
"rewrite_Goto_safe cs = rewrite_Goto_safe_internal cs cs"
(*use rewrite_Goto_chain_safe whenever possible!*)
qualified definition rewrite_Goto :: "(string \<times> 'a rule list) list \<Rightarrow> (string \<times> 'a rule list) list" where
"rewrite_Goto cs = the (rewrite_Goto_safe cs)"
private lemma step_IH_cong: "(\<And>s. \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs1, s\<rangle> \<Rightarrow> t = \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs2, s\<rangle> \<Rightarrow> t) \<Longrightarrow>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r#rs1, s\<rangle> \<Rightarrow> t = \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>r#rs2, s\<rangle> \<Rightarrow> t"
apply(rule iffI)
apply(erule seqE_cons)
apply(rule seq'_cons)
apply simp_all
apply(drule not_no_matching_Goto_cases)
apply(simp; fail)
apply(elim exE conjE, rename_tac rs1a m chain rs2a)
apply(subgoal_tac "r = Rule m (Goto chain)")
prefer 2
subgoal by (simp add: Cons_eq_append_conv)
apply(thin_tac "[r] = _ @ Rule m (Goto chain) # _")
apply simp
apply (metis decision decisionD seq_cons_Goto_t state.exhaust)
apply(erule seqE_cons)
apply(rule seq'_cons)
apply simp_all
apply(drule not_no_matching_Goto_cases)
apply(simp; fail)
apply(elim exE conjE, rename_tac rs1a m chain rs2a)
apply(subgoal_tac "r = Rule m (Goto chain)")
prefer 2
subgoal by (simp add: Cons_eq_append_conv)
apply(thin_tac "[r] = _ @ Rule m (Goto chain) # _")
apply simp
apply (metis decision decisionD seq_cons_Goto_t state.exhaust)
done
private lemma terminal_chain_decision:
"terminal_chain rs \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, Undecided\<rangle> \<Rightarrow> t \<Longrightarrow> \<exists>X. t = Decision X"
apply(induction rs arbitrary: t rule: terminal_chain.induct)
apply simp_all
apply(auto dest: iptables_goto_bigstepD)[3]
apply(erule seqE_cons, simp_all, blast dest: iptables_goto_bigstepD)+ (*6s*)
done
private lemma terminal_chain_Goto_decision: "\<Gamma> chain = Some rs \<Longrightarrow> terminal_chain rs \<Longrightarrow> matches \<gamma> m p \<Longrightarrow>
\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)], s\<rangle> \<Rightarrow> t \<Longrightarrow> \<exists>X. t = Decision X"
apply(cases s)
apply(drule gotoD, simp_all)
apply(elim exE conjE, simp_all)
using terminal_chain_decision apply fast
by (meson decisionD)
qualified theorem rewrite_Goto_chain_safe:
"rewrite_Goto_chain_safe \<Gamma> rs = Some rs' \<Longrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs', s\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, s\<rangle> \<Rightarrow> t"
proof(induction \<Gamma> rs arbitrary: rs' s rule: rewrite_Goto_chain_safe.induct)
case 1 thus ?case by (simp split: option.split_asm if_split_asm)
next
case (2 \<Gamma> m chain rs)
from 2(2) obtain z x2 where "\<Gamma> chain = Some x2" and "terminal_chain x2"
and "rs' = Rule m (Call chain) # z"
and "Some z = rewrite_Goto_chain_safe \<Gamma> rs"
by(auto split: option.split_asm if_split_asm)
from 2(1) \<open>\<Gamma> chain = Some x2\<close> \<open>terminal_chain x2\<close> \<open>Some z = rewrite_Goto_chain_safe \<Gamma> rs\<close>
have IH: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>z, s\<rangle> \<Rightarrow> t = \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, s\<rangle> \<Rightarrow> t" for s by simp
have "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>Rule m (Call chain) # z, Undecided\<rangle> \<Rightarrow> t \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>Rule m (Goto chain) # rs, Undecided\<rangle> \<Rightarrow> t"
(is "?lhs \<longleftrightarrow> ?rhs")
proof(intro iffI)
assume ?lhs
with IH obtain ti where ti1: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Call chain)], Undecided\<rangle> \<Rightarrow> ti" and ti2: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>rs, ti\<rangle> \<Rightarrow> t"
by(auto elim: seqE_cons)
show ?rhs
proof(cases "matches \<gamma> m p")
case False
from replace_Goto_with_Call_in_terminal_chain \<open>\<Gamma> chain = Some x2\<close> \<open>terminal_chain x2\<close>
have " \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Call chain)], Undecided\<rangle> \<Rightarrow> ti \<longleftrightarrow> \<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)], Undecided\<rangle> \<Rightarrow> ti"
by fast
with False ti1 ti2 show ?thesis by(rule_tac t=ti in seq'_cons) simp+
next
case True
from ti1 \<open>\<Gamma> chain = Some x2\<close> \<open>terminal_chain x2\<close>
have g: "\<Gamma>,\<gamma>,p\<turnstile>\<^sub>g \<langle>[Rule m (Goto chain)], Undecided\<rangle> \<Rightarrow> ti"
by(subst(asm) replace_Goto_with_Call_in_terminal_chain[symmetric]) simp+
with True \<open>\<Gamma> chain = Some x2\<close> \<open>terminal_chain x2\<close> obtain X where X: "ti = Decision X"
by(blast dest: terminal_chain_Goto_decision)
with this ti2 have "t = Decision X"
by(simp add: decisionD)
with g X True ti2 \<open>\<Gamma> chain = Some x2\<close> \<open>terminal_chain x2\<close> show ?thesis
apply(simp)
apply(rule seq_cons_Goto_t, simp_all)
done
qed
next
assume ?rhs
with IH \<open>\<Gamma> chain = Some x2\<close> \<open>terminal_chain x2\<close> \<open>Some z = rewrite_Goto_chain_safe \<Gamma> rs\<close> show ?lhs
apply -
apply(erule seqE_cons)
subgoal for ti
apply simp_all
apply(rule_tac t=ti in seq'_cons)
apply simp_all
using replace_Goto_with_Call_in_terminal_chain by fast
apply simp
apply(frule(3) terminal_chain_Goto_decision)
apply(subst(asm) replace_Goto_with_Call_in_terminal_chain, simp_all)
apply(rule seq'_cons, simp_all)
apply(elim exE)
by (simp add: decision)
qed
with \<open>rs' = Rule m (Call chain) # z\<close> show ?case
apply -
apply(rule just_show_all_bigstep_semantics_equalities_with_start_Undecided)
by simp
qed(auto cong: step_IH_cong)
text\<open>Example: The semantics are actually defined (for this example).\<close>
lemma defines "\<gamma> \<equiv> (\<lambda>_ _. True)" and "m \<equiv> MatchAny"
shows "[''FORWARD'' \<mapsto> [Rule m Log, Rule m (Call ''foo''), Rule m Drop],
''foo'' \<mapsto> [Rule m Log, Rule m (Goto ''bar''), Rule m Reject],
''bar'' \<mapsto> [Rule m (Goto ''baz''), Rule m Reject],
''baz'' \<mapsto> [(Rule m Accept)]],
\<gamma>,p\<turnstile>\<^sub>g\<langle>[Rule MatchAny (Call ''FORWARD'')], Undecided\<rangle> \<Rightarrow> (Decision FinalAllow)"
apply(subgoal_tac "matches \<gamma> m p")
prefer 2
apply(simp add: \<gamma>_def m_def; fail)
apply(rule call_result)
apply(auto)
apply(rule_tac t=Undecided in seq_cons)
apply(auto intro: log)
apply(rule_tac t="Decision FinalAllow" in seq_cons)
apply(auto intro: decision)
apply(rule call_result)
apply(simp)+
apply(rule_tac t=Undecided in seq_cons)
apply(auto intro: log)
apply(rule goto_decision)
apply(simp)+
apply(rule goto_decision)
apply(simp)+
apply(auto intro: accept)
done
end
end
diff --git a/thys/Jinja/J/execute_Bigstep.thy b/thys/Jinja/J/execute_Bigstep.thy
--- a/thys/Jinja/J/execute_Bigstep.thy
+++ b/thys/Jinja/J/execute_Bigstep.thy
@@ -1,239 +1,238 @@
(* Title: Jinja/J/execute_Bigstep.thy
Author: Tobias Nipkow
Copyright 2004 Technische Universitaet Muenchen
*)
section \<open>Code Generation For BigStep\<close>
theory execute_Bigstep
imports
BigStep Examples
"HOL-Library.Code_Target_Numeral"
begin
inductive map_val :: "expr list \<Rightarrow> val list \<Rightarrow> bool"
where
Nil: "map_val [] []"
| Cons: "map_val xs ys \<Longrightarrow> map_val (Val y # xs) (y # ys)"
inductive map_val2 :: "expr list \<Rightarrow> val list \<Rightarrow> expr list \<Rightarrow> bool"
where
Nil: "map_val2 [] [] []"
| Cons: "map_val2 xs ys zs \<Longrightarrow> map_val2 (Val y # xs) (y # ys) zs"
| Throw: "map_val2 (throw e # xs) [] (throw e # xs)"
theorem map_val_conv: "(xs = map Val ys) = map_val xs ys"
(*<*)
proof -
have "\<And>ys. xs = map Val ys \<Longrightarrow> map_val xs ys"
apply (induct xs type:list)
apply (case_tac ys)
apply simp
apply (rule map_val.Nil)
apply simp
apply (case_tac ys)
apply simp
apply simp
- apply (erule conjE)
apply (rule map_val.Cons)
apply simp
done
moreover have "map_val xs ys \<Longrightarrow> xs = map Val ys"
by (erule map_val.induct) simp+
ultimately show ?thesis ..
qed
(*>*)
theorem map_val2_conv:
"(xs = map Val ys @ throw e # zs) = map_val2 xs ys (throw e # zs)"
(*<*)
proof -
have "\<And>ys. xs = map Val ys @ throw e # zs \<Longrightarrow> map_val2 xs ys (throw e # zs)"
apply (induct xs type:list)
apply (case_tac ys)
apply simp
apply simp
apply simp
apply (case_tac ys)
apply simp
apply (rule map_val2.Throw)
apply simp
apply (rule map_val2.Cons)
apply simp
done
moreover have "map_val2 xs ys (throw e # zs) \<Longrightarrow> xs = map Val ys @ throw e # zs"
by (erule map_val2.induct) simp+
ultimately show ?thesis ..
qed
(*>*)
lemma CallNull2:
"\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>null,s\<^sub>1\<rangle>; P \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>evs,s\<^sub>2\<rangle>; map_val evs vs \<rbrakk>
\<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>THROW NullPointer,s\<^sub>2\<rangle>"
apply(rule CallNull, assumption+)
apply(simp add: map_val_conv[symmetric])
done
lemma CallParamsThrow2:
"\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>Val v,s\<^sub>1\<rangle>; P \<turnstile> \<langle>es,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>evs,s\<^sub>2\<rangle>;
map_val2 evs vs (throw ex # es'') \<rbrakk>
\<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(es),s\<^sub>0\<rangle> \<Rightarrow> \<langle>throw ex,s\<^sub>2\<rangle>"
apply(rule eval_evals.CallParamsThrow, assumption+)
apply(simp add: map_val2_conv[symmetric])
done
lemma Call2:
"\<lbrakk> P \<turnstile> \<langle>e,s\<^sub>0\<rangle> \<Rightarrow> \<langle>addr a,s\<^sub>1\<rangle>; P \<turnstile> \<langle>ps,s\<^sub>1\<rangle> [\<Rightarrow>] \<langle>evs,(h\<^sub>2,l\<^sub>2)\<rangle>;
map_val evs vs;
h\<^sub>2 a = Some(C,fs); P \<turnstile> C sees M:Ts\<rightarrow>T = (pns,body) in D;
length vs = length pns; l\<^sub>2' = [this\<mapsto>Addr a, pns[\<mapsto>]vs];
P \<turnstile> \<langle>body,(h\<^sub>2,l\<^sub>2')\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>3)\<rangle> \<rbrakk>
\<Longrightarrow> P \<turnstile> \<langle>e\<bullet>M(ps),s\<^sub>0\<rangle> \<Rightarrow> \<langle>e',(h\<^sub>3,l\<^sub>2)\<rangle>"
apply(rule Call, assumption+)
apply(simp add: map_val_conv[symmetric])
apply assumption+
done
code_pred
(modes: i \<Rightarrow> o \<Rightarrow> bool)
map_val
.
code_pred
(modes: i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool)
map_val2
.
lemmas [code_pred_intro] =
eval_evals.New eval_evals.NewFail
eval_evals.Cast eval_evals.CastNull eval_evals.CastFail eval_evals.CastThrow
eval_evals.Val eval_evals.Var
eval_evals.BinOp eval_evals.BinOpThrow1 eval_evals.BinOpThrow2
eval_evals.LAss eval_evals.LAssThrow
eval_evals.FAcc eval_evals.FAccNull eval_evals.FAccThrow
eval_evals.FAss eval_evals.FAssNull
eval_evals.FAssThrow1 eval_evals.FAssThrow2
eval_evals.CallObjThrow
declare CallNull2 [code_pred_intro CallNull2]
declare CallParamsThrow2 [code_pred_intro CallParamsThrow2]
declare Call2 [code_pred_intro Call2]
lemmas [code_pred_intro] =
eval_evals.Block
eval_evals.Seq eval_evals.SeqThrow
eval_evals.CondT eval_evals.CondF eval_evals.CondThrow
eval_evals.WhileF eval_evals.WhileT
eval_evals.WhileCondThrow
declare eval_evals.WhileBodyThrow [code_pred_intro WhileBodyThrow2]
lemmas [code_pred_intro] =
eval_evals.Throw eval_evals.ThrowNull
eval_evals.ThrowThrow
eval_evals.Try eval_evals.TryCatch eval_evals.TryThrow
eval_evals.Nil eval_evals.Cons eval_evals.ConsThrow
code_pred
(modes: i \<Rightarrow> i \<Rightarrow> i \<Rightarrow> o \<Rightarrow> o \<Rightarrow> bool as execute)
eval
proof -
case eval
from eval.prems show thesis
proof(cases (no_simp))
case CallNull thus ?thesis
by(rule eval.CallNull2[OF refl])(simp add: map_val_conv[symmetric])
next
case CallParamsThrow thus ?thesis
by(rule eval.CallParamsThrow2[OF refl])(simp add: map_val2_conv[symmetric])
next
case Call thus ?thesis
by -(rule eval.Call2[OF refl], simp_all add: map_val_conv[symmetric])
next
case WhileBodyThrow thus ?thesis by(rule eval.WhileBodyThrow2[OF refl])
qed(assumption|erule (4) eval.that[OF refl]|erule (3) eval.that[OF refl])+
next
case evals
from evals.prems show thesis
by(cases (no_simp))(assumption|erule (3) evals.that[OF refl])+
qed
notation execute ("_ \<turnstile> ((1\<langle>_,/_\<rangle>) \<Rightarrow>/ \<langle>'_, '_\<rangle>)" [51,0,0] 81)
definition "test1 = [] \<turnstile> \<langle>testExpr1,(Map.empty,Map.empty)\<rangle> \<Rightarrow> \<langle>_,_\<rangle>"
definition "test2 = [] \<turnstile> \<langle>testExpr2,(Map.empty,Map.empty)\<rangle> \<Rightarrow> \<langle>_,_\<rangle>"
definition "test3 = [] \<turnstile> \<langle>testExpr3,(Map.empty,Map.empty(''V''\<mapsto>Intg 77))\<rangle> \<Rightarrow> \<langle>_,_\<rangle>"
definition "test4 = [] \<turnstile> \<langle>testExpr4,(Map.empty,Map.empty)\<rangle> \<Rightarrow> \<langle>_,_\<rangle>"
definition "test5 = [(''Object'',('''',[],[])),(''C'',(''Object'',[(''F'',Integer)],[]))] \<turnstile> \<langle>testExpr5,(Map.empty,Map.empty)\<rangle> \<Rightarrow> \<langle>_,_\<rangle>"
definition "test6 = [(''Object'',('''',[],[])), classI] \<turnstile> \<langle>testExpr6,(Map.empty,Map.empty)\<rangle> \<Rightarrow> \<langle>_,_\<rangle>"
definition "V = ''V''"
definition "C = ''C''"
definition "F = ''F''"
ML_val \<open>
val SOME ((@{code Val} (@{code Intg} (@{code int_of_integer} 5)), _), _) = Predicate.yield @{code test1};
val SOME ((@{code Val} (@{code Intg} (@{code int_of_integer} 11)), _), _) = Predicate.yield @{code test2};
val SOME ((@{code Val} (@{code Intg} (@{code int_of_integer} 83)), _), _) = Predicate.yield @{code test3};
val SOME ((_, (_, l)), _) = Predicate.yield @{code test4};
val SOME (@{code Intg} (@{code int_of_integer} 6)) = l @{code V};
val SOME ((_, (h, _)), _) = Predicate.yield @{code test5};
val SOME (c, fs) = h (@{code nat_of_integer} 1);
val SOME (obj, _) = h (@{code nat_of_integer} 0);
val SOME (@{code Intg} i) = fs (@{code F}, @{code C});
@{assert} (c = @{code C} andalso obj = @{code Object} andalso i = @{code int_of_integer} 42);
val SOME ((@{code Val} (@{code Intg} (@{code int_of_integer} 160)), _), _) = Predicate.yield @{code test6};
\<close>
definition "test7 = [classObject, classL] \<turnstile> \<langle>testExpr_BuildList, (Map.empty,Map.empty)\<rangle> \<Rightarrow> \<langle>_,_\<rangle>"
definition "L = ''L''"
definition "N = ''N''"
ML_val \<open>
val SOME ((_, (h, _)), _) = Predicate.yield @{code test7};
val SOME (_, fs1) = h (@{code nat_of_integer} 0);
val SOME (_, fs2) = h (@{code nat_of_integer} 1);
val SOME (_, fs3) = h (@{code nat_of_integer} 2);
val SOME (_, fs4) = h (@{code nat_of_integer} 3);
val F = @{code "F"};
val L = @{code "L"};
val N = @{code "N"};
@{assert} (fs1 (F, L) = SOME (@{code Intg} (@{code int_of_integer} 1)) andalso
fs1 (N, L) = SOME (@{code Addr} (@{code nat_of_integer} 1)) andalso
fs2 (F, L) = SOME (@{code Intg} (@{code int_of_integer} 2)) andalso
fs2 (N, L) = SOME (@{code Addr} (@{code nat_of_integer} 2)) andalso
fs3 (F, L) = SOME (@{code Intg} (@{code int_of_integer} 3)) andalso
fs3 (N, L) = SOME (@{code Addr} (@{code nat_of_integer} 3)) andalso
fs4 (F, L) = SOME (@{code Intg} (@{code int_of_integer} 4)) andalso
fs4 (N, L) = SOME @{code Null});
\<close>
definition "test8 = [classObject, classA] \<turnstile> \<langle>testExpr_ClassA, (Map.empty,Map.empty)\<rangle> \<Rightarrow> \<langle>_,_\<rangle>"
definition "i = ''int''"
definition "t = ''test''"
definition "A = ''A''"
ML_val \<open>
val SOME ((_, (h, l)), _) = Predicate.yield @{code test8};
val SOME (_, fs1) = h (@{code nat_of_integer} 0);
val SOME (_, fs2) = h (@{code nat_of_integer} 1);
val i = @{code "i"};
val t = @{code "t"};
val A = @{code "A"};
@{assert} (fs1 (i, A) = SOME (@{code Intg} (@{code int_of_integer} 10)) andalso
fs1 (t, A) = SOME @{code Null} andalso
fs2 (i, A) = SOME (@{code Intg} (@{code int_of_integer} 50)) andalso
fs2 (t, A) = SOME @{code Null});
\<close>
end
diff --git a/thys/JinjaThreads/J/Deadlocked.thy b/thys/JinjaThreads/J/Deadlocked.thy
--- a/thys/JinjaThreads/J/Deadlocked.thy
+++ b/thys/JinjaThreads/J/Deadlocked.thy
@@ -1,462 +1,462 @@
(* Title: JinjaThreads/J/Deadlocked.thy
Author: Andreas Lochbihler
*)
section \<open>Preservation of Deadlock\<close>
theory Deadlocked
imports
ProgressThreaded
begin
context J_progress begin
lemma red_wt_hconf_hext:
assumes wf: "wf_J_prog P"
and hconf: "hconf H"
and tconf: "P,H \<turnstile> t \<surd>t"
shows "\<lbrakk> convert_extTA extNTA,P,t \<turnstile> \<langle>e, s\<rangle> -ta\<rightarrow> \<langle>e', s'\<rangle>; P,E,H \<turnstile> e : T; hext H (hp s) \<rbrakk>
\<Longrightarrow> \<exists>ta' e' s'. convert_extTA extNTA,P,t \<turnstile> \<langle>e, (H, lcl s)\<rangle> -ta'\<rightarrow> \<langle>e', s'\<rangle> \<and>
collect_locks \<lbrace>ta\<rbrace>\<^bsub>l\<^esub> = collect_locks \<lbrace>ta'\<rbrace>\<^bsub>l\<^esub> \<and> collect_cond_actions \<lbrace>ta\<rbrace>\<^bsub>c\<^esub> = collect_cond_actions \<lbrace>ta'\<rbrace>\<^bsub>c\<^esub> \<and>
collect_interrupts \<lbrace>ta\<rbrace>\<^bsub>i\<^esub> = collect_interrupts \<lbrace>ta'\<rbrace>\<^bsub>i\<^esub>"
and "\<lbrakk> convert_extTA extNTA,P,t \<turnstile> \<langle>es, s\<rangle> [-ta\<rightarrow>] \<langle>es', s'\<rangle>; P,E,H \<turnstile> es [:] Ts; hext H (hp s) \<rbrakk>
\<Longrightarrow> \<exists>ta' es' s'. convert_extTA extNTA,P,t \<turnstile> \<langle>es, (H, lcl s)\<rangle> [-ta'\<rightarrow>] \<langle>es', s'\<rangle> \<and>
collect_locks \<lbrace>ta\<rbrace>\<^bsub>l\<^esub> = collect_locks \<lbrace>ta'\<rbrace>\<^bsub>l\<^esub> \<and> collect_cond_actions \<lbrace>ta\<rbrace>\<^bsub>c\<^esub> = collect_cond_actions \<lbrace>ta'\<rbrace>\<^bsub>c\<^esub> \<and>
collect_interrupts \<lbrace>ta\<rbrace>\<^bsub>i\<^esub> = collect_interrupts \<lbrace>ta'\<rbrace>\<^bsub>i\<^esub>"
proof(induct arbitrary: E T and E Ts rule: red_reds.inducts)
case (RedNew h' a h C l)
thus ?case
by(cases "allocate H (Class_type C) = {}")(fastforce simp add: ta_upd_simps intro: RedNewFail red_reds.RedNew)+
next
case (RedNewFail h C l)
thus ?case
by(cases "allocate H (Class_type C) = {}")(fastforce simp add: ta_upd_simps intro: red_reds.RedNewFail RedNew)+
next
case NewArrayRed thus ?case by(fastforce intro: red_reds.intros)
next
case (RedNewArray i h' a h T l E T')
thus ?case
by(cases "allocate H (Array_type T (nat (sint i))) = {}")(fastforce simp add: ta_upd_simps intro: red_reds.RedNewArray RedNewArrayFail)+
next
case RedNewArrayNegative thus ?case by(fastforce intro: red_reds.intros)
next
case (RedNewArrayFail i h T l E T')
thus ?case
by(cases "allocate H (Array_type T (nat (sint i))) = {}")(fastforce simp add: ta_upd_simps intro: RedNewArray red_reds.RedNewArrayFail)+
next
case CastRed thus ?case by(fastforce intro: red_reds.intros)
next
case (RedCast s v U T E T')
from \<open>P,E,H \<turnstile> Cast T (Val v) : T'\<close> show ?case
proof(rule WTrt_elim_cases)
fix T''
- assume wt: "P,E,H \<turnstile> Val v : T''" "T = T'"
+ assume wt: "P,E,H \<turnstile> Val v : T''" "T' = T"
thus ?thesis
by(cases "P \<turnstile> T'' \<le> T")(fastforce intro: red_reds.RedCast red_reds.RedCastFail)+
qed
next
case (RedCastFail s v U T E T')
from \<open>P,E,H \<turnstile> Cast T (Val v) : T'\<close>
obtain T'' where "P,E,H \<turnstile> Val v : T''" "T = T'" by auto
thus ?case
by(cases "P \<turnstile> T'' \<le> T")(fastforce intro: red_reds.RedCast red_reds.RedCastFail)+
next
case InstanceOfRed thus ?case by(fastforce intro: red_reds.intros)
next
case RedInstanceOf thus ?case
using [[hypsubst_thin = true]]
by auto((rule exI conjI red_reds.RedInstanceOf)+, auto)
next
case BinOpRed1 thus ?case by(fastforce intro: red_reds.intros)
next
case BinOpRed2 thus ?case by(fastforce intro: red_reds.intros)
next
case RedBinOp thus ?case by(fastforce intro: red_reds.intros)
next
case RedBinOpFail thus ?case by(fastforce intro: red_reds.intros)
next
case RedVar thus ?case by(fastforce intro: red_reds.intros)
next
case LAssRed thus ?case by(fastforce intro: red_reds.intros)
next
case RedLAss thus ?case by(fastforce intro: red_reds.intros)
next
case AAccRed1 thus ?case by(fastforce intro: red_reds.intros)
next
case AAccRed2 thus ?case by(fastforce intro: red_reds.intros)
next
case RedAAccNull thus ?case by(fastforce intro: red_reds.intros)
next
case RedAAccBounds thus ?case
by(fastforce intro: red_reds.RedAAccBounds dest: hext_arrD)
next
case (RedAAcc h a T n i v l E T')
from \<open>P,E,H \<turnstile> addr a\<lfloor>Val (Intg i)\<rceil> : T'\<close>
have wt: "P,E,H \<turnstile> addr a : T'\<lfloor>\<rceil>" by(auto)
with \<open>H \<unlhd> hp (h, l)\<close> \<open>typeof_addr h a = \<lfloor>Array_type T n\<rfloor>\<close>
have Ha: "typeof_addr H a = \<lfloor>Array_type T n\<rfloor>" by(auto dest: hext_arrD)
with \<open>0 <=s i\<close> \<open>sint i < int n\<close>
have "nat (sint i) < n" by(metis nat_less_iff sint_0 word_sle_def)
with Ha have "P,H \<turnstile> a@ACell (nat (sint i)) : T"
by(auto intro: addr_loc_type.intros)
from heap_read_total[OF hconf this]
obtain v where "heap_read H a (ACell (nat (sint i))) v" by blast
with Ha \<open>0 <=s i\<close> \<open>sint i < int n\<close> show ?case
by(fastforce intro: red_reds.RedAAcc simp add: ta_upd_simps)
next
case AAssRed1 thus ?case by(fastforce intro: red_reds.intros)
next
case AAssRed2 thus ?case by(fastforce intro: red_reds.intros)
next
case AAssRed3 thus ?case by(fastforce intro: red_reds.intros)
next
case RedAAssNull thus ?case by(fastforce intro: red_reds.intros)
next
case RedAAssBounds thus ?case by(fastforce intro: red_reds.RedAAssBounds dest: hext_arrD)
next
case (RedAAssStore s a T n i w U E T')
from \<open>P,E,H \<turnstile> addr a\<lfloor>Val (Intg i)\<rceil> := Val w : T'\<close>
obtain T'' T''' where wt: "P,E,H \<turnstile> addr a : T''\<lfloor>\<rceil>"
and wtw: "P,E,H \<turnstile> Val w : T'''" by auto
with \<open>H \<unlhd> hp s\<close> \<open>typeof_addr (hp s) a = \<lfloor>Array_type T n\<rfloor>\<close>
have Ha: "typeof_addr H a = \<lfloor>Array_type T n\<rfloor>" by(auto dest: hext_arrD)
from \<open>typeof\<^bsub>hp s\<^esub> w = \<lfloor>U\<rfloor>\<close> wtw \<open>H \<unlhd> hp s\<close> have "typeof\<^bsub>H\<^esub> w = \<lfloor>U\<rfloor>"
by(auto dest: type_of_hext_type_of)
with Ha \<open>0 <=s i\<close> \<open>sint i < int n\<close> \<open>\<not> P \<turnstile> U \<le> T\<close> show ?case
by(fastforce intro: red_reds.RedAAssStore)
next
case (RedAAss h a T n i w U h' l E T')
from \<open>P,E,H \<turnstile> addr a\<lfloor>Val (Intg i)\<rceil> := Val w : T'\<close>
obtain T'' T''' where wt: "P,E,H \<turnstile> addr a : T''\<lfloor>\<rceil>"
and wtw: "P,E,H \<turnstile> Val w : T'''" by auto
with \<open>H \<unlhd> hp (h, l)\<close> \<open>typeof_addr h a = \<lfloor>Array_type T n\<rfloor>\<close>
have Ha: "typeof_addr H a = \<lfloor>Array_type T n\<rfloor>" by(auto dest: hext_arrD)
from \<open>typeof\<^bsub>h\<^esub> w = \<lfloor>U\<rfloor>\<close> wtw \<open>H \<unlhd> hp (h, l)\<close> have "typeof\<^bsub>H\<^esub> w = \<lfloor>U\<rfloor>"
by(auto dest: type_of_hext_type_of)
moreover
with \<open>P \<turnstile> U \<le> T\<close> have conf: "P,H \<turnstile> w :\<le> T"
by(auto simp add: conf_def)
from \<open>0 <=s i\<close> \<open>sint i < int n\<close>
have "nat (sint i) < n"
by (metis nat_less_iff sint_0 word_sle_def)
with Ha have "P,H \<turnstile> a@ACell (nat (sint i)) : T"
by(auto intro: addr_loc_type.intros)
from heap_write_total[OF hconf this conf]
obtain H' where "heap_write H a (ACell (nat (sint i))) w H'" ..
ultimately show ?case using \<open>0 <=s i\<close> \<open>sint i < int n\<close> Ha \<open>P \<turnstile> U \<le> T\<close>
by(fastforce simp del: split_paired_Ex intro: red_reds.RedAAss)
next
case ALengthRed thus ?case by(fastforce intro: red_reds.intros)
next
case (RedALength h a T n l E T')
from \<open>P,E,H \<turnstile> addr a\<bullet>length : T'\<close>
obtain T'' where [simp]: "T' = Integer"
and wta: "P,E,H \<turnstile> addr a : T''\<lfloor>\<rceil>" by(auto)
then obtain n'' where "typeof_addr H a = \<lfloor>Array_type T'' n''\<rfloor>" by(auto)
thus ?case by(fastforce intro: red_reds.RedALength)
next
case RedALengthNull show ?case by(fastforce intro: red_reds.RedALengthNull)
next
case FAccRed thus ?case by(fastforce intro: red_reds.intros)
next
case (RedFAcc h a D F v l E T)
from \<open>P,E,H \<turnstile> addr a\<bullet>F{D} : T\<close> obtain U C' fm
where wt: "P,E,H \<turnstile> addr a : U"
and icto: "class_type_of' U = \<lfloor>C'\<rfloor>"
and has: "P \<turnstile> C' has F:T (fm) in D"
by(auto)
then obtain hU where Ha: "typeof_addr H a = \<lfloor>hU\<rfloor>" "U = ty_of_htype hU" by(auto)
with icto \<open>P \<turnstile> C' has F:T (fm) in D\<close> have "P,H \<turnstile> a@CField D F : T"
by(auto intro: addr_loc_type.intros)
from heap_read_total[OF hconf this]
obtain v where "heap_read H a (CField D F) v" by blast
thus ?case by(fastforce intro: red_reds.RedFAcc simp add: ta_upd_simps)
next
case RedFAccNull thus ?case by(fastforce intro: red_reds.intros)
next
case FAssRed1 thus ?case by(fastforce intro: red_reds.intros)
next
case FAssRed2 thus ?case by(fastforce intro: red_reds.intros)
next
case RedFAssNull thus ?case by(fastforce intro: red_reds.intros)
next
case (RedFAss h a D F v h' l E T)
from \<open>P,E,H \<turnstile> addr a\<bullet>F{D} := Val v : T\<close> obtain U C' T' T2 fm
where wt: "P,E,H \<turnstile> addr a : U"
and icto: "class_type_of' U = \<lfloor>C'\<rfloor>"
and has: "P \<turnstile> C' has F:T' (fm) in D"
and wtv: "P,E,H \<turnstile> Val v : T2"
and T2T: "P \<turnstile> T2 \<le> T'" by(auto)
moreover from wt obtain hU where Ha: "typeof_addr H a = \<lfloor>hU\<rfloor>" "U = ty_of_htype hU" by(auto)
with icto has have adal: "P,H \<turnstile> a@CField D F : T'" by(auto intro: addr_loc_type.intros)
from wtv T2T have "P,H \<turnstile> v :\<le> T'" by(auto simp add: conf_def)
from heap_write_total[OF hconf adal this]
obtain h' where "heap_write H a (CField D F) v h'" ..
thus ?case by(fastforce intro: red_reds.RedFAss)
next
case CASRed1 thus ?case by(fastforce intro: red_reds.intros)
next
case CASRed2 thus ?case by(fastforce intro: red_reds.intros)
next
case CASRed3 thus ?case by(fastforce intro: red_reds.intros)
next
case CASNull thus ?case by(fastforce intro: red_reds.intros)
next
case (RedCASSucceed h a D F v v' h' l)
note split_paired_Ex[simp del]
from RedCASSucceed.prems(1) obtain T' fm T2 T3 U C where *:
"T = Boolean" "class_type_of' U = \<lfloor>C\<rfloor>" "P \<turnstile> C has F:T' (fm) in D"
"volatile fm" "P \<turnstile> T2 \<le> T'" "P \<turnstile> T3 \<le> T'"
"P,E,H \<turnstile> Val v : T2" "P,E,H \<turnstile> Val v' : T3" "P,E,H \<turnstile> addr a : U" by auto
then have adal: "P,H \<turnstile> a@CField D F : T'" by(auto intro: addr_loc_type.intros)
from heap_read_total[OF hconf this] obtain v'' where v': "heap_read H a (CField D F) v''" by blast
show ?case
proof(cases "v'' = v")
case True
from * have "P,H \<turnstile> v' :\<le> T'" by(auto simp add: conf_def)
from heap_write_total[OF hconf adal this] True * v'
show ?thesis by(fastforce intro: red_reds.RedCASSucceed)
next
case False
then show ?thesis using * v' by(fastforce intro: RedCASFail)
qed
next
case (RedCASFail h a D F v'' v v' l)
note split_paired_Ex[simp del]
from RedCASFail.prems(1) obtain T' fm T2 T3 U C where *:
"T = Boolean" "class_type_of' U = \<lfloor>C\<rfloor>" "P \<turnstile> C has F:T' (fm) in D"
"volatile fm" "P \<turnstile> T2 \<le> T'" "P \<turnstile> T3 \<le> T'"
"P,E,H \<turnstile> Val v : T2" "P,E,H \<turnstile> Val v' : T3" "P,E,H \<turnstile> addr a : U" by auto
then have adal: "P,H \<turnstile> a@CField D F : T'" by(auto intro: addr_loc_type.intros)
from heap_read_total[OF hconf this] obtain v''' where v'': "heap_read H a (CField D F) v'''" by blast
show ?case
proof(cases "v''' = v")
case True
from * have "P,H \<turnstile> v' :\<le> T'" by(auto simp add: conf_def)
from heap_write_total[OF hconf adal this] True * v''
show ?thesis by(fastforce intro: red_reds.RedCASSucceed)
next
case False
then show ?thesis using * v'' by(fastforce intro: red_reds.RedCASFail)
qed
next
case CallObj thus ?case by(fastforce intro: red_reds.intros)
next
case CallParams thus ?case by(fastforce intro: red_reds.intros)
next
case (RedCall s a U M Ts T pns body D vs E T')
from \<open>P,E,H \<turnstile> addr a\<bullet>M(map Val vs) : T'\<close>
obtain U' C' Ts' meth D' Ts''
where wta: "P,E,H \<turnstile> addr a : U'"
and icto: "class_type_of' U' = \<lfloor>C'\<rfloor>"
and sees: "P \<turnstile> C' sees M: Ts'\<rightarrow>T' = meth in D'"
and wtes: "P,E,H \<turnstile> map Val vs [:] Ts''"
and widens: "P \<turnstile> Ts'' [\<le>] Ts'" by auto
from wta obtain hU' where Ha: "typeof_addr H a = \<lfloor>hU'\<rfloor>" "U' = ty_of_htype hU'" by(auto)
moreover from \<open>typeof_addr (hp s) a = \<lfloor>U\<rfloor>\<close> \<open>H \<unlhd> hp s\<close> Ha
have [simp]: "U = hU'" by(auto dest: typeof_addr_hext_mono)
from wtes have "length vs = length Ts''"
by(auto intro: map_eq_imp_length_eq)
moreover from widens have "length Ts'' = length Ts'"
by(auto dest: widens_lengthD)
moreover from sees icto sees \<open>P \<turnstile> class_type_of U sees M: Ts\<rightarrow>T = \<lfloor>(pns, body)\<rfloor> in D\<close> Ha
have [simp]: "meth = \<lfloor>(pns, body)\<rfloor>" by(auto dest: sees_method_fun)
with sees wf have "wf_mdecl wf_J_mdecl P D' (M, Ts', T', \<lfloor>(pns, body)\<rfloor>)"
by(auto intro: sees_wf_mdecl)
hence "length pns = length Ts'" by(simp add: wf_mdecl_def)
ultimately show ?case using sees icto
by(fastforce intro: red_reds.RedCall)
next
case (RedCallExternal s a U M Ts T' D vs ta va h' ta' e' s')
from \<open>P,E,H \<turnstile> addr a\<bullet>M(map Val vs) : T\<close>
obtain U' C' Ts' meth D' Ts''
where wta: "P,E,H \<turnstile> addr a : U'" and icto: "class_type_of' U' = \<lfloor>C'\<rfloor>"
and sees: "P \<turnstile> C' sees M: Ts'\<rightarrow>T = meth in D'"
and wtvs: "P,E,H \<turnstile> map Val vs [:] Ts''"
and sub: "P \<turnstile> Ts'' [\<le>] Ts'" by auto
from wta \<open>typeof_addr (hp s) a = \<lfloor>U\<rfloor>\<close> \<open>hext H (hp s)\<close> have [simp]: "U' = ty_of_htype U"
by(auto dest: typeof_addr_hext_mono)
with icto have [simp]: "C' = class_type_of U" by(auto)
from sees \<open>P \<turnstile> class_type_of U sees M: Ts\<rightarrow>T' = Native in D\<close>
have [simp]: "meth = Native" by(auto dest: sees_method_fun)
with wta sees icto wtvs sub have "P,H \<turnstile> a\<bullet>M(vs) : T"
by(cases U)(auto 4 4 simp add: external_WT'_iff)
from red_external_wt_hconf_hext[OF wf \<open>P,t \<turnstile> \<langle>a\<bullet>M(vs),hp s\<rangle> -ta\<rightarrow>ext \<langle>va,h'\<rangle>\<close> \<open>H \<unlhd> hp s\<close> this tconf hconf]
wta icto sees \<open>ta' = convert_extTA extNTA ta\<close> \<open>e' = extRet2J (addr a\<bullet>M(map Val vs)) va\<close> \<open>s' = (h', lcl s)\<close>
show ?case by(cases U)(auto 4 5 intro: red_reds.RedCallExternal simp del: split_paired_Ex)
next
case RedCallNull thus ?case by(fastforce intro: red_reds.intros)
next
case (BlockRed e h l V vo ta e' h' l' T E T')
note IH = BlockRed.hyps(2)
from IH[of "E(V \<mapsto> T)" T'] \<open>P,E,H \<turnstile> {V:T=vo; e} : T'\<close> \<open>hext H (hp (h, l))\<close>
show ?case by(fastforce dest: red_reds.BlockRed)
next
case RedBlock thus ?case by(fastforce intro: red_reds.intros)
next
case SynchronizedRed1 thus ?case by(fastforce intro: red_reds.intros)
next
case SynchronizedNull thus ?case by(fastforce intro: red_reds.intros)
next
case LockSynchronized thus ?case by(fastforce intro: red_reds.intros)
next
case SynchronizedRed2 thus ?case by(fastforce intro: red_reds.intros)
next
case UnlockSynchronized thus ?case by(fastforce intro: red_reds.intros)
next
case SeqRed thus ?case by(fastforce intro: red_reds.intros)
next
case RedSeq thus ?case by(fastforce intro: red_reds.intros)
next
case CondRed thus ?case by(fastforce intro: red_reds.intros)
next
case RedCondT thus ?case by(fastforce intro: red_reds.intros)
next
case RedCondF thus ?case by(fastforce intro: red_reds.intros)
next
case RedWhile thus ?case by(fastforce intro: red_reds.intros)
next
case ThrowRed thus ?case by(fastforce intro: red_reds.intros)
next
case RedThrowNull thus ?case by(fastforce intro: red_reds.intros)
next
case TryRed thus ?case by(fastforce intro: red_reds.intros)
next
case RedTry thus ?case by(fastforce intro: red_reds.intros)
next
case (RedTryCatch s a D C V e2 E T)
from \<open>P,E,H \<turnstile> try Throw a catch(C V) e2 : T\<close>
obtain T' where "P,E,H \<turnstile> addr a : T'" by auto
with \<open>typeof_addr (hp s) a = \<lfloor>Class_type D\<rfloor>\<close> \<open>hext H (hp s)\<close>
have Ha: "typeof_addr H a = \<lfloor>Class_type D\<rfloor>"
by(auto dest: typeof_addr_hext_mono)
with \<open>P \<turnstile> D \<preceq>\<^sup>* C\<close> show ?case
by(fastforce intro: red_reds.RedTryCatch)
next
case (RedTryFail s a D C V e2 E T)
from \<open>P,E,H \<turnstile> try Throw a catch(C V) e2 : T\<close>
obtain T' where "P,E,H \<turnstile> addr a : T'" by auto
with \<open>typeof_addr (hp s) a = \<lfloor>Class_type D\<rfloor>\<close> \<open>hext H (hp s)\<close>
have Ha: "typeof_addr H a = \<lfloor>Class_type D\<rfloor>"
by(auto dest: typeof_addr_hext_mono)
with \<open>\<not> P \<turnstile> D \<preceq>\<^sup>* C\<close> show ?case
by(fastforce intro: red_reds.RedTryFail)
next
case ListRed1 thus ?case by(fastforce intro: red_reds.intros)
next
case ListRed2 thus ?case by(fastforce intro: red_reds.intros)
next
case NewArrayThrow thus ?case by(fastforce intro: red_reds.intros)
next
case CastThrow thus ?case by(fastforce intro: red_reds.intros)
next
case InstanceOfThrow thus ?case by(fastforce intro: red_reds.intros)
next
case BinOpThrow1 thus ?case by(fastforce intro: red_reds.intros)
next
case BinOpThrow2 thus ?case by(fastforce intro: red_reds.intros)
next
case LAssThrow thus ?case by(fastforce intro: red_reds.intros)
next
case AAccThrow1 thus ?case by(fastforce intro: red_reds.intros)
next
case AAccThrow2 thus ?case by(fastforce intro: red_reds.intros)
next
case AAssThrow1 thus ?case by(fastforce intro: red_reds.intros)
next
case AAssThrow2 thus ?case by(fastforce intro: red_reds.intros)
next
case AAssThrow3 thus ?case by(fastforce intro: red_reds.intros)
next
case ALengthThrow thus ?case by(fastforce intro: red_reds.intros)
next
case FAccThrow thus ?case by(fastforce intro: red_reds.intros)
next
case FAssThrow1 thus ?case by(fastforce intro: red_reds.intros)
next
case FAssThrow2 thus ?case by(fastforce intro: red_reds.intros)
next
case CASThrow thus ?case by(fastforce intro: red_reds.intros)
next
case CASThrow2 thus ?case by(fastforce intro: red_reds.intros)
next
case CASThrow3 thus ?case by(fastforce intro: red_reds.intros)
next
case CallThrowObj thus ?case by(fastforce intro: red_reds.intros)
next
case CallThrowParams thus ?case by(fastforce intro: red_reds.intros)
next
case BlockThrow thus ?case by(fastforce intro: red_reds.intros)
next
case SynchronizedThrow1 thus ?case by(fastforce intro: red_reds.intros)
next
case SynchronizedThrow2 thus ?case by(fastforce intro: red_reds.intros)
next
case SeqThrow thus ?case by(fastforce intro: red_reds.intros)
next
case CondThrow thus ?case by(fastforce intro: red_reds.intros)
next
case ThrowThrow thus ?case by(fastforce intro: red_reds.intros)
qed
lemma can_lock_devreserp:
"\<lbrakk> wf_J_prog P; red_mthr.can_sync P t (e, l) h' L; P,E,h \<turnstile> e : T; P,h \<turnstile> t \<surd>t; hconf h; h \<unlhd> h' \<rbrakk>
\<Longrightarrow> red_mthr.can_sync P t (e, l) h L"
apply(erule red_mthr.can_syncE)
apply(clarsimp)
apply(drule red_wt_hconf_hext, assumption+)
apply(simp)
apply(fastforce intro!: red_mthr.can_syncI)
done
end
context J_typesafe begin
lemma preserve_deadlocked:
assumes wf: "wf_J_prog P"
shows "preserve_deadlocked final_expr (mred P) convert_RA ({s. sync_es_ok (thr s) (shr s) \<and> lock_ok (locks s) (thr s)} \<inter> {s. \<exists>Es. sconf_type_ts_ok Es (thr s) (shr s)} \<inter> {s. def_ass_ts_ok (thr s) (shr s)})"
(is "preserve_deadlocked _ _ _ ?wf_state")
proof(unfold_locales)
show inv: "invariant3p (mredT P) ?wf_state"
by(intro invariant3p_IntI invariant3p_sync_es_ok_lock_ok[OF wf] lifting_inv.invariant3p_ts_inv[OF lifting_inv_sconf_subject_ok[OF wf]] lifting_wf.invariant3p_ts_ok[OF lifting_wf_def_ass[OF wf]])
fix s t' ta' s' t x ln
assume wfs: "s \<in> ?wf_state"
and redT: "P \<turnstile> s -t'\<triangleright>ta'\<rightarrow> s'"
and tst: "thr s t = \<lfloor>(x, ln)\<rfloor>"
from redT have hext: "shr s \<unlhd> shr s'" by(rule redT_hext_incr)
from inv redT wfs have wfs': "s' \<in> ?wf_state" by(rule invariant3pD)
from redT tst obtain x' ln' where ts't: "thr s' t= \<lfloor>(x', ln')\<rfloor>"
by(cases "thr s' t")(cases s, cases s', auto dest: red_mthr.redT_thread_not_disappear)
from wfs tst obtain E T where wt: "P,E,shr s \<turnstile> fst x : T"
and hconf: "hconf (shr s)"
and da: "\<D> (fst x) \<lfloor>dom (snd x)\<rfloor>"
and tconf: "P,shr s \<turnstile> t \<surd>t"
by(force dest: ts_invD ts_okD simp add: type_ok_def sconf_def sconf_type_ok_def)
from wt hext have wt': "P,E,shr s' \<turnstile> fst x : T" by(rule WTrt_hext_mono)
from wfs' ts't have hconf': "hconf (shr s')"
by(auto dest: ts_invD simp add: type_ok_def sconf_def sconf_type_ok_def)
{
assume cs: "red_mthr.must_sync P t x (shr s)"
from cs have "\<not> final (fst x)" by(auto elim!: red_mthr.must_syncE simp add: split_beta)
from progress[OF wf_prog_wwf_prog[OF wf] hconf' wt' da this, of "extTA2J P" t]
obtain e' h x' ta where "P,t \<turnstile> \<langle>fst x,(shr s', snd x)\<rangle> -ta\<rightarrow> \<langle>e', (h, x')\<rangle>" by auto
with red_ta_satisfiable[OF this]
show "red_mthr.must_sync P t x (shr s')"
by-(rule red_mthr.must_syncI, fastforce simp add: split_beta)
next
fix LT
assume "red_mthr.can_sync P t x (shr s') LT"
with can_lock_devreserp[OF wf _ wt tconf hconf hext, of "snd x" LT]
show "\<exists>LT'\<subseteq>LT. red_mthr.can_sync P t x (shr s) LT'" by auto
}
qed
end
end
diff --git a/thys/JinjaThreads/J/TypeSafe.thy b/thys/JinjaThreads/J/TypeSafe.thy
--- a/thys/JinjaThreads/J/TypeSafe.thy
+++ b/thys/JinjaThreads/J/TypeSafe.thy
@@ -1,886 +1,886 @@
(* Title: JinjaThreads/J/SmallTypeSafe.thy
Author: Tobias Nipkow, Andreas Lochbihler
*)
section \<open>Type Safety Proof\<close>
theory TypeSafe
imports
Progress
DefAssPreservation
begin
subsection\<open>Basic preservation lemmas\<close>
text\<open>First two easy preservation lemmas.\<close>
theorem (in J_conf_read)
shows red_preserves_hconf:
"\<lbrakk> extTA,P,t \<turnstile> \<langle>e,s\<rangle> -ta\<rightarrow> \<langle>e',s'\<rangle>; P,E,hp s \<turnstile> e : T; hconf (hp s) \<rbrakk> \<Longrightarrow> hconf (hp s')"
and reds_preserves_hconf:
"\<lbrakk> extTA,P,t \<turnstile> \<langle>es,s\<rangle> [-ta\<rightarrow>] \<langle>es',s'\<rangle>; P,E,hp s \<turnstile> es [:] Ts; hconf (hp s) \<rbrakk> \<Longrightarrow> hconf (hp s')"
proof (induct arbitrary: T E and Ts E rule: red_reds.inducts)
case RedNew thus ?case
by(auto intro: hconf_heap_ops_mono)
next
case RedNewFail thus ?case
by(auto intro: hconf_heap_ops_mono)
next
case RedNewArray thus ?case
by(auto intro: hconf_heap_ops_mono)
next
case RedNewArrayFail thus ?case
by(auto intro: hconf_heap_ops_mono)
next
case (RedAAss h a U n i v U' h' l)
from \<open>sint i < int n\<close> \<open>0 <=s i\<close>
have "nat (sint i) < n" by(metis nat_less_iff sint_0 word_sle_def)
thus ?case using RedAAss
by(fastforce elim: hconf_heap_write_mono intro: addr_loc_type.intros simp add: conf_def)
next
case RedFAss thus ?case
by(fastforce elim: hconf_heap_write_mono intro: addr_loc_type.intros simp add: conf_def)
next
case RedCASSucceed thus ?case
by(fastforce elim: hconf_heap_write_mono intro: addr_loc_type.intros simp add: conf_def)
next
case (RedCallExternal s a U M Ts T' D vs ta va h' ta' e' s')
hence "P,hp s \<turnstile> a\<bullet>M(vs) : T"
by(fastforce simp add: external_WT'_iff dest: sees_method_fun)
with RedCallExternal show ?case by(auto dest: external_call_hconf)
qed auto
theorem (in J_heap) red_preserves_lconf:
"\<lbrakk> extTA,P,t \<turnstile> \<langle>e,s\<rangle> -ta\<rightarrow> \<langle>e',s'\<rangle>; P,E,hp s \<turnstile> e:T; P,hp s \<turnstile> lcl s (:\<le>) E \<rbrakk> \<Longrightarrow> P,hp s' \<turnstile> lcl s' (:\<le>) E"
and reds_preserves_lconf:
"\<lbrakk> extTA,P,t \<turnstile> \<langle>es,s\<rangle> [-ta\<rightarrow>] \<langle>es',s'\<rangle>; P,E,hp s \<turnstile> es[:]Ts; P,hp s \<turnstile> lcl s (:\<le>) E \<rbrakk> \<Longrightarrow> P,hp s' \<turnstile> lcl s' (:\<le>) E"
proof(induct arbitrary: T E and Ts E rule:red_reds.inducts)
case RedNew thus ?case
by(fastforce intro:lconf_hext hext_heap_ops simp del: fun_upd_apply)
next
case RedNewFail thus ?case
by(auto intro:lconf_hext hext_heap_ops simp del: fun_upd_apply)
next
case RedNewArray thus ?case
by(fastforce intro:lconf_hext hext_heap_ops simp del: fun_upd_apply)
next
case RedNewArrayFail thus ?case
by(fastforce intro:lconf_hext hext_heap_ops simp del: fun_upd_apply)
next
case RedLAss thus ?case
by(fastforce elim: lconf_upd simp add: conf_def simp del: fun_upd_apply )
next
case RedAAss thus ?case
by(fastforce intro:lconf_hext hext_heap_ops simp del: fun_upd_apply)
next
case RedFAss thus ?case
by(fastforce intro:lconf_hext hext_heap_ops simp del: fun_upd_apply)
next
case RedCASSucceed thus ?case
by(fastforce intro:lconf_hext hext_heap_ops simp del: fun_upd_apply)
next
case (BlockRed e h x V vo ta e' h' x' T T' E)
note red = \<open>extTA,P,t \<turnstile> \<langle>e,(h, x(V := vo))\<rangle> -ta\<rightarrow> \<langle>e',(h', x')\<rangle>\<close>
note IH = \<open>\<And>T E. \<lbrakk>P,E,hp (h, x(V := vo)) \<turnstile> e : T;
P,hp (h, x(V := vo)) \<turnstile> lcl (h, x(V := vo)) (:\<le>) E\<rbrakk>
\<Longrightarrow> P,hp (h', x') \<turnstile> lcl (h', x') (:\<le>) E\<close>
note wt = \<open>P,E,hp (h, x) \<turnstile> {V:T=vo; e} : T'\<close>
note lconf = \<open>P,hp (h, x) \<turnstile> lcl (h, x) (:\<le>) E\<close>
from lconf_hext[OF lconf[simplified] red_hext_incr[OF red, simplified]]
have "P,h' \<turnstile> x (:\<le>) E" .
moreover from wt have "P,E(V\<mapsto>T),h \<turnstile> e : T'" by(cases vo, auto)
moreover from lconf wt have "P,h \<turnstile> x(V := vo) (:\<le>) E(V \<mapsto> T)"
by(cases vo)(simp add: lconf_def,auto intro: lconf_upd2 simp add: conf_def)
ultimately have "P,h' \<turnstile> x' (:\<le>) E(V\<mapsto>T)"
by(auto intro: IH[simplified])
with \<open>P,h' \<turnstile> x (:\<le>) E\<close> show ?case
by(auto simp add: lconf_def split: if_split_asm)
next
case (RedCallExternal s a U M Ts T' D vs ta va h' ta' e' s')
from \<open>P,t \<turnstile> \<langle>a\<bullet>M(vs),hp s\<rangle> -ta\<rightarrow>ext \<langle>va,h'\<rangle>\<close> have "hp s \<unlhd> h'" by(rule red_external_hext)
with \<open>s' = (h', lcl s)\<close> \<open>P,hp s \<turnstile> lcl s (:\<le>) E\<close> show ?case by(auto intro: lconf_hext)
qed auto
text\<open>Combining conformance of heap and local variables:\<close>
definition (in J_heap_conf_base) sconf :: "env \<Rightarrow> ('addr, 'heap) Jstate \<Rightarrow> bool" ("_ \<turnstile> _ \<surd>" [51,51]50)
where "E \<turnstile> s \<surd> \<equiv> let (h,l) = s in hconf h \<and> P,h \<turnstile> l (:\<le>) E \<and> preallocated h"
context J_conf_read begin
lemma red_preserves_sconf:
"\<lbrakk> extTA,P,t \<turnstile> \<langle>e,s\<rangle> -tas\<rightarrow> \<langle>e',s'\<rangle>; P,E,hp s \<turnstile> e : T; E \<turnstile> s \<surd> \<rbrakk> \<Longrightarrow> E \<turnstile> s' \<surd>"
apply(auto dest: red_preserves_hconf red_preserves_lconf simp add:sconf_def)
apply(fastforce dest: red_hext_incr intro: preallocated_hext)
done
lemma reds_preserves_sconf:
"\<lbrakk> extTA,P,t \<turnstile> \<langle>es,s\<rangle> [-ta\<rightarrow>] \<langle>es',s'\<rangle>; P,E,hp s \<turnstile> es [:] Ts; E \<turnstile> s \<surd> \<rbrakk> \<Longrightarrow> E \<turnstile> s' \<surd>"
apply(auto dest: reds_preserves_hconf reds_preserves_lconf simp add: sconf_def)
apply(fastforce dest: reds_hext_incr intro: preallocated_hext)
done
end
lemma (in J_heap_base) wt_external_call:
"\<lbrakk> conf_extRet P h va T; P,E,h \<turnstile> e : T \<rbrakk> \<Longrightarrow> \<exists>T'. P,E,h \<turnstile> extRet2J e va : T' \<and> P \<turnstile> T' \<le> T"
by(cases va)(auto simp add: conf_def)
subsection "Subject reduction"
theorem (in J_conf_read) assumes wf: "wf_J_prog P"
shows subject_reduction:
"\<lbrakk> extTA,P,t \<turnstile> \<langle>e,s\<rangle> -ta\<rightarrow> \<langle>e',s'\<rangle>; E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e:T; P,hp s \<turnstile> t \<surd>t \<rbrakk>
\<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> e':T' \<and> P \<turnstile> T' \<le> T"
and subjects_reduction:
"\<lbrakk> extTA,P,t \<turnstile> \<langle>es,s\<rangle> [-ta\<rightarrow>] \<langle>es',s'\<rangle>; E \<turnstile> s \<surd>; P,E,hp s \<turnstile> es[:]Ts; P,hp s \<turnstile> t \<surd>t \<rbrakk>
\<Longrightarrow> \<exists>Ts'. P,E,hp s' \<turnstile> es'[:]Ts' \<and> P \<turnstile> Ts' [\<le>] Ts"
proof (induct arbitrary: T E and Ts E rule:red_reds.inducts)
case RedNew
thus ?case by(auto dest: allocate_SomeD)
next
case RedNewFail thus ?case unfolding sconf_def
by(fastforce intro:typeof_OutOfMemory preallocated_heap_ops simp add: xcpt_subcls_Throwable[OF _ wf])
next
case NewArrayRed
thus ?case by fastforce
next
case RedNewArray
thus ?case by(auto dest: allocate_SomeD)
next
case RedNewArrayNegative thus ?case unfolding sconf_def
by(fastforce intro: preallocated_heap_ops simp add: xcpt_subcls_Throwable[OF _ wf])
next
case RedNewArrayFail thus ?case unfolding sconf_def
by(fastforce intro:typeof_OutOfMemory preallocated_heap_ops simp add: xcpt_subcls_Throwable[OF _ wf])
next
case (CastRed e s ta e' s' C T E)
have esse: "extTA,P,t \<turnstile> \<langle>e,s\<rangle> -ta\<rightarrow> \<langle>e',s'\<rangle>"
and IH: "\<And>T E. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T"
and hconf: "E \<turnstile> s \<surd>"
and wtc: "P,E,hp s \<turnstile> Cast C e : T" by fact+
thus ?case
proof(clarsimp)
fix T'
- assume wte: "P,E,hp s \<turnstile> e : T'" "is_type P T"
+ assume wte: "P,E,hp s \<turnstile> e : T'" "is_type P C"
from wte and hconf and IH and \<open>P,hp s \<turnstile> t \<surd>t\<close> have "\<exists>U. P,E,hp s' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T'" by simp
then obtain U where wtee: "P,E,hp s' \<turnstile> e' : U" and UsTT: "P \<turnstile> U \<le> T'" by blast
- from wtee \<open>is_type P T\<close> have "P,E,hp s' \<turnstile> Cast T e' : T" by(rule WTrtCast)
- thus "\<exists>T'. P,E,hp s' \<turnstile> Cast T e' : T' \<and> P \<turnstile> T' \<le> T" by blast
+ from wtee \<open>is_type P C\<close> have "P,E,hp s' \<turnstile> Cast C e' : C" by(rule WTrtCast)
+ thus "\<exists>T'. P,E,hp s' \<turnstile> Cast C e' : T' \<and> P \<turnstile> T' \<le> C" by blast
qed
next
case RedCast thus ?case
by(clarsimp simp add: is_refT_def)
next
case RedCastFail thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case (InstanceOfRed e s ta e' s' U T E)
have IH: "\<And>T E. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T"
and hconf: "E \<turnstile> s \<surd>"
and wtc: "P,E,hp s \<turnstile> e instanceof U : T"
and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
from wtc obtain T' where "P,E,hp s \<turnstile> e : T'" by auto
from IH[OF hconf this tconf] obtain T'' where "P,E,hp s' \<turnstile> e' : T''" by auto
with wtc show ?case by auto
next
case RedInstanceOf thus ?case
by(clarsimp)
next
case (BinOpRed1 e\<^sub>1 s ta e\<^sub>1' s' bop e\<^sub>2 T E)
have red: "extTA,P,t \<turnstile> \<langle>e\<^sub>1, s\<rangle> -ta\<rightarrow> \<langle>e\<^sub>1', s'\<rangle>"
and IH: "\<And>T E. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e\<^sub>1:T; P,hp s \<turnstile> t \<surd>t\<rbrakk>
\<Longrightarrow> \<exists>U. P,E,hp s' \<turnstile> e\<^sub>1' : U \<and> P \<turnstile> U \<le> T"
and conf: "E \<turnstile> s \<surd>" and wt: "P,E,hp s \<turnstile> e\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T"
and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
from wt obtain T1 T2 where wt1: "P,E,hp s \<turnstile> e\<^sub>1 : T1"
and wt2: "P,E,hp s \<turnstile> e\<^sub>2 : T2" and wtbop: "P \<turnstile> T1\<guillemotleft>bop\<guillemotright>T2 : T" by auto
from IH[OF conf wt1 tconf] obtain T1' where wt1': "P,E,hp s' \<turnstile> e\<^sub>1' : T1'"
and sub: "P \<turnstile> T1' \<le> T1" by blast
from WTrt_binop_widen_mono[OF wtbop sub widen_refl]
obtain T' where wtbop': "P \<turnstile> T1'\<guillemotleft>bop\<guillemotright>T2 : T'" and sub': "P \<turnstile> T' \<le> T" by blast
from wt1' WTrt_hext_mono[OF wt2 red_hext_incr[OF red]] wtbop'
have "P,E,hp s' \<turnstile> e\<^sub>1' \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T'" by(rule WTrtBinOp)
with sub' show ?case by blast
next
case (BinOpRed2 e\<^sub>2 s ta e\<^sub>2' s' v\<^sub>1 bop T E)
have red: "extTA,P,t \<turnstile> \<langle>e\<^sub>2,s\<rangle> -ta\<rightarrow> \<langle>e\<^sub>2',s'\<rangle>" by fact
have IH: "\<And>E T. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e\<^sub>2:T; P,hp s \<turnstile> t \<surd>t\<rbrakk>
\<Longrightarrow> \<exists>U. P,E,hp s' \<turnstile> e\<^sub>2' : U \<and> P \<turnstile> U \<le> T"
and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
have conf: "E \<turnstile> s \<surd>" and wt: "P,E,hp s \<turnstile> (Val v\<^sub>1) \<guillemotleft>bop\<guillemotright> e\<^sub>2 : T" by fact+
from wt obtain T1 T2 where wt1: "P,E,hp s \<turnstile> Val v\<^sub>1 : T1"
and wt2: "P,E,hp s \<turnstile> e\<^sub>2 : T2" and wtbop: "P \<turnstile> T1\<guillemotleft>bop\<guillemotright>T2 : T" by auto
from IH[OF conf wt2 tconf] obtain T2' where wt2': "P,E,hp s' \<turnstile> e\<^sub>2' : T2'"
and sub: "P \<turnstile> T2' \<le> T2" by blast
from WTrt_binop_widen_mono[OF wtbop widen_refl sub]
obtain T' where wtbop': "P \<turnstile> T1\<guillemotleft>bop\<guillemotright>T2' : T'" and sub': "P \<turnstile> T' \<le> T" by blast
from WTrt_hext_mono[OF wt1 red_hext_incr[OF red]] wt2' wtbop'
have "P,E,hp s' \<turnstile> Val v\<^sub>1 \<guillemotleft>bop\<guillemotright> e\<^sub>2' : T'" by(rule WTrtBinOp)
with sub' show ?case by blast
next
case (RedBinOp bop v1 v2 v s)
from \<open>E \<turnstile> s \<surd>\<close> have preh: "preallocated (hp s)" by(cases s)(simp add: sconf_def)
from \<open>P,E,hp s \<turnstile> Val v1 \<guillemotleft>bop\<guillemotright> Val v2 : T\<close> obtain T1 T2
where "typeof\<^bsub>hp s\<^esub> v1 = \<lfloor>T1\<rfloor>" "typeof\<^bsub>hp s\<^esub> v2 = \<lfloor>T2\<rfloor>" "P \<turnstile> T1\<guillemotleft>bop\<guillemotright>T2 : T" by auto
with wf preh have "P,hp s \<turnstile> v :\<le> T" using \<open>binop bop v1 v2 = \<lfloor>Inl v\<rfloor>\<close>
by(rule binop_type)
thus ?case by(auto simp add: conf_def)
next
case (RedBinOpFail bop v1 v2 a s)
from \<open>E \<turnstile> s \<surd>\<close> have preh: "preallocated (hp s)" by(cases s)(simp add: sconf_def)
from \<open>P,E,hp s \<turnstile> Val v1 \<guillemotleft>bop\<guillemotright> Val v2 : T\<close> obtain T1 T2
where "typeof\<^bsub>hp s\<^esub> v1 = \<lfloor>T1\<rfloor>" "typeof\<^bsub>hp s\<^esub> v2 = \<lfloor>T2\<rfloor>" "P \<turnstile> T1\<guillemotleft>bop\<guillemotright>T2 : T" by auto
with wf preh have "P,hp s \<turnstile> Addr a :\<le> Class Throwable" using \<open>binop bop v1 v2 = \<lfloor>Inr a\<rfloor>\<close>
by(rule binop_type)
thus ?case by(auto simp add: conf_def)
next
case RedVar thus ?case by (fastforce simp:sconf_def lconf_def conf_def)
next
case LAssRed thus ?case by(blast intro:widen_trans)
next
case RedLAss thus ?case by fastforce
next
case (AAccRed1 a s ta a' s' i T E)
have IH: "\<And>E T. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> a : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> a' : T' \<and> P \<turnstile> T' \<le> T"
and assa: "extTA,P,t \<turnstile> \<langle>a,s\<rangle> -ta\<rightarrow> \<langle>a',s'\<rangle>"
and wt: "P,E,hp s \<turnstile> a\<lfloor>i\<rceil> : T"
and hconf: "E \<turnstile> s \<surd>"
and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
from wt have wti: "P,E,hp s \<turnstile> i : Integer" by auto
from wti red_hext_incr[OF assa] have wti': "P,E,hp s' \<turnstile> i : Integer" by - (rule WTrt_hext_mono)
{ assume wta: "P,E,hp s \<turnstile> a : T\<lfloor>\<rceil>"
from IH[OF hconf wta tconf]
obtain U where wta': "P,E,hp s' \<turnstile> a' : U" and UsubT: "P \<turnstile> U \<le> T\<lfloor>\<rceil>" by fastforce
with wta' wti' have ?case by(cases U, auto simp add: widen_Array) }
moreover
{ assume wta: "P,E,hp s \<turnstile> a : NT"
from IH[OF hconf wta tconf] have "P,E,hp s' \<turnstile> a' : NT" by fastforce
from this wti' have ?case
by(fastforce intro:WTrtAAccNT) }
ultimately show ?case using wt by auto
next
case (AAccRed2 i s ta i' s' a T E)
have IH: "\<And>E T. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> i : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> i' : T' \<and> P \<turnstile> T' \<le> T"
and issi: "extTA,P,t \<turnstile> \<langle>i,s\<rangle> -ta\<rightarrow> \<langle>i',s'\<rangle>"
and wt: "P,E,hp s \<turnstile> Val a\<lfloor>i\<rceil> : T"
and sconf: "E \<turnstile> s \<surd>"
and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
from wt have wti: "P,E,hp s \<turnstile> i : Integer" by auto
from wti IH sconf tconf have wti': "P,E,hp s' \<turnstile> i' : Integer" by blast
from wt show ?case
proof (rule WTrt_elim_cases)
assume wta: "P,E,hp s \<turnstile> Val a : T\<lfloor>\<rceil>"
from wta red_hext_incr[OF issi] have wta': "P,E,hp s' \<turnstile> Val a : T\<lfloor>\<rceil>" by (rule WTrt_hext_mono)
from wta' wti' show ?case by(fastforce)
next
assume wta: "P,E,hp s \<turnstile> Val a : NT"
from wta red_hext_incr[OF issi] have wta': "P,E,hp s' \<turnstile> Val a : NT" by (rule WTrt_hext_mono)
from wta' wti' show ?case
by(fastforce elim:WTrtAAccNT)
qed
next
case RedAAccNull thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case RedAAccBounds thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case (RedAAcc h a T n i v l T' E)
from \<open>E \<turnstile> (h, l) \<surd>\<close> have "hconf h" by(clarsimp simp add: sconf_def)
from \<open>0 <=s i\<close> \<open>sint i < int n\<close>
have "nat (sint i) < n" by(metis nat_less_iff sint_0 word_sle_def)
with \<open>typeof_addr h a = \<lfloor>Array_type T n\<rfloor>\<close> have "P,h \<turnstile> a@ACell (nat (sint i)) : T"
by(auto intro: addr_loc_type.intros)
from heap_read_conf[OF \<open>heap_read h a (ACell (nat (sint i))) v\<close> this] \<open>hconf h\<close>
have "P,h \<turnstile> v :\<le> T" by simp
thus ?case using RedAAcc by(auto simp add: conf_def)
next
case (AAssRed1 a s ta a' s' i e T E)
have IH: "\<And>E T. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> a : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> a' : T' \<and> P \<turnstile> T' \<le> T"
and assa: "extTA,P,t \<turnstile> \<langle>a,s\<rangle> -ta\<rightarrow> \<langle>a',s'\<rangle>"
and wt: "P,E,hp s \<turnstile> a\<lfloor>i\<rceil> := e : T"
and sconf: "E \<turnstile> s \<surd>"
and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
from wt have void: "T = Void" by blast
from wt have wti: "P,E,hp s \<turnstile> i : Integer" by auto
from wti red_hext_incr[OF assa] have wti': "P,E,hp s' \<turnstile> i : Integer" by - (rule WTrt_hext_mono)
{ assume wta: "P,E,hp s \<turnstile> a : NT"
from IH[OF sconf wta tconf] have wta': "P,E,hp s' \<turnstile> a' : NT" by fastforce
from wt wta obtain V where wte: "P,E,hp s \<turnstile> e : V" by(auto)
from wte red_hext_incr[OF assa] have wte': "P,E,hp s' \<turnstile> e : V" by - (rule WTrt_hext_mono)
from wta' wti' wte' void have ?case
by(fastforce elim: WTrtAAssNT) }
moreover
{ fix U
assume wta: "P,E,hp s \<turnstile> a : U\<lfloor>\<rceil>"
from IH[OF sconf wta tconf]
obtain U' where wta': "P,E,hp s' \<turnstile> a' : U'" and UsubT: "P \<turnstile> U' \<le> U\<lfloor>\<rceil>" by fastforce
with wta' have ?case
proof(cases U')
case NT
assume UNT: "U' = NT"
from UNT wt wta obtain V where wte: "P,E,hp s \<turnstile> e : V" by(auto)
from wte red_hext_incr[OF assa] have wte': "P,E,hp s' \<turnstile> e : V" by - (rule WTrt_hext_mono)
from wta' UNT wti' wte' void show ?thesis
by(fastforce elim: WTrtAAssNT)
next
case (Array A)
have UA: "U' = A\<lfloor>\<rceil>" by fact
with UA UsubT wt wta obtain V where wte: "P,E,hp s \<turnstile> e : V" by auto
from wte red_hext_incr[OF assa] have wte': "P,E,hp s' \<turnstile> e : V" by - (rule WTrt_hext_mono)
with wta' wte' UA wti' void show ?thesis by (fast elim:WTrtAAss)
qed(simp_all add: widen_Array) }
ultimately show ?case using wt by blast
next
case (AAssRed2 i s ta i' s' a e T E)
have IH: "\<And>E T. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> i : T; P,hp s \<turnstile> t \<surd>t \<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> i' : T' \<and> P \<turnstile> T' \<le> T"
and issi: "extTA,P,t \<turnstile> \<langle>i,s\<rangle> -ta\<rightarrow> \<langle>i',s'\<rangle>"
and wt: "P,E,hp s \<turnstile> Val a\<lfloor>i\<rceil> := e : T"
and sconf: "E \<turnstile> s \<surd>" and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
from wt have void: "T = Void" by blast
from wt have wti: "P,E,hp s \<turnstile> i : Integer" by auto
from IH[OF sconf wti tconf] have wti': "P,E,hp s' \<turnstile> i' : Integer" by fastforce
from wt show ?case
proof(rule WTrt_elim_cases)
fix U T'
assume wta: "P,E,hp s \<turnstile> Val a : U\<lfloor>\<rceil>"
and wte: "P,E,hp s \<turnstile> e : T'"
from wte red_hext_incr[OF issi] have wte': "P,E,hp s' \<turnstile> e : T'" by - (rule WTrt_hext_mono)
from wta red_hext_incr[OF issi] have wta': "P,E,hp s' \<turnstile> Val a : U\<lfloor>\<rceil>" by - (rule WTrt_hext_mono)
from wta' wti' wte' void show ?case by (fastforce elim:WTrtAAss)
next
fix T'
assume wta: "P,E,hp s \<turnstile> Val a : NT"
and wte: "P,E,hp s \<turnstile> e : T'"
from wte red_hext_incr[OF issi] have wte': "P,E,hp s' \<turnstile> e : T'" by - (rule WTrt_hext_mono)
from wta red_hext_incr[OF issi] have wta': "P,E,hp s' \<turnstile> Val a : NT" by - (rule WTrt_hext_mono)
from wta' wti' wte' void show ?case by (fastforce elim:WTrtAAss)
qed
next
case (AAssRed3 e s ta e' s' a i T E)
have IH: "\<And>E T. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T"
and issi: "extTA,P,t \<turnstile> \<langle>e,s\<rangle> -ta\<rightarrow> \<langle>e',s'\<rangle>"
and wt: "P,E,hp s \<turnstile> Val a\<lfloor>Val i\<rceil> := e : T"
and sconf: "E \<turnstile> s \<surd>" and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
from wt have void: "T = Void" by blast
from wt have wti: "P,E,hp s \<turnstile> Val i : Integer" by auto
from wti red_hext_incr[OF issi] have wti': "P,E,hp s' \<turnstile> Val i : Integer" by - (rule WTrt_hext_mono)
from wt show ?case
proof(rule WTrt_elim_cases)
fix U T'
assume wta: "P,E,hp s \<turnstile> Val a : U\<lfloor>\<rceil>"
and wte: "P,E,hp s \<turnstile> e : T'"
from wta red_hext_incr[OF issi] have wta': "P,E,hp s' \<turnstile> Val a : U\<lfloor>\<rceil>" by - (rule WTrt_hext_mono)
from IH[OF sconf wte tconf]
obtain V where wte': "P,E,hp s' \<turnstile> e' : V" by fastforce
from wta' wti' wte' void show ?case by (fastforce elim:WTrtAAss)
next
fix T'
assume wta: "P,E,hp s \<turnstile> Val a : NT"
and wte: "P,E,hp s \<turnstile> e : T'"
from wta red_hext_incr[OF issi] have wta': "P,E,hp s' \<turnstile> Val a : NT" by - (rule WTrt_hext_mono)
from IH[OF sconf wte tconf]
obtain V where wte': "P,E,hp s' \<turnstile> e' : V" by fastforce
from wta' wti' wte' void show ?case by (fastforce elim:WTrtAAss)
qed
next
case RedAAssNull thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case RedAAssBounds thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case RedAAssStore thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case RedAAss thus ?case
by(auto simp del:fun_upd_apply)
next
case (ALengthRed a s ta a' s' T E)
note IH = \<open>\<And>T'. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> a : T'; P,hp s \<turnstile> t \<surd>t\<rbrakk>
\<Longrightarrow> \<exists>T''. P,E,hp s' \<turnstile> a' : T'' \<and> P \<turnstile> T'' \<le> T'\<close>
from \<open>P,E,hp s \<turnstile> a\<bullet>length : T\<close>
show ?case
proof(rule WTrt_elim_cases)
fix T'
assume [simp]: "T = Integer"
and wta: "P,E,hp s \<turnstile> a : T'\<lfloor>\<rceil>"
from wta \<open>E \<turnstile> s \<surd>\<close> IH \<open>P,hp s \<turnstile> t \<surd>t\<close>
obtain T'' where wta': "P,E,hp s' \<turnstile> a' : T''"
and sub: "P \<turnstile> T'' \<le> T'\<lfloor>\<rceil>" by blast
from sub have "P,E,hp s' \<turnstile> a'\<bullet>length : Integer"
unfolding widen_Array
proof(rule disjE)
assume "T'' = NT"
with wta' show ?thesis by(auto)
next
assume "\<exists>V. T'' = V\<lfloor>\<rceil> \<and> P \<turnstile> V \<le> T'"
then obtain V where "T'' = V\<lfloor>\<rceil>" "P \<turnstile> V \<le> T'" by blast
with wta' show ?thesis by -(rule WTrtALength, simp)
qed
thus ?thesis by(simp)
next
assume "P,E,hp s \<turnstile> a : NT"
with \<open>E \<turnstile> s \<surd>\<close> IH \<open>P,hp s \<turnstile> t \<surd>t\<close>
obtain T'' where wta': "P,E,hp s' \<turnstile> a' : T''"
and sub: "P \<turnstile> T'' \<le> NT" by blast
from sub have "T'' = NT" by auto
with wta' show ?thesis by(auto)
qed
next
case (RedALength h a T n l T' E)
from \<open>P,E,hp (h, l) \<turnstile> addr a\<bullet>length : T'\<close> \<open>typeof_addr h a = \<lfloor>Array_type T n\<rfloor>\<close>
have [simp]: "T' = Integer" by(auto)
thus ?case by(auto)
next
case RedALengthNull thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case (FAccRed e s ta e' s' F D T E)
have IH: "\<And>E T. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e : T; P,hp s \<turnstile> t \<surd>t\<rbrakk>
\<Longrightarrow> \<exists>U. P,E,hp s' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
and conf: "E \<turnstile> s \<surd>" and wt: "P,E,hp s \<turnstile> e\<bullet>F{D} : T"
and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
\<comment> \<open>Now distinguish the two cases how wt can have arisen.\<close>
{ fix T' C fm
assume wte: "P,E,hp s \<turnstile> e : T'"
and icto: "class_type_of' T' = \<lfloor>C\<rfloor>"
and has: "P \<turnstile> C has F:T (fm) in D"
from IH[OF conf wte tconf]
obtain U where wte': "P,E,hp s' \<turnstile> e' : U" and UsubC: "P \<turnstile> U \<le> T'" by auto
\<comment> \<open>Now distinguish what @{term U} can be.\<close>
with UsubC have ?case
proof(cases "U = NT")
case True
thus ?thesis using wte' by(blast intro:WTrtFAccNT widen_refl)
next
case False
with icto UsubC obtain C' where icto': "class_type_of' U = \<lfloor>C'\<rfloor>"
and C'subC: "P \<turnstile> C' \<preceq>\<^sup>* C"
by(rule widen_is_class_type_of)
from has_field_mono[OF has C'subC] wte' icto'
show ?thesis by(auto intro!:WTrtFAcc)
qed }
moreover
{ assume "P,E,hp s \<turnstile> e : NT"
hence "P,E,hp s' \<turnstile> e' : NT" using IH[OF conf _ tconf] by fastforce
hence ?case by(fastforce intro:WTrtFAccNT widen_refl) }
ultimately show ?case using wt by blast
next
case RedFAcc thus ?case unfolding sconf_def
by(fastforce dest: heap_read_conf intro: addr_loc_type.intros simp add: conf_def)
next
case RedFAccNull thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case (FAssRed1 e s ta e' s' F D e\<^sub>2)
have red: "extTA,P,t \<turnstile> \<langle>e,s\<rangle> -ta\<rightarrow> \<langle>e',s'\<rangle>"
and IH: "\<And>E T. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e : T; P,hp s \<turnstile> t \<surd>t\<rbrakk>
\<Longrightarrow> \<exists>U. P,E,hp s' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
and conf: "E \<turnstile> s \<surd>" and wt: "P,E,hp s \<turnstile> e\<bullet>F{D}:=e\<^sub>2 : T"
and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
from wt have void: "T = Void" by blast
\<comment> \<open>We distinguish if @{term e} has type @{term NT} or a Class type\<close>
{ assume "P,E,hp s \<turnstile> e : NT"
hence "P,E,hp s' \<turnstile> e' : NT" using IH[OF conf _ tconf] by fastforce
moreover obtain T\<^sub>2 where "P,E,hp s \<turnstile> e\<^sub>2 : T\<^sub>2" using wt by auto
from this red_hext_incr[OF red] have "P,E,hp s' \<turnstile> e\<^sub>2 : T\<^sub>2"
by(rule WTrt_hext_mono)
ultimately have ?case using void by(blast intro!:WTrtFAssNT)
}
moreover
{ fix T' C TF T\<^sub>2 fm
assume wt\<^sub>1: "P,E,hp s \<turnstile> e : T'" and icto: "class_type_of' T' = \<lfloor>C\<rfloor>" and wt\<^sub>2: "P,E,hp s \<turnstile> e\<^sub>2 : T\<^sub>2"
and has: "P \<turnstile> C has F:TF (fm) in D" and sub: "P \<turnstile> T\<^sub>2 \<le> TF"
obtain U where wt\<^sub>1': "P,E,hp s' \<turnstile> e' : U" and UsubC: "P \<turnstile> U \<le> T'"
using IH[OF conf wt\<^sub>1 tconf] by blast
have wt\<^sub>2': "P,E,hp s' \<turnstile> e\<^sub>2 : T\<^sub>2"
by(rule WTrt_hext_mono[OF wt\<^sub>2 red_hext_incr[OF red]])
\<comment> \<open>Is @{term U} the null type or a class type?\<close>
have ?case
proof(cases "U = NT")
case True
with wt\<^sub>1' wt\<^sub>2' void show ?thesis by(blast intro!:WTrtFAssNT)
next
case False
with icto UsubC obtain C' where icto': "class_type_of' U = \<lfloor>C'\<rfloor>"
and "subclass": "P \<turnstile> C' \<preceq>\<^sup>* C" by(rule widen_is_class_type_of)
have "P \<turnstile> C' has F:TF (fm) in D" by(rule has_field_mono[OF has "subclass"])
with wt\<^sub>1' show ?thesis using wt\<^sub>2' sub void icto' by(blast intro:WTrtFAss)
qed }
ultimately show ?case using wt by blast
next
case (FAssRed2 e\<^sub>2 s ta e\<^sub>2' s' v F D T E)
have red: "extTA,P,t \<turnstile> \<langle>e\<^sub>2,s\<rangle> -ta\<rightarrow> \<langle>e\<^sub>2',s'\<rangle>"
and IH: "\<And>E T. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e\<^sub>2 : T; P,hp s \<turnstile> t \<surd>t\<rbrakk>
\<Longrightarrow> \<exists>U. P,E,hp s' \<turnstile> e\<^sub>2' : U \<and> P \<turnstile> U \<le> T"
and conf: "E \<turnstile> s \<surd>" and wt: "P,E,hp s \<turnstile> Val v\<bullet>F{D}:=e\<^sub>2 : T"
and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
from wt have [simp]: "T = Void" by auto
from wt show ?case
proof (rule WTrt_elim_cases)
fix U C TF T\<^sub>2 fm
assume wt\<^sub>1: "P,E,hp s \<turnstile> Val v : U"
and icto: "class_type_of' U = \<lfloor>C\<rfloor>"
and has: "P \<turnstile> C has F:TF (fm) in D"
and wt\<^sub>2: "P,E,hp s \<turnstile> e\<^sub>2 : T\<^sub>2" and TsubTF: "P \<turnstile> T\<^sub>2 \<le> TF"
have wt\<^sub>1': "P,E,hp s' \<turnstile> Val v : U"
by(rule WTrt_hext_mono[OF wt\<^sub>1 red_hext_incr[OF red]])
obtain T\<^sub>2' where wt\<^sub>2': "P,E,hp s' \<turnstile> e\<^sub>2' : T\<^sub>2'" and T'subT: "P \<turnstile> T\<^sub>2' \<le> T\<^sub>2"
using IH[OF conf wt\<^sub>2 tconf] by blast
have "P,E,hp s' \<turnstile> Val v\<bullet>F{D}:=e\<^sub>2' : Void"
by(rule WTrtFAss[OF wt\<^sub>1' icto has wt\<^sub>2' widen_trans[OF T'subT TsubTF]])
thus ?case by auto
next
fix T\<^sub>2 assume null: "P,E,hp s \<turnstile> Val v : NT" and wt\<^sub>2: "P,E,hp s \<turnstile> e\<^sub>2 : T\<^sub>2"
from null have "v = Null" by simp
moreover
obtain T\<^sub>2' where "P,E,hp s' \<turnstile> e\<^sub>2' : T\<^sub>2' \<and> P \<turnstile> T\<^sub>2' \<le> T\<^sub>2"
using IH[OF conf wt\<^sub>2 tconf] by blast
ultimately show ?thesis by(fastforce intro:WTrtFAssNT)
qed
next
case RedFAss thus ?case by(auto simp del:fun_upd_apply)
next
case RedFAssNull thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case (CASRed1 e s ta e' s' D F e2 e3)
from CASRed1.prems(2) consider (NT) T2 T3 where
"P,E,hp s \<turnstile> e : NT" "T = Boolean" "P,E,hp s \<turnstile> e2 : T2" "P,E,hp s \<turnstile> e3 : T3"
| (RefT) U T' C fm T2 T3 where
"P,E,hp s \<turnstile> e : U" "T = Boolean" "class_type_of' U = \<lfloor>C\<rfloor>" "P \<turnstile> C has F:T' (fm) in D"
"P,E,hp s \<turnstile> e2 : T2" "P \<turnstile> T2 \<le> T'" "P,E,hp s \<turnstile> e3 : T3" "P \<turnstile> T3 \<le> T'" "volatile fm" by fastforce
thus ?case
proof cases
case NT
have "P,E,hp s' \<turnstile> e' : NT" using CASRed1.hyps(2)[OF CASRed1.prems(1) NT(1) CASRed1.prems(3)] by auto
moreover from NT CASRed1.hyps(1)[THEN red_hext_incr]
have "P,E,hp s' \<turnstile> e2 : T2" "P,E,hp s' \<turnstile> e3 : T3" by(auto intro: WTrt_hext_mono)
ultimately show ?thesis using NT by(auto intro: WTrtCASNT)
next
case RefT
from CASRed1.hyps(2)[OF CASRed1.prems(1) RefT(1) CASRed1.prems(3)]
obtain U' where wt1: "P,E,hp s' \<turnstile> e' : U'" "P \<turnstile> U' \<le> U" by blast
from RefT CASRed1.hyps(1)[THEN red_hext_incr]
have wt2: "P,E,hp s' \<turnstile> e2 : T2" and wt3: "P,E,hp s' \<turnstile> e3 : T3" by(auto intro: WTrt_hext_mono)
show ?thesis
proof(cases "U' = NT")
case True
with RefT wt1 wt2 wt3 show ?thesis by(auto intro: WTrtCASNT)
next
case False
with RefT(3) wt1 obtain C' where icto': "class_type_of' U' = \<lfloor>C'\<rfloor>"
and "subclass": "P \<turnstile> C' \<preceq>\<^sup>* C" by(blast intro: widen_is_class_type_of)
have "P \<turnstile> C' has F:T' (fm) in D" by(rule has_field_mono[OF RefT(4) "subclass"])
with RefT wt1 wt2 wt3 icto' show ?thesis by(auto intro!: WTrtCAS)
qed
qed
next
case (CASRed2 e s ta e' s' v D F e3)
consider (Null) "v = Null" | (Val) U C T' fm T2 T3 where
"class_type_of' U = \<lfloor>C\<rfloor>" "P \<turnstile> C has F:T' (fm) in D" "volatile fm"
"P,E,hp s \<turnstile> e : T2" "P \<turnstile> T2 \<le> T'" "P,E,hp s \<turnstile> e3 : T3" "P \<turnstile> T3 \<le> T'" "T = Boolean"
"typeof\<^bsub>hp s\<^esub> v = \<lfloor>U\<rfloor>" using CASRed2.prems(2) by auto
then show ?case
proof cases
case Null
then show ?thesis using CASRed2
by(force dest: red_hext_incr intro: WTrt_hext_mono WTrtCASNT)
next
case Val
from CASRed2.hyps(1) have hext: "hp s \<unlhd> hp s'" by(auto dest: red_hext_incr)
with Val(9) have "typeof\<^bsub>hp s'\<^esub> v = \<lfloor>U\<rfloor>" by(rule type_of_hext_type_of)
moreover from CASRed2.hyps(2)[OF CASRed2.prems(1) Val(4) CASRed2.prems(3)] Val(5)
obtain T2' where "P,E,hp s' \<turnstile> e' : T2'" "P \<turnstile> T2' \<le> T'" by(auto intro: widen_trans)
moreover from Val(6) hext have "P,E,hp s' \<turnstile> e3 : T3" by(rule WTrt_hext_mono)
ultimately show ?thesis using Val by(auto intro: WTrtCAS)
qed
next
case (CASRed3 e s ta e' s' v D F v')
consider (Null) "v = Null" | (Val) U C T' fm T2 T3 where
"T = Boolean" "class_type_of' U = \<lfloor>C\<rfloor>" "P \<turnstile> C has F:T' (fm) in D" "volatile fm"
"P \<turnstile> T2 \<le> T'" "P,E,hp s \<turnstile> e : T3" "P \<turnstile> T3 \<le> T'"
"typeof\<^bsub>hp s\<^esub> v = \<lfloor>U\<rfloor>" "typeof\<^bsub>hp s\<^esub> v' = \<lfloor>T2\<rfloor>"
using CASRed3.prems(2) by auto
then show ?case
proof cases
case Null
then show ?thesis using CASRed3
by(force dest: red_hext_incr intro: type_of_hext_type_of WTrtCASNT)
next
case Val
from CASRed3.hyps(1) have hext: "hp s \<unlhd> hp s'" by(auto dest: red_hext_incr)
with Val(8,9) have "typeof\<^bsub>hp s'\<^esub> v = \<lfloor>U\<rfloor>" "typeof\<^bsub>hp s'\<^esub> v' = \<lfloor>T2\<rfloor>"
by(blast intro: type_of_hext_type_of)+
moreover from CASRed3.hyps(2)[OF CASRed3.prems(1) Val(6) CASRed3.prems(3)] Val(7)
obtain T3' where "P,E,hp s' \<turnstile> e' : T3'" "P \<turnstile> T3' \<le> T'" by(auto intro: widen_trans)
ultimately show ?thesis using Val by(auto intro: WTrtCAS)
qed
next
case CASNull thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case (CallObj e s ta e' s' M es T E)
have red: "extTA,P,t \<turnstile> \<langle>e,s\<rangle> -ta\<rightarrow> \<langle>e',s'\<rangle>"
and IH: "\<And>E T. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e : T; P,hp s \<turnstile> t \<surd>t\<rbrakk>
\<Longrightarrow> \<exists>U. P,E,hp s' \<turnstile> e' : U \<and> P \<turnstile> U \<le> T"
and conf: "E \<turnstile> s \<surd>" and wt: "P,E,hp s \<turnstile> e\<bullet>M(es) : T"
and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
\<comment> \<open>We distinguish if @{term e} has type @{term NT} or a Class type\<close>
from wt show ?case
proof(rule WTrt_elim_cases)
fix T' C Ts meth D Us
assume wte: "P,E,hp s \<turnstile> e : T'" and icto: "class_type_of' T' = \<lfloor>C\<rfloor>"
and "method": "P \<turnstile> C sees M:Ts\<rightarrow>T = meth in D"
and wtes: "P,E,hp s \<turnstile> es [:] Us" and subs: "P \<turnstile> Us [\<le>] Ts"
obtain U where wte': "P,E,hp s' \<turnstile> e' : U" and UsubC: "P \<turnstile> U \<le> T'"
using IH[OF conf wte tconf] by blast
show ?thesis
proof(cases "U = NT")
case True
moreover have "P,E,hp s' \<turnstile> es [:] Us"
by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]])
ultimately show ?thesis using wte' by(blast intro!:WTrtCallNT)
next
case False
with icto UsubC obtain C'
where icto': "class_type_of' U = \<lfloor>C'\<rfloor>" and "subclass": "P \<turnstile> C' \<preceq>\<^sup>* C"
by(rule widen_is_class_type_of)
obtain Ts' T' meth' D'
where method': "P \<turnstile> C' sees M:Ts'\<rightarrow>T' = meth' in D'"
and subs': "P \<turnstile> Ts [\<le>] Ts'" and sub': "P \<turnstile> T' \<le> T"
using Call_lemma[OF "method" "subclass" wf] by fast
have wtes': "P,E,hp s' \<turnstile> es [:] Us"
by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]])
show ?thesis using wtes' wte' icto' subs method' subs' sub' by(blast intro:widens_trans)
qed
next
fix Ts
assume "P,E,hp s \<turnstile> e:NT"
hence "P,E,hp s' \<turnstile> e' : NT" using IH[OF conf _ tconf] by fastforce
moreover
fix Ts assume wtes: "P,E,hp s \<turnstile> es [:] Ts"
have "P,E,hp s' \<turnstile> es [:] Ts"
by(rule WTrts_hext_mono[OF wtes red_hext_incr[OF red]])
ultimately show ?thesis by(blast intro!:WTrtCallNT)
qed
next
case (CallParams es s ta es' s' v M T E)
have reds: "extTA,P,t \<turnstile> \<langle>es,s\<rangle> [-ta\<rightarrow>] \<langle>es',s'\<rangle>"
and IH: "\<And>Ts E. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> es [:] Ts; P,hp s \<turnstile> t \<surd>t\<rbrakk>
\<Longrightarrow> \<exists>Ts'. P,E,hp s' \<turnstile> es' [:] Ts' \<and> P \<turnstile> Ts' [\<le>] Ts"
and conf: "E \<turnstile> s \<surd>" and wt: "P,E,hp s \<turnstile> Val v\<bullet>M(es) : T"
and tconf: "P,hp s \<turnstile> t \<surd>t" by fact+
from wt show ?case
proof (rule WTrt_elim_cases)
fix U C Ts meth D Us
assume wte: "P,E,hp s \<turnstile> Val v : U" and icto: "class_type_of' U = \<lfloor>C\<rfloor>"
and "P \<turnstile> C sees M:Ts\<rightarrow>T = meth in D"
and wtes: "P,E,hp s \<turnstile> es [:] Us" and "P \<turnstile> Us [\<le>] Ts"
moreover have "P,E,hp s' \<turnstile> Val v : U"
by(rule WTrt_hext_mono[OF wte reds_hext_incr[OF reds]])
moreover obtain Us' where "P,E,hp s' \<turnstile> es' [:] Us'" "P \<turnstile> Us' [\<le>] Us"
using IH[OF conf wtes tconf] by blast
ultimately show ?thesis by(fastforce intro:WTrtCall widens_trans)
next
fix Us
assume null: "P,E,hp s \<turnstile> Val v : NT" and wtes: "P,E,hp s \<turnstile> es [:] Us"
from null have "v = Null" by simp
moreover
obtain Us' where "P,E,hp s' \<turnstile> es' [:] Us' \<and> P \<turnstile> Us' [\<le>] Us"
using IH[OF conf wtes tconf] by blast
ultimately show ?thesis by(fastforce intro:WTrtCallNT)
qed
next
case (RedCall s a U M Ts T pns body D vs T' E)
have hp: "typeof_addr (hp s) a = \<lfloor>U\<rfloor>"
and "method": "P \<turnstile> class_type_of U sees M: Ts\<rightarrow>T = \<lfloor>(pns,body)\<rfloor> in D"
and wt: "P,E,hp s \<turnstile> addr a\<bullet>M(map Val vs) : T'" by fact+
obtain Ts' where wtes: "P,E,hp s \<turnstile> map Val vs [:] Ts'"
and subs: "P \<turnstile> Ts' [\<le>] Ts" and T'isT: "T' = T"
using wt "method" hp wf by(auto 4 3 dest: sees_method_fun)
from wtes subs have length_vs: "length vs = length Ts"
by(auto simp add: WTrts_conv_list_all2 dest!: list_all2_lengthD)
have UsubD: "P \<turnstile> ty_of_htype U \<le> Class (class_type_of U)"
by(cases U)(simp_all add: widen_array_object)
from sees_wf_mdecl[OF wf "method"] obtain T''
where wtabody: "P,[this#pns [\<mapsto>] Class D#Ts] \<turnstile> body :: T''"
and T''subT: "P \<turnstile> T'' \<le> T" and length_pns: "length pns = length Ts"
by(fastforce simp:wf_mdecl_def simp del:map_upds_twist)
from wtabody have "P,Map.empty(this#pns [\<mapsto>] Class D#Ts),hp s \<turnstile> body : T''"
by(rule WT_implies_WTrt)
hence "P,E(this#pns [\<mapsto>] Class D#Ts),hp s \<turnstile> body : T''"
by(rule WTrt_env_mono) simp
hence "P,E,hp s \<turnstile> blocks (this#pns) (Class D#Ts) (Addr a#vs) body : T''"
using wtes subs hp sees_method_decl_above[OF "method"] length_vs length_pns UsubD
by(auto simp add:wt_blocks rel_list_all2_Cons2 intro: widen_trans)
with T''subT T'isT show ?case by blast
next
case (RedCallExternal s a U M Ts T' D vs ta va h' ta' e' s')
from \<open>P,t \<turnstile> \<langle>a\<bullet>M(vs),hp s\<rangle> -ta\<rightarrow>ext \<langle>va,h'\<rangle>\<close> have "hp s \<unlhd> h'" by(rule red_external_hext)
with \<open>P,E,hp s \<turnstile> addr a\<bullet>M(map Val vs) : T\<close>
have "P,E,h' \<turnstile> addr a\<bullet>M(map Val vs) : T" by(rule WTrt_hext_mono)
moreover from \<open>typeof_addr (hp s) a = \<lfloor>U\<rfloor>\<close> \<open>P \<turnstile> class_type_of U sees M: Ts\<rightarrow>T' = Native in D\<close> \<open>P,E,hp s \<turnstile> addr a\<bullet>M(map Val vs) : T\<close>
have "P,hp s \<turnstile> a\<bullet>M(vs) : T'"
by(fastforce simp add: external_WT'_iff dest: sees_method_fun)
ultimately show ?case using RedCallExternal
by(auto 4 3 intro: red_external_conf_extRet[OF wf] intro!: wt_external_call simp add: sconf_def dest: sees_method_fun[where C="class_type_of U"])
next
case RedCallNull thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case (BlockRed e h x V vo ta e' h' x' T T' E)
note IH = \<open>\<And>T E. \<lbrakk>E \<turnstile> (h, x(V := vo)) \<surd>; P,E,hp (h, x(V := vo)) \<turnstile> e : T; P,hp (h, x(V := vo)) \<turnstile> t \<surd>t\<rbrakk>
\<Longrightarrow> \<exists>T'. P,E,hp (h', x') \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T\<close>[simplified]
from \<open>P,E,hp (h, x) \<turnstile> {V:T=vo; e} : T'\<close> have "P,E(V\<mapsto>T),h \<turnstile> e : T'" by(cases vo, auto)
moreover from \<open>E \<turnstile> (h, x) \<surd>\<close> \<open>P,E,hp (h, x) \<turnstile> {V:T=vo; e} : T'\<close>
have "(E(V \<mapsto> T)) \<turnstile> (h, x(V := vo)) \<surd>"
by(cases vo)(simp add: lconf_def sconf_def,auto simp add: sconf_def conf_def intro: lconf_upd2)
ultimately obtain T'' where wt': "P,E(V\<mapsto>T),h' \<turnstile> e' : T''" "P \<turnstile> T'' \<le> T'" using \<open>P,hp (h, x) \<turnstile> t \<surd>t\<close>
by(auto dest: IH)
{ fix v
assume vo: "x' V = \<lfloor>v\<rfloor>"
from \<open>(E(V \<mapsto> T)) \<turnstile> (h, x(V := vo)) \<surd>\<close> \<open>extTA,P,t \<turnstile> \<langle>e,(h, x(V := vo))\<rangle> -ta\<rightarrow> \<langle>e',(h', x')\<rangle>\<close> \<open>P,E(V\<mapsto>T),h \<turnstile> e : T'\<close>
have "P,h' \<turnstile> x' (:\<le>) (E(V \<mapsto> T))" by(auto simp add: sconf_def dest: red_preserves_lconf)
with vo have "\<exists>T'. typeof\<^bsub>h'\<^esub> v = \<lfloor>T'\<rfloor> \<and> P \<turnstile> T' \<le> T" by(fastforce simp add: sconf_def lconf_def conf_def)
then obtain T' where "typeof\<^bsub>h'\<^esub> v = \<lfloor>T'\<rfloor>" "P \<turnstile> T' \<le> T" by blast
hence ?case using wt' vo by(auto) }
moreover
{ assume "x' V = None" with wt' have ?case by(auto) }
ultimately show ?case by blast
next
case RedBlock thus ?case by auto
next
case (SynchronizedRed1 o' s ta o'' s' e T E)
have red: "extTA,P,t \<turnstile> \<langle>o',s\<rangle> -ta\<rightarrow> \<langle>o'',s'\<rangle>" by fact
have IH: "\<And>T E. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> o' : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> o'' : T' \<and> P \<turnstile> T' \<le> T" by fact
have conf: "E \<turnstile> s \<surd>" by fact
have wt: "P,E,hp s \<turnstile> sync(o') e : T" by fact+
thus ?case
proof(rule WTrt_elim_cases)
fix To
assume wto: "P,E,hp s \<turnstile> o' : To"
and refT: "is_refT To"
and wte: "P,E,hp s \<turnstile> e : T"
from IH[OF conf wto \<open>P,hp s \<turnstile> t \<surd>t\<close>] obtain To' where "P,E,hp s' \<turnstile> o'' : To'" and sub: "P \<turnstile> To' \<le> To" by auto
moreover have "P,E,hp s' \<turnstile> e : T"
by(rule WTrt_hext_mono[OF wte red_hext_incr[OF red]])
moreover have "is_refT To'" using refT sub by(auto intro: widen_refT)
ultimately show ?thesis by(auto)
qed
next
case SynchronizedNull thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case LockSynchronized thus ?case by(auto)
next
case (SynchronizedRed2 e s ta e' s' a T E)
have red: "extTA,P,t \<turnstile> \<langle>e,s\<rangle> -ta\<rightarrow> \<langle>e',s'\<rangle>" by fact
have IH: "\<And>T E. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T" by fact
have conf: "E \<turnstile> s \<surd>" by fact
have wt: "P,E,hp s \<turnstile> insync(a) e : T" by fact
thus ?case
proof(rule WTrt_elim_cases)
fix Ta
assume "P,E,hp s \<turnstile> e : T"
and hpa: "typeof_addr (hp s) a = \<lfloor>Ta\<rfloor>"
from \<open>P,E,hp s \<turnstile> e : T\<close> conf \<open>P,hp s \<turnstile> t \<surd>t\<close> obtain T'
where "P,E,hp s' \<turnstile> e' : T'" "P \<turnstile> T' \<le> T" by(blast dest: IH)
moreover from red have hext: "hp s \<unlhd> hp s'" by(auto dest: red_hext_incr)
with hpa have "P,E,hp s' \<turnstile> addr a : ty_of_htype Ta"
by(auto intro: typeof_addr_hext_mono)
ultimately show ?thesis by auto
qed
next
case UnlockSynchronized thus ?case by(auto)
next
case SeqRed thus ?case
apply(auto)
apply(drule WTrt_hext_mono[OF _ red_hext_incr], assumption)
by auto
next
case (CondRed b s ta b' s' e1 e2 T E)
have red: "extTA,P,t \<turnstile> \<langle>b,s\<rangle> -ta\<rightarrow> \<langle>b',s'\<rangle>" by fact
have IH: "\<And>T E. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> b : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> b' : T' \<and> P \<turnstile> T' \<le> T" by fact
have conf: "E \<turnstile> s \<surd>" by fact
have wt: "P,E,hp s \<turnstile> if (b) e1 else e2 : T" by fact
thus ?case
proof(rule WTrt_elim_cases)
fix T1 T2
assume wtb: "P,E,hp s \<turnstile> b : Boolean"
and wte1: "P,E,hp s \<turnstile> e1 : T1"
and wte2: "P,E,hp s \<turnstile> e2 : T2"
and lub: "P \<turnstile> lub(T1, T2) = T"
from IH[OF conf wtb \<open>P,hp s \<turnstile> t \<surd>t\<close>] have "P,E,hp s' \<turnstile> b' : Boolean" by(auto)
moreover have "P,E,hp s' \<turnstile> e1 : T1"
by(rule WTrt_hext_mono[OF wte1 red_hext_incr[OF red]])
moreover have "P,E,hp s' \<turnstile> e2 : T2"
by(rule WTrt_hext_mono[OF wte2 red_hext_incr[OF red]])
ultimately show ?thesis using lub by auto
qed
next
case (ThrowRed e s ta e' s' T E)
have IH: "\<And>T E. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T" by fact
have conf: "E \<turnstile> s \<surd>" by fact
have wt: "P,E,hp s \<turnstile> throw e : T" by fact
then obtain T'
where wte: "P,E,hp s \<turnstile> e : T'"
and nobject: "P \<turnstile> T' \<le> Class Throwable" by auto
from IH[OF conf wte \<open>P,hp s \<turnstile> t \<surd>t\<close>] obtain T''
where wte': "P,E,hp s' \<turnstile> e' : T''"
and PT'T'': "P \<turnstile> T'' \<le> T'" by blast
from nobject PT'T'' have "P \<turnstile> T'' \<le> Class Throwable"
by(auto simp add: widen_Class)(erule notE, rule rtranclp_trans)
hence "P,E,hp s' \<turnstile> throw e' : T" using wte' PT'T''
by -(erule WTrtThrow)
thus ?case by(auto)
next
case RedThrowNull thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case (TryRed e s ta e' s' C V e2 T E)
have red: "extTA,P,t \<turnstile> \<langle>e,s\<rangle> -ta\<rightarrow> \<langle>e',s'\<rangle>" by fact
have IH: "\<And>T E. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T" by fact
have conf: "E \<turnstile> s \<surd>" by fact
have wt: "P,E,hp s \<turnstile> try e catch(C V) e2 : T" by fact
thus ?case
proof(rule WTrt_elim_cases)
fix T1
assume wte: "P,E,hp s \<turnstile> e : T1"
and wte2: "P,E(V \<mapsto> Class C),hp s \<turnstile> e2 : T"
and sub: "P \<turnstile> T1 \<le> T"
from IH[OF conf wte \<open>P,hp s \<turnstile> t \<surd>t\<close>] obtain T1' where "P,E,hp s' \<turnstile> e' : T1'" and "P \<turnstile> T1' \<le> T1" by(auto)
moreover have "P,E(V \<mapsto> Class C),hp s' \<turnstile> e2 : T"
by(rule WTrt_hext_mono[OF wte2 red_hext_incr[OF red]])
ultimately show ?thesis using sub by(auto elim: widen_trans)
qed
next
case RedTryFail thus ?case unfolding sconf_def
by(fastforce simp add: xcpt_subcls_Throwable[OF _ wf])
next
case RedSeq thus ?case by auto
next
case RedCondT thus ?case by(auto dest: is_lub_upper)
next
case RedCondF thus ?case by(auto dest: is_lub_upper)
next
case RedWhile thus ?case by(fastforce)
next
case RedTry thus ?case by auto
next
case RedTryCatch thus ?case by(fastforce)
next
case (ListRed1 e s ta e' s' es Ts E)
note IH = \<open>\<And>T E. \<lbrakk>E \<turnstile> s \<surd>; P,E,hp s \<turnstile> e : T; P,hp s \<turnstile> t \<surd>t\<rbrakk> \<Longrightarrow> \<exists>T'. P,E,hp s' \<turnstile> e' : T' \<and> P \<turnstile> T' \<le> T\<close>
from \<open>P,E,hp s \<turnstile> e # es [:] Ts\<close> obtain T Ts' where "Ts = T # Ts'" "P,E,hp s \<turnstile> e : T" "P,E,hp s \<turnstile> es [:] Ts'" by auto
with IH[of E T] \<open>E \<turnstile> s \<surd>\<close> WTrts_hext_mono[OF \<open>P,E,hp s \<turnstile> es [:] Ts'\<close> red_hext_incr[OF \<open>extTA,P,t \<turnstile> \<langle>e,s\<rangle> -ta\<rightarrow> \<langle>e',s'\<rangle>\<close>]]
show ?case using \<open>P,hp s \<turnstile> t \<surd>t\<close> by(auto simp add: list_all2_Cons2 intro: widens_refl)
next
case ListRed2 thus ?case
by(fastforce dest: hext_typeof_mono[OF reds_hext_incr])
qed(fastforce)+
end
diff --git a/thys/Key_Agreement_Strong_Adversaries/Implem_asymmetric.thy b/thys/Key_Agreement_Strong_Adversaries/Implem_asymmetric.thy
--- a/thys/Key_Agreement_Strong_Adversaries/Implem_asymmetric.thy
+++ b/thys/Key_Agreement_Strong_Adversaries/Implem_asymmetric.thy
@@ -1,511 +1,511 @@
(*******************************************************************************
Project: Refining Authenticated Key Agreement with Strong Adversaries
Module: Implem_asymmetric.thy (Isabelle/HOL 2016-1)
ID: $Id: Implem_asymmetric.thy 132885 2016-12-23 18:41:32Z csprenge $
Author: Joseph Lallemand, INRIA Nancy <joseph.lallemand@loria.fr>
Christoph Sprenger, ETH Zurich <sprenger@inf.ethz.ch>
Asymmetric channel implementation (local interpretation) using
public-key encryption and signatures.
Copyright (c) 2015-2016 Joseph Lallemand and Christoph Sprenger
Licence: LGPL
*******************************************************************************)
section \<open>Asymmetric Implementation of Channel Messages\<close>
theory Implem_asymmetric
imports Implem
begin
(**************************************************************************************************)
subsection \<open>Implementation of channel messages\<close>
(**************************************************************************************************)
fun implem_asym :: "chan \<Rightarrow> msg" where
"implem_asym (Insec A B M) = \<langle>InsecTag, Agent A, Agent B, M\<rangle>"
|"implem_asym (Confid A B M) = Aenc \<langle>Agent A, M\<rangle> (pubK B)"
|"implem_asym (Auth A B M) = Sign \<langle>Agent B, M\<rangle> (priK A)"
|"implem_asym (Secure A B M) = Sign (Aenc \<langle>SecureTag, Agent A, M\<rangle> (pubK B)) (priK A)"
text \<open>
First step: @{locale "basic_implem"}.
Trivial as there are no assumption, this locale just defines some useful abbreviations and valid.
\<close>
interpretation asym: basic_implem implem_asym
done
text \<open>Second step: @{locale "semivalid_implem"}.
Here we prove some basic properties such as injectivity and some properties about the
interaction of sets of implementation messages with @{term analz}; these properties are
proved as separate lemmas as the proofs are more complex.
\<close>
text \<open>Auxiliary: simpler definitions of the \<open>implSets\<close> for the proofs, using the
\<open>msgSet\<close> definitions.
\<close>
abbreviation implInsecSet_aux :: "msg set \<Rightarrow> msg set"
where "implInsecSet_aux G \<equiv> PairSet Tags (PairSet AgentSet (PairSet AgentSet G))"
abbreviation implAuthSet_aux :: "msg set \<Rightarrow> msg set"
where "implAuthSet_aux G \<equiv> SignSet (PairSet AgentSet G) (range priK)"
abbreviation implConfidSet_aux :: "(agent * agent) set \<Rightarrow> msg set \<Rightarrow> msg set"
where "implConfidSet_aux Ag G \<equiv> AencSet (PairSet AgentSet G) (pubK` (Ag `` UNIV))"
abbreviation implSecureSet_aux :: "(agent * agent) set \<Rightarrow> msg set \<Rightarrow> msg set"
where "implSecureSet_aux Ag G \<equiv>
SignSet (AencSet (PairSet Tags (PairSet AgentSet G)) (pubK` (Ag `` UNIV))) (range priK)"
text \<open>These auxiliary definitions are overapproximations.\<close>
lemma implInsecSet_implInsecSet_aux: "asym.implInsecSet G \<subseteq> implInsecSet_aux G"
by auto
lemma implAuthSet_implAuthSet_aux: "asym.implAuthSet G \<subseteq> implAuthSet_aux G"
by auto
lemma implConfidSet_implConfidSet_aux: "asym.implConfidSet Ag G \<subseteq> implConfidSet_aux Ag G"
by (auto, blast)
lemma implSecureSet_implSecureSet_aux: "asym.implSecureSet Ag G \<subseteq> implSecureSet_aux Ag G"
by (auto, blast)
lemmas implSet_implSet_aux =
implInsecSet_implInsecSet_aux implAuthSet_implAuthSet_aux
implConfidSet_implConfidSet_aux implSecureSet_implSecureSet_aux
declare Enc_keys_clean_msgSet_Un [intro]
(**************************************************************************************************)
subsection \<open>Lemmas to pull implementation sets out of @{term analz}\<close>
(**************************************************************************************************)
text \<open>
All these proofs are similar:
\begin{enumerate}
\item prove the lemma for the @{term "implSet_aux"} and with the set added outside of
@{term analz} given explicitly,
\item prove the lemma for the @{term "implSet_aux"} but with payload, and
\item prove the lemma for the @{term "implSet"}.
\end{enumerate}
There are two cases for the confidential and secure messages:
the general case (the payloads stay in @{term analz}) and the case where the key is unknown
(the messages cannot be opened and are completely removed from the @{term analz}).
\<close>
subsubsection \<open>Pull @{term PairAgentSet} out of \<open>analz\<close>\<close>
(**************************************************************************************************)
lemma analz_Un_PairAgentSet:
shows
"analz (PairSet AgentSet G \<union> H) \<subseteq> PairSet AgentSet G \<union> AgentSet \<union> analz (G \<union> H)"
proof -
have "analz (PairSet AgentSet G \<union> H) \<subseteq> PairSet AgentSet G \<union> analz (AgentSet \<union> G \<union> H)"
by (rule analz_Un_PairSet)
also have "... \<subseteq> PairSet AgentSet G \<union> AgentSet \<union> analz (G \<union> H)"
apply (simp only: Un_assoc)
apply (intro Un_mono analz_Un_AgentSet, fast)
done
finally show ?thesis .
qed
subsubsection \<open>Pull @{term implInsecSet} out of @{term analz}\<close>
(**************************************************************************************************)
lemma analz_Un_implInsecSet_aux_aux:
assumes "Enc_keys_clean (G \<union> H)"
shows "analz (implInsecSet_aux G \<union> H) \<subseteq> implInsecSet_aux G \<union> Tags \<union> synth (analz (G \<union> H))"
proof -
have "analz (implInsecSet_aux G \<union> H) \<subseteq>
implInsecSet_aux G \<union> analz (Tags \<union> PairSet AgentSet (PairSet AgentSet G) \<union> H)"
by (rule analz_Un_PairSet)
also have "... \<subseteq> implInsecSet_aux G \<union> Tags \<union> analz (PairSet AgentSet (PairSet AgentSet G) \<union> H)"
using assms
apply -
apply (simp only: Un_assoc, rule Un_mono, fast)
apply (rule analz_Un_Tag, blast)
done
also have "... \<subseteq> implInsecSet_aux G \<union> Tags \<union> PairSet AgentSet (PairSet AgentSet G) \<union> AgentSet
\<union> analz (PairSet AgentSet G \<union> H)"
apply -
apply (simp only: Un_assoc, (rule Un_mono, fast)+)
apply (simp only: Un_assoc [symmetric], rule analz_Un_PairAgentSet)
done
also have
"... \<subseteq> implInsecSet_aux G \<union> Tags \<union> PairSet AgentSet (PairSet AgentSet G) \<union> AgentSet
\<union> PairSet AgentSet G \<union> AgentSet \<union> analz (G \<union> H)"
apply -
apply (simp only: Un_assoc, (rule Un_mono, fast)+)
apply (simp only: Un_assoc [symmetric], rule analz_Un_PairAgentSet)
done
also have "... \<subseteq> implInsecSet_aux G \<union> Tags \<union> synth (analz (G \<union> H))"
apply -
apply (simp only: Un_assoc, (rule Un_mono, fast)+, auto)
done
finally show ?thesis .
qed
lemma analz_Un_implInsecSet_aux:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implInsecSet_aux G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
by (rule subset_trans [OF analz_Un_implInsecSet_aux_aux], auto)
lemma analz_Un_implInsecSet:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (asym.implInsecSet G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
apply (rule subset_trans [of _ "analz (implInsecSet_aux G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
apply (blast dest: analz_Un_implInsecSet_aux)
done
subsection \<open>Pull @{term implConfidSet} out of @{term analz}\<close>
(**************************************************************************************************)
lemma analz_Un_implConfidSet_aux_aux:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implConfidSet_aux Ag G \<union> H) \<subseteq>
implConfidSet_aux Ag G \<union> PairSet AgentSet G \<union>
synth (analz (G \<union> H))"
apply (rule subset_trans [OF analz_Un_AencSet], blast, blast)
apply (simp only: Un_assoc, rule Un_mono, simp)
apply (rule subset_trans [OF analz_Un_PairAgentSet], blast)
done
lemma analz_Un_implConfidSet_aux:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implConfidSet_aux Ag G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
by (rule subset_trans [OF analz_Un_implConfidSet_aux_aux], auto)
lemma analz_Un_implConfidSet:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (asym.implConfidSet Ag G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
apply (rule subset_trans [of _ "analz (implConfidSet_aux Ag G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
using analz_Un_implConfidSet_aux apply blast
done
text \<open>Pull @{term implConfidSet} out of @{term analz}, 2nd case where the agents are honest.\<close>
lemma analz_Un_implConfidSet_aux_aux_2:
"Enc_keys_clean H \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (implConfidSet_aux Ag G \<union> H) \<subseteq> implConfidSet_aux Ag G \<union> synth (analz H)"
apply (rule subset_trans [OF analz_Un_AencSet2], simp)
apply (auto dest:analz_into_parts)
done
lemma analz_Un_implConfidSet_aux_2:
"Enc_keys_clean H \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (implConfidSet_aux Ag G \<union> H) \<subseteq> synth (analz H) \<union> -payload"
by (rule subset_trans [OF analz_Un_implConfidSet_aux_aux_2], auto)
lemma analz_Un_implConfidSet_2:
"Enc_keys_clean H \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (asym.implConfidSet Ag G \<union> H) \<subseteq> synth (analz H) \<union> -payload"
apply (rule subset_trans [of _ "analz (implConfidSet_aux Ag G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
using analz_Un_implConfidSet_aux_2 apply auto
done
subsection \<open>Pull @{term implAuthSet} out of @{term analz}\<close>
(**************************************************************************************************)
lemma analz_Un_implAuthSet_aux_aux:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implAuthSet_aux G \<union> H) \<subseteq> implAuthSet_aux G \<union> synth (analz (G \<union> H))"
apply (rule subset_trans [OF analz_Un_SignSet], blast, blast)
apply (rule Un_mono, blast)
apply (rule subset_trans [OF analz_Un_PairAgentSet], blast)
done
lemma analz_Un_implAuthSet_aux:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implAuthSet_aux G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
by (rule subset_trans [OF analz_Un_implAuthSet_aux_aux], auto)
lemma analz_Un_implAuthSet:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (asym.implAuthSet G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
apply (rule subset_trans [of _ "analz (implAuthSet_aux G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
using analz_Un_implAuthSet_aux apply blast
done
subsection \<open>Pull @{term implSecureSet} out of @{term analz}\<close>
(**************************************************************************************************)
lemma analz_Un_implSecureSet_aux_aux:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implSecureSet_aux Ag G \<union> H) \<subseteq>
implSecureSet_aux Ag G \<union> AencSet (PairSet Tags (PairSet AgentSet G)) (pubK` (Ag`` UNIV)) \<union>
PairSet Tags (PairSet AgentSet G) \<union> Tags \<union> PairSet AgentSet G \<union>
synth (analz (G \<union> H))"
apply (rule subset_trans [OF analz_Un_SignSet], blast, blast)
apply (simp only:Un_assoc, rule Un_mono, simp)
apply (rule subset_trans [OF analz_Un_AencSet], blast, blast)
apply (rule Un_mono, simp)
apply (rule subset_trans [OF analz_Un_PairSet], rule Un_mono, simp, simp only: Un_assoc)
apply (rule subset_trans [OF analz_Un_Tag], blast)
apply (rule Un_mono, simp)
apply (rule subset_trans [OF analz_Un_PairAgentSet], blast)
done
lemma analz_Un_implSecureSet_aux:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implSecureSet_aux Ag G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
by (rule subset_trans [OF analz_Un_implSecureSet_aux_aux], auto)
lemma analz_Un_implSecureSet:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (asym.implSecureSet Ag G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
apply (rule subset_trans [of _ "analz (implSecureSet_aux Ag G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
using analz_Un_implSecureSet_aux apply blast
done
text \<open>
Pull @{term implSecureSet} out of @{term analz}, 2nd case, where the agents are honest.
\<close>
lemma analz_Un_implSecureSet_aux_aux_2:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (implSecureSet_aux Ag G \<union> H) \<subseteq>
implSecureSet_aux Ag G \<union> AencSet (PairSet Tags (PairSet AgentSet G)) (pubK` (Ag`` UNIV)) \<union>
synth (analz H)"
apply (rule subset_trans [OF analz_Un_SignSet], blast, blast)
apply (simp only: Un_assoc, rule Un_mono, simp)
apply (rule subset_trans [OF analz_Un_AencSet2], simp)
apply (auto dest: analz_into_parts)
done
lemma analz_Un_implSecureSet_aux_2:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (implSecureSet_aux Ag G \<union> H) \<subseteq> synth (analz H) \<union> -payload"
by (rule subset_trans [OF analz_Un_implSecureSet_aux_aux_2], auto)
lemma analz_Un_implSecureSet_2:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (asym.implSecureSet Ag G \<union> H) \<subseteq>
synth (analz H) \<union> -payload"
apply (rule subset_trans [of _ "analz (implSecureSet_aux Ag G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
using analz_Un_implSecureSet_aux_2 apply auto
done
declare Enc_keys_clean_msgSet_Un [rule del]
subsection \<open>Locale interpretations\<close>
(**************************************************************************************************)
interpretation asym: semivalid_implem implem_asym
proof (unfold_locales)
fix x A B M x' A' B' M'
show "implem_asym (Chan x A B M) = implem_asym (Chan x' A' B' M') \<longleftrightarrow>
x = x' \<and> A = A' \<and> B = B' \<and> M = M'"
by (cases x, (cases x', auto)+)
next
fix M' M x x' A A' B B'
assume "M' \<in> payload" "implem_asym (Chan x A B M) \<in> parts {implem_asym (Chan x' A' B' M')}"
then show "x = x' \<and> A = A' \<and> B = B' \<and> M = M'"
by (cases "x", auto,(cases x', auto)+)
next
fix I
assume "I \<subseteq> asym.valid"
then show "Enc_keys_clean I"
proof (simp add: Enc_keys_clean_def, intro allI impI)
fix X Y
assume "Enc X Y \<in> parts I"
obtain x A B M where "M \<in> payload" and "Enc X Y \<in> parts {implem_asym (Chan x A B M)}"
using parts_singleton [OF \<open>Enc X Y \<in> parts I\<close>] \<open>I \<subseteq> asym.valid\<close>
by (auto elim!: asym.validE)
then show "Y \<in> range LtK \<or> Y \<in> payload" by (cases x, auto)
qed
next
fix Z
show "composed (implem_asym Z)"
proof (cases Z, simp)
fix x A B M
show "composed (implem_asym (Chan x A B M))" by (cases x, auto)
qed
next
fix x A B M
show "implem_asym (Chan x A B M) \<notin> payload"
by (cases x, auto)
next
fix X K
assume "X \<in> asym.valid"
then obtain x A B M where "M \<in> payload" "X = implem_asym (Chan x A B M)"
by (auto elim: asym.validE)
then show "LtK K \<notin> parts {X}"
by (cases x, auto)
next
fix G H
assume "G \<subseteq> payload" "Enc_keys_clean H"
hence "Enc_keys_clean (G \<union> H)" by (auto intro: Enc_keys_clean_Un)
then show "analz ({implem_asym (Insec A B M) |A B M. M \<in> G} \<union> H) \<subseteq>
synth (analz (G \<union> H)) \<union> - payload"
by (rule analz_Un_implInsecSet)
next
fix G H
assume "G \<subseteq> payload" "Enc_keys_clean H"
hence "Enc_keys_clean (G \<union> H)" by (auto intro: Enc_keys_clean_Un)
then show "analz ({implem_asym (Auth A B M) |A B M. M \<in> G} \<union> H) \<subseteq>
synth (analz (G \<union> H)) \<union> - payload"
by (rule analz_Un_implAuthSet)
next
fix G H Ag
assume "G \<subseteq> payload" "Enc_keys_clean H"
hence "Enc_keys_clean (G \<union> H)" by (auto intro: Enc_keys_clean_Un)
then show "analz ({implem_asym (Confid A B M) |A B M. (A, B) \<in> Ag \<and> M \<in> G} \<union> H) \<subseteq>
synth (analz (G \<union> H)) \<union> - payload"
by (rule analz_Un_implConfidSet)
next
fix G H Ag
assume "G \<subseteq> payload" "Enc_keys_clean H"
hence "Enc_keys_clean (G \<union> H)" by (auto intro: Enc_keys_clean_Un)
then show "analz ({implem_asym (Secure A B M) |A B M. (A, B) \<in> Ag \<and> M \<in> G} \<union> H) \<subseteq>
synth (analz (G \<union> H)) \<union> - payload"
by (rule analz_Un_implSecureSet)
next
fix G H Ag
assume "G \<subseteq> payload" (*unused*)
assume "Enc_keys_clean H"
moreover assume "Ag \<inter> broken (parts H \<inter> range LtK) = {}"
ultimately show "analz ({implem_asym (Confid A B M) |A B M. (A, B) \<in> Ag \<and> M \<in> G} \<union> H) \<subseteq>
synth (analz H) \<union> - payload"
by (rule analz_Un_implConfidSet_2)
next
fix G H Ag
assume "G \<subseteq> payload" "Enc_keys_clean H"
hence "Enc_keys_clean (G \<union> H)" by (auto intro: Enc_keys_clean_Un)
moreover assume "Ag \<inter> broken (parts H \<inter> range LtK) = {}"
ultimately
show "analz ({implem_asym (Secure A B M) |A B M. (A, B) \<in> Ag \<and> M \<in> G} \<union> H) \<subseteq>
synth (analz H) \<union> - payload"
by (rule analz_Un_implSecureSet_2)
qed
text \<open>
Third step: @{locale "valid_implem"}. The lemmas giving conditions on $M$, $A$ and $B$ for
@{prop [display] "implXXX A B M \<in> synth (analz Z)"}.
\<close>
lemma implInsec_synth_analz:
"H \<subseteq> payload \<union> asym.valid \<union> range LtK \<union> Tags \<Longrightarrow>
asym.implInsec A B M \<in> synth (analz H) \<Longrightarrow>
asym.implInsec A B M \<in> I \<or> M \<in> synth (analz H)"
apply (erule synth.cases, auto)
done
lemma implConfid_synth_analz:
"H \<subseteq> payload \<union> asym.valid \<union> range LtK \<union> Tags \<Longrightarrow>
asym.implConfid A B M \<in> synth (analz H) \<Longrightarrow>
asym.implConfid A B M \<in> H \<or> M \<in> synth (analz H)"
apply (erule synth.cases, auto)
apply (frule asym.analz_valid [where x=confid], auto)
done
lemma implAuth_synth_analz:
"H \<subseteq> payload \<union> asym.valid \<union> range LtK \<union> Tags \<Longrightarrow>
asym.implAuth A B M \<in> synth (analz H) \<Longrightarrow>
asym.implAuth A B M \<in> H \<or> (M \<in> synth (analz H) \<and> (A, B) \<in> broken H)"
apply (erule synth.cases, auto)
apply (frule asym.analz_valid [where x=auth], auto)
apply (frule asym.analz_valid [where x=auth], auto)
apply (blast dest: asym.analz_LtKeys)
done
lemma implSecure_synth_analz:
"H \<subseteq> payload \<union> asym.valid \<union> range LtK \<union> Tags \<Longrightarrow>
asym.implSecure A B M \<in> synth (analz H) \<Longrightarrow>
asym.implSecure A B M \<in> H \<or> (M \<in> synth (analz H) \<and> (A, B) \<in> broken H)"
-proof (erule synth.cases, simp_all)
+using [[simproc del: defined_all]] proof (erule synth.cases, simp_all)
fix X
assume H:"H \<subseteq> payload \<union> asym.valid \<union> range LtK \<union> Tags"
assume H':"Sign (Aenc \<langle>SecureTag, Agent A, M\<rangle> (pubK B)) (priK A) = X"
" X \<in> analz H"
hence "asym.implSecure A B M \<in> analz H" by auto
with H have "asym.implSecure A B M \<in> H" by (rule asym.analz_valid)
with H' show "X \<in> H \<or> M \<in> synth (analz H) \<and> (A, B) \<in> broken H"
by auto
next
fix X Y
assume H:"H \<subseteq> payload \<union> asym.valid \<union> range LtK \<union> Tags"
assume H':"Aenc \<langle>SecureTag, Agent A, M\<rangle> (pubK B) = X \<and> priK A = Y"
"X \<in> synth (analz H)" "Y \<in> synth (analz H)"
hence "priK A \<in> analz H" by auto
with H have HAgents:"(A, B) \<in> broken H " by (auto dest: asym.analz_LtKeys)
from H' have "Aenc \<langle>SecureTag, Agent A, M\<rangle> (pubK B) \<in> synth (analz H)" by auto
then have "Aenc \<langle>SecureTag, Agent A, M\<rangle> (pubK B) \<in> analz H \<or>
M \<in> synth (analz H)"
by (rule synth.cases, auto)
then show "Sign X Y \<in> H \<or> M \<in> synth (analz H) \<and> (A, B) \<in> broken H"
proof
assume "M \<in> synth (analz H)"
with HAgents show ?thesis by auto
next
assume "Aenc \<langle>SecureTag, Agent A, M\<rangle> (pubK B) \<in> analz H"
hence "Aenc \<langle>SecureTag, Agent A, M\<rangle> (pubK B) \<in> parts H" by (rule analz_into_parts)
from H obtain Z where
"Z \<in> H" and H'':"Aenc \<langle>SecureTag, Agent A, M\<rangle> (pubK B) \<in> parts {Z}"
using parts_singleton [OF \<open>Aenc \<langle>SecureTag, Agent A, M\<rangle> (pubK B) \<in> parts H\<close>]
by blast
moreover with H have "Z \<in> asym.valid" by (auto dest!: subsetD)
moreover with H'' have "Z = asym.implSecure A B M"
by (auto) (erule asym.valid_cases, auto)
ultimately have "asym.implSecure A B M \<in> H" by auto
with H' show ?thesis by auto
qed
qed
interpretation asym: valid_implem implem_asym
proof (unfold_locales)
fix H A B M
assume "H \<subseteq> payload \<union> asym.valid \<union> range LtK \<union> Tags"
"implem_asym (Insec A B M) \<in> synth (analz H)"
then show "implem_asym (Insec A B M) \<in> H \<or> M \<in> synth (analz H)"
by (rule implInsec_synth_analz)
next
fix H A B M
assume "H \<subseteq> payload \<union> asym.valid \<union> range LtK \<union> Tags"
"implem_asym (Confid A B M) \<in> synth (analz H)"
then show "implem_asym (Confid A B M) \<in> H \<or> M \<in> synth (analz H)"
by (rule implConfid_synth_analz)
next
fix H A B M
assume "H \<subseteq> payload \<union> asym.valid \<union> range LtK \<union> Tags"
"implem_asym (Auth A B M) \<in> synth (analz H)"
then show "implem_asym (Auth A B M) \<in> H \<or>
M \<in> synth (analz H) \<and> (A, B) \<in> broken H"
by (rule implAuth_synth_analz)
next
fix H A B M
assume "H \<subseteq> payload \<union> asym.valid \<union> range LtK \<union> Tags"
"implem_asym (Secure A B M) \<in> synth (analz H)"
then show "implem_asym (Secure A B M) \<in> H \<or>
M \<in> synth (analz H) \<and> (A, B) \<in> broken H"
by (rule implSecure_synth_analz)
qed
end
diff --git a/thys/Key_Agreement_Strong_Adversaries/Implem_symmetric.thy b/thys/Key_Agreement_Strong_Adversaries/Implem_symmetric.thy
--- a/thys/Key_Agreement_Strong_Adversaries/Implem_symmetric.thy
+++ b/thys/Key_Agreement_Strong_Adversaries/Implem_symmetric.thy
@@ -1,615 +1,615 @@
(*******************************************************************************
Project: Refining Authenticated Key Agreement with Strong Adversaries
Module: Implem_symmetric.thy (Isabelle/HOL 2016-1)
ID: $Id: Implem_symmetric.thy 132885 2016-12-23 18:41:32Z csprenge $
Author: Joseph Lallemand, INRIA Nancy <joseph.lallemand@loria.fr>
Christoph Sprenger, ETH Zurich <sprenger@inf.ethz.ch>
Symmetric channel implementation (locale interpretation) using
symmetric encryption and HMACs
Copyright (c) 2015-2016 Joseph Lallemand and Christoph Sprenger
Licence: LGPL
*******************************************************************************)
section \<open>Symmetric Implementation of Channel Messages\<close>
theory Implem_symmetric
imports Implem
begin
(**************************************************************************************************)
subsection \<open>Implementation of channel messages\<close>
(**************************************************************************************************)
fun implem_sym :: "chan \<Rightarrow> msg" where
"implem_sym (Insec A B M) = \<langle>InsecTag, Agent A, Agent B, M\<rangle>"
|"implem_sym (Confid A B M) = Enc \<langle>ConfidTag, M\<rangle> (shrK A B)"
|"implem_sym (Auth A B M) = \<langle>M, hmac \<langle>AuthTag, M\<rangle> (shrK A B)\<rangle>"
|"implem_sym (Secure A B M) = Enc \<langle>SecureTag, M\<rangle> (shrK A B)"
text \<open>
First step: @{locale "basic_implem"}.
Trivial as there are no assumption, this locale just defines some useful abbreviations and valid.
\<close>
interpretation sym: basic_implem implem_sym
done
text \<open>Second step: @{locale "semivalid_implem"}.
Here we prove some basic properties such as injectivity and some properties about the
interaction of sets of implementation messages with @{term analz}; these properties are
proved as separate lemmas as the proofs are more complex.
\<close>
text \<open>Auxiliary: simpler definitions of the \<open>implSets\<close> for the proofs, using the
\<open>msgSet\<close> definitions.
\<close>
abbreviation implInsecSet_aux :: "msg set \<Rightarrow> msg set" where
"implInsecSet_aux G \<equiv> PairSet Tags (PairSet (range Agent) (PairSet (range Agent) G))"
abbreviation implAuthSet_aux :: "msg set \<Rightarrow> msg set" where
"implAuthSet_aux G \<equiv> PairSet G (HashSet (PairSet (PairSet Tags G) (range (case_prod shrK))))"
abbreviation implConfidSet_aux :: "(agent * agent) set \<Rightarrow> msg set \<Rightarrow> msg set" where
"implConfidSet_aux Ag G \<equiv> EncSet (PairSet Tags G) (case_prod shrK`Ag)"
abbreviation implSecureSet_aux :: "(agent * agent) set \<Rightarrow> msg set \<Rightarrow> msg set" where
"implSecureSet_aux Ag G \<equiv> EncSet (PairSet Tags G) (case_prod shrK`Ag)"
text \<open>These auxiliary definitions are overapproximations.\<close>
lemma implInsecSet_implInsecSet_aux: "sym.implInsecSet G \<subseteq> implInsecSet_aux G"
by auto
lemma implAuthSet_implAuthSet_aux: "sym.implAuthSet G \<subseteq> implAuthSet_aux G"
by (auto, auto)
lemma implConfidSet_implConfidSet_aux: "sym.implConfidSet Ag G \<subseteq> implConfidSet_aux Ag G"
by (auto)
lemma implSecureSet_implSecureSet_aux: "sym.implSecureSet Ag G \<subseteq> implSecureSet_aux Ag G"
by (auto)
lemmas implSet_implSet_aux =
implInsecSet_implInsecSet_aux implAuthSet_implAuthSet_aux
implConfidSet_implConfidSet_aux implSecureSet_implSecureSet_aux
declare Enc_keys_clean_msgSet_Un [intro]
(**************************************************************************************************)
subsection \<open>Lemmas to pull implementation sets out of @{term analz}\<close>
(**************************************************************************************************)
text \<open>
All these proofs are similar:
\begin{enumerate}
\item prove the lemma for the @{term "implSet_aux"} and with the set added outside of
@{term analz} given explicitly,
\item prove the lemma for the @{term "implSet_aux"} but with payload, and
\item prove the lemma for the @{term "implSet"}.
\end{enumerate}
There are two cases for the confidential and secure messages:
the general case (the payloads stay in @{term analz}) and the case where the key is unknown
(the messages cannot be opened and are completely removed from the @{term analz}).
\<close>
subsubsection \<open>Pull @{term implInsecSet} out of @{term analz}\<close>
(**************************************************************************************************)
lemma analz_Un_implInsecSet_aux_1:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implInsecSet_aux G \<union> H) \<subseteq>
implInsecSet_aux G \<union> Tags \<union>
PairSet (range Agent) (PairSet (range Agent) G) \<union>
PairSet (range Agent) G \<union>
analz (range Agent \<union> G \<union> (range Agent \<union> H))"
proof -
assume H:"Enc_keys_clean (G \<union> H)"
have "analz (implInsecSet_aux G \<union> H) \<subseteq> implInsecSet_aux G \<union>
analz (Tags \<union> PairSet (range Agent) (PairSet (range Agent) G) \<union> H)"
by (rule analz_Un_PairSet)
also have "... = implInsecSet_aux G \<union>
analz (Tags \<union> (PairSet (range Agent) (PairSet (range Agent) G) \<union> H))"
by (simp only: Un_assoc)
also have "... \<subseteq> implInsecSet_aux G \<union>
(Tags \<union> analz (PairSet (range Agent) (PairSet (range Agent) G) \<union> H))"
by (rule Un_mono, blast, rule analz_Un_Tag, blast intro: H)
also have "... = implInsecSet_aux G \<union> Tags \<union>
analz (PairSet (range Agent) (PairSet (range Agent) G) \<union> H)"
by auto
also have "... \<subseteq> implInsecSet_aux G \<union> Tags \<union> (PairSet (range Agent) (PairSet (range Agent) G) \<union>
analz (range Agent \<union> PairSet (range Agent) G \<union> H))"
by (rule Un_mono, blast, rule analz_Un_PairSet)
also have "... = implInsecSet_aux G \<union> Tags \<union> PairSet (range Agent) (PairSet (range Agent) G) \<union>
analz (PairSet (range Agent) G \<union> (range Agent \<union> H))"
by (auto simp add: Un_assoc Un_commute)
also have "... \<subseteq> implInsecSet_aux G \<union> Tags \<union> PairSet (range Agent) (PairSet (range Agent) G) \<union>
(PairSet (range Agent) G \<union> analz (range Agent \<union> G \<union> (range Agent \<union> H)))"
by (rule Un_mono, blast, rule analz_Un_PairSet)
also have "... = implInsecSet_aux G \<union> Tags \<union> (PairSet (range Agent) (PairSet (range Agent) G) \<union>
PairSet (range Agent) G) \<union> analz (range Agent \<union> G \<union> (range Agent \<union> H))"
by (simp only: Un_assoc Un_commute)
finally show ?thesis by auto
qed
lemma analz_Un_implInsecSet_aux_2:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implInsecSet_aux G \<union> H) \<subseteq>
implInsecSet_aux G \<union> Tags \<union>
synth (analz (G \<union> H))"
proof -
assume H:"Enc_keys_clean (G \<union> H)"
have HH:"PairSet (range Agent) (PairSet (range Agent) G) \<union>
PairSet (range Agent) G \<subseteq> synth (analz (G \<union> H))"
by auto
have HHH:"analz (range Agent \<union> G \<union> (range Agent \<union> H)) \<subseteq> synth (analz (G \<union> H))"
proof -
have "analz (range Agent \<union> G \<union> (range Agent \<union> H)) \<subseteq>
synth (analz (range Agent \<union> G \<union> (range Agent \<union> H)))"
by auto
also have "... = synth (analz (synth (range Agent \<union> G \<union> (range Agent \<union> H))))" by auto
also have "... \<subseteq> synth (analz (synth (G \<union> H)))"
proof (rule synth_analz_mono)
have "range Agent \<union> G \<union> (range Agent \<union> H) \<subseteq> synth (G \<union> H)" by auto
then have "synth (range Agent \<union> G \<union> (range Agent \<union> H)) \<subseteq> synth (synth (G \<union> H))"
by (rule synth_mono)
then show "synth (range Agent \<union> G \<union> (range Agent \<union> H)) \<subseteq> synth (G \<union> H)" by auto
qed
also have "... = synth (analz (G \<union> H))" by auto
finally show ?thesis .
qed
from H have
"analz (implInsecSet_aux G \<union> H) \<subseteq>
implInsecSet_aux G \<union> Tags \<union> PairSet (range Agent) (PairSet (range Agent) G) \<union>
PairSet (range Agent) G \<union> analz (range Agent \<union> G \<union> (range Agent \<union> H))"
by (rule analz_Un_implInsecSet_aux_1)
also have "... = implInsecSet_aux G \<union> Tags \<union>
(PairSet (range Agent) (PairSet (range Agent) G) \<union>
PairSet (range Agent) G) \<union> analz (range Agent \<union> G \<union> (range Agent \<union> H))"
by (simp only: Un_assoc Un_commute)
also have "... \<subseteq> implInsecSet_aux G \<union> Tags \<union> synth (analz (G \<union> H)) \<union>
synth (analz (G \<union> H))"
by ((rule Un_mono)+, auto simp add: HH HHH)
finally show ?thesis by auto
qed
lemma analz_Un_implInsecSet_aux_3:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implInsecSet_aux G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
by (rule subset_trans [OF analz_Un_implInsecSet_aux_2], auto)
lemma analz_Un_implInsecSet:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (sym.implInsecSet G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
apply (rule subset_trans [of _ "analz (implInsecSet_aux G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
using analz_Un_implInsecSet_aux_3 apply blast
done
subsection \<open>Pull @{term implConfidSet} out of @{term analz}\<close>
(**************************************************************************************************)
lemma analz_Un_implConfidSet_aux_1:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implConfidSet_aux Ag G \<union> H) \<subseteq>
implConfidSet_aux Ag G \<union> PairSet Tags G \<union> Tags \<union>
analz (G \<union> H)"
proof -
assume H:"Enc_keys_clean (G \<union> H)"
have "analz (implConfidSet_aux Ag G \<union> H) \<subseteq>
implConfidSet_aux Ag G \<union> analz (PairSet Tags G \<union> H)"
by (rule analz_Un_EncSet, fast, blast intro: H)
also have "... \<subseteq> implConfidSet_aux Ag G \<union> (PairSet Tags G \<union> analz (Tags \<union> G \<union> H))"
by (rule Un_mono, blast, rule analz_Un_PairSet)
also have "... = implConfidSet_aux Ag G \<union> PairSet Tags G \<union> analz (Tags \<union> (G \<union> H))"
by (simp only: Un_assoc)
also have "... \<subseteq> implConfidSet_aux Ag G \<union> PairSet Tags G \<union> (Tags \<union> analz (G \<union> H))"
by (rule Un_mono, blast, rule analz_Un_Tag, blast intro: H)
finally show ?thesis by auto
qed
lemma analz_Un_implConfidSet_aux_2:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implConfidSet_aux Ag G \<union> H) \<subseteq>
implConfidSet_aux Ag G \<union> PairSet Tags G \<union> Tags \<union>
synth (analz (G \<union> H))"
proof -
assume H:"Enc_keys_clean (G \<union> H)"
from H have "analz (implConfidSet_aux Ag G \<union> H) \<subseteq>
implConfidSet_aux Ag G \<union> PairSet Tags G \<union> Tags \<union> analz (G \<union> H)"
by (rule analz_Un_implConfidSet_aux_1)
also have "... \<subseteq> implConfidSet_aux Ag G \<union> PairSet Tags G \<union> Tags \<union> synth (analz (G \<union> H))"
by auto
finally show ?thesis by auto
qed
lemma analz_Un_implConfidSet_aux_3:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implConfidSet_aux Ag G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
by (rule subset_trans [OF analz_Un_implConfidSet_aux_2], auto)
lemma analz_Un_implConfidSet:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (sym.implConfidSet Ag G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
apply (rule subset_trans [of _ "analz (implConfidSet_aux Ag G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
using analz_Un_implConfidSet_aux_3 apply blast
done
text \<open>Pull @{term implConfidSet} out of @{term analz}, 2nd case where the agents are honest.\<close>
lemma analz_Un_implConfidSet_2_aux_1:
"Enc_keys_clean H \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (implConfidSet_aux Ag G \<union> H) \<subseteq> implConfidSet_aux Ag G \<union> synth (analz H)"
apply (rule subset_trans [OF analz_Un_EncSet2], simp)
apply (auto dest:analz_into_parts)
done
lemma analz_Un_implConfidSet_2_aux_3:
"Enc_keys_clean H \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (implConfidSet_aux Ag G \<union> H) \<subseteq> synth (analz H) \<union> -payload"
by (rule subset_trans [OF analz_Un_implConfidSet_2_aux_1], auto)
lemma analz_Un_implConfidSet_2:
"Enc_keys_clean H \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (sym.implConfidSet Ag G \<union> H) \<subseteq> synth (analz H) \<union> -payload"
apply (rule subset_trans [of _ "analz (implConfidSet_aux Ag G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
using analz_Un_implConfidSet_2_aux_3 apply auto
done
subsection \<open>Pull @{term implSecureSet} out of @{term analz}\<close>
(**************************************************************************************************)
lemma analz_Un_implSecureSet_aux_1:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implSecureSet_aux Ag G \<union> H) \<subseteq>
implSecureSet_aux Ag G \<union> PairSet Tags G \<union> Tags \<union>
analz (G \<union> H)"
proof -
assume H:"Enc_keys_clean (G \<union> H)"
have "analz (implSecureSet_aux Ag G \<union> H) \<subseteq>
implSecureSet_aux Ag G \<union> analz (PairSet Tags G \<union> H)"
by (rule analz_Un_EncSet, fast, blast intro: H)
also have "... \<subseteq> implSecureSet_aux Ag G \<union> (PairSet Tags G \<union> analz (Tags \<union> G \<union> H))"
by (rule Un_mono, blast, rule analz_Un_PairSet)
also have "... = implSecureSet_aux Ag G \<union> PairSet Tags G \<union> analz (Tags \<union> (G \<union> H))"
by (simp only: Un_assoc)
also have "... \<subseteq> implSecureSet_aux Ag G \<union> PairSet Tags G \<union> (Tags \<union> analz (G \<union> H))"
by (rule Un_mono, blast, rule analz_Un_Tag, blast intro: H)
finally show ?thesis by auto
qed
lemma analz_Un_implSecureSet_aux_2:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implSecureSet_aux Ag G \<union> H) \<subseteq>
implSecureSet_aux Ag G \<union> PairSet Tags G \<union> Tags \<union>
synth (analz (G \<union> H))"
proof -
assume H:"Enc_keys_clean (G \<union> H)"
from H have "analz (implSecureSet_aux Ag G \<union> H) \<subseteq>
implSecureSet_aux Ag G \<union> PairSet Tags G \<union> Tags \<union> analz (G \<union> H)"
by (rule analz_Un_implSecureSet_aux_1)
also have "... \<subseteq> implSecureSet_aux Ag G \<union> PairSet Tags G \<union> Tags \<union> synth (analz (G \<union> H))"
by auto
finally show ?thesis by auto
qed
lemma analz_Un_implSecureSet_aux_3:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implSecureSet_aux Ag G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
by (rule subset_trans [OF analz_Un_implSecureSet_aux_2], auto)
lemma analz_Un_implSecureSet:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (sym.implSecureSet Ag G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
apply (rule subset_trans [of _ "analz (implSecureSet_aux Ag G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
using analz_Un_implSecureSet_aux_3 apply blast
done
text \<open>
Pull @{term implSecureSet} out of @{term analz}, 2nd case, where the agents are honest.
\<close>
lemma analz_Un_implSecureSet_2_aux_1:
"Enc_keys_clean H \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (implSecureSet_aux Ag G \<union> H) \<subseteq> implSecureSet_aux Ag G \<union> synth (analz H)"
apply (rule subset_trans [OF analz_Un_EncSet2], simp)
apply (auto dest:analz_into_parts)
done
lemma analz_Un_implSecureSet_2_aux_3:
"Enc_keys_clean H \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (implSecureSet_aux Ag G \<union> H) \<subseteq> synth (analz H) \<union> -payload"
by (rule subset_trans [OF analz_Un_implSecureSet_2_aux_1], auto)
lemma analz_Un_implSecureSet_2:
"Enc_keys_clean H \<Longrightarrow>
Ag \<inter> broken (parts H \<inter> range LtK) = {} \<Longrightarrow>
analz (sym.implSecureSet Ag G \<union> H) \<subseteq> synth (analz H) \<union> -payload"
apply (rule subset_trans [of _ "analz (implSecureSet_aux Ag G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
using analz_Un_implSecureSet_2_aux_3 apply auto
done
subsection \<open>Pull @{term implAuthSet} out of @{term analz}\<close>
(**************************************************************************************************)
lemma analz_Un_implAuthSet_aux_1:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implAuthSet_aux G \<union> H) \<subseteq>
implAuthSet_aux G \<union> HashSet (PairSet (PairSet Tags G) (range (case_prod shrK))) \<union>
analz (G \<union> H)"
proof -
assume H:"Enc_keys_clean (G \<union> H)"
have "analz (implAuthSet_aux G \<union> H) \<subseteq> implAuthSet_aux G \<union>
analz (G \<union> HashSet (PairSet (PairSet Tags G) (range (case_prod shrK))) \<union> H)"
by (rule analz_Un_PairSet)
also have "... = implAuthSet_aux G \<union>
analz (HashSet (PairSet (PairSet Tags G) (range (case_prod shrK))) \<union> G \<union> H)"
by (simp only: Un_assoc Un_commute)
also have "... = implAuthSet_aux G \<union>
analz (HashSet (PairSet (PairSet Tags G) (range (case_prod shrK))) \<union> (G \<union> H))"
by (simp only: Un_assoc)
also have
"... \<subseteq> implAuthSet_aux G \<union>
(HashSet (PairSet (PairSet Tags G) (range (case_prod shrK))) \<union>
analz (G \<union> H))"
by (rule Un_mono, blast, rule analz_Un_HashSet, blast intro: H, auto)
also have "... = implAuthSet_aux G \<union>
HashSet (PairSet (PairSet Tags G) (range (case_prod shrK))) \<union>
analz (G \<union> H)"
by auto
finally show ?thesis by auto
qed
lemma analz_Un_implAuthSet_aux_2:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implAuthSet_aux G \<union> H) \<subseteq>
implAuthSet_aux G \<union> HashSet (PairSet (PairSet Tags G) (range (case_prod shrK))) \<union>
synth (analz (G \<union> H))"
proof -
assume H:"Enc_keys_clean (G \<union> H)"
from H have
"analz (implAuthSet_aux G \<union> H) \<subseteq>
implAuthSet_aux G \<union>
HashSet (PairSet (PairSet Tags G) (range (case_prod shrK))) \<union>
analz (G \<union> H)"
by (rule analz_Un_implAuthSet_aux_1)
also have
"... \<subseteq> implAuthSet_aux G \<union>
HashSet (PairSet (PairSet Tags G) (range (case_prod shrK))) \<union>
synth (analz (G \<union> H))"
by auto
finally show ?thesis by auto
qed
lemma analz_Un_implAuthSet_aux_3:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (implAuthSet_aux G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
by (rule subset_trans [OF analz_Un_implAuthSet_aux_2], auto)
lemma analz_Un_implAuthSet:
"Enc_keys_clean (G \<union> H) \<Longrightarrow>
analz (sym.implAuthSet G \<union> H) \<subseteq> synth (analz (G \<union> H)) \<union> -payload"
apply (rule subset_trans [of _ "analz (implAuthSet_aux G \<union> H)" _])
apply (rule analz_mono, rule Un_mono, blast intro!: implSet_implSet_aux, simp)
using analz_Un_implAuthSet_aux_3 apply blast
done
declare Enc_keys_clean_msgSet_Un [rule del]
subsection \<open>Locale interpretations\<close>
(**************************************************************************************************)
interpretation sym: semivalid_implem implem_sym
proof (unfold_locales)
fix x A B M x' A' B' M'
show "implem_sym (Chan x A B M) = implem_sym (Chan x' A' B' M') \<longleftrightarrow>
x = x' \<and> A = A' \<and> B = B' \<and> M = M'"
by (cases x, (cases x', auto)+)
next
fix M' M x x' A A' B B'
assume H: "M' \<in> payload"
then have A1: "\<And>y. y \<in> parts {M'} \<Longrightarrow> y \<in> payload"
and A2: "\<And>y. M' = y \<Longrightarrow> y \<in> payload" by auto
assume "implem_sym (Chan x A B M) \<in> parts {implem_sym (Chan x' A' B' M')}"
then show "x = x' \<and> A = A' \<and> B = B' \<and> M = M'"
by (cases "x", (cases x', auto dest!: A1 A2)+)
next
fix I
assume "I \<subseteq> sym.valid"
then show "Enc_keys_clean I"
proof (simp add: Enc_keys_clean_def, intro allI impI)
fix X Y
assume "Enc X Y \<in> parts I"
obtain x A B M where "M \<in> payload" and "Enc X Y \<in> parts {implem_sym (Chan x A B M)}"
using parts_singleton [OF \<open>Enc X Y \<in> parts I\<close>] \<open>I \<subseteq> sym.valid\<close>
by (auto elim!: sym.validE)
then show "Y \<in> range LtK \<or> Y \<in> payload" by (cases x, auto)
qed
next
fix Z
show "composed (implem_sym Z)"
proof (cases Z, simp)
fix x A B M
show "composed (implem_sym (Chan x A B M))" by (cases x, auto)
qed
next
fix x A B M
show "implem_sym (Chan x A B M) \<notin> payload"
by (cases x, auto)
next
fix X K
assume "X \<in> sym.valid"
then obtain x A B M where "M \<in> payload" "X = implem_sym (Chan x A B M)"
by (auto elim: sym.validE)
then show "LtK K \<notin> parts {X}"
by (cases x, auto)
next
fix G H
assume "G \<subseteq> payload" "Enc_keys_clean H"
hence "Enc_keys_clean (G \<union> H)" by (auto intro: Enc_keys_clean_Un)
then show "analz ({implem_sym (Insec A B M) |A B M. M \<in> G} \<union> H) \<subseteq>
synth (analz (G \<union> H)) \<union> - payload"
by (rule analz_Un_implInsecSet)
next
fix G H
assume "G \<subseteq> payload" "Enc_keys_clean H"
hence "Enc_keys_clean (G \<union> H)" by (auto intro: Enc_keys_clean_Un)
then show "analz ({implem_sym (Auth A B M) |A B M. M \<in> G} \<union> H) \<subseteq>
synth (analz (G \<union> H)) \<union> - payload"
by (rule analz_Un_implAuthSet)
next
fix G H Ag
assume "G \<subseteq> payload" "Enc_keys_clean H"
hence "Enc_keys_clean (G \<union> H)" by (auto intro: Enc_keys_clean_Un)
then show "analz ({implem_sym (Confid A B M) |A B M. (A, B) \<in> Ag \<and> M \<in> G} \<union> H) \<subseteq>
synth (analz (G \<union> H)) \<union> - payload"
by (rule analz_Un_implConfidSet)
next
fix G H Ag
assume "G \<subseteq> payload" "Enc_keys_clean H"
hence "Enc_keys_clean (G \<union> H)" by (auto intro: Enc_keys_clean_Un)
then show "analz ({implem_sym (Secure A B M) |A B M. (A, B) \<in> Ag \<and> M \<in> G} \<union> H) \<subseteq>
synth (analz (G \<union> H)) \<union> - payload"
by (rule analz_Un_implSecureSet)
next
fix G H Ag
assume "Enc_keys_clean H"
hence "Enc_keys_clean H" by auto
moreover assume "Ag \<inter> broken (parts H \<inter> range LtK) = {}"
ultimately show "analz ({implem_sym (Confid A B M) |A B M. (A, B) \<in> Ag \<and> M \<in> G} \<union> H) \<subseteq>
synth (analz H) \<union> - payload"
by (rule analz_Un_implConfidSet_2)
next
fix G H Ag
assume "Enc_keys_clean H"
moreover assume "Ag \<inter> broken (parts H \<inter> range LtK) = {}"
ultimately show "analz ({implem_sym (Secure A B M) |A B M. (A, B) \<in> Ag \<and> M \<in> G} \<union> H) \<subseteq>
synth (analz H) \<union> - payload"
by (rule analz_Un_implSecureSet_2)
qed
text \<open>
Third step: @{locale "valid_implem"}.
The lemmas giving conditions on $M$, $A$ and $B$ for
@{prop "implXXX A B M \<in> synth (analz Z)"}.
\<close>
lemma implInsec_synth_analz:
"H \<subseteq> payload \<union> sym.valid \<union> range LtK \<union> Tags \<Longrightarrow>
sym.implInsec A B M \<in> synth (analz H) \<Longrightarrow>
sym.implInsec A B M \<in> H \<or> M \<in> synth (analz H)"
apply (erule synth.cases, auto)
done
lemma implConfid_synth_analz:
"H \<subseteq> payload \<union> sym.valid \<union> range LtK \<union> Tags \<Longrightarrow>
sym.implConfid A B M \<in> synth (analz H) \<Longrightarrow>
sym.implConfid A B M \<in> H \<or> M \<in> synth (analz H)"
apply (erule synth.cases, auto)
\<comment> \<open>1 subgoal\<close>
apply (frule sym.analz_valid [where x=confid], auto)
done
lemma implAuth_synth_analz:
"H \<subseteq> payload \<union> sym.valid \<union> range LtK \<union> Tags \<Longrightarrow>
sym.implAuth A B M \<in> synth (analz H) \<Longrightarrow>
sym.implAuth A B M \<in> H \<or> (M \<in> synth (analz H) \<and> (A, B) \<in> broken H)"
-proof (erule synth.cases, simp_all)
+using [[simproc del: defined_all]] proof (erule synth.cases, simp_all)
fix X
assume H: "H \<subseteq> payload \<union> sym.valid \<union> range LtK \<union> Tags"
assume H':"\<langle>M, hmac \<langle>AuthTag, M\<rangle> (shrK A B)\<rangle> = X" " X \<in> analz H"
hence "sym.implAuth A B M \<in> analz H" by auto
with H have "sym.implAuth A B M \<in> H" by (rule sym.analz_valid)
with H' show "X \<in> H \<or> M \<in> synth (analz H) \<and> (A, B) \<in> broken H"
by auto
next
fix X Y
assume H:"H \<subseteq> payload \<union> sym.valid \<union> range LtK \<union> Tags"
assume H':"M = X \<and> hmac \<langle>AuthTag, M\<rangle> (shrK A B) = Y" "Y \<in> synth (analz H)"
hence "hmac \<langle>AuthTag, M\<rangle> (shrK A B) \<in> synth (analz H)" by auto
then have "hmac \<langle>AuthTag, M\<rangle> (shrK A B) \<in> analz H \<or>
shrK A B \<in> synth (analz H)"
by (rule synth.cases, auto)
then show "\<langle>X, hmac \<langle>AuthTag, X\<rangle> (shrK A B)\<rangle> \<in> H \<or> (A, B) \<in> broken H"
proof
assume "shrK A B \<in> synth (analz H)"
with H have "(A, B) \<in> broken H" by (auto dest:sym.analz_LtKeys)
then show ?thesis by auto
next
assume "hmac \<langle>AuthTag, M\<rangle> (shrK A B) \<in> analz H"
hence "hmac \<langle>AuthTag, M\<rangle> (shrK A B) \<in> parts H" by (rule analz_into_parts)
with H have "hmac \<langle>AuthTag, M\<rangle> (shrK A B) \<in> parts H"
by (auto dest!:payload_parts elim!:payload_Hash)
from H obtain Z where "Z \<in> H" and H'':"hmac \<langle>AuthTag, M\<rangle> (shrK A B) \<in> parts {Z}"
using parts_singleton [OF \<open>hmac \<langle>AuthTag, M\<rangle> (shrK A B) \<in> parts H\<close>] by blast
moreover with H have "Z \<in> sym.valid" by (auto dest!: subsetD)
moreover with H'' have "Z = sym.implAuth A B M"
by (auto) (erule sym.valid_cases, auto)
ultimately have "sym.implAuth A B M \<in> H" by auto
with H' show ?thesis by auto
qed
qed
lemma implSecure_synth_analz:
"H \<subseteq> payload \<union> sym.valid \<union> range LtK \<union> Tags \<Longrightarrow>
sym.implSecure A B M \<in> synth (analz H) \<Longrightarrow>
sym.implSecure A B M \<in> H \<or> (M \<in> synth (analz H) \<and> (A, B) \<in> broken H)"
apply (erule synth.cases, auto)
(* 1 subgoal*)
apply (frule sym.analz_valid [where x=secure], auto)
apply (frule sym.analz_valid [where x=secure], auto)
apply (auto dest:sym.analz_LtKeys)
done
interpretation sym: valid_implem implem_sym
proof (unfold_locales)
fix H A B M
assume "H \<subseteq> payload \<union> sym.valid \<union> range LtK \<union> Tags"
"implem_sym (Insec A B M) \<in> synth (analz H)"
then show "implem_sym (Insec A B M) \<in> H \<or> M \<in> synth (analz H)"
by (rule implInsec_synth_analz)
next
fix H A B M
assume "H \<subseteq> payload \<union> sym.valid \<union> range LtK \<union> Tags"
"implem_sym (Confid A B M) \<in> synth (analz H)"
then show "implem_sym (Confid A B M) \<in> H \<or> M \<in> synth (analz H)"
by (rule implConfid_synth_analz)
next
fix H A B M
assume "H \<subseteq> payload \<union> sym.valid \<union> range LtK \<union> Tags"
"implem_sym (Auth A B M) \<in> synth (analz H)"
then show "implem_sym (Auth A B M) \<in> H \<or>
M \<in> synth (analz H) \<and> (A, B) \<in> broken H"
by (rule implAuth_synth_analz)
next
fix H A B M
assume "H \<subseteq> payload \<union> sym.valid \<union> range LtK \<union> Tags"
"implem_sym (Secure A B M) \<in> synth (analz H)"
then show "implem_sym (Secure A B M) \<in> H \<or>
M \<in> synth (analz H) \<and> (A, B) \<in> broken H"
by (rule implSecure_synth_analz)
qed
end
diff --git a/thys/Key_Agreement_Strong_Adversaries/pfslvl2.thy b/thys/Key_Agreement_Strong_Adversaries/pfslvl2.thy
--- a/thys/Key_Agreement_Strong_Adversaries/pfslvl2.thy
+++ b/thys/Key_Agreement_Strong_Adversaries/pfslvl2.thy
@@ -1,992 +1,994 @@
(*******************************************************************************
Project: Refining Authenticated Key Agreement with Strong Adversaries
Module: pfslvl2.thy (Isabelle/HOL 2016-1)
ID: $Id: pfslvl2.thy 133183 2017-01-31 13:55:43Z csprenge $
Author: Joseph Lallemand, INRIA Nancy <joseph.lallemand@loria.fr>
Christoph Sprenger, ETH Zurich <sprenger@inf.ethz.ch>
Level-2 protocol using ephemeral asymmetric keys to achieve forward secrecy.
Copyright (c) 2015-2016 Joseph Lallemand and Christoph Sprenger
Licence: LGPL
*******************************************************************************)
section \<open>Key Transport Protocol with PFS (L2)\<close>
theory pfslvl2
imports pfslvl1 Channels
begin
declare domIff [simp, iff del]
(**************************************************************************************************)
subsection \<open>State and Events\<close>
(**************************************************************************************************)
text \<open>initial compromise\<close>
consts
bad_init :: "agent set"
specification (bad_init)
bad_init_spec: "test_owner \<notin> bad_init \<and> test_partner \<notin> bad_init"
by auto
text \<open>level 2 state\<close>
record l2_state =
l1_state +
chan :: "chan set"
bad :: "agent set"
type_synonym l2_obs = "l2_state"
type_synonym
l2_pred = "l2_state set"
type_synonym
l2_trans = "(l2_state \<times> l2_state) set"
text \<open>attacker events\<close>
definition
l2_dy_fake_msg :: "msg \<Rightarrow> l2_trans"
where
"l2_dy_fake_msg m \<equiv> {(s,s').
\<comment> \<open>guards\<close>
m \<in> dy_fake_msg (bad s) (ik s) (chan s) \<and>
\<comment> \<open>actions\<close>
s' = s\<lparr>ik := {m} \<union> ik s\<rparr>
}"
definition
l2_dy_fake_chan :: "chan \<Rightarrow> l2_trans"
where
"l2_dy_fake_chan M \<equiv> {(s,s').
\<comment> \<open>guards\<close>
M \<in> dy_fake_chan (bad s) (ik s) (chan s)\<and>
\<comment> \<open>actions\<close>
s' = s\<lparr>chan := {M} \<union> chan s\<rparr>
}"
text \<open>partnering\<close>
fun
role_comp :: "role_t \<Rightarrow> role_t"
where
"role_comp Init = Resp"
| "role_comp Resp = Init"
definition
matching :: "frame \<Rightarrow> frame \<Rightarrow> bool"
where
"matching sigma sigma' \<equiv> \<forall> x. x \<in> dom sigma \<inter> dom sigma' \<longrightarrow> sigma x = sigma' x"
definition
partner_runs :: "rid_t \<Rightarrow> rid_t \<Rightarrow> bool"
where
"partner_runs R R' \<equiv>
role (guessed_runs R) = role_comp (role (guessed_runs R')) \<and>
owner (guessed_runs R) = partner (guessed_runs R') \<and>
owner (guessed_runs R') = partner (guessed_runs R) \<and>
matching (guessed_frame R) (guessed_frame R')
"
lemma role_comp_inv [simp]:
"role_comp (role_comp x) = x"
by (cases x, auto)
lemma role_comp_inv_eq:
"y = role_comp x \<longleftrightarrow> x = role_comp y"
by (auto elim!: role_comp.elims [OF sym])
definition
partners :: "rid_t set"
where
"partners \<equiv> {R. partner_runs test R}"
lemma test_not_partner [simp]:
"test \<notin> partners"
by (auto simp add: partners_def partner_runs_def, cases "role (guessed_runs test)", auto)
lemma matching_symmetric:
"matching sigma sigma' \<Longrightarrow> matching sigma' sigma"
by (auto simp add: matching_def)
lemma partner_symmetric:
"partner_runs R R' \<Longrightarrow> partner_runs R' R"
by (auto simp add: partner_runs_def matching_symmetric)
lemma partner_unique:
"partner_runs R R'' \<Longrightarrow> partner_runs R R' \<Longrightarrow> R' = R''"
proof -
assume H':"partner_runs R R'"
then have Hm': "matching (guessed_frame R) (guessed_frame R')"
by (auto simp add: partner_runs_def)
assume H'':"partner_runs R R''"
then have Hm'': "matching (guessed_frame R) (guessed_frame R'')"
by (auto simp add: partner_runs_def)
show ?thesis
proof (cases "role (guessed_runs R')")
case Init
with H' partner_symmetric [OF H''] have Hrole:"role (guessed_runs R) = Resp"
"role (guessed_runs R'') = Init"
by (auto simp add: partner_runs_def)
with Init Hm' have "guessed_frame R xpkE = Some (epubKF (R'$kE))"
by (simp add: matching_def)
moreover from Hrole Hm'' have "guessed_frame R xpkE = Some (epubKF (R''$kE))"
by (simp add: matching_def)
ultimately show ?thesis by simp
next
case Resp
with H' partner_symmetric [OF H''] have Hrole:"role (guessed_runs R) = Init"
"role (guessed_runs R'') = Resp"
by (auto simp add: partner_runs_def)
with Resp Hm' have "guessed_frame R xsk = Some (NonceF (R'$sk))"
by (simp add: matching_def)
moreover from Hrole Hm'' have "guessed_frame R xsk = Some (NonceF (R''$sk))"
by (simp add: matching_def)
ultimately show ?thesis by simp
qed
qed
lemma partner_test:
"R \<in> partners \<Longrightarrow> partner_runs R R' \<Longrightarrow> R' = test"
by (auto intro!:partner_unique simp add:partners_def partner_symmetric)
text \<open>compromising events\<close>
definition
l2_lkr_others :: "agent \<Rightarrow> l2_trans"
where
"l2_lkr_others A \<equiv> {(s,s').
\<comment> \<open>guards\<close>
A \<noteq> test_owner \<and>
A \<noteq> test_partner \<and>
\<comment> \<open>actions\<close>
s' = s\<lparr>bad := {A} \<union> bad s\<rparr>
}"
definition
l2_lkr_actor :: "agent \<Rightarrow> l2_trans"
where
"l2_lkr_actor A \<equiv> {(s,s').
\<comment> \<open>guards\<close>
A = test_owner \<and>
A \<noteq> test_partner \<and>
\<comment> \<open>actions\<close>
s' = s\<lparr>bad := {A} \<union> bad s\<rparr>
}"
definition
l2_lkr_after :: "agent \<Rightarrow> l2_trans"
where
"l2_lkr_after A \<equiv> {(s,s').
\<comment> \<open>guards\<close>
test_ended s \<and>
\<comment> \<open>actions\<close>
s' = s\<lparr>bad := {A} \<union> bad s\<rparr>
}"
definition
l2_skr :: "rid_t \<Rightarrow> msg \<Rightarrow> l2_trans"
where
"l2_skr R K \<equiv> {(s,s').
\<comment> \<open>guards\<close>
R \<noteq> test \<and> R \<notin> partners \<and>
in_progress (progress s R) xsk \<and>
guessed_frame R xsk = Some K \<and>
\<comment> \<open>actions\<close>
s' = s\<lparr>ik := {K} \<union> ik s\<rparr>
}"
text \<open>protocol events\<close>
definition
l2_step1 :: "rid_t \<Rightarrow> agent \<Rightarrow> agent \<Rightarrow> l2_trans"
where
"l2_step1 Ra A B \<equiv> {(s, s').
\<comment> \<open>guards:\<close>
Ra \<notin> dom (progress s) \<and>
guessed_runs Ra = \<lparr>role=Init, owner=A, partner=B\<rparr> \<and>
\<comment> \<open>actions:\<close>
s' = s\<lparr>
progress := (progress s)(Ra \<mapsto> {xpkE, xskE}),
chan := {Auth A B (\<langle>Number 0, epubKF (Ra$kE)\<rangle>)} \<union> (chan s)
\<rparr>
}"
definition
l2_step2 :: "rid_t \<Rightarrow> agent \<Rightarrow> agent \<Rightarrow> msg \<Rightarrow> l2_trans"
where
"l2_step2 Rb A B KE \<equiv> {(s, s').
\<comment> \<open>guards:\<close>
guessed_runs Rb = \<lparr>role=Resp, owner=B, partner=A\<rparr> \<and>
Rb \<notin> dom (progress s) \<and>
guessed_frame Rb xpkE = Some KE \<and>
Auth A B \<langle>Number 0, KE\<rangle> \<in> chan s \<and>
\<comment> \<open>actions:\<close>
s' = s\<lparr>
progress := (progress s)(Rb \<mapsto> {xpkE, xsk}),
chan := {Auth B A (Aenc (NonceF (Rb$sk)) KE)} \<union> (chan s),
signals := if can_signal s A B then
addSignal (signals s) (Running A B \<langle>KE, NonceF (Rb$sk)\<rangle>)
else
signals s,
secret := {x. x = NonceF (Rb$sk) \<and> Rb = test} \<union> secret s
\<rparr>
}"
definition
l2_step3 :: "rid_t \<Rightarrow> agent \<Rightarrow> agent \<Rightarrow> msg \<Rightarrow> l2_trans"
where
"l2_step3 Ra A B K \<equiv> {(s, s').
\<comment> \<open>guards:\<close>
guessed_runs Ra = \<lparr>role=Init, owner=A, partner=B\<rparr> \<and>
progress s Ra = Some {xpkE, xskE} \<and>
guessed_frame Ra xsk = Some K \<and>
Auth B A (Aenc K (epubKF (Ra$kE))) \<in> chan s \<and>
\<comment> \<open>actions:\<close>
s' = s\<lparr> progress := (progress s)(Ra \<mapsto> {xpkE, xskE, xsk}),
signals := if can_signal s A B then
addSignal (signals s) (Commit A B \<langle>epubKF (Ra$kE),K\<rangle>)
else
signals s,
secret := {x. x = K \<and> Ra = test} \<union> secret s
\<rparr>
}"
text \<open>specification\<close>
definition
l2_init :: "l2_state set"
where
"l2_init \<equiv> { \<lparr>
ik = {},
secret = {},
progress = Map.empty,
signals = \<lambda>x. 0,
chan = {},
bad = bad_init
\<rparr>}"
definition
l2_trans :: "l2_trans" where
"l2_trans \<equiv> (\<Union>m M KE Rb Ra A B K.
l2_step1 Ra A B \<union>
l2_step2 Rb A B KE \<union>
l2_step3 Ra A B m \<union>
l2_dy_fake_chan M \<union>
l2_dy_fake_msg m \<union>
l2_lkr_others A \<union>
l2_lkr_after A \<union>
l2_skr Ra K \<union>
Id
)"
definition
l2 :: "(l2_state, l2_obs) spec" where
"l2 \<equiv> \<lparr>
init = l2_init,
trans = l2_trans,
obs = id
\<rparr>"
lemmas l2_loc_defs =
l2_step1_def l2_step2_def l2_step3_def
l2_def l2_init_def l2_trans_def
l2_dy_fake_chan_def l2_dy_fake_msg_def
l2_lkr_after_def l2_lkr_others_def l2_skr_def
lemmas l2_defs = l2_loc_defs ik_dy_def
lemmas l2_nostep_defs = l2_def l2_init_def l2_trans_def
lemma l2_obs_id [simp]: "obs l2 = id"
by (simp add: l2_def)
text \<open>Once a run is finished, it stays finished, therefore if the test is not finished at some
point then it was not finished before either\<close>
declare domIff [iff]
lemma l2_run_ended_trans:
"run_ended (progress s R) \<Longrightarrow>
(s, s') \<in> trans l2 \<Longrightarrow>
run_ended (progress s' R)"
apply (auto simp add: l2_nostep_defs)
apply (simp add: l2_defs, fast ?)+
done
declare domIff [iff del]
lemma l2_can_signal_trans:
"can_signal s' A B \<Longrightarrow>
(s, s') \<in> trans l2 \<Longrightarrow>
can_signal s A B"
by (auto simp add: can_signal_def l2_run_ended_trans)
(**************************************************************************************************)
subsection \<open>Invariants\<close>
(**************************************************************************************************)
subsubsection \<open>inv1\<close>
(**************************************************************************************************)
text \<open>If @{term "can_signal s A B"} (i.e., @{term "A"}, @{term "B"} are the test session
agents and the test is not finished), then @{term "A"}, @{term "B"} are honest.\<close>
definition
l2_inv1 :: "l2_state set"
where
"l2_inv1 \<equiv> {s. \<forall>A B.
can_signal s A B \<longrightarrow>
A \<notin> bad s \<and> B \<notin> bad s
}"
lemmas l2_inv1I = l2_inv1_def [THEN setc_def_to_intro, rule_format]
lemmas l2_inv1E [elim] = l2_inv1_def [THEN setc_def_to_elim, rule_format]
lemmas l2_inv1D = l2_inv1_def [THEN setc_def_to_dest, rule_format, rotated 1, simplified]
lemma l2_inv1_init [iff]:
"init l2 \<subseteq> l2_inv1"
by (auto simp add: l2_def l2_init_def l2_inv1_def can_signal_def bad_init_spec)
lemma l2_inv1_trans [iff]:
"{l2_inv1} trans l2 {> l2_inv1}"
proof (auto simp add: PO_hoare_defs intro!: l2_inv1I del: conjI)
fix s' s :: l2_state
fix A B
assume HI:"s \<in> l2_inv1"
assume HT:"(s, s') \<in> trans l2"
assume "can_signal s' A B"
with HT have HS:"can_signal s A B"
by (auto simp add: l2_can_signal_trans)
with HI have "A \<notin> bad s \<and> B \<notin> bad s"
by fast
with HS HT show "A \<notin> bad s' \<and> B \<notin> bad s'"
by (auto simp add: l2_nostep_defs can_signal_def)
(simp_all add: l2_defs)
qed
lemma PO_l2_inv1 [iff]: "reach l2 \<subseteq> l2_inv1"
by (rule inv_rule_basic) (auto)
subsubsection \<open>inv2 (authentication guard)\<close>
(**************************************************************************************************)
text \<open>
If @{term "Auth A B (\<langle>Number 0, KE\<rangle>) \<in> chan s"} and @{term "A"}, @{term "B"} are honest
then the message has indeed been sent by an initiator run (with the right agents etc.)\<close>
definition
l2_inv2 :: "l2_state set"
where
"l2_inv2 \<equiv> {s. \<forall> A B KE.
Auth A B \<langle>Number 0, KE\<rangle> \<in> chan s \<longrightarrow>
A \<notin> bad s \<and> B \<notin> bad s \<longrightarrow>
(\<exists> Ra.
guessed_runs Ra = \<lparr>role=Init, owner=A, partner=B\<rparr> \<and>
in_progress (progress s Ra) xpkE \<and>
KE = epubKF (Ra$kE))
}"
lemmas l2_inv2I = l2_inv2_def [THEN setc_def_to_intro, rule_format]
lemmas l2_inv2E [elim] = l2_inv2_def [THEN setc_def_to_elim, rule_format]
lemmas l2_inv2D = l2_inv2_def [THEN setc_def_to_dest, rule_format, rotated 1, simplified]
lemma l2_inv2_init [iff]:
"init l2 \<subseteq> l2_inv2"
by (auto simp add: l2_def l2_init_def l2_inv2_def)
lemma l2_inv2_trans [iff]:
"{l2_inv2} trans l2 {> l2_inv2}"
apply (auto simp add: PO_hoare_defs l2_nostep_defs intro!: l2_inv2I)
apply (auto simp add: l2_defs dy_fake_chan_def)
apply force+
done
lemma PO_l2_inv2 [iff]: "reach l2 \<subseteq> l2_inv2"
by (rule inv_rule_basic) (auto)
subsubsection \<open>inv3 (authentication guard)\<close>
(**************************************************************************************************)
text \<open>If @{term "Auth B A (Aenc K (epubKF (Ra $ kE))) \<in> chan s"}
and @{term "A"}, @{term "B"} are honest then the message has indeed been sent by a
responder run (etc).\<close>
definition
l2_inv3 :: "l2_state set"
where
"l2_inv3 \<equiv> {s. \<forall> Ra A B K.
Auth B A (Aenc K (epubKF (Ra $ kE))) \<in> chan s \<longrightarrow>
A \<notin> bad s \<and> B \<notin> bad s \<longrightarrow>
(\<exists> Rb.
guessed_runs Rb = \<lparr>role=Resp, owner=B, partner=A\<rparr> \<and>
progress s Rb = Some {xpkE, xsk} \<and>
guessed_frame Rb xpkE = Some (epubKF (Ra$kE))\<and>
K = NonceF (Rb$sk)
)
}"
lemmas l2_inv3I = l2_inv3_def [THEN setc_def_to_intro, rule_format]
lemmas l2_inv3E [elim] = l2_inv3_def [THEN setc_def_to_elim, rule_format]
lemmas l2_inv3D = l2_inv3_def [THEN setc_def_to_dest, rule_format, rotated 1, simplified]
lemma l2_inv3_init [iff]:
"init l2 \<subseteq> l2_inv3"
by (auto simp add: l2_def l2_init_def l2_inv3_def)
lemma l2_inv3_trans [iff]:
"{l2_inv3} trans l2 {> l2_inv3}"
apply (auto simp add: PO_hoare_defs l2_nostep_defs intro!: l2_inv3I)
apply (auto simp add: l2_defs dy_fake_chan_def)
apply (simp_all add: domIff)
apply force+
done
lemma PO_l2_inv3 [iff]: "reach l2 \<subseteq> l2_inv3"
by (rule inv_rule_basic) (auto)
subsubsection \<open>inv4\<close>
(**************************************************************************************************)
text \<open>If the test run is finished and has the session key generated by a run,
then this run is also finished.\<close>
definition
l2_inv4 :: "l2_state set"
where
"l2_inv4 \<equiv> {s. \<forall>Rb.
in_progress (progress s test) xsk \<longrightarrow>
guessed_frame test xsk = Some (NonceF (Rb$sk)) \<longrightarrow>
progress s Rb = Some {xpkE, xsk}
}"
lemmas l2_inv4I = l2_inv4_def [THEN setc_def_to_intro, rule_format]
lemmas l2_inv4E [elim] = l2_inv4_def [THEN setc_def_to_elim, rule_format]
lemmas l2_inv4D = l2_inv4_def [THEN setc_def_to_dest, rule_format, rotated 1, simplified]
lemma l2_inv4_init [iff]:
"init l2 \<subseteq> l2_inv4"
by (auto simp add: l2_def l2_init_def l2_inv4_def)
lemma l2_inv4_trans [iff]:
"{l2_inv4 \<inter> l2_inv3 \<inter> l2_inv1} trans l2 {> l2_inv4}"
apply (auto simp add: PO_hoare_defs l2_nostep_defs intro!: l2_inv4I)
apply (auto simp add: l2_defs dy_fake_chan_def)
apply (auto dest!: l2_inv4D simp add: domIff can_signal_def)
apply (auto dest!: l2_inv3D intro!:l2_inv1D simp add: can_signal_def)
done
lemma PO_l2_inv4 [iff]: "reach l2 \<subseteq> l2_inv4"
by (rule_tac J="l2_inv3 \<inter> l2_inv1" in inv_rule_incr) (auto)
subsubsection \<open>inv5\<close>
(**************************************************************************************************)
text \<open>The only confidential or secure messages on the channel have been put there
by the attacker.\<close>
definition
l2_inv5 :: "l2_state set"
where
"l2_inv5 \<equiv> {s. \<forall>A B M.
(Confid A B M \<in> chan s \<or> Secure A B M \<in> chan s) \<longrightarrow>
M \<in> dy_fake_msg (bad s) (ik s) (chan s)
}"
lemmas l2_inv5I = l2_inv5_def [THEN setc_def_to_intro, rule_format]
lemmas l2_inv5E [elim] = l2_inv5_def [THEN setc_def_to_elim, rule_format]
lemmas l2_inv5D = l2_inv5_def [THEN setc_def_to_dest, rule_format, rotated 1, simplified]
lemma l2_inv5_init [iff]:
"init l2 \<subseteq> l2_inv5"
by (auto simp add: l2_def l2_init_def l2_inv5_def)
lemma l2_inv5_trans [iff]:
"{l2_inv5} trans l2 {> l2_inv5}"
apply (auto simp add: PO_hoare_defs l2_nostep_defs intro!: l2_inv5I)
apply (auto simp add: l2_defs dy_fake_chan_def intro: l2_inv5D dy_fake_msg_monotone)
done
lemma PO_l2_inv5 [iff]: "reach l2 \<subseteq> l2_inv5"
by (rule inv_rule_basic) (auto)
subsubsection \<open>inv6\<close>
(**************************************************************************************************)
text \<open>If an initiator @{term "Ra"} knows a session key @{term "K"}, then the attacker
knows @{term "Aenc K (epubKF (Ra$kE))"}.\<close>
definition
l2_inv6 :: "l2_state set"
where
"l2_inv6 \<equiv> {s. \<forall>Ra K.
role (guessed_runs Ra) = Init \<longrightarrow>
in_progress (progress s Ra) xsk \<longrightarrow>
guessed_frame Ra xsk = Some K \<longrightarrow>
Aenc K (epubKF (Ra$kE)) \<in> extr (bad s) (ik s) (chan s)
}"
lemmas l2_inv6I = l2_inv6_def [THEN setc_def_to_intro, rule_format]
lemmas l2_inv6E [elim] = l2_inv6_def [THEN setc_def_to_elim, rule_format]
lemmas l2_inv6D = l2_inv6_def [THEN setc_def_to_dest, rule_format, rotated 1, simplified]
lemma l2_inv6_init [iff]:
"init l2 \<subseteq> l2_inv6"
by (auto simp add: l2_def l2_init_def l2_inv6_def)
lemma l2_inv6_trans [iff]:
"{l2_inv6} trans l2 {> l2_inv6}"
apply (auto simp add: PO_hoare_defs l2_nostep_defs intro!: l2_inv6I)
apply (auto simp add: l2_defs dy_fake_chan_def dest: l2_inv6D)
done
lemma PO_l2_inv6 [iff]: "reach l2 \<subseteq> l2_inv6"
by (rule inv_rule_basic) (auto)
subsubsection \<open>inv7\<close>
(**************************************************************************************************)
text \<open>Form of the messages in @{term "extr (bad s) (ik s) (chan s)"} =
@{term "synth (analz (generators))"}.\<close>
abbreviation
"generators \<equiv> range epubK \<union>
{Aenc (NonceF (Rb $ sk)) (epubKF (Ra$kE)) |Ra Rb. \<exists> A B.
guessed_runs Ra = \<lparr>role=Init, owner=A, partner=B\<rparr> \<and>
guessed_runs Rb = \<lparr>role=Resp, owner=B, partner=A\<rparr> \<and>
guessed_frame Rb xpkE = Some (epubKF (Ra$kE))} \<union>
{NonceF (R $ sk) |R. R \<noteq> test \<and> R \<notin> partners}"
lemma analz_generators: "analz generators = generators"
by (rule, rule, erule analz.induct, auto)
definition
l2_inv7 :: "l2_state set"
where
"l2_inv7 \<equiv> {s.
extr (bad s) (ik s) (chan s) \<subseteq>
synth (analz (generators))
}"
lemmas l2_inv7I = l2_inv7_def [THEN setc_def_to_intro, rule_format]
lemmas l2_inv7E [elim] = l2_inv7_def [THEN setc_def_to_elim, rule_format]
lemmas l2_inv7D = l2_inv7_def [THEN setc_def_to_dest, rule_format, rotated 1, simplified]
lemma l2_inv7_init [iff]:
"init l2 \<subseteq> l2_inv7"
by (auto simp add: l2_def l2_init_def l2_inv7_def)
lemma l2_inv7_step1:
"{l2_inv7} l2_step1 Ra A B {> l2_inv7}"
apply (auto simp add: PO_hoare_defs l2_defs intro!: l2_inv7I)
apply (auto dest: l2_inv7D [THEN [2] rev_subsetD])+
done
lemma l2_inv7_step2:
"{l2_inv1 \<inter> l2_inv2 \<inter> l2_inv4 \<inter> l2_inv7} l2_step2 Rb A B KE {> l2_inv7}"
proof (auto simp add: PO_hoare_defs l2_nostep_defs intro!: l2_inv7I)
fix s' s :: l2_state
fix x
assume Hx:"x \<in> extr (bad s') (ik s') (chan s')"
assume Hi:"s \<in> l2_inv7"
assume Hi':"s \<in> l2_inv1"
assume Hi'':"s \<in> l2_inv2"
assume Hi''':"s \<in> l2_inv4"
assume Hs:"(s, s') \<in> l2_step2 Rb A B KE"
from Hx Hi Hs show " x \<in> synth (analz (generators))"
proof (auto simp add: l2_defs dest: l2_inv7D [THEN [2] rev_subsetD])
txt \<open>first case: @{term "can_signal s A B"}, which implies that @{term "A"}, @{term "B"} are
honest, and therefore the public key received by @{term "B"} is not from the attacker,
which proves that the
message added to the channel is in @{term "{z. \<exists>x k. z = Aenc x (epubKF k)}"}\<close>
assume Hc:"Auth A B \<langle>Number 0, KE\<rangle> \<in> chan s"
assume HRb:"guessed_runs Rb = \<lparr>role = Resp, owner = B, partner = A\<rparr>"
"guessed_frame Rb xpkE = Some KE"
assume Hcs: "can_signal s A B"
from Hcs Hi' have "A \<notin> bad s \<and> B \<notin> bad s"
by auto
with Hc Hi'' obtain Ra where "KE = epubKF (Ra$kE)"
and "guessed_runs Ra = \<lparr>role=Init, owner=A, partner=B\<rparr>"
by (auto dest: l2_inv2D)
with HRb show "Aenc (NonceF (Rb $ sk)) KE \<in> synth (analz generators)"
by blast
next
txt \<open>second case: @{term "\<not> can_signal s A B"}. We show that @{term "Rb"} is not test and
not a partner:
- @{term "Rb"} is not test because in that case test is not finished and
@{term "A"}, @{term "B"} are the test agents, thus
@{term "can_signal s A B"}
- @{term "Rb"} is not a partner for the same reason
therefore the message added to the channel can be constructed from
@{term "{NonceF (R $ sk) |R. R \<noteq> test \<and> R \<notin> partners}"}\<close>
assume Hc:"Auth A B \<langle>Number 0, KE\<rangle> \<in> chan s"
assume Hcs:"\<not> can_signal s A B"
assume HRb:"Rb \<notin> dom (progress s)"
"guessed_runs Rb = \<lparr>role = Resp, owner = B, partner = A\<rparr>"
from Hcs HRb have "Rb \<noteq> test"
by (auto simp add: can_signal_def domIff)
moreover from HRb Hi''' Hcs have "Rb \<notin> partners"
by (clarify, auto simp add: partners_def partner_runs_def can_signal_def matching_def domIff)
moreover from Hc Hi have "KE \<in> synth (analz (generators))"
by auto
ultimately show "Aenc (NonceF (Rb $ sk)) KE \<in> synth (analz (generators))"
by blast
qed
qed
lemma l2_inv7_step3:
"{l2_inv7} l2_step3 Rb A B K {> l2_inv7}"
by (auto simp add: PO_hoare_defs l2_defs intro!: l2_inv7I dest: l2_inv7D [THEN [2] rev_subsetD])
lemma l2_inv7_dy_fake_msg:
"{l2_inv7} l2_dy_fake_msg M {> l2_inv7}"
by (auto simp add: PO_hoare_defs l2_defs extr_insert_IK_eq
intro!: l2_inv7I
elim!: l2_inv7E dy_fake_msg_extr [THEN [2] rev_subsetD])
lemma l2_inv7_dy_fake_chan:
"{l2_inv7} l2_dy_fake_chan M {> l2_inv7}"
by (auto simp add: PO_hoare_defs l2_defs
intro!: l2_inv7I
dest: dy_fake_chan_extr_insert [THEN [2] rev_subsetD]
elim!: l2_inv7E dy_fake_msg_extr [THEN [2] rev_subsetD])
lemma l2_inv7_lkr_others:
"{l2_inv7 \<inter> l2_inv5} l2_lkr_others A {> l2_inv7}"
apply (auto simp add: PO_hoare_defs l2_defs
intro!: l2_inv7I
dest!: extr_insert_bad [THEN [2] rev_subsetD]
elim!: l2_inv7E l2_inv5E)
apply (auto dest: dy_fake_msg_extr [THEN [2] rev_subsetD])
done
lemma l2_inv7_lkr_after:
"{l2_inv7 \<inter> l2_inv5} l2_lkr_after A {> l2_inv7}"
apply (auto simp add: PO_hoare_defs l2_defs
intro!: l2_inv7I
dest!: extr_insert_bad [THEN [2] rev_subsetD]
elim!: l2_inv7E l2_inv5E)
apply (auto dest: dy_fake_msg_extr [THEN [2] rev_subsetD])
done
lemma l2_inv7_skr:
"{l2_inv7 \<inter> l2_inv6} l2_skr R K {> l2_inv7}"
proof (auto simp add: PO_hoare_defs l2_defs extr_insert_IK_eq intro!: l2_inv7I,
auto elim: l2_inv7D [THEN subsetD])
fix s
assume HRtest: "R \<noteq> test" "R \<notin> partners"
assume Hi: "s \<in> l2_inv7"
assume Hi': "s \<in> l2_inv6"
assume HRsk: "in_progress (progress s R) xsk" "guessed_frame R xsk = Some K"
show "K \<in> synth (analz generators)"
proof (cases "role (guessed_runs R)")
txt \<open>first case: @{term "R"} is the initiator, then @{term "Aenc K epk"}
is in @{term "extr (bad s) (ik s) (chan s)"} (by invariant)
therefore either @{term "K \<in> synth (analz generators)"} which proves the goal
or @{term "Aenc K epk \<in> generators"}, which means that
@{term "K = NonceF (Rb$sk)"} where @{term "R"} and @{term "Rb"} are matching
and since @{term "R"} is not partner or test, neither is
@{term "Rb"}, and therefore @{term "K \<in> synth (analz (generators))"}\<close>
assume HRI: "role (guessed_runs R) = Init"
with HRsk Hi Hi' have "Aenc K (epubKF (R$kE)) \<in> synth (analz generators)"
by (auto dest!: l2_inv7D)
then have "K \<in> synth (analz generators) \<or> Aenc K (epubKF (R$kE)) \<in> generators"
by (rule synth.cases, simp_all, simp add: analz_generators)
with HRsk show ?thesis
proof auto
fix Rb A B
assume HR: "guessed_runs R = \<lparr>role = Init, owner = A, partner = B\<rparr>"
"guessed_frame R xsk = Some (NonceF (Rb $ sk))"
assume HRb: "guessed_runs Rb = \<lparr>role = Resp, owner = B, partner = A\<rparr>"
"guessed_frame Rb xpkE = Some (epubKF (R $ kE))"
from HR HRb have "partner_runs Rb R"
by (auto simp add: partner_runs_def matching_def)
with HRtest have "Rb \<noteq> test \<and> Rb \<notin> partners"
by (auto dest: partner_test, simp add:partners_def)
then show "NonceF (Rb $ sk) \<in> analz generators"
by blast
qed
next
txt \<open>second case: @{term "R"} is the Responder, then @{term "K"} is @{term "R$sk"}
which is in @{term "synth (analz (generators))"}
since @{term "R"} is not test or partner\<close>
assume HRI: "role (guessed_runs R) = Resp"
with HRsk HRtest show ?thesis
by auto
qed
qed
lemmas l2_inv7_trans_aux =
l2_inv7_step1 l2_inv7_step2 l2_inv7_step3
l2_inv7_dy_fake_msg l2_inv7_dy_fake_chan
l2_inv7_lkr_others l2_inv7_lkr_after l2_inv7_skr
lemma l2_inv7_trans [iff]:
"{l2_inv7 \<inter> l2_inv1 \<inter> l2_inv2 \<inter> l2_inv4 \<inter> l2_inv5 \<inter> l2_inv6} trans l2 {> l2_inv7}"
by (auto simp add: l2_nostep_defs intro:l2_inv7_trans_aux)
lemma PO_l2_inv7 [iff]: "reach l2 \<subseteq> l2_inv7"
by (rule_tac J="l2_inv1 \<inter> l2_inv2 \<inter> l2_inv4 \<inter> l2_inv5 \<inter> l2_inv6" in inv_rule_incr) (auto)
lemma l2_inv7_aux:
"NonceF (R$sk) \<in> analz (ik s) \<Longrightarrow> s \<in> l2_inv7 \<Longrightarrow> R \<noteq> test \<and> R \<notin> partners"
proof -
assume H:"s \<in> l2_inv7" and H':"NonceF (R$sk) \<in> analz (ik s)"
then have H'':"NonceF (R$sk) \<in> analz (extr (bad s) (ik s) (chan s))"
by (auto elim: analz_monotone)
from H have "analz (extr (bad s) (ik s) (chan s)) \<subseteq> analz (synth (analz generators))"
by (blast dest: analz_mono intro: l2_inv7D)
with H'' have "NonceF (R$sk) \<in> analz generators"
by auto
then have "NonceF (R$sk) \<in> generators"
by (simp add: analz_generators)
then show ?thesis
by auto
qed
subsubsection \<open>inv8\<close>
text \<open>Form of the secrets = nonces generated by test or partners\<close>
(**************************************************************************************************)
definition
l2_inv8 :: "l2_state set"
where
"l2_inv8 \<equiv> {s.
secret s \<subseteq> {NonceF (R$sk) | R. R = test \<or> R \<in> partners}
}"
lemmas l2_inv8I = l2_inv8_def [THEN setc_def_to_intro, rule_format]
lemmas l2_inv8E [elim] = l2_inv8_def [THEN setc_def_to_elim, rule_format]
lemmas l2_inv8D = l2_inv8_def [THEN setc_def_to_dest, rule_format, rotated 1, simplified]
lemma l2_inv8_init [iff]:
"init l2 \<subseteq> l2_inv8"
by (auto simp add: l2_def l2_init_def l2_inv8_def)
lemma l2_inv8_trans [iff]:
"{l2_inv8 \<inter> l2_inv1 \<inter> l2_inv3} trans l2 {> l2_inv8}"
+supply [[simproc del: defined_all]]
apply (auto simp add: PO_hoare_defs l2_nostep_defs intro!: l2_inv8I)
apply (auto simp add: l2_defs dy_fake_chan_def)
apply (auto simp add: partners_def partner_runs_def matching_def dest!: l2_inv3D)
apply (simp_all add: can_signal_def)
done
lemma PO_l2_inv8 [iff]: "reach l2 \<subseteq> l2_inv8"
by (rule_tac J="l2_inv1 \<inter> l2_inv3" in inv_rule_incr) (auto)
(**************************************************************************************************)
subsection \<open>Refinement\<close>
(**************************************************************************************************)
text \<open>mediator function\<close>
definition
med12s :: "l2_obs \<Rightarrow> l1_obs"
where
"med12s t \<equiv> \<lparr>
ik = ik t,
secret = secret t,
progress = progress t,
signals = signals t
\<rparr>"
text \<open>relation between states\<close>
definition
R12s :: "(l1_state * l2_state) set"
where
"R12s \<equiv> {(s,s').
s = med12s s'
}"
lemmas R12s_defs = R12s_def med12s_def
lemma can_signal_R12 [simp]:
"(s1, s2) \<in> R12s \<Longrightarrow>
can_signal s1 A B \<longleftrightarrow> can_signal s2 A B"
by (auto simp add: can_signal_def R12s_defs)
text \<open>protocol events\<close>
lemma l2_step1_refines_step1:
"{R12s} l1_step1 Ra A B, l2_step1 Ra A B {>R12s}"
by (auto simp add: PO_rhoare_defs R12s_defs l1_step1_def l2_step1_def)
lemma l2_step2_refines_step2:
"{R12s \<inter> UNIV \<times> (l2_inv1 \<inter> l2_inv2 \<inter> l2_inv7)}
l1_step2 Rb A B KE, l2_step2 Rb A B KE
{>R12s}"
apply (auto simp add: PO_rhoare_defs R12s_defs l1_step2_def, simp_all add: l2_step2_def)
apply (auto dest!: l2_inv7_aux l2_inv2D)
done
text \<open>auxiliary lemma needed to prove that the nonce received by the test in step 3
comes from a partner\<close>
lemma l2_step3_partners:
"guessed_runs test = \<lparr>role = Init, owner = A, partner = B\<rparr> \<Longrightarrow>
guessed_frame test xsk = Some (NonceF (Rb$sk)) \<Longrightarrow>
guessed_runs Rb = \<lparr>role = Resp, owner = B, partner = A\<rparr> \<Longrightarrow>
guessed_frame Rb xpkE = Some (epubKF (test $ kE)) \<Longrightarrow>
Rb \<in> partners"
by (auto simp add: partners_def partner_runs_def matching_def)
lemma l2_step3_refines_step3:
"{R12s \<inter> UNIV \<times> (l2_inv1 \<inter> l2_inv3 \<inter> l2_inv7)}
l1_step3 Ra A B K, l2_step3 Ra A B K
{>R12s}"
+supply [[simproc del: defined_all]]
apply (auto simp add: PO_rhoare_defs R12s_defs l1_step3_def, simp_all add: l2_step3_def)
apply (auto dest!: l2_inv3D l2_inv7_aux intro:l2_step3_partners)
apply (auto simp add: can_signal_def)
done
text \<open>attacker events\<close>
lemma l2_dy_fake_chan_refines_skip:
"{R12s} Id, l2_dy_fake_chan M {>R12s}"
by (auto simp add: PO_rhoare_defs R12s_defs l2_defs)
lemma l2_dy_fake_msg_refines_learn:
"{R12s \<inter> UNIV \<times> l2_inv7 \<inter> UNIV \<times> l2_inv8} l1_learn m, l2_dy_fake_msg m {>R12s}"
apply (auto simp add: PO_rhoare_defs R12s_defs l2_loc_defs l1_defs)
apply (drule Fake_insert_dy_fake_msg, erule l2_inv7D)
apply (auto simp add: analz_generators dest!: l2_inv8D)
apply (drule subsetD, simp, drule subsetD, simp, auto) (*change this ?*)
done
text \<open>compromising events\<close>
lemma l2_lkr_others_refines_skip:
"{R12s} Id, l2_lkr_others A {>R12s}"
by (auto simp add: PO_rhoare_defs R12s_defs l2_loc_defs l1_defs)
lemma l2_lkr_after_refines_skip:
"{R12s} Id, l2_lkr_after A {>R12s}"
by (auto simp add: PO_rhoare_defs R12s_defs l2_loc_defs l1_defs)
lemma l2_skr_refines_learn:
"{R12s \<inter> UNIV \<times> l2_inv7 \<inter> UNIV \<times> l2_inv6 \<inter> UNIV \<times> l2_inv8} l1_learn K, l2_skr R K {>R12s}"
proof (auto simp add: PO_rhoare_defs R12s_defs l2_loc_defs l1_defs)
fix s :: l2_state fix x
assume H:"s \<in> l2_inv7" "s \<in> l2_inv6"
"R \<notin> partners" "R \<noteq> test" "run_ended (progress s R)" "guessed_frame R xsk = Some K"
assume Hx:"x \<in> synth (analz (insert K (ik s)))"
assume "x \<in> secret s" "s \<in> l2_inv8"
then obtain R where "x = NonceF (R$sk)" and "R = test \<or> R \<in> partners"
by auto
moreover from H have "s \<lparr>ik := insert K (ik s)\<rparr> \<in> l2_inv7"
by (auto intro: hoare_apply [OF l2_inv7_skr] simp add: l2_defs)
ultimately show False using Hx
by (auto dest: l2_inv7_aux [rotated 1])
qed
text \<open>refinement proof\<close>
lemmas l2_trans_refines_l1_trans =
l2_dy_fake_msg_refines_learn l2_dy_fake_chan_refines_skip
l2_lkr_others_refines_skip l2_lkr_after_refines_skip l2_skr_refines_learn
l2_step1_refines_step1 l2_step2_refines_step2 l2_step3_refines_step3
lemma l2_refines_init_l1 [iff]:
"init l2 \<subseteq> R12s `` (init l1)"
by (auto simp add: R12s_defs l1_defs l2_loc_defs)
lemma l2_refines_trans_l1 [iff]:
"{R12s \<inter> (UNIV \<times> (l2_inv1 \<inter> l2_inv2 \<inter> l2_inv3 \<inter> l2_inv6 \<inter> l2_inv7 \<inter> l2_inv8))} trans l1, trans l2 {> R12s}"
by (auto 0 3 simp add: l1_def l2_def l1_trans_def l2_trans_def
intro!: l2_trans_refines_l1_trans)
lemma PO_obs_consistent_R12s [iff]:
"obs_consistent R12s med12s l1 l2"
by (auto simp add: obs_consistent_def R12s_def med12s_def l2_defs)
lemma l2_refines_l1 [iff]:
"refines
(R12s \<inter>
(reach l1 \<times> (l2_inv1 \<inter> l2_inv2 \<inter> l2_inv3 \<inter> l2_inv4 \<inter> l2_inv5 \<inter> l2_inv6 \<inter> l2_inv7 \<inter> l2_inv8)))
med12s l1 l2"
by (rule Refinement_using_invariants, auto)
lemma l2_implements_l1 [iff]:
"implements med12s l1 l2"
by (rule refinement_soundness) (auto)
subsection \<open>Derived invariants\<close>
(**************************************************************************************************)
text \<open>
We want to prove @{term "l2_secrecy"}:
@{term "dy_fake_msg (bad s) (ik s) (chan s) \<inter> secret s = {}"}
but by refinement we only get @{term "l2_partial_secrecy"}:
@{term "synth (analz (ik s)) \<inter> secret s = {}"}
This is fine, since a message in
@{term "dy_fake_msg (bad s) (ik s) (chan s)"} could be added to @{term "ik s"},
and @{term "l2_partial_secrecy"} would still hold for this new state.
\<close>
definition
l2_partial_secrecy :: "('a l2_state_scheme) set"
where
"l2_partial_secrecy \<equiv> {s. synth (analz (ik s)) \<inter> secret s = {}}"
lemma l2_obs_partial_secrecy [iff]: "oreach l2 \<subseteq> l2_partial_secrecy"
apply (rule external_invariant_translation
[OF l1_obs_secrecy _ l2_implements_l1])
apply (auto simp add: med12s_def s0_secrecy_def l2_partial_secrecy_def)
done
lemma l2_oreach_dy_fake_msg:
"s \<in> oreach l2 \<Longrightarrow> x \<in> dy_fake_msg (bad s) (ik s) (chan s) \<Longrightarrow> s \<lparr>ik := insert x (ik s)\<rparr> \<in> oreach l2"
apply (auto simp add: oreach_def, rule, simp_all, simp add: l2_def l2_trans_def l2_dy_fake_msg_def)
apply blast
done
definition
l2_secrecy :: "('a l2_state_scheme) set"
where
"l2_secrecy \<equiv> {s. dy_fake_msg (bad s) (ik s) (chan s) \<inter> secret s = {}}"
lemma l2_obs_secrecy [iff]: "oreach l2 \<subseteq> l2_secrecy"
apply (auto simp add:l2_secrecy_def)
apply (drule l2_oreach_dy_fake_msg, simp_all)
apply (drule l2_obs_partial_secrecy [THEN [2] rev_subsetD], simp add: l2_partial_secrecy_def)
apply blast
done
lemma l2_secrecy [iff]: "reach l2 \<subseteq> l2_secrecy"
by (rule external_to_internal_invariant [OF l2_obs_secrecy], auto)
abbreviation "l2_iagreement \<equiv> l1_iagreement"
lemma l2_obs_iagreement [iff]: "oreach l2 \<subseteq> l2_iagreement"
apply (rule external_invariant_translation
[OF l1_obs_iagreement _ l2_implements_l1])
apply (auto simp add: med12s_def l1_iagreement_def)
done
lemma l2_iagreement [iff]: "reach l2 \<subseteq> l2_iagreement"
by (rule external_to_internal_invariant [OF l2_obs_iagreement], auto)
end
diff --git a/thys/Koenigsberg_Friendship/KoenigsbergBridge.thy b/thys/Koenigsberg_Friendship/KoenigsbergBridge.thy
--- a/thys/Koenigsberg_Friendship/KoenigsbergBridge.thy
+++ b/thys/Koenigsberg_Friendship/KoenigsbergBridge.thy
@@ -1,990 +1,986 @@
(*
Title:KoenigsbergBridge.thy
Author:Wenda Li
*)
theory KoenigsbergBridge imports MoreGraph
begin
section\<open>Definition of Eulerian trails and circuits\<close>
definition (in valid_unMultigraph) is_Eulerian_trail:: "'v\<Rightarrow>('v,'w) path\<Rightarrow>'v\<Rightarrow> bool" where
"is_Eulerian_trail v ps v'\<equiv> is_trail v ps v' \<and> edges (rem_unPath ps G) = {}"
definition (in valid_unMultigraph) is_Eulerian_circuit:: "'v \<Rightarrow> ('v,'w) path \<Rightarrow> 'v \<Rightarrow> bool" where
"is_Eulerian_circuit v ps v'\<equiv> (v=v') \<and> (is_Eulerian_trail v ps v')"
section\<open>Necessary conditions for Eulerian trails and circuits\<close>
lemma (in valid_unMultigraph) euclerian_rev:
"is_Eulerian_trail v' (rev_path ps) v=is_Eulerian_trail v ps v' "
proof -
have "is_trail v' (rev_path ps) v=is_trail v ps v'"
by (metis is_trail_rev)
moreover have "edges (rem_unPath (rev_path ps) G)=edges (rem_unPath ps G)"
by (metis rem_unPath_graph)
ultimately show ?thesis unfolding is_Eulerian_trail_def by auto
qed
(*Necessary conditions for Eulerian circuits*)
theorem (in valid_unMultigraph) euclerian_cycle_ex:
assumes "is_Eulerian_circuit v ps v'" "finite V" "finite E"
shows "\<forall>v\<in>V. even (degree v G)"
proof -
obtain v ps v' where cycle:"is_Eulerian_circuit v ps v'" using assms by auto
hence "edges (rem_unPath ps G) = {}"
unfolding is_Eulerian_circuit_def is_Eulerian_trail_def
by simp
moreover have "nodes (rem_unPath ps G)=nodes G" by auto
ultimately have "rem_unPath ps G = G \<lparr>edges:={}\<rparr>" by auto
hence "num_of_odd_nodes (rem_unPath ps G) = 0" by (metis assms(2) odd_nodes_no_edge)
moreover have "v=v'"
by (metis \<open>is_Eulerian_circuit v ps v'\<close> is_Eulerian_circuit_def)
hence "num_of_odd_nodes (rem_unPath ps G)=num_of_odd_nodes G"
by (metis assms(2) assms(3) cycle is_Eulerian_circuit_def
is_Eulerian_trail_def rem_UnPath_cycle)
ultimately have "num_of_odd_nodes G=0" by auto
moreover have "finite(odd_nodes_set G)"
using \<open>finite V\<close> unfolding odd_nodes_set_def by auto
ultimately have "odd_nodes_set G = {}" unfolding num_of_odd_nodes_def by auto
thus ?thesis unfolding odd_nodes_set_def by auto
qed
(*Necessary conditions for Eulerian trails*)
theorem (in valid_unMultigraph) euclerian_path_ex:
assumes "is_Eulerian_trail v ps v'" "finite V" "finite E"
shows "(\<forall>v\<in>V. even (degree v G)) \<or> (num_of_odd_nodes G =2)"
proof -
obtain v ps v' where path:"is_Eulerian_trail v ps v'" using assms by auto
hence "edges (rem_unPath ps G) = {}"
unfolding is_Eulerian_trail_def
by simp
moreover have "nodes (rem_unPath ps G)=nodes G" by auto
ultimately have "rem_unPath ps G = G \<lparr>edges:={}\<rparr>" by auto
hence odd_nodes: "num_of_odd_nodes (rem_unPath ps G) = 0"
by (metis assms(2) odd_nodes_no_edge)
have "v\<noteq>v' \<Longrightarrow> ?thesis"
proof (cases "even(degree v' G)")
case True
assume "v\<noteq>v'"
have "is_trail v ps v'" by (metis is_Eulerian_trail_def path)
hence "num_of_odd_nodes (rem_unPath ps G) = num_of_odd_nodes G
+ (if even (degree v G) then 2 else 0)"
using rem_UnPath_even True \<open>finite V\<close> \<open>finite E\<close> \<open>v\<noteq>v'\<close> by auto
hence "num_of_odd_nodes G + (if even (degree v G) then 2 else 0)=0"
using odd_nodes by auto
hence "num_of_odd_nodes G = 0" by auto
moreover have "finite(odd_nodes_set G)"
using \<open>finite V\<close> unfolding odd_nodes_set_def by auto
ultimately have "odd_nodes_set G = {}" unfolding num_of_odd_nodes_def by auto
thus ?thesis unfolding odd_nodes_set_def by auto
next
case False
assume "v\<noteq>v'"
have "is_trail v ps v'" by (metis is_Eulerian_trail_def path)
hence "num_of_odd_nodes (rem_unPath ps G) = num_of_odd_nodes G
+ (if odd (degree v G) then -2 else 0)"
using rem_UnPath_odd False \<open>finite V\<close> \<open>finite E\<close> \<open>v\<noteq>v'\<close> by auto
hence odd_nodes_if: "num_of_odd_nodes G + (if odd (degree v G) then -2 else 0)=0"
using odd_nodes by auto
have "odd (degree v G) \<Longrightarrow> ?thesis"
proof -
assume "odd (degree v G)"
hence "num_of_odd_nodes G = 2" using odd_nodes_if by auto
thus ?thesis by simp
qed
moreover have "even(degree v G) \<Longrightarrow> ?thesis"
proof -
assume "even (degree v G)"
hence "num_of_odd_nodes G = 0" using odd_nodes_if by auto
moreover have "finite(odd_nodes_set G)"
using \<open>finite V\<close> unfolding odd_nodes_set_def by auto
ultimately have "odd_nodes_set G = {}" unfolding num_of_odd_nodes_def by auto
thus ?thesis unfolding odd_nodes_set_def by auto
qed
ultimately show ?thesis by auto
qed
moreover have "v=v'\<Longrightarrow> ?thesis"
by (metis assms(2) assms(3) euclerian_cycle_ex is_Eulerian_circuit_def path)
ultimately show ?thesis by auto
qed
section\<open>Specific case of the Konigsberg Bridge Problem\<close>
(*to denote the four landmasses*)
datatype kon_node = a | b | c | d
(*to denote the seven bridges*)
datatype kon_bridge = ab1 | ab2 | ac1 | ac2 | ad1 | bd1 | cd1
definition kon_graph :: "(kon_node,kon_bridge) graph" where
"kon_graph\<equiv>\<lparr>nodes={a,b,c,d},
edges={(a,ab1,b), (b,ab1,a),
(a,ab2,b), (b,ab2,a),
(a,ac1,c), (c,ac1,a),
(a,ac2,c), (c,ac2,a),
(a,ad1,d), (d,ad1,a),
(b,bd1,d), (d,bd1,b),
(c,cd1,d), (d,cd1,c)} \<rparr>"
instantiation kon_node :: enum
begin
definition [simp]: "enum_class.enum =[a,b,c,d]"
definition [simp]: "enum_class.enum_all P \<longleftrightarrow> P a \<and> P b \<and> P c \<and> P d"
definition [simp]:"enum_class.enum_ex P \<longleftrightarrow> P a \<or> P b \<or> P c \<or> P d"
instance proof qed (auto,(case_tac x,auto)+)
end
instantiation kon_bridge :: enum
begin
definition [simp]:"enum_class.enum =[ab1,ab2,ac1,ac2,ad1,cd1,bd1]"
definition [simp]:"enum_class.enum_all P \<longleftrightarrow> P ab1 \<and> P ab2 \<and> P ac1 \<and> P ac2 \<and> P ad1 \<and> P bd1
\<and> P cd1"
definition [simp]:"enum_class.enum_ex P \<longleftrightarrow> P ab1 \<or> P ab2 \<or> P ac1 \<or> P ac2 \<or> P ad1 \<or> P bd1
\<or> P cd1"
instance proof qed (auto,(case_tac x,auto)+)
end
interpretation kon_graph: valid_unMultigraph kon_graph
proof (unfold_locales)
show "fst ` edges kon_graph \<subseteq> nodes kon_graph" by eval
next
show "snd ` snd ` edges kon_graph \<subseteq> nodes kon_graph" by eval
next
have " \<forall>v w u'. ((v, w, u') \<in> edges kon_graph) = ((u', w, v) \<in> edges kon_graph)"
by eval
thus "\<And>v w u'. ((v, w, u') \<in> edges kon_graph) = ((u', w, v) \<in> edges kon_graph)" by simp
next
have "\<forall>v w. (v, w, v) \<notin> edges kon_graph" by eval
thus "\<And>v w. (v, w, v) \<notin> edges kon_graph" by simp
qed
(*The specific case of the Konigsberg Bridge Problem does not have a solution*)
theorem "\<not>kon_graph.is_Eulerian_trail v1 p v2"
proof
assume "kon_graph.is_Eulerian_trail v1 p v2"
moreover have "finite (nodes kon_graph)" by (metis finite_code)
moreover have "finite (edges kon_graph)" by (metis finite_code)
ultimately have contra:
"(\<forall>v\<in>nodes kon_graph. even (degree v kon_graph)) \<or>(num_of_odd_nodes kon_graph =2)"
by (metis kon_graph.euclerian_path_ex)
have "odd(degree a kon_graph)" by eval
moreover have "odd(degree b kon_graph)" by eval
moreover have "odd(degree c kon_graph)" by eval
moreover have "odd(degree d kon_graph)" by eval
ultimately have "\<not>(num_of_odd_nodes kon_graph =2)" by eval
moreover have "\<not>(\<forall>v\<in>nodes kon_graph. even (degree v kon_graph))" by eval
ultimately show False using contra by auto
qed
section\<open>Sufficient conditions for Eulerian trails and circuits\<close>
lemma (in valid_unMultigraph) eulerian_cons:
assumes
"valid_unMultigraph.is_Eulerian_trail (del_unEdge v0 w v1 G) v1 ps v2"
"(v0,w,v1)\<in> E"
shows "is_Eulerian_trail v0 ((v0,w,v1)#ps) v2"
proof -
have valid:"valid_unMultigraph (del_unEdge v0 w v1 G)"
using valid_unMultigraph_axioms by auto
hence distinct:"valid_unMultigraph.is_trail (del_unEdge v0 w v1 G) v1 ps v2"
using assms unfolding valid_unMultigraph.is_Eulerian_trail_def[OF valid]
by auto
hence "set ps \<subseteq> edges (del_unEdge v0 w v1 G)"
using valid_unMultigraph.path_in_edges[OF valid] by auto
moreover have "(v0,w,v1)\<notin>edges (del_unEdge v0 w v1 G)"
unfolding del_unEdge_def by auto
moreover have "(v1,w,v0)\<notin>edges (del_unEdge v0 w v1 G)"
unfolding del_unEdge_def by auto
ultimately have "(v0,w,v1)\<notin>set ps" "(v1,w,v0)\<notin>set ps" by auto
moreover have "is_trail v1 ps v2"
using distinct_path_intro[OF distinct] .
ultimately have "is_trail v0 ((v0,w,v1)#ps) v2"
using \<open>(v0,w,v1)\<in> E\<close> by auto
moreover have "edges (rem_unPath ps (del_unEdge v0 w v1 G)) ={}"
using assms unfolding valid_unMultigraph.is_Eulerian_trail_def[OF valid]
by auto
hence "edges (rem_unPath ((v0,w,v1)#ps) G)={}"
by (metis rem_unPath.simps(2))
ultimately show ?thesis unfolding is_Eulerian_trail_def by auto
qed
lemma (in valid_unMultigraph) eulerian_cons':
assumes
"valid_unMultigraph.is_Eulerian_trail (del_unEdge v2 w v3 G) v1 ps v2"
"(v2,w,v3)\<in> E"
shows "is_Eulerian_trail v1 (ps@[(v2,w,v3)]) v3"
proof -
have valid:"valid_unMultigraph (del_unEdge v3 w v2 G)"
using valid_unMultigraph_axioms del_unEdge_valid by auto
have "del_unEdge v2 w v3 G=del_unEdge v3 w v2 G"
by (metis delete_edge_sym)
hence "valid_unMultigraph.is_Eulerian_trail (del_unEdge v3 w v2 G) v2
(rev_path ps) v1" using assms valid_unMultigraph.euclerian_rev[OF valid]
by auto
hence "is_Eulerian_trail v3 ((v3,w,v2)#(rev_path ps)) v1"
using eulerian_cons by (metis assms(2) corres)
hence "is_Eulerian_trail v1 (rev_path((v3,w,v2)#(rev_path ps))) v3"
using euclerian_rev by auto
moreover have "rev_path((v3,w,v2)#(rev_path ps)) = rev_path(rev_path ps)@[(v2,w,v3)]"
unfolding rev_path_def by auto
hence "rev_path((v3,w,v2)#(rev_path ps))=ps@[(v2,w,v3)]" by auto
ultimately show ?thesis by auto
qed
lemma eulerian_split:
assumes "nodes G1 \<inter> nodes G2 = {}" "edges G1 \<inter> edges G2={}"
"valid_unMultigraph G1" "valid_unMultigraph G2"
"valid_unMultigraph.is_Eulerian_trail G1 v1 ps1 v1'"
"valid_unMultigraph.is_Eulerian_trail G2 v2 ps2 v2'"
shows "valid_unMultigraph.is_Eulerian_trail \<lparr>nodes=nodes G1 \<union> nodes G2,
edges=edges G1 \<union> edges G2 \<union> {(v1',w,v2),(v2,w,v1')}\<rparr> v1 (ps1@(v1',w,v2)#ps2) v2'"
proof -
have "valid_graph G1" using \<open>valid_unMultigraph G1\<close> valid_unMultigraph_def by auto
have "valid_graph G2" using \<open>valid_unMultigraph G2\<close> valid_unMultigraph_def by auto
obtain G where G:"G=\<lparr>nodes=nodes G1 \<union> nodes G2, edges=edges G1 \<union> edges G2
\<union> {(v1',w,v2),(v2,w,v1')}\<rparr>"
by metis
have "v1'\<in>nodes G1"
by (metis (full_types) \<open>valid_graph G1\<close> assms(3) assms(5) valid_graph.is_path_memb
valid_unMultigraph.is_trail_intro valid_unMultigraph.is_Eulerian_trail_def)
moreover have "v2\<in>nodes G2"
by (metis (full_types) \<open>valid_graph G2\<close> assms(4) assms(6) valid_graph.is_path_memb
valid_unMultigraph.is_trail_intro valid_unMultigraph.is_Eulerian_trail_def)
+ moreover have \<open>ba \<in> nodes G1\<close> if \<open>(aa, ab, ba) \<in> edges G1\<close>
+ for aa ab ba
+ using that
+ by (meson \<open>valid_graph G1\<close> valid_graph.E_validD(2))
ultimately have "valid_unMultigraph \<lparr>nodes=nodes G1 \<union> nodes G2, edges=edges G1 \<union> edges G2 \<union>
{(v1',w,v2),(v2,w,v1')}\<rparr>"
using
valid_unMultigraph.corres[OF \<open>valid_unMultigraph G1\<close>]
valid_unMultigraph.no_id[OF \<open>valid_unMultigraph G1\<close>]
valid_unMultigraph.corres[OF \<open>valid_unMultigraph G2\<close>]
valid_unMultigraph.no_id[OF \<open>valid_unMultigraph G2\<close>]
valid_graph.E_validD[OF \<open>valid_graph G1\<close>]
valid_graph.E_validD[OF \<open>valid_graph G2\<close>]
\<open>nodes G1 \<inter> nodes G2 = {}\<close>
- proof (unfold_locales,auto)
- fix aa ab ba
- assume "(aa, ab, ba) \<in> edges G1"
- thus "ba \<in> nodes G1" by (metis \<open>\<And>v' v e. (v, e, v') \<in> edges G1 \<Longrightarrow> v' \<in> nodes G1\<close>)
- next
- fix aa ab ba
- assume "ba \<notin> nodes G2" "(aa, ab, ba) \<in> edges G2"
- thus "ba \<in> nodes G1" by (metis \<open>valid_graph G2\<close> valid_graph.E_validD(2))
- qed
+ by unfold_locales auto
hence valid: "valid_unMultigraph G" using G by auto
hence valid':"valid_graph G" using valid_unMultigraph_def by auto
moreover have "valid_unMultigraph.is_trail G v1 (ps1@((v1',w,v2)#ps2)) v2'"
proof -
have ps1_G:"valid_unMultigraph.is_trail G v1 ps1 v1'"
proof -
have "valid_unMultigraph.is_trail G1 v1 ps1 v1'" using assms
by (metis valid_unMultigraph.is_Eulerian_trail_def)
moreover have "edges G1 \<subseteq> edges G" by (metis G UnI1 Un_assoc select_convs(2) subrelI)
moreover have "nodes G1 \<subseteq> nodes G" by (metis G inf_sup_absorb le_iff_inf select_convs(1))
ultimately show ?thesis
using distinct_path_subset[of G1 G,OF \<open>valid_unMultigraph G1\<close> valid] by auto
qed
have ps2_G:"valid_unMultigraph.is_trail G v2 ps2 v2'"
proof -
have "valid_unMultigraph.is_trail G2 v2 ps2 v2'" using assms
by (metis valid_unMultigraph.is_Eulerian_trail_def)
moreover have "edges G2 \<subseteq> edges G" by (metis G inf_sup_ord(3) le_supE select_convs(2))
moreover have "nodes G2 \<subseteq> nodes G" by (metis G inf_sup_ord(4) select_convs(1))
ultimately show ?thesis
using distinct_path_subset[of G2 G,OF \<open>valid_unMultigraph G2\<close> valid] by auto
qed
have "valid_graph.is_path G v1 (ps1@((v1',w,v2)#ps2)) v2'"
proof -
have "valid_graph.is_path G v1 ps1 v1'"
by (metis ps1_G valid valid_unMultigraph.is_trail_intro)
moreover have "valid_graph.is_path G v2 ps2 v2'"
by (metis ps2_G valid valid_unMultigraph.is_trail_intro)
moreover have "(v1',w,v2) \<in> edges G"
using G by auto
ultimately show ?thesis
using valid_graph.is_path_split'[OF valid',of v1 ps1 v1' w v2 ps2 v2'] by auto
qed
moreover have "distinct (ps1@((v1',w,v2)#ps2))"
proof -
have "distinct ps1" by (metis ps1_G valid valid_unMultigraph.is_trail_path)
moreover have "distinct ps2"
by (metis ps2_G valid valid_unMultigraph.is_trail_path)
moreover have "set ps1 \<inter> set ps2 = {}"
proof -
have "set ps1 \<subseteq>edges G1"
by (metis assms(3) assms(5) valid_unMultigraph.is_Eulerian_trail_def
valid_unMultigraph.path_in_edges)
moreover have "set ps2 \<subseteq> edges G2"
by (metis assms(4) assms(6) valid_unMultigraph.is_Eulerian_trail_def
valid_unMultigraph.path_in_edges)
ultimately show ?thesis using \<open>edges G1 \<inter> edges G2={}\<close> by auto
qed
moreover have "(v1',w,v2)\<notin>edges G1"
using \<open>v2 \<in> nodes G2\<close> \<open>valid_graph G1\<close>
by (metis Int_iff all_not_in_conv assms(1) valid_graph.E_validD(2))
hence "(v1',w,v2)\<notin>set ps1"
by (metis (full_types) assms(3) assms(5) subsetD valid_unMultigraph.path_in_edges
valid_unMultigraph.is_Eulerian_trail_def )
moreover have "(v1',w,v2)\<notin>edges G2"
using \<open>v1' \<in> nodes G1\<close> \<open>valid_graph G2\<close>
by (metis assms(1) disjoint_iff_not_equal valid_graph.E_validD(1))
hence "(v1',w,v2)\<notin>set ps2"
by (metis (full_types) assms(4) assms(6) in_mono valid_unMultigraph.path_in_edges
valid_unMultigraph.is_Eulerian_trail_def )
ultimately show ?thesis using distinct_append by auto
qed
moreover have "set (ps1@((v1',w,v2)#ps2)) \<inter> set (rev_path (ps1@((v1',w,v2)#ps2))) = {}"
proof -
have "set ps1 \<inter> set (rev_path ps1) = {}"
by (metis ps1_G valid valid_unMultigraph.is_trail_path)
moreover have "set (rev_path ps2) \<subseteq> edges G2"
by (metis assms(4) assms(6) valid_unMultigraph.is_trail_rev
valid_unMultigraph.is_Eulerian_trail_def valid_unMultigraph.path_in_edges)
hence "set ps1 \<inter> set (rev_path ps2) = {}"
using assms
valid_unMultigraph.path_in_edges[OF \<open>valid_unMultigraph G1\<close>, of v1 ps1 v1']
valid_unMultigraph.path_in_edges[OF \<open>valid_unMultigraph G2\<close>, of v2 ps2 v2']
unfolding valid_unMultigraph.is_Eulerian_trail_def[OF \<open>valid_unMultigraph G1\<close>]
valid_unMultigraph.is_Eulerian_trail_def[OF \<open>valid_unMultigraph G2\<close>]
by auto
moreover have "set ps2 \<inter> set (rev_path ps2) = {}"
by (metis ps2_G valid valid_unMultigraph.is_trail_path)
moreover have "set (rev_path ps1) \<subseteq>edges G1"
by (metis assms(3) assms(5) valid_unMultigraph.is_Eulerian_trail_def
valid_unMultigraph.path_in_edges valid_unMultigraph.euclerian_rev)
hence "set ps2 \<inter> set (rev_path ps1) = {}"
by (metis calculation(2) distinct_append distinct_rev_path ps1_G ps2_G rev_path_append
rev_path_double valid valid_unMultigraph.is_trail_path)
moreover have "(v2,w,v1')\<notin>set (ps1@((v1',w,v2)#ps2))"
proof -
have "(v2,w,v1')\<notin>edges G1"
using \<open>v2 \<in> nodes G2\<close> \<open>valid_graph G1\<close>
by (metis Int_iff all_not_in_conv assms(1) valid_graph.E_validD(1))
hence "(v2,w,v1')\<notin>set ps1"
by (metis assms(3) assms(5) split_list valid_unMultigraph.is_trail_split'
valid_unMultigraph.is_Eulerian_trail_def)
moreover have "(v2,w,v1')\<notin>edges G2"
using \<open>v1' \<in> nodes G1\<close> \<open>valid_graph G2\<close>
by (metis IntI assms(1) empty_iff valid_graph.E_validD(2))
hence "(v2,w,v1')\<notin>set ps2"
by (metis (full_types) assms(4) assms(6) in_mono valid_unMultigraph.path_in_edges
valid_unMultigraph.is_Eulerian_trail_def)
moreover have "(v2,w,v1')\<noteq>(v1',w,v2)"
using \<open>v1' \<in> nodes G1\<close> \<open>v2 \<in> nodes G2\<close>
by (metis IntI Pair_inject assms(1) assms(5) bex_empty)
ultimately show ?thesis by auto
qed
ultimately show ?thesis using rev_path_append by auto
qed
ultimately show ?thesis using valid_unMultigraph.is_trail_path[OF valid]
by auto
qed
moreover have "edges (rem_unPath (ps1@((v1',w,v2)#ps2)) G)= {}"
proof -
have "edges (rem_unPath (ps1@((v1',w,v2)#ps2)) G)=edges G -
(set (ps1@((v1',w,v2)#ps2)) \<union> set (rev_path (ps1@((v1',w,v2)#ps2))))"
by (metis rem_unPath_edges)
also have "...=edges G - (set ps1 \<union> set ps2 \<union> set (rev_path ps1) \<union> set (rev_path ps2)
\<union> {(v1',w,v2),(v2,w,v1')})" using rev_path_append by auto
finally have "edges (rem_unPath (ps1@((v1',w,v2)#ps2)) G) = edges G - (set ps1 \<union>
set ps2 \<union> set (rev_path ps1) \<union> set (rev_path ps2) \<union> {(v1',w,v2),(v2,w,v1')})" .
moreover have "edges (rem_unPath ps1 G1)={}"
by (metis assms(3) assms(5) valid_unMultigraph.is_Eulerian_trail_def)
hence "edges G1 - (set ps1 \<union> set (rev_path ps1))={}"
by (metis rem_unPath_edges)
moreover have "edges (rem_unPath ps2 G2)={}"
by (metis assms(4) assms(6) valid_unMultigraph.is_Eulerian_trail_def)
hence "edges G2 - (set ps2 \<union> set (rev_path ps2))={}"
by (metis rem_unPath_edges)
ultimately show ?thesis using G by auto
qed
ultimately show ?thesis by (metis G valid valid_unMultigraph.is_Eulerian_trail_def)
qed
lemma (in valid_unMultigraph) eulerian_sufficient:
assumes "finite V" "finite E" "connected" "V\<noteq>{}"
shows "num_of_odd_nodes G = 2 \<Longrightarrow>
(\<exists>v\<in>V.\<exists>v'\<in>V.\<exists>ps. odd(degree v G)\<and>odd(degree v' G)\<and>(v\<noteq>v')\<and>is_Eulerian_trail v ps v')"
and "num_of_odd_nodes G=0 \<Longrightarrow> (\<forall>v\<in>V.\<exists>ps. is_Eulerian_circuit v ps v)"
using \<open>finite E\<close> \<open>finite V\<close> valid_unMultigraph_axioms \<open>V\<noteq>{}\<close> \<open>connected\<close>
proof (induct "card E" arbitrary: G rule: less_induct)
case less
assume "finite (edges G)" and "finite (nodes G)" and "valid_unMultigraph G" and "nodes G\<noteq>{}"
and "valid_unMultigraph.connected G" and "num_of_odd_nodes G = 2"
have "valid_graph G" using \<open>valid_unMultigraph G\<close> valid_unMultigraph_def by auto
obtain n1 n2 where
n1: "n1\<in>nodes G" "odd(degree n1 G)"
and n2: "n2\<in>nodes G" "odd(degree n2 G)"
and "n1\<noteq>n2" unfolding num_of_odd_nodes_def odd_nodes_set_def
proof -
have "\<forall>S. card S=2 \<longrightarrow> (\<exists>n1 n2. n1\<in>S\<and>n2\<in>S\<and>n1\<noteq>n2)"
by (metis card_eq_0_iff equals0I even_card' even_numeral zero_neq_numeral)
then obtain t1 t2
where "t1\<in>{v \<in> nodes G. odd (degree v G)}" "t2\<in>{v \<in> nodes G. odd (degree v G)}" "t1\<noteq>t2"
using \<open>num_of_odd_nodes G = 2\<close> unfolding num_of_odd_nodes_def odd_nodes_set_def
by force
thus ?thesis by (metis (lifting) that mem_Collect_eq)
qed
have even_except_two:"\<And>n. n\<in>nodes G\<Longrightarrow> n\<noteq>n1 \<Longrightarrow> n\<noteq>n2 \<Longrightarrow> even(degree n G)"
proof (rule ccontr)
fix n assume "n \<in> nodes G" "n \<noteq> n1" "n \<noteq> n2" "odd (degree n G)"
have "n\<in> odd_nodes_set G"
by (metis (mono_tags) \<open>n \<in> nodes G\<close> \<open>odd (degree n G)\<close> mem_Collect_eq odd_nodes_set_def)
moreover have "n1 \<in> odd_nodes_set G"
by (metis (mono_tags) mem_Collect_eq n1(1) n1(2) odd_nodes_set_def)
moreover have "n2 \<in> odd_nodes_set G"
using n2(1) n2(2) unfolding odd_nodes_set_def by auto
ultimately have "{n,n1,n2}\<subseteq> odd_nodes_set G" by auto
moreover have "card{n,n1,n2} \<ge>3" using \<open>n1\<noteq>n2\<close> \<open>n\<noteq>n1\<close> \<open>n\<noteq>n2\<close> by auto
moreover have "finite (odd_nodes_set G)"
using \<open>finite (nodes G)\<close> unfolding odd_nodes_set_def by auto
ultimately have "card (odd_nodes_set G) \<ge> 3"
using card_mono[of "odd_nodes_set G" "{n, n1, n2}"] by auto
thus False using \<open>num_of_odd_nodes G = 2\<close> unfolding num_of_odd_nodes_def by auto
qed
have "{e \<in> edges G. fst e = n1}\<noteq>{}"
using n1
by (metis (full_types) degree_def empty_iff finite.emptyI odd_card)
then obtain v' w where "(n1,w,v')\<in>edges G" by auto
have "v'=n2 \<Longrightarrow> (\<exists>v\<in>nodes G. \<exists>v'\<in>nodes G.\<exists>ps. odd (degree v G) \<and> odd (degree v' G) \<and> v \<noteq> v'
\<and> valid_unMultigraph.is_Eulerian_trail G v ps v')"
proof (cases "valid_unMultigraph.connected (del_unEdge n1 w n2 G)")
assume "v'=n2"
assume conneted':"valid_unMultigraph.connected (del_unEdge n1 w n2 G)"
moreover have "num_of_odd_nodes (del_unEdge n1 w n2 G) = 0"
using \<open>(n1, w, v') \<in> edges G\<close> \<open>finite (edges G)\<close> \<open>finite (nodes G)\<close> \<open>v' = n2\<close>
\<open>num_of_odd_nodes G = 2\<close> \<open>valid_unMultigraph G\<close> del_UnEdge_odd_odd n1(2) n2(2)
by force
moreover have "finite (edges (del_unEdge n1 w n2 G))"
using \<open>finite (edges G)\<close> by auto
moreover have "finite (nodes (del_unEdge n1 w n2 G))"
using \<open>finite (nodes G)\<close> by auto
moreover have "edges G - {(n1,w,n2),(n2,w,n1)} \<subset> edges G"
using Diff_iff Diff_subset \<open>(n1, w, v') \<in> edges G\<close> \<open>v' = n2\<close>
by fast
hence "card (edges (del_unEdge n1 w n2 G)) < card (edges G)"
using \<open>finite (edges G)\<close> psubset_card_mono[of "edges G" "edges G - {(n1,w,n2),(n2,w,n1)}"]
unfolding del_unEdge_def by auto
moreover have "valid_unMultigraph (del_unEdge n1 w n2 G)"
using \<open>valid_unMultigraph G\<close> del_unEdge_valid by auto
moreover have "nodes (del_unEdge n1 w n2 G) \<noteq> {}"
by (metis (full_types) del_UnEdge_node empty_iff n1(1))
ultimately have "\<forall>v\<in>nodes (del_unEdge n1 w n2 G). \<exists>ps. valid_unMultigraph.is_Eulerian_circuit
(del_unEdge n1 w n2 G) v ps v"
using less.hyps[of "del_unEdge n1 w n2 G"] by auto
thus ?thesis using eulerian_cons
by (metis \<open>(n1, w, v') \<in> edges G\<close> \<open>n1 \<noteq> n2\<close> \<open>v' = n2\<close> \<open>valid_unMultigraph G\<close>
\<open>valid_unMultigraph (del_unEdge n1 w n2 G)\<close> del_UnEdge_node n1(1) n1(2) n2(1) n2(2)
valid_unMultigraph.eulerian_cons valid_unMultigraph.is_Eulerian_circuit_def)
next
assume "v'=n2"
assume not_conneted:"\<not>valid_unMultigraph.connected (del_unEdge n1 w n2 G)"
have valid0:"valid_unMultigraph (del_unEdge n1 w n2 G)"
using \<open>valid_unMultigraph G\<close> del_unEdge_valid by auto
hence valid0':"valid_graph (del_unEdge n1 w n2 G)"
using valid_unMultigraph_def by auto
have all_even:"\<forall>n\<in>nodes (del_unEdge n1 w n2 G). even(degree n (del_unEdge n1 w n2 G))"
proof -
have "even (degree n1 (del_unEdge n1 w n2 G))"
using \<open>(n1, w, v') \<in> edges G\<close> \<open>finite (edges G)\<close> \<open>v' = n2\<close> \<open>valid_unMultigraph G\<close> n1
by (auto simp add: valid_unMultigraph.corres)
moreover have "even (degree n2 (del_unEdge n1 w n2 G))"
using \<open>(n1, w, v') \<in> edges G\<close> \<open>finite (edges G)\<close> \<open>v' = n2\<close> \<open>valid_unMultigraph G\<close> n2
by (auto simp add: valid_unMultigraph.corres)
moreover have "\<And>n. n \<in> nodes (del_unEdge n1 w n2 G) \<Longrightarrow> n \<noteq> n1 \<Longrightarrow> n \<noteq> n2 \<Longrightarrow>
even (degree n (del_unEdge n1 w n2 G))"
using valid_unMultigraph.degree_frame[OF \<open>valid_unMultigraph G\<close>,
of _ n1 n2 w] even_except_two
by (metis (no_types) \<open>finite (edges G)\<close> del_unEdge_def empty_iff insert_iff
select_convs(1))
ultimately show ?thesis by auto
qed
have "(n1,w,n2)\<in>edges G" by (metis \<open>(n1, w, v') \<in> edges G\<close> \<open>v' = n2\<close>)
hence "(n2,w,n1)\<in>edges G" by (metis \<open>valid_unMultigraph G\<close> valid_unMultigraph.corres)
obtain G1 G2 where
G1_nodes: "nodes G1={n. \<exists>ps. valid_graph.is_path (del_unEdge n1 w n2 G) n ps n1}"
and G1_edges: "edges G1={(n,e,n'). (n,e,n')\<in>edges (del_unEdge n1 w n2 G)
\<and> n\<in>nodes G1 \<and> n'\<in>nodes G1}"
and G2_nodes:"nodes G2={n. \<exists>ps. valid_graph.is_path (del_unEdge n1 w n2 G) n ps n2}"
and G2_edges:"edges G2={(n,e,n'). (n,e,n')\<in>edges (del_unEdge n1 w n2 G) \<and> n\<in>nodes G2
\<and> n'\<in>nodes G2}"
and G1_G2_edges_union:"edges G1 \<union> edges G2 = edges (del_unEdge n1 w n2 G)"
and "edges G1 \<inter> edges G2={}"
and G1_G2_nodes_union:"nodes G1 \<union> nodes G2=nodes (del_unEdge n1 w n2 G)"
and "nodes G1 \<inter> nodes G2={}"
and "valid_unMultigraph G1"
and "valid_unMultigraph G2"
and "valid_unMultigraph.connected G1"
and "valid_unMultigraph.connected G2"
using valid_unMultigraph.connectivity_split[OF \<open>valid_unMultigraph G\<close>
\<open>valid_unMultigraph.connected G\<close> \<open>\<not> valid_unMultigraph.connected (del_unEdge n1 w n2 G)\<close>
\<open>(n1, w, n2) \<in> edges G\<close> ] .
have "edges (del_unEdge n1 w n2 G) \<subset> edges G"
unfolding del_unEdge_def using \<open>(n1, w, n2)\<in>edges G\<close> \<open>(n2, w, n1)\<in>edges G\<close> by auto
hence "card (edges G1) < card (edges G)" using G1_G2_edges_union
by (metis (full_types) \<open>finite (edges G)\<close> inf_sup_absorb less_infI2 psubset_card_mono)
moreover have "finite (edges G1)"
using G1_G2_edges_union \<open>finite (edges G)\<close>
by (metis \<open>edges (del_unEdge n1 w n2 G) \<subset> edges G\<close> finite_Un less_imp_le rev_finite_subset)
moreover have "nodes G1 \<subseteq> nodes (del_unEdge n1 w n2 G)"
by (metis G1_G2_nodes_union Un_upper1)
hence "finite (nodes G1)"
using \<open>finite (nodes G)\<close> del_UnEdge_node rev_finite_subset by auto
moreover have "n1 \<in> nodes G1"
proof -
have "n1\<in>nodes (del_unEdge n1 w n2 G)" using \<open>n1\<in>nodes G\<close> by auto
hence "valid_graph.is_path (del_unEdge n1 w n2 G) n1 [] n1"
using valid0' by (metis valid_graph.is_path_simps(1))
thus ?thesis using G1_nodes by auto
qed
hence "nodes G1 \<noteq> {}" by auto
moreover have "num_of_odd_nodes G1 = 0"
proof -
have "valid_graph G2" using \<open>valid_unMultigraph G2\<close> valid_unMultigraph_def by auto
hence "\<forall>n\<in>nodes G1. degree n G1 = degree n (del_unEdge n1 w n2 G)"
using sub_graph_degree_frame[of G2 G1 "(del_unEdge n1 w n2 G)"]
by (metis G1_G2_edges_union \<open>nodes G1 \<inter> nodes G2 = {}\<close>)
hence "\<forall>n\<in>nodes G1. even(degree n G1)" using all_even
by (metis G1_G2_nodes_union Un_iff)
thus ?thesis
unfolding num_of_odd_nodes_def odd_nodes_set_def
by (metis (lifting) Collect_empty_eq card_eq_0_iff)
qed
ultimately have "\<forall>v\<in>nodes G1. \<exists>ps. valid_unMultigraph.is_Eulerian_circuit G1 v ps v"
using less.hyps[of G1] \<open>valid_unMultigraph G1\<close> \<open>valid_unMultigraph.connected G1\<close>
by auto
then obtain ps1 where ps1:"valid_unMultigraph.is_Eulerian_trail G1 n1 ps1 n1"
using \<open>n1\<in>nodes G1\<close>
by (metis (full_types) \<open>valid_unMultigraph G1\<close> valid_unMultigraph.is_Eulerian_circuit_def)
have "card (edges G2) < card (edges G)"
using G1_G2_edges_union \<open>edges (del_unEdge n1 w n2 G) \<subset> edges G\<close>
by (metis (full_types) \<open>finite (edges G)\<close> inf_sup_ord(4) le_less_trans psubset_card_mono)
moreover have "finite (edges G2)"
using G1_G2_edges_union \<open>finite (edges G)\<close>
by (metis \<open>edges (del_unEdge n1 w n2 G) \<subset> edges G\<close> finite_Un less_imp_le rev_finite_subset)
moreover have "nodes G2 \<subseteq> nodes (del_unEdge n1 w n2 G)"
by (metis G1_G2_nodes_union Un_upper2)
hence "finite (nodes G2)"
using \<open>finite (nodes G)\<close> del_UnEdge_node rev_finite_subset by auto
moreover have "n2 \<in> nodes G2"
proof -
have "n2\<in>nodes (del_unEdge n1 w n2 G)"
using \<open>n2\<in>nodes G\<close> by auto
hence "valid_graph.is_path (del_unEdge n1 w n2 G) n2 [] n2"
using valid0' by (metis valid_graph.is_path_simps(1))
thus ?thesis using G2_nodes by auto
qed
hence "nodes G2 \<noteq> {}" by auto
moreover have "num_of_odd_nodes G2 = 0"
proof -
have "valid_graph G1" using \<open>valid_unMultigraph G1\<close> valid_unMultigraph_def by auto
hence "\<forall>n\<in>nodes G2. degree n G2 = degree n (del_unEdge n1 w n2 G)"
using sub_graph_degree_frame[of G1 G2 "(del_unEdge n1 w n2 G)"]
by (metis G1_G2_edges_union \<open>nodes G1 \<inter> nodes G2 = {}\<close> inf_commute sup_commute)
hence "\<forall>n\<in>nodes G2. even(degree n G2)" using all_even
by (metis G1_G2_nodes_union Un_iff)
thus ?thesis
unfolding num_of_odd_nodes_def odd_nodes_set_def
by (metis (lifting) Collect_empty_eq card_eq_0_iff)
qed
ultimately have "\<forall>v\<in>nodes G2. \<exists>ps. valid_unMultigraph.is_Eulerian_circuit G2 v ps v"
using less.hyps[of G2] \<open>valid_unMultigraph G2\<close> \<open>valid_unMultigraph.connected G2\<close>
by auto
then obtain ps2 where ps2:"valid_unMultigraph.is_Eulerian_trail G2 n2 ps2 n2"
using \<open>n2\<in>nodes G2\<close>
by (metis (full_types) \<open>valid_unMultigraph G2\<close> valid_unMultigraph.is_Eulerian_circuit_def)
have "\<lparr>nodes = nodes G1 \<union> nodes G2, edges = edges G1 \<union> edges G2 \<union> {(n1, w, n2),
(n2, w, n1)}\<rparr>=G"
proof -
have "edges (del_unEdge n1 w n2 G) \<union> {(n1, w, n2),(n2, w, n1)} =edges G"
using \<open>(n1,w,n2)\<in>edges G\<close> \<open>(n2,w,n1)\<in>edges G\<close>
unfolding del_unEdge_def by auto
moreover have "nodes (del_unEdge n1 w n2 G)=nodes G"
unfolding del_unEdge_def by auto
ultimately have "\<lparr>nodes = nodes (del_unEdge n1 w n2 G), edges =
edges (del_unEdge n1 w n2 G) \<union> {(n1, w, n2), (n2, w, n1)}\<rparr>=G"
by auto
moreover have "\<lparr>nodes = nodes G1 \<union> nodes G2, edges = edges G1 \<union> edges G2 \<union>
{(n1, w, n2),(n2, w, n1)}\<rparr>=\<lparr>nodes = nodes (del_unEdge n1 w n2 G),edges
= edges (del_unEdge n1 w n2 G) \<union> {(n1, w, n2), (n2, w, n1)}\<rparr>"
by (metis G1_G2_edges_union G1_G2_nodes_union)
ultimately show ?thesis by auto
qed
moreover have "valid_unMultigraph.is_Eulerian_trail \<lparr>nodes = nodes G1 \<union> nodes G2,
edges = edges G1 \<union> edges G2 \<union> {(n1, w, n2), (n2, w, n1)}\<rparr> n1 (ps1 @ (n1, w, n2) # ps2) n2"
using eulerian_split[of G1 G2 n1 ps1 n1 n2 ps2 n2 w]
by (metis \<open>edges G1 \<inter> edges G2 = {}\<close> \<open>nodes G1 \<inter> nodes G2 = {}\<close> \<open>valid_unMultigraph G1\<close>
\<open>valid_unMultigraph G2\<close> ps1 ps2)
ultimately show ?thesis by (metis \<open>n1 \<noteq> n2\<close> n1(1) n1(2) n2(1) n2(2))
qed
moreover have "v'\<noteq>n2 \<Longrightarrow> (\<exists>v\<in>nodes G. \<exists>v'\<in>nodes G.\<exists>ps. odd (degree v G) \<and> odd (degree v' G)
\<and> v \<noteq> v' \<and> valid_unMultigraph.is_Eulerian_trail G v ps v')"
proof (cases "valid_unMultigraph.connected (del_unEdge n1 w v' G)")
case True
assume "v' \<noteq> n2"
assume connected':"valid_unMultigraph.connected (del_unEdge n1 w v' G)"
have "n1 \<in> nodes (del_unEdge n1 w v' G)" by (metis del_UnEdge_node n1(1))
hence even_n1:"even(degree n1 (del_unEdge n1 w v' G))"
using valid_unMultigraph.del_UnEdge_even[OF \<open>valid_unMultigraph G\<close> \<open>(n1, w, v') \<in> edges G\<close>
\<open>finite (edges G)\<close>] \<open>odd (degree n1 G)\<close>
unfolding odd_nodes_set_def by auto
moreover have odd_n2:"odd(degree n2 (del_unEdge n1 w v' G))"
using valid_unMultigraph.degree_frame[OF \<open>valid_unMultigraph G\<close> \<open>finite (edges G)\<close>,
of n2 n1 v' w] \<open>n1 \<noteq> n2\<close> \<open>v' \<noteq> n2\<close>
by (metis empty_iff insert_iff n2(2))
moreover have "even (degree v' G)"
using even_except_two[of v']
by (metis (full_types) \<open>(n1, w, v') \<in> edges G\<close> \<open>v' \<noteq> n2\<close> \<open>valid_graph G\<close>
\<open>valid_unMultigraph G\<close> valid_graph.E_validD(2) valid_unMultigraph.no_id)
hence odd_v':"odd(degree v' (del_unEdge n1 w v' G))"
using valid_unMultigraph.del_UnEdge_even'[OF \<open>valid_unMultigraph G\<close> \<open>(n1, w, v') \<in> edges G\<close>
\<open>finite (edges G)\<close>]
unfolding odd_nodes_set_def by auto
ultimately have two_odds:"num_of_odd_nodes (del_unEdge n1 w v' G) = 2"
by (metis (lifting) \<open>v' \<noteq> n2\<close> \<open>valid_graph G\<close> \<open>valid_unMultigraph G\<close>
\<open>(n1, w, v') \<in> edges G\<close> \<open>finite (edges G)\<close> \<open>finite (nodes G)\<close> \<open>num_of_odd_nodes G = 2\<close>
del_UnEdge_odd_even even_except_two n1(2) valid_graph.E_validD(2))
moreover have valid0:"valid_unMultigraph (del_unEdge n1 w v' G)"
using del_unEdge_valid \<open>valid_unMultigraph G\<close> by auto
moreover have " edges G - {(n1, w, v'), (v', w, n1)} \<subset> edges G"
using \<open>(n1,w,v')\<in>edges G\<close> by auto
hence "card (edges (del_unEdge n1 w v' G)) < card (edges G)"
using \<open>finite (edges G)\<close> unfolding del_unEdge_def
by (metis (hide_lams, no_types) psubset_card_mono select_convs(2))
moreover have "finite (edges (del_unEdge n1 w v' G))"
unfolding del_unEdge_def
by (metis (full_types) \<open>finite (edges G)\<close> finite_Diff select_convs(2))
moreover have "finite (nodes (del_unEdge n1 w v' G))"
unfolding del_unEdge_def by (metis \<open>finite (nodes G)\<close> select_convs(1))
moreover have "nodes (del_unEdge n1 w v' G) \<noteq> {}"
by (metis (full_types) del_UnEdge_node empty_iff n1(1))
ultimately obtain s t ps where
s: "s\<in>nodes (del_unEdge n1 w v' G)" "odd (degree s (del_unEdge n1 w v' G))"
and t:"t\<in>nodes (del_unEdge n1 w v' G)" "odd (degree t (del_unEdge n1 w v' G))"
and "s \<noteq> t"
and s_ps_t: "valid_unMultigraph.is_Eulerian_trail (del_unEdge n1 w v' G) s ps t"
using connected' less.hyps[of "(del_unEdge n1 w v' G)"] by auto
hence "(s=n2\<and>t=v')\<or>(s=v'\<and>t=n2)"
using odd_n2 odd_v' two_odds \<open>finite (edges G)\<close>\<open>valid_unMultigraph G\<close>
by (metis (mono_tags) del_UnEdge_node empty_iff even_except_two even_n1 insert_iff
valid_unMultigraph.degree_frame)
moreover have "s=n2\<Longrightarrow>t=v'\<Longrightarrow>?thesis"
by (metis \<open>(n1, w, v') \<in> edges G\<close> \<open>n1 \<noteq> n2\<close> \<open>valid_unMultigraph G\<close> n1(1) n1(2) n2(1) n2(2)
s_ps_t valid0 valid_unMultigraph.euclerian_rev valid_unMultigraph.eulerian_cons)
moreover have "s=v'\<Longrightarrow>t=n2\<Longrightarrow>?thesis"
by (metis \<open>(n1, w, v') \<in> edges G\<close> \<open>n1 \<noteq> n2\<close> \<open>valid_unMultigraph G\<close> n1(1) n1(2) n2(1) n2(2)
s_ps_t valid_unMultigraph.eulerian_cons)
ultimately show ?thesis by auto
next
case False
assume "v'\<noteq>n2"
assume not_conneted:"\<not>valid_unMultigraph.connected (del_unEdge n1 w v' G)"
have "(v',w,n1)\<in>edges G" using \<open>(n1,w,v')\<in>edges G\<close>
by (metis \<open>valid_unMultigraph G\<close> valid_unMultigraph.corres)
have valid0:"valid_unMultigraph (del_unEdge n1 w v' G)"
using \<open>valid_unMultigraph G\<close> del_unEdge_valid by auto
hence valid0':"valid_graph (del_unEdge n1 w v' G)"
using valid_unMultigraph_def by auto
have even_n1:"even(degree n1 (del_unEdge n1 w v' G))"
using valid_unMultigraph.del_UnEdge_even[OF \<open>valid_unMultigraph G\<close> \<open>(n1,w,v')\<in>edges G\<close>
\<open>finite (edges G)\<close>] n1
unfolding odd_nodes_set_def by auto
moreover have odd_n2:"odd(degree n2 (del_unEdge n1 w v' G))"
using \<open>n1 \<noteq> n2\<close> \<open>v' \<noteq> n2\<close> n2 valid_unMultigraph.degree_frame[OF \<open>valid_unMultigraph G\<close>
\<open>finite (edges G)\<close>, of n2 n1 v' w]
by auto
moreover have "v'\<noteq>n1"
using valid_unMultigraph.no_id[OF \<open>valid_unMultigraph G\<close>] \<open>(n1,w,v')\<in>edges G\<close> by auto
hence odd_v':"odd(degree v' (del_unEdge n1 w v' G))"
using \<open>v' \<noteq> n2\<close> even_except_two[of v']
valid_graph.E_validD(2)[OF \<open>valid_graph G\<close> \<open>(n1, w, v') \<in> edges G\<close>]
valid_unMultigraph.del_UnEdge_even'[OF \<open>valid_unMultigraph G\<close> \<open>(n1, w, v') \<in> edges G\<close>
\<open>finite (edges G)\<close> ]
unfolding odd_nodes_set_def by auto
ultimately have even_except_two':"\<And>n. n\<in>nodes (del_unEdge n1 w v' G)\<Longrightarrow> n\<noteq>n2
\<Longrightarrow> n\<noteq>v'\<Longrightarrow> even(degree n (del_unEdge n1 w v' G))"
using del_UnEdge_node[of _ n1 w v' G] even_except_two valid_unMultigraph.degree_frame[OF
\<open>valid_unMultigraph G\<close> \<open>finite (edges G)\<close>, of _ n1 v' w]
by force
obtain G1 G2 where
G1_nodes: "nodes G1={n. \<exists>ps. valid_graph.is_path (del_unEdge n1 w v' G) n ps n1}"
and G1_edges: "edges G1={(n,e,n'). (n,e,n')\<in>edges (del_unEdge n1 w v' G) \<and> n\<in>nodes G1
\<and> n'\<in>nodes G1}"
and G2_nodes:"nodes G2={n. \<exists>ps. valid_graph.is_path (del_unEdge n1 w v' G) n ps v'}"
and G2_edges:"edges G2={(n,e,n'). (n,e,n')\<in>edges (del_unEdge n1 w v' G) \<and> n\<in>nodes G2
\<and> n'\<in>nodes G2}"
and G1_G2_edges_union:"edges G1 \<union> edges G2 = edges (del_unEdge n1 w v' G)"
and "edges G1 \<inter> edges G2={}"
and G1_G2_nodes_union:"nodes G1 \<union> nodes G2=nodes (del_unEdge n1 w v' G)"
and "nodes G1 \<inter> nodes G2={}"
and "valid_unMultigraph G1"
and "valid_unMultigraph G2"
and "valid_unMultigraph.connected G1"
and "valid_unMultigraph.connected G2"
using valid_unMultigraph.connectivity_split[OF \<open>valid_unMultigraph G\<close>
\<open>valid_unMultigraph.connected G\<close> not_conneted \<open>(n1,w,v')\<in>edges G\<close>]
.
have "n2\<in>nodes G2" using extend_distinct_path
proof -
have "finite (edges (del_unEdge n1 w v' G))"
unfolding del_unEdge_def using \<open>finite (edges G)\<close> by auto
moreover have "num_of_odd_nodes (del_unEdge n1 w v' G) = 2"
by (metis \<open>(n1, w, v') \<in> edges G\<close> \<open>(v', w, n1) \<in> edges G\<close> \<open>num_of_odd_nodes G = 2\<close>
\<open>v' \<noteq> n2\<close> \<open>valid_graph G\<close> del_UnEdge_even_odd delete_edge_sym even_except_two
\<open>finite (edges G)\<close> \<open>finite (nodes G)\<close> \<open>valid_unMultigraph G\<close>
n1(2) valid_graph.E_validD(2) valid_unMultigraph.no_id)
ultimately have "\<exists>ps. valid_unMultigraph.is_trail (del_unEdge n1 w v' G) n2 ps v'"
using valid_unMultigraph.path_between_odds[OF valid0,of n2 v',OF odd_n2 odd_v'] \<open>v'\<noteq>n2\<close>
by auto
hence "\<exists>ps. valid_graph.is_path (del_unEdge n1 w v' G) n2 ps v'"
by (metis valid0 valid_unMultigraph.is_trail_intro)
thus ?thesis using G2_nodes by auto
qed
have "v'\<in>nodes G2"
proof -
have "valid_graph.is_path (del_unEdge n1 w v' G) v' [] v'"
by (metis (full_types) \<open>(n1, w, v') \<in> edges G\<close> \<open>valid_graph G\<close> del_UnEdge_node
valid0' valid_graph.E_validD(2) valid_graph.is_path_simps(1))
thus ?thesis by (metis (lifting) G2_nodes mem_Collect_eq)
qed
have edges_subset:"edges (del_unEdge n1 w v' G) \<subset> edges G"
using \<open>(n1,w,v')\<in>edges G\<close> \<open>(v',w,n1)\<in>edges G\<close>
unfolding del_unEdge_def by auto
hence "card (edges G1) < card (edges G)"
by (metis G1_G2_edges_union inf_sup_absorb \<open>finite (edges G)\<close> less_infI2 psubset_card_mono)
moreover have "finite (edges G1)"
by (metis (full_types) G1_G2_edges_union edges_subset finite_Un finite_subset
\<open>finite (edges G)\<close> less_imp_le)
moreover have "finite (nodes G1)"
using G1_G2_nodes_union \<open>finite (nodes G)\<close>
unfolding del_unEdge_def
by (metis (full_types) finite_Un select_convs(1))
moreover have "n1\<in>nodes G1"
proof -
have "valid_graph.is_path (del_unEdge n1 w v' G) n1 [] n1"
by (metis (full_types) del_UnEdge_node n1(1) valid0' valid_graph.is_path_simps(1))
thus ?thesis by (metis (lifting) G1_nodes mem_Collect_eq)
qed
moreover hence "nodes G1 \<noteq> {}" by auto
moreover have "num_of_odd_nodes G1 = 0"
proof -
have "\<forall>n\<in>nodes G1. even(degree n (del_unEdge n1 w v' G))"
using even_except_two' odd_v' odd_n2 \<open>n2\<in>nodes G2\<close> \<open>nodes G1 \<inter> nodes G2 = {}\<close>
\<open>v'\<in>nodes G2\<close>
by (metis (full_types) G1_G2_nodes_union Un_iff disjoint_iff_not_equal)
moreover have "valid_graph G2"
using \<open>valid_unMultigraph G2\<close> valid_unMultigraph_def
by auto
ultimately have "\<forall>n\<in>nodes G1. even(degree n G1)"
using sub_graph_degree_frame[of G2 G1 "del_unEdge n1 w v' G"]
by (metis G1_G2_edges_union \<open>nodes G1 \<inter> nodes G2 = {}\<close>)
thus ?thesis unfolding num_of_odd_nodes_def odd_nodes_set_def
by (metis (lifting) card_eq_0_iff empty_Collect_eq)
qed
ultimately obtain ps1 where ps1:"valid_unMultigraph.is_Eulerian_trail G1 n1 ps1 n1"
using \<open>valid_unMultigraph G1\<close> \<open>valid_unMultigraph.connected G1\<close> less.hyps[of G1]
by (metis valid_unMultigraph.is_Eulerian_circuit_def)
have "card (edges G2) < card (edges G)"
by (metis G1_G2_edges_union \<open>finite (edges G)\<close> edges_subset inf_sup_absorb less_infI2
psubset_card_mono sup_commute)
moreover have "finite (edges G2)"
by (metis (full_types) G1_G2_edges_union edges_subset finite_Un \<open>finite (edges G)\<close> less_le
rev_finite_subset)
moreover have "finite (nodes G2)"
by (metis (mono_tags) G1_G2_nodes_union del_UnEdge_node le_sup_iff \<open>finite (nodes G)\<close>
rev_finite_subset subsetI)
moreover have "nodes G2 \<noteq> {}" using \<open>v'\<in>nodes G2\<close> by auto
moreover have "num_of_odd_nodes G2 = 2"
proof -
have "\<forall>n\<in>nodes G2. n\<notin>{n2,v'}\<longrightarrow>even(degree n (del_unEdge n1 w v' G))"
using even_except_two'
by (metis (full_types) G1_G2_nodes_union Un_iff insert_iff)
moreover have "valid_graph G1"
using \<open>valid_unMultigraph G1\<close> valid_unMultigraph_def by auto
ultimately have "\<forall>n\<in>nodes G2. n\<notin>{n2,v'}\<longrightarrow>even(degree n G2)"
using sub_graph_degree_frame[of G1 G2 "del_unEdge n1 w v' G"]
by (metis G1_G2_edges_union Int_commute Un_commute \<open>nodes G1 \<inter> nodes G2 = {}\<close>)
hence "\<forall>n\<in>nodes G2. n\<notin>{n2,v'}\<longrightarrow>n\<notin>{v \<in> nodes G2. odd (degree v G2)}"
by (metis (lifting) mem_Collect_eq)
moreover have "odd(degree n2 G2)"
using sub_graph_degree_frame[of G1 G2 "del_unEdge n1 w v' G"]
by (metis (hide_lams, no_types) G1_G2_edges_union \<open>nodes G1 \<inter> nodes G2 = {}\<close>
\<open>valid_graph G1\<close> \<open>n2 \<in> nodes G2\<close> inf_assoc inf_bot_right inf_sup_absorb
odd_n2 sup_bot_right sup_commute)
hence "n2\<in>{v \<in> nodes G2. odd (degree v G2)}"
by (metis (lifting) \<open>n2 \<in> nodes G2\<close> mem_Collect_eq)
moreover have "odd(degree v' G2)"
using sub_graph_degree_frame[of G1 G2 "del_unEdge n1 w v' G"]
by (metis G1_G2_edges_union Int_commute Un_commute \<open>nodes G1 \<inter> nodes G2 = {}\<close>
\<open>v' \<in> nodes G2\<close> \<open>valid_graph G1\<close> odd_v')
hence "v'\<in>{v \<in> nodes G2. odd (degree v G2)}"
by (metis (full_types) Collect_conj_eq Collect_mem_eq Int_Collect \<open>v' \<in> nodes G2\<close>)
ultimately have "{v \<in> nodes G2. odd (degree v G2)}={n2,v'}"
using \<open>finite (nodes G2)\<close> by (induct G2,auto)
thus ?thesis using \<open>v'\<noteq>n2\<close>
unfolding num_of_odd_nodes_def odd_nodes_set_def by auto
qed
ultimately obtain s t ps2 where
s: "s\<in>nodes G2" "odd (degree s G2)"
and t:"t\<in>nodes G2" "odd (degree t G2)"
and "s \<noteq> t"
and s_ps2_t: "valid_unMultigraph.is_Eulerian_trail G2 s ps2 t"
using \<open>valid_unMultigraph G2\<close> \<open>valid_unMultigraph.connected G2\<close> less.hyps[of G2]
by auto
moreover have "valid_graph G1"
using \<open>valid_unMultigraph G1\<close> valid_unMultigraph_def by auto
ultimately have "(s=n2\<and>t=v')\<or>(s=v'\<and>t=n2)"
using odd_n2 odd_v' even_except_two'
sub_graph_degree_frame[of G1 G2 "(del_unEdge n1 w v' G)"]
by (metis G1_G2_edges_union G1_G2_nodes_union UnI1 \<open>nodes G1 \<inter> nodes G2 = {}\<close> inf_commute
sup.commute)
moreover have merge_G1_G2:"\<lparr>nodes = nodes G1 \<union> nodes G2, edges = edges G1 \<union> edges G2 \<union>
{(n1, w,v'),(v', w, n1)}\<rparr>=G"
proof -
have "edges (del_unEdge n1 w v' G) \<union> {(n1, w, v'),(v', w, n1)} =edges G"
using \<open>(n1,w,v')\<in>edges G\<close> \<open>(v',w,n1)\<in>edges G\<close>
unfolding del_unEdge_def by auto
moreover have "nodes (del_unEdge n1 w v' G)=nodes G"
unfolding del_unEdge_def by auto
ultimately have "\<lparr>nodes = nodes (del_unEdge n1 w v' G), edges =
edges (del_unEdge n1 w v' G) \<union> {(n1, w, v'), (v', w, n1)}\<rparr>=G"
by auto
moreover have "\<lparr>nodes = nodes G1 \<union> nodes G2, edges = edges G1 \<union> edges G2 \<union>
{(n1, w, v'),(v', w, n1)}\<rparr>=\<lparr>nodes = nodes (del_unEdge n1 w v' G),edges
= edges (del_unEdge n1 w v' G) \<union> {(n1, w, v'), (v', w, n1)}\<rparr>"
by (metis G1_G2_edges_union G1_G2_nodes_union)
ultimately show ?thesis by auto
qed
moreover have "s=n2\<Longrightarrow>t=v'\<Longrightarrow>?thesis"
using eulerian_split[of G1 G2 n1 ps1 n1 v' "(rev_path ps2)" n2 w] merge_G1_G2
by (metis \<open>edges G1 \<inter> edges G2 = {}\<close> \<open>n1 \<noteq> n2\<close> \<open>nodes G1 \<inter> nodes G2 = {}\<close>
\<open>valid_unMultigraph G1\<close> \<open>valid_unMultigraph G2\<close> n1(1) n1(2) n2(1) n2(2) ps1 s_ps2_t
valid_unMultigraph.euclerian_rev)
moreover have "s=v'\<Longrightarrow>t=n2\<Longrightarrow>?thesis"
using eulerian_split[of G1 G2 n1 ps1 n1 v' ps2 n2 w] merge_G1_G2
by (metis \<open>edges G1 \<inter> edges G2 = {}\<close> \<open>n1 \<noteq> n2\<close> \<open>nodes G1 \<inter> nodes G2 = {}\<close>
\<open>valid_unMultigraph G1\<close> \<open>valid_unMultigraph G2\<close> n1(1) n1(2) n2(1) n2(2) ps1 s_ps2_t)
ultimately show ?thesis by auto
qed
ultimately show "\<exists>v\<in>nodes G. \<exists>v'\<in>nodes G.\<exists>ps. odd (degree v G) \<and> odd (degree v' G) \<and> v \<noteq> v'
\<and> valid_unMultigraph.is_Eulerian_trail G v ps v'"
by auto
next
case less
assume "finite (edges G)" and "finite (nodes G)" and "valid_unMultigraph G" and "nodes G\<noteq>{}"
and "valid_unMultigraph.connected G" and "num_of_odd_nodes G = 0"
show "\<forall>v\<in>nodes G. \<exists>ps. valid_unMultigraph.is_Eulerian_circuit G v ps v"
proof (rule,cases "card (nodes G)=1")
fix v assume "v\<in>nodes G"
assume " card (nodes G) = 1 "
hence "nodes G={v}"
using \<open>v \<in> nodes G\<close> card_Suc_eq[of "nodes G" 0] empty_iff insert_iff[of _ v]
by auto
have "edges G={}"
proof (rule ccontr)
assume "edges G \<noteq> {}"
then obtain e1 e2 e3 where e:"(e1,e2,e3)\<in>edges G" by (metis ex_in_conv prod_cases3)
hence "e1=e3" using \<open>nodes G={v}\<close>
by (metis (hide_lams, no_types) append_Nil2 valid_unMultigraph.is_trail_rev
valid_unMultigraph.is_trail.simps(1) \<open>valid_unMultigraph G\<close> singletonE
valid_unMultigraph.is_trail_split valid_unMultigraph.singleton_distinct_path)
thus False by (metis e \<open>valid_unMultigraph G\<close> valid_unMultigraph.no_id)
qed
hence "valid_unMultigraph.is_Eulerian_circuit G v [] v"
by (metis \<open>nodes G = {v}\<close> insert_subset \<open>valid_unMultigraph G\<close> rem_unPath.simps(1)
subsetI valid_unMultigraph.is_trail.simps(1)
valid_unMultigraph.is_Eulerian_circuit_def
valid_unMultigraph.is_Eulerian_trail_def)
thus "\<exists>ps. valid_unMultigraph.is_Eulerian_circuit G v ps v" by auto
next
fix v assume "v\<in>nodes G"
assume "card (nodes G) \<noteq> 1"
moreover have "card (nodes G)\<noteq>0" using \<open>nodes G\<noteq>{}\<close>
by (metis card_eq_0_iff \<open>finite (nodes G)\<close>)
ultimately have "card (nodes G) \<ge>2" by auto
then obtain n where "card (nodes G) = Suc (Suc n)"
by (metis le_iff_add add_2_eq_Suc)
hence "\<exists>n\<in>nodes G. n\<noteq>v" by (auto dest!: card_eq_SucD)
then obtain v' w where "(v,w,v')\<in>edges G"
proof -
assume pre:"\<And>w v'. (v, w, v') \<in> edges G \<Longrightarrow> thesis"
assume "\<exists>n\<in>nodes G. n \<noteq> v"
then obtain ps where ps:"\<exists>v'. valid_graph.is_path G v ps v' \<and> ps\<noteq>Nil"
using valid_unMultigraph_def
by (metis (full_types) \<open>v \<in> nodes G\<close> \<open>valid_unMultigraph G\<close> valid_graph.is_path.simps(1)
\<open>valid_unMultigraph.connected G\<close> valid_unMultigraph.connected_def)
then obtain v0 w v' where "\<exists>ps'. ps=Cons (v0,w,v') ps'" by (metis neq_Nil_conv prod_cases3)
hence "v0=v"
using valid_unMultigraph_def
by (metis \<open>valid_unMultigraph G\<close> ps valid_graph.is_path.simps(2))
hence "(v,w,v')\<in>edges G"
using valid_unMultigraph_def
by (metis \<open>\<exists>ps'. ps = (v0, w, v') # ps'\<close> \<open>valid_unMultigraph G\<close> ps
valid_graph.is_path.simps(2))
thus ?thesis by (metis pre)
qed
have all_even:"\<forall>x\<in>nodes G. even(degree x G)"
using \<open>finite (nodes G)\<close> \<open>num_of_odd_nodes G = 0\<close>
unfolding num_of_odd_nodes_def odd_nodes_set_def by auto
have odd_v: "odd (degree v (del_unEdge v w v' G))"
using \<open>v \<in> nodes G\<close> all_even valid_unMultigraph.del_UnEdge_even[OF \<open>valid_unMultigraph G\<close>
\<open>(v, w, v') \<in> edges G\<close> \<open>finite (edges G)\<close>]
unfolding odd_nodes_set_def by auto
have odd_v': "odd (degree v' (del_unEdge v w v' G))"
using valid_unMultigraph.del_UnEdge_even'[OF \<open>valid_unMultigraph G\<close> \<open>(v, w, v') \<in> edges G\<close>
\<open>finite (edges G)\<close>]
all_even valid_graph.E_validD(2)[OF _ \<open>(v, w, v') \<in> edges G\<close>]
\<open>valid_unMultigraph G\<close>
unfolding valid_unMultigraph_def odd_nodes_set_def
by auto
have valid_unMulti:"valid_unMultigraph (del_unEdge v w v' G)"
by (metis del_unEdge_valid \<open>valid_unMultigraph G\<close>)
moreover have valid_graph: "valid_graph (del_unEdge v w v' G)"
using valid_unMultigraph_def del_undirected
by (metis \<open>valid_unMultigraph G\<close> delete_edge_valid)
moreover have fin_E': "finite(edges (del_unEdge v w v' G))"
using \<open>finite(edges G)\<close> unfolding del_unEdge_def by auto
moreover have fin_V': "finite(nodes (del_unEdge v w v' G))"
using \<open>finite(nodes G)\<close> unfolding del_unEdge_def by auto
moreover have less_card:"card(edges (del_unEdge v w v' G))<card(edges G)"
unfolding del_unEdge_def using \<open>(v,w,v')\<in>edges G\<close>
by (metis Diff_insert2 card_Diff2_less \<open>finite (edges G)\<close> \<open>valid_unMultigraph G\<close>
select_convs(2) valid_unMultigraph.corres)
moreover have "num_of_odd_nodes (del_unEdge v w v' G) = 2"
using \<open>valid_unMultigraph G\<close> \<open>num_of_odd_nodes G = 0\<close> \<open>v \<in> nodes G\<close> all_even
del_UnEdge_even_even[OF \<open>valid_unMultigraph G\<close> \<open>finite (edges G)\<close> \<open>finite (nodes G)\<close>
\<open>(v, w, v') \<in> edges G\<close>] valid_graph.E_validD(2)[OF _ \<open>(v, w, v') \<in> edges G\<close>]
unfolding valid_unMultigraph_def
by auto
moreover have "valid_unMultigraph.connected (del_unEdge v w v' G)"
using \<open>finite (edges G)\<close> \<open>finite (nodes G)\<close> \<open>valid_unMultigraph G\<close>
\<open>valid_unMultigraph.connected G\<close>
by (metis \<open>(v, w, v') \<in> edges G\<close> all_even valid_unMultigraph.del_unEdge_even_connectivity)
moreover have "nodes(del_unEdge v w v' G)\<noteq>{}"
by (metis \<open>v \<in> nodes G\<close> del_UnEdge_node emptyE)
ultimately obtain n1 n2 ps where
n1_n2:
"n1\<in>nodes (del_unEdge v w v' G)"
"n2\<in>nodes (del_unEdge v w v' G)"
"odd (degree n1 (del_unEdge v w v' G))"
"odd (degree n2 (del_unEdge v w v' G))"
"n1\<noteq>n2"
and
ps_eulerian:
"valid_unMultigraph.is_Eulerian_trail (del_unEdge v w v' G) n1 ps n2"
by (metis \<open>num_of_odd_nodes (del_unEdge v w v' G) = 2\<close> less.hyps(1))
have "n1=v\<Longrightarrow>n2=v'\<Longrightarrow>valid_unMultigraph.is_Eulerian_circuit G v (ps@[(v',w,v)]) v"
using ps_eulerian
by (metis \<open>(v, w, v') \<in> edges G\<close> delete_edge_sym \<open>valid_unMultigraph G\<close>
valid_unMultigraph.corres valid_unMultigraph.eulerian_cons'
valid_unMultigraph.is_Eulerian_circuit_def)
moreover have "n1=v'\<Longrightarrow>n2=v\<Longrightarrow>\<exists>ps. valid_unMultigraph.is_Eulerian_circuit G v ps v"
by (metis \<open>(v, w, v') \<in> edges G\<close> \<open>valid_unMultigraph G\<close> ps_eulerian
valid_unMultigraph.eulerian_cons valid_unMultigraph.is_Eulerian_circuit_def)
moreover have "(n1=v\<and>n2=v')\<or>(n2=v\<and>n1=v')"
by (metis (mono_tags) all_even del_UnEdge_node insert_iff \<open>finite (edges G)\<close>
\<open>valid_unMultigraph G\<close> n1_n2(1) n1_n2(2) n1_n2(3) n1_n2(4) n1_n2(5) singletonE
valid_unMultigraph.degree_frame)
ultimately show "\<exists>ps. valid_unMultigraph.is_Eulerian_circuit G v ps v" by auto
qed
qed
end
diff --git a/thys/Kruskal/Kruskal_Impl.thy b/thys/Kruskal/Kruskal_Impl.thy
--- a/thys/Kruskal/Kruskal_Impl.thy
+++ b/thys/Kruskal/Kruskal_Impl.thy
@@ -1,311 +1,309 @@
section "Kruskal Implementation"
theory Kruskal_Impl
imports Kruskal_Refine Refine_Imperative_HOL.IICF
begin
subsection \<open>Refinement III: concrete edges\<close>
text \<open>Given a concrete representation of edges and their endpoints as a pair, we refine
Kruskal's algorithm to work on these concrete edges.\<close>
locale Kruskal_concrete = Kruskal_interface E V vertices joins forest connected weight
for E V vertices joins forest connected and weight :: "'edge \<Rightarrow> int" +
fixes
\<alpha> :: "'cedge \<Rightarrow> 'edge"
and endpoints :: "'cedge \<Rightarrow> ('a*'a) nres"
assumes
endpoints_refine: "\<alpha> xi = x \<Longrightarrow> endpoints xi \<le> \<Down> Id (a_endpoints x)"
begin
definition wsorted' where "wsorted' == sorted_wrt (\<lambda>x y. weight (\<alpha> x) \<le> weight (\<alpha> y))"
lemma wsorted_map\<alpha>[simp]: "wsorted' s \<Longrightarrow> wsorted (map \<alpha> s)"
by(auto simp: wsorted'_def sorted_wrt_map)
definition "obtain_sorted_carrier' == SPEC (\<lambda>L. wsorted' L \<and> \<alpha> ` set L = E)"
abbreviation concrete_edge_rel :: "('cedge \<times> 'edge) set" where
"concrete_edge_rel \<equiv> br \<alpha> (\<lambda>_. True)"
lemma obtain_sorted_carrier'_refine:
"(obtain_sorted_carrier', obtain_sorted_carrier) \<in> \<langle>\<langle>concrete_edge_rel\<rangle>list_rel\<rangle>nres_rel"
unfolding obtain_sorted_carrier'_def obtain_sorted_carrier_def
apply refine_vcg
apply (auto intro!: RES_refine simp: )
subgoal for s apply(rule exI[where x="map \<alpha> s"])
by(auto simp: map_in_list_rel_conv in_br_conv)
done
definition kruskal2
where "kruskal2 \<equiv> do {
l \<leftarrow> obtain_sorted_carrier';
let initial_union_find = per_init V;
(per, spanning_forest) \<leftarrow> nfoldli l (\<lambda>_. True)
(\<lambda>ce (uf, T). do {
ASSERT (\<alpha> ce \<in> E);
(a,b) \<leftarrow> endpoints ce;
ASSERT (a\<in>V \<and> b\<in>V \<and> a \<in> Domain uf \<and> b \<in> Domain uf );
if \<not> per_compare uf a b then
do {
let uf = per_union uf a b;
ASSERT (ce \<notin> set T);
RETURN (uf, T@[ce])
}
else
RETURN (uf,T)
}) (initial_union_find, []);
RETURN spanning_forest
}"
lemma lst_graph_rel_empty[simp]: "([], {}) \<in> \<langle>concrete_edge_rel\<rangle>list_set_rel"
unfolding list_set_rel_def apply(rule relcompI[where b="[]"])
by (auto simp add: in_br_conv)
lemma loop_initial_rel:
"((per_init V, []), per_init V, {}) \<in> Id \<times>\<^sub>r \<langle>concrete_edge_rel\<rangle>list_set_rel"
by simp
lemma concrete_edge_rel_list_set_rel:
"(a, b) \<in> \<langle>concrete_edge_rel\<rangle>list_set_rel \<Longrightarrow> \<alpha> ` (set a) = b"
by (auto simp: in_br_conv list_set_rel_def dest: list_relD2)
theorem kruskal2_refine: "(kruskal2, kruskal1)\<in>\<langle>\<langle>concrete_edge_rel\<rangle>list_set_rel\<rangle>nres_rel"
unfolding kruskal1_def kruskal2_def Let_def
apply (refine_rcg obtain_sorted_carrier'_refine[THEN nres_relD]
endpoints_refine loop_initial_rel)
by (auto intro!: list_set_rel_append
dest: concrete_edge_rel_list_set_rel
simp: in_br_conv)
end
subsection \<open>Refinement to Imperative/HOL with Sepref-Tool\<close>
text \<open>Given implementations for the operations of getting a list of concrete edges
and getting the endpoints of a concrete edge we synthesize Kruskal in Imperative/HOL.\<close>
locale Kruskal_Impl = Kruskal_concrete E V vertices joins forest connected weight \<alpha> endpoints
for E V vertices joins forest connected and weight :: "'edge \<Rightarrow> int"
and \<alpha> and endpoints :: "nat \<times> int \<times> nat \<Rightarrow> (nat \<times> nat) nres"
+
fixes getEdges :: "(nat \<times> int \<times> nat) list nres"
and getEdges_impl :: "(nat \<times> int \<times> nat) list Heap"
and superE :: "(nat \<times> int \<times> nat) set"
and endpoints_impl :: "(nat \<times> int \<times> nat) \<Rightarrow> (nat \<times> nat) Heap"
assumes
getEdges_refine: "getEdges \<le> SPEC (\<lambda>L. \<alpha> ` set L = E
\<and> (\<forall>(a,wv,b)\<in>set L. weight (\<alpha> (a,wv,b)) = wv) \<and> set L \<subseteq> superE)"
and
getEdges_impl: "(uncurry0 getEdges_impl, uncurry0 getEdges)
\<in> unit_assn\<^sup>k \<rightarrow>\<^sub>a list_assn (nat_assn \<times>\<^sub>a int_assn \<times>\<^sub>a nat_assn)"
and
max_node_is_Max_V: "E = \<alpha> ` set la \<Longrightarrow> max_node la = Max (insert 0 V)"
and
endpoints_impl: "( endpoints_impl, endpoints)
\<in> (nat_assn \<times>\<^sub>a int_assn \<times>\<^sub>a nat_assn)\<^sup>k \<rightarrow>\<^sub>a (nat_assn \<times>\<^sub>a nat_assn)"
begin
lemma this_loc: "Kruskal_Impl E V vertices joins forest connected weight
\<alpha> endpoints getEdges getEdges_impl superE endpoints_impl" by unfold_locales
subsubsection \<open>Refinement IV: given an edge set\<close>
text \<open>We now assume to have an implementation of the operation to obtain a list of the edges of
a graph. By sorting this list we refine @{term obtain_sorted_carrier'}.\<close>
definition "obtain_sorted_carrier'' = do {
l \<leftarrow> SPEC (\<lambda>L. \<alpha> ` set L = E
\<and> (\<forall>(a,wv,b)\<in>set L. weight (\<alpha> (a,wv,b)) = wv) \<and> set L \<subseteq> superE);
SPEC (\<lambda>L. sorted_wrt edges_less_eq L \<and> set L = set l)
}"
lemma wsorted'_sorted_wrt_edges_less_eq:
assumes "\<forall>(a,wv,b)\<in>set s. weight (\<alpha> (a,wv,b)) = wv"
"sorted_wrt edges_less_eq s"
shows "wsorted' s"
using assms apply -
unfolding wsorted'_def unfolding edges_less_eq_def
apply(rule sorted_wrt_mono_rel )
by (auto simp: case_prod_beta)
lemma obtain_sorted_carrier''_refine:
"(obtain_sorted_carrier'', obtain_sorted_carrier') \<in> \<langle>Id\<rangle>nres_rel"
unfolding obtain_sorted_carrier''_def obtain_sorted_carrier'_def
apply refine_vcg
apply(auto simp: in_br_conv wsorted'_sorted_wrt_edges_less_eq
distinct_map map_in_list_rel_conv)
done
definition "obtain_sorted_carrier''' =
do {
l \<leftarrow> getEdges;
RETURN (quicksort_by_rel edges_less_eq [] l, max_node l)
}"
definition "add_size_rel = br fst (\<lambda>(l,n). n= Max (insert 0 V))"
lemma obtain_sorted_carrier'''_refine:
"(obtain_sorted_carrier''', obtain_sorted_carrier'') \<in> \<langle>add_size_rel\<rangle>nres_rel"
unfolding obtain_sorted_carrier'''_def obtain_sorted_carrier''_def
apply (refine_rcg getEdges_refine)
by (auto intro!: RETURN_SPEC_refine simp: quicksort_by_rel_distinct sort_edges_correct
add_size_rel_def in_br_conv max_node_is_Max_V
dest!: distinct_mapI)
lemmas osc_refine = obtain_sorted_carrier'''_refine[FCOMP obtain_sorted_carrier''_refine,
to_foparam, simplified]
definition kruskal3 :: "(nat \<times> int \<times> nat) list nres"
where "kruskal3 \<equiv> do {
(sl,mn) \<leftarrow> obtain_sorted_carrier''';
let initial_union_find = per_init' (mn + 1);
(per, spanning_forest) \<leftarrow> nfoldli sl (\<lambda>_. True)
(\<lambda>ce (uf, T). do {
ASSERT (\<alpha> ce \<in> E);
(a,b) \<leftarrow> endpoints ce;
ASSERT (a \<in> Domain uf \<and> b \<in> Domain uf);
if \<not> per_compare uf a b then
do {
let uf = per_union uf a b;
ASSERT (ce\<notin>set T);
RETURN (uf, T@[ce])
}
else
RETURN (uf,T)
}) (initial_union_find, []);
RETURN spanning_forest
}"
lemma endpoints_spec: "endpoints ce \<le> SPEC (\<lambda>_. True)"
by(rule order.trans[OF endpoints_refine], auto)
lemma kruskal3_subset:
shows "kruskal3 \<le>\<^sub>n SPEC (\<lambda>T. distinct T \<and> set T \<subseteq> superE )"
unfolding kruskal3_def obtain_sorted_carrier'''_def
apply (refine_vcg getEdges_refine[THEN leof_lift] endpoints_spec[THEN leof_lift]
nfoldli_leof_rule[where I="\<lambda>_ _ (_, T). distinct T \<and> set T \<subseteq> superE "])
apply auto
subgoal
by (metis append_self_conv in_set_conv_decomp set_quicksort_by_rel subset_iff)
- subgoal
- by blast
done
definition per_supset_rel :: "('a per \<times> 'a per) set" where
"per_supset_rel
\<equiv> {(p1,p2). p1 \<inter> Domain p2 \<times> Domain p2 = p2 \<and> p1 - (Domain p2 \<times> Domain p2) \<subseteq> Id}"
lemma per_supset_rel_dom: "(p1, p2) \<in> per_supset_rel \<Longrightarrow> Domain p1 \<supseteq> Domain p2"
by (auto simp: per_supset_rel_def)
lemma per_supset_compare:
"(p1, p2) \<in> per_supset_rel \<Longrightarrow> x1\<in>Domain p2 \<Longrightarrow> x2\<in>Domain p2
\<Longrightarrow> per_compare p1 x1 x2 \<longleftrightarrow> per_compare p2 x1 x2"
by (auto simp: per_supset_rel_def)
lemma per_supset_union: "(p1, p2) \<in> per_supset_rel \<Longrightarrow> x1\<in>Domain p2 \<Longrightarrow> x2\<in>Domain p2 \<Longrightarrow>
(per_union p1 x1 x2, per_union p2 x1 x2) \<in> per_supset_rel"
apply (clarsimp simp: per_supset_rel_def per_union_def Domain_unfold )
apply (intro subsetI conjI)
apply blast
apply force
done
lemma per_initN_refine: "(per_init' (Max (insert 0 V) + 1), per_init V) \<in> per_supset_rel"
unfolding per_supset_rel_def per_init'_def per_init_def max_node_def
by (auto simp: less_Suc_eq_le )
theorem kruskal3_refine: "(kruskal3, kruskal2)\<in>\<langle>Id\<rangle>nres_rel"
unfolding kruskal2_def kruskal3_def Let_def
apply (refine_rcg osc_refine[THEN nres_relD] )
supply RELATESI[where R="per_supset_rel::(nat per \<times> _) set", refine_dref_RELATES]
apply refine_dref_type
subgoal by (simp add: add_size_rel_def in_br_conv)
subgoal using per_initN_refine by (simp add: add_size_rel_def in_br_conv)
by (auto simp add: add_size_rel_def in_br_conv per_supset_compare per_supset_union
dest: per_supset_rel_dom
simp del: per_compare_def )
subsubsection \<open>Synthesis of Kruskal by SepRef\<close>
lemma [sepref_import_param]: "(sort_edges,sort_edges)\<in>\<langle>Id\<times>\<^sub>rId\<times>\<^sub>rId\<rangle>list_rel \<rightarrow>\<langle>Id\<times>\<^sub>rId\<times>\<^sub>rId\<rangle>list_rel"
by simp
lemma [sepref_import_param]: "(max_node, max_node) \<in> \<langle>Id\<times>\<^sub>rId\<times>\<^sub>rId\<rangle>list_rel \<rightarrow> nat_rel" by simp
sepref_register "getEdges" :: "(nat \<times> int \<times> nat) list nres"
sepref_register "endpoints" :: "(nat \<times> int \<times> nat) \<Rightarrow> (nat*nat) nres"
declare getEdges_impl [sepref_fr_rules]
declare endpoints_impl [sepref_fr_rules]
schematic_goal kruskal_impl:
"(uncurry0 ?c, uncurry0 kruskal3 ) \<in> (unit_assn)\<^sup>k \<rightarrow>\<^sub>a list_assn (nat_assn \<times>\<^sub>a int_assn \<times>\<^sub>a nat_assn)"
unfolding kruskal3_def obtain_sorted_carrier'''_def
unfolding sort_edges_def[symmetric]
apply (rewrite at "nfoldli _ _ _ (_,rewrite_HOLE)" HOL_list.fold_custom_empty)
by sepref
concrete_definition (in -) kruskal uses Kruskal_Impl.kruskal_impl
prepare_code_thms (in -) kruskal_def
lemmas kruskal_refine = kruskal.refine[OF this_loc]
abbreviation "MSF == minBasis"
abbreviation "SpanningForest == basis"
lemmas SpanningForest_def = basis_def
lemmas MSF_def = minBasis_def
lemmas kruskal3_ref_spec_ = kruskal3_refine[FCOMP kruskal2_refine, FCOMP kruskal1_refine,
FCOMP kruskal0_refine,
FCOMP minWeightBasis_refine]
lemma kruskal3_ref_spec':
"(uncurry0 kruskal3, uncurry0 (SPEC (\<lambda>r. MSF (\<alpha> ` set r)))) \<in> unit_rel \<rightarrow>\<^sub>f \<langle>Id\<rangle>nres_rel"
unfolding fref_def
apply auto
apply(rule nres_relI)
apply(rule order.trans[OF kruskal3_ref_spec_[unfolded fref_def, simplified, THEN nres_relD]])
by (auto simp: conc_fun_def list_set_rel_def in_br_conv dest!: list_relD2)
lemma kruskal3_ref_spec:
"(uncurry0 kruskal3,
uncurry0 (SPEC (\<lambda>r. distinct r \<and> set r \<subseteq> superE \<and> MSF (\<alpha> ` set r))))
\<in> unit_rel \<rightarrow>\<^sub>f \<langle>Id\<rangle>nres_rel"
unfolding fref_def
apply auto
apply(rule nres_relI)
apply simp
using SPEC_rule_conj_leofI2[OF kruskal3_subset kruskal3_ref_spec'
[unfolded fref_def, simplified, THEN nres_relD, simplified]]
by simp
lemma [fcomp_norm_simps]: "list_assn (nat_assn \<times>\<^sub>a int_assn \<times>\<^sub>a nat_assn) = id_assn"
by (auto simp: list_assn_pure_conv)
lemmas kruskal_ref_spec = kruskal_refine[FCOMP kruskal3_ref_spec]
text \<open>The final correctness lemma for Kruskal's algorithm. \<close>
lemma kruskal_correct_forest:
shows "<emp> kruskal getEdges_impl endpoints_impl ()
<\<lambda>r. \<up>( distinct r \<and> set r \<subseteq> superE \<and> MSF (set (map \<alpha> r)))>\<^sub>t"
proof -
show ?thesis
using kruskal_ref_spec[to_hnr]
unfolding hn_refine_def
apply clarsimp
apply (erule cons_post_rule)
by (sep_auto simp: hn_ctxt_def pure_def list_set_rel_def in_br_conv dest: list_relD)
qed
end \<comment> \<open>locale @{text Kruskal_Impl}\<close>
end
\ No newline at end of file
diff --git a/thys/LambdaAuth/Results.thy b/thys/LambdaAuth/Results.thy
--- a/thys/LambdaAuth/Results.thy
+++ b/thys/LambdaAuth/Results.thy
@@ -1,1557 +1,1559 @@
(* Author: Matthias Brun, ETH Zürich, 2019 *)
(* Author: Dmitriy Traytel, ETH Zürich, 2019 *)
section \<open>Formalization of Miller et al.'s~\cite{adsg} Main Results\<close>
(*<*)
theory Results
imports Agreement
begin
(*>*)
lemma judge_imp_agree:
assumes "\<Gamma> \<turnstile> e : \<tau>"
shows "\<Gamma> \<turnstile> e, e, e : \<tau>"
using assms by (induct \<Gamma> e \<tau>) (auto simp: fresh_Pair)
subsection \<open>Lemma 2.1\<close>
lemma lemma2_1:
assumes "\<Gamma> \<turnstile> e, eP, eV : \<tau>"
shows "\<lparr>eP\<rparr> = eV"
using assms by (induct \<Gamma> e eP eV \<tau>) (simp_all add: Abs1_eq)
subsection \<open>Counterexample to Lemma 2.2\<close>
lemma lemma2_2_false:
fixes x :: var
assumes "\<And>\<Gamma> e eP eV \<tau> eP' eV'. \<lbrakk> \<Gamma> \<turnstile> e, eP, eV : \<tau>; \<Gamma> \<turnstile> e, eP', eV' : \<tau> \<rbrakk> \<Longrightarrow> eP = eP' \<and> eV = eV'"
shows False
proof -
have a1: "{$$} \<turnstile> Prj1 (Pair Unit Unit), Prj1 (Pair Unit Unit), Prj1 (Pair Unit Unit) : One"
by fastforce
also have a2: "{$$} \<turnstile> Prj1 (Pair Unit Unit), Prj1 (Pair Unit (Hashed (hash Unit) Unit)), Prj1 (Pair Unit (Hash (hash Unit))) : One"
by fastforce
from a1 a2 have "Prj1 (Pair Unit Unit) = Prj1 (Pair Unit (Hash (hash Unit)))"
by (auto dest: assms)
then show False
by simp
qed
lemma smallstep_ideal_deterministic:
"\<lless>[], t\<ggreater> I\<rightarrow> \<lless>[], u\<ggreater> \<Longrightarrow> \<lless>[], t\<ggreater> I\<rightarrow> \<lless>[], u'\<ggreater> \<Longrightarrow> u = u'"
proof (nominal_induct "[]::proofstream" t I "[]::proofstream" u avoiding: u' rule: smallstep.strong_induct)
case (s_App1 e\<^sub>1 e\<^sub>1' e\<^sub>2)
from s_App1(3) value_no_step[OF s_App1(1)] show ?case
by (auto dest!: s_App1(2))
next
case (s_App2 v\<^sub>1 e\<^sub>2 e\<^sub>2')
from s_App2(4) value_no_step[OF s_App2(2)] value_no_step[OF _ s_App2(1)] show ?case
by (force dest!: s_App2(3))
next
case (s_AppLam v x e)
from s_AppLam(5,1,3) value_no_step[OF _ s_AppLam(2)] show ?case
proof (cases rule: s_App_inv)
case (AppLam y e')
obtain z :: var where "atom z \<sharp> (x, e, y, e')"
by (metis obtain_fresh)
with AppLam s_AppLam(1,3) show ?thesis
by (auto simp: fresh_Pair intro: box_equals[OF _ subst_term_perm[symmetric, of z] subst_term_perm[symmetric, of z]])
qed (auto dest: value_no_step)
next
case (s_AppRec v x e)
from s_AppRec(5,1,3) value_no_step[OF _ s_AppRec(2)] show ?case
proof (cases rule: s_App_inv)
case (AppRec y e')
obtain z :: var where "atom z \<sharp> (x, e, y, e')"
by (metis obtain_fresh)
with AppRec(1-4) AppRec(5)[simplified] s_AppRec(1,3) show ?thesis
- by (auto simp: fresh_Pair AppRec(5)
- intro: box_equals[OF _ subst_term_perm[symmetric, of z] subst_term_perm[symmetric, of z]])
+ apply (auto simp: fresh_Pair AppRec(1))
+ apply (rule box_equals[OF _ subst_term_perm[symmetric, of z] subst_term_perm[symmetric, of z]])
+ using AppRec(1) apply auto
+ done
qed (auto dest: value_no_step)
next
case (s_Let1 x e\<^sub>1 e\<^sub>1' e\<^sub>2)
from s_Let1(1,2,3,8) value_no_step[OF s_Let1(6)] show ?case
by (auto dest: s_Let1(7))
next
case (s_Let2 v x e)
from s_Let2(1,3,5) value_no_step[OF _ s_Let2(2)] show ?case
by force
next
case (s_Inj1 e e')
from s_Inj1(2,3) show ?case
by auto
next
case (s_Inj2 e e')
from s_Inj2(2,3) show ?case
by auto
next
case (s_Case e e' e\<^sub>1 e\<^sub>2)
from s_Case(2,3) value_no_step[OF s_Case(1)] show ?case
by auto
next
case (s_Pair1 e\<^sub>1 e\<^sub>1' e\<^sub>2)
from s_Pair1(2,3) value_no_step[OF s_Pair1(1)] show ?case
by auto
next
case (s_Pair2 v\<^sub>1 e\<^sub>2 e\<^sub>2')
from s_Pair2(3,4) value_no_step[OF _ s_Pair2(1)] value_no_step[OF s_Pair2(2)] show ?case
by force
next
case (s_Prj1 e e')
from s_Prj1(2,3) value_no_step[OF s_Prj1(1)] show ?case
by auto
next
case (s_Prj2 e e')
from s_Prj2(2,3) value_no_step[OF s_Prj2(1)] show ?case
by auto
next
case (s_Unroll e e')
from s_Unroll(2,3) value_no_step[OF s_Unroll(1)] show ?case
by auto
next
case (s_Roll e e')
from s_Roll(2,3) show ?case
by auto
next
case (s_Auth e e')
from s_Auth(2,3) value_no_step[OF s_Auth(1)] show ?case
by auto
next
case (s_Unauth e e')
from s_Unauth(2,3) value_no_step[OF s_Unauth(1)] show ?case
by auto
qed (auto 0 3 dest: value_no_step)
lemma smallsteps_ideal_deterministic:
"\<lless>[], t\<ggreater> I\<rightarrow>i \<lless>[], u\<ggreater> \<Longrightarrow> \<lless>[], t\<ggreater> I\<rightarrow>i \<lless>[], u'\<ggreater> \<Longrightarrow> u = u'"
proof (induct "[]::proofstream" t I i "[]::proofstream" u arbitrary: u' rule: smallsteps.induct)
case (s_Tr e\<^sub>1 i \<pi>\<^sub>2 e\<^sub>2 e\<^sub>3)
from s_Tr(4) show ?case
proof (cases rule: smallsteps.cases)
case _: (s_Tr i \<pi>\<^sub>4 e\<^sub>4)
with s_Tr(1,3) s_Tr(2)[of e\<^sub>4] show ?thesis
using smallstepsI_ps_emptyD(2)[of e\<^sub>1 i \<pi>\<^sub>4 e\<^sub>4] smallstepI_ps_eq[of \<pi>\<^sub>2 e\<^sub>2 "[]" e\<^sub>3]
by (auto intro!: smallstep_ideal_deterministic elim!: smallstepI_ps_emptyD)
qed simp
qed auto
subsection \<open>Lemma 2.3\<close>
lemma lemma2_3:
assumes "\<Gamma> \<turnstile> e, eP, eV : \<tau>"
shows "erase_env \<Gamma> \<turnstile>\<^sub>W e : erase \<tau>"
using assms unfolding erase_env_def
proof (nominal_induct \<Gamma> e eP eV \<tau> rule: agree.strong_induct)
case (a_HashI v vP \<tau> h \<Gamma>)
then show ?case
by (induct \<Gamma>) (auto simp: judge_weak_weakening_2 fmmap_fmupd judge_weak_fresh_env_fresh_term fresh_tyenv_None)
qed (simp_all add: fresh_fmmap_erase_fresh fmmap_fmupd judge_weak_fresh_env_fresh_term)
subsection \<open>Lemma 2.4\<close>
lemma lemma2_4[dest]:
assumes "\<Gamma> \<turnstile> e, eP, eV : \<tau>"
shows "value e \<and> value eP \<and> value eV \<or> \<not> value e \<and> \<not> value eP \<and> \<not> value eV"
using assms by (nominal_induct \<Gamma> e eP eV \<tau> rule: agree.strong_induct) auto
subsection \<open>Lemma 3\<close>
lemma lemma3_general:
fixes \<Gamma> :: tyenv and vs vPs vVs :: tenv
assumes "\<Gamma> \<turnstile> e : \<tau>" "A |\<subseteq>| fmdom \<Gamma>"
and "fmdom vs = A" "fmdom vPs = A" "fmdom vVs = A"
and "\<forall>x. x |\<in>| A \<longrightarrow> (\<exists>\<tau>' v vP h.
\<Gamma> $$ x = Some (AuthT \<tau>') \<and>
vs $$ x = Some v \<and>
vPs $$ x = Some (Hashed h vP) \<and>
vVs $$ x = Some (Hash h) \<and>
{$$} \<turnstile> v, Hashed h vP, Hash h : (AuthT \<tau>'))"
shows "fmdrop_fset A \<Gamma> \<turnstile> psubst_term e vs, psubst_term e vPs, psubst_term e vVs : \<tau>"
using assms
proof (nominal_induct \<Gamma> e \<tau> avoiding: vs vPs vVs A rule: judge.strong_induct)
case (j_Unit \<Gamma>)
then show ?case
by auto
next
case (j_Var \<Gamma> x \<tau>)
with j_Var show ?case
proof (cases "x |\<in>| A")
case True
with j_Var show ?thesis
by (fastforce dest!: spec[of _ x] intro: agree_weakening_env)
next
case False
with j_Var show ?thesis
by (auto simp add: a_Var dest!: fmdomI split: option.splits)
qed
next
case (j_Lam x \<Gamma> \<tau>\<^sub>1 e \<tau>\<^sub>2)
from j_Lam(4) have [simp]: "A |-| {|x|} = A"
by (simp add: fresh_fset_fminus)
from j_Lam(5,8-) have "fmdrop_fset A \<Gamma>(x $$:= \<tau>\<^sub>1) \<turnstile> psubst_term e vs, psubst_term e vPs, psubst_term e vVs : \<tau>\<^sub>2"
by (intro j_Lam(7)[of A vs vPs vVs]) (auto simp: fresh_tyenv_None)
with j_Lam(1-5) show ?case
by (auto simp: fresh_fmdrop_fset)
next
case (j_App \<Gamma> e \<tau>\<^sub>1 \<tau>\<^sub>2 e')
then have "fmdrop_fset A \<Gamma> \<turnstile> psubst_term e vs, psubst_term e vPs, psubst_term e vVs : Fun \<tau>\<^sub>1 \<tau>\<^sub>2"
and "fmdrop_fset A \<Gamma> \<turnstile> psubst_term e' vs, psubst_term e' vPs, psubst_term e' vVs : \<tau>\<^sub>1"
by simp_all
then show ?case
by auto
next
case (j_Let x \<Gamma> e\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 \<tau>\<^sub>2)
from j_Let(4) have [simp]: "A |-| {|x|} = A"
by (simp add: fresh_fset_fminus)
from j_Let(8,11-) have "fmdrop_fset A \<Gamma> \<turnstile> psubst_term e\<^sub>1 vs, psubst_term e\<^sub>1 vPs, psubst_term e\<^sub>1 vVs : \<tau>\<^sub>1"
by simp
moreover from j_Let(4,5,11-) have "fmdrop_fset A \<Gamma>(x $$:= \<tau>\<^sub>1) \<turnstile> psubst_term e\<^sub>2 vs, psubst_term e\<^sub>2 vPs, psubst_term e\<^sub>2 vVs : \<tau>\<^sub>2"
by (intro j_Let(10)) (auto simp: fresh_tyenv_None)
ultimately show ?case using j_Let(1-6)
by (auto simp: fresh_fmdrop_fset fresh_Pair fresh_psubst)
next
case (j_Rec x \<Gamma> y \<tau>\<^sub>1 \<tau>\<^sub>2 e')
from j_Rec(4) have [simp]: "A |-| {|x|} = A"
by (simp add: fresh_fset_fminus)
from j_Rec(9,14-) have "fmdrop_fset A \<Gamma>(x $$:= Fun \<tau>\<^sub>1 \<tau>\<^sub>2) \<turnstile> psubst_term (Lam y e') vs, psubst_term (Lam y e') vPs, psubst_term (Lam y e') vVs : Fun \<tau>\<^sub>1 \<tau>\<^sub>2"
by (intro j_Rec(13)) (auto simp: fresh_tyenv_None)
with j_Rec(1-11) show ?case
by (auto simp: fresh_fmdrop_fset)
next
case (j_Case \<Gamma> e \<tau>\<^sub>1 \<tau>\<^sub>2 e\<^sub>1 \<tau> e\<^sub>2)
then have "fmdrop_fset A \<Gamma> \<turnstile> psubst_term e vs, psubst_term e vPs, psubst_term e vVs : Sum \<tau>\<^sub>1 \<tau>\<^sub>2"
and "fmdrop_fset A \<Gamma> \<turnstile> psubst_term e\<^sub>1 vs, psubst_term e\<^sub>1 vPs, psubst_term e\<^sub>1 vVs : Fun \<tau>\<^sub>1 \<tau>"
and "fmdrop_fset A \<Gamma> \<turnstile> psubst_term e\<^sub>2 vs, psubst_term e\<^sub>2 vPs, psubst_term e\<^sub>2 vVs : Fun \<tau>\<^sub>2 \<tau>"
by simp_all
then show ?case
by auto
next
case (j_Prj1 \<Gamma> e \<tau>\<^sub>1 \<tau>\<^sub>2)
then have "fmdrop_fset A \<Gamma> \<turnstile> psubst_term e vs, psubst_term e vPs, psubst_term e vVs : Prod \<tau>\<^sub>1 \<tau>\<^sub>2"
by simp
then show ?case
by auto
next
case (j_Prj2 \<Gamma> e \<tau>\<^sub>1 \<tau>\<^sub>2)
then have "fmdrop_fset A \<Gamma> \<turnstile> psubst_term e vs, psubst_term e vPs, psubst_term e vVs : Prod \<tau>\<^sub>1 \<tau>\<^sub>2"
by simp
then show ?case
by auto
next
case (j_Roll \<alpha> \<Gamma> e \<tau>)
then have "fmdrop_fset A \<Gamma> \<turnstile> psubst_term e vs, psubst_term e vPs, psubst_term e vVs : subst_type \<tau> (Mu \<alpha> \<tau>) \<alpha>"
by simp
with j_Roll(4,5) show ?case
by (auto simp: fresh_fmdrop_fset)
next
case (j_Unroll \<alpha> \<Gamma> e \<tau>)
then have "fmdrop_fset A \<Gamma> \<turnstile> psubst_term e vs, psubst_term e vPs, psubst_term e vVs : Mu \<alpha> \<tau>"
by simp
with j_Unroll(4,5) show ?case
by (auto simp: fresh_fmdrop_fset)
qed auto
lemmas lemma3 = lemma3_general[where A = "fmdom \<Gamma>" and \<Gamma> = \<Gamma>, simplified] for \<Gamma>
subsection \<open>Lemma 4\<close>
lemma lemma4:
assumes "\<Gamma>(x $$:= \<tau>') \<turnstile> e, eP, eV : \<tau>"
and "{$$} \<turnstile> v, vP, vV : \<tau>'"
and "value v" "value vP" "value vV"
shows "\<Gamma> \<turnstile> e[v / x], eP[vP / x], eV[vV / x] : \<tau>"
using assms
proof (nominal_induct "\<Gamma>(x $$:= \<tau>')" e eP eV \<tau> avoiding: x \<Gamma> rule: agree.strong_induct)
case a_Unit
then show ?case by auto
next
case (a_Var y \<tau>)
then show ?case
proof (induct \<Gamma>)
case fmempty
then show ?case by (metis agree.a_Var fmupd_lookup option.sel subst_term.simps(2))
next
case (fmupd x y \<Gamma>)
then show ?case
using agree_weakening_2 fresh_tyenv_None by auto
qed
next
case (a_Lam y \<tau>\<^sub>1 e eP eV \<tau>\<^sub>2)
from a_Lam(1,2,5,6,7-) show ?case
using agree_empty_fresh by (auto simp: fmupd_reorder_neq)
next
case (a_App v\<^sub>1 v\<^sub>2 v\<^sub>1P v\<^sub>2P v\<^sub>1V v\<^sub>2V \<tau>\<^sub>1 \<tau>\<^sub>2)
from a_App(5-) show ?case
by (auto intro: a_App(2,4))
next
case (a_Let y e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 eP\<^sub>2 eV\<^sub>2 \<tau>\<^sub>2)
then show ?case
using agree_empty_fresh by (auto simp: fmupd_reorder_neq intro!: agree.a_Let[where x=y])
next
case (a_Rec y z \<tau>\<^sub>1 \<tau>\<^sub>2 e eP eV)
from a_Rec(10) have "\<forall>a::var. atom a \<sharp> v" "\<forall>a::var. atom a \<sharp> vP" "\<forall>a::var. atom a \<sharp> vV"
by auto
with a_Rec(1-8,10-) show ?case
using a_Rec(9)[OF fmupd_reorder_neq]
by (auto simp: fmupd_reorder_neq intro!: agree.a_Rec[where x=y])
next
case (a_Case v v\<^sub>1 v\<^sub>2 vP v\<^sub>1P v\<^sub>2P vV v\<^sub>1V v\<^sub>2V \<tau>\<^sub>1 \<tau>\<^sub>2 \<tau>)
from a_Case(7-) show ?case
by (auto intro: a_Case(2,4,6))
next
case (a_HashI v vP \<tau> h)
then have "atom x \<sharp> v" "atom x \<sharp> vP" by auto
with a_HashI show ?case by auto
qed auto
subsection \<open>Lemma 5: Single-Step Correctness\<close>
lemma lemma5:
assumes "{$$} \<turnstile> e, eP, eV : \<tau>"
and "\<lless> [], e \<ggreater> I\<rightarrow> \<lless> [], e' \<ggreater>"
obtains eP' eV' \<pi>
where "{$$} \<turnstile> e', eP', eV' : \<tau>" "\<forall>\<pi>\<^sub>P. \<lless> \<pi>\<^sub>P, eP \<ggreater> P\<rightarrow> \<lless> \<pi>\<^sub>P @ \<pi>, eP' \<ggreater>" "\<forall>\<pi>'. \<lless> \<pi> @ \<pi>', eV \<ggreater> V\<rightarrow> \<lless> \<pi>', eV' \<ggreater>"
proof (atomize_elim, insert assms, nominal_induct "{$$}::tyenv" e eP eV \<tau> avoiding: e' rule: agree.strong_induct)
case (a_App e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau>\<^sub>1 \<tau>\<^sub>2 e\<^sub>2 eP\<^sub>2 eV\<^sub>2)
from a_App(5) show ?case
proof (cases rule: s_App_inv)
case (App1 e\<^sub>1')
with a_App(2) obtain eP\<^sub>1' eV\<^sub>1' \<pi> where *: "{$$} \<turnstile> e\<^sub>1', eP\<^sub>1', eV\<^sub>1' : Fun \<tau>\<^sub>1 \<tau>\<^sub>2"
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, eP\<^sub>1\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, eP\<^sub>1'\<ggreater>" "\<forall>\<pi>'. \<lless>\<pi> @ \<pi>', eV\<^sub>1\<ggreater> V\<rightarrow> \<lless>\<pi>', eV\<^sub>1'\<ggreater>"
by blast
show ?thesis
proof (intro exI conjI)
from * App1 a_App(1,3,5-)
show "{$$} \<turnstile> e', App eP\<^sub>1' eP\<^sub>2, App eV\<^sub>1' eV\<^sub>2 : \<tau>\<^sub>2"
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, App eP\<^sub>1 eP\<^sub>2\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, App eP\<^sub>1' eP\<^sub>2\<ggreater>"
"\<forall>\<pi>'. \<lless>\<pi> @ \<pi>', App eV\<^sub>1 eV\<^sub>2\<ggreater> V\<rightarrow> \<lless>\<pi>', App eV\<^sub>1' eV\<^sub>2\<ggreater>"
by auto
qed
next
case (App2 e\<^sub>2')
with a_App(4) obtain eP\<^sub>2' eV\<^sub>2' \<pi> where *: "{$$} \<turnstile> e\<^sub>2', eP\<^sub>2', eV\<^sub>2' : \<tau>\<^sub>1"
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, eP\<^sub>2\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, eP\<^sub>2'\<ggreater>" "\<forall>\<pi>'. \<lless>\<pi> @ \<pi>', eV\<^sub>2\<ggreater> V\<rightarrow> \<lless>\<pi>', eV\<^sub>2'\<ggreater>"
by blast
show ?thesis
proof (intro exI conjI)
from * App2 a_App(1,3,5-)
show "{$$} \<turnstile> e', App eP\<^sub>1 eP\<^sub>2', App eV\<^sub>1 eV\<^sub>2' : \<tau>\<^sub>2"
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, App eP\<^sub>1 eP\<^sub>2\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, App eP\<^sub>1 eP\<^sub>2'\<ggreater>"
"\<forall>\<pi>'. \<lless>\<pi> @ \<pi>', App eV\<^sub>1 eV\<^sub>2\<ggreater> V\<rightarrow> \<lless>\<pi>', App eV\<^sub>1 eV\<^sub>2'\<ggreater>"
by auto
qed
next
case (AppLam x e)
from a_App(1)[unfolded \<open>e\<^sub>1 = Lam x e\<close>] show ?thesis
proof (cases rule: a_Lam_inv_I[case_names _ Lam])
case (Lam eP eV)
show ?thesis
proof (intro exI conjI)
from Lam a_App(3,5) AppLam show "{$$} \<turnstile> e', eP[eP\<^sub>2 / x], eV[eV\<^sub>2 / x] : \<tau>\<^sub>2"
by (auto intro: lemma4)
from Lam a_App(3,5) AppLam show "\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, App eP\<^sub>1 eP\<^sub>2\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ [], eP[eP\<^sub>2 / x]\<ggreater>"
by (intro allI iffD1[OF smallstepP_ps_prepend[where \<pi> = "[]", simplified]])
(auto simp: fresh_Pair intro!: s_AppLam[where v=eP\<^sub>2])
from Lam a_App(3,5) AppLam show "\<forall>\<pi>'. \<lless>[] @ \<pi>', App eV\<^sub>1 eV\<^sub>2\<ggreater> V\<rightarrow> \<lless>\<pi>', eV[eV\<^sub>2 / x]\<ggreater>"
by (intro allI iffD1[OF smallstepV_ps_append[where \<pi>' = "[]", simplified]])
(auto simp: fresh_Pair intro!: s_AppLam[where v=eV\<^sub>2])
qed
qed simp
next
case (AppRec x e)
from a_App(1)[unfolded \<open>e\<^sub>1 = Rec x e\<close>] show ?thesis
proof (cases rule: a_Rec_inv_I[consumes 1, case_names _ Rec])
case (Rec y e'' eP' eV')
from Rec(5,4) show ?thesis
proof (cases rule: a_Lam_inv_I[consumes 1, case_names _ Lam])
case (Lam eP'' eV'')
show ?thesis
proof (intro conjI exI[of _ "[]"] exI)
let ?eP = "App (Lam y eP''[Rec x (Lam y eP'') / x]) eP\<^sub>2"
let ?eV = "App (Lam y eV''[Rec x (Lam y eV'') / x]) eV\<^sub>2"
from a_App(3) AppRec have [simp]: "value eP\<^sub>2" "atom x \<sharp> eP\<^sub>2" "value eV\<^sub>2" "atom x \<sharp> eV\<^sub>2"
by (auto simp: fresh_Pair)
from Lam a_App(3,5-) AppRec Rec show "{$$} \<turnstile> e', ?eP, ?eV : \<tau>\<^sub>2"
unfolding term.eq_iff Abs1_eq(3)
by (auto simp: fmupd_reorder_neq
intro!: agree.a_App[where \<Gamma>="{$$}"] a_Lam[where x=y] lemma4)
from Lam a_App(3,5-) AppRec Rec show "\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, App eP\<^sub>1 eP\<^sub>2\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ [], ?eP\<ggreater>"
unfolding term.eq_iff Abs1_eq(3)
using s_AppRec[where v=eP\<^sub>2 and x=x and \<pi>="[]" and e="Lam y eP''" and uv=P]
by (intro allI iffD1[OF smallstepP_ps_prepend[of "[]", simplified]])
(auto simp: fresh_Pair simp del: s_AppRec)
from Lam a_App(3,5-) AppRec Rec show "\<forall>\<pi>'. \<lless>[] @ \<pi>', App eV\<^sub>1 eV\<^sub>2\<ggreater> V\<rightarrow> \<lless>\<pi>', ?eV\<ggreater>"
unfolding term.eq_iff Abs1_eq(3)
using s_AppRec[where v=eV\<^sub>2 and x=x and \<pi>="[]" and e="Lam y eV''" and uv=V]
by (intro allI iffD1[OF smallstepV_ps_append[of _ _ "[]", simplified]])
(auto simp: fresh_Pair simp del: s_AppRec)
qed
qed (simp add: fresh_fmap_update)
qed simp
qed
next
case (a_Let x e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 eP\<^sub>2 eV\<^sub>2 \<tau>\<^sub>2)
then have "atom x \<sharp> (e\<^sub>1, [])" by auto
with a_Let(10) show ?case
proof (cases rule: s_Let_inv)
case Let1
show ?thesis
proof (intro conjI exI[of _ "[]"] exI)
from a_Let(6,8) Let1 show "{$$} \<turnstile> e', eP\<^sub>2[eP\<^sub>1 / x], eV\<^sub>2[eV\<^sub>1 / x] : \<tau>\<^sub>2"
by (auto intro: lemma4)
from a_Let(4,6) Let1 show "\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, Let eP\<^sub>1 x eP\<^sub>2\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ [], eP\<^sub>2[eP\<^sub>1 / x]\<ggreater>"
by (intro allI iffD1[OF smallstepP_ps_prepend[of "[]", simplified]] s_Let2) auto
from a_Let(5,6) Let1 show "\<forall>\<pi>'. \<lless>[] @ \<pi>', Let eV\<^sub>1 x eV\<^sub>2\<ggreater> V\<rightarrow> \<lless>\<pi>', eV\<^sub>2[eV\<^sub>1 / x]\<ggreater>"
by (intro allI iffD1[OF smallstepV_ps_append[of _ _ "[]", simplified]] s_Let2) auto
qed
next
case (Let2 e\<^sub>1')
moreover
from Let2 a_Let(7) obtain eP\<^sub>1' eV\<^sub>1' \<pi>
where ih: "{$$} \<turnstile> e\<^sub>1', eP\<^sub>1', eV\<^sub>1' : \<tau>\<^sub>1"
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, eP\<^sub>1\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, eP\<^sub>1'\<ggreater>"
"\<forall>\<pi>'. \<lless>\<pi> @ \<pi>', eV\<^sub>1\<ggreater> V\<rightarrow> \<lless>\<pi>', eV\<^sub>1'\<ggreater>"
by (blast dest: spec[of _ "[]"])
then have [simp]: "atom x \<sharp> ({$$}, e\<^sub>1', eP\<^sub>1', eV\<^sub>1')"
using agree_empty_fresh by auto
from ih a_Let(4) have [simp]: "atom x \<sharp> \<pi>"
using fresh_Nil fresh_append fresh_ps_smallstep_P by blast
from a_Let ih have agree: "{$$} \<turnstile> Let e\<^sub>1' x e\<^sub>2, Let eP\<^sub>1' x eP\<^sub>2, Let eV\<^sub>1' x eV\<^sub>2 : \<tau>\<^sub>2"
by auto
moreover from a_Let(4,5) ih(1) spec[OF ih(2), of "[]", simplified]
have "\<lless>\<pi>', Let eP\<^sub>1 x eP\<^sub>2\<ggreater> P\<rightarrow> \<lless>\<pi>' @ \<pi>, Let eP\<^sub>1' x eP\<^sub>2\<ggreater>" for \<pi>'
by (intro iffD1[OF smallstepP_ps_prepend[of "[]", simplified]] s_Let1) (auto simp: fresh_Pair)
moreover from a_Let(4,5) ih(1) spec[OF ih(3), of "[]", simplified]
have "\<lless>\<pi> @ \<pi>', Let eV\<^sub>1 x eV\<^sub>2\<ggreater> V\<rightarrow> \<lless>\<pi>', Let eV\<^sub>1' x eV\<^sub>2\<ggreater>" for \<pi>'
by (intro iffD1[OF smallstepV_ps_append[of \<pi> _ "[]", simplified]] s_Let1) (auto simp: fresh_Pair)
ultimately show ?thesis
by blast
qed
next
case (a_Case e eP eV \<tau>\<^sub>1 \<tau>\<^sub>2 e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau> e\<^sub>2 eP\<^sub>2 eV\<^sub>2)
from a_Case(7) show ?case
proof (cases rule: s_Case_inv)
case (Case e'')
with a_Case(2)[of e''] obtain eP'' eV'' \<pi> where ih: "{$$} \<turnstile> e'', eP'', eV'' : Syntax.Sum \<tau>\<^sub>1 \<tau>\<^sub>2"
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, eP\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, eP''\<ggreater>" "\<forall>\<pi>'. \<lless>\<pi> @ \<pi>', eV\<ggreater> V\<rightarrow> \<lless>\<pi>', eV''\<ggreater>"
by blast
show ?thesis
proof (intro conjI exI[of _ \<pi>] exI)
from ih(1) a_Case(3,5) Case show "{$$} \<turnstile> e', Case eP'' eP\<^sub>1 eP\<^sub>2, Case eV'' eV\<^sub>1 eV\<^sub>2 : \<tau>"
by auto
from a_Case(5) spec[OF ih(2), of "[]", simplified]
show "\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, Case eP eP\<^sub>1 eP\<^sub>2\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, Case eP'' eP\<^sub>1 eP\<^sub>2\<ggreater>"
by (intro allI iffD1[OF smallstepP_ps_prepend[of "[]", simplified]] s_Case) auto
from a_Case(5) spec[OF ih(3), of "[]", simplified]
show "\<forall>\<pi>'. \<lless>\<pi> @ \<pi>', Case eV eV\<^sub>1 eV\<^sub>2\<ggreater> V\<rightarrow> \<lless>\<pi>', Case eV'' eV\<^sub>1 eV\<^sub>2\<ggreater>"
by (intro allI iffD1[OF smallstepV_ps_append[of _ _ "[]", simplified]] s_Case) auto
qed
next
case (Inj1 v)
from a_Case(1)[unfolded \<open>e = Inj1 v\<close>] show ?thesis
proof (cases rule: a_Inj1_inv_I[consumes 1, case_names Case])
case (Case vP vV)
with a_Case(3,5) Inj1 show ?thesis
proof (intro conjI exI[of _ "[]"] exI)
from Case a_Case(3,5) Inj1 show "{$$} \<turnstile> e', App eP\<^sub>1 vP, App eV\<^sub>1 vV : \<tau>"
by auto
qed auto
qed
next
case (Inj2 v)
from a_Case(1)[unfolded \<open>e = Inj2 v\<close>] show ?thesis
proof (cases rule: a_Inj2_inv_I[consumes 1, case_names Case])
case (Case vP vV)
with a_Case(3,5) Inj2 show ?thesis
proof (intro conjI exI[of _ "[]"] exI)
from Case a_Case(3,5) Inj2 show "{$$} \<turnstile> e', App eP\<^sub>2 vP, App eV\<^sub>2 vV : \<tau>"
by auto
qed auto
qed
qed
next
case (a_Prj1 e eP eV \<tau>\<^sub>1 \<tau>\<^sub>2)
from a_Prj1(3) show ?case
proof (cases rule: s_Prj1_inv)
case (Prj1 e'')
then show ?thesis
by (blast dest!: a_Prj1(2))
next
case (PrjPair1 v\<^sub>2)
from a_Prj1(1)[unfolded \<open>e = Syntax.Pair e' v\<^sub>2\<close>] show ?thesis
proof (cases rule: a_Pair_inv_I[consumes 1, case_names Pair])
case (Pair eP\<^sub>1 eV\<^sub>1 eP\<^sub>2 eV\<^sub>2)
with PrjPair1 show ?thesis
proof (intro conjI exI[of _ "[]"] exI)
show "{$$} \<turnstile> e', eP\<^sub>1, eV\<^sub>1 : \<tau>\<^sub>1"
by (rule Pair)
qed auto
qed
qed
next
case (a_Prj2 e eP eV \<tau>\<^sub>1 \<tau>\<^sub>2)
from a_Prj2(3) show ?case
proof (cases rule: s_Prj2_inv)
case (Prj2 e'')
then show ?thesis
by (blast dest!: a_Prj2(2))
next
case (PrjPair2 v\<^sub>1)
from a_Prj2(1)[unfolded \<open>e = Syntax.Pair v\<^sub>1 e'\<close>] show ?thesis
proof (cases rule: a_Pair_inv_I[consumes 1, case_names Pair])
case (Pair eP\<^sub>1 eV\<^sub>1 eP\<^sub>2 eV\<^sub>2)
with PrjPair2 show ?thesis
proof (intro conjI exI[of _ "[]"] exI)
show "{$$} \<turnstile> e', eP\<^sub>2, eV\<^sub>2 : \<tau>\<^sub>2"
by (rule Pair)
qed auto
qed
qed
next
case (a_Roll \<alpha> e eP eV \<tau>)
from a_Roll(5) show ?case
proof (cases rule: s_Roll_inv)
case (Roll e'')
with a_Roll(4) obtain eP'' eV'' \<pi> where *: "{$$} \<turnstile> e'', eP'', eV'' : subst_type \<tau> (Mu \<alpha> \<tau>) \<alpha>"
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, eP\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, eP''\<ggreater>" "\<forall>\<pi>'. \<lless>\<pi> @ \<pi>', eV\<ggreater> V\<rightarrow> \<lless>\<pi>', eV''\<ggreater>"
by blast
show ?thesis
proof (intro exI conjI)
from * Roll
show "{$$} \<turnstile> e', Roll eP'', Roll eV'' : Mu \<alpha> \<tau>"
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, Roll eP\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, Roll eP''\<ggreater>"
"\<forall>\<pi>'. \<lless>\<pi> @ \<pi>', Roll eV\<ggreater> V\<rightarrow> \<lless>\<pi>', Roll eV''\<ggreater>"
by auto
qed
qed
next
case (a_Unroll \<alpha> e eP eV \<tau>)
from a_Unroll(5) show ?case
proof (cases rule: s_Unroll_inv)
case (Unroll e'')
with a_Unroll(4) obtain eP'' eV'' \<pi> where *: "{$$} \<turnstile> e'', eP'', eV'' : Mu \<alpha> \<tau>"
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, eP\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, eP''\<ggreater>" "\<forall>\<pi>'. \<lless>\<pi> @ \<pi>', eV\<ggreater> V\<rightarrow> \<lless>\<pi>', eV''\<ggreater>"
by blast
show ?thesis
proof (intro exI conjI)
from * Unroll
show "{$$} \<turnstile> e', Unroll eP'', Unroll eV'' : subst_type \<tau> (Mu \<alpha> \<tau>) \<alpha>"
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, Unroll eP\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, Unroll eP''\<ggreater>"
"\<forall>\<pi>'. \<lless>\<pi> @ \<pi>', Unroll eV\<ggreater> V\<rightarrow> \<lless>\<pi>', Unroll eV''\<ggreater>"
by auto
qed
next
case UnrollRoll
with a_Unroll(3)[unfolded \<open>e = Roll e'\<close>] show ?thesis
proof (cases rule: a_Roll_inv_I[case_names Roll])
case (Roll eP' eV')
with UnrollRoll show ?thesis
proof (intro conjI exI[of _ "[]"] exI)
show "{$$} \<turnstile> e', eP', eV' : subst_type \<tau> (Mu \<alpha> \<tau>) \<alpha>" by fact
qed auto
qed
qed
next
case (a_Auth e eP eV \<tau>)
from a_Auth(1) have [simp]: "atom x \<sharp> eP" for x :: var
using agree_empty_fresh by simp
from a_Auth(3) show ?case
proof (cases rule: s_AuthI_inv)
case (Auth e'')
then show ?thesis
by (blast dest!: a_Auth(2))
next
case AuthI
with a_Auth(1) have "value eP" "value eV" by auto
with a_Auth(1) AuthI(2) show ?thesis
proof (intro conjI exI[of _ "[]"] exI)
from a_Auth(1) AuthI(2) \<open>value eP\<close>
show "{$$} \<turnstile> e', Hashed (hash \<lparr>eP\<rparr>) eP, Hash (hash \<lparr>eP\<rparr>) : AuthT \<tau>"
by (auto dest: lemma2_1 simp: fresh_shallow)
qed (auto dest: lemma2_1 simp: fresh_shallow)
qed
next
case (a_Unauth e eP eV \<tau>)
from a_Unauth(1) have eP_closed: "closed eP"
using agree_empty_fresh by simp
from a_Unauth(3) show ?case
proof (cases rule: s_UnauthI_inv)
case (Unauth e')
then show ?thesis
by (blast dest!: a_Unauth(2))
next
case UnauthI
with a_Unauth(1) have "value eP" "value eV" by auto
from a_Unauth(1) show ?thesis
proof (cases rule: a_AuthT_value_inv[case_names _ _ _ Unauth])
case (Unauth vP')
show ?thesis
proof (intro conjI exI[of _ "[\<lparr>vP'\<rparr>]"] exI)
from Unauth(1,2) UnauthI(2) a_Unauth(1)
show "{$$} \<turnstile> e', vP', \<lparr>vP'\<rparr> : \<tau>"
by (auto simp: fresh_shallow)
then have "closed vP'"
by auto
with Unauth(1,2) a_Unauth(1) show
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, Unauth eP\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ [\<lparr>vP'\<rparr>], vP'\<ggreater>"
"\<forall>\<pi>'. \<lless>[\<lparr>vP'\<rparr>] @ \<pi>', Unauth eV\<ggreater> V\<rightarrow> \<lless>\<pi>', \<lparr>vP'\<rparr>\<ggreater>"
by (auto simp: fresh_shallow)
qed
qed (auto simp: \<open>value e\<close> \<open>value eP\<close> \<open>value eV\<close>)
qed
next
case (a_Pair e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 eP\<^sub>2 eV\<^sub>2 \<tau>\<^sub>2)
from a_Pair(5) show ?case
proof (cases rule: s_Pair_inv)
case (Pair1 e\<^sub>1')
with a_Pair(1,3) show ?thesis
by (blast dest!: a_Pair(2))
next
case (Pair2 e\<^sub>2')
with a_Pair(1,3) show ?thesis
by (blast dest!: a_Pair(4))
qed
next
case (a_Inj1 e eP eV \<tau>\<^sub>1 \<tau>\<^sub>2)
from a_Inj1(3) show ?case
proof (cases rule: s_Inj1_inv)
case (Inj1 e')
with a_Inj1(1) show ?thesis
by (blast dest!: a_Inj1(2))
qed
next
case (a_Inj2 e eP eV \<tau>\<^sub>2 \<tau>\<^sub>1)
from a_Inj2(3) show ?case
proof (cases rule: s_Inj2_inv)
case (Inj2 e'')
with a_Inj2(1) show ?thesis
by (blast dest!: a_Inj2(2))
qed
qed (simp, meson value.intros value_no_step)+
subsection \<open>Lemma 6: Single-Step Security\<close>
lemma lemma6:
assumes "{$$} \<turnstile> e, eP, eV : \<tau>"
and "\<lless> \<pi>\<^sub>A, eV \<ggreater> V\<rightarrow> \<lless> \<pi>', eV' \<ggreater>"
obtains e' eP' \<pi>
where "\<lless> [], e \<ggreater> I\<rightarrow> \<lless> [], e' \<ggreater>" "\<forall>\<pi>\<^sub>P. \<lless> \<pi>\<^sub>P, eP \<ggreater> P\<rightarrow> \<lless> \<pi>\<^sub>P @ \<pi>, eP' \<ggreater>"
and "{$$} \<turnstile> e', eP', eV' : \<tau> \<and> \<pi>\<^sub>A = \<pi> @ \<pi>' \<or>
(\<exists>s s'. closed s \<and> closed s' \<and> \<pi> = [s] \<and> \<pi>\<^sub>A = [s'] @ \<pi>' \<and> s \<noteq> s' \<and> hash s = hash s')"
proof (atomize_elim, insert assms, nominal_induct "{$$}::tyenv" e eP eV \<tau> avoiding: \<pi>\<^sub>A \<pi>' eV' rule: agree.strong_induct)
case (a_App e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau>\<^sub>1 \<tau>\<^sub>2 e\<^sub>2 eP\<^sub>2 eV\<^sub>2)
from a_App(5) show ?case
proof (cases rule: s_App_inv)
case (App1 eV\<^sub>1')
with a_App(1,3) show ?thesis
by (blast dest!: a_App(2))
next
case (App2 e\<^sub>2')
with a_App(1,3) show ?thesis
by (blast dest!: a_App(4))
next
case (AppLam x eV'')
from a_App(1)[unfolded \<open>eV\<^sub>1 = Lam x eV''\<close>] show ?thesis
proof (cases rule: a_Lam_inv_V[case_names Lam])
case (Lam e'' eP'')
with a_App(3) AppLam show ?thesis
proof (intro conjI exI[of _ "[]"] exI disjI1)
from Lam a_App(3) AppLam show "{$$} \<turnstile> e''[e\<^sub>2 / x], eP''[eP\<^sub>2 / x], eV' : \<tau>\<^sub>2"
by (auto intro: lemma4)
qed (auto 0 3 simp: fresh_Pair intro!: s_AppLam[where \<pi>="[]"]
intro: iffD1[OF smallstepP_ps_prepend[of "[]" _ "[]", simplified]])
qed
next
case (AppRec x eV'')
from a_App(1)[unfolded \<open>eV\<^sub>1 = Rec x eV''\<close>] show ?thesis
proof (cases rule: a_Rec_inv_V[case_names _ Rec])
case (Rec y e''' eP''' eV''')
with a_App(1,3) AppRec show ?thesis
proof (intro conjI exI[of _ "[]"] exI disjI1)
let ?e = "App (Lam y e'''[Rec x (Lam y e''') / x]) e\<^sub>2"
let ?eP = "App (Lam y eP'''[Rec x (Lam y eP''') / x]) eP\<^sub>2"
from Rec a_App(3) AppRec show "{$$} \<turnstile> ?e, ?eP, eV' : \<tau>\<^sub>2"
by (auto simp del: subst_term.simps(3) intro!: agree.a_App[where \<Gamma>="{$$}"] lemma4)
qed (auto 0 3 simp del: subst_term.simps(3) simp: fresh_Pair intro!: s_AppRec[where \<pi>="[]"]
intro: iffD1[OF smallstepP_ps_prepend[of "[]" _ "[]", simplified]])
qed simp
qed
next
case (a_Let x e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 eP\<^sub>2 eV\<^sub>2 \<tau>\<^sub>2)
then have "atom x \<sharp> (eV\<^sub>1, \<pi>\<^sub>A)" by auto
with a_Let(12) show ?case
proof (cases rule: s_Let_inv)
case Let1
with a_Let(5,6,8,10) show ?thesis
proof (intro conjI exI[of _ "[]"] exI disjI1)
from Let1 a_Let(5,6,8,10) show "{$$} \<turnstile> e\<^sub>2[e\<^sub>1 / x], eP\<^sub>2[eP\<^sub>1 / x], eV' : \<tau>\<^sub>2"
by (auto intro: lemma4)
qed (auto 0 3 intro: iffD1[OF smallstepP_ps_prepend[of "[]" _ "[]", simplified]])
next
case (Let2 eV\<^sub>1')
with a_Let(9)[of \<pi>\<^sub>A \<pi>' eV\<^sub>1'] obtain e\<^sub>1' \<pi> eP\<^sub>1' s s' where
ih: "\<lless>[], e\<^sub>1\<ggreater> I\<rightarrow> \<lless>[], e\<^sub>1'\<ggreater>" "\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, eP\<^sub>1\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, eP\<^sub>1'\<ggreater>"
"{$$} \<turnstile> e\<^sub>1', eP\<^sub>1', eV\<^sub>1' : \<tau>\<^sub>1 \<and> \<pi>\<^sub>A = \<pi> @ \<pi>' \<or>
closed s \<and> closed s' \<and> \<pi> = [s] \<and> \<pi>\<^sub>A = [s'] @ \<pi>' \<and> s \<noteq> s' \<and> hash s = hash s'"
by blast
with a_Let(5,6) have fresh: "atom x \<sharp> e\<^sub>1'" "atom x \<sharp> eP\<^sub>1'"
using fresh_smallstep_I fresh_smallstep_P by blast+
from ih a_Let(2,6) have "atom x \<sharp> \<pi>"
using fresh_append fresh_ps_smallstep_P by blast
with Let2 a_Let(1-8,10,12-) fresh ih show ?thesis
proof (intro conjI exI[of _ "\<pi>"] exI)
from \<open>atom x \<sharp> \<pi>\<close> Let2 a_Let(1-8,10,12-) fresh ih
show "{$$} \<turnstile> Let e\<^sub>1' x e\<^sub>2, Let eP\<^sub>1' x eP\<^sub>2, eV' : \<tau>\<^sub>2 \<and> \<pi>\<^sub>A = \<pi> @ \<pi>' \<or>
(\<exists>s s'. closed s \<and> closed s' \<and> \<pi> = [s] \<and> \<pi>\<^sub>A = [s'] @ \<pi>' \<and> s \<noteq> s' \<and> hash s = hash s')"
by auto
qed (auto dest: spec[of _ "[]"] intro!: iffD1[OF smallstepP_ps_prepend, of "[]", simplified])
qed
next
case (a_Case e eP eV \<tau>\<^sub>1 \<tau>\<^sub>2 e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau> e\<^sub>2 eP\<^sub>2 eV\<^sub>2)
from a_Case(7) show ?case
proof (cases rule: s_Case_inv)
case (Case eV'')
from a_Case(2)[OF Case(2)] show ?thesis
proof (elim exE disjE conjE, goal_cases ok collision)
case (ok e'' \<pi> eP'')
with Case a_Case(3,5) show ?case by blast
next
case (collision e'' \<pi> eP'' s s')
with Case a_Case(3,5) show ?case
proof (intro exI[of _ "[s]"] exI conjI disjI2)
from Case a_Case(3,5) collision show "\<lless>[], Case e e\<^sub>1 e\<^sub>2\<ggreater> I\<rightarrow> \<lless>[], Case e'' e\<^sub>1 e\<^sub>2\<ggreater>"
"\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, Case eP eP\<^sub>1 eP\<^sub>2\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ [s], Case eP'' eP\<^sub>1 eP\<^sub>2\<ggreater>"
by auto
from collision show "closed s" "closed s'" "s \<noteq> s'" "hash s = hash s'" by auto
qed simp
qed
next
case (Inj1 vV)
from a_Case(1)[unfolded \<open>eV = Inj1 vV\<close>] show ?thesis
proof (cases rule: a_Inj1_inv_V[consumes 1, case_names Inj])
case (Inj v' vP')
with Inj1 show ?thesis
proof (intro conjI exI[of _ "[]"] exI disjI1)
from a_Case(3) Inj Inj1 show "{$$} \<turnstile> App e\<^sub>1 v', App eP\<^sub>1 vP', eV' : \<tau>"
by auto
qed auto
qed
next
case (Inj2 vV)
from a_Case(1)[unfolded \<open>eV = Inj2 vV\<close>] show ?thesis
proof (cases rule: a_Inj2_inv_V[consumes 1, case_names Inj])
case (Inj v' vP')
with Inj2 show ?thesis
proof (intro conjI exI[of _ "[]"] exI disjI1)
from a_Case(5) Inj Inj2 show "{$$} \<turnstile> App e\<^sub>2 v', App eP\<^sub>2 vP', eV' : \<tau>"
by auto
qed auto
qed
qed
next
case (a_Prj1 e eP eV \<tau>\<^sub>1 \<tau>\<^sub>2)
from a_Prj1(3) show ?case
proof (cases rule: s_Prj1_inv)
case (Prj1 eV'')
then show ?thesis
by (blast dest!: a_Prj1(2))
next
case (PrjPair1 v\<^sub>2)
with a_Prj1(1) show ?thesis
proof (cases rule: a_Prod_inv[consumes 1, case_names _ _ _ _ Pair])
case (Pair e\<^sub>1 eP\<^sub>1 eV\<^sub>1 e\<^sub>2 eP\<^sub>2 eV\<^sub>2)
with PrjPair1 a_Prj1(1) show ?thesis
proof (intro conjI exI[of _ "[]"] exI disjI1)
from Pair PrjPair1 a_Prj1(1) show "{$$} \<turnstile> e\<^sub>1, eP\<^sub>1, eV' : \<tau>\<^sub>1"
by auto
qed auto
qed simp_all
qed
next
case (a_Prj2 e eP eV \<tau>\<^sub>1 \<tau>\<^sub>2)
from a_Prj2(3) show ?case
proof (cases rule: s_Prj2_inv)
case (Prj2 eV'')
then show ?thesis
by (blast dest!: a_Prj2(2))
next
case (PrjPair2 v\<^sub>2)
with a_Prj2(1) show ?thesis
proof (cases rule: a_Prod_inv[consumes 1, case_names _ _ _ _ Pair])
case (Pair e\<^sub>1 eP\<^sub>1 eV\<^sub>1 e\<^sub>2 eP\<^sub>2 eV\<^sub>2)
with PrjPair2 a_Prj2(1) show ?thesis
proof (intro conjI exI[of _ "[]"] exI disjI1)
from Pair PrjPair2 a_Prj2(1) show "{$$} \<turnstile> e\<^sub>2, eP\<^sub>2, eV' : \<tau>\<^sub>2"
by auto
qed auto
qed simp_all
qed
next
case (a_Roll \<alpha> e eP eV \<tau>)
from a_Roll(7) show ?case
proof (cases rule: s_Roll_inv)
case (Roll eV'')
from a_Roll(6)[OF Roll(2)] obtain e'' \<pi> eP'' where ih:
"\<lless>[], e\<ggreater> I\<rightarrow> \<lless>[], e''\<ggreater>" "\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, eP\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, eP''\<ggreater>"
"{$$} \<turnstile> e'', eP'', eV'' : subst_type \<tau> (Mu \<alpha> \<tau>) \<alpha> \<and> \<pi>\<^sub>A = \<pi> @ \<pi>' \<or>
(\<exists>s s'. closed s \<and> closed s' \<and> \<pi> = [s] \<and> \<pi>\<^sub>A = [s'] @ \<pi>' \<and> s \<noteq> s' \<and> hash s = hash s')"
by blast
with Roll show ?thesis
proof (intro exI[of _ "\<pi>"] exI conjI)
from ih Roll show "{$$} \<turnstile> Roll e'', Roll eP'', eV' : Mu \<alpha> \<tau> \<and> \<pi>\<^sub>A = \<pi> @ \<pi>' \<or>
(\<exists>s s'. closed s \<and> closed s' \<and> \<pi> = [s] \<and> \<pi>\<^sub>A = [s'] @ \<pi>' \<and> s \<noteq> s' \<and> hash s = hash s')"
by auto
qed auto
qed
next
case (a_Unroll \<alpha> e eP eV \<tau>)
from a_Unroll(7) show ?case
proof (cases rule: s_Unroll_inv)
case (Unroll eV'')
from a_Unroll(6)[OF Unroll(2)] obtain e'' \<pi> eP'' where ih:
"\<lless>[], e\<ggreater> I\<rightarrow> \<lless>[], e''\<ggreater>" "\<forall>\<pi>\<^sub>P. \<lless>\<pi>\<^sub>P, eP\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>, eP''\<ggreater>"
"{$$} \<turnstile> e'', eP'', eV'' : Mu \<alpha> \<tau> \<and> \<pi>\<^sub>A = \<pi> @ \<pi>' \<or>
(\<exists>s s'. closed s \<and> closed s' \<and> \<pi> = [s] \<and> \<pi>\<^sub>A = [s'] @ \<pi>' \<and> s \<noteq> s' \<and> hash s = hash s')"
by blast
with Unroll show ?thesis
proof (intro exI[of _ "\<pi>"] exI conjI)
from ih Unroll show "{$$} \<turnstile> Unroll e'', Unroll eP'', eV' : subst_type \<tau> (Mu \<alpha> \<tau>) \<alpha> \<and> \<pi>\<^sub>A = \<pi> @ \<pi>' \<or>
(\<exists>s s'. closed s \<and> closed s' \<and> \<pi> = [s] \<and> \<pi>\<^sub>A = [s'] @ \<pi>' \<and> s \<noteq> s' \<and> hash s = hash s')"
by auto
qed auto
next
case UnrollRoll
with a_Unroll(5) show ?thesis
by fastforce
qed
next
case (a_Auth e eP eV \<tau>)
from a_Auth(1) have [simp]: "atom x \<sharp> eP" for x :: var
using agree_empty_fresh by simp
from a_Auth(3) show ?case
proof (cases rule: s_AuthV_inv)
case (Auth eV'')
from a_Auth(2)[OF Auth(2)] show ?thesis
proof (elim exE disjE conjE, goal_cases ok collision)
case (ok e'' \<pi> eP'')
with Auth show ?case
proof (intro conjI exI[of _ "\<pi>"] exI disjI1)
from ok Auth show "{$$} \<turnstile> Auth e'', Auth eP'', eV' : AuthT \<tau>"
by auto
qed auto
next
case (collision e'' \<pi> eP'' s s')
then show ?case by blast
qed
next
case AuthV
with a_Auth(1) show ?thesis
proof (intro exI[of _ "[]"] exI conjI disjI1)
from a_Auth(1) AuthV show "{$$} \<turnstile> e, Hashed (hash \<lparr>eP\<rparr>) eP, eV' : AuthT \<tau>"
by (auto dest: lemma2_1)
qed (auto simp: fresh_shallow)
qed
next
case (a_Unauth e eP eV \<tau>)
from a_Unauth(3) show ?case
proof (cases rule: s_UnauthV_inv)
case (Unauth e')
then show ?thesis
by (blast dest!: a_Unauth(2))
next
case UnauthV
from a_Unauth(1)[unfolded \<open>eV = Hash (hash eV')\<close>] UnauthV a_Unauth(1) show ?thesis
proof (cases rule: a_AuthT_value_inv[consumes 1, case_names _ _ _ Hashed])
case (Hashed vP')
with UnauthV a_Unauth(1) show ?thesis
proof (intro exI[of _ "[\<lparr>vP'\<rparr>]"] exI conjI)
from Hashed UnauthV a_Unauth(1) show "{$$} \<turnstile> e, vP', eV' : \<tau> \<and> \<pi>\<^sub>A = [\<lparr>vP'\<rparr>] @ \<pi>' \<or>
(\<exists>s s'. closed s \<and> closed s' \<and> [\<lparr>vP'\<rparr>] = [s] \<and> \<pi>\<^sub>A = [s'] @ \<pi>' \<and> s \<noteq> s' \<and> hash s = hash s')"
by (fastforce elim: a_HashI_inv[where \<Gamma>="{$$}"])
qed auto
qed auto
qed
next
case (a_Pair e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 eP\<^sub>2 eV\<^sub>2 \<tau>\<^sub>2)
from a_Pair(5) show ?case
proof (cases rule: s_Pair_inv)
case (Pair1 eV\<^sub>1')
with a_Pair(3) show ?thesis
using a_Pair(2)[of \<pi>\<^sub>A \<pi>' eV\<^sub>1'] by blast
next
case (Pair2 eV\<^sub>2')
with a_Pair(1) show ?thesis
using a_Pair(4)[of \<pi>\<^sub>A \<pi>' eV\<^sub>2'] by blast
qed
next
case (a_Inj1 e eP eV \<tau>\<^sub>1 \<tau>\<^sub>2)
from a_Inj1(3) show ?case
proof (cases rule: s_Inj1_inv)
case (Inj1 eV'')
then show ?thesis
using a_Inj1(2)[of \<pi>\<^sub>A \<pi>' eV''] by blast
qed
next
case (a_Inj2 e eP eV \<tau>\<^sub>2 \<tau>\<^sub>1)
from a_Inj2(3) show ?case
proof (cases rule: s_Inj2_inv)
case (Inj2 eV'')
with a_Inj2(1) show ?thesis
using a_Inj2(2)[of \<pi>\<^sub>A \<pi>' eV''] by blast
qed
qed (simp, meson value.intros value_no_step)+
subsection \<open>Theorem 1: Correctness\<close>
lemma theorem1_correctness:
assumes "{$$} \<turnstile> e, eP, eV : \<tau>"
and "\<lless> [], e \<ggreater> I\<rightarrow>i \<lless> [], e' \<ggreater>"
obtains eP' eV' \<pi>
where "\<lless> [], eP \<ggreater> P\<rightarrow>i \<lless> \<pi>, eP' \<ggreater>"
"\<lless> \<pi>, eV \<ggreater> V\<rightarrow>i \<lless> [], eV' \<ggreater>"
"{$$} \<turnstile> e', eP', eV' : \<tau>"
using assms(2,1)
proof (atomize_elim, induct "[]::proofstream" e I i "[]::proofstream" e' rule: smallsteps.induct)
case (s_Id e)
then show ?case by auto
next
case (s_Tr e\<^sub>1 i \<pi>\<^sub>2 e\<^sub>2 e\<^sub>3)
then have "\<pi>\<^sub>2 = []" by (auto dest: smallstepI_ps_eq)
with s_Tr obtain eP\<^sub>2 \<pi> eV\<^sub>2 where ih:
"\<lless>[], eP\<ggreater> P\<rightarrow>i \<lless>\<pi>, eP\<^sub>2\<ggreater>" "\<lless>\<pi>, eV\<ggreater> V\<rightarrow>i \<lless>[], eV\<^sub>2\<ggreater>" "{$$} \<turnstile> e\<^sub>2, eP\<^sub>2, eV\<^sub>2 : \<tau>"
by (atomize_elim, intro s_Tr(2)) auto
moreover obtain eP\<^sub>3 eV\<^sub>3 \<pi>' where agree: "{$$} \<turnstile> e\<^sub>3, eP\<^sub>3, eV\<^sub>3 : \<tau>"
and "\<lless>\<pi>\<^sub>P, eP\<^sub>2\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>', eP\<^sub>3\<ggreater>" "\<lless>\<pi>' @ \<pi>'', eV\<^sub>2\<ggreater> V\<rightarrow> \<lless>\<pi>'', eV\<^sub>3\<ggreater>"
for \<pi>\<^sub>P \<pi>'' using lemma5[OF ih(3) s_Tr(3)[unfolded \<open>\<pi>\<^sub>2 = []\<close>], of thesis] by blast
ultimately have "\<lless>[], eP\<ggreater> P\<rightarrow>i + 1 \<lless>\<pi> @ \<pi>', eP\<^sub>3\<ggreater>" "\<lless>\<pi> @ \<pi>', eV\<ggreater> V\<rightarrow>i + 1 \<lless>[], eV\<^sub>3\<ggreater>"
by (auto simp: smallstepsV_ps_append[of _ _ _ "[]", simplified, symmetric]
intro!: smallsteps.s_Tr[of "\<pi> @ \<pi>'"])
with agree show ?case by blast
qed
subsection \<open>Counterexamples to Theorem 1: Security\<close>
text \<open>Counterexample using administrative normal form.\<close>
lemma security_false:
assumes agree: "\<And>e eP eV \<tau> \<pi>A i \<pi>' eV'. \<lbrakk> {$$} \<turnstile> e, eP, eV : \<tau>; \<lless> \<pi>A, eV \<ggreater> V\<rightarrow>i \<lless> \<pi>', eV' \<ggreater> \<rbrakk> \<Longrightarrow>
\<exists>e' eP' \<pi> j \<pi>0 s s'. (\<lless> [], e \<ggreater> I\<rightarrow>i \<lless> [], e' \<ggreater> \<and> \<lless> [], eP \<ggreater> P\<rightarrow>i \<lless> \<pi>, eP' \<ggreater> \<and> (\<pi>A = \<pi> @ \<pi>') \<and> {$$} \<turnstile> e', eP', eV' : \<tau>) \<or>
(j \<le> i \<and> \<lless> [], eP \<ggreater> P\<rightarrow>j \<lless> \<pi>0 @ [s], eP' \<ggreater> \<and> (\<pi>A = \<pi>0 @ [s'] @ \<pi>') \<and> s \<noteq> s' \<and> hash s = hash s')"
and collision: "hash (Inj1 Unit) = hash (Inj2 Unit)"
and no_collision_with_Unit: "\<And>t. hash Unit = hash t \<Longrightarrow> t = Unit"
shows False
proof -
define i where "i = Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc 0)))))))"
obtain x y z :: var where fresh: "atom y \<sharp> x" "atom z \<sharp> (x, y)"
by (metis obtain_fresh)
define t where "t = Let (Let (Auth (Inj1 Unit)) y (Unauth (Var y))) x (Let (Let (Auth Unit) z (Unauth (Var z))) y (Var x))"
note fresh_Cons[simp]
have prover: "\<lless> [], t \<ggreater> P\<rightarrow>i \<lless> [Inj1 Unit, Unit], Inj1 Unit \<ggreater>" \<comment> \<open>prover execution\<close>
unfolding i_def t_def Suc_eq_plus1 using fresh
apply -
apply (rule smallsteps.intros)+
apply (rule s_Let1[rotated])
apply (rule s_Let1[rotated])
apply (rule s_AuthP[rotated])
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthP)
apply simp
apply simp
apply simp
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let1[rotated])
apply (rule s_AuthP[rotated])
apply simp
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthP)
apply simp
apply simp
apply simp
apply (rule s_Let2[of Unit y _ "Inj1 Unit", simplified])
apply simp
done
have verifier1: "\<lless> [Inj1 Unit, Unit], t \<ggreater> V\<rightarrow>i \<lless> [], Inj1 Unit \<ggreater>" \<comment> \<open>verifier execution\<close>
unfolding i_def t_def Suc_eq_plus1 using fresh
apply -
apply (rule smallsteps.intros)+
apply (rule s_Let1[rotated])
apply (rule s_Let1[rotated])
apply (rule s_AuthV[rotated])
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthV)
apply simp
apply simp
apply simp
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let1[rotated])
apply (rule s_AuthV[rotated])
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthV)
apply simp
apply simp
apply simp
apply (rule s_Let2[of Unit y _ "Inj1 Unit", simplified])
apply simp
done
have verifier2: "\<lless> [Inj2 Unit, Unit], t \<ggreater> V\<rightarrow>i \<lless> [], Inj2 Unit \<ggreater>" \<comment> \<open>verifier execution with proofstream containing collision\<close>
unfolding i_def t_def Suc_eq_plus1 using fresh
apply -
apply (rule smallsteps.intros)+
apply (rule s_Let1[rotated])
apply (rule s_Let1[rotated])
apply (rule s_AuthV[rotated])
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthV)
apply simp
apply (simp add: collision)
apply simp
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let1[rotated])
apply (rule s_AuthV[rotated])
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthV)
apply simp
apply simp
apply simp
apply (rule s_Let2[of Unit y _ "Inj2 Unit", simplified])
apply simp
done
have judge: "{$$} \<turnstile> t : Sum One One"
unfolding t_def using fresh
by (force simp: fresh_Pair fresh_fmap_update)
have ideal_deterministic: "e = Inj1 Unit" if "\<lless>[], t\<ggreater> I\<rightarrow>i \<lless>[], e\<ggreater>" for e
apply (rule smallsteps_ideal_deterministic[OF that])
unfolding i_def t_def Suc_eq_plus1 using fresh
apply -
apply (rule smallsteps.intros)+
apply (rule s_Let1[rotated])
apply (rule s_Let1[rotated])
apply (rule s_AuthI[rotated])
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthI)
apply simp
apply simp
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let1[rotated])
apply (rule s_AuthI[rotated])
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthI)
apply simp
apply simp
apply (rule s_Let2[of Unit y _ "Inj1 Unit", simplified])
apply simp
done
from agree[OF judge_imp_agree[OF judge] verifier2] collision prover verifier1 show False
proof safe
fix e' eP'
assume agree: "{$$} \<turnstile> e', eP', Inj2 Unit : Sum One One"
assume assm: "\<lless>[], t\<ggreater> I\<rightarrow>i \<lless>[], e'\<ggreater>"
then have "e' = Inj1 Unit"
by (simp add: ideal_deterministic)
with agree show False
by auto
qed (auto dest!: no_collision_with_Unit[OF sym])
qed
text \<open>Alternative, shorter counterexample not in administrative normal form.\<close>
lemma security_false_alt:
assumes agree: "\<And>e eP eV \<tau> \<pi>A i \<pi>' eV'. \<lbrakk> {$$} \<turnstile> e, eP, eV : \<tau>; \<lless> \<pi>A, eV \<ggreater> V\<rightarrow>i \<lless> \<pi>', eV' \<ggreater> \<rbrakk> \<Longrightarrow>
\<exists>e' eP' \<pi> j \<pi>0 s s'. (\<lless> [], e \<ggreater> I\<rightarrow>i \<lless> [], e' \<ggreater> \<and> \<lless> [], eP \<ggreater> P\<rightarrow>i \<lless> \<pi>, eP' \<ggreater> \<and> (\<pi>A = \<pi> @ \<pi>') \<and> {$$} \<turnstile> e', eP', eV' : \<tau>) \<or>
(j \<le> i \<and> \<lless> [], eP \<ggreater> P\<rightarrow>j \<lless> \<pi>0 @ [s], eP' \<ggreater> \<and> (\<pi>A = \<pi>0 @ [s'] @ \<pi>') \<and> s \<noteq> s' \<and> hash s = hash s')"
and collision: "hash (Inj1 Unit) = hash (Inj2 Unit)"
and no_collision_with_Unit: "\<And>t. hash Unit = hash t \<Longrightarrow> t = Unit"
shows False
proof -
define i where "i = Suc (Suc (Suc (Suc (Suc (Suc 0)))))"
obtain x y z :: var where fresh: "atom y \<sharp> x" "atom z \<sharp> (x, y)"
by (metis obtain_fresh)
define t where "t = Let (Unauth (Auth (Inj1 Unit))) x (Let (Unauth (Auth Unit)) y (Var x))"
note fresh_Cons[simp]
have prover: "\<lless> [], t \<ggreater> P\<rightarrow>i \<lless> [Inj1 Unit, Unit], Inj1 Unit \<ggreater>" \<comment> \<open>prover execution\<close>
unfolding i_def t_def Suc_eq_plus1 using fresh
apply -
apply (rule smallsteps.intros)+
apply (rule s_Let1[rotated])
apply (rule s_Unauth)
apply (rule s_AuthP[rotated])
apply simp
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthP)
apply simp
apply simp
apply simp
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Unauth)
apply (rule s_AuthP[rotated])
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthP)
apply simp
apply simp
apply simp
apply (rule s_Let2[of Unit y _ "Inj1 Unit", simplified])
apply simp
done
have verifier1: "\<lless> [Inj1 Unit, Unit], t \<ggreater> V\<rightarrow>i \<lless> [], Inj1 Unit \<ggreater>" \<comment> \<open>verifier execution\<close>
unfolding i_def t_def Suc_eq_plus1 using fresh
apply -
apply (rule smallsteps.intros)+
apply (rule s_Let1[rotated])
apply (rule s_Unauth)
apply (rule s_AuthV[rotated])
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthV)
apply simp
apply simp
apply simp
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Unauth)
apply (rule s_AuthV[rotated])
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthV)
apply simp
apply simp
apply simp
apply (rule s_Let2[of Unit y _ "Inj1 Unit", simplified])
apply simp
done
have verifier2: "\<lless> [Inj2 Unit, Unit], t \<ggreater> V\<rightarrow>i \<lless> [], Inj2 Unit \<ggreater>" \<comment> \<open>verifier execution with proofstream containing collision\<close>
unfolding i_def t_def Suc_eq_plus1 using fresh
apply -
apply (rule smallsteps.intros)+
apply (rule s_Let1[rotated])
apply (rule s_Unauth)
apply (rule s_AuthV[rotated])
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthV)
apply simp
apply (simp add: collision)
apply simp
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Unauth)
apply (rule s_AuthV[rotated])
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthV)
apply simp
apply simp
apply simp
apply (rule s_Let2[of Unit y _ "Inj2 Unit", simplified])
apply simp
done
have judge: "{$$} \<turnstile> t : Sum One One"
unfolding t_def using fresh
by (force simp: fresh_Pair fresh_fmap_update)
have ideal_deterministic: "e = Inj1 Unit" if "\<lless>[], t\<ggreater> I\<rightarrow>i \<lless>[], e\<ggreater>" for e
apply (rule smallsteps_ideal_deterministic[OF that])
unfolding i_def t_def Suc_eq_plus1 using fresh
apply -
apply (rule smallsteps.intros)+
apply (rule s_Let1[rotated])
apply (rule s_Unauth)
apply (rule s_AuthI[rotated])
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthI)
apply simp
apply simp
apply (rule s_Let2)
apply simp
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_Unauth)
apply (rule s_AuthI[rotated])
apply simp
apply simp
apply (rule s_Let1[rotated])
apply (rule s_UnauthI)
apply simp
apply simp
apply (rule s_Let2[of Unit y _ "Inj1 Unit", simplified])
apply simp
done
from agree[OF judge_imp_agree[OF judge] verifier2] collision prover verifier1 show False
proof safe
fix e' eP'
assume agree: "{$$} \<turnstile> e', eP', Inj2 Unit : Sum One One"
assume assm: "\<lless>[], t\<ggreater> I\<rightarrow>i \<lless>[], e'\<ggreater>"
then have "e' = Inj1 Unit"
by (simp add: ideal_deterministic)
with agree show False
by auto
qed (auto dest!: no_collision_with_Unit[OF sym])
qed
subsection \<open>Corrected Theorem 1: Security\<close>
lemma theorem1_security:
assumes "{$$} \<turnstile> e, eP, eV : \<tau>"
and "\<lless> \<pi>\<^sub>A, eV \<ggreater> V\<rightarrow>i \<lless> \<pi>', eV' \<ggreater>"
shows "(\<exists>e' eP' \<pi>. \<lless> [], e \<ggreater> I\<rightarrow>i \<lless> [], e' \<ggreater> \<and> \<lless> [], eP \<ggreater> P\<rightarrow>i \<lless> \<pi>, eP' \<ggreater> \<and> \<pi>\<^sub>A = \<pi> @ \<pi>' \<and> {$$} \<turnstile> e', eP', eV' : \<tau>) \<or>
(\<exists>eP' j \<pi>\<^sub>0 \<pi>\<^sub>0' s s'. j \<le> i \<and> \<lless> [], eP \<ggreater> P\<rightarrow>j \<lless> \<pi>\<^sub>0 @ [s], eP' \<ggreater> \<and> \<pi>\<^sub>A = \<pi>\<^sub>0 @ [s'] @ \<pi>\<^sub>0' @ \<pi>' \<and> s \<noteq> s' \<and> hash s = hash s' \<and> closed s \<and> closed s')"
using assms(2,1) proof (induct \<pi>\<^sub>A eV V i \<pi>' eV' rule: smallsteps.induct)
case (s_Id \<pi> e)
then show ?case by blast
next
case (s_Tr \<pi>\<^sub>1 eV\<^sub>1 i \<pi>\<^sub>2 eV\<^sub>2 \<pi>\<^sub>3 eV\<^sub>3)
then obtain e\<^sub>2 \<pi> eP\<^sub>2 j \<pi>\<^sub>0 \<pi>\<^sub>0' s s'
where "\<lless>[], e\<ggreater> I\<rightarrow>i \<lless>[], e\<^sub>2\<ggreater> \<and> \<lless>[], eP\<ggreater> P\<rightarrow>i \<lless>\<pi>, eP\<^sub>2\<ggreater> \<and> \<pi>\<^sub>1 = \<pi> @ \<pi>\<^sub>2 \<and> {$$} \<turnstile> e\<^sub>2, eP\<^sub>2, eV\<^sub>2 : \<tau> \<or>
j \<le> i \<and> \<lless>[], eP\<ggreater> P\<rightarrow>j \<lless>\<pi>\<^sub>0 @ [s], eP\<^sub>2\<ggreater> \<and> closed s \<and> closed s' \<and> \<pi>\<^sub>1 = \<pi>\<^sub>0 @ [s'] @ \<pi>\<^sub>0' @ \<pi>\<^sub>2 \<and> s \<noteq> s' \<and> hash s = hash s'"
by blast
then show ?case
proof (elim disjE conjE, goal_cases ok collision)
case ok
obtain e\<^sub>3 eP\<^sub>3 \<pi>' where
"\<lless>[], e\<^sub>2\<ggreater> I\<rightarrow> \<lless>[], e\<^sub>3\<ggreater>" "\<lless>\<pi>\<^sub>P, eP\<^sub>2\<ggreater> P\<rightarrow> \<lless>\<pi>\<^sub>P @ \<pi>', eP\<^sub>3\<ggreater>"
"{$$} \<turnstile> e\<^sub>3, eP\<^sub>3, eV\<^sub>3 : \<tau> \<and> \<pi>\<^sub>2 = \<pi>' @ \<pi>\<^sub>3 \<or>
(\<exists>s s'. closed s \<and> closed s' \<and> \<pi>' = [s] \<and> \<pi>\<^sub>2 = [s'] @ \<pi>\<^sub>3 \<and> s \<noteq> s' \<and> hash s = hash s')"
for \<pi>\<^sub>P using lemma6[OF ok(4) s_Tr(3), of thesis] by blast
then show ?case
proof (elim disjE conjE exE, goal_cases ok2 collision)
case ok2
with s_Tr(1,3-) ok show ?case
by auto
next
case (collision s s')
then show ?case
proof (intro disjI2 exI conjI)
from ok collision show "\<lless>[], eP\<ggreater> P\<rightarrow>i + 1 \<lless>\<pi> @ [s], eP\<^sub>3\<ggreater>"
by (elim smallsteps.s_Tr) auto
from ok collision show "\<pi>\<^sub>1 = \<pi> @ [s'] @ [] @ \<pi>\<^sub>3"
by simp
qed simp_all
qed
next
case collision
from s_Tr(3) collision show ?case
proof (elim smallstepV_consumes_proofstream, intro disjI2 exI conjI)
fix \<pi>\<^sub>0''
assume *: "\<pi>\<^sub>2 = \<pi>\<^sub>0'' @ \<pi>\<^sub>3"
from collision * show "\<pi>\<^sub>1 = \<pi>\<^sub>0 @ [s'] @ (\<pi>\<^sub>0' @ \<pi>\<^sub>0'') @ \<pi>\<^sub>3"
by simp
qed simp_all
qed
qed
subsection \<open>Remark 1\<close>
lemma remark1_single:
assumes "{$$} \<turnstile> e, eP, eV : \<tau>"
and "\<lless> \<pi>P, eP \<ggreater> P\<rightarrow> \<lless> \<pi>P @ \<pi>, eP' \<ggreater>"
obtains e' eV' where "{$$} \<turnstile> e', eP', eV' : \<tau> \<and> \<lless> [], e \<ggreater> I\<rightarrow> \<lless> [], e' \<ggreater> \<and> \<lless> \<pi>, eV \<ggreater> V\<rightarrow> \<lless> [], eV' \<ggreater>"
proof (atomize_elim, insert assms, nominal_induct "{$$}::tyenv" e eP eV \<tau> avoiding: \<pi>P \<pi> eP' rule: agree.strong_induct)
case (a_App e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau>\<^sub>1 \<tau>\<^sub>2 e\<^sub>2 eP\<^sub>2 eV\<^sub>2)
from a_App(5) show ?case
proof (cases rule: s_App_inv)
case (App1 eP\<^sub>1')
with a_App(2,3) show ?thesis by blast
next
case (App2 eP\<^sub>2')
with a_App(1,4) show ?thesis by blast
next
case (AppLam x eP)
from a_App(1)[unfolded \<open>eP\<^sub>1 = Lam x eP\<close>] show ?thesis
proof (cases rule: a_Lam_inv_P[case_names Lam])
case (Lam v' vV')
with a_App(3) AppLam show ?thesis
by (auto 0 4 simp: fresh_Pair del: s_AppLam intro!: s_AppLam lemma4)
qed
next
case (AppRec x e)
from a_App(1)[unfolded \<open>eP\<^sub>1 = Rec x e\<close>] show ?thesis
proof (cases rule: a_Rec_inv_P[case_names _ Rec])
case (Rec y e'' eP'' eV'')
show ?thesis
proof (intro exI conjI)
let ?e = "App (Lam y (e''[Rec x (Lam y e'') / x])) e\<^sub>2"
let ?eV = "App (Lam y (eV''[Rec x (Lam y eV'') / x])) eV\<^sub>2"
from a_App(3) Rec AppRec show "{$$} \<turnstile> ?e, eP', ?eV : \<tau>\<^sub>2"
by (auto intro!: agree.a_App[where \<Gamma>="{$$}"] lemma4
simp del: subst_term.simps(3) simp: subst_term.simps(3)[symmetric])
from a_App(3) Rec AppRec show "\<lless>[], App e\<^sub>1 e\<^sub>2\<ggreater> I\<rightarrow> \<lless>[], ?e\<ggreater>"
by (auto intro!: s_AppRec[where v=e\<^sub>2]
simp del: subst_term.simps(3) simp: subst_term.simps(3)[symmetric] fresh_Pair)
from a_App(3) Rec AppRec show "\<lless>\<pi>, App eV\<^sub>1 eV\<^sub>2\<ggreater> V\<rightarrow> \<lless>[], ?eV\<ggreater>"
by (auto intro!: s_AppRec[where v=eV\<^sub>2]
simp del: subst_term.simps(3) simp: subst_term.simps(3)[symmetric] fresh_Pair)
qed
qed simp
qed
next
case (a_Let x e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 eP\<^sub>2 eV\<^sub>2 \<tau>\<^sub>2)
then have "atom x \<sharp> (eP\<^sub>1, \<pi>P)" by auto
with a_Let(12) show ?case
proof (cases rule: s_Let_inv)
case Let1
with a_Let show ?thesis
by (intro exI[where P = "\<lambda>x. \<exists>y. (Q x y)" for Q, OF exI, of _ "e\<^sub>2[e\<^sub>1 / x]" "eV\<^sub>2[eV\<^sub>1 / x]"])
(auto intro!: lemma4)
next
case (Let2 eP\<^sub>1')
with a_Let(9) obtain e\<^sub>1' eV\<^sub>1'
where ih: "{$$} \<turnstile> e\<^sub>1', eP\<^sub>1', eV\<^sub>1' : \<tau>\<^sub>1" "\<lless>[], e\<^sub>1\<ggreater> I\<rightarrow> \<lless>[], e\<^sub>1'\<ggreater>" "\<lless>\<pi>, eV\<^sub>1\<ggreater> V\<rightarrow> \<lless>[], eV\<^sub>1'\<ggreater>"
by blast
from a_Let Let2 have "\<not> value e\<^sub>1" "\<not> value eP\<^sub>1" "\<not> value eV\<^sub>1" by auto
with Let2 a_Let(2,5,7,10) ih show ?thesis
by (intro exI[where P = "\<lambda>x. \<exists>y. (Q x y)" for Q, OF exI, of _ "Let e\<^sub>1' x e\<^sub>2" "Let eV\<^sub>1' x eV\<^sub>2"])
(fastforce simp: fresh_Pair del: agree.a_Let intro!: agree.a_Let)
qed
next
case (a_Case e eP eV \<tau>\<^sub>1 \<tau>\<^sub>2 e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau> e\<^sub>2 eP\<^sub>2 eV\<^sub>2)
from a_Case(7) show ?case
proof (cases rule: s_Case_inv)
case (Case eP'')
with a_Case(2,3,5) show ?thesis by blast
next
case (Inj1 v)
with a_Case(1,3,5) show ?thesis by blast
next
case (Inj2 v)
with a_Case(1,3,5) show ?thesis by blast
qed
next
case (a_Prj1 e eP eV \<tau>\<^sub>1 \<tau>\<^sub>2 \<pi>P \<pi> eP')
from a_Prj1(3) show ?case
proof (cases rule: s_Prj1_inv)
case (Prj1 eP'')
with a_Prj1(2) show ?thesis by blast
next
case (PrjPair1 v\<^sub>2)
with a_Prj1(1) show ?thesis by fastforce
qed
next
case (a_Prj2 v vP vV \<tau>\<^sub>1 \<tau>\<^sub>2)
from a_Prj2(3) show ?case
proof (cases rule: s_Prj2_inv)
case (Prj2 eP'')
with a_Prj2(2) show ?thesis by blast
next
case (PrjPair2 v\<^sub>2)
with a_Prj2(1) show ?thesis by fastforce
qed
next
case (a_Roll \<alpha> e eP eV \<tau>)
from a_Roll(7) show ?case
proof (cases rule: s_Roll_inv)
case (Roll eP'')
with a_Roll(4,5,6) show ?thesis by blast
qed
next
case (a_Unroll \<alpha> e eP eV \<tau>)
from a_Unroll(7) show ?case
proof (cases rule: s_Unroll_inv)
case (Unroll eP'')
with a_Unroll(5,6) show ?thesis by fastforce
next
case UnrollRoll
with a_Unroll(5) show ?thesis by blast
qed
next
case (a_Auth e eP eV \<tau>)
from a_Auth(3) show ?case
proof (cases rule: s_AuthP_inv)
case (Auth eP'')
with a_Auth(3) show ?thesis
by (auto dest!: a_Auth(2)[of \<pi>P \<pi> eP''])
next
case AuthP
with a_Auth(1) show ?thesis
by (auto 0 4 simp: lemma2_1 intro: exI[of _ "Hash (hash \<lparr>eP\<rparr>)"] exI[of _ e])
qed
next
case (a_Unauth e eP eV \<tau>)
from a_Unauth(1) have eP_closed: "closed eP"
using agree_empty_fresh by simp
from a_Unauth(3) show ?case
proof (cases rule: s_UnauthP_inv)
case (Unauth e')
with a_Unauth(2) show ?thesis
by blast
next
case (UnauthP h)
with a_Unauth(1,3) eP_closed show ?thesis
by (force intro: a_AuthT_value_inv[OF a_Unauth(1)] simp: fresh_shallow)
qed
next
case (a_Inj1 e eP eV \<tau>\<^sub>1 \<tau>\<^sub>2)
from a_Inj1(3) show ?case
proof (cases rule: s_Inj1_inv)
case (Inj1 eP'')
with a_Inj1(1,2) show ?thesis by blast
qed
next
case (a_Inj2 e eP eV \<tau>\<^sub>2 \<tau>\<^sub>1)
from a_Inj2(3) show ?case
proof (cases rule: s_Inj2_inv)
case (Inj2 eP'')
with a_Inj2(1,2) show ?thesis by blast
qed
next
case (a_Pair e\<^sub>1 eP\<^sub>1 eV\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 eP\<^sub>2 eV\<^sub>2 \<tau>\<^sub>2)
from a_Pair(5) show ?case
proof (cases rule: s_Pair_inv)
case (Pair1 eP\<^sub>1')
with a_Pair(1,2,3) show ?thesis by blast
next
case (Pair2 eP\<^sub>2')
with a_Pair(1,3,4) show ?thesis by blast
qed
qed (auto dest: value_no_step)
lemma remark1:
assumes "{$$} \<turnstile> e, eP, eV : \<tau>"
and "\<lless> \<pi>\<^sub>P, eP \<ggreater> P\<rightarrow>i \<lless> \<pi>\<^sub>P @ \<pi>, eP' \<ggreater>"
obtains e' eV'
where "{$$} \<turnstile> e', eP', eV' : \<tau>" "\<lless> [], e \<ggreater> I\<rightarrow>i \<lless> [], e' \<ggreater>" "\<lless> \<pi>, eV \<ggreater> V\<rightarrow>i \<lless> [], eV' \<ggreater>"
using assms(2,1)
proof (atomize_elim, nominal_induct \<pi>\<^sub>P eP P i "\<pi>\<^sub>P @ \<pi>" eP' arbitrary: \<pi> rule: smallsteps.strong_induct)
case (s_Id e \<pi>P)
then show ?case
using s_Id_inv by blast
next
case (s_Tr \<pi>\<^sub>1 eP\<^sub>1 i \<pi>\<^sub>2 eP\<^sub>2 eP\<^sub>3)
from s_Tr obtain \<pi>' \<pi>'' where ps: "\<pi>\<^sub>2 = \<pi>\<^sub>1 @ \<pi>'" "\<pi> = \<pi>' @ \<pi>''"
by (force elim: smallstepP_generates_proofstream smallstepsP_generates_proofstream)
with s_Tr obtain e\<^sub>2 eV\<^sub>2 where ih: "{$$} \<turnstile> e\<^sub>2, eP\<^sub>2, eV\<^sub>2 : \<tau>"
"\<lless>[], e\<ggreater> I\<rightarrow>i \<lless>[], e\<^sub>2\<ggreater>" "\<lless>\<pi>', eV\<ggreater> V\<rightarrow>i \<lless>[], eV\<^sub>2\<ggreater>"
by atomize_elim (auto elim: s_Tr(2)[of \<pi>'])
moreover
obtain e\<^sub>3 eV\<^sub>3 where agree: "{$$} \<turnstile> e\<^sub>3, eP\<^sub>3, eV\<^sub>3 : \<tau>" and
"\<lless>[], e\<^sub>2\<ggreater> I\<rightarrow> \<lless>[], e\<^sub>3\<ggreater>" "\<lless>\<pi>'', eV\<^sub>2\<ggreater> V\<rightarrow> \<lless>[], eV\<^sub>3\<ggreater>"
by (rule remark1_single[OF ih(1) iffD2[OF smallstepP_ps_prepend s_Tr(3)[unfolded ps]]]) blast
ultimately have "\<lless>[], e\<ggreater> I\<rightarrow>i + 1 \<lless>[], e\<^sub>3\<ggreater>" "\<lless>\<pi>, eV\<ggreater> V\<rightarrow>i + 1 \<lless>[], eV\<^sub>3\<ggreater>"
by (auto simp: smallstepsV_ps_append[of _ _ _ "[]", simplified, symmetric] ps
intro!: smallsteps.s_Tr[where m=V and \<pi>\<^sub>1="\<pi>' @ \<pi>''" and \<pi>\<^sub>2=\<pi>''])
with agree show ?case
by blast
qed
(*<*)
end
(*>*)
\ No newline at end of file
diff --git a/thys/LambdaAuth/Semantics.thy b/thys/LambdaAuth/Semantics.thy
--- a/thys/LambdaAuth/Semantics.thy
+++ b/thys/LambdaAuth/Semantics.thy
@@ -1,1149 +1,1148 @@
(* Author: Matthias Brun, ETH Zürich, 2019 *)
(* Author: Dmitriy Traytel, ETH Zürich, 2019 *)
section \<open>Semantics of $\lambda\bullet$\<close>
(*<*)
theory Semantics
imports FMap_Lemmas Syntax
begin
(*>*)
text \<open>Avoid clash with substitution notation.\<close>
no_notation inverse_divide (infixl "'/" 70)
text \<open>Help automated provers with smallsteps.\<close>
declare One_nat_def[simp del]
subsection \<open>Equivariant Hash Function\<close>
consts hash_real :: "term \<Rightarrow> hash"
nominal_function map_fixed :: "var \<Rightarrow> var list \<Rightarrow> term \<Rightarrow> term" where
"map_fixed fp l Unit = Unit" |
"map_fixed fp l (Var y) = (if y \<in> set l then (Var y) else (Var fp))" |
"atom y \<sharp> (fp, l) \<Longrightarrow> map_fixed fp l (Lam y t) = (Lam y ((map_fixed fp (y # l) t)))" |
"atom y \<sharp> (fp, l) \<Longrightarrow> map_fixed fp l (Rec y t) = (Rec y ((map_fixed fp (y # l) t)))" |
"map_fixed fp l (Inj1 t) = (Inj1 ((map_fixed fp l t)))" |
"map_fixed fp l (Inj2 t) = (Inj2 ((map_fixed fp l t)))" |
"map_fixed fp l (Pair t1 t2) = (Pair ((map_fixed fp l t1)) ((map_fixed fp l t2)))" |
"map_fixed fp l (Roll t) = (Roll ((map_fixed fp l t)))" |
"atom y \<sharp> (fp, l) \<Longrightarrow> map_fixed fp l (Let t1 y t2) = (Let ((map_fixed fp l t1)) y ((map_fixed fp (y # l) t2)))" |
"map_fixed fp l (App t1 t2) = (App ((map_fixed fp l t1)) ((map_fixed fp l t2)))" |
"map_fixed fp l (Case t1 t2 t3) = (Case ((map_fixed fp l t1)) ((map_fixed fp l t2)) ((map_fixed fp l t3)))" |
"map_fixed fp l (Prj1 t) = (Prj1 ((map_fixed fp l t)))" |
"map_fixed fp l (Prj2 t) = (Prj2 ((map_fixed fp l t)))" |
"map_fixed fp l (Unroll t) = (Unroll ((map_fixed fp l t)))" |
"map_fixed fp l (Auth t) = (Auth ((map_fixed fp l t)))" |
"map_fixed fp l (Unauth t) = (Unauth ((map_fixed fp l t)))" |
"map_fixed fp l (Hash h) = (Hash h)" |
"map_fixed fp l (Hashed h t) = (Hashed h ((map_fixed fp l t)))"
- using [[simproc del: alpha_lst]]
+ using [[simproc del: alpha_lst defined_all]]
subgoal by (simp add: eqvt_def map_fixed_graph_aux_def)
subgoal by (erule map_fixed_graph.induct) (auto simp: fresh_star_def fresh_at_base)
apply clarify
subgoal for P fp l t
by (rule term.strong_exhaust[of t P "(fp, l)"]) (auto simp: fresh_star_def fresh_Pair)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal for y fp l t ya fpa la ta
apply (erule conjE)+
apply (erule Abs_lst1_fcb2'[where c = "(fp, l)"])
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
subgoal for y fp l t ya fpa la ta
apply (erule conjE)+
apply (erule Abs_lst1_fcb2'[where c = "(fp, l)"])
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
subgoal for y fp l t ya fpa la ta
apply (erule conjE)+
apply (erule Abs_lst1_fcb2'[where c = "(fp, l)"])
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
done
nominal_termination (eqvt)
by lexicographic_order
definition hash where
"hash t = hash_real (map_fixed undefined [] t)"
lemma permute_map_list: "p \<bullet> l = map (\<lambda>x. p \<bullet> x) l"
by (induct l) auto
lemma map_fixed_eqvt: "p \<bullet> l = l \<Longrightarrow> map_fixed v l (p \<bullet> t) = map_fixed v l t"
proof (nominal_induct t avoiding: v l p rule: term.strong_induct)
case (Var x)
then show ?case
by (auto simp: term.supp supp_at_base permute_map_list list_eq_iff_nth_eq in_set_conv_nth)
next
case (Lam y e)
from Lam(1,2,3,5) Lam(4)[of p "y # l" v] show ?case
by (auto simp: fresh_perm)
next
case (Rec y e)
from Rec(1,2,3,5) Rec(4)[of p "y # l" v] show ?case
by (auto simp: fresh_perm)
next
case (Let e' y e)
from Let(1,2,3,6) Let(4)[of p l v] Let(5)[of p "y # l" v] show ?case
by (auto simp: fresh_perm)
qed (auto simp: permute_pure)
lemma hash_eqvt[eqvt]: "p \<bullet> hash t = hash (p \<bullet> t)"
unfolding permute_pure hash_def by (auto simp: map_fixed_eqvt)
lemma map_fixed_idle: "{x. \<not> atom x \<sharp> t} \<subseteq> set l \<Longrightarrow> map_fixed v l t = t"
proof (nominal_induct t avoiding: v l rule: term.strong_induct)
case (Var x)
then show ?case
by (auto simp: subset_iff fresh_at_base)
next
case (Lam y e)
from Lam(1,2,4) Lam(3)[of "y # l" v] show ?case
by (auto simp: fresh_Pair Abs1_eq)
next
case (Rec y e)
from Rec(1,2,4) Rec(3)[of "y # l" v] show ?case
by (auto simp: fresh_Pair Abs1_eq)
next
case (Let e' y e)
from Let(1,2,5) Let(3)[of l v] Let(4)[of "y # l" v] show ?case
by (auto simp: fresh_Pair Abs1_eq)
qed (auto simp: subset_iff)
lemma map_fixed_idle_closed:
"closed t \<Longrightarrow> map_fixed undefined [] t = t"
by (rule map_fixed_idle) auto
lemma map_fixed_inj_closed:
"closed t \<Longrightarrow> closed u \<Longrightarrow> map_fixed undefined [] t = map_fixed undefined [] u \<Longrightarrow> t = u"
by (rule box_equals[OF _ map_fixed_idle_closed map_fixed_idle_closed])
lemma hash_eq_hash_real_closed:
assumes "closed t"
shows "hash t = hash_real t"
unfolding hash_def map_fixed_idle_closed[OF assms] ..
subsection \<open>Substitution\<close>
nominal_function subst_term :: "term \<Rightarrow> term \<Rightarrow> var \<Rightarrow> term" ("_[_ '/ _]" [250, 200, 200] 250) where
"Unit[t' / x] = Unit" |
"(Var y)[t' / x] = (if x = y then t' else Var y)" |
"atom y \<sharp> (x, t') \<Longrightarrow> (Lam y t)[t' / x] = Lam y (t[t' / x])" |
"atom y \<sharp> (x, t') \<Longrightarrow> (Rec y t)[t' / x] = Rec y (t[t' / x])" |
"(Inj1 t)[t' / x] = Inj1 (t[t' / x])" |
"(Inj2 t)[t' / x] = Inj2 (t[t' / x])" |
"(Pair t1 t2)[t' / x] = Pair (t1[t' / x]) (t2[t' / x]) " |
"(Roll t)[t' / x] = Roll (t[t' / x])" |
"atom y \<sharp> (x, t') \<Longrightarrow> (Let t1 y t2)[t' / x] = Let (t1[t' / x]) y (t2[t' / x])" |
"(App t1 t2)[t' / x] = App (t1[t' / x]) (t2[t' / x])" |
"(Case t1 t2 t3)[t' / x] = Case (t1[t' / x]) (t2[t' / x]) (t3[t' / x])" |
"(Prj1 t)[t' / x] = Prj1 (t[t' / x])" |
"(Prj2 t)[t' / x] = Prj2 (t[t' / x])" |
"(Unroll t)[t' / x] = Unroll (t[t' / x])" |
"(Auth t)[t' / x] = Auth (t[t' / x])" |
"(Unauth t)[t' / x] = Unauth (t[t' / x])" |
"(Hash h)[t' / x] = Hash h" |
"(Hashed h t)[t' / x] = Hashed h (t[t' / x])"
- using [[simproc del: alpha_lst]]
+ using [[simproc del: alpha_lst defined_all]]
subgoal by (simp add: eqvt_def subst_term_graph_aux_def)
subgoal by (erule subst_term_graph.induct) (auto simp: fresh_star_def fresh_at_base)
apply clarify
subgoal for P a b t
by (rule term.strong_exhaust[of a P "(b, t)"]) (auto simp: fresh_star_def fresh_Pair)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
done
nominal_termination (eqvt)
by lexicographic_order
type_synonym tenv = "(var, term) fmap"
nominal_function psubst_term :: "term \<Rightarrow> tenv \<Rightarrow> term" where
"psubst_term Unit f = Unit" |
"psubst_term (Var y) f = (case f $$ y of Some t \<Rightarrow> t | None \<Rightarrow> Var y)" |
"atom y \<sharp> f \<Longrightarrow> psubst_term (Lam y t) f = Lam y (psubst_term t f)" |
"atom y \<sharp> f \<Longrightarrow> psubst_term (Rec y t) f = Rec y (psubst_term t f)" |
"psubst_term (Inj1 t) f = Inj1 (psubst_term t f)" |
"psubst_term (Inj2 t) f = Inj2 (psubst_term t f)" |
"psubst_term (Pair t1 t2) f = Pair (psubst_term t1 f) (psubst_term t2 f) " |
"psubst_term (Roll t) f = Roll (psubst_term t f)" |
"atom y \<sharp> f \<Longrightarrow> psubst_term (Let t1 y t2) f = Let (psubst_term t1 f) y (psubst_term t2 f)" |
"psubst_term (App t1 t2) f = App (psubst_term t1 f) (psubst_term t2 f)" |
"psubst_term (Case t1 t2 t3) f = Case (psubst_term t1 f) (psubst_term t2 f) (psubst_term t3 f)" |
"psubst_term (Prj1 t) f = Prj1 (psubst_term t f)" |
"psubst_term (Prj2 t) f = Prj2 (psubst_term t f)" |
"psubst_term (Unroll t) f = Unroll (psubst_term t f)" |
"psubst_term (Auth t) f = Auth (psubst_term t f)" |
"psubst_term (Unauth t) f = Unauth (psubst_term t f)" |
"psubst_term (Hash h) f = Hash h" |
"psubst_term (Hashed h t) f = Hashed h (psubst_term t f)"
- using [[simproc del: alpha_lst]]
+ using [[simproc del: alpha_lst defined_all]]
subgoal by (simp add: eqvt_def psubst_term_graph_aux_def)
subgoal by (erule psubst_term_graph.induct) (auto simp: fresh_star_def fresh_at_base)
apply clarify
subgoal for P a b
by (rule term.strong_exhaust[of a P b]) (auto simp: fresh_star_def fresh_Pair)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal by clarify
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
done
nominal_termination (eqvt)
by lexicographic_order
nominal_function subst_type :: "ty \<Rightarrow> ty \<Rightarrow> tvar \<Rightarrow> ty" where
"subst_type One t' x = One" |
"subst_type (Fun t1 t2) t' x = Fun (subst_type t1 t' x) (subst_type t2 t' x)" |
"subst_type (Sum t1 t2) t' x = Sum (subst_type t1 t' x) (subst_type t2 t' x)" |
"subst_type (Prod t1 t2) t' x = Prod (subst_type t1 t' x) (subst_type t2 t' x)" |
"atom y \<sharp> (t', x) \<Longrightarrow> subst_type (Mu y t) t' x = Mu y (subst_type t t' x)" |
"subst_type (Alpha y) t' x = (if y = x then t' else Alpha y)" |
"subst_type (AuthT t) t' x = AuthT (subst_type t t' x)"
- using [[simproc del: alpha_lst]]
+ using [[simproc del: alpha_lst defined_all]]
subgoal by (simp add: eqvt_def subst_type_graph_aux_def)
subgoal by (erule subst_type_graph.induct) (auto simp: fresh_star_def fresh_at_base)
apply clarify
subgoal for P a
by (rule ty.strong_exhaust[of a P]) (auto simp: fresh_star_def)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal
apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff fresh_Pair)
done
done
nominal_termination (eqvt)
by lexicographic_order
lemma fresh_subst_term: "atom x \<sharp> t[t' / x'] \<longleftrightarrow> (x = x' \<or> atom x \<sharp> t) \<and> (atom x' \<sharp> t \<or> atom x \<sharp> t')"
by (nominal_induct t avoiding: t' x x' rule: term.strong_induct) (auto simp add: fresh_at_base)
lemma term_fresh_subst[simp]: "atom x \<sharp> t \<Longrightarrow> atom x \<sharp> s \<Longrightarrow> (atom (x::var)) \<sharp> t[s / y]"
by (nominal_induct t avoiding: s y rule: term.strong_induct) (auto)
lemma term_subst_idle[simp]: "atom y \<sharp> t \<Longrightarrow> t[s / y] = t"
by (nominal_induct t avoiding: s y rule: term.strong_induct) (auto simp: fresh_Pair fresh_at_base)
lemma term_subst_subst: "atom y1 \<noteq> atom y2 \<Longrightarrow> atom y1 \<sharp> s2 \<Longrightarrow> t[s1 / y1][s2 / y2] = t[s2 / y2][s1[s2 / y2] / y1]"
by (nominal_induct t avoiding: y1 y2 s1 s2 rule: term.strong_induct) auto
lemma fresh_psubst:
fixes x :: var
assumes "atom x \<sharp> e" "atom x \<sharp> vs"
shows "atom x \<sharp> psubst_term e vs"
using assms
by (induct e vs rule: psubst_term.induct)
(auto simp: fresh_at_base elim: fresh_fmap_fresh_Some split: option.splits)
lemma fresh_subst_type:
"atom \<alpha> \<sharp> subst_type \<tau> \<tau>' \<alpha>' \<longleftrightarrow> ((\<alpha> = \<alpha>' \<or> atom \<alpha> \<sharp> \<tau>) \<and> (atom \<alpha>' \<sharp> \<tau> \<or> atom \<alpha> \<sharp> \<tau>'))"
by (nominal_induct \<tau> avoiding: \<alpha> \<alpha>' \<tau>' rule: ty.strong_induct) (auto simp add: fresh_at_base)
lemma type_fresh_subst[simp]: "atom x \<sharp> t \<Longrightarrow> atom x \<sharp> s \<Longrightarrow> (atom (x::tvar)) \<sharp> subst_type t s y"
by (nominal_induct t avoiding: s y rule: ty.strong_induct) (auto)
lemma type_subst_idle[simp]: "atom y \<sharp> t \<Longrightarrow> subst_type t s y = t"
by (nominal_induct t avoiding: s y rule: ty.strong_induct) (auto simp: fresh_Pair fresh_at_base)
lemma type_subst_subst: "atom y1 \<noteq> atom y2 \<Longrightarrow> atom y1 \<sharp> s2 \<Longrightarrow>
subst_type (subst_type t s1 y1) s2 y2 = subst_type (subst_type t s2 y2) (subst_type s1 s2 y2) y1"
by (nominal_induct t avoiding: y1 y2 s1 s2 rule: ty.strong_induct) auto
subsection \<open>Weak Typing Judgement\<close>
type_synonym tyenv = "(var, ty) fmap"
inductive judge_weak :: "tyenv \<Rightarrow> term \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile>\<^sub>W _ : _" [150,0,150] 149) where
jw_Unit: "\<Gamma> \<turnstile>\<^sub>W Unit : One" |
jw_Var: "\<lbrakk> \<Gamma> $$ x = Some \<tau> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Var x : \<tau>" |
jw_Lam: "\<lbrakk> atom x \<sharp> \<Gamma>; \<Gamma>(x $$:= \<tau>\<^sub>1) \<turnstile>\<^sub>W e : \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Lam x e : Fun \<tau>\<^sub>1 \<tau>\<^sub>2" |
jw_App: "\<lbrakk> \<Gamma> \<turnstile>\<^sub>W e : Fun \<tau>\<^sub>1 \<tau>\<^sub>2; \<Gamma> \<turnstile>\<^sub>W e' : \<tau>\<^sub>1 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W App e e' : \<tau>\<^sub>2" |
jw_Let: "\<lbrakk> atom x \<sharp> (\<Gamma>, e\<^sub>1); \<Gamma> \<turnstile>\<^sub>W e\<^sub>1 : \<tau>\<^sub>1; \<Gamma>(x $$:= \<tau>\<^sub>1) \<turnstile>\<^sub>W e\<^sub>2 : \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Let e\<^sub>1 x e\<^sub>2 : \<tau>\<^sub>2" |
jw_Rec: "\<lbrakk> atom x \<sharp> \<Gamma>; atom y \<sharp> (\<Gamma>,x); \<Gamma>(x $$:= Fun \<tau>\<^sub>1 \<tau>\<^sub>2) \<turnstile>\<^sub>W Lam y e : Fun \<tau>\<^sub>1 \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Rec x (Lam y e) : Fun \<tau>\<^sub>1 \<tau>\<^sub>2" |
jw_Inj1: "\<lbrakk> \<Gamma> \<turnstile>\<^sub>W e : \<tau>\<^sub>1 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Inj1 e : Sum \<tau>\<^sub>1 \<tau>\<^sub>2" |
jw_Inj2: "\<lbrakk> \<Gamma> \<turnstile>\<^sub>W e : \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Inj2 e : Sum \<tau>\<^sub>1 \<tau>\<^sub>2" |
jw_Case: "\<lbrakk> \<Gamma> \<turnstile>\<^sub>W e : Sum \<tau>\<^sub>1 \<tau>\<^sub>2; \<Gamma> \<turnstile>\<^sub>W e\<^sub>1 : Fun \<tau>\<^sub>1 \<tau>; \<Gamma> \<turnstile>\<^sub>W e\<^sub>2 : Fun \<tau>\<^sub>2 \<tau> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Case e e\<^sub>1 e\<^sub>2 : \<tau>" |
jw_Pair: "\<lbrakk> \<Gamma> \<turnstile>\<^sub>W e\<^sub>1 : \<tau>\<^sub>1; \<Gamma> \<turnstile>\<^sub>W e\<^sub>2 : \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Pair e\<^sub>1 e\<^sub>2 : Prod \<tau>\<^sub>1 \<tau>\<^sub>2" |
jw_Prj1: "\<lbrakk> \<Gamma> \<turnstile>\<^sub>W e : Prod \<tau>\<^sub>1 \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Prj1 e : \<tau>\<^sub>1" |
jw_Prj2: "\<lbrakk> \<Gamma> \<turnstile>\<^sub>W e : Prod \<tau>\<^sub>1 \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Prj2 e : \<tau>\<^sub>2" |
jw_Roll: "\<lbrakk> atom \<alpha> \<sharp> \<Gamma>; \<Gamma> \<turnstile>\<^sub>W e : subst_type \<tau> (Mu \<alpha> \<tau>) \<alpha> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Roll e : Mu \<alpha> \<tau>" |
jw_Unroll: "\<lbrakk> atom \<alpha> \<sharp> \<Gamma>; \<Gamma> \<turnstile>\<^sub>W e : Mu \<alpha> \<tau> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Unroll e : subst_type \<tau> (Mu \<alpha> \<tau>) \<alpha>" |
jw_Auth: "\<lbrakk> \<Gamma> \<turnstile>\<^sub>W e : \<tau> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Auth e : \<tau>" |
jw_Unauth: "\<lbrakk> \<Gamma> \<turnstile>\<^sub>W e : \<tau> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile>\<^sub>W Unauth e : \<tau>"
declare judge_weak.intros[simp]
declare judge_weak.intros[intro]
equivariance judge_weak
nominal_inductive judge_weak
avoids jw_Lam: x
| jw_Rec: x and y
| jw_Let: x
| jw_Roll: \<alpha>
| jw_Unroll: \<alpha>
by (auto simp: fresh_subst_type fresh_Pair)
text \<open>Inversion rules for typing judgment.\<close>
inductive_cases jw_Unit_inv[elim]: "\<Gamma> \<turnstile>\<^sub>W Unit : \<tau>"
inductive_cases jw_Var_inv[elim]: "\<Gamma> \<turnstile>\<^sub>W Var x : \<tau>"
lemma jw_Lam_inv[elim]:
assumes "\<Gamma> \<turnstile>\<^sub>W Lam x e : \<tau>"
and "atom x \<sharp> \<Gamma>"
obtains \<tau>\<^sub>1 \<tau>\<^sub>2 where "\<tau> = Fun \<tau>\<^sub>1 \<tau>\<^sub>2" "(\<Gamma>(x $$:= \<tau>\<^sub>1)) \<turnstile>\<^sub>W e : \<tau>\<^sub>2"
using assms proof (atomize_elim, nominal_induct \<Gamma> "Lam x e" \<tau> avoiding: x e rule: judge_weak.strong_induct)
case (jw_Lam x \<Gamma> \<tau>\<^sub>1 t \<tau>\<^sub>2 y u)
then show ?case
by (auto simp: perm_supp_eq elim!:
iffD1[OF Abs_lst1_fcb2'[where f = "\<lambda>x t (\<Gamma>, \<tau>\<^sub>1, \<tau>\<^sub>2). (\<Gamma>(x $$:= \<tau>\<^sub>1)) \<turnstile>\<^sub>W t : \<tau>\<^sub>2"
and c = "(\<Gamma>, \<tau>\<^sub>1, \<tau>\<^sub>2)" and a = x and b = y and x = t and y = u, unfolded prod.case],
rotated -1])
qed
lemma swap_permute_swap: "atom x \<sharp> \<pi> \<Longrightarrow> atom y \<sharp> \<pi> \<Longrightarrow> (x \<leftrightarrow> y) \<bullet> \<pi> \<bullet> (x \<leftrightarrow> y) \<bullet> t = \<pi> \<bullet> t"
by (subst permute_eqvt) (auto simp: flip_fresh_fresh)
lemma jw_Rec_inv[elim]:
assumes "\<Gamma> \<turnstile>\<^sub>W Rec x t : \<tau>"
and "atom x \<sharp> \<Gamma>"
obtains y e \<tau>\<^sub>1 \<tau>\<^sub>2 where "atom y \<sharp> (\<Gamma>,x)" "t = Lam y e" "\<tau> = Fun \<tau>\<^sub>1 \<tau>\<^sub>2" "\<Gamma>(x $$:= Fun \<tau>\<^sub>1 \<tau>\<^sub>2) \<turnstile>\<^sub>W Lam y e : Fun \<tau>\<^sub>1 \<tau>\<^sub>2"
using [[simproc del: alpha_lst]] assms
proof (atomize_elim, nominal_induct \<Gamma> "Rec x t" \<tau> avoiding: x t rule: judge_weak.strong_induct)
case (jw_Rec x \<Gamma> y \<tau>\<^sub>1 \<tau>\<^sub>2 e z t)
then show ?case
proof (nominal_induct t avoiding: y x z rule: term.strong_induct)
case (Lam y' e')
show ?case
proof (intro exI conjI)
from Lam.prems show "atom y \<sharp> (\<Gamma>, z)" by simp
from Lam.hyps(1-3) Lam.prems show "Lam y' e' = Lam y ((y' \<leftrightarrow> y) \<bullet> e')"
by (subst term.eq_iff(3), intro Abs_lst_eq_flipI) (simp add: fresh_at_base)
from Lam.hyps(1-3) Lam.prems show "\<Gamma>(z $$:= Fun \<tau>\<^sub>1 \<tau>\<^sub>2) \<turnstile>\<^sub>W Lam y ((y' \<leftrightarrow> y) \<bullet> e') : Fun \<tau>\<^sub>1 \<tau>\<^sub>2"
by (elim judge_weak.eqvt[of "\<Gamma>(x $$:= Fun \<tau>\<^sub>1 \<tau>\<^sub>2)" "Lam y e" "Fun \<tau>\<^sub>1 \<tau>\<^sub>2" "(x \<leftrightarrow> z)", elim_format])
(simp add: perm_supp_eq Abs1_eq_iff fresh_at_base swap_permute_swap fresh_perm flip_commute)
qed simp
qed (simp_all add: Abs1_eq_iff)
qed
inductive_cases jw_Inj1_inv[elim]: "\<Gamma> \<turnstile>\<^sub>W Inj1 e : \<tau>"
inductive_cases jw_Inj2_inv[elim]: "\<Gamma> \<turnstile>\<^sub>W Inj2 e : \<tau>"
inductive_cases jw_Pair_inv[elim]: "\<Gamma> \<turnstile>\<^sub>W Pair e\<^sub>1 e\<^sub>2 : \<tau>"
lemma jw_Let_inv[elim]:
assumes "\<Gamma> \<turnstile>\<^sub>W Let e\<^sub>1 x e\<^sub>2 : \<tau>\<^sub>2"
and "atom x \<sharp> (e\<^sub>1, \<Gamma>)"
obtains \<tau>\<^sub>1 where "\<Gamma> \<turnstile>\<^sub>W e\<^sub>1 : \<tau>\<^sub>1" "\<Gamma>(x $$:= \<tau>\<^sub>1) \<turnstile>\<^sub>W e\<^sub>2 : \<tau>\<^sub>2"
using assms proof (atomize_elim, nominal_induct \<Gamma> "Let e\<^sub>1 x e\<^sub>2" \<tau>\<^sub>2 avoiding: e\<^sub>1 x e\<^sub>2 rule: judge_weak.strong_induct)
case (jw_Let x \<Gamma> e\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 \<tau>\<^sub>2 x' e\<^sub>2')
then show ?case
by (auto simp: fresh_Pair perm_supp_eq elim!:
iffD1[OF Abs_lst1_fcb2'[where f = "\<lambda>x t (\<Gamma>, \<tau>\<^sub>1, \<tau>\<^sub>2). \<Gamma>(x $$:= \<tau>\<^sub>1) \<turnstile>\<^sub>W t : \<tau>\<^sub>2"
and c = "(\<Gamma>, \<tau>\<^sub>1, \<tau>\<^sub>2)" and a = x and b = x' and x = e\<^sub>2 and y = e\<^sub>2', unfolded prod.case],
rotated -1])
qed
inductive_cases jw_Prj1_inv[elim]: "\<Gamma> \<turnstile>\<^sub>W Prj1 e : \<tau>\<^sub>1"
inductive_cases jw_Prj2_inv[elim]: "\<Gamma> \<turnstile>\<^sub>W Prj2 e : \<tau>\<^sub>2"
inductive_cases jw_App_inv[elim]: "\<Gamma> \<turnstile>\<^sub>W App e e' : \<tau>\<^sub>2"
inductive_cases jw_Case_inv[elim]: "\<Gamma> \<turnstile>\<^sub>W Case e e\<^sub>1 e\<^sub>2 : \<tau>"
inductive_cases jw_Auth_inv[elim]: "\<Gamma> \<turnstile>\<^sub>W Auth e : \<tau>"
inductive_cases jw_Unauth_inv[elim]: "\<Gamma> \<turnstile>\<^sub>W Unauth e : \<tau>"
lemma subst_type_perm_eq:
assumes "atom b \<sharp> t"
shows "subst_type t (Mu a t) a = subst_type ((a \<leftrightarrow> b) \<bullet> t) (Mu b ((a \<leftrightarrow> b) \<bullet> t)) b"
using assms proof -
have f: "atom a \<sharp> subst_type t (Mu a t) a" by (rule iffD2[OF fresh_subst_type]) simp
have "atom b \<sharp> subst_type t (Mu a t) a" by (auto simp: assms)
with f have "subst_type t (Mu a t) a = (a \<leftrightarrow> b) \<bullet> subst_type t (Mu a t) a"
by (simp add: flip_fresh_fresh)
then show "subst_type t (Mu a t) a = subst_type ((a \<leftrightarrow> b) \<bullet> t) (Mu b ((a \<leftrightarrow> b) \<bullet> t)) b"
by simp
qed
lemma jw_Roll_inv[elim]:
assumes "\<Gamma> \<turnstile>\<^sub>W Roll e : \<tau>"
and "atom \<alpha> \<sharp> (\<Gamma>, \<tau>)"
obtains \<tau>' where "\<tau> = Mu \<alpha> \<tau>'" "\<Gamma> \<turnstile>\<^sub>W e : subst_type \<tau>' (Mu \<alpha> \<tau>') \<alpha>"
using assms [[simproc del: alpha_lst]]
proof (atomize_elim, nominal_induct \<Gamma> "Roll e" \<tau> avoiding: e \<alpha> rule: judge_weak.strong_induct)
case (jw_Roll \<alpha> \<Gamma> e \<tau> \<alpha>')
then show ?case
by (auto simp: perm_supp_eq fresh_Pair fresh_at_base subst_type.eqvt
intro!: exI[of _ "(\<alpha> \<leftrightarrow> \<alpha>') \<bullet> \<tau>"] Abs_lst_eq_flipI dest: judge_weak.eqvt[of _ _ _ "(\<alpha> \<leftrightarrow> \<alpha>')"])
qed
lemma jw_Unroll_inv[elim]:
assumes "\<Gamma> \<turnstile>\<^sub>W Unroll e : \<tau>"
and "atom \<alpha> \<sharp> (\<Gamma>, \<tau>)"
obtains \<tau>' where "\<tau> = subst_type \<tau>' (Mu \<alpha> \<tau>') \<alpha>" "\<Gamma> \<turnstile>\<^sub>W e : Mu \<alpha> \<tau>'"
using assms proof (atomize_elim, nominal_induct \<Gamma> "Unroll e" \<tau> avoiding: e \<alpha> rule: judge_weak.strong_induct)
case (jw_Unroll \<alpha> \<Gamma> e \<tau> \<alpha>')
then show ?case
by (auto simp: perm_supp_eq fresh_Pair subst_type_perm_eq fresh_subst_type
intro!: exI[of _ "(\<alpha> \<leftrightarrow> \<alpha>') \<bullet> \<tau>"] dest: judge_weak.eqvt[of _ _ _ "(\<alpha> \<leftrightarrow> \<alpha>')"])
qed
text \<open>Additional inversion rules based on type rather than term.\<close>
inductive_cases jw_Prod_inv[elim]: "{$$} \<turnstile>\<^sub>W e : Prod \<tau>\<^sub>1 \<tau>\<^sub>2"
inductive_cases jw_Sum_inv[elim]: "{$$} \<turnstile>\<^sub>W e : Sum \<tau>\<^sub>1 \<tau>\<^sub>2"
lemma jw_Fun_inv[elim]:
assumes "{$$} \<turnstile>\<^sub>W v : Fun \<tau>\<^sub>1 \<tau>\<^sub>2" "value v"
obtains e x where "v = Lam x e \<or> v = Rec x e" "atom x \<sharp> (c::term)"
using assms [[simproc del: alpha_lst]]
proof (atomize_elim, nominal_induct "{$$} :: tyenv" v "Fun \<tau>\<^sub>1 \<tau>\<^sub>2" avoiding: \<tau>\<^sub>1 \<tau>\<^sub>2 rule: judge_weak.strong_induct)
case (jw_Lam x \<tau>\<^sub>1 e \<tau>\<^sub>2)
then obtain x' where "atom (x'::var) \<sharp> (c, e)" using finite_supp obtain_fresh' by blast
then have "[[atom x]]lst. e = [[atom x']]lst. (x \<leftrightarrow> x') \<bullet> e \<and> atom x' \<sharp> c"
by (simp add: Abs_lst_eq_flipI fresh_Pair)
then show ?case
by auto
next
case (jw_Rec x y \<tau>\<^sub>1 \<tau>\<^sub>2 e')
obtain x' where "atom (x'::var) \<sharp> (c, Lam y e')" using finite_supp obtain_fresh by blast
then have "[[atom x]]lst. Lam y e' = [[atom x']]lst. (x \<leftrightarrow> x') \<bullet> (Lam y e') \<and> atom x' \<sharp> c"
using Abs_lst_eq_flipI fresh_Pair by blast
then show ?case
by auto
qed simp_all
lemma jw_Mu_inv[elim]:
assumes "{$$} \<turnstile>\<^sub>W v : Mu \<alpha> \<tau>" "value v"
obtains v' where "v = Roll v'"
using assms by (atomize_elim, nominal_induct "{$$} :: tyenv" v "Mu \<alpha> \<tau>" rule: judge_weak.strong_induct) simp_all
subsection \<open>Erasure of Authenticated Types\<close>
nominal_function erase :: "ty \<Rightarrow> ty" where
"erase One = One" |
"erase (Fun \<tau>\<^sub>1 \<tau>\<^sub>2) = Fun (erase \<tau>\<^sub>1) (erase \<tau>\<^sub>2)" |
"erase (Sum \<tau>\<^sub>1 \<tau>\<^sub>2) = Sum (erase \<tau>\<^sub>1) (erase \<tau>\<^sub>2)" |
"erase (Prod \<tau>\<^sub>1 \<tau>\<^sub>2) = Prod (erase \<tau>\<^sub>1) (erase \<tau>\<^sub>2)" |
"erase (Mu \<alpha> \<tau>) = Mu \<alpha> (erase \<tau>)" |
"erase (Alpha \<alpha>) = Alpha \<alpha>" |
"erase (AuthT \<tau>) = erase \<tau>"
using [[simproc del: alpha_lst]]
subgoal by (simp add: eqvt_def erase_graph_aux_def)
subgoal by (erule erase_graph.induct) (auto simp: fresh_star_def fresh_at_base)
subgoal for P x
by (rule ty.strong_exhaust[of x P x]) (auto simp: fresh_star_def)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
done
nominal_termination (eqvt)
by lexicographic_order
lemma fresh_erase_fresh:
assumes "atom x \<sharp> \<tau>"
shows "atom x \<sharp> erase \<tau>"
using assms by (induct \<tau> rule: ty.induct) auto
lemma fresh_fmmap_erase_fresh:
assumes "atom x \<sharp> \<Gamma>"
shows "atom x \<sharp> fmmap erase \<Gamma>"
using assms by transfer simp
lemma erase_subst_type_shift[simp]:
"erase (subst_type \<tau> \<tau>' \<alpha>) = subst_type (erase \<tau>) (erase \<tau>') \<alpha>"
by (induct \<tau> \<tau>' \<alpha> rule: subst_type.induct) (auto simp: fresh_Pair fresh_erase_fresh)
definition erase_env :: "tyenv \<Rightarrow> tyenv" where
"erase_env = fmmap erase"
subsection \<open>Strong Typing Judgement\<close>
inductive judge :: "tyenv \<Rightarrow> term \<Rightarrow> ty \<Rightarrow> bool" ("_ \<turnstile> _ : _" [150,0,150] 149) where
j_Unit: "\<Gamma> \<turnstile> Unit : One" |
j_Var: "\<lbrakk> \<Gamma> $$ x = Some \<tau> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Var x : \<tau>" |
j_Lam: "\<lbrakk> atom x \<sharp> \<Gamma>; \<Gamma>(x $$:= \<tau>\<^sub>1) \<turnstile> e : \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Lam x e : Fun \<tau>\<^sub>1 \<tau>\<^sub>2" |
j_App: "\<lbrakk> \<Gamma> \<turnstile> e : Fun \<tau>\<^sub>1 \<tau>\<^sub>2; \<Gamma> \<turnstile> e' : \<tau>\<^sub>1 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> App e e' : \<tau>\<^sub>2" |
j_Let: "\<lbrakk> atom x \<sharp> (\<Gamma>, e\<^sub>1); \<Gamma> \<turnstile> e\<^sub>1 : \<tau>\<^sub>1; \<Gamma>(x $$:= \<tau>\<^sub>1) \<turnstile> e\<^sub>2 : \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Let e\<^sub>1 x e\<^sub>2 : \<tau>\<^sub>2" |
j_Rec: "\<lbrakk> atom x \<sharp> \<Gamma>; atom y \<sharp> (\<Gamma>,x); \<Gamma>(x $$:= Fun \<tau>\<^sub>1 \<tau>\<^sub>2) \<turnstile> Lam y e' : Fun \<tau>\<^sub>1 \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Rec x (Lam y e') : Fun \<tau>\<^sub>1 \<tau>\<^sub>2" |
j_Inj1: "\<lbrakk> \<Gamma> \<turnstile> e : \<tau>\<^sub>1 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Inj1 e : Sum \<tau>\<^sub>1 \<tau>\<^sub>2" |
j_Inj2: "\<lbrakk> \<Gamma> \<turnstile> e : \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Inj2 e : Sum \<tau>\<^sub>1 \<tau>\<^sub>2" |
j_Case: "\<lbrakk> \<Gamma> \<turnstile> e : Sum \<tau>\<^sub>1 \<tau>\<^sub>2; \<Gamma> \<turnstile> e\<^sub>1 : Fun \<tau>\<^sub>1 \<tau>; \<Gamma> \<turnstile> e\<^sub>2 : Fun \<tau>\<^sub>2 \<tau> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Case e e\<^sub>1 e\<^sub>2 : \<tau>" |
j_Pair: "\<lbrakk> \<Gamma> \<turnstile> e\<^sub>1 : \<tau>\<^sub>1; \<Gamma> \<turnstile> e\<^sub>2 : \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Pair e\<^sub>1 e\<^sub>2 : Prod \<tau>\<^sub>1 \<tau>\<^sub>2" |
j_Prj1: "\<lbrakk> \<Gamma> \<turnstile> e : Prod \<tau>\<^sub>1 \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Prj1 e : \<tau>\<^sub>1" |
j_Prj2: "\<lbrakk> \<Gamma> \<turnstile> e : Prod \<tau>\<^sub>1 \<tau>\<^sub>2 \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Prj2 e : \<tau>\<^sub>2" |
j_Roll: "\<lbrakk> atom \<alpha> \<sharp> \<Gamma>; \<Gamma> \<turnstile> e : subst_type \<tau> (Mu \<alpha> \<tau>) \<alpha> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Roll e : Mu \<alpha> \<tau>" |
j_Unroll: "\<lbrakk> atom \<alpha> \<sharp> \<Gamma>; \<Gamma> \<turnstile> e : Mu \<alpha> \<tau> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Unroll e : subst_type \<tau> (Mu \<alpha> \<tau>) \<alpha>" |
j_Auth: "\<lbrakk> \<Gamma> \<turnstile> e : \<tau> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Auth e : AuthT \<tau>" |
j_Unauth: "\<lbrakk> \<Gamma> \<turnstile> e : AuthT \<tau> \<rbrakk>
\<Longrightarrow> \<Gamma> \<turnstile> Unauth e : \<tau>"
declare judge.intros[intro]
equivariance judge
nominal_inductive judge
avoids j_Lam: x
| j_Rec: x and y
| j_Let: x
| j_Roll: \<alpha>
| j_Unroll: \<alpha>
by (auto simp: fresh_subst_type fresh_Pair)
lemma judge_imp_judge_weak:
assumes "\<Gamma> \<turnstile> e : \<tau>"
shows "erase_env \<Gamma> \<turnstile>\<^sub>W e : erase \<tau>"
using assms unfolding erase_env_def
by (induct \<Gamma> e \<tau> rule: judge.induct) (simp_all add: fresh_Pair fresh_fmmap_erase_fresh fmmap_fmupd)
subsection \<open>Shallow Projection\<close>
nominal_function shallow :: "term \<Rightarrow> term" ("\<lparr>_\<rparr>") where
"\<lparr>Unit\<rparr> = Unit" |
"\<lparr>Var v\<rparr> = Var v" |
"\<lparr>Lam x e\<rparr> = Lam x \<lparr>e\<rparr>" |
"\<lparr>Rec x e\<rparr> = Rec x \<lparr>e\<rparr>" |
"\<lparr>Inj1 e\<rparr> = Inj1 \<lparr>e\<rparr>" |
"\<lparr>Inj2 e\<rparr> = Inj2 \<lparr>e\<rparr>" |
"\<lparr>Pair e\<^sub>1 e\<^sub>2\<rparr> = Pair \<lparr>e\<^sub>1\<rparr> \<lparr>e\<^sub>2\<rparr>" |
"\<lparr>Roll e\<rparr> = Roll \<lparr>e\<rparr>" |
"\<lparr>Let e\<^sub>1 x e\<^sub>2\<rparr> = Let \<lparr>e\<^sub>1\<rparr> x \<lparr>e\<^sub>2\<rparr>" |
"\<lparr>App e\<^sub>1 e\<^sub>2\<rparr> = App \<lparr>e\<^sub>1\<rparr> \<lparr>e\<^sub>2\<rparr>" |
"\<lparr>Case e e\<^sub>1 e\<^sub>2\<rparr> = Case \<lparr>e\<rparr> \<lparr>e\<^sub>1\<rparr> \<lparr>e\<^sub>2\<rparr>" |
"\<lparr>Prj1 e\<rparr> = Prj1 \<lparr>e\<rparr>" |
"\<lparr>Prj2 e\<rparr> = Prj2 \<lparr>e\<rparr>" |
"\<lparr>Unroll e\<rparr> = Unroll \<lparr>e\<rparr>" |
"\<lparr>Auth e\<rparr> = Auth \<lparr>e\<rparr>" |
"\<lparr>Unauth e\<rparr> = Unauth \<lparr>e\<rparr>" |
\<comment> \<open>No rule is defined for Hash, but: "[..] preserving that structure in every case but that of <h, v> [..]"\<close>
"\<lparr>Hash h\<rparr> = Hash h" |
"\<lparr>Hashed h e\<rparr> = Hash h"
using [[simproc del: alpha_lst]]
subgoal by (simp add: eqvt_def shallow_graph_aux_def)
subgoal by (erule shallow_graph.induct) (auto simp: fresh_star_def fresh_at_base)
subgoal for P a
by (rule term.strong_exhaust[of a P a]) (auto simp: fresh_star_def)
apply (simp_all add: fresh_star_def fresh_at_base)
subgoal
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
subgoal
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
subgoal
- apply (erule conjE)
apply (erule Abs_lst1_fcb2')
apply (simp_all add: eqvt_at_def)
apply (simp_all add: perm_supp_eq Abs_fresh_iff)
done
done
nominal_termination (eqvt)
by lexicographic_order
lemma fresh_shallow: "atom x \<sharp> e \<Longrightarrow> atom x \<sharp> \<lparr>e\<rparr>"
by (induct e rule: term.induct) auto
subsection \<open>Small-step Semantics\<close>
datatype mode = I | P | V \<comment> \<open>Ideal, Prover and Verifier modes\<close>
instantiation mode :: pure
begin
definition permute_mode :: "perm \<Rightarrow> mode \<Rightarrow> mode" where
"permute_mode \<pi> h = h"
instance proof qed (auto simp: permute_mode_def)
end
type_synonym proofstream = "term list"
inductive smallstep :: "proofstream \<Rightarrow> term \<Rightarrow> mode \<Rightarrow> proofstream \<Rightarrow> term \<Rightarrow> bool" ("\<lless>_, _\<ggreater> _\<rightarrow> \<lless>_, _\<ggreater>") where
s_App1: "\<lbrakk> \<lless> \<pi>, e\<^sub>1 \<ggreater> m\<rightarrow> \<lless> \<pi>', e\<^sub>1' \<ggreater> \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, App e\<^sub>1 e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', App e\<^sub>1' e\<^sub>2 \<ggreater>" |
s_App2: "\<lbrakk> value v\<^sub>1; \<lless> \<pi>, e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', e\<^sub>2' \<ggreater> \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, App v\<^sub>1 e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', App v\<^sub>1 e\<^sub>2' \<ggreater>" |
s_AppLam: "\<lbrakk> value v; atom x \<sharp> (v,\<pi>) \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, App (Lam x e) v \<ggreater> _\<rightarrow> \<lless> \<pi>, e[v / x] \<ggreater>" |
s_AppRec: "\<lbrakk> value v; atom x \<sharp> (v,\<pi>) \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, App (Rec x e) v \<ggreater> _\<rightarrow> \<lless> \<pi>, App (e[(Rec x e) / x]) v \<ggreater>" |
s_Let1: "\<lbrakk> atom x \<sharp> (e\<^sub>1,e\<^sub>1',\<pi>,\<pi>'); \<lless> \<pi>, e\<^sub>1 \<ggreater> m\<rightarrow> \<lless> \<pi>', e\<^sub>1' \<ggreater> \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Let e\<^sub>1 x e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', Let e\<^sub>1' x e\<^sub>2 \<ggreater>" |
s_Let2: "\<lbrakk> value v; atom x \<sharp> (v,\<pi>) \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Let v x e \<ggreater> _\<rightarrow> \<lless> \<pi>, e[v / x] \<ggreater>" |
s_Inj1: "\<lbrakk> \<lless> \<pi>, e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater> \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Inj1 e \<ggreater> m\<rightarrow> \<lless> \<pi>', Inj1 e' \<ggreater>" |
s_Inj2: "\<lbrakk> \<lless> \<pi>, e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater> \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Inj2 e \<ggreater> m\<rightarrow> \<lless> \<pi>', Inj2 e' \<ggreater>" |
s_Case: "\<lbrakk> \<lless> \<pi>, e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater> \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Case e e\<^sub>1 e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', Case e' e\<^sub>1 e\<^sub>2 \<ggreater>" |
\<comment> \<open>Case rules are different from paper to account for recursive functions.\<close>
s_CaseInj1: "\<lbrakk> value v \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Case (Inj1 v) e\<^sub>1 e\<^sub>2 \<ggreater> _\<rightarrow> \<lless> \<pi>, App e\<^sub>1 v \<ggreater>" |
s_CaseInj2: "\<lbrakk> value v \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Case (Inj2 v) e\<^sub>1 e\<^sub>2 \<ggreater> _\<rightarrow> \<lless> \<pi>, App e\<^sub>2 v \<ggreater>" |
s_Pair1: "\<lbrakk> \<lless> \<pi>, e\<^sub>1 \<ggreater> m\<rightarrow> \<lless> \<pi>', e\<^sub>1' \<ggreater> \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Pair e\<^sub>1 e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', Pair e\<^sub>1' e\<^sub>2 \<ggreater>" |
s_Pair2: "\<lbrakk> value v\<^sub>1; \<lless> \<pi>, e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', e\<^sub>2' \<ggreater> \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Pair v\<^sub>1 e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', Pair v\<^sub>1 e\<^sub>2' \<ggreater>" |
s_Prj1: "\<lbrakk> \<lless> \<pi>, e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater> \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Prj1 e \<ggreater> m\<rightarrow> \<lless> \<pi>', Prj1 e' \<ggreater>" |
s_Prj2: "\<lbrakk> \<lless> \<pi>, e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater> \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Prj2 e \<ggreater> m\<rightarrow> \<lless> \<pi>', Prj2 e' \<ggreater>" |
s_PrjPair1: "\<lbrakk> value v\<^sub>1; value v\<^sub>2 \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Prj1 (Pair v\<^sub>1 v\<^sub>2) \<ggreater> _\<rightarrow> \<lless> \<pi>, v\<^sub>1 \<ggreater>" |
s_PrjPair2: "\<lbrakk> value v\<^sub>1; value v\<^sub>2 \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Prj2 (Pair v\<^sub>1 v\<^sub>2) \<ggreater> _\<rightarrow> \<lless> \<pi>, v\<^sub>2 \<ggreater>" |
s_Unroll: "\<lless> \<pi>, e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>
\<Longrightarrow> \<lless> \<pi>, Unroll e \<ggreater> m\<rightarrow> \<lless> \<pi>', Unroll e' \<ggreater>" |
s_Roll: "\<lless> \<pi>, e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>
\<Longrightarrow> \<lless> \<pi>, Roll e \<ggreater> m\<rightarrow> \<lless> \<pi>', Roll e' \<ggreater>" |
s_UnrollRoll:"\<lbrakk> value v \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Unroll (Roll v) \<ggreater> _\<rightarrow> \<lless> \<pi>, v \<ggreater>" |
\<comment> \<open>Mode-specific rules\<close>
s_Auth: "\<lless> \<pi>, e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>
\<Longrightarrow> \<lless> \<pi>, Auth e \<ggreater> m\<rightarrow> \<lless> \<pi>', Auth e' \<ggreater>" |
s_Unauth: "\<lless> \<pi>, e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>
\<Longrightarrow> \<lless> \<pi>, Unauth e \<ggreater> m\<rightarrow> \<lless> \<pi>', Unauth e' \<ggreater>" |
s_AuthI: "\<lbrakk> value v \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Auth v \<ggreater> I\<rightarrow> \<lless> \<pi>, v \<ggreater>" |
s_UnauthI: "\<lbrakk> value v \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Unauth v \<ggreater> I\<rightarrow> \<lless> \<pi>, v \<ggreater>" |
s_AuthP: "\<lbrakk> closed \<lparr>v\<rparr>; value v \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Auth v \<ggreater> P\<rightarrow> \<lless> \<pi>, Hashed (hash \<lparr>v\<rparr>) v \<ggreater>" |
s_UnauthP: "\<lbrakk> value v \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Unauth (Hashed h v) \<ggreater> P\<rightarrow> \<lless> \<pi> @ [\<lparr>v\<rparr>], v \<ggreater>" |
s_AuthV: "\<lbrakk> closed v; value v \<rbrakk>
\<Longrightarrow> \<lless> \<pi>, Auth v \<ggreater> V\<rightarrow> \<lless> \<pi>, Hash (hash v) \<ggreater>" |
s_UnauthV: "\<lbrakk> closed s\<^sub>0; hash s\<^sub>0 = h \<rbrakk>
\<Longrightarrow> \<lless> s\<^sub>0#\<pi>, Unauth (Hash h) \<ggreater> V\<rightarrow> \<lless> \<pi>, s\<^sub>0 \<ggreater>"
declare smallstep.intros[simp]
declare smallstep.intros[intro]
equivariance smallstep
nominal_inductive smallstep
avoids s_AppLam: x
| s_AppRec: x
| s_Let1: x
| s_Let2: x
by (auto simp add: fresh_Pair fresh_subst_term)
inductive smallsteps :: "proofstream \<Rightarrow> term \<Rightarrow> mode \<Rightarrow> nat \<Rightarrow> proofstream \<Rightarrow> term \<Rightarrow> bool" ("\<lless>_, _\<ggreater> _\<rightarrow>_ \<lless>_, _\<ggreater>") where
s_Id: "\<lless> \<pi>, e \<ggreater> _\<rightarrow>0 \<lless> \<pi>, e \<ggreater>" |
s_Tr: "\<lbrakk> \<lless> \<pi>\<^sub>1, e\<^sub>1 \<ggreater> m\<rightarrow>i \<lless> \<pi>\<^sub>2, e\<^sub>2 \<ggreater>; \<lless> \<pi>\<^sub>2, e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>\<^sub>3, e\<^sub>3 \<ggreater> \<rbrakk>
\<Longrightarrow> \<lless> \<pi>\<^sub>1, e\<^sub>1 \<ggreater> m\<rightarrow>(i+1) \<lless> \<pi>\<^sub>3, e\<^sub>3 \<ggreater>"
declare smallsteps.intros[simp]
declare smallsteps.intros[intro]
equivariance smallsteps
nominal_inductive smallsteps .
lemma steps_1_step[simp]: "\<lless> \<pi>, e \<ggreater> m\<rightarrow>1 \<lless> \<pi>', e' \<ggreater> = \<lless> \<pi>, e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>" (is "?L \<longleftrightarrow> ?R")
proof
assume ?L
then show ?R
proof (induct \<pi> e m "1::nat" \<pi>' e' rule: smallsteps.induct)
case (s_Tr \<pi>\<^sub>1 e\<^sub>1 m i \<pi>\<^sub>2 e\<^sub>2 \<pi>\<^sub>3 e\<^sub>3)
then show ?case
by (induct \<pi>\<^sub>1 e\<^sub>1 m i \<pi>\<^sub>2 e\<^sub>2 rule: smallsteps.induct) auto
qed simp
qed (auto intro: s_Tr[where i=0, simplified])
text \<open>Inversion rules for smallstep(s) predicates.\<close>
lemma value_no_step[intro]:
assumes "\<lless> \<pi>\<^sub>1, v \<ggreater> m\<rightarrow> \<lless> \<pi>\<^sub>2, t \<ggreater>" "value v"
shows "False"
using assms by (induct \<pi>\<^sub>1 v m \<pi>\<^sub>2 t rule: smallstep.induct) auto
lemma subst_term_perm:
assumes "atom x' \<sharp> (x, e)"
shows "e[v / x] = ((x \<leftrightarrow> x') \<bullet> e)[v / x']"
using assms [[simproc del: alpha_lst]]
by (nominal_induct e avoiding: x x' v rule: term.strong_induct)
(auto simp: fresh_Pair fresh_at_base(2) permute_hash_def)
inductive_cases s_Unit_inv[elim]: "\<lless> \<pi>\<^sub>1, Unit \<ggreater> m\<rightarrow> \<lless> \<pi>\<^sub>2, v \<ggreater>"
inductive_cases s_App_inv[consumes 1, case_names App1 App2 AppLam AppRec, elim]: "\<lless> \<pi>, App v\<^sub>1 v\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', e \<ggreater>"
lemma s_Let_inv':
assumes "\<lless> \<pi>, Let e\<^sub>1 x e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
and "atom x \<sharp> (e\<^sub>1,\<pi>)"
obtains e\<^sub>1' where "(e' = e\<^sub>2[e\<^sub>1 / x] \<and> value e\<^sub>1 \<and> \<pi> = \<pi>') \<or> (\<lless> \<pi>, e\<^sub>1 \<ggreater> m\<rightarrow> \<lless> \<pi>', e\<^sub>1' \<ggreater> \<and> e' = Let e\<^sub>1' x e\<^sub>2 \<and> \<not> value e\<^sub>1)"
using assms [[simproc del: alpha_lst]]
by (atomize_elim, induct \<pi> "Let e\<^sub>1 x e\<^sub>2" m \<pi>' e' rule: smallstep.induct)
(auto simp: fresh_Pair fresh_subst_term perm_supp_eq elim: Abs_lst1_fcb2')
lemma s_Let_inv[consumes 2, case_names Let1 Let2, elim]:
assumes "\<lless> \<pi>, Let e\<^sub>1 x e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
and "atom x \<sharp> (e\<^sub>1,\<pi>)"
and "e' = e\<^sub>2[e\<^sub>1 / x] \<and> value e\<^sub>1 \<and> \<pi> = \<pi>' \<Longrightarrow> Q"
and "\<And>e\<^sub>1'. \<lless> \<pi>, e\<^sub>1 \<ggreater> m\<rightarrow> \<lless> \<pi>', e\<^sub>1' \<ggreater> \<and> e' = Let e\<^sub>1' x e\<^sub>2 \<and> \<not> value e\<^sub>1 \<Longrightarrow> Q"
shows "Q"
using assms by (auto elim: s_Let_inv')
inductive_cases s_Case_inv[consumes 1, case_names Case Inj1 Inj2, elim]:
"\<lless> \<pi>, Case e e\<^sub>1 e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_Prj1_inv[consumes 1, case_names Prj1 PrjPair1, elim]:
"\<lless> \<pi>, Prj1 e \<ggreater> m\<rightarrow> \<lless> \<pi>', v \<ggreater>"
inductive_cases s_Prj2_inv[consumes 1, case_names Prj2 PrjPair2, elim]:
"\<lless> \<pi>, Prj2 e \<ggreater> m\<rightarrow> \<lless> \<pi>', v \<ggreater>"
inductive_cases s_Pair_inv[consumes 1, case_names Pair1 Pair2, elim]:
"\<lless> \<pi>, Pair e\<^sub>1 e\<^sub>2 \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_Inj1_inv[consumes 1, case_names Inj1, elim]:
"\<lless> \<pi>, Inj1 e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_Inj2_inv[consumes 1, case_names Inj2, elim]:
"\<lless> \<pi>, Inj2 e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_Roll_inv[consumes 1, case_names Roll, elim]:
"\<lless> \<pi>, Roll e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_Unroll_inv[consumes 1, case_names Unroll UnrollRoll, elim]:
"\<lless> \<pi>, Unroll e \<ggreater> m\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_AuthI_inv[consumes 1, case_names Auth AuthI, elim]:
"\<lless> \<pi>, Auth e \<ggreater> I\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_UnauthI_inv[consumes 1, case_names Unauth UnauthI, elim]:
"\<lless> \<pi>, Unauth e \<ggreater> I\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_AuthP_inv[consumes 1, case_names Auth AuthP, elim]:
"\<lless> \<pi>, Auth e \<ggreater> P\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_UnauthP_inv[consumes 1, case_names Unauth UnauthP, elim]:
"\<lless> \<pi>, Unauth e \<ggreater> P\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_AuthV_inv[consumes 1, case_names Auth AuthV, elim]:
"\<lless> \<pi>, Auth e \<ggreater> V\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_UnauthV_inv[consumes 1, case_names Unauth UnauthV, elim]:
"\<lless> \<pi>, Unauth e \<ggreater> V\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
inductive_cases s_Id_inv[elim]: "\<lless> \<pi>\<^sub>1, e\<^sub>1 \<ggreater> m\<rightarrow>0 \<lless> \<pi>\<^sub>2, e\<^sub>2 \<ggreater>"
inductive_cases s_Tr_inv[elim]: "\<lless> \<pi>\<^sub>1, e\<^sub>1 \<ggreater> m\<rightarrow>i \<lless> \<pi>\<^sub>3, e\<^sub>3 \<ggreater>"
text \<open>Freshness with smallstep.\<close>
lemma fresh_smallstep_I:
fixes x :: var
assumes "\<lless> \<pi>, e \<ggreater> I\<rightarrow> \<lless> \<pi>', e' \<ggreater>" "atom x \<sharp> e"
shows "atom x \<sharp> e'"
using assms by (induct \<pi> e I \<pi>' e' rule: smallstep.induct) (auto simp: fresh_subst_term)
lemma fresh_smallstep_P:
fixes x :: var
assumes "\<lless> \<pi>, e \<ggreater> P\<rightarrow> \<lless> \<pi>', e' \<ggreater>" "atom x \<sharp> e"
shows "atom x \<sharp> e'"
using assms by (induct \<pi> e P \<pi>' e' rule: smallstep.induct) (auto simp: fresh_subst_term)
lemma fresh_smallsteps_I:
fixes x :: var
assumes "\<lless> \<pi>, e \<ggreater> I\<rightarrow>i \<lless> \<pi>', e' \<ggreater>" "atom x \<sharp> e"
shows "atom x \<sharp> e'"
using assms by (induct \<pi> e I i \<pi>' e' rule: smallsteps.induct) (simp_all add: fresh_smallstep_I)
lemma fresh_ps_smallstep_P:
fixes x :: var
assumes "\<lless> \<pi>, e \<ggreater> P\<rightarrow> \<lless> \<pi>', e' \<ggreater>" "atom x \<sharp> e" "atom x \<sharp> \<pi>"
shows "atom x \<sharp> \<pi>'"
using assms proof (induct \<pi> e P \<pi>' e' rule: smallstep.induct)
case (s_UnauthP v \<pi> h)
then show ?case
by (simp add: fresh_Cons fresh_append fresh_shallow)
qed auto
text \<open>Proofstream lemmas.\<close>
lemma smallstepI_ps_eq:
assumes "\<lless> \<pi>, e \<ggreater> I\<rightarrow> \<lless> \<pi>', e' \<ggreater>"
shows "\<pi> = \<pi>'"
using assms by (induct \<pi> e I \<pi>' e' rule: smallstep.induct) auto
lemma smallstepI_ps_emptyD:
"\<lless>\<pi>, e\<ggreater> I\<rightarrow> \<lless>[], e'\<ggreater> \<Longrightarrow> \<lless>[], e\<ggreater> I\<rightarrow> \<lless>[], e'\<ggreater>"
"\<lless>[], e\<ggreater> I\<rightarrow> \<lless>\<pi>, e'\<ggreater> \<Longrightarrow> \<lless>[], e\<ggreater> I\<rightarrow> \<lless>[], e'\<ggreater>"
using smallstepI_ps_eq by force+
lemma smallstepsI_ps_eq:
assumes "\<lless>\<pi>, e\<ggreater> I\<rightarrow>i \<lless>\<pi>', e'\<ggreater>"
shows "\<pi> = \<pi>'"
using assms by (induct \<pi> e I i \<pi>' e' rule: smallsteps.induct) (auto dest: smallstepI_ps_eq)
lemma smallstepsI_ps_emptyD:
"\<lless>\<pi>, e\<ggreater> I\<rightarrow>i \<lless>[], e'\<ggreater> \<Longrightarrow> \<lless>[], e\<ggreater> I\<rightarrow>i \<lless>[], e'\<ggreater>"
"\<lless>[], e\<ggreater> I\<rightarrow>i \<lless>\<pi>, e'\<ggreater> \<Longrightarrow> \<lless>[], e\<ggreater> I\<rightarrow>i \<lless>[], e'\<ggreater>"
using smallstepsI_ps_eq by force+
lemma smallstepV_consumes_proofstream:
assumes "\<lless> \<pi>\<^sub>1, eV \<ggreater> V\<rightarrow> \<lless> \<pi>\<^sub>2, eV' \<ggreater>"
obtains \<pi> where "\<pi>\<^sub>1 = \<pi> @ \<pi>\<^sub>2"
using assms by (induct \<pi>\<^sub>1 eV V \<pi>\<^sub>2 eV' rule: smallstep.induct) auto
lemma smallstepsV_consumes_proofstream:
assumes "\<lless> \<pi>\<^sub>1, eV \<ggreater> V\<rightarrow>i \<lless> \<pi>\<^sub>2, eV' \<ggreater>"
obtains \<pi> where "\<pi>\<^sub>1 = \<pi> @ \<pi>\<^sub>2"
using assms by (induct \<pi>\<^sub>1 eV V i \<pi>\<^sub>2 eV' rule: smallsteps.induct)
(auto elim: smallstepV_consumes_proofstream)
lemma smallstepP_generates_proofstream:
assumes "\<lless> \<pi>\<^sub>1, eP \<ggreater> P\<rightarrow> \<lless> \<pi>\<^sub>2, eP' \<ggreater>"
obtains \<pi> where "\<pi>\<^sub>2 = \<pi>\<^sub>1 @ \<pi>"
using assms by (induct \<pi>\<^sub>1 eP P \<pi>\<^sub>2 eP' rule: smallstep.induct) auto
lemma smallstepsP_generates_proofstream:
assumes "\<lless> \<pi>\<^sub>1, eP \<ggreater> P\<rightarrow>i \<lless> \<pi>\<^sub>2, eP' \<ggreater>"
obtains \<pi> where "\<pi>\<^sub>2 = \<pi>\<^sub>1 @ \<pi>"
using assms by (induct \<pi>\<^sub>1 eP P i \<pi>\<^sub>2 eP' rule: smallsteps.induct)
(auto elim: smallstepP_generates_proofstream)
lemma smallstepV_ps_append:
"\<lless> \<pi>, eV \<ggreater> V\<rightarrow> \<lless> \<pi>', eV' \<ggreater> \<longleftrightarrow> \<lless> \<pi> @ X, eV \<ggreater> V\<rightarrow> \<lless> \<pi>' @ X, eV' \<ggreater>" (is "?L \<longleftrightarrow> ?R")
proof (rule iffI)
assume ?L then show ?R
by (nominal_induct \<pi> eV V \<pi>' eV' avoiding: X rule: smallstep.strong_induct)
(auto simp: fresh_append fresh_Pair)
next
assume ?R then show ?L
by (nominal_induct "\<pi> @ X" eV V "\<pi>' @ X" eV' avoiding: X rule: smallstep.strong_induct)
(auto simp: fresh_append fresh_Pair)
qed
lemma smallstepV_ps_to_suffix:
assumes "\<lless>\<pi>, e\<ggreater> V\<rightarrow> \<lless>\<pi>' @ X, e'\<ggreater>"
obtains \<pi>'' where "\<pi> = \<pi>'' @ X"
using assms
by (induct \<pi> e V "\<pi>' @ X" e' rule: smallstep.induct) auto
lemma smallstepsV_ps_append:
"\<lless> \<pi>, eV \<ggreater> V\<rightarrow>i \<lless> \<pi>', eV' \<ggreater> \<longleftrightarrow> \<lless> \<pi> @ X, eV \<ggreater> V\<rightarrow>i \<lless> \<pi>' @ X, eV' \<ggreater>" (is "?L \<longleftrightarrow> ?R")
proof (rule iffI)
assume ?L then show ?R
proof (induct \<pi> eV V i \<pi>' eV' rule: smallsteps.induct)
case (s_Tr \<pi>\<^sub>1 e\<^sub>1 i \<pi>\<^sub>2 e\<^sub>2 \<pi>\<^sub>3 e\<^sub>3)
then show ?case
by (auto simp: iffD1[OF smallstepV_ps_append])
qed simp
next
assume ?R then show ?L
proof (induct "\<pi> @ X" eV V i "\<pi>' @ X" eV' arbitrary: \<pi>' rule: smallsteps.induct)
case (s_Tr e\<^sub>1 i \<pi>\<^sub>2 e\<^sub>2 e\<^sub>3)
from s_Tr(3) obtain \<pi>''' where "\<pi>\<^sub>2 = \<pi>''' @ X"
by (auto elim: smallstepV_ps_to_suffix)
with s_Tr show ?case
by (auto dest: iffD2[OF smallstepV_ps_append])
qed simp
qed
lemma smallstepP_ps_prepend:
"\<lless> \<pi>, eP \<ggreater> P\<rightarrow> \<lless> \<pi>', eP' \<ggreater> \<longleftrightarrow> \<lless> X @ \<pi>, eP \<ggreater> P\<rightarrow> \<lless> X @ \<pi>', eP' \<ggreater>" (is "?L \<longleftrightarrow> ?R")
proof (rule iffI)
assume ?L then show ?R
proof (nominal_induct \<pi> eP P \<pi>' eP' avoiding: X rule: smallstep.strong_induct)
case (s_UnauthP v \<pi> h)
then show ?case
by (subst append_assoc[symmetric, of X \<pi> "[\<lparr>v\<rparr>]"]) (erule smallstep.s_UnauthP)
qed (auto simp: fresh_append fresh_Pair)
next
assume ?R then show ?L
by (nominal_induct "X @ \<pi>" eP P "X @ \<pi>'" eP' avoiding: X rule: smallstep.strong_induct)
(auto simp: fresh_append fresh_Pair)
qed
lemma smallstepsP_ps_prepend:
"\<lless> \<pi>, eP \<ggreater> P\<rightarrow>i \<lless> \<pi>', eP' \<ggreater> \<longleftrightarrow> \<lless> X @ \<pi>, eP \<ggreater> P\<rightarrow>i \<lless> X @ \<pi>', eP' \<ggreater>" (is "?L \<longleftrightarrow> ?R")
proof (rule iffI)
assume ?L then show ?R
proof (induct \<pi> eP P i \<pi>' eP' rule: smallsteps.induct)
case (s_Tr \<pi>\<^sub>1 e\<^sub>1 i \<pi>\<^sub>2 e\<^sub>2 \<pi>\<^sub>3 e\<^sub>3)
then show ?case
by (auto simp: iffD1[OF smallstepP_ps_prepend])
qed simp
next
assume ?R then show ?L
proof (induct "X @ \<pi>" eP P i "X @ \<pi>'" eP' arbitrary: \<pi>' rule: smallsteps.induct)
case (s_Tr e\<^sub>1 i \<pi>\<^sub>2 e\<^sub>2 e\<^sub>3)
then obtain \<pi>'' where \<pi>'': "\<pi>\<^sub>2 = X @ \<pi> @ \<pi>''"
by (auto elim: smallstepsP_generates_proofstream)
then have "\<lless>\<pi>, e\<^sub>1\<ggreater> P\<rightarrow>i \<lless>\<pi> @ \<pi>'', e\<^sub>2\<ggreater>"
by (auto dest: s_Tr(2))
with \<pi>'' s_Tr(1,3) show ?case
by (auto dest: iffD2[OF smallstepP_ps_prepend])
qed simp
qed
subsection \<open>Type Progress\<close>
lemma type_progress:
assumes "{$$} \<turnstile>\<^sub>W e : \<tau>"
shows "value e \<or> (\<exists>e'. \<lless> [], e \<ggreater> I\<rightarrow> \<lless> [], e' \<ggreater>)"
using assms proof (nominal_induct "{$$} :: tyenv" e \<tau> rule: judge_weak.strong_induct)
case (jw_Let x e\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 \<tau>\<^sub>2)
then show ?case
by (auto 0 3 simp: fresh_smallstep_I elim!: s_Let2[of e\<^sub>2]
intro: exI[where P="\<lambda>e. \<lless>_, _\<ggreater> _\<rightarrow> \<lless>_, e\<ggreater>", OF s_Let1, rotated])
next
case (jw_Prj1 v \<tau>\<^sub>1 \<tau>\<^sub>2)
then show ?case
by (auto elim!: jw_Prod_inv[of v \<tau>\<^sub>1 \<tau>\<^sub>2])
next
case (jw_Prj2 v \<tau>\<^sub>1 \<tau>\<^sub>2)
then show ?case
by (auto elim!: jw_Prod_inv[of v \<tau>\<^sub>1 \<tau>\<^sub>2])
next
case (jw_App e \<tau>\<^sub>1 \<tau>\<^sub>2 e')
then show ?case
by (auto 0 4 elim: jw_Fun_inv[of _ _ _ e'])
next
case (jw_Case v v\<^sub>1 v\<^sub>2 \<tau>\<^sub>1 \<tau>\<^sub>2 \<tau>)
then show ?case
by (auto 0 4 elim: jw_Sum_inv[of _ v\<^sub>1 v\<^sub>2])
qed fast+
subsection \<open>Weak Type Preservation\<close>
lemma fresh_tyenv_None:
fixes \<Gamma> :: tyenv
shows "atom x \<sharp> \<Gamma> \<longleftrightarrow> \<Gamma> $$ x = None" (is "?L \<longleftrightarrow> ?R")
proof
assume assm: ?L show ?R
proof (rule ccontr)
assume "\<Gamma> $$ x \<noteq> None"
then obtain \<tau> where "\<Gamma> $$ x = Some \<tau>" by blast
with assm have "\<forall>a :: var. atom a \<sharp> \<Gamma> \<longrightarrow> \<Gamma> $$ a = Some \<tau>"
using fmap_freshness_lemma_unique[OF exI, of x \<Gamma>]
by (simp add: fresh_Pair fresh_Some) metis
then have "{a :: var. atom a \<sharp> \<Gamma>} \<subseteq> fmdom' \<Gamma>"
by (auto simp: image_iff Ball_def fmlookup_dom'_iff)
moreover
{ assume "infinite {a :: var. \<not> atom a \<sharp> \<Gamma>}"
then have "infinite {a :: var. atom a \<in> supp \<Gamma>}"
unfolding fresh_def by auto
then have "infinite (supp \<Gamma>)"
by (rule contrapos_nn)
(auto simp: image_iff inv_f_f[of atom] inj_on_def
elim!: finite_surj[of _ _ "inv atom"] bexI[rotated])
then have False
using finite_supp[of \<Gamma>] by blast
}
then have "infinite {a :: var. atom a \<sharp> \<Gamma>}"
by auto
ultimately show False
using finite_subset[of "{a. atom a \<sharp> \<Gamma>}" "fmdom' \<Gamma>"] unfolding fmdom'_alt_def
by auto
qed
next
assume ?R then show ?L
proof (induct \<Gamma> arbitrary: x)
case (fmupd y z \<Gamma>)
then show ?case
by (cases "y = x") (auto intro: fresh_fmap_update)
qed simp
qed
lemma judge_weak_fresh_env_fresh_term[dest]:
fixes a :: var
assumes "\<Gamma> \<turnstile>\<^sub>W e : \<tau>" "atom a \<sharp> \<Gamma>"
shows "atom a \<sharp> e"
using assms proof (nominal_induct \<Gamma> e \<tau> avoiding: a rule: judge_weak.strong_induct)
case (jw_Var \<Gamma> x \<tau>)
then show ?case
by (cases "a = x") (auto simp: fresh_tyenv_None)
qed (simp_all add: fresh_Cons fresh_fmap_update)
lemma judge_weak_weakening_1:
assumes "\<Gamma> \<turnstile>\<^sub>W e : \<tau>" "atom y \<sharp> e"
shows "\<Gamma>(y $$:= \<tau>') \<turnstile>\<^sub>W e : \<tau>"
using assms proof (nominal_induct \<Gamma> e \<tau> avoiding: y \<tau>' rule: judge_weak.strong_induct)
case (jw_Lam x \<Gamma> \<tau>\<^sub>1 e \<tau>\<^sub>2)
from jw_Lam(5)[of y \<tau>'] jw_Lam(1-4,6) show ?case
by (auto simp add: fresh_at_base fmupd_reorder_neq fresh_fmap_update)
next
case (jw_App v v' \<Gamma> \<tau>\<^sub>1 \<tau>\<^sub>2)
then show ?case
by (force simp add: fresh_at_base fmupd_reorder_neq fresh_fmap_update)
next
case (jw_Let x \<Gamma> e\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 \<tau>\<^sub>2)
from jw_Let(6)[of y \<tau>'] jw_Let(8)[of y \<tau>'] jw_Let(1-5,7,9) show ?case
by (auto simp add: fresh_at_base fmupd_reorder_neq fresh_fmap_update)
next
case (jw_Rec x \<Gamma> z \<tau>\<^sub>1 \<tau>\<^sub>2 e')
from jw_Rec(9)[of y \<tau>'] jw_Rec(1-8,10) show ?case
by (auto simp add: fresh_at_base fmupd_reorder_neq fresh_fmap_update fresh_Pair)
next
case (jw_Case v v\<^sub>1 v\<^sub>2 \<Gamma> \<tau>\<^sub>1 \<tau>\<^sub>2 \<tau>)
then show ?case
by (fastforce simp add: fresh_at_base fmupd_reorder_neq fresh_fmap_update)
next
case (jw_Roll \<alpha> \<Gamma> v \<tau>)
then show ?case
by (simp add: fresh_fmap_update)
next
case (jw_Unroll \<alpha> \<Gamma> v \<tau>)
then show ?case
by (simp add: fresh_fmap_update)
qed auto
lemma judge_weak_weakening_2:
assumes "\<Gamma> \<turnstile>\<^sub>W e : \<tau>" "atom y \<sharp> \<Gamma>"
shows "\<Gamma>(y $$:= \<tau>') \<turnstile>\<^sub>W e : \<tau>"
proof -
from assms have "atom y \<sharp> e"
by (rule judge_weak_fresh_env_fresh_term)
with assms show "\<Gamma>(y $$:= \<tau>') \<turnstile>\<^sub>W e : \<tau>" by (simp add: judge_weak_weakening_1)
qed
lemma judge_weak_weakening_env:
assumes "{$$} \<turnstile>\<^sub>W e : \<tau>"
shows "\<Gamma> \<turnstile>\<^sub>W e : \<tau>"
using assms proof (induct \<Gamma>)
case fmempty
then show ?case by assumption
next
case (fmupd x y \<Gamma>)
then show ?case
by (simp add: fresh_tyenv_None judge_weak_weakening_2)
qed
lemma value_subst_value:
assumes "value e" "value e'"
shows "value (e[e' / x])"
using assms by (induct e e' x rule: subst_term.induct) auto
lemma judge_weak_subst[intro]:
assumes "\<Gamma>(a $$:= \<tau>') \<turnstile>\<^sub>W e : \<tau>" "{$$} \<turnstile>\<^sub>W e' : \<tau>'"
shows "\<Gamma> \<turnstile>\<^sub>W e[e' / a] : \<tau>"
using assms proof (nominal_induct "\<Gamma>(a $$:= \<tau>')" e \<tau> avoiding: a e' \<Gamma> rule: judge_weak.strong_induct)
case (jw_Var x \<tau>)
then show ?case
by (auto simp: judge_weak_weakening_env)
next
case (jw_Lam x \<tau>\<^sub>1 e \<tau>\<^sub>2)
then show ?case
by (fastforce simp: fmupd_reorder_neq)
next
case (jw_Rec x y \<tau>\<^sub>1 \<tau>\<^sub>2 e)
then show ?case
by (fastforce simp: fmupd_reorder_neq)
next
case (jw_Let x e\<^sub>1 \<tau>\<^sub>1 e\<^sub>2 \<tau>\<^sub>2)
then show ?case
by (fastforce simp: fmupd_reorder_neq)
qed fastforce+
lemma type_preservation:
assumes "\<lless> [], e \<ggreater> I\<rightarrow> \<lless> [], e' \<ggreater>" "{$$} \<turnstile>\<^sub>W e : \<tau>"
shows "{$$} \<turnstile>\<^sub>W e' : \<tau>"
using assms [[simproc del: alpha_lst]]
proof (nominal_induct "[]::proofstream" e I "[]::proofstream" e' arbitrary: \<tau> rule: smallstep.strong_induct)
case (s_AppLam v x e)
then show ?case by force
next
case (s_AppRec v x e)
then show ?case
by (elim jw_App_inv jw_Rec_inv) (auto 0 3 simp del: subst_term.simps)
next
case (s_Let1 x e\<^sub>1 e\<^sub>1' e\<^sub>2)
from s_Let1(1,2,7) show ?case
by (auto intro: s_Let1(6) del: jw_Let_inv elim!: jw_Let_inv)
next
case (s_Unroll e e')
then obtain \<alpha>::tvar where "atom \<alpha> \<sharp> \<tau>"
using obtain_fresh by blast
with s_Unroll show ?case
by (auto elim: jw_Unroll_inv[where \<alpha> = \<alpha>])
next
case (s_Roll e e')
then obtain \<alpha>::tvar where "atom \<alpha> \<sharp> \<tau>"
using obtain_fresh by blast
with s_Roll show ?case
by (auto elim: jw_Roll_inv[where \<alpha> = \<alpha>])
next
case (s_UnrollRoll v)
then obtain \<alpha>::tvar where "atom \<alpha> \<sharp> \<tau>"
using obtain_fresh by blast
with s_UnrollRoll show ?case
by (fastforce simp: Abs1_eq(3) elim: jw_Roll_inv[where \<alpha> = \<alpha>] jw_Unroll_inv[where \<alpha> = \<alpha>])
qed fastforce+
subsection \<open>Corrected Lemma 1 from Miller et al.~\cite{adsg}: Weak Type Soundness\<close>
lemma type_soundness:
assumes "{$$} \<turnstile>\<^sub>W e : \<tau>"
shows "value e \<or> (\<exists>e'. \<lless> [], e \<ggreater> I\<rightarrow> \<lless> [], e' \<ggreater> \<and> {$$} \<turnstile>\<^sub>W e' : \<tau>)"
proof (cases "value e")
case True
then show ?thesis by simp
next
case False
with assms obtain e' where "\<lless>[], e\<ggreater> I\<rightarrow> \<lless>[], e'\<ggreater>" by (auto dest: type_progress)
with assms show ?thesis
by (auto simp: type_preservation)
qed
(*<*)
end
(*>*)
diff --git a/thys/LightweightJava/Lightweight_Java_Equivalence.thy b/thys/LightweightJava/Lightweight_Java_Equivalence.thy
--- a/thys/LightweightJava/Lightweight_Java_Equivalence.thy
+++ b/thys/LightweightJava/Lightweight_Java_Equivalence.thy
@@ -1,893 +1,894 @@
(* Title: Lightweight Java, the equivalence functions
Authors: Rok Strnisa <rok@strnisa.com>, 2006
Matthew Parkinson <matt@matthewp.com>, 2006
Maintainer:
*)
theory Lightweight_Java_Equivalence imports Lightweight_Java_Definition begin
(* BEGIN: HELPER FUNCTIONS *)
lemma map_id[simp]: "map id list = list" by (induct list) auto
lemma id_map_two[simp]: "map (\<lambda>(x,y). (x,y)) list = list" by (induct list) auto
lemma id_image_two[simp]: "(\<lambda>(x,y). (x,y)) ` set list = set list" by (induct list) auto
lemma map_fst[simp]: "map (\<lambda>(x, y). x) list = map fst list" by (induct list) auto
lemma map_snd[simp]: "map (\<lambda>(x, y). y) list = map snd list" by (induct list) auto
lemma zip_map_map_two [simp]: "zip (map fst list) (map snd list) = list"
by (induct list) auto
lemma concat_map_singlton [simp]: "concat (map (\<lambda>e. [e]) list) = list"
by (induct list) simp_all
lemma list_all_map_P [simp]: "list_all (\<lambda>b. b) (map (\<lambda>x. P x) list) = (\<forall>x \<in> set list. P x)"
by (induct list) simp_all
lemma dom_single[simp]: "a : dom (Map.empty(k \<mapsto> v)) = (a = k)"
by (simp add: dom_def)
lemma predicted_lu[rule_format, simp]:
"x \<in> set list \<longrightarrow> map_of (map (\<lambda>key. (key, value)) list) x = Some value"
by (induct list) auto
lemma key_in_map1[simp]: "k \<notin> dom M' \<Longrightarrow> (M ++ M') k = M k"
apply (subgoal_tac "M' k = None")
apply (simp add: map_add_def, force simp add: domI)
done
lemma forall_cons [rule_format]: "(\<forall>x \<in> set (s#S). P x) \<and> y \<in> set S \<longrightarrow> P y"
by (induct_tac S) simp_all
lemma mem_cong[rule_format]:
"x \<in> set list \<longrightarrow> (f x \<in> set (map f list))"
by (induct list) auto
lemma forall_union: "\<lbrakk>\<forall>a \<in> dom A. P (A a); \<forall>b \<in> dom B. P (B b)\<rbrakk> \<Longrightarrow> \<forall>x \<in> dom A \<union> dom B. P ((B ++ A) x)"
apply(safe)
apply(force)
apply(drule_tac x = x in bspec, simp add: domI)
apply(case_tac "A x = None")
apply(force simp add: map_add_def)
by (force)
(* END: HELPER FUNCTIONS *)
definition
class_name_f :: "cld \<Rightarrow> dcl"
where
"class_name_f cld =
(case cld of cld_def dcl cl fds mds \<Rightarrow> dcl)"
lemma [simp]: "(class_name cld dcl) = (class_name_f cld = dcl)"
by (force simp add: class_name_f_def split: cld.splits
intro: class_nameI elim: class_name.cases)
definition
superclass_name_f :: "cld \<Rightarrow> cl"
where
"superclass_name_f cld =
(case cld of cld_def dcl cl fds mds \<Rightarrow> cl)"
lemma [simp]: "(superclass_name cld cl) = (superclass_name_f cld = cl)"
by (force simp add: superclass_name_f_def split: cld.splits
intro: superclass_nameI elim: superclass_name.cases)
definition
class_fields_f :: "cld \<Rightarrow> fds"
where
"class_fields_f cld =
(case cld of cld_def dcl cl fds mds \<Rightarrow> fds)"
lemma [simp]: "(class_fields cld fds) = (class_fields_f cld = fds)"
by (force simp add: class_fields_f_def split: cld.splits
intro: class_fieldsI elim: class_fields.cases)
definition
class_methods_f :: "cld \<Rightarrow> meth_defs"
where
"class_methods_f cld =
(case cld of cld_def dcl cl fds mds \<Rightarrow> mds)"
lemma [simp]: "(class_methods cld fds) = (class_methods_f cld = fds)"
by (force simp add: class_methods_f_def split: cld.splits
intro: class_methodsI elim: class_methods.cases)
definition
method_name_f :: "meth_def \<Rightarrow> meth"
where
"method_name_f md =
(case md of meth_def_def meth_sig meth_body \<Rightarrow>
(case meth_sig of meth_sig_def cl meth vds \<Rightarrow> meth))"
lemma [simp]: "(method_name md m) = (method_name_f md = m)"
by (force simp add: method_name_f_def split: meth_def.splits meth_sig.splits
intro: method_nameI elim: method_name.cases)
definition
distinct_names_f :: "P \<Rightarrow> bool"
where
"distinct_names_f P =
(distinct (map class_name_f P))"
lemma distinct_names_map[rule_format]:
"(\<forall>x\<in>set cld_dcl_list. case_prod (\<lambda>cld. (=) (class_name_f cld)) x) \<and> distinct (map snd cld_dcl_list)
\<longrightarrow> distinct_names_f (map fst cld_dcl_list)"
apply(induct cld_dcl_list)
apply(clarsimp simp add: distinct_names_f_def)+ apply(force)
done
lemma [simp]: "(distinct_names P) = (distinct_names_f P)"
apply(rule)
apply(erule distinct_names.cases) apply(clarsimp simp add: distinct_names_map)
apply(simp add: distinct_names_f_def)
apply(rule_tac cld_dcl_list = "map (\<lambda>cld. (cld, class_name_f cld)) P" in dn_defI)
apply(simp) apply(induct P) apply(simp) apply(simp)
apply(clarsimp)
apply(clarsimp) apply(induct P) apply(simp) apply(force)
done
primrec
find_cld_f :: "P \<Rightarrow> ctx \<Rightarrow> fqn \<Rightarrow> ctxcld_opt"
where
"find_cld_f [] ctx fqn = None" |
"find_cld_f (cld#clds) ctx fqn =
(case cld of cld_def dcl cl fds mds \<Rightarrow>
(case fqn of fqn_def dcl' \<Rightarrow>
(if dcl = dcl' then Some (ctx, cld) else find_cld_f clds ctx fqn)))"
lemma [simp]: "(find_cld P ctx fqn ctxcld_opt) = (find_cld_f P ctx fqn = ctxcld_opt)"
apply(rule)
apply(induct rule: find_cld.induct) apply(simp+)
apply(clarsimp)
apply(induct P)
apply(simp add: fc_emptyI)
apply(case_tac fqn) apply(rename_tac cld clds dcl) apply(case_tac cld)
apply(clarsimp) apply(rename_tac dcl' cl' vds' mds') apply(rule)
apply(clarsimp) apply(rule fc_cons_trueI) apply(simp) apply(force)
apply(force intro: fc_cons_falseI[simplified])
done
lemma find_to_mem[rule_format]:
"find_cld_f P ctx fqn = Some (ctx', cld) \<longrightarrow> cld : set P"
apply(induct P) by (clarsimp split: cld.splits fqn.splits)+
lemma find_cld_name_eq[rule_format]:
"\<forall>ctxcld. find_cld_f P ctx (fqn_def dcl) = Some ctxcld \<longrightarrow> (\<exists>cl fds mds. (ctx, cld_def dcl cl fds mds) = ctxcld)"
apply(induct P) apply(simp) apply(clarsimp split: cld.splits fqn.splits) done
primrec
find_type_f :: "P \<Rightarrow> ctx \<Rightarrow> cl \<Rightarrow> ty_opt"
where
"find_type_f P ctx cl_object = Some ty_top" |
"find_type_f P ctx (cl_fqn fqn) =
(case fqn of fqn_def dcl \<Rightarrow>
(case find_cld_f P ctx fqn of None \<Rightarrow> None | Some (ctx', cld) \<Rightarrow>
Some (ty_def ctx' dcl)))"
lemma [simp]: "(find_type P ctx cl ty_opt) = (find_type_f P ctx cl = ty_opt)"
apply(rule)
apply(force elim: find_type.cases split: fqn.splits)
apply(case_tac cl) apply(force intro: ft_objI)
apply(rename_tac fqn)
apply(case_tac fqn) apply(clarsimp)
apply(split option.splits) apply(rule)
apply(force intro: ft_nullI)
by (force intro: ft_dclI)
lemma mem_remove: "cld : set P \<Longrightarrow> length (remove1 cld P) < length P"
apply(induct P) by(simp, force split: if_split_asm)
lemma finite_program[rule_format, intro]:
"\<forall>P cld. (\<exists>ctx ctx' fqn. find_cld_f P ctx fqn = Some (ctx', cld)) \<longrightarrow>
length (remove1 cld P) < length P"
apply(clarsimp) apply(drule find_to_mem) by (simp add: mem_remove)
lemma path_length_eq[rule_format]:
"path_length P ctx cl nn \<Longrightarrow> \<forall>nn'. path_length P ctx cl nn' \<longrightarrow> nn = nn'"
apply(erule path_length.induct)
apply(clarsimp) apply(erule path_length.cases) apply(simp) apply(simp)
apply(rule) apply(rule)
apply(erule_tac ?a4.0 = nn' in path_length.cases) apply(clarify)
apply(clarsimp)
done
lemma fpr_termination[iff]:
"\<forall>P cld ctx ctx' fqn. find_cld_f P ctx fqn = Some (ctx', cld) \<and> acyclic_clds P
\<longrightarrow> The (path_length P ctx' (superclass_name_f cld)) < The (path_length P ctx (cl_fqn fqn))"
apply(clarsimp)
apply(erule acyclic_clds.cases) apply(clarsimp) apply(rename_tac P)
apply(erule_tac x = ctx in allE)
apply(erule_tac x = fqn in allE)
apply(clarsimp)
apply(rule theI2) apply(simp) apply(simp add: path_length_eq)
apply(erule path_length.cases) apply(simp) apply(clarsimp)
apply(rule theI2) apply(simp) apply(simp add: path_length_eq)
apply(drule_tac path_length_eq, simp)
apply(erule path_length.cases) apply(simp) apply(clarsimp)
apply(drule_tac path_length_eq, simp) apply(simp)
done
function
find_path_rec_f :: "P \<Rightarrow> ctx \<Rightarrow> cl \<Rightarrow> ctxclds \<Rightarrow> ctxclds_opt"
where
"find_path_rec_f P ctx cl_object path = Some path" |
"find_path_rec_f P ctx (cl_fqn fqn) path =
(if ~(acyclic_clds P) then None else
(case find_cld_f P ctx fqn of None \<Rightarrow> None | Some (ctx', cld) \<Rightarrow>
find_path_rec_f P ctx'
(superclass_name_f cld) (path @ [(ctx',cld)])))"
by pat_completeness auto
termination
by (relation "measure (\<lambda>(P, ctx, cl, path). (THE nn. path_length P ctx cl nn))") auto
lemma [simp]: "(find_path_rec P ctx cl path path_opt) = (find_path_rec_f P ctx cl path = path_opt)"
apply(rule)
apply(erule find_path_rec.induct) apply(simp)+
apply(induct rule: find_path_rec_f.induct)
apply(clarsimp simp add: fpr_objI)
apply(clarsimp) apply(rule)
apply(simp add: fpr_nullI)
apply(clarsimp split: option.splits)
apply(rule fpr_nullI)
apply(simp add: fpr_nullI)
apply(rule fpr_fqnI) apply(force)+
done
definition
find_path_f :: "P \<Rightarrow> ctx \<Rightarrow> cl \<Rightarrow> ctxclds_opt"
where
"find_path_f P ctx cl = find_path_rec_f P ctx cl []"
lemma [simp]: "(find_path P ctx cl path_opt) = (find_path_f P ctx cl = path_opt)"
apply(rule)
apply(erule find_path.cases) apply(unfold find_path_f_def) apply(simp)
apply(simp add: fp_defI)
done
primrec
find_path_ty_f :: "P \<Rightarrow> ty \<Rightarrow> ctxclds_opt"
where
"find_path_ty_f P ty_top = Some []" |
"find_path_ty_f P (ty_def ctx dcl) =
find_path_f P ctx (cl_fqn (fqn_def dcl))"
lemma [simp]: "(find_path_ty P ty ctxclds_opt) = (find_path_ty_f P ty = ctxclds_opt)"
apply(rule)
apply(force elim: find_path_ty.cases)
apply(case_tac ty)
apply(clarsimp simp add: fpty_objI)
apply(clarsimp simp add: fpty_dclI)
done
primrec
fields_in_path_f :: "ctxclds \<Rightarrow> fs"
where
"fields_in_path_f [] = []" |
"fields_in_path_f (ctxcld#ctxclds) =
(map (\<lambda>fd. case fd of fd_def cl f \<Rightarrow> f) (class_fields_f (snd ctxcld)))
@ fields_in_path_f ctxclds"
lemma cl_f_list_map: "map (case_fd (\<lambda>cl f. f)) (map (\<lambda>(x, y). fd_def x y) cl_f_list) = map (\<lambda>(cl_XXX, f_XXX). f_XXX) cl_f_list"
by (induct cl_f_list, auto)
lemma fip_ind_to_f: "\<forall>fs. fields_in_path clds fs \<longrightarrow> fields_in_path_f clds = fs"
apply(induct clds)
apply(clarsimp, erule fields_in_path.cases) apply(simp) apply(clarsimp)
apply(clarsimp)
apply(erule fields_in_path.cases) apply(simp) by(clarsimp simp add: cl_f_list_map)
lemma fd_map_split: "map (case_fd (\<lambda>cl f. f)) (map (\<lambda>(x, y). fd_def x y) list) = map (\<lambda>(cl, f). f) list"
apply(induct list) apply(simp) apply(clarsimp) done
lemma fd_map_split': "map (\<lambda>(x, y). fd_def x y) (map (case_fd Pair) list) = list"
apply(induct list) apply(simp split: fd.splits)+ done
lemma fd_map_split'': "map ((\<lambda>(x, y). fd_def x y) \<circ> case_fd Pair) list = list"
apply(induct list) apply(simp split: fd.splits)+ done
lemma [simp]: "\<forall>fs. (fields_in_path ctxclds fs) = (fields_in_path_f ctxclds = fs)"
apply(induct ctxclds)
apply(rule) apply(rule)
apply(force elim: fields_in_path.cases)
apply(simp add: fip_emptyI)
apply(clarsimp)
apply(rule)
apply(erule fields_in_path.cases)
apply(simp)
apply(simp add: fip_ind_to_f fd_map_split)
apply(clarsimp)
apply(rule_tac cld = b and ctxcld_list = ctxclds
and cl_f_list = "map (case_fd (\<lambda>cl f. (cl, f))) (class_fields_f b)" in fip_consI[simplified])
apply(simp add: fd_map_split'')
apply(simp add: fd_map_split')
apply(clarsimp split: fd.splits)
done
definition
fields_f :: "P \<Rightarrow> ty \<Rightarrow> fs option"
where
"fields_f P ty =
(case find_path_ty_f P ty of None \<Rightarrow> None | Some ctxclds \<Rightarrow>
Some (fields_in_path_f ctxclds))"
lemma [simp]: "\<forall>fs_opt. (fields P ty fs_opt) = (fields_f P ty = fs_opt)"
apply(rule) apply(rule)
apply(case_tac fs_opt)
apply(clarsimp) apply(erule fields.cases)
apply(clarsimp) apply(simp add: fields_f_def)
apply(clarsimp)
apply(erule fields.cases)
apply(simp)
apply(clarsimp) apply(simp add: fields_f_def)
apply(simp add: fields_f_def) apply(case_tac "find_path_ty_f P ty") apply(simp) apply(simp add: fields_noneI[simplified])
apply(clarsimp) apply(case_tac "find_path_ty_f P ty") apply(simp) apply(simp)
apply(rule fields_someI) apply(simp) apply(simp)
done
primrec
methods_in_path_f :: "clds \<Rightarrow> meths"
where
"methods_in_path_f [] = []" |
"methods_in_path_f (cld#clds) =
map (\<lambda>md. case md of meth_def_def meth_sig meth_body \<Rightarrow>
case meth_sig of meth_sig_def cl meth vds \<Rightarrow> meth)
(class_methods_f cld) @ methods_in_path_f clds"
lemma meth_def_map[THEN mp]:
"(\<forall>x \<in> set list. (\<lambda>(md, cl, m, vds, mb). md = meth_def_def (meth_sig_def cl m vds) mb) x)
\<longrightarrow> map (case_meth_def (\<lambda>ms mb. case ms of meth_sig_def cl m vds \<Rightarrow> m)) (map (\<lambda>(md, cl, m, vds, mb). md) list) = map (\<lambda>(md, cl, m, vds, mb). m) list"
by (induct list, auto)
lemma meth_def_map':
"map ((\<lambda>(md, cl, m, vds, mb). md) \<circ> (\<lambda>md. case md of meth_def_def (meth_sig_def cl m vds) mb \<Rightarrow> (md, cl, m, vds, mb))) list = list"
apply(induct list) by (auto split: meth_def.splits meth_sig.splits)
lemma [simp]: "\<forall>meths. (methods_in_path clds meths) = (methods_in_path_f clds = meths)"
apply(induct clds)
apply(clarsimp) apply(rule)
apply(erule methods_in_path.cases) apply(simp) apply(clarsimp)
apply(clarsimp) apply(rule mip_emptyI)
apply(clarsimp) apply(rule)
apply(erule methods_in_path.cases) apply(simp)
apply(force)
apply(rule_tac
meth_def_cl_meth_vds_meth_body_list =
"(case a of cld_def dcl cl fds mds \<Rightarrow>
(map (\<lambda>md. (case md of meth_def_def ms mb \<Rightarrow>
(case ms of meth_sig_def cl m vds \<Rightarrow> (md, cl, m, vds, mb)))) mds))" in
mip_consI[simplified])
apply(clarsimp) apply(case_tac a) apply(simp add: class_methods_f_def meth_def_map')
apply(clarsimp) apply(clarsimp split: meth_def.splits meth_sig.splits)
apply(simp)
apply(clarsimp) apply(case_tac a) apply(clarsimp simp add: class_methods_f_def split: meth_def.splits meth_sig.splits)
done
definition
methods_f :: "P \<Rightarrow> ty \<Rightarrow> meths option"
where
"methods_f P ty =
(case find_path_ty_f P ty of None \<Rightarrow> None | Some ctxclds \<Rightarrow>
Some (methods_in_path_f (map (\<lambda>(ctx, cld). cld) ctxclds)))"
lemma [simp]: "(methods P ty meths) = (methods_f P ty = Some meths)"
apply(rule)
apply(erule methods.cases) apply(clarsimp simp add: methods_f_def)
apply(simp add: methods_f_def) apply(split option.splits) apply(simp)
apply(clarsimp) apply(rule methods_methodsI) apply(simp) apply(clarsimp)
done
primrec
ftype_in_fds_f :: "P \<Rightarrow> ctx \<Rightarrow> fds \<Rightarrow> f \<Rightarrow> ty_opt_bot"
where
"ftype_in_fds_f P ctx [] f = ty_opt_bot_opt None" |
"ftype_in_fds_f P ctx (fd#fds) f =
(case fd of fd_def cl f' \<Rightarrow> (if f = f' then
(case find_type_f P ctx cl of None \<Rightarrow> ty_opt_bot_bot | Some ty \<Rightarrow>
ty_opt_bot_opt (Some ty)) else ftype_in_fds_f P ctx fds f))"
lemma [simp]: "(ftype_in_fds P ctx fds f ty_opt) = (ftype_in_fds_f P ctx fds f = ty_opt)"
apply(rule)
apply(induct fds) apply(erule ftype_in_fds.cases) apply(simp) apply(simp) apply(simp) apply(simp)
apply(erule ftype_in_fds.cases) apply(simp) apply(simp) apply(simp) apply(clarsimp)
apply(induct fds)
apply(clarsimp) apply(rule ftif_emptyI)
apply(rename_tac fd fds) apply(case_tac fd, rename_tac cl f', clarsimp) apply(rule)
apply(clarsimp) apply(case_tac "find_type_f P ctx cl")
apply(simp add: ftif_cons_botI[simplified])
apply(simp add: ftif_cons_trueI[simplified])
apply(force intro: ftif_cons_falseI[simplified])
done
primrec
ftype_in_path_f :: "P \<Rightarrow> ctxclds \<Rightarrow> f \<Rightarrow> ty_opt"
where
"ftype_in_path_f P [] f = None" |
"ftype_in_path_f P (ctxcld#ctxclds) f =
(case ctxcld of (ctx, cld) \<Rightarrow>
(case ftype_in_fds_f P ctx (class_fields_f cld) f of
ty_opt_bot_bot \<Rightarrow> None | ty_opt_bot_opt ty_opt \<Rightarrow>
(case ty_opt of Some ty \<Rightarrow> Some ty | None \<Rightarrow>
ftype_in_path_f P ctxclds f)))"
lemma [simp]: "(ftype_in_path P ctxclds f ty_opt) = (ftype_in_path_f P ctxclds f = ty_opt)"
apply(rule)
apply(induct rule: ftype_in_path.induct) apply(simp+)
apply(induct ctxclds) apply(simp) apply(rule ftip_emptyI)
apply(hypsubst_thin)
apply(clarsimp) apply(case_tac "ftype_in_fds_f P a (class_fields_f b) f")
apply(rename_tac ty_opt) apply(case_tac ty_opt)
apply(simp add: ftip_cons_falseI[simplified])
apply(simp add: ftip_cons_trueI[simplified])
apply(simp add: ftip_cons_botI[simplified])
done
definition
ftype_f :: "P \<Rightarrow> ty \<Rightarrow> f \<Rightarrow> ty_opt"
where
"ftype_f P ty f =
(case find_path_ty_f P ty of None \<Rightarrow> None | Some ctxclds \<Rightarrow>
ftype_in_path_f P ctxclds f)"
lemma [simp]: "(ftype P ty f ty') = (ftype_f P ty f = Some ty')"
apply(rule)
apply(induct rule: ftype.induct) apply(clarsimp simp add: ftype_f_def)
apply(clarsimp simp add: ftype_f_def split: option.splits) apply(rule ftypeI) apply(simp+)
done
primrec
find_meth_def_in_list_f :: "meth_defs \<Rightarrow> meth \<Rightarrow> meth_def_opt"
where
"find_meth_def_in_list_f [] m = None" |
"find_meth_def_in_list_f (md#mds) m =
(case md of meth_def_def ms mb \<Rightarrow>
(case ms of meth_sig_def cl m' vds \<Rightarrow>
(if m = m' then Some md else find_meth_def_in_list_f mds m)))"
lemma [simp]: "(find_meth_def_in_list mds m md_opt) = (find_meth_def_in_list_f mds m = md_opt)"
apply(rule)
apply(induct rule: find_meth_def_in_list.induct) apply(simp+)
apply(induct mds) apply(simp) apply(rule fmdil_emptyI)
apply(clarsimp) apply(clarsimp split: meth_def.split meth_sig.split) apply(rule)
apply(clarsimp) apply(rule fmdil_cons_trueI[simplified]) apply(force)
apply(clarsimp) apply(rule fmdil_cons_falseI[simplified]) apply(force+)
done
primrec
find_meth_def_in_path_f :: "ctxclds \<Rightarrow> meth \<Rightarrow> ctxmeth_def_opt"
where
fmdip_empty: "find_meth_def_in_path_f [] m = None" |
fmdip_cons: "find_meth_def_in_path_f (ctxcld#ctxclds) m =
(case ctxcld of (ctx, cld) \<Rightarrow>
(case find_meth_def_in_list_f (class_methods_f cld) m of
Some md \<Rightarrow> Some (ctx, md) |
None \<Rightarrow> find_meth_def_in_path_f ctxclds m))"
lemma [simp]: "(find_meth_def_in_path ctxclds m ctxmeth_def_opt) = (find_meth_def_in_path_f ctxclds m = ctxmeth_def_opt)"
apply(rule)
apply(induct rule: find_meth_def_in_path.induct) apply(simp+)
apply(induct ctxclds) apply(simp) apply(rule fmdip_emptyI)
apply(clarsimp) apply(case_tac "find_meth_def_in_list_f (class_methods_f b) m") apply(clarsimp)
apply(simp add: fmdip_cons_falseI[simplified])
apply(simp add: fmdip_cons_trueI[simplified])
done
definition
find_meth_def_f :: "P \<Rightarrow> ty \<Rightarrow> meth \<Rightarrow> ctxmeth_def_opt"
where
"find_meth_def_f P ty m =
(case find_path_ty_f P ty of None \<Rightarrow> None | Some ctxclds \<Rightarrow>
find_meth_def_in_path_f ctxclds m)"
lemma [simp]: "(find_meth_def P ty m ctxmd_opt) = (find_meth_def_f P ty m = ctxmd_opt)"
apply(rule)
apply(induct rule: find_meth_def.induct) apply(simp add: find_meth_def_f_def)+
apply(clarsimp simp add: find_meth_def_f_def split: option.splits)
apply(simp add: fmd_nullI)
apply(simp add: fmd_optI)
done
primrec
lift_opts :: "'a option list \<Rightarrow> 'a list option"
where
"lift_opts [] = Some []" |
"lift_opts (opt#opts) =
(case opt of None \<Rightarrow> None | Some v \<Rightarrow>
(case lift_opts opts of None \<Rightarrow> None | Some vs \<Rightarrow> Some (v#vs)))"
definition
mtype_f :: "P \<Rightarrow> ty \<Rightarrow> meth \<Rightarrow> mty option"
where
"mtype_f P ty m =
(case find_meth_def_f P ty m of None \<Rightarrow> None | Some (ctx, md) \<Rightarrow>
(case md of meth_def_def ms mb \<Rightarrow>
(case ms of meth_sig_def cl m' vds \<Rightarrow>
(case find_type_f P ctx cl of None \<Rightarrow> None | Some ty' \<Rightarrow>
(case lift_opts (map (\<lambda>vd. case vd of vd_def clk vark \<Rightarrow>
find_type_f P ctx clk) vds) of None \<Rightarrow> None | Some tys \<Rightarrow>
Some (mty_def tys ty'))))))"
lemma lift_opts_ind[rule_format]:
"(\<forall>x\<in>set list. (\<lambda>(cl, var, ty). find_type_f P ctx cl = Some ty) x)
\<longrightarrow> lift_opts (map (case_vd (\<lambda>clk vark. find_type_f P ctx clk) \<circ> (\<lambda>(cl, var, ty). vd_def cl var)) list) = Some (map (\<lambda>(cl, var, ty). ty) list)"
by (induct list, auto)
lemma find_md_m_match'[rule_format]:
"find_meth_def_in_list_f mds m = Some (meth_def_def (meth_sig_def cl m' vds) mb) \<longrightarrow> m' = m"
apply(induct mds) apply(simp) apply(clarsimp split: meth_def.splits meth_sig.splits) done
lemma find_md_m_match:
"find_meth_def_in_path_f path m = Some (ctx, meth_def_def (meth_sig_def cl m' vds) mb) \<longrightarrow> m' = m"
apply(induct path) apply(simp) apply(clarsimp split: option.splits) by(rule find_md_m_match')
lemma vds_map_length:
"length (map (case_vd (\<lambda>clk vark. find_type_f P ctx clk)) vds) = length vds"
by (induct vds, auto)
lemma lift_opts_length[rule_format]:
"\<forall>tys. lift_opts ty_opts = Some tys \<longrightarrow> length ty_opts = length tys"
apply(induct ty_opts) apply(simp) by(clarsimp split: option.splits)
lemma vds_tys_length_eq[rule_format]:
"lift_opts (map (case_vd (\<lambda>clk vark. find_type_f P ctx clk)) vds) = Some tys \<longrightarrow> length vds = length tys"
apply(rule) apply(drule lift_opts_length) apply(simp add: vds_map_length) done
lemma vds_tys_length_eq'[rule_format]:
"\<forall>tys. length vds = length tys \<longrightarrow> vds = map (\<lambda>(cl, var, ty). vd_def cl var) (map (\<lambda>(vd, ty). case vd of vd_def cl var \<Rightarrow> (cl, var, ty)) (zip vds tys))"
apply(induct vds) apply(simp) apply(clarsimp) apply(case_tac a) apply(clarsimp)
apply(case_tac tys) apply(simp) apply(clarsimp) done
lemma vds_tys_length_eq''[rule_format]:
"\<forall>vds. length vds = length tys \<longrightarrow> tys = map ((\<lambda>(cl, var, ty). ty) \<circ> (\<lambda>(vd, ty). case vd of vd_def cl var \<Rightarrow> (cl, var, ty))) (zip vds tys)"
apply(induct tys) apply(simp) apply(clarsimp) apply(case_tac vds) apply(clarsimp) apply(clarsimp)
apply(split vd.splits) apply(simp) done
lemma lift_opts_find_type[rule_format]:
"\<forall>tys. lift_opts (map (case_vd (\<lambda>clk vark. find_type_f P ctx clk)) vds) = Some tys
\<longrightarrow> (\<forall>(vd, ty) \<in> set (zip vds tys). case vd of vd_def cl var \<Rightarrow> find_type_f P ctx cl = Some ty)"
apply(induct vds) apply(simp) apply(clarsimp split: vd.splits option.splits) apply(rename_tac cl' var)
apply(drule_tac x = "(vd_def cl' var, b)" in bspec, simp) apply(force) done
lemma [simp]: "(mtype P ty m mty) = (mtype_f P ty m = Some mty)"
apply(rule)
apply(induct rule: mtype.induct) apply(clarsimp simp add: mtype_f_def)
apply(split option.splits) apply(rule) apply(clarsimp)
apply(rule_tac x = "map (\<lambda>(cl, var, ty). ty) cl_var_ty_list" in exI)
apply(simp add: lift_opts_ind) apply(clarsimp) apply(simp add: lift_opts_ind)
apply(clarsimp simp add: mtype_f_def split: option.splits meth_def.splits meth_sig.splits)
apply(rename_tac ctx mb cl m' vds ty' tys)
apply(rule_tac ctx = ctx and cl = cl and meth_body = mb and ty' = ty'
and cl_var_ty_list = "map (\<lambda>(vd, ty). case vd of vd_def cl var \<Rightarrow> (cl, var, ty)) (zip vds tys)"
and meth_def = "meth_def_def (meth_sig_def cl m' vds) mb" in mtypeI[simplified])
apply(simp) apply(clarsimp) apply(rule) apply(clarsimp simp add: find_meth_def_f_def split: option.splits)
apply(simp add: find_md_m_match) apply(drule vds_tys_length_eq) apply(frule vds_tys_length_eq') apply(clarsimp)
apply(simp) apply(clarsimp) apply(split vd.splits) apply(clarsimp) apply(drule lift_opts_find_type) apply(simp)
apply(clarsimp)
apply(clarsimp) apply(drule vds_tys_length_eq) apply(simp add: vds_tys_length_eq'')
done
definition
is_sty_one :: "P \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool option"
where
"is_sty_one P ty ty' =
(case find_path_ty_f P ty of None \<Rightarrow> None | Some ctxclds \<Rightarrow>
(case ty' of ty_top \<Rightarrow> Some True | ty_def ctx' dcl' \<Rightarrow>
Some ((ctx', dcl') : set (map (\<lambda>(ctx, cld). (ctx, class_name_f cld)) ctxclds)) ))"
lemma class_name_mem_map[rule_format]:
"(ctx, cld, class_name_f cld) \<in> set ctx_cld_dcl_list
\<Longrightarrow> (ctx, class_name_f cld)
\<in> ((\<lambda>(ctx, cld). (ctx, class_name_f cld)) \<circ> (\<lambda>(ctx, cld, dcl). (ctx, cld))) `
set ctx_cld_dcl_list"
by (induct ctx_cld_dcl_list, auto)
lemma map_map_three:
" ctxclds = map ((\<lambda>(ctx, cld, dcl). (ctx, cld)) \<circ> (\<lambda>(ctx, cld). (ctx, cld, class_name_f cld))) ctxclds"
by (induct ctxclds, auto)
lemma mem_el_map[rule_format]:
"(ctx, dcl) \<in> set ctxclds
\<Longrightarrow> (ctx, class_name_f dcl)
\<in> (\<lambda>(ctx_XXX, cld_XXX, y). (ctx_XXX, y)) `
set (map (\<lambda>(ctx, cld). (ctx, cld, class_name_f cld)) ctxclds)"
by (induct ctxclds, auto)
lemma [simp]: "(sty_one P ty ty') = (is_sty_one P ty ty' = Some True)"
apply(rule)
apply(induct rule: sty_one.induct)
apply(simp add: is_sty_one_def)
apply(clarsimp simp add: is_sty_one_def) apply(rename_tac ctx cld dcl)
apply(drule_tac x = "(ctx, cld, dcl)" in bspec, simp) apply(clarsimp)
using case_prod_conv image_iff apply fastforce
apply(clarsimp simp add: is_sty_one_def split: option.splits ty.splits)
apply(simp add: sty_objI)
apply(rename_tac ctxclds ctx dcl)
apply(rule_tac ctx_cld_dcl_list = "map (\<lambda>(ctx, cld). (ctx, cld, class_name_f cld)) ctxclds" in sty_dclI[simplified])
apply(clarsimp) apply(rule map_map_three)
apply(force)
apply(rule mem_el_map, assumption)
done
lemma path_append[rule_format]:
"find_path_rec_f P ctx cl path' = Some path \<Longrightarrow> \<exists>path''. path = path' @ path''"
apply(induct rule: find_path_rec_f.induct)
apply(clarsimp)
apply(force split: if_split_asm option.splits)
done
lemma all_in_path_found'[rule_format]:
"find_path_rec_f P ctx cl path' = Some path \<longrightarrow>
(\<forall>ctxcld \<in> set path. ctxcld \<in> set path' \<or> (\<exists>ctx' fqn'. find_cld_f P ctx' fqn' = Some ctxcld))"
apply(induct rule: find_path_rec_f.induct)
apply(clarsimp)
apply(rule)
apply(force split: if_split_asm option.splits)
done
lemma all_in_path_found:
"\<lbrakk>find_path_f P ctx cl = Some path; ctxcld \<in> set path\<rbrakk> \<Longrightarrow> \<exists>ctx' fqn'. find_cld_f P ctx' fqn' = Some ctxcld"
by (unfold find_path_f_def, simp only: all_in_path_found'[of _ _ _ "[]", simplified])
lemma fpr_target_is_head':
"find_path_rec_f P ctx cl path' = Some path \<longrightarrow>
(\<forall>fqn ctxcld. cl = cl_fqn fqn \<and> find_cld_f P ctx fqn = Some ctxcld \<longrightarrow>
(\<exists>path''. path = path' @ ctxcld # path''))"
apply(induct_tac P ctx cl path' rule: find_path_rec_f.induct)
apply(simp)
apply(clarsimp split: if_split_asm option.splits)
apply(case_tac "superclass_name_f b") apply(clarsimp)
apply(clarsimp split: if_split_asm option.splits)
apply(force)
done
lemma fpr_target_is_head:
"find_path_f P ctx (cl_fqn fqn) = Some path \<Longrightarrow> \<exists>ctxcld. find_cld_f P ctx fqn = Some ctxcld \<and> (\<exists>path''. path = ctxcld # path'')"
apply(unfold find_path_f_def) apply(frule fpr_target_is_head'[of _ _ _ "[]", THEN mp]) apply(clarsimp split: option.splits) done
lemma fpr_sub_path':
"find_path_rec_f P ctx cl path' = Some path \<longrightarrow>
(\<forall>fqn ctxcld path'' path_fqn.
cl = cl_fqn fqn \<and> find_cld_f P ctx fqn = Some ctxcld \<and>
find_path_rec_f P (fst ctxcld) (superclass_name_f (snd ctxcld)) path'' = Some path_fqn \<longrightarrow>
(\<exists>path'''. path_fqn = path'' @ path''' \<and> path = path' @ ctxcld # path'''))"
apply(induct_tac P ctx cl path' rule: find_path_rec_f.induct)
apply(simp)
apply(clarsimp split: if_split_asm option.splits)
apply(case_tac "superclass_name_f b") apply(simp add: find_path_f_def)
apply(clarsimp split: if_split_asm option.splits)
apply(frule_tac path = path in path_append) apply(clarsimp) apply(force)
done
lemma fpr_sub_path:
"\<lbrakk>find_path_f P ctx (cl_fqn fqn) = Some path; find_cld_f P ctx fqn = Some ctxcld;
find_path_f P (fst ctxcld) (superclass_name_f (snd ctxcld)) = Some path'\<rbrakk>
\<Longrightarrow> path = ctxcld # path'"
by (unfold find_path_f_def, force intro: fpr_sub_path'[rule_format, of _ _ _ "[]" _ _ _ "[]", simplified])
lemma fpr_sub_path_simp:
"\<lbrakk>find_path_rec_f P ctx (superclass_name_f cld) path'' = Some path_fqn; find_cld_f P ctx fqn = Some (ctx, cld); acyclic_clds P;
find_path_rec_f P ctx (superclass_name_f cld) (path' @ [(ctx, cld)]) = Some path\<rbrakk>
\<Longrightarrow> \<exists>path'''. path_fqn = path'' @ path''' \<and> path = path' @ (ctx, cld) # path'''"
apply(cut_tac P = P and ctx = ctx and cl = "cl_fqn fqn" and path' = path' and path = path in fpr_sub_path')
apply(clarsimp split: option.splits if_split_asm)
done
lemma fpr_same_suffix'[rule_format]:
"find_path_rec_f P ctx cl prefix = Some path \<longrightarrow>
(\<forall>suffix prefix'. path = prefix @ suffix \<longrightarrow>
find_path_rec_f P ctx cl prefix' = Some (prefix' @ suffix))"
apply(induct_tac P ctx cl prefix rule: find_path_rec_f.induct)
apply(clarsimp)
apply(clarsimp split: option.splits)
apply(frule path_append) apply(force)
done
lemma fpr_same_suffix:
"find_path_rec_f P ctx cl prefix = Some path \<longrightarrow>
(\<forall>suffix prefix' suffix'. path = prefix @ suffix \<and>
find_path_rec_f P ctx cl prefix' = Some (prefix' @ suffix')
\<longrightarrow> suffix = suffix')"
apply(induct_tac P ctx cl prefix rule: find_path_rec_f.induct)
apply(clarsimp)
by (metis fpr_same_suffix' option.inject same_append_eq)
lemma fpr_mid_path'[rule_format]:
"find_path_rec_f P ctx cl path' = Some path \<longrightarrow>
(\<forall>ctxcld \<in> set path.
ctxcld \<in> set path' \<or>
(\<forall>path_fqn. find_path_rec_f P (fst ctxcld) (cl_fqn (fqn_def (class_name_f (snd ctxcld)))) path'' = Some path_fqn \<longrightarrow>
(\<forall>path'''. path_fqn = path'' @ path''' \<longrightarrow> (\<exists>path_rest. path = path_rest @ path'''))))"
+supply [[simproc del: defined_all]]
apply(induct_tac P ctx cl path' rule: find_path_rec_f.induct)
apply(simp)
apply(clarsimp split: option.splits)
apply(frule find_cld_name_eq) apply(clarsimp)
apply(rename_tac path' ctx' cld' cld'' ctx'' path''' cl fds mds)
apply(subgoal_tac "find_path_rec_f P ctx' (superclass_name_f cld') (path' @ [(ctx', cld')]) = Some path \<longrightarrow>
(\<forall>ctxcld\<in>set path.
ctxcld = (ctx', cld') \<or>
ctxcld \<in> set path' \<or>
(\<forall>path_fqn. case_option None (case_prod (\<lambda>ctx' cld. find_path_rec_f P ctx' (superclass_name_f cld) (path'' @ [(ctx', cld)])))
(find_cld_f P (fst ctxcld) (fqn_def (class_name_f (snd ctxcld)))) =
Some path_fqn \<longrightarrow>
(\<forall>path'''. path_fqn = path'' @ path''' \<longrightarrow> (\<exists>path_rest. path = path_rest @ path'''))))")
apply(erule impE) apply simp
apply(drule_tac x = "(ctx'', cld'')" in bspec, simp) apply(clarsimp)
apply(simp add: superclass_name_f_def)
apply(case_tac cld') apply(rename_tac dcl' cl' fds' mds') apply(clarsimp simp add: class_name_f_def)
apply(case_tac fqn) apply(rename_tac dcl'') apply(clarsimp)
apply(frule find_cld_name_eq) apply(clarsimp)
apply(frule path_append) apply(frule_tac path = "path'' @ path'''" in path_append) apply(clarsimp)
apply(rule_tac x = path' in exI) apply(clarsimp)
apply(frule_tac suffix = path''a and prefix' = "path'' @ [(ctx', cld_def dcl' cl' fds' mds')]" and
suffix' = path''aa in fpr_same_suffix[rule_format]) apply(simp)
apply(force)
apply(simp)
done
lemma fpr_mid_path:
"\<lbrakk>find_path_f P ctx cl = Some path; (ctx', cld') \<in> set path;
find_path_f P ctx' (cl_fqn (fqn_def (class_name_f cld'))) = Some path'\<rbrakk>
\<Longrightarrow> \<exists>path''. path = path'' @ path'"
apply(cut_tac P = P and ctx = ctx and cl = cl and path' = "[]" and path = path and ctxcld = "(ctx', cld')" and path'' = "[]" in fpr_mid_path')
apply(unfold find_path_f_def, assumption) apply(assumption) apply(simp)
done
lemma fpr_first_in_path'[rule_format]:
"find_path_rec_f P ctx cl path' = Some path \<longrightarrow>
(\<forall>fqn ctxcld. cl = cl_fqn fqn \<and> find_cld_f P ctx fqn = Some ctxcld \<longrightarrow> ctxcld \<in> set path)"
apply(case_tac cl)
apply(simp)
apply(clarsimp) apply(drule path_append) apply(force)
done
lemma fpr_first_in_path:
"\<lbrakk>find_path_f P ctx (cl_fqn fqn) = Some path; find_cld_f P ctx fqn = Some ctxcld\<rbrakk> \<Longrightarrow> ctxcld \<in> set path"
by (unfold find_path_f_def, force intro: fpr_first_in_path'[of _ _ _ "[]", simplified])
lemma cld_for_path:
"find_path_f P ctx (cl_fqn fqn) = Some path \<Longrightarrow> \<exists>ctxcld. find_cld_f P ctx fqn = Some ctxcld"
apply(unfold find_path_f_def) apply(clarsimp split: if_split_asm option.splits) done
lemma ctx_cld_ctx_dcl[rule_format]:
"(ctx, cld_def dcl cl fds mds) \<in> set path \<longrightarrow> (ctx, dcl) \<in> (\<lambda>(ctx, cld). (ctx, class_name_f cld)) ` set path"
by (induct path, auto simp add: class_name_f_def)
lemma ctx_dcl_ctx_cld[rule_format]:
"(ctx, dcl) \<in> (\<lambda>(ctx, cld). (ctx, class_name_f cld)) ` set path \<longrightarrow> (\<exists>cl fds mds. (ctx, cld_def dcl cl fds mds) \<in> set path)"
apply(induct path) apply(simp)
apply(auto) apply(case_tac b) apply(force simp add: class_name_f_def)
done
lemma ctx_dcl_mem_path:
"find_path_f P ctx (cl_fqn (fqn_def dcl)) = Some path \<Longrightarrow> (ctx, dcl) \<in> (\<lambda>(ctx, cld). (ctx, class_name_f cld)) ` set path"
apply(frule cld_for_path) apply(erule exE)
apply(frule fpr_first_in_path) apply(assumption)
apply(frule find_cld_name_eq) apply(elim exE)
apply(clarsimp simp add: ctx_cld_ctx_dcl)
done
lemma sty_reflexiveI:
"is_sty_one P ty ty' = Some True \<Longrightarrow> is_sty_one P ty ty = Some True"
apply(simp add: is_sty_one_def split: option.splits) apply(case_tac ty) apply(simp) apply(clarsimp)
apply(drule ctx_dcl_mem_path) apply(force)
done
lemma sty_transitiveI:
"\<lbrakk>is_sty_one P ty ty' = Some True; is_sty_one P ty' ty'' = Some True\<rbrakk>
\<Longrightarrow> is_sty_one P ty ty'' = Some True"
apply(clarsimp simp add: is_sty_one_def split: ty.splits option.splits)
apply(rename_tac path ctx dcl path' ctx' dcl')
apply(case_tac ty) apply(simp) apply(clarsimp) apply(rename_tac ctx dcl ctx' cld' ctx'' cld'')
apply(frule fpr_mid_path) apply(simp) apply(simp) apply(force)
done
definition
is_sty_many :: "P \<Rightarrow> tys \<Rightarrow> tys \<Rightarrow> bool option"
where
"is_sty_many P tys tys' =
(if length tys \<noteq> length tys' then None else
(case lift_opts (map (\<lambda>(ty, ty'). is_sty_one P ty ty') (zip tys tys'))
of None \<Rightarrow> None | Some bools \<Rightarrow> Some (list_all id bools)))"
lemma lift_opts_exists:
"\<forall>x\<in>set ty_ty'_list. (\<lambda>(ty, ty'). is_sty_one P ty ty' = Some True) x \<Longrightarrow> \<exists>bools. lift_opts (map (\<lambda>(ty, ty'). is_sty_one P ty ty') ty_ty'_list) = Some bools"
by (induct ty_ty'_list, auto)
lemma lift_opts_all_true[rule_format]:
"\<forall>bools. (\<forall>x\<in>set ty_ty'_list. (\<lambda>(ty, ty'). is_sty_one P ty ty' = Some True) x) \<and>
lift_opts (map (\<lambda>(ty, ty'). is_sty_one P ty ty') ty_ty'_list) = Some bools
\<longrightarrow> list_all id bools"
by (induct ty_ty'_list, auto split: option.splits)
lemma tys_tys'_list: "\<And>bools ty ty'. \<lbrakk>lift_opts (map (\<lambda>(x, y). is_sty_one P x y) tys_tys'_list) = Some bools; length tys = length tys'; list_all id bools; (ty, ty') \<in> set tys_tys'_list\<rbrakk> \<Longrightarrow> is_sty_one P ty ty' = Some True"
apply(induct tys_tys'_list)
apply(simp)
apply(clarsimp split: option.splits)
by force
lemma [simp]: "(sty_many P tys tys') = (is_sty_many P tys tys' = Some True)"
apply(rule)
apply(erule sty_many.cases) apply(clarsimp simp add: is_sty_many_def split: option.splits) apply(rule)
apply(simp add: lift_opts_exists)
apply(force intro: lift_opts_all_true)
apply(clarsimp simp add: is_sty_many_def split: option.splits if_split_asm)
apply(rule_tac ty_ty'_list = "zip tys tys'" in sty_manyI)
apply(simp add: map_fst_zip[THEN sym])
apply(simp add: map_snd_zip[THEN sym])
apply(clarsimp)
apply(simp add: tys_tys'_list)
done
(* TODO: here will go all the other definitions for functions equivalent to
relations concerning well-formedness *)
definition
tr_x :: "T \<Rightarrow> x \<Rightarrow> x"
where
"tr_x T x = (case T x of None \<Rightarrow> x | Some x' \<Rightarrow> x')"
definition
tr_var :: "T \<Rightarrow> var \<Rightarrow> var"
where
"tr_var T var = (case tr_x T (x_var var) of x_this \<Rightarrow> var | x_var var' \<Rightarrow> var')"
primrec
tr_s_f :: "T \<Rightarrow> s \<Rightarrow> s" and
tr_ss_f :: "T \<Rightarrow> s list \<Rightarrow> s list"
where
"tr_s_f T (s_block ss) = s_block (tr_ss_f T ss)" |
"tr_s_f T (s_ass var x) = s_ass (tr_var T var) (tr_x T x)" |
"tr_s_f T (s_read var x f) = s_read (tr_var T var) (tr_x T x) f" |
"tr_s_f T (s_write x f y) = s_write (tr_x T x) f (tr_x T y)" |
"tr_s_f T (s_if x y s1 s2) = s_if (tr_x T x) (tr_x T y) (tr_s_f T s1) (tr_s_f T s2)" |
"tr_s_f T (s_call var x m ys) = s_call (tr_var T var) (tr_x T x) m (map (tr_x T) ys)" |
"tr_s_f T (s_new var ctx cl) = s_new (tr_var T var) ctx cl" |
"tr_ss_f T [] = []" |
"tr_ss_f T (s#ss) = tr_s_f T s # tr_ss_f T ss"
lemma [simp]: "(\<forall>x\<in>set s_s'_list. case x of (s_XXX, s_') \<Rightarrow> tr_s T s_XXX s_' \<and> tr_s_f T s_XXX = s_') \<longrightarrow>
tr_ss_f T (map fst s_s'_list) = map snd s_s'_list"
by (induct s_s'_list, auto)
lemma [simp]: "(\<forall>x\<in>set y_y'_list. case_prod (\<lambda>y_XXX. (=) (case T y_XXX of None \<Rightarrow> y_XXX | Some x' \<Rightarrow> x')) x) \<longrightarrow> map (tr_x T) (map fst y_y'_list) = map snd y_y'_list"
apply(induct y_y'_list) apply(simp) apply(clarsimp) apply(simp only: tr_x_def) done
lemma set_zip_tr[simp]: "(s, s') \<in> set (zip ss (tr_ss_f T ss)) \<longrightarrow> s' = tr_s_f T s" by (induct ss, auto)
lemma [iff]: "length ss = length (tr_ss_f T ss)" by (induct ss, auto)
lemma tr_ss_map:
"tr_ss_f T (map fst s_s'_list) = map snd s_s'_list \<and> (\<forall>x\<in>set s_s'_list. case_prod (tr_s T) x) \<and>
(a, b) \<in> set s_s'_list \<longrightarrow> tr_s T a (tr_s_f T a)"
apply(induct s_s'_list) by auto
lemma tr_f_to_rel: "\<forall>s'. tr_s_f T s = s' \<longrightarrow> tr_s T s s'"
apply(induct s rule: tr_s_f.induct)
apply(simp)
apply(clarsimp) apply(rule tr_s_var_assignI)
apply(clarsimp simp add: tr_x_def tr_var_def split: option.splits)
apply(simp add: tr_x_def)
apply(clarsimp) apply(rule tr_s_field_readI)
apply(clarsimp simp add: tr_x_def tr_var_def split: option.splits)
apply(simp add: tr_x_def)
apply(clarsimp) apply(rule tr_s_field_writeI)
apply(simp add: tr_x_def) apply(simp add: tr_x_def)
apply(clarsimp simp only: tr_s_f.simps) apply(rule tr_s_ifI)
apply(simp only: tr_x_def) apply(simp only: tr_x_def) apply(simp) apply(simp)
apply(clarsimp) apply(rule tr_s_newI)
apply(clarsimp simp add: tr_x_def tr_var_def split: option.splits)
apply(clarsimp)
apply(rename_tac var x m ys)
apply(cut_tac T = T and var = var and var' = "tr_var T var" and x = x and
x' = "tr_x T x" and y_y'_list = "zip ys (map (tr_x T) ys)" and
meth = m in tr_s_mcallI)
apply(clarsimp simp add: tr_x_def tr_var_def split: option.splits)
apply(simp only: tr_x_def) apply(clarsimp simp add: set_zip tr_x_def) apply(simp)
apply(simp)
apply(cut_tac T = T and s_s'_list = "[]" in tr_s_blockI[simplified])
apply(simp)
apply(clarsimp)
apply(clarsimp) apply(rename_tac s ss)
apply(cut_tac T = T and s_s'_list = "zip (s # ss) (tr_s_f T s # tr_ss_f T ss)" in tr_s_blockI[simplified])
apply(clarsimp) apply(rename_tac x x')
apply(frule set_zip_tr[THEN mp]) apply(clarsimp)
apply(erule in_set_zipE)
apply(erule_tac ?a1.0="T" and ?a2.0="s_block ss" and ?a3.0="s_block (tr_ss_f T ss)" in tr_s.cases)
apply(clarsimp)
apply(rule tr_ss_map[THEN mp], force)
apply(clarsimp)+
done
lemma tr_rel_f_eq: "((tr_s T s s') = (tr_s_f T s = s'))"
apply(rule)
apply(erule tr_s.induct)
apply(force simp add: tr_x_def tr_var_def split: option.splits)+
apply(cut_tac T = T and s = s in tr_f_to_rel) apply(simp)
done
end
diff --git a/thys/LightweightJava/Lightweight_Java_Proof.thy b/thys/LightweightJava/Lightweight_Java_Proof.thy
--- a/thys/LightweightJava/Lightweight_Java_Proof.thy
+++ b/thys/LightweightJava/Lightweight_Java_Proof.thy
@@ -1,1813 +1,1819 @@
(* Title: Lightweight Java, the proof
Authors: Rok Strnisa <rok@strnisa.com>, 2006
Matthew Parkinson <matt@matthewp.com>, 2006
Maintainer:
*)
theory Lightweight_Java_Proof
imports Lightweight_Java_Equivalence "HOL-Library.Infinite_Set"
begin
lemmas wf_intros = wf_object_wf_varstate_wf_heap_wf_config_wf_stmt_wf_meth_wf_class_common_wf_class_wf_program.intros [simplified]
lemmas wf_induct = wf_object_wf_varstate_wf_heap_wf_config_wf_stmt_wf_meth_wf_class_common_wf_class_wf_program.induct [simplified]
lemmas wf_config_normalI = wf_object_wf_varstate_wf_heap_wf_config_wf_stmt_wf_meth_wf_class_common_wf_class_wf_program.wf_allI [simplified]
lemmas wf_objectE = wf_object.cases[simplified]
lemmas wf_varstateE = wf_varstate.cases[simplified]
lemmas wf_heapE = wf_heap.cases[simplified]
lemmas wf_configE = wf_config.cases[simplified]
lemmas wf_stmtE = wf_stmt.cases[simplified]
lemmas wf_methE = wf_meth.cases[simplified]
lemmas wf_class_cE = wf_class_common.cases[simplified]
lemmas wf_classE = wf_class.cases[simplified]
lemmas wf_programE = wf_program.cases[simplified]
lemma wf_subvarstateI:
"\<lbrakk>wf_varstate P \<Gamma> H L; \<Gamma> x' = Some ty;
wf_object P H (Some v) (Some ty)\<rbrakk>
\<Longrightarrow> wf_varstate P \<Gamma> H (L (x' \<mapsto> v))"
apply(erule wf_varstateE) apply(rule wf_varstateI) apply(simp) apply(clarsimp)
done
lemma finite_super_varstate:
"finite (dom ((L ++ map_of (map (\<lambda>(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list))(x' \<mapsto> v_oid oid))) = finite (dom L)"
apply(induct y_cl_var_var'_v_list)
apply(clarsimp) apply(simp add: dom_def)
apply(subgoal_tac "{a. a \<noteq> x' \<longrightarrow> (\<exists>y. L a = Some y)} = {a. a = x' \<or> (\<exists>y. L a = Some y)}")
apply(simp add: Collect_disj_eq) apply(force)
apply(clarsimp simp add: map_add_def dom_def split: if_split_asm)
apply(subgoal_tac "{aa. aa \<noteq> x_var a \<longrightarrow> aa \<noteq> x' \<longrightarrow> (\<exists>y. case_option (L aa) Some (map_of (map (\<lambda>(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list) aa) = Some y)} =
{aa. aa = x_var a \<or> (aa = x' \<or> (\<exists>y. case_option (L aa) Some (map_of (map (\<lambda>(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list) aa) = Some y))}")
apply(simp add: Collect_disj_eq) apply(force)
done
lemma fst_map_elem:
"(y_k, ty_k, var_k, var'_k, v_k) \<in> set y_ty_var_var'_v_list \<Longrightarrow>
x_var var'_k \<in> fst ` (\<lambda>(y, ty, var, var', v). (x_var var', ty)) ` set y_ty_var_var'_v_list"
by (force)
lemma map_add_x_var[THEN mp]:
"var' \<in> set (map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list) \<longrightarrow>
x_var var' \<in> set (map (\<lambda>(y, cl, var, var', v). x_var var) y_cl_var_var'_v_list)"
by (induct y_cl_var_var'_v_list, auto)
lemma map_of_map_and_x[simp]:
"\<lbrakk>\<Gamma> x = Some ty_x; L x' = None; \<forall>x\<in>set y_cl_var_var'_v_list. (\<lambda>(y, cl, var, var', v). L (x_var var') = None) x;
\<forall>x\<in>Map.dom \<Gamma>. wf_object Pa H (L x) (\<Gamma> x)\<rbrakk> \<Longrightarrow>
(if x = x' then Some ty
else (\<Gamma> ++ map_of (map (\<lambda>((y, cl, var, var', v), y', ty). (x_var var', ty)) (zip y_cl_var_var'_v_list y_ty_list))) x) =
Some ty_x"
apply(subgoal_tac "x \<noteq> x'") apply(subgoal_tac "\<forall>(y, cl, var, var', v) \<in> set y_cl_var_var'_v_list. x_var var' \<noteq> x")
apply(case_tac "map_of (map (\<lambda>((y, cl, var, var', v), y', ty). (x_var var', ty)) (zip y_cl_var_var'_v_list y_ty_list)) x")
apply(clarsimp simp add: map_add_def)
apply(clarsimp) apply(drule map_of_SomeD) apply(clarsimp) apply(rename_tac y cl var var' v y' ty)
apply(drule_tac x = "(y, cl, var, var', v)" in bspec) apply(force simp add: set_zip)
apply(drule_tac x = "x_var var'" in bspec, simp add: domI) apply(erule wf_objectE, simp+)
apply(force elim: wf_objectE)+
done
lemma wf_stmt_in_G':
"(L x' = None \<and> (\<forall>x\<in>set y_cl_var_var'_v_list. (\<lambda>(y, cl, var, var', v). L (x_var var') = None) x)
\<and> (\<forall>x\<in>dom \<Gamma>. wf_object P H (L x) (\<Gamma> x)) \<and> wf_stmt P \<Gamma> s \<longrightarrow>
wf_stmt P ((\<Gamma> ++ map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)))(x' \<mapsto> ty)) s) \<and>
(L x' = None \<and> (\<forall>x\<in>set y_cl_var_var'_v_list. (\<lambda>(y, cl, var, var', v). L (x_var var') = None) x)
\<and> (\<forall>x\<in>dom \<Gamma>. wf_object P H (L x) (\<Gamma> x)) \<and> (\<forall>s\<in>set ss. wf_stmt P \<Gamma> s) \<longrightarrow>
(\<forall>s\<in>set ss. wf_stmt P ((\<Gamma> ++ map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)))(x' \<mapsto> ty)) s))"
apply(rule tr_s_f_tr_ss_f.induct)
apply(clarsimp, erule wf_stmtE, simp_all, simp add: wf_blockI)
apply(clarify, erule wf_stmtE, simp_all)
apply(erule sty_option.cases, force intro: wf_var_assignI sty_optionI)
apply(clarify, erule wf_stmtE, simp_all)
apply(erule sty_option.cases, force intro: wf_field_readI sty_optionI)
apply(clarify, erule wf_stmtE, simp_all)
apply(erule sty_option.cases, force intro: wf_field_writeI sty_optionI)
apply(clarify, erule wf_stmtE, simp_all)
apply(clarsimp)
apply(rule wf_ifI)
apply(erule disjE)
apply(rule disjI1, erule sty_option.cases, force intro: sty_optionI)
apply(rule disjI2, erule sty_option.cases, force intro: sty_optionI)
apply(assumption+)
apply(clarify, erule wf_stmtE, simp_all)
apply(erule sty_option.cases, force intro: wf_newI sty_optionI)
apply(clarify, erule wf_stmtE, simp_all)
apply(rule wf_mcallI, simp, simp, simp)
apply(clarsimp split del: if_split) apply(rename_tac y_k ty_k)
apply(drule_tac x = "(y_k, ty_k)" in bspec, simp, clarify)
apply(erule_tac ?a3.0 = "Some ty_k" in sty_option.cases)
apply(force intro: sty_optionI)
apply(clarify)
apply(erule sty_option.cases)
apply(rule sty_optionI, simp+)
done
lemma map_three_in_four:
"var_k \<in> set (map (\<lambda>(y, cl, var, u). var) y_cl_var_var'_v_list)
\<longrightarrow> (\<exists>y cl u. (y, cl, var_k, u) \<in> set y_cl_var_var'_v_list)"
by (induct y_cl_var_var'_v_list, auto)
lemma map_var:
"\<lbrakk>(cl_k, var_k, ty_k) \<in> set cl_var_ty_list;
map (\<lambda>(y, cl, var, u). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list\<rbrakk>
\<Longrightarrow> \<exists>y_k cl'_k u_k. (y_k, cl'_k, var_k, u_k) \<in> set y_cl_var_var'_v_list"
apply(subgoal_tac "var_k \<in> set (map (\<lambda>(y, cl, var, u). var) y_cl_var_var'_v_list)")
by (force simp add: map_three_in_four)+
lemma length_one_in_two:
"length y_ty_list = length (map fst y_ty_list)"
by (induct y_ty_list, auto)
lemma length_two_in_two:
"length y_ty_list = length (map snd y_ty_list)"
by (induct y_ty_list, auto)
lemma length_two_in_three:
"length cl_var_ty_list = length (map (\<lambda>(cl, var, ty). var) cl_var_ty_list)"
by (induct cl_var_ty_list, auto)
lemma length_three_in_three:
"length cl_var_ty_list = length (map (\<lambda>(cl, var, ty). ty) cl_var_ty_list)"
by (induct cl_var_ty_list, auto)
lemma length_three_in_four:
"length y_cl_var_var'_v_list = length (map (\<lambda>(y, cl, var, u). var) y_cl_var_var'_v_list)"
by (induct y_cl_var_var'_v_list, auto)
lemma length_one_in_five:
"length y_cl_var_var'_v_list = length (map (\<lambda>(y, cl, var, var', v). y) y_cl_var_var'_v_list)"
by (induct y_cl_var_var'_v_list, auto)
lemma length_three_in_five:
"length y_cl_var_var'_v_list = length (map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list)"
by (induct y_cl_var_var'_v_list, auto)
lemma map_length_list:
"length (map (\<lambda>((y, y'), yy, ty). (y', ty)) list) = length list"
by (induct list, auto)
lemma map_length_y:
"map (\<lambda>(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list \<Longrightarrow>
length y_cl_var_var'_v_list = length y_ty_list"
by (simp only: length_one_in_five length_one_in_two)
lemma map_length_y':
"map fst y_y'_list = map fst y_ty_list \<Longrightarrow> length y_y'_list = length y_ty_list"
by (simp only: length_one_in_two length_two_in_two)
lemma map_length_var:
"map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list \<Longrightarrow>
length y_cl_var_var'_v_list = length cl_var_ty_list"
by (simp only: length_three_in_five length_two_in_three)
lemma map_length_ty:
"map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list \<Longrightarrow>
length cl_var_ty_list = length y_ty_list"
by (simp only: length_three_in_three length_two_in_two)
lemma map_length_var_ty:
"\<lbrakk>map (\<lambda>(y, cl, var, u). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list;
map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list\<rbrakk> \<Longrightarrow>
length y_cl_var_var'_v_list = length y_ty_list"
apply(rule_tac s = "length cl_var_ty_list" in trans)
apply (simp only: length_three_in_four length_two_in_three)
apply(simp only: length_three_in_three length_two_in_two)
done
lemma fun_eq_fst:
"(fst \<circ> (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y))) = (\<lambda>((y, cl, var, var', v), y', y). (x_var var'))"
by (simp add: fun_eq_iff)
lemma fst_zip_eq:
"length y_cl_var_var'_v_list = length y_ty_list \<Longrightarrow>
(map fst (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list))) =
(map (\<lambda>(y, cl, var, var', v). x_var var') y_cl_var_var'_v_list)"
apply(simp)
apply(simp add: fun_eq_fst)
apply(rule nth_equalityI)
apply(force)
apply(clarsimp)
apply(frule nth_mem)
apply(subgoal_tac "\<exists>x. (y_cl_var_var'_v_list ! i) = x")
apply(force)
apply(rule_tac x = "y_cl_var_var'_v_list ! i" in exI)
apply(simp)
done
lemma distinct_x_var:
"distinct (map (\<lambda>(y, cl, var, var', v). x_var var) y_cl_var_var'_v_list) =
distinct (map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list)"
apply (induct y_cl_var_var'_v_list)
by (force simp add: contrapos_np)+
lemma distinct_x_var':
"distinct (map (\<lambda>(y, cl, var, var', v). x_var var') y_cl_var_var'_v_list) =
distinct (map (\<lambda>(y, cl, var, var', v). var') y_cl_var_var'_v_list)"
apply(induct y_cl_var_var'_v_list)
by (force simp add: contrapos_np)+
lemma map_fst_two:
"map fst (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) =
map (\<lambda>(y, cl, var, var', v). x_var var) y_cl_var_var'_v_list"
by (induct y_cl_var_var'_v_list, auto)
lemma map_fst_two':
"map fst (map (\<lambda>(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list) =
map (\<lambda>(y, cl, var, var', y). x_var var') y_cl_var_var'_v_list"
by (induct y_cl_var_var'_v_list, auto)
lemma map_fst_var':
"distinct (map fst (map (\<lambda>(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list)) =
distinct (map (\<lambda>(y, cl, var, var', v). var') y_cl_var_var'_v_list)"
by (simp only: map_fst_two' distinct_x_var')
lemma distinct_var:
"\<lbrakk>distinct (map (\<lambda>(cl, var, ty). var) cl_var_ty_list);
map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list\<rbrakk> \<Longrightarrow>
distinct (map fst (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list))"
by (simp only: map_fst_two distinct_x_var)
lemma distinct_var':
"\<lbrakk>map (\<lambda>(y, cl, var, u). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list;
map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list\<rbrakk> \<Longrightarrow>
distinct (map fst (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list))) =
distinct (map (\<lambda>(y, cl, var, var', v). var') y_cl_var_var'_v_list)"
by (simp only: map_length_var_ty fst_zip_eq distinct_x_var')
lemma weaken_list_var:
"map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list =
map (\<lambda>(y, cl, var, u). var) y_cl_var_var'_v_list"
by (induct y_cl_var_var'_v_list, auto)
lemma weaken_list_fd:
"map (\<lambda>(y, cl, var, var', v). vd_def cl var) list = map (\<lambda>(y, cl, var, u). vd_def cl var) list"
by (induct list, auto)
lemma weaken_list_y:
"map (\<lambda>(y, cl, var, var', v). y) y_cl_var_var'_v_list =
map fst y_cl_var_var'_v_list"
by (induct y_cl_var_var'_v_list, auto)
lemma wf_cld_from_lu:
"\<lbrakk>wf_program P; find_cld P ctx fqn (Some (ctx', cld'))\<rbrakk> \<Longrightarrow> wf_class P cld'"
apply(clarsimp) apply(drule find_to_mem) apply(erule wf_programE) by (simp add: member_def)
lemma meth_def_in_set[rule_format]:
"find_meth_def_in_list meth_defs meth (Some meth_def) \<longrightarrow> meth_def \<in> set meth_defs"
apply(clarsimp) apply(induct meth_defs) by (auto split: meth_def.splits meth_sig.splits if_split_asm)
lemma find_meth_def_in_list_mem[rule_format, simp]:
"find_meth_def_in_list_f meth_defs meth = Some meth_def \<longrightarrow> meth_def \<in> set meth_defs"
apply(induct meth_defs) apply(auto split: meth_def.splits meth_sig.splits) done
lemma find_meth_def_in_path_ex_cld[rule_format]:
"find_meth_def_in_path_f ctxclds meth = Some (ctx, meth_def)
\<longrightarrow> (\<exists>cld meth_defs. (ctx, cld) \<in> set ctxclds \<and> class_methods_f cld = meth_defs \<and> meth_def \<in> set meth_defs)"
apply(induct ctxclds) apply(simp) apply(clarsimp)
apply(clarsimp split: option.splits) apply(rule_tac x = cld in exI) apply(force)
apply(force)
done
lemma map_ctx_cld_dcl_two[simp]:
"ctxclds = map (\<lambda>(ctx, cld, dcl). (ctx, cld)) (map (\<lambda>(ctx, cld). (ctx, cld, classname_f cld)) ctxclds)"
by (induct ctxclds, auto)
lemma clds_in_path_exist:
"find_path_f P ctx' fqn = Some path \<Longrightarrow> (\<forall>(ctx, cld) \<in> set path. cld \<in> set P)"
apply(clarify)
apply(drule all_in_path_found) apply(simp) apply(clarsimp)
apply(drule find_to_mem) apply(simp add: member_def)
done
lemma wf_meth_defs_in_wf_class[rule_format]:
"meth_def \<in> set (class_methods_f cld) \<and> wf_class P cld
\<longrightarrow> wf_meth P (ty_def ctx_def (class_name_f cld)) meth_def"
apply(clarsimp)
apply(erule wf_classE) apply(clarsimp simp: class_methods_f_def)
apply(erule wf_class_cE) apply(clarsimp simp: class_name_f_def)
apply(drule(1)bspec, simp)
done
lemma map_ctxclds: "ctxclds = map ((\<lambda>(ctx_XXX, cld_XXX, dcl_XXX). (ctx_XXX, cld_XXX)) \<circ> (\<lambda>(ctx, cld). (ctx, cld, class_name_f cld))) ctxclds"
by (induct ctxclds, auto)
lemma wf_method_from_find'[rule_format]:
"wf_program P \<and> find_path_ty_f P (ty_def ctxa dcl) = Some ctxclds \<and> find_meth_def_in_path_f ctxclds meth = Some (ctx, meth_def)
\<longrightarrow> (\<exists>dcla. is_sty_one P (ty_def ctxa dcl) (ty_def ctx dcla) = Some True \<and> wf_meth P (ty_def ctx_def dcla) meth_def)"
apply(clarsimp) apply(drule find_meth_def_in_path_ex_cld) apply(clarsimp)
apply(rule_tac x = "class_name_f cld" in exI)
apply(rule)
apply(rule_tac ctx_cld_dcl_list = "map (\<lambda>(ctx, cld). (ctx, cld, class_name_f cld)) ctxclds" in sty_dclI[simplified])
apply(simp add: map_ctxclds)
apply(clarsimp)
apply(force)
apply(drule clds_in_path_exist)
apply(simp)
apply(erule wf_programE)
apply(drule_tac x = "(ctx, cld)" in bspec, simp)
apply(drule_tac x = cld in bspec, simp)
apply(simp add: wf_meth_defs_in_wf_class)
done
(* fix the ctx problem here\<dots> *)
lemma wf_method_from_find:
"\<lbrakk>wf_program P; find_meth_def P ty_x_d meth (Some (ctx, meth_def))\<rbrakk> \<Longrightarrow>
\<exists>dcl. sty_one P ty_x_d (ty_def ctx dcl) \<and> wf_meth P (ty_def ctx_def dcl) meth_def"
apply(erule find_meth_def.cases, clarsimp+)
apply(induct ty_x_d) apply(simp) apply(rename_tac ctx' dcl' ctxclds)
apply(cut_tac P = P and ctxa = ctx' and dcl = dcl' and ctxclds = ctxclds and
meth = meth and ctx = ctx and meth_def = meth_def in wf_method_from_find') apply(simp)
apply(simp)
done
lemma find_type_lift_opts[rule_format]:
"(\<forall>x\<in>set cl_var_ty_list. (\<lambda>(cl_k, var_k, ty_k). find_type_f P ctx cl_k = Some ty_k) x) \<longrightarrow>
lift_opts (map (case_vd (\<lambda>clk vark. find_type_f P ctx clk)) (map (\<lambda>(cl_k, var_k, ty_k). vd_def cl_k var_k) cl_var_ty_list)) =
Some (map (\<lambda>(cl, var, ty). ty) cl_var_ty_list)"
by (induct cl_var_ty_list, auto)
lemma find_md_in_pre_path[rule_format]:
"find_meth_def_in_path_f path m = Some ctxmd \<longrightarrow> (\<forall>path'. find_meth_def_in_path_f (path @ path') m = Some ctxmd)"
by (induct path, auto split: option.splits)
lemma lift_opts_map:
"\<forall>x\<in>set cl_var_ty_list. (\<lambda>(cl_k, var_k, ty_k). find_type_f P ctx cl_k = Some ty_k) x \<Longrightarrow>
lift_opts (map (case_vd (\<lambda>clk vark. find_type_f P ctx clk)) (map (\<lambda>(cl_k, var_k, ty_k). vd_def cl_k var_k) cl_var_ty_list)) = Some (map (\<lambda>(cl, var, ty). ty) cl_var_ty_list)"
by (induct cl_var_ty_list, auto)
lemma m_in_list[rule_format]:
"(\<forall>x\<in>set meth_def_meth_list. case_prod (\<lambda>meth_def. (=) (method_name_f meth_def)) x) \<and>
find_meth_def_in_list_f (map fst meth_def_meth_list) m = Some md \<longrightarrow>
m \<in> snd ` set meth_def_meth_list"
by (induct meth_def_meth_list, auto simp add: method_name_f_def split: meth_def.splits meth_sig.splits)
lemma m_image_eq[rule_format]:
"meth_def_def (meth_sig_def cl m list2) meth_body \<in> set meth_defs \<longrightarrow>
m \<in> case_meth_def (\<lambda>meth_sig meth_body. case meth_sig of meth_sig_def cl meth vds \<Rightarrow> meth) ` set meth_defs"
by (induct meth_defs, auto)
lemma fmdip_to_mip[rule_format]:
"find_meth_def_in_path_f path m = Some ctxmd \<longrightarrow> m \<in> set (methods_in_path_f (map snd path))"
apply(induct path) apply(simp) apply(clarsimp simp add: class_methods_f_def split: option.splits cld.splits)
apply(rename_tac aa bb)
apply(case_tac aa) apply(rename_tac meth_sig meth_body) apply(case_tac meth_sig) apply(clarsimp) apply(frule find_md_m_match') apply(clarsimp)
apply(frule find_meth_def_in_list_mem) apply(frule m_image_eq) apply(assumption)
done
lemma lift_opts_for_all[rule_format]:
"lift_opts (map (case_vd (\<lambda>clk vark. find_type_f P ctx clk)) (map (\<lambda>(cl, var, ty). vd_def cl var) cl_var_ty_list)) = None \<and>
(\<forall>x\<in>set cl_var_ty_list. (\<lambda>(cl, var, ty). find_type_f P ctx_def cl = Some ty) x) \<longrightarrow> False"
apply(induct cl_var_ty_list) apply(simp) apply(clarsimp) apply(case_tac ctx) apply(simp) done
lemma mty_preservation'''':
"\<lbrakk>find_cld_f P ctx (fqn_def dcl') = Some (ctx, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds');
find_cld_f P ctx (fqn_def dcl'') = Some (ctx, cld_def dcl'' cl'' fds'' mds'');
find_path_rec_f P ctx cl'' [(ctx, cld_def dcl'' cl'' fds'' mds'')] = Some path;
find_meth_def_in_path_f path m = Some (ctx'', meth_def_def (meth_sig_def cl_r m' vds) mb);
find_type_f P ctx'' cl_r = Some ty_r; lift_opts (map (case_vd (\<lambda>clk vark. find_type_f P ctx'' clk)) vds) = Some tys;
lift_opts (map (case_vd (\<lambda>clk vark. find_type_f P ctx' clk)) vds') = Some tys'; find_type_f P ctx' cl_r' = Some ty_r';
find_meth_def_in_path_f ((ctx, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds') # path) m = Some (ctx', meth_def_def (meth_sig_def cl_r' m'' vds') mb');
wf_program P\<rbrakk>
\<Longrightarrow> tys' = tys \<and> ty_r' = ty_r"
using [[hypsubst_thin = true]]
apply(clarsimp simp add: class_methods_f_def split: option.splits)
apply(drule wf_cld_from_lu) apply(simp) apply(erule wf_classE) apply(erule wf_class_cE) apply(clarsimp)
apply(subgoal_tac "m \<in> snd ` set meth_def_meth_list")
apply(clarsimp) apply(rename_tac md_m m) apply(drule_tac x = "(md_m, m)" in bspec, simp)+ apply(clarsimp split: option.splits)
apply(case_tac ctx') apply(clarsimp)
apply(case_tac md_m) apply(rename_tac meth_sig meth_body) apply(case_tac meth_sig) apply(rename_tac ms mb_m cl_r_m m vds_m) apply(clarsimp simp add: method_name_f_def)
apply(erule_tac Q = "tys' = tys \<longrightarrow> ty_r' \<noteq> ty_r" in contrapos_pp) apply(clarsimp)
apply(clarsimp simp add: methods_f_def find_path_f_def superclass_name_f_def split: option.splits if_split_asm)
apply(frule fmdip_to_mip) apply(clarsimp) apply(rename_tac m mty mty') apply(drule_tac x = "(m, mty, mty')" in bspec, simp)+
apply(clarsimp) apply(frule_tac f = snd in imageI) apply(simp) apply(thin_tac "m \<in> snd ` set meth_def_meth_list") apply(clarsimp)
apply(clarsimp simp add: mtype_f_def split: option.splits meth_def.splits meth_sig.splits)
apply(clarsimp simp add: find_meth_def_f_def find_path_f_def superclass_name_f_def split: option.splits if_split_asm)
apply(frule path_append) apply(rename_tac ad) apply(frule_tac path = ad in path_append) apply(clarsimp)
apply(frule_tac prefix = "[(ctx_def, cld_def dcl'' cl'' fds'' mds'')]" and suffix = path'' and
prefix' = "[(ctx_def, cld_def dcl' (cl_fqn (fqn_def dcl'')) (map (\<lambda>(cl, f, ty). fd_def cl f) cl_f_ty_list) (map fst meth_def_meth_list)),
(ctx_def, cld_def dcl'' cl'' fds'' mds'')]" in fpr_same_suffix[rule_format]) apply(simp)
apply(clarsimp simp add: class_methods_f_def)
apply(rule m_in_list) apply(simp)
done
lemma mty_preservation'''[rule_format]:
"\<forall>suffix dcl. find_path_rec_f P ctx cl prefix = Some (prefix @ suffix) \<longrightarrow> wf_program P \<and> cl = cl_fqn (fqn_def dcl) \<longrightarrow>
(\<forall>(ctx', cld') \<in> set suffix.
(\<forall>prefix' suffix'. find_path_rec_f P ctx' (cl_fqn (fqn_def (class_name_f cld'))) prefix' = Some (prefix' @ suffix') \<and>
mtype_f P (ty_def ctx' (class_name_f cld')) m = Some mty \<longrightarrow> mtype_f P (ty_def ctx dcl) m = Some mty))"
+supply [[simproc del: defined_all]]
apply(induct_tac P ctx cl prefix rule: find_path_rec_f.induct)
apply(simp)
apply(clarsimp split: option.splits)
apply(frule path_append) apply(clarify)
apply(clarsimp)
apply(erule disjE)
apply(clarify)
apply(clarsimp simp add: class_name_f_def split: cld.splits)
apply(frule find_cld_name_eq) apply(frule_tac ctx = a in find_cld_name_eq)
apply(clarsimp simp add: class_name_f_def split: cld.splits)
apply(case_tac b) apply(rename_tac dcl' cl' fds' mds')
apply(clarsimp simp add: superclass_name_f_def) apply(case_tac cl') apply(simp) apply(rename_tac fqn) apply(case_tac fqn) apply(rename_tac dcl'') apply(clarsimp)
apply(clarsimp split: option.splits)
apply(subgoal_tac "\<forall>aa b ab bb. (a = ab \<and> cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds' = bb \<and> aa = ab \<and> b = bb)
\<longrightarrow> (\<forall>suffix. find_path_rec_f P ab (case bb of cld_def dcl cl fds mds \<Rightarrow> cl) (path @ [(ab, bb)]) = Some (path @ (ab, bb) # suffix) \<longrightarrow>
(\<forall>dcl. (case bb of cld_def dcl cl fds mds \<Rightarrow> cl) = cl_fqn (fqn_def dcl) \<longrightarrow>
(\<forall>x\<in>set suffix.
(\<lambda>(ctx', cld').
(\<exists>prefix' suffix'.
case_option None (case_prod (\<lambda>ctx' cld. find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)]))) (find_cld_f P ctx' (fqn_def (class_name_f cld'))) =
Some (prefix' @ suffix')) \<and>
mtype_f P (ty_def ctx' (class_name_f cld')) m = Some mty \<longrightarrow>
mtype_f P (ty_def ab dcl) m = Some mty)
x)))")
defer
apply(force)
apply(thin_tac "\<And>aa b x y.
\<lbrakk>a = aa \<and> cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds' = b; x = aa \<and> y = b\<rbrakk>
\<Longrightarrow> \<forall>suffix.
find_path_rec_f P aa (case b of cld_def dcl cl fds mds \<Rightarrow> cl) (path @ [(aa, b)]) =
Some (path @ (aa, b) # suffix) \<longrightarrow>
(\<forall>dcl. (case b of cld_def dcl cl fds mds \<Rightarrow> cl) = cl_fqn (fqn_def dcl) \<longrightarrow>
(\<forall>x\<in>set suffix.
case x of
(ctx', cld') \<Rightarrow>
(\<exists>prefix' suffix'.
case_option None
(\<lambda>(ctx', cld).
find_path_rec_f P ctx' (superclass_name_f cld)
(prefix' @ [(ctx', cld)]))
(find_cld_f P ctx' (fqn_def (class_name_f cld'))) =
Some (prefix' @ suffix')) \<and>
mtype_f P (ty_def ctx' (class_name_f cld')) m = Some mty \<longrightarrow>
mtype_f P (ty_def aa dcl) m = Some mty))")
apply(clarsimp)
apply(drule_tac x = "(aa, ba)" in bspec, simp)
apply(clarsimp)
apply(subgoal_tac "(\<exists>prefix' suffix'. find_path_rec_f P ab (case bb of cld_def dcl cl fds mds \<Rightarrow> cl) (prefix' @ [(ab, bb)]) = Some (prefix' @ suffix'))")
apply(erule impE) apply(simp add: superclass_name_f_def) apply(thin_tac "\<exists>prefix suffix'. P prefix suffix'" for P)
apply(thin_tac "mtype_f P (ty_def aa (class_name_f ba)) m = Some mty")
apply(frule find_cld_name_eq) apply(clarsimp split: option.splits if_split_asm)
apply(clarsimp simp add: mtype_f_def split: option.splits meth_def.splits meth_sig.splits)
apply(rule)
apply(clarsimp simp add: find_meth_def_f_def) apply(subgoal_tac "\<exists>path'. find_path_f P a (cl_fqn (fqn_def dcl')) = Some path'")
apply(case_tac b) apply(clarsimp simp add: find_path_f_def superclass_name_f_def split: if_split_asm option.splits)
apply(rename_tac meth_body cl list_char1 list_vd ty list_ty list_char2 cla list_fd list_md list_p1 path')
apply(frule_tac path = path' in path_append) apply(rename_tac ag ah) apply(frule_tac path = ag in path_append) apply(clarify)
apply(frule_tac suffix = path''a and suffix' = path''b and prefix' = "[(ac, cld_def list_char2 cla list_fd list_md)]" in fpr_same_suffix[rule_format], force)
apply(simp) apply(clarify) apply(clarsimp simp add: class_methods_f_def split: option.splits)
apply(clarsimp simp add: find_path_f_def superclass_name_f_def split: option.splits if_split_asm)
apply(frule_tac path' = "(path @ [(a, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds'), (ac, b)])" in path_append) apply(clarsimp)
apply(frule_tac suffix = path''a and prefix' = "[(a, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds'), (ac, b)]" in fpr_same_suffix') back apply(simp) apply(simp)
apply(clarsimp) apply(rule)
apply(clarsimp) apply(frule_tac ty_x_d = "(ty_def a dcl')" in wf_method_from_find[simplified]) apply(simp) apply(clarsimp)
apply(erule wf_methE) apply(rename_tac ag ah ai aj ak al am an ao ap aq ar as at au av) apply(case_tac ag) apply(clarsimp)
apply(clarsimp) apply(rule)
apply(clarsimp) apply(frule_tac ty_x_d = "(ty_def a dcl')" in wf_method_from_find[simplified]) apply(simp) apply(clarsimp)
apply(erule wf_methE) apply(clarsimp) apply(cut_tac lift_opts_for_all) apply(assumption) apply(rule) apply(simp) apply(assumption)
apply(clarsimp)
apply(thin_tac "find_path_rec_f P ab (case bb of cld_def dcl cl fds mds \<Rightarrow> cl) (prefix' @ [(ab, bb)]) = Some (prefix' @ suffix')")
apply(thin_tac "find_cld_f P aa (fqn_def (class_name_f ba)) = Some (ab, bb)") apply(thin_tac "(aa, ba) \<in> set path''")
apply(case_tac b) apply(frule_tac dcl = dcl'' in find_cld_name_eq) apply(clarsimp simp add: superclass_name_f_def)
apply(rename_tac ctx ctx'' mb cl_r m' vds ty_r tys ctx' cl_r' m'' vds' ty_r' tys' mb' dcl'' cl'' fds'' mds'')
apply(clarsimp simp add: find_meth_def_f_def find_path_f_def superclass_name_f_def split: option.splits if_split_asm)
apply(thin_tac "find_path_rec_f P ctx cl'' (path @ [(ctx, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds'), (ctx, cld_def dcl'' cl'' fds'' mds'')]) =
Some (path @ (ctx, cld_def dcl' (cl_fqn (fqn_def dcl'')) fds' mds') # path'')") apply(clarsimp) apply(rename_tac path'' path')
apply(frule_tac path = path' in path_append) apply(frule_tac path = path'' in path_append) apply(clarify)
apply(frule_tac suffix = path''a and suffix' = path''b and prefix' = "[(ctx, cld_def dcl'' cl'' fds'' mds'')]" in fpr_same_suffix[rule_format], force)
apply(clarsimp simp del: find_meth_def_in_path_f.simps) apply(rename_tac path'')
apply(frule_tac ctx' = ctx' and m'' = m'' and cl'' = cl'' and fds'' = fds'' and ctx'' = ctx'' and
path = "(ctx, cld_def dcl'' cl'' fds'' mds'') # path''" in mty_preservation'''', simp+)
apply(rule_tac x = prefix' in exI) apply(clarsimp simp add: superclass_name_f_def)
done
lemma mty_preservation'':
"\<lbrakk>wf_program P; find_path_f P ctx (cl_fqn (fqn_def dcl)) = Some path; (ctx', cld') \<in> set path\<rbrakk> \<Longrightarrow>
\<forall>prefix' suffix'.
find_path_rec_f P ctx' (cl_fqn (fqn_def (class_name_f cld'))) prefix' = Some (prefix' @ suffix') \<and>
mtype_f P (ty_def ctx' (class_name_f cld')) m = Some mty \<longrightarrow> mtype_f P (ty_def ctx dcl) m = Some mty"
apply(cut_tac x = "(ctx', cld')" in mty_preservation'''[of _ _ _ "[]", simplified])
apply(unfold find_path_f_def) apply(assumption) apply(simp+) done
lemma mty_preservation':
"\<lbrakk>wf_program P; find_path_f P ctx (cl_fqn (fqn_def dcl)) = Some path; (ctx', cld') \<in> set path;
find_path_f P ctx' (cl_fqn (fqn_def (class_name_f cld'))) = Some path'; mtype_f P (ty_def ctx' (class_name_f cld')) m = Some mty\<rbrakk>
\<Longrightarrow> mtype_f P (ty_def ctx dcl) m = Some mty"
apply(cut_tac path = path and suffix' = path' in mty_preservation''[rule_format, of _ _ _ _ _ _ "[]"])
apply(assumption+) apply(rule) apply(simp add: find_path_f_def) apply(assumption+) done
lemma lift_opts_for_all_true[rule_format]:
"\<forall>y_ty_list.
(\<forall>x\<in>set cl_var_ty_list. (\<lambda>(cl_k, var_k, ty_k). find_type_f P ctx cl_k = Some ty_k) x) \<and>
lift_opts (map (case_vd (\<lambda>clk vark. find_type_f P ctx clk)) (map (\<lambda>(cl_k, var_k, ty_k). vd_def cl_k var_k) cl_var_ty_list)) = Some (map snd y_ty_list) \<longrightarrow>
map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list"
by (induct cl_var_ty_list, auto split: option.splits)
lemma mty_preservation:
"\<lbrakk>wf_program P;
find_meth_def_f P ty_x_d meth = Some (ctx, meth_def_def (meth_sig_def cl_r meth (map (\<lambda>(cl_k, var_k, ty_k). vd_def cl_k var_k) cl_var_ty_list)) meth_body);
find_type_f P ctx cl_r = Some ty_r_d; \<forall>x\<in>set cl_var_ty_list. (\<lambda>(cl_k, var_k, ty_k). find_type_f P ctx cl_k = Some ty_k) x;
mtype_f P ty_x_s meth = Some (mty_def (map snd y_ty_list) ty_r_s); is_sty_one P ty_x_d ty_x_s = Some True\<rbrakk>
\<Longrightarrow> ty_r_d = ty_r_s \<and> map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list"
apply(case_tac "ty_x_d = ty_x_s")
apply(clarsimp) apply(clarsimp simp add: mtype_f_def find_type_lift_opts[simplified] split: option.splits)
apply(simp add: is_sty_one_def split: option.splits)
apply(case_tac ty_x_s) apply(clarsimp simp add: mtype_f_def find_meth_def_f_def)
apply(clarsimp) apply(rename_tac path_d ctx_s cld_s)
apply(case_tac ty_x_d) apply(clarsimp) apply(clarsimp) apply(rename_tac ctx_d dcl_d)
apply(clarsimp simp add: find_meth_def_f_def split: option.splits)
apply(subgoal_tac "\<exists>path_s. find_path_f P ctx_s (cl_fqn (fqn_def (class_name_f cld_s))) = Some path_s")
apply(clarify) apply(frule_tac path = path_d and path' = path_s in mty_preservation') apply(simp+)
apply(clarsimp simp add: mtype_f_def split: option.splits meth_def.splits meth_sig.splits)
apply(clarsimp simp add: find_meth_def_f_def)
apply(rule_tac P = P and ctx = ctx in lift_opts_for_all_true) apply(simp)
apply(clarsimp simp add: mtype_f_def find_meth_def_f_def split: option.splits)
done
lemma map_not_mem:
"\<lbrakk>(y_k, ty_k, var_k, var'_k, v_k) \<in> set y_ty_var_var'_v_list; x' \<notin> (\<lambda>(y, ty, var, var', v). x_var var') ` set y_ty_var_var'_v_list\<rbrakk>
\<Longrightarrow> x_var var'_k \<noteq> x'"
by (force)
lemma map_fst:
"map fst (map (\<lambda>(y, ty, var, var', v). (x_var var', ty)) y_ty_var_var'_v_list) = map (\<lambda>(y, ty, var, var', v). x_var var') y_ty_var_var'_v_list"
by (force)
lemma supertype_top:
"sty_one P ty ty' \<Longrightarrow> ty = ty_top \<longrightarrow> ty' = ty_top"
by (induct rule: sty_one.induct, auto)
lemma top_not_subtype[rule_format]:
"sty_one P ty_top (ty_def ctx dcl) \<longrightarrow> False"
by (rule, drule supertype_top, simp)
lemma f_in_found_fds:
"ftype_in_fds_f P ctx fds f = ty_opt_bot_opt (Some ty_f) \<longrightarrow> f \<in> case_fd (\<lambda>cl f. f) ` set fds"
by (induct fds, auto split: fd.splits)
lemma ftype_fields[rule_format]:
"ftype_in_path_f P path_s f = Some ty_f \<longrightarrow> f \<in> set (fields_in_path_f path_s)"
apply(induct path_s) apply(simp) apply(clarsimp split: option.splits)
apply(case_tac b) apply(clarsimp) apply(rename_tac ctx path_s dcl cl fds mds)
apply(clarsimp simp add: class_fields_f_def split: ty_opt_bot.splits option.splits) apply(simp add: f_in_found_fds)
done
lemma no_field_hiding''':
"\<forall>suffix. find_path_rec_f P ctx cl prefix = Some (prefix @ suffix) \<longrightarrow> wf_program P \<longrightarrow>
(\<forall>(ctx, cld) \<in> set suffix.
(\<forall>prefix' suffix'. find_path_rec_f P ctx (cl_fqn (fqn_def (class_name_f cld))) prefix' = Some (prefix' @ suffix') \<and>
f \<in> set (fields_in_path_f suffix') \<longrightarrow> f \<in> set (fields_in_path_f suffix)))"
+supply [[simproc del: defined_all]]
apply(induct_tac P ctx cl prefix rule: find_path_rec_f.induct)
apply(simp)
apply(clarsimp split: option.splits)
apply(subgoal_tac "\<forall>aa ba aaa bb.
(a = aaa \<and> b = bb \<and> aa = aaa \<and> ba = bb)
\<longrightarrow> (\<forall>suffix. find_path_rec_f P aaa (superclass_name_f bb) (path @ [(aaa, bb)]) = Some (path @ (aaa, bb) # suffix) \<longrightarrow>
(\<forall>x\<in>set suffix.
(\<lambda>(ctx, cld). (\<exists>prefix' suffix'.
case_option None (case_prod (\<lambda>ctx' cld. find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)]))) (find_cld_f P ctx (fqn_def (class_name_f cld))) =
Some (prefix' @ suffix') \<and>
f \<in> set (fields_in_path_f suffix')) \<longrightarrow>
f \<in> set (fields_in_path_f suffix))
x))")
defer
apply(thin_tac _)+
apply(force)
apply(thin_tac "\<And>aa ba x y.
\<lbrakk>a = aa \<and> b = ba; x = aa \<and> y = ba\<rbrakk>
\<Longrightarrow> \<forall>suffix.
find_path_rec_f P aa (superclass_name_f ba) (path @ [(aa, ba)]) =
Some (path @ (aa, ba) # suffix) \<longrightarrow>
(\<forall>x\<in>set suffix.
case x of
(ctx, cld) \<Rightarrow>
(\<exists>prefix' suffix'.
case_option None
(\<lambda>(ctx', cld).
find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)]))
(find_cld_f P ctx (fqn_def (class_name_f cld))) =
Some (prefix' @ suffix') \<and>
f \<in> set (fields_in_path_f suffix')) \<longrightarrow>
f \<in> set (fields_in_path_f suffix))")
apply(clarsimp)
apply(frule path_append) apply(clarify)
apply(erule_tac x = "tl suffix" in allE)
apply(clarsimp)
apply(erule disjE)
apply(clarify)
apply(case_tac fqn)
apply(clarsimp simp del: fmdip_cons simp add: class_name_f_def split: cld.splits)
apply(frule find_cld_name_eq) apply(frule_tac ctx = a in find_cld_name_eq)
apply(clarsimp simp del: fmdip_cons simp add: class_name_f_def split: cld.splits)
apply (rename_tac list1 cl list2 list3)
apply(frule_tac path = "prefix' @ suffix'" in path_append) apply(clarsimp simp del: fmdip_cons)
apply(frule_tac prefix = "path @ [(ab, cld_def list1 cl list2 list3)]" and suffix = path'' and
prefix' = "prefix' @ [(ab, cld_def list1 cl list2 list3)]" and suffix' = path''a in fpr_same_suffix[rule_format])
apply(clarsimp) apply(simp)
apply(drule_tac x = "(aa, ba)" in bspec, simp)
apply(force)
done
lemma no_field_hiding'':
"\<lbrakk>find_path_f P ctx' cl = Some suffix; wf_program P; (ctx, cld) \<in> set suffix\<rbrakk>
\<Longrightarrow> \<forall>prefix' suffix'.
find_path_rec_f P ctx (cl_fqn (fqn_def (class_name_f cld))) prefix' = Some (prefix' @ suffix') \<and> f \<in> set (fields_in_path_f suffix') \<longrightarrow>
f \<in> set (fields_in_path_f suffix)"
apply(cut_tac x = "(ctx, cld)" in no_field_hiding'''[of _ _ _ "[]", rule_format]) apply(simp add: find_path_f_def)+ done
lemma no_field_hiding':
"\<lbrakk>find_path_f P ctx' cl = Some path; wf_program P; (ctx, cld) \<in> set path;
find_path_f P ctx (cl_fqn (fqn_def (class_name_f cld))) = Some path'; f \<in> set (fields_in_path_f path')\<rbrakk>
\<Longrightarrow> f \<in> set (fields_in_path_f path)"
apply(cut_tac no_field_hiding''[rule_format, of _ _ _ _ _ _ "[]"]) apply(assumption+) apply(unfold find_path_f_def) apply(rule) apply(simp+) done
lemma no_field_hiding:
"\<lbrakk>wf_program P; is_sty_one P ty_x_d ty_x_s = Some True; ftype_f P ty_x_s f = Some ty_f; fields_f P ty_x_d = Some fields_oid\<rbrakk>
\<Longrightarrow> f \<in> set fields_oid"
apply(case_tac ty_x_s) apply(simp add: ftype_f_def)
apply(simp add: is_sty_one_def split: option.splits)
apply(case_tac ty_x_d) apply(simp) apply(clarsimp) apply(rename_tac ctx_s dcl_s path_d ctx_d dcl_d)
apply(simp add: ftype_f_def split: option.splits)
apply(clarsimp simp add: fields_f_def) apply(drule ftype_fields)
apply(frule no_field_hiding') apply(simp add: ftype_fields class_name_f_def)+
done
lemma ftype_pre_path[rule_format]:
"ftype_in_path_f P path_s f = Some ty_f \<longrightarrow> ftype_in_path_f P (path_s @ path'') f = Some ty_f"
by (induct path_s, auto split: option.splits ty_opt_bot.splits)
lemma fields_non_intersection[rule_format]:
" (\<lambda>(cl, f, ty). f) ` set cl_f_ty_list \<inter> set (fields_in_path_f path'') = {}
\<and> f \<in> set (fields_in_path_f path'')
\<longrightarrow> ftype_in_fds_f P ctx_def (map (\<lambda>(cl, f, ty). fd_def cl f) cl_f_ty_list) f = ty_opt_bot_opt None"
by (induct cl_f_ty_list, auto)
lemma f_notin_list[rule_format]:
"f \<notin> set (map (\<lambda>(cl, f, ty). f) cl_f_ty_list) \<longrightarrow> ftype_in_fds_f P ctx_def (map (\<lambda>(cl, f, ty). fd_def cl f) cl_f_ty_list) f = ty_opt_bot_opt None"
by (induct cl_f_ty_list, auto)
declare find_path_rec_f.simps[simp del]
lemma ftype_preservation''''':
"\<lbrakk>find_path_rec_f P ctx cl (prefix @ [cld']) = Some (prefix @ cld' # path'');
find_type_f P ctx_def cl = Some ty; fields_f P ty = Some fs;
ftype_in_path_f P path'' f = Some ty_f;
(set (map (\<lambda>(cl, f, ty). f) cl_f_ty_list)) \<inter> (set fs) = {}\<rbrakk>
\<Longrightarrow> ftype_in_fds_f P ctx (map (\<lambda>(cl, f, ty). fd_def cl f) cl_f_ty_list) f = ty_opt_bot_opt None"
apply(clarsimp simp add: fields_f_def split: option.splits) apply(rename_tac path)
apply(case_tac cl) apply(simp add: find_path_rec_f.simps) apply(clarsimp split: fqn.splits option.splits)
apply(rename_tac dcl ctx'' cld'') apply(case_tac ctx'') apply(case_tac ctx) apply(clarsimp simp add: find_path_f_def)
apply(frule_tac prefix = "prefix @ [cld']" and suffix = path'' and prefix' = "[]" in fpr_same_suffix[rule_format]) apply(simp)
apply(clarsimp) apply(drule ftype_fields) apply(cut_tac fields_non_intersection) apply(force) apply(force)
done
declare find_path_rec_f.simps[simp]
lemma ftype_preservation'''':
"\<forall>suffix. find_path_rec_f P ctx cl prefix = Some (prefix @ suffix) \<longrightarrow> wf_program P \<longrightarrow>
(\<forall>(ctx, cld) \<in> set suffix.
(\<forall>prefix' suffix'. find_path_rec_f P ctx (cl_fqn (fqn_def (class_name_f cld))) prefix' = Some (prefix' @ suffix') \<and>
ftype_in_path_f P suffix' f = Some ty \<longrightarrow> ftype_in_path_f P suffix f = Some ty))"
+supply [[simproc del: defined_all]]
apply(induct_tac P ctx cl prefix rule: find_path_rec_f.induct)
apply(simp)
apply(clarsimp split: option.splits)
apply(subgoal_tac "\<forall>aa ba aaa bb.
(a = aaa \<and> b = bb \<and> aa = aaa \<and> ba = bb)
\<longrightarrow> (\<forall>suffix. find_path_rec_f P aaa (superclass_name_f bb) (path @ [(aaa, bb)]) = Some (path @ (aaa, bb) # suffix) \<longrightarrow>
(\<forall>x\<in>set suffix.
(\<lambda>(ctx, cld). (\<exists>prefix' suffix'.
case_option None (case_prod (\<lambda>ctx' cld. find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)]))) (find_cld_f P ctx (fqn_def (class_name_f cld))) =
Some (prefix' @ suffix') \<and>
ftype_in_path_f P suffix' f = Some ty) \<longrightarrow>
ftype_in_path_f P suffix f = Some ty)
x))")
defer
apply(thin_tac _)+
apply(force)
apply(clarsimp)
apply(frule path_append) apply(clarify)
apply(erule_tac x = "tl suffix" in allE)
apply(clarsimp)
apply(erule disjE)
apply(clarify)
apply(case_tac fqn)
apply(clarsimp simp add: class_name_f_def split: cld.splits)
apply(frule find_cld_name_eq) apply(frule_tac ctx = a in find_cld_name_eq)
apply(clarsimp simp add: class_name_f_def split: cld.splits)
apply(rename_tac list1 cl list2 list3)
apply(frule_tac path = "prefix' @ suffix'" in path_append) apply(clarsimp)
apply(frule_tac prefix = "path @ [(ab, cld_def list1 cl list2 list3)]" and suffix = path'' and
prefix' = "prefix' @ [(ab, cld_def list1 cl list2 list3)]" and suffix' = path''a in fpr_same_suffix[rule_format])
apply(clarsimp) apply(clarsimp)
apply(drule_tac x = "(aa, ba)" in bspec, simp)
apply(clarsimp)
apply(subgoal_tac
"(\<exists>prefix' suffix'.
find_path_rec_f P ab (superclass_name_f bb) (prefix' @ [(ab, bb)]) = Some (prefix' @ suffix') \<and> ftype_in_path_f P suffix' f = Some ty)")
apply(clarsimp)
apply(subgoal_tac "ftype_in_fds_f P a (class_fields_f b) f = ty_opt_bot_opt None") apply(simp)
apply(frule_tac cld' = b in wf_cld_from_lu) apply(simp) apply(erule wf_classE) apply(erule wf_class_cE)
apply(clarsimp simp add: superclass_name_f_def class_fields_f_def)
apply(rule ftype_preservation''''') apply(simp+)
apply(force)
done
lemma ftype_preservation''':
"\<lbrakk>find_path_f P ctx' cl = Some suffix; wf_program P; (ctx, cld) \<in> set suffix\<rbrakk>
\<Longrightarrow> \<forall>prefix' suffix'.
find_path_rec_f P ctx (cl_fqn (fqn_def (class_name_f cld))) prefix' = Some (prefix' @ suffix') \<and> ftype_in_path_f P suffix' f = Some ty \<longrightarrow>
ftype_in_path_f P suffix f = Some ty"
apply(cut_tac x = "(ctx, cld)" in ftype_preservation''''[of _ _ _ "[]", rule_format]) apply(simp add: find_path_f_def)+ done
lemma ftype_preservation'':
"\<lbrakk>find_path_f P ctx' cl = Some path; wf_program P; (ctx, cld) \<in> set path;
find_path_f P ctx (cl_fqn (fqn_def (class_name_f cld))) = Some path'; ftype_in_path_f P path' f = Some ty\<rbrakk>
\<Longrightarrow> ftype_in_path_f P path f = Some ty"
apply(cut_tac ftype_preservation'''[rule_format, of _ _ _ _ _ _ "[]"]) apply(assumption+) apply(unfold find_path_f_def) apply(rule) apply(simp+) done
lemma ftype_preservation'[simplified]:
"\<lbrakk>wf_program P; sty_one P ty_x_d ty_x_s; ftype P ty_x_s f ty_f\<rbrakk>
\<Longrightarrow> ftype P ty_x_d f ty_f"
apply(case_tac ty_x_s) apply(simp add: ftype_f_def)
apply(simp add: is_sty_one_def split: option.splits)
apply(case_tac ty_x_d) apply(simp) apply(clarsimp) apply(rename_tac ctx_s dcl_s path_d ctx_d dcl_d)
apply(simp add: ftype_f_def split: option.splits)
apply(rename_tac path_s cl_s fds_s mds_s)
apply(rule ftype_preservation'') apply(simp add: class_name_f_def)+
done
lemma ftype_preservation[simplified]:
"\<lbrakk>wf_program P;
\<forall>x\<in>dom \<Gamma>. wf_object P H (L x) (\<Gamma> x); \<Gamma> x = Some ty_x_s; ftype P ty_x_s f ty_f_s;
L x = Some (v_oid oid_x); H oid_x = Some (ty_x_d, fs_oid); ftype P ty_x_d f ty_f_d\<rbrakk>
\<Longrightarrow> ty_f_s = ty_f_d"
apply(drule_tac x = x in bspec, simp add: domI)
apply(erule wf_objectE)
apply(simp)
apply(clarsimp)
apply(erule sty_option.cases)
apply(clarsimp)
apply(simp add: ftype_preservation')
done
lemma map_f:
"\<forall>cl_f_ty_list. map (\<lambda>(cl, f). fd_def cl f) cl_f_list = map (\<lambda>(cl, f, ty). fd_def cl f) cl_f_ty_list \<longrightarrow>
map (\<lambda>(cl, f). f) cl_f_list = map (\<lambda>(cl, f, ty). f) cl_f_ty_list"
by (induct cl_f_list, auto)
lemma fields_to_owner[rule_format]:
"f \<in> set (fields_in_path_f path) \<longrightarrow> (\<exists>cld \<in> set (map snd path). \<exists>cl_f. fd_def cl_f f \<in> set (class_fields_f cld))"
apply(induct path) apply(simp) apply(clarsimp) apply(rename_tac cld path fd)
apply(clarsimp simp add: class_fields_f_def split: option.splits cld.splits fd.splits)
apply(force)
done
lemma ftype_in_fds_none[rule_format]:
"f \<notin> (\<lambda>(cl, f, ty). f) ` set cl_f_ty_list \<longrightarrow> ftype_in_fds_f P ctx (map (\<lambda>(cl, f, ty). fd_def cl f) cl_f_ty_list) f = ty_opt_bot_opt None"
by (induct cl_f_ty_list, auto)
lemma ftype_in_fds_some[rule_format]:
"(cl_f, f, ty_f) \<in> set cl_f_ty_list \<and> distinct (map (\<lambda>(cl, f, ty). f) cl_f_ty_list) \<and> find_type_f P ctx cl_f = Some ty_f
\<longrightarrow> ftype_in_fds_f P ctx (map (\<lambda>(cl, f, ty). fd_def cl f) cl_f_ty_list) f = ty_opt_bot_opt (Some ty_f)"
apply(induct cl_f_ty_list) apply(simp) apply(clarsimp) apply(safe) apply(force) apply(simp)
apply(frule ftype_in_fds_none) apply(force)
done
lemma no_ty_bot[THEN mp]:
"(\<forall>x\<in>set cl_f_ty_list. (\<lambda>(cl, f, ty). find_type_f P ctx_def cl = Some ty) x)
\<longrightarrow> ftype_in_fds_f P ctx' (map (\<lambda>(cl, f, ty). fd_def cl f) cl_f_ty_list) f \<noteq> ty_opt_bot_bot"
apply(induct cl_f_ty_list) apply(simp) apply(clarsimp split: if_split_asm option.splits)
apply(case_tac ctx', simp) (* should get rid of ctx dependency here *)
done
lemma field_has_type'[rule_format]:
"(cl_f, f, ty_f) \<in> set cl_f_ty_list \<and>
(\<forall>(ctx, cld) \<in> set path. wf_class P cld) \<and>
(ctx, cld_def dcl cl (map (\<lambda>(cl, f, ty). fd_def cl f) cl_f_ty_list) mds) \<in> set path \<and>
distinct (map (\<lambda>(cl, f, ty). f) cl_f_ty_list) \<and>
find_type_f P ctx cl_f = Some ty_f
\<longrightarrow> (\<exists>ty'. ftype_in_path_f P path f = Some ty')"
apply(induct path) apply(simp)
apply(simp) apply(rule)
apply(elim conjE) apply(erule disjE) apply(clarsimp)
apply(subgoal_tac
"ftype_in_fds_f P ctx (class_fields_f (cld_def dcl cl (map (\<lambda>(cl, f, ty). fd_def cl f) cl_f_ty_list) mds)) f = ty_opt_bot_opt (Some ty_f)") apply(simp)
apply(simp add: class_fields_f_def) apply(rule ftype_in_fds_some) apply(force)
apply(case_tac a) apply(rename_tac ctx' cld') apply(simp) apply(case_tac "ftype_in_fds_f P ctx' (class_fields_f cld') f")
apply(rename_tac ty_opt) apply(case_tac ty_opt)
apply(clarsimp)
apply(clarsimp)
apply(clarsimp)
apply(erule wf_classE) apply(erule wf_class_cE) apply(clarsimp simp add: class_fields_f_def)
apply(frule no_ty_bot) apply(force)
done
lemma mem_class_wf:
"\<lbrakk>wf_program P; cld \<in> set P\<rbrakk> \<Longrightarrow> wf_class P cld"
apply(erule wf_programE) apply(drule_tac x = cld in bspec, simp) apply(clarsimp) done
lemma field_has_type:
"\<lbrakk>wf_program P; fields P ty_oid (Some f_list); f \<in> set f_list\<rbrakk> \<Longrightarrow> \<exists>ty. ftype P ty_oid f ty"
apply(simp) apply(clarsimp simp add: fields_f_def split: option.splits) apply(rename_tac path)
apply(case_tac ty_oid) apply(simp) apply(clarsimp) apply(rename_tac dcl)
apply(frule fields_to_owner) apply(clarsimp) apply(rename_tac ctx' cld cl_f)
apply(frule clds_in_path_exist) apply(drule_tac x = "(ctx', cld)" in bspec, simp) apply(clarsimp)
apply(frule mem_class_wf, simp)
apply(erule wf_classE) apply(erule wf_class_cE) apply(clarsimp simp add: class_fields_f_def)
apply(rename_tac cl_f ty_f) apply(drule_tac x = "(cl_f, f, ty_f)" in bspec, simp) apply(clarsimp)
apply(simp add: ftype_f_def)
apply(rule_tac cl_f = cl_f and ty_f = ty_f and cl_f_ty_list = cl_f_ty_list and ctx = ctx' and
dcl = dclb and cl = cla and mds = "(map fst meth_def_meth_list)" in field_has_type')
apply(simp) apply(case_tac ctx') apply(simp)
apply(clarify) apply(drule_tac cl = "cl_fqn (fqn_def dcl)" and ctxcld = "(a, b)" in all_in_path_found) apply(simp)
apply(clarsimp) apply(rule wf_cld_from_lu) apply(simp) apply(force)
done
lemma exists_three_in_five:
"var_k \<in> set (map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list) \<longrightarrow>
(\<exists>y cl var' v. (y, cl, var_k, var', v) \<in> set y_cl_var_var'_v_list)"
by (induct y_cl_var_var'_v_list, auto)
lemma map_xx[rule_format]:
"(y, cl, var_k, var', v) \<in> set y_cl_var_var'_v_list \<longrightarrow>
(x_var var_k, x_var var') \<in> set (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list)"
by (induct y_cl_var_var'_v_list, auto)
lemma map_of_vd[rule_format]:
"\<forall>cl_var_ty_list.
map (\<lambda>(y, cl, var, var', v). vd_def cl var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). vd_def cl var) cl_var_ty_list
\<longrightarrow> map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list"
by (induct y_cl_var_var'_v_list, auto)
lemma mem_vd:
"vd_def cl_k var_k \<in> set (map (\<lambda>(y, cl, var, var', v). vd_def cl var) y_cl_var_var'_v_list) \<longrightarrow>
var_k \<in> set (map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list)"
by (induct y_cl_var_var'_v_list, auto)
lemma exists_var':
"\<lbrakk>distinct (map (\<lambda>(cl, var, ty). var) cl_var_ty_list);
map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list;
map_of (map (\<lambda>(cl, var, ty). (x_var var, ty)) cl_var_ty_list) var_k = Some ty_k\<rbrakk>
\<Longrightarrow> \<exists>var'_k. map_of (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) var_k = Some (x_var var'_k)"
apply(drule map_of_SomeD) apply(clarsimp, hypsubst_thin) apply(rename_tac cl_k var_k)
apply(subgoal_tac "var_k \<in> set (map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list)")
apply(simp (no_asm_use)) apply(clarify) apply(rename_tac y_k cl_k' var_k var'_k v_k)
apply(rule_tac x = var'_k in exI)
apply(rule map_of_is_SomeI)
apply(drule distinct_var, assumption+)
apply(rule map_xx, assumption)
by force
lemma map_y:
"\<lbrakk>((y_k, cl_k, var_k, var'_k, v_k), y'_k, ty_k) \<in> set (zip y_cl_var_var'_v_list y_ty_list);
map (\<lambda>(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list\<rbrakk>
\<Longrightarrow> (y_k, ty_k) \<in> set y_ty_list"
apply(clarsimp simp add: set_zip)
apply(drule_tac f = "\<lambda>(y, cl, var, var', v). y" in arg_cong)
apply(drule_tac f1 = "\<lambda>(y, cl, var, var', v). y" in nth_map[THEN sym])
apply(frule_tac f = "\<lambda>(y, ty). y" in arg_cong)
apply(case_tac "y_ty_list ! i")
apply(force simp add: in_set_conv_nth)
done
lemma set_zip_var'_v:
"\<lbrakk>((y_k, cl_k, var_k, var'_k, v_k), y'_k, ty_k) \<in> set (zip y_cl_var_var'_v_list y_ty_list)\<rbrakk>
\<Longrightarrow> (x_var var'_k, v_k) \<in> set (map (\<lambda>(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list)"
apply(subgoal_tac "(y_k, cl_k, var_k, var'_k, v_k) \<in> set y_cl_var_var'_v_list")
by (force simp add: set_zip)+
lemma map_of_map_zip_none:
"\<lbrakk>map (\<lambda>(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list;
map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)) x_G = None\<rbrakk>
\<Longrightarrow> map_of (map (\<lambda>(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list) x_G = None"
apply(drule map_length_y)
apply(clarsimp simp add: map_of_eq_None_iff)
apply(erule contrapos_np)
apply(simp only: set_map[THEN sym] fst_zip_eq)
by force
lemma map_of_map_zip_some:
"\<lbrakk>distinct (map (\<lambda>(y, cl, var, var', v). var') y_cl_var_var'_v_list);
map (\<lambda>(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list;
map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)) x_G = Some ty_k\<rbrakk>
\<Longrightarrow> \<exists>v_k. map_of (map (\<lambda>(y, cl, var, var', y). (x_var var', y)) y_cl_var_var'_v_list) x_G = Some v_k"
apply(drule map_of_SomeD)
apply(clarsimp) apply(rename_tac y cl var var' v y')
apply(subgoal_tac "(y, cl, var, var', v) \<in> set y_cl_var_var'_v_list")
apply(rule_tac x = v in exI)
apply(rule map_of_is_SomeI)
apply(simp only: map_fst_two' distinct_x_var')
apply(force)
by (force simp add: set_zip)
lemma x'_not_in_list:
"\<lbrakk>(y_k, cl_k, var_k, var'_k, v_k) \<in> set y_cl_var_var'_v_list;
x' \<notin> (\<lambda>(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list\<rbrakk>
\<Longrightarrow> x_var var'_k \<noteq> x'"
by (erule contrapos_nn, force)
lemma nth_in_set:
"\<lbrakk>i < length y_ty_list; map fst y_ty_list ! i = y_k\<rbrakk> \<Longrightarrow> y_k \<in> set (map fst y_ty_list)"
apply(subgoal_tac "\<exists>i<length (map fst y_ty_list). map fst y_ty_list ! i = y_k")
apply(simp only: in_set_conv_nth[THEN sym])
apply(rule_tac x = i in exI) apply(clarsimp)
done
lemma map_one_in_two:
"y_k \<in> set (map fst y_ty_list) \<longrightarrow> (\<exists>ty. (y_k, ty) \<in> set y_ty_list)"
by (induct y_ty_list, auto)
lemma exists_i: (* can probably simplify a lot *)
"\<lbrakk>distinct (map (\<lambda>(cl, var, ty). var) cl_var_ty_list);
(y_k, cl_ka, var_k, var'_k, v_k) \<in> set y_cl_var_var'_v_list;
(cl_k, var_k, ty_k) \<in> set cl_var_ty_list;
map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list\<rbrakk> \<Longrightarrow>
\<exists>i < length cl_var_ty_list. cl_var_ty_list ! i = (cl_k, var_k, ty_k) \<and> y_cl_var_var'_v_list ! i = (y_k, cl_ka, var_k, var'_k, v_k)"
apply(simp only: in_set_conv_nth)
apply(erule_tac P = "\<lambda>i. i < length cl_var_ty_list \<and> cl_var_ty_list ! i = (cl_k, var_k, ty_k)" in exE) apply(erule conjE)
apply(rule_tac x = x in exI)
apply(simp)
apply(drule_tac f = "\<lambda>(cl, var, ty). var" in arg_cong)
apply(frule_tac f1 = "\<lambda>(cl, var, ty). var" in nth_map[THEN sym])
apply(drule_tac t = "map (\<lambda>(cl, var, ty). var) cl_var_ty_list ! x" in sym)
apply(drule_tac s = "(\<lambda>(cl, var, ty). var) (cl_var_ty_list ! x)" in trans)
apply(simp)
apply(thin_tac "(\<lambda>(cl, var, ty). var) (cl_var_ty_list ! x) = (\<lambda>(cl, var, ty). var) (cl_k, var_k, ty_k)")
apply(drule sym)
apply(simp)
apply(erule exE) apply(rename_tac j) apply(erule conjE)
apply(simp only: distinct_conv_nth)
apply(erule_tac x = x in allE)
apply(drule_tac s = "map (\<lambda>(cl, var, ty). var) cl_var_ty_list" in sym)
apply(drule map_length_var)
apply(simp)
apply(erule_tac x = j in allE)
apply(simp)
done
lemma y_ty_match:
"\<lbrakk>i < length cl_var_ty_list; cl_var_ty_list ! i = (cl_k, var_k, ty_k);
y_cl_var_var'_v_list ! i = (y_k, cl_ka, var_k, var'_k, v_k);
map (\<lambda>(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list;
map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list\<rbrakk> \<Longrightarrow>
y_ty_list ! i = (y_k, ty_k)"
apply(drule_tac f = "\<lambda>(cl, var, ty). ty" in arg_cong)
apply(frule_tac f1 = "\<lambda>(cl, var, ty). ty" in nth_map[THEN sym])
apply(drule_tac t = "map (\<lambda>(cl, var, ty). ty) cl_var_ty_list ! i" in sym)
apply(frule map_length_y) apply(frule map_length_ty)
apply(drule_tac t = "length y_ty_list" in sym) apply(simp)
apply(drule_tac f = "\<lambda>(y, cl, var, var', v). y" in arg_cong)
apply(frule_tac f1 = "\<lambda>(y, cl, var, var', v). y" in nth_map[THEN sym])
apply(drule_tac t = "map (\<lambda>(y, cl, var, var', v). y) y_cl_var_var'_v_list ! i" in sym)
apply(case_tac "y_ty_list ! i") apply(simp)
done
lemma map_of_ty_k:
"\<lbrakk>distinct (map (\<lambda>(cl, var, ty). var) cl_var_ty_list);
distinct (map (\<lambda>(y, cl, var, var', v). var') y_cl_var_var'_v_list);
(y_k, cl_ka, var_k, var'_k, v_k) \<in> set y_cl_var_var'_v_list;
(cl_k, var_k, ty_k) \<in> set cl_var_ty_list;
map (\<lambda>(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list;
map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list;
map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list\<rbrakk>
\<Longrightarrow> map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)) (x_var var'_k) = Some ty_k"
apply(frule map_length_y)
apply(rule map_of_is_SomeI)
apply(simp only: fst_zip_eq)
apply(simp only: distinct_x_var')
apply(clarsimp)
apply(frule map_length_var)
apply(frule exists_i, simp+) apply(erule exE)
apply(drule_tac t = "length y_ty_list" in sym) apply(clarsimp)
apply(frule y_ty_match, assumption+)
apply(subgoal_tac "zip y_cl_var_var'_v_list y_ty_list ! i = (y_cl_var_var'_v_list ! i, y_ty_list ! i)")
apply(drule_tac f = "\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)" in arg_cong)
apply(subgoal_tac "i < length (zip y_cl_var_var'_v_list y_ty_list)")
apply(frule_tac f1 = "\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)" in nth_map[THEN sym])
apply(drule_tac t = "map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list) ! i" in sym)
apply(drule_tac s = "(\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list ! i)" in trans)
apply(simp)
apply(subgoal_tac "\<exists>j<length (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)). map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list) ! j = (x_var var'_k, ty_k)")
apply(simp only: in_set_conv_nth[THEN sym])
apply(simp)
apply(rule_tac x = i in exI) apply(simp)
apply(force)
apply(rule nth_zip) apply(simp) apply(simp)
done
lemma map_of_ty_k':
"\<lbrakk>distinct (map (\<lambda>(cl, var, ty). var) cl_var_ty_list); distinct (map (\<lambda>(y, cl, var, var', v). var') y_cl_var_var'_v_list);
x' \<notin> (\<lambda>(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list; map fst y_cl_var_var'_v_list = map fst y_ty_list;
map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list; map (\<lambda>(y, cl, var, u). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list;
map_of (map (\<lambda>(cl, var, y). (x_var var, y)) cl_var_ty_list) (x_var var) = Some ty_var\<rbrakk>
\<Longrightarrow> (if x_var (case_option var (case_x (\<lambda>var'. var') var) (map_of (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) (x_var var))) = x' then Some ty_x_m
else (\<Gamma> ++ map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)))
(x_var (case_option var (case_x (\<lambda>var'. var') var)
(map_of (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) (x_var var))))) =
Some ty_var"
apply(frule exists_var') apply(simp add: weaken_list_var) apply(simp) apply(clarsimp split del: if_split) apply(clarsimp)
apply(drule map_of_SomeD)+ apply(clarsimp split del: if_split) apply(frule x'_not_in_list, simp+)
apply(frule map_of_ty_k) apply(simp add: weaken_list_y weaken_list_var)+
done
lemma var_not_var'_x':
"\<lbrakk>\<Gamma> (x_var var) = Some ty_var; L (x_var var) = Some v_var;
\<forall>x\<in>set y_cl_var_var'_v_list. L (x_var ((\<lambda>(y, cl, var, var', v). var') x)) = None; L x' = None\<rbrakk>
\<Longrightarrow> (\<Gamma> ++ map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)) ++ Map.empty(x' \<mapsto> ty_x_d)) (x_var var) = Some ty_var"
apply(subgoal_tac "x_var var \<noteq> x'")
apply(simp add: map_add_def)
apply(subgoal_tac "map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)) (x_var var) = None")
apply(simp)
apply(clarsimp simp add: map_of_eq_None_iff set_zip)
apply(rename_tac y_k cl_k var_k v_k y'_k ty_k i)
apply(drule_tac x = "(y_k, cl_k, var_k, var, v_k)" in bspec, simp)
apply(force)
by force
lemma same_el:
"\<lbrakk>distinct (map (\<lambda>(y, cl, var, var', v). var') y_cl_var_var'_v_list);
i < length y_cl_var_var'_v_list;
(y_i, cl_i, var_i, var'_i, v_i) = y_cl_var_var'_v_list ! i;
(y_j, cl_j, var_j, var'_i, v_j) \<in> set y_cl_var_var'_v_list\<rbrakk>
\<Longrightarrow> (y_i, cl_i, var_i, var'_i, v_i) = (y_j, cl_j, var_j, var'_i, v_j)"
apply(simp only: in_set_conv_nth) apply(clarsimp) apply(rename_tac j)
apply(simp only: distinct_conv_nth)
apply(erule_tac x = i in allE) apply(simp)
apply(erule_tac x = j in allE) apply(simp)
apply(drule sym) apply(simp)
done
lemma same_y:
"\<lbrakk>map (\<lambda>(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list;
i < length y_cl_var_var'_v_list;
(y'_k, ty_k) = y_ty_list ! i; y_cl_var_var'_v_list ! i = (y_k, cl_k, var_k, var'_k, v_k)\<rbrakk>
\<Longrightarrow> y'_k = y_k"
apply(drule_tac f = "\<lambda>(y, cl, var, var', v). y" in arg_cong)
apply(frule_tac f1 = "\<lambda>(y, cl, var, var', v). y" in nth_map[THEN sym])
apply(drule_tac f = fst in arg_cong) apply(simp)
apply(frule map_length_y) apply(simp)
done
lemma map_map_zip_y':
"map fst y_y'_list = map fst y_ty_list \<Longrightarrow>
map snd y_y'_list = map fst (map (\<lambda>((y, y'), yy, ty). (y', ty)) (zip y_y'_list y_ty_list))"
apply(rule nth_equalityI)
apply(induct y_y'_list) apply(simp) apply(clarsimp) apply(frule map_length_y') apply(simp)
apply(clarsimp)
apply(case_tac "y_y'_list ! i")
apply(case_tac "y_ty_list ! i")
apply(clarsimp)
apply(subgoal_tac "i < length (map (\<lambda>((y, y'), yy, ty). (y', ty)) (zip y_y'_list y_ty_list))")
apply(frule_tac f = fst and xs = "map (\<lambda>((y, y'), yy, ty). (y', ty)) (zip y_y'_list y_ty_list)" in nth_map)
apply(simp)
apply(subgoal_tac "length (map (\<lambda>((y, y'), yy, ty). (y', ty)) list) = length list" for list)
apply(simp) apply(frule map_length_y') apply(simp)
apply(simp only: map_length_list)
done
lemma map_map_zip_ty:
"map fst y_y'_list = map fst y_ty_list \<Longrightarrow>
map snd (map (\<lambda>((y, y'), yy, ty). (y', ty)) (zip y_y'_list y_ty_list)) = map snd y_ty_list"
apply(frule map_length_y')
apply(rule nth_equalityI)
apply(induct y_ty_list) apply(simp) apply(simp)
apply(clarsimp)
apply(case_tac "y_y'_list ! i")
apply(case_tac "y_ty_list ! i")
apply(simp)
done
lemma map_y_yy:
"\<lbrakk>i < length y_y'_list; i < length y_ty_list;
map fst y_y'_list = map fst y_ty_list;
y_y'_list ! i = (y, y'); y_ty_list ! i = (yy, ty)\<rbrakk> \<Longrightarrow>
y = yy"
apply(drule_tac f = fst in arg_cong)
apply(frule_tac f1 = fst in nth_map[THEN sym])
by simp
lemma var_k_of_ty_k:
"\<lbrakk>distinct (map (\<lambda>(cl, var, ty). var) cl_var_ty_list); distinct (map (\<lambda>(y, cl, var, var', v). var') y_cl_var_var'_v_list);
x' \<notin> (\<lambda>(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list;
map fst y_cl_var_var'_v_list = map fst y_ty_list;
map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list;
map (\<lambda>(y, cl, var, u). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list;
(if x = x_this then Some ty_x_m else map_of (map (\<lambda>(cl, var, y). (x_var var, y)) cl_var_ty_list) x) = Some ty_x\<rbrakk>
\<Longrightarrow> (if (case if x = x_this then Some x' else map_of (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) x of None \<Rightarrow> x | Some x' \<Rightarrow> x') = x'
then Some ty_x_m
else (\<Gamma> ++ map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)))
(case if x = x_this then Some x' else map_of (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) x of None \<Rightarrow> x | Some x' \<Rightarrow> x')) =
Some ty_x"
apply(clarsimp) apply(rule) apply(rule) apply(simp) apply(clarsimp)
apply(frule exists_var') apply(simp add: weaken_list_var) apply(simp) apply(clarsimp split del: if_split)
apply(drule map_of_SomeD)+ apply(clarsimp split del: if_split) apply(frule x'_not_in_list, simp+)
apply(frule map_of_ty_k) apply(simp add: weaken_list_y weaken_list_var)+
done
lemma x_var_not_this: "(if x_var var = x_this then Some x' else Q) = Q" by simp
lemma subtype_through_tr:
"\<lbrakk>distinct (map (\<lambda>(cl, var, ty). var) cl_var_ty_list); distinct (map (\<lambda>(y, cl, var, var', v). var') y_cl_var_var'_v_list);
x' \<notin> (\<lambda>(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list;
map fst y_cl_var_var'_v_list = map fst y_ty_list;
map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list;
map (\<lambda>(y, cl, var, u). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list;
is_sty_one P ty_x ty_var = Some True; (if x = x_this then Some ty_x_m else map_of (map (\<lambda>(cl, var, y). (x_var var, y)) cl_var_ty_list) x) = Some ty_x;
map_of (map (\<lambda>(cl, var, y). (x_var var, y)) cl_var_ty_list) (x_var var) = Some ty_var\<rbrakk>
\<Longrightarrow> sty_option P (if (case if x = x_this then Some x' else map_of (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) x of None \<Rightarrow> x | Some x' \<Rightarrow> x') =
x'
then Some ty_x_m
else (\<Gamma> ++ map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)))
(case if x = x_this then Some x' else map_of (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) x of None \<Rightarrow> x | Some x' \<Rightarrow> x'))
(if x_var (case_option var (case_x (\<lambda>var'. var') var) (map_of (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) (x_var var))) = x' then Some ty_x_m
else (\<Gamma> ++ map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)))
(x_var (case_option var (case_x (\<lambda>var'. var') var)
(if x_var var = x_this then Some x' else map_of (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) (x_var var)))))"
apply(rule_tac ty = ty_x and ty' = ty_var in sty_optionI)
apply(simp add: var_k_of_ty_k)
apply(simp only: x_var_not_this) apply(rule map_of_ty_k') apply(simp+)
done
lemma subtypes_through_tr:
"\<lbrakk>distinct (map (\<lambda>(cl, var, ty). var) cl_var_ty_list); distinct (map (\<lambda>(y, cl, var, var', v). var') y_cl_var_var'_v_list);
x' \<notin> (\<lambda>(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list; map fst y_cl_var_var'_v_list = map fst y_ty_list;
map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list; map (\<lambda>(y, cl, var, u). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list;
\<forall>x\<in>set y_y'_list. case_prod (\<lambda>y. (=) (case if y = x_this then Some x' else map_of (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list) y of
None \<Rightarrow> y | Some x' \<Rightarrow> x')) x;
\<forall>x\<in>set y_ty_lista. (\<lambda>(y, ty). sty_option P (if y = x_this then Some ty_x_m else map_of (map (\<lambda>(cl, var, y). (x_var var, y)) cl_var_ty_list) y) (Some ty)) x;
map fst y_y'_list = map fst y_ty_lista; (y, y') = y_y'_list ! i; (yy, ty) = y_ty_lista ! i; i < length y_y'_list; i < length y_ty_lista\<rbrakk> \<Longrightarrow>
sty_option P (if y' = x' then Some ty_x_m else (\<Gamma> ++ map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list))) y') (Some ty)"
apply(drule_tac x = "(y, y')" in bspec, simp) apply(drule_tac x = "(yy, ty)" in bspec, simp) apply(clarsimp split del: if_split)
apply(frule_tac y_ty_list = y_ty_lista in map_y_yy, assumption+) apply(erule sty_option.cases) apply(clarsimp split del: if_split)
apply(rule_tac ty = tya in sty_optionI, (simp add: var_k_of_ty_k)+)
done
lemma map_override_get:
"(\<Gamma> ++ (\<lambda>u. if u = x' then Some ty_x_d else None)) x' = Some ty_x_d"
apply(simp add: map_add_def)
done
lemma s_induct':
"\<lbrakk>\<And>ss. \<forall>s\<in>set ss. P s \<Longrightarrow> P (s_block ss); \<And>list x. P (s_ass list x); \<And>list1 x list2. P (s_read list1 x list2); \<And>x1 list x2. P (s_write x1 list x2);
\<And>x1 x2 s1 s2. \<lbrakk>P s1; P s2\<rbrakk> \<Longrightarrow> P (s_if x1 x2 s1 s2); \<And>list1 x list2 list3. P (s_call list1 x list2 list3); \<And>list ctx cl. P (s_new list ctx cl)\<rbrakk>
\<Longrightarrow> P s"
by (metis s.induct)
lemma wf_tr_stmt_in_G':
"\<forall>s'. distinct (map (\<lambda>(cl, var, ty). var) cl_var_ty_list)
\<and> distinct (map (\<lambda>(y, cl, var, var', v). var') y_cl_var_var'_v_list)
\<and> x' \<notin> (\<lambda>(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list
\<and> map (\<lambda>(y, cl, var, var', v). y) y_cl_var_var'_v_list = map fst y_ty_list
\<and> map (\<lambda>(cl, var, ty). ty) cl_var_ty_list = map snd y_ty_list
\<and> map (\<lambda>(y, cl, var, var', v). var) y_cl_var_var'_v_list = map (\<lambda>(cl, var, ty). var) cl_var_ty_list
\<and> is_sty_one P ty_x_d ty_x_m = Some True
\<and> wf_stmt P (map_of (map (\<lambda>(cl, var, ty). (x_var var, ty)) cl_var_ty_list)(x_this \<mapsto> ty_x_m)) s'
\<and> tr_s (map_of (map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var')) y_cl_var_var'_v_list)(x_this \<mapsto> x')) s' s''
\<longrightarrow> wf_stmt P ((\<Gamma> ++ map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)))(x' \<mapsto> ty_x_m)) s''"
apply(rule s_induct')
apply(clarsimp)
apply(erule tr_s.cases, simp_all) apply(clarsimp)
apply(erule wf_stmtE, simp_all) apply(clarsimp)
apply(rule wf_blockI) apply(clarsimp) apply(rename_tac s'_s''_list s' s'')
apply(drule_tac x = "(s', s'')" in bspec, simp)+ apply(clarsimp)
apply(clarsimp) apply(rename_tac var x s')
apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all)
apply(erule sty_option.cases) apply(rule wf_var_assignI) apply(clarsimp split del: if_split)
apply(rule_tac ty_x = ty in subtype_through_tr, simp+)
apply(clarsimp) apply(rename_tac var x f s')
apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all)
apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rename_tac f ty_x var x ty_f ty_var)
apply(rule wf_field_readI) apply(simp add: var_k_of_ty_k) apply(simp)
apply(rule_tac ty' = ty_var in sty_optionI) apply(simp) apply(rule map_of_ty_k', simp+)
apply(clarsimp) apply(rename_tac x f y s')
apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all)
apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rename_tac f ty_x x y ty_y ty_f)
apply(rule wf_field_writeI) apply(simp add: var_k_of_ty_k) apply(simp)
apply(rule sty_optionI, (simp add: var_k_of_ty_k)+)
apply(clarsimp) apply(rename_tac x y s1 s2 s')
apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all) apply(rule wf_ifI)
apply(erule disjE)
apply(rule disjI1) apply(erule sty_option.cases) apply(clarsimp split del: if_split)
apply(rule sty_optionI, (simp add: var_k_of_ty_k split del: if_split)+)
apply(rule disjI2) apply(erule sty_option.cases) apply(clarsimp split del: if_split)
apply(rule sty_optionI, (simp add: var_k_of_ty_k)+)
apply(clarsimp)
apply(clarsimp)
apply(clarsimp) apply(rename_tac var x meth ys s')
apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all) apply(clarsimp split del: if_split)
apply(rule_tac y_ty_list = "map (\<lambda>((y, y'), (yy, ty)). (y', ty)) (zip y_y'_list y_ty_lista)" in wf_mcallI[simplified])
apply(force simp add: map_map_zip_y'[simplified])
apply(simp add: var_k_of_ty_k)
apply(force simp add: map_map_zip_ty[simplified])
apply(clarsimp simp add: set_zip split del: if_split) apply(simp add: subtypes_through_tr)
apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rule sty_optionI, (simp add: map_of_ty_k')+)
apply(clarsimp) apply(rename_tac var ctx cl s')
apply(erule tr_s.cases, simp+) apply(erule wf_stmtE, simp_all) apply(clarsimp)
apply(rule wf_newI) apply(simp)
apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rule sty_optionI, (simp add: map_of_ty_k')+)
done
lemma wf_object_heap:
"\<lbrakk>wf_object P H v_opt (Some ty'); is_sty_one Pa ty_y_s ty_f = Some True; H oid_x = Some (ty_x_d, fs_oid)\<rbrakk>
\<Longrightarrow> wf_object P (H(oid_x \<mapsto> (ty_x_d, fs_oid(f \<mapsto> v)))) v_opt (Some ty')"
apply(erule wf_objectE) apply(clarsimp)
apply(rule wf_nullI, simp)
apply(erule sty_option.cases) apply(clarsimp split: option.splits)
apply(rule wf_objectI) apply(force intro: sty_optionI)
done
lemma not_dom_is_none:
"((\<lambda>(y, cl, var, var', v). x_var var') ` set y_cl_var_var'_v_list \<inter> Map.dom L = {})
\<Longrightarrow> (\<forall>x\<in>set y_cl_var_var'_v_list. (\<lambda>(y, cl, var, var', v). L (x_var var') = None) x)"
by force
lemma type_to_val:
"\<lbrakk>wf_varstate P \<Gamma> H L; sty_option P (\<Gamma> x) (\<Gamma> y)\<rbrakk> \<Longrightarrow> \<exists>v w. L x = Some v \<and> L y = Some w"
apply(erule sty_option.cases) apply(clarsimp) apply(erule wf_varstateE)
apply(frule_tac x = x in bspec, simp add: domI) apply(drule_tac x = y in bspec, simp add: domI)
apply(elim wf_objectE) by (simp+)
lemma find_path_fields[THEN mp]:
"find_path_ty_f P ty = Some path \<longrightarrow> (\<exists>fs. fields_in_path_f path = fs)"
by (force)
lemma replicate_eq_length:
"replicate x c = replicate y c \<Longrightarrow> x = y"
apply(subgoal_tac "length (replicate x c) = length (replicate y c)") apply(thin_tac _)
apply(subgoal_tac "length (replicate x c) = x")
apply(subgoal_tac "length (replicate y c) = y")
by simp_all
lemma string_infinite:
"infinite (UNIV :: string set)"
apply(simp add: infinite_iff_countable_subset)
apply(rule_tac x = "\<lambda>n. replicate n c" in exI) apply(rule linorder_injI)
apply(clarify) apply(frule replicate_eq_length) apply(simp)
done
lemma x_infinite:
"infinite (UNIV :: x set)"
apply(simp add: infinite_iff_countable_subset)
apply(rule_tac x = "\<lambda>n. if n = 0 then x_this else x_var (replicate n c)" in exI)
apply(rule linorder_injI) apply(clarsimp)
done
lemma fresh_oid:
"wf_heap P H \<Longrightarrow> (\<exists>oid. oid \<notin> dom H)"
apply(erule wf_heapE)
apply(cut_tac infinite_UNIV_nat) apply(frule_tac A = "dom H" in ex_new_if_finite)
apply(assumption) apply(simp)
done
lemma fresh_x:
"finite A \<Longrightarrow> \<exists>x::x. x \<notin> A"
apply(cut_tac x_infinite) apply(frule_tac A = A in ex_new_if_finite)
apply(assumption) apply(simp)
done
lemma fresh_var_not_in_list:
"finite (A::x set) \<Longrightarrow> \<exists>var::var. x_var var \<notin> A \<and> var \<notin> set list"
apply(cut_tac x_infinite) apply(subgoal_tac "finite (insert x_this A)")
apply(frule_tac F = "insert x_this A" and G = "x_var ` set list" in finite_UnI) apply(simp)
apply(frule_tac A = "insert x_this A \<union> x_var ` set list" in ex_new_if_finite)
apply(assumption) apply(erule exE) apply(clarsimp) apply(case_tac a) apply(force+)
done
lemma fresh_vars[rule_format]:
"finite (A::x set) \<longrightarrow> (\<exists>list::var list. length list = i \<and> x_var` set list \<inter> A = {} \<and> distinct list)"
apply(induct i) apply(simp)
apply(clarsimp) apply(frule fresh_var_not_in_list) apply(erule exE)
apply(rule_tac x = "var#list" in exI) apply(simp)
done
lemma fresh_x_not_in_vars':
"finite A \<Longrightarrow> \<exists>x'. x' \<notin> A \<and> x' \<notin> (x_var ` set vars')"
apply(subgoal_tac "finite (x_var ` set vars')") apply(frule_tac G = "x_var ` set vars'" in finite_UnI) apply(assumption)
apply(cut_tac x_infinite) apply(frule_tac A = "A \<union> x_var ` set vars'" in ex_new_if_finite)
apply(simp) apply(simp) apply(simp)
done
lemma lift_opts_ft_length[rule_format]:
"\<forall>tys. lift_opts (map (case_vd (\<lambda>clk vark. find_type_f P ctx_o clk)) vds) = Some tys \<longrightarrow> length tys = length vds"
by (induct vds, auto split: vd.splits option.splits)
lemma fpr_mem_has_path'[rule_format]:
"\<forall>suffix. find_path_rec_f P ctx cl prefix = Some (prefix @ suffix) \<longrightarrow>
(\<forall>(ctx', cld') \<in> set suffix. \<forall>prefix'.
(\<exists>suffix'. find_path_rec_f P ctx' (cl_fqn (fqn_def (class_name_f cld'))) prefix' = Some suffix'))"
+supply [[simproc del: defined_all]]
apply(induct_tac P ctx cl prefix rule: find_path_rec_f.induct)
apply(simp)
apply(clarsimp split: option.splits)
apply(subgoal_tac "\<forall>aa ba aaa bb.
(a = aaa \<and> b = bb \<and> aa = aaa \<and> ba = bb)
\<longrightarrow> (\<forall>suffix. find_path_rec_f P aaa (superclass_name_f bb) (path @ [(aaa, bb)]) = Some (path @ (aaa, bb) # suffix) \<longrightarrow>
(\<forall>x\<in>set suffix.
(\<lambda>(ctx', cld').
\<forall>prefix'. \<exists>suffix'. case_option None (case_prod (\<lambda>ctx' cld. find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)]))) (find_cld_f P ctx' (fqn_def (class_name_f cld'))) =
Some suffix')
x))")
defer
apply(thin_tac _)+
apply(force)
apply(thin_tac "\<And>aa ba x y.
\<lbrakk>a = aa \<and> b = ba; x = aa \<and> y = ba\<rbrakk>
\<Longrightarrow> \<forall>suffix.
find_path_rec_f P aa (superclass_name_f ba) (path @ [(aa, ba)]) =
Some (path @ (aa, ba) # suffix) \<longrightarrow>
(\<forall>x\<in>set suffix.
case x of
(ctx', cld') \<Rightarrow>
\<forall>prefix'.
\<exists>suffix'.
case_option None
(\<lambda>(ctx', cld).
find_path_rec_f P ctx' (superclass_name_f cld) (prefix' @ [(ctx', cld)]))
(find_cld_f P ctx' (fqn_def (class_name_f cld'))) =
Some suffix')")
apply(clarsimp)
apply(frule path_append) apply(clarify)
apply(erule_tac x = "tl suffix" in allE)
apply(clarsimp)
apply(erule disjE)
apply(clarify)
apply(clarsimp simp add: class_name_f_def split: cld.splits)
apply(rename_tac list1 cl list2 list3)
apply(case_tac fqn) apply(clarsimp) apply(frule find_cld_name_eq)
apply(clarsimp simp add: class_name_f_def split: cld.splits)
apply(frule_tac suffix = path'' and prefix' = " prefix' @ [(a, cld_def list1 cl list2 list3)]" in fpr_same_suffix') apply(simp) apply(simp)
apply(drule_tac x = "(aa, ba)" in bspec, simp) apply(clarsimp split: option.splits)
done
lemma fpr_mem_has_path:
"\<lbrakk>find_path_f P ctx cl = Some suffix; (ctx', cld') \<in> set suffix\<rbrakk> \<Longrightarrow>
\<exists>suffix'. find_path_rec_f P ctx' (cl_fqn (fqn_def (class_name_f cld'))) [] = Some suffix'"
apply(cut_tac x = "(ctx', cld')" in fpr_mem_has_path'[of _ _ _ "[]"]) apply(simp add: find_path_f_def)+ done
lemma mtype_to_find_md:
"\<lbrakk>wf_program P; mtype_f P ty_x_s m = Some (mty_def tys ty_r); is_sty_one P ty_x_d ty_x_s = Some True\<rbrakk> \<Longrightarrow>
\<exists>ctx cl_r vds ss' y. find_meth_def_f P ty_x_d m = Some (ctx, meth_def_def (meth_sig_def cl_r m vds) (meth_body_def ss' y))
\<and> length tys = length vds"
apply(clarsimp simp add: is_sty_one_def split: option.splits ty.splits)
apply(clarsimp simp add: mtype_f_def find_meth_def_f_def)
apply(case_tac ty_x_d) apply(simp) apply(clarsimp) apply(rename_tac path_d ctx_s dcl_s ctx_d dcl_d)
apply(rename_tac cl_s fds_s mds_s)
apply(frule fpr_mem_has_path) apply(simp) apply(clarsimp simp del: find_path_rec_f.simps)
apply(frule_tac path = path_d and m = m and mty = "mty_def tys ty_r" and path' = suffix' in mty_preservation')
apply(simp add: class_name_f_def find_path_f_def)+
apply(clarsimp simp add: mtype_f_def split: option.splits if_split_asm meth_def.splits meth_sig.splits)
apply(rename_tac meth_body a b c d e f g)
apply(case_tac meth_body) apply(simp) apply(clarsimp simp add: lift_opts_ft_length)
apply(clarsimp simp add: find_meth_def_f_def split: option.splits)
apply(frule find_md_m_match[rule_format]) apply(assumption)
done
lemma exist_lifted_values:
"\<lbrakk>\<forall>x\<in>dom \<Gamma>. wf_object P H (L x) (\<Gamma> x); \<forall>x\<in>set y_ty_list. (\<lambda>(y, ty). sty_option P (\<Gamma> y) (Some ty)) x\<rbrakk>
\<Longrightarrow> \<exists>vs. lift_opts (map (\<lambda>(y, ty). L y) y_ty_list) = Some vs"
apply(induct y_ty_list) apply(simp) apply(clarsimp)
apply(erule sty_option.cases) apply(clarsimp) apply(drule_tac x = a in bspec, simp add: domI)
apply(erule wf_objectE) apply(simp) apply(simp)
done
lemma [iff]:
"length (tr_ss_f (map_of (zip (map (case_vd (\<lambda>cl. x_var)) vds) (map x_var vars'))(x_this \<mapsto> x')) ss') = length ss'"
by (induct ss', auto)
lemma collect_suc_eq_lt[simp]:
"{f i |i::nat. i < Suc n} = {f i |i. i = 0} \<union> {f (Suc i) |i. i < n}"
apply(induct n) apply(simp)
apply(clarsimp)
apply(simp add: less_Suc_eq)
apply(simp add: image_Collect[THEN sym])
apply(simp add: Collect_disj_eq)
apply(blast)
done
lemma [iff]:
"\<forall>y_ty_list vars vars' vs. length y_ty_list = length vds \<and> length vars = length vds \<and> length vars' = length vds \<and> length vs = length vds \<longrightarrow>
map (\<lambda>(y, cl, var, var', v). vd_def cl var) (zip (map fst y_ty_list) (zip (map (case_vd (\<lambda>cl var. cl)) vds) (zip (map (case_vd (\<lambda>cl var. var)) vds) (zip vars' vs)))) = vds"
apply(induct vds) apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv split: vd.splits)
done
lemma vars'_eq[rule_format]:
"\<forall>y_ty_list vds vars vs. length y_ty_list = length vars' \<and> length vds = length vars' \<and> length vars = length vars' \<and> length vs = length vars' \<longrightarrow>
(\<lambda>(y, cl, var, var', v). x_var var') ` set (zip (map fst y_ty_list) (zip (map (case_vd (\<lambda>cl var. cl)) vds) (zip vars (zip vars' vs)))) = x_var ` set vars'"
apply(induct vars') apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv)
done
lemma [iff]:
"\<forall>y_ty_list vds vars vs. length y_ty_list = length vars' \<and> length vds = length vars' \<and> length vars = length vars' \<and> length vs = length vars' \<longrightarrow>
map (\<lambda>(y, cl, var, var', v). var') (zip (map fst y_ty_list) (zip (map (case_vd (\<lambda>cl var. cl)) vds) (zip vars (zip vars' vs)))) = vars'"
apply(induct vars') apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv)
done
lemma [iff]:
"\<forall>y_ty_list vds vars vs. length y_ty_list = length vars' \<and> length vds = length vars' \<and> length vars = length vars' \<and> length vs = length vars' \<longrightarrow>
map (\<lambda>(y, cl, var, var', v). (x_var var, x_var var'))
(zip (map fst y_ty_list) (zip (map (case_vd (\<lambda>cl var. cl)) vds) (zip (map (case_vd (\<lambda>cl var. var)) vds) (zip vars' vs)))) =
zip (map (case_vd (\<lambda>cl. x_var)) vds) (map x_var vars')"
apply(induct vars') apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv split: vd.splits)
done
lemma [iff]:
"\<forall>vds vars vars' vs. length vds = length y_ty_list \<and> length vars = length y_ty_list \<and> length vars' = length y_ty_list \<and> length vs = length y_ty_list \<longrightarrow>
map (\<lambda>(y, cl, var, var', v). y) (zip (map fst y_ty_list) (zip (map (case_vd (\<lambda>cl var. cl)) vds) (zip vars (zip vars' vs)))) = map fst y_ty_list"
apply(induct y_ty_list) apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv)
done
lemma lift_opts_mapping:
"\<forall>vds vars vars' vs. length vds = length y_ty_list \<and> length vars = length y_ty_list \<and> length vars' = length y_ty_list \<and> length vs = length y_ty_list \<longrightarrow>
lift_opts (map (\<lambda>(y, ty). L y) y_ty_list) = Some vs \<longrightarrow>
(\<forall>x\<in>set (zip (map fst y_ty_list) (zip (map (case_vd (\<lambda>cl var. cl)) vds) (zip (map (case_vd (\<lambda>cl var. var)) vds) (zip vars' vs)))). (\<lambda>(y, cl, var, var', v). L y = Some v) x)"
apply(induct y_ty_list) apply(simp) apply(clarsimp simp add: set_zip length_Suc_conv split: vd.splits option.splits)
apply(rename_tac y1 y_ty_list v2 vs cl1 var1 var'1 v1 vds vars var'2 vars' i cl2 var2)
apply(erule_tac x = vds in allE) apply(erule_tac x = vars in allE) apply(erule_tac x = vars' in allE)
apply(simp) apply(case_tac i) apply(simp) apply(rename_tac j) apply(clarsimp)
apply(erule_tac x = "map fst y_ty_list ! j" in allE) apply(clarsimp)
apply(case_tac "zip (map (case_vd (\<lambda>cl var. cl)) vds) (zip (map (case_vd (\<lambda>cl var. var)) vds) (zip vars' vs)) ! j") apply(rename_tac cl1 var1 var'1 v1)
apply(erule_tac x = cl1 in allE) apply(erule_tac x = var1 in allE) apply(erule_tac x = var'1 in allE) apply(erule_tac x = v1 in allE)
apply(clarsimp) apply(erule_tac x = j in allE) apply(simp)
done
lemma length_y_ty_list_vs[rule_format]:
"\<forall>vs. lift_opts (map (\<lambda>(y, ty). L y) y_ty_list) = Some vs \<longrightarrow> length y_ty_list = length vs"
by (induct y_ty_list, auto split: option.splits)
lemma translation_eq:
"\<forall>x\<in>set (zip (tr_ss_f (map_of (zip (map (case_vd (\<lambda>cl. x_var)) vds) (map x_var vars'))(x_this \<mapsto> x')) ss') ss').
(\<lambda>(s'', s'). tr_s (map_of (zip (map (case_vd (\<lambda>cl. x_var)) vds) (map x_var vars'))(x_this \<mapsto> x')) s' s'') x"
apply(induct ss') apply(simp add: tr_rel_f_eq)+ done
theorem progress:
"\<lbrakk>wf_config \<Gamma> (config_normal P L H S); S \<noteq> []\<rbrakk> \<Longrightarrow> \<exists>config'. r_stmt (config_normal P L H S) config'"
+supply [[simproc del: defined_all]]
apply(case_tac S)
apply(simp)
apply(clarsimp) apply(rename_tac s ss)
apply(case_tac s)
apply(erule_tac[1-7] wf_configE) apply(simp_all)
\<comment> \<open>block\<close>
apply(force intro: r_blockI[simplified])
\<comment> \<open>variable assignment\<close>
apply(clarsimp, erule wf_stmtE, simp_all, clarsimp)
apply(frule type_to_val, simp) apply(clarify) apply(frule r_var_assignI) apply(force)
\<comment> \<open>field read\<close>
apply(clarsimp, erule wf_stmtE, simp_all, clarsimp)
apply(erule wf_varstateE)
apply(drule_tac x = x in bspec, simp add: domI)
apply(erule wf_objectE) apply(clarsimp) apply(frule r_field_read_npeI) apply(force)
apply(clarsimp) apply(erule_tac ?a3.0 = "Some ty" in sty_option.cases) apply(clarsimp split: option.splits)
apply(erule wf_heapE) apply(drule_tac x = oid in bspec, simp add: domI) apply(clarsimp)
apply(rename_tac x oid ty_x_s ty_x_d fields_oid fs)
apply(frule no_field_hiding, simp+) apply(drule_tac x = f in bspec, simp) apply(clarsimp)
apply(erule wf_objectE) apply(clarsimp, frule_tac H = H in r_field_readI, simp, force)+
\<comment> \<open>field write\<close>
apply(clarsimp, erule wf_stmtE, simp_all, clarsimp)
apply(erule wf_varstateE) apply(frule_tac x = x in bspec, simp add: domI)
apply(erule wf_objectE) apply(clarsimp) apply(frule r_field_write_npeI) apply(force)
apply(clarsimp) apply(erule sty_option.cases) apply(clarsimp) apply(rename_tac ty_y ty_f)
apply(drule_tac x = y in bspec, simp add: domI) apply(clarsimp)
apply(erule sty_option.cases) apply(clarsimp split: option.splits)
apply(erule wf_objectE) apply(clarsimp)
apply(frule_tac H = H and y = y in r_field_writeI, simp, force)+
\<comment> \<open>conditional branch\<close>
apply(clarsimp, erule wf_stmtE, simp_all, clarsimp) apply(erule disjE)
apply(frule type_to_val, simp, clarify) apply(case_tac "v = w")
apply(frule_tac y = y in r_if_trueI, force+)
apply(frule_tac y = y in r_if_falseI, force+)
apply(frule type_to_val, simp, clarify) apply(case_tac "v = w")
apply(frule_tac y = y and v = w in r_if_trueI, force+)
apply(frule_tac y = y and v = w in r_if_falseI, force+)
\<comment> \<open>object creation\<close>
apply(clarsimp, erule wf_stmtE, simp_all, clarsimp) apply(rename_tac cl ctx ty var)
apply(erule sty_option.cases) apply(clarsimp) apply(rename_tac ty_cl ty_var)
apply(simp add: is_sty_one_def split: option.splits) apply(rename_tac path)
apply(frule find_path_fields) apply(erule exE)
apply(frule fresh_oid) apply(erule exE)
apply(frule_tac H = H and L = L and var = var and s_list = ss and f_list = fs in r_newI[simplified])
apply(clarsimp simp add: fields_f_def split: option.splits) apply(assumption) apply(simp) apply(force)
\<comment> \<open>method call\<close>
apply(clarsimp, erule wf_stmtE, simp_all, clarsimp)
apply(rename_tac ss y_ty_list x ty_x_s m ty_r_s var)
apply(erule wf_varstateE) apply(frule_tac x = x in bspec, simp add: domI) apply(clarsimp)
apply(erule wf_objectE) apply(clarsimp) apply(frule r_mcall_npeI) apply(force)
apply(elim sty_option.cases) apply(clarsimp split: option.splits)
apply(rename_tac ty_r_s ty_var_s ty_x_s ty_x_d fs_x)
apply(frule mtype_to_find_md, simp+) apply(clarsimp)
apply(frule_tac A = "dom L" and i = "length vds" in fresh_vars) apply(clarsimp) apply(rename_tac vars')
apply(frule exist_lifted_values) apply(simp) apply(clarify)
apply(frule_tac vars' = vars' in fresh_x_not_in_vars') apply(erule exE) apply(erule conjE)
apply(subgoal_tac "\<exists>vars. vars = map (\<lambda>vd. case vd of vd_def cl var \<Rightarrow> var) vds")
apply(erule exE) apply(subgoal_tac "length vars = length vds")
apply(frule length_y_ty_list_vs)
apply(subgoal_tac "\<exists>T. T = (map_of (zip (map (\<lambda>vd. case vd of vd_def cl var \<Rightarrow> x_var var) vds) (map x_var vars')))(x_this \<mapsto> x')")
apply(erule exE)
apply(frule_tac H = H and P = P and meth = m and ctx = ctx and cl = cl_r and y = y and ty = ty_x_d and y_cl_var_var'_v_list =
"zip (map fst y_ty_list) (zip (map (\<lambda>vd. case vd of vd_def cl var \<Rightarrow> cl) vds) (zip vars (zip vars' vs)))"
and s''_s'_list = "zip (tr_ss_f T ss') ss'" and var = var and s_list = ss and x' = x' and T = T
in r_mcallI[simplified])
apply(force) apply(simp) apply(simp) apply(simp add: vars'_eq) apply(simp) apply(assumption) apply(simp add: vars'_eq)
apply(cut_tac L = L and y_ty_list = y_ty_list in lift_opts_mapping) apply(erule_tac x = vds in allE)
apply(erule_tac x = vars in allE) apply(erule_tac x = vars' in allE) apply(erule_tac x = vs in allE)
apply(simp) apply(simp) apply(simp) apply(simp add: translation_eq) apply(simp) apply(force)
by force+
theorem wf_preservation:
"\<And>\<Gamma> config.
\<lbrakk>wf_config \<Gamma> config; r_stmt config config'\<rbrakk>
\<Longrightarrow> \<exists>\<Gamma>'. \<Gamma> \<subseteq>\<^sub>m \<Gamma>' \<and> wf_config \<Gamma>' config'"
+supply [[simproc del: defined_all]]
apply(erule r_stmt.cases)
(* s_read_npe, s_write_npe *)
(* \<Gamma>' = \<Gamma> for all cases except for mcall (case 11) *)
apply(rule_tac[1-10] x = \<Gamma> in exI)
apply(erule_tac[1-11] wf_configE)
apply(simp_all)
(* s_if s_block *)
apply(clarsimp) apply(erule wf_stmtE) apply(simp_all)
apply(force intro: wf_config_normalI)
(* s_ass *)
apply(clarsimp) apply(erule wf_stmtE) apply(simp_all)
apply(rule wf_config_normalI, assumption+)
apply(erule sty_option.cases)
apply(rule wf_subvarstateI, assumption+, simp)
apply(erule wf_varstateE)
apply(drule_tac x = x in bspec, simp add: domI)
apply(erule wf_objectE)
apply(simp add: wf_nullI)
apply(clarsimp)
apply(rule wf_objectI)
apply(erule sty_option.cases, simp)
apply(rule sty_optionI, simp+)
apply(rule_tac ty' = ty'a in sty_transitiveI, assumption+)
(* s_read *)
apply(force intro: wf_intros)
apply(clarsimp split: option.splits) apply(erule wf_stmtE) apply(simp_all)
apply(rule wf_config_normalI, assumption+)
apply(clarsimp)
apply(rename_tac L oid H v s_list ty_oid fs_oid \<Gamma> x ty_x P f ty_f var)
apply(erule sty_option.cases)
apply(rule wf_subvarstateI, assumption, simp)
apply(clarsimp)
apply(case_tac v)
apply(clarify)
apply(rule wf_nullI, simp)
apply(clarify)
apply(rename_tac ty_f ty_var P oid_v)
apply(erule wf_heapE)
apply(drule_tac x = oid in bspec, simp add: domI)
apply(clarsimp)
apply(rename_tac H P fs_oid)
apply(erule wf_varstateE)
apply(frule_tac x = x in bspec, simp add: domI)
apply(clarsimp)
apply(erule wf_objectE)
apply(simp)
apply(erule sty_option.cases)
apply(clarsimp split: option.splits)
apply(rename_tac L \<Gamma> H oid_x ty_x_d ty_x_s P)
apply(frule_tac ty_x_s = ty_x_s in no_field_hiding, simp+)
apply(drule_tac x = f in bspec, simp)
apply(clarsimp)
apply(frule ftype_preservation[simplified], assumption+)
apply(clarsimp)
apply(rename_tac ty_f)
apply(erule wf_objectE)
apply(simp)
apply(clarsimp)
apply(erule sty_option.cases)
apply(clarsimp split: option.splits)
apply(rename_tac H oid_v ty_f P ty_v fs_v)
apply(rule wf_objectI)
apply(rule sty_optionI, simp+)
apply(rule sty_transitiveI, assumption+)
(* s_write *)
apply(force intro: wf_intros)
apply(clarsimp, hypsubst_thin) apply(erule wf_stmtE) apply(simp_all)
apply(clarsimp split: option.splits) apply(rule)
apply(erule wf_varstateE) apply(clarsimp) apply(rename_tac \<Gamma> P H)
apply(drule_tac x = xa in bspec, simp add: domI)
apply(erule wf_objectE) apply(simp) apply(clarsimp)
apply(elim sty_option.cases) apply(simp)
apply(erule sty_option.cases)
apply(clarsimp)
apply(rule wf_config_normalI)
apply(assumption)
apply(rename_tac L oid_x v H ss \<Gamma> x ty_x_s f y ty_y_s ty_f P ty_x_d fs_oid)
apply(erule wf_heapE)
apply(rule wf_heapI) apply(simp)
apply(rule ballI) apply(clarsimp simp add: map_add_def) apply(rule)
apply(clarsimp) apply(drule_tac x = oid_x in bspec, simp add: domI) apply(clarsimp)
apply(drule_tac x = fa in bspec, simp) apply(clarsimp)
apply(rule)
apply(clarsimp) apply(case_tac v) apply(clarsimp simp add: wf_nullI) apply(clarsimp)
apply(rename_tac H P fields_x_d ty_f_d oid_v)
apply(erule wf_varstateE) apply(clarsimp)
apply(frule ftype_preservation, assumption+) apply(clarsimp)
apply(drule_tac x = y in bspec, simp add: domI) apply(clarsimp)
apply(erule_tac ?a4.0 = "Some ty_y_s" in wf_objectE) apply(simp) apply(clarsimp)
apply(erule sty_option.cases) apply(clarsimp split: option.splits)
apply(rule wf_objectI) apply(force intro: sty_optionI sty_transitiveI)
apply(clarsimp) apply(case_tac "fs_oid fa") apply(clarsimp) apply(force elim: wf_objectE)
apply(case_tac a) apply(force intro: wf_nullI)
apply(clarsimp) apply(rename_tac H P fields_x_d f ty_f_d oid_v)
apply(erule wf_varstateE) apply(clarsimp)
apply(drule_tac x = y in bspec, simp add: domI) apply(clarsimp)
apply(erule wf_objectE) apply(simp) apply(clarsimp)
apply(rule wf_objectI) apply(erule sty_option.cases) apply(clarsimp split: option.splits)
apply(safe) apply(clarsimp) apply(force intro: sty_optionI) apply(force intro: sty_optionI)
apply(drule_tac x = oid in bspec) apply(clarsimp)
apply(clarsimp) apply(drule_tac x = fa in bspec, simp)
apply(clarsimp simp add: wf_object_heap)
apply(rename_tac L oid_x v H ss \<Gamma> x ty_x_s f y ty_y_s ty_f P ty_x_d fs_x)
apply(erule wf_varstateE)
apply(rule wf_varstateI) apply(simp)
apply(clarsimp)
apply(rename_tac L \<Gamma> P H x' y')
apply(drule_tac x = x' in bspec, simp add: domI)
apply(clarsimp)
apply(erule wf_objectE)
apply(clarsimp)
apply(rule wf_nullI, simp)
apply(clarsimp)
apply(erule sty_option.cases)
apply(clarsimp split: option.splits)
apply(rename_tac H oid_x' ty_x'_s P ty_x'_d fs_x')
apply(rule wf_objectI)
apply(rule sty_optionI)
apply(clarsimp)
apply(rule conjI)
apply(clarsimp)
apply(simp)
apply(clarsimp)
apply(simp)
apply(simp)
(* s_if *)
apply(erule wf_stmtE) apply(simp_all)
apply(force intro: wf_config_normalI)
apply(erule wf_stmtE) apply(simp_all)
apply(force intro: wf_config_normalI)
(* s_new *)
apply(erule contrapos_np)
apply(erule wf_stmtE, simp_all, clarsimp)
apply(erule sty_option.cases, clarsimp)
apply(rule wf_config_normalI)
apply(assumption)
apply(rename_tac H L ss ctx cl \<Gamma> var ty' ty_var P)
apply(erule wf_heapE)
apply(rule wf_heapI) apply(simp)
apply(safe) apply(simp add: map_upd_Some_unfold) apply(split if_split_asm)
apply(clarsimp) apply(frule field_has_type, simp+)
apply(erule exE) apply(rule_tac x = ty in exI) apply(simp) apply(force intro: wf_nullI)
apply(drule_tac x = oida in bspec, simp add: domI) apply(clarsimp split: option.splits)
apply(drule_tac x = f in bspec, simp) apply(safe) apply(rule_tac x = ty'a in exI) apply(simp)
apply(erule wf_objectE) apply(clarsimp) apply(force intro: wf_nullI)
apply(erule sty_option.cases) apply(clarsimp split: option.splits)
apply(rule wf_objectI) apply(clarsimp) apply(rule conjI) apply(force)
apply(force intro: sty_optionI)
apply(rename_tac oid_new H L ss ctx cl \<Gamma> var ty_new ty_var P)
apply(erule wf_varstateE) apply(clarsimp)
apply(rule wf_varstateI) apply(simp) apply(rule ballI) apply(drule_tac x = x in bspec, simp)
apply(clarsimp) apply(rule conjI)
apply(clarsimp) apply(rule wf_objectI) apply(clarsimp) apply(force intro: sty_optionI)
apply(rule) apply(erule wf_objectE) apply(clarsimp) apply(force intro: wf_nullI)
apply(erule sty_option.cases) apply(clarsimp split: option.splits)
apply(rule wf_objectI) apply(clarsimp) apply(rule) apply(rule) apply(force)
apply(rule) apply(force intro: sty_optionI)
(* s_call *)
apply(force intro: wf_all_exI)
(* setting the new type environment *)
apply(case_tac "H oid", simp) apply(clarsimp)
apply(erule wf_stmtE, simp_all, clarsimp)
apply(frule wf_method_from_find[simplified]) apply(simp) apply(safe)
apply(erule_tac Q = "\<forall>\<Gamma>'. \<Gamma> \<subseteq>\<^sub>m \<Gamma>' \<longrightarrow>
~ (wf_config \<Gamma>' (config_normal Pa ((L ++ map_of (map (\<lambda>(y_XXX, cl_XXX, var_XXX, var_', y). (x_var var_', y)) y_cl_var_var'_v_list))(x' \<mapsto> v_oid oid)) H
(map fst s''_s'_list @
s_ass vara
(case if y = x_this then Some x'
else map_of (map (\<lambda>(y_XXX, cl_XXX, var_XXX, var_', v_XXX). (x_var var_XXX, x_var var_')) y_cl_var_var'_v_list) y of
None \<Rightarrow> y | Some x' \<Rightarrow> x') #
s_list)))" in contrapos_pp) apply(simp only: not_all)
apply(rule_tac x = "\<Gamma> ++
map_of (map (\<lambda>((y, cl, var, var', v), (y', ty)). (x_var var', ty)) (zip y_cl_var_var'_v_list y_ty_list)) ++
Map.empty (x' \<mapsto> ty_def ctx dcl)" in exI)
apply(clarsimp split del: if_split) apply(rule)
apply(clarsimp simp add: map_le_def split del: if_split) apply(rule sym)
apply(rule_tac L = L and Pa = Pa and H = H in map_of_map_and_x) apply(assumption+) apply(simp add: not_dom_is_none) apply(erule wf_varstateE) apply(clarsimp)
apply(rename_tac ss ty_x_d fs_x y_ty_list \<Gamma> x ty_x_s P meth ty_r var dcl)
apply(erule sty_option.cases) apply(clarsimp split del: if_split) apply(rename_tac ty_r ty_var P)
apply(erule wf_varstateE, clarify) apply(rename_tac L \<Gamma> P H)
apply(frule_tac x = x in bspec, simp add: domI)
apply(frule_tac x = "x_var var" in bspec, simp add: domI)
apply(clarsimp split del: if_split)
apply(erule wf_objectE, simp, clarsimp split del: if_split) apply(rename_tac P H oid)
apply(erule sty_option.cases, clarsimp split del: if_split) apply(rename_tac ty_x_d ty_x_s P)
apply(subgoal_tac "x_var var \<in> dom \<Gamma>")
apply(clarify) apply(rename_tac v_var)
apply(drule not_dom_is_none)
apply(rule wf_config_normalI)
apply(assumption)
apply(assumption)
(* varstate *)
apply(rule wf_varstateI) apply(simp only: finite_super_varstate) apply(clarsimp) apply(rule)
apply(rule wf_objectI) apply(simp) apply(rule sty_optionI, simp+)
apply(clarsimp, hypsubst_thin) apply(rule) apply(rule) apply(rule wf_objectI) apply(clarsimp) apply(rule sty_optionI, simp+)
apply(rule) apply(erule disjE)
apply(clarsimp) apply(drule map_of_SomeD) apply(clarsimp) apply(rename_tac y_k cl_k var_k var'_k v_k y'_k ty_k)
apply(frule_tac x = "(y_k, cl_k, var_k, var'_k, v_k)" in bspec) apply(force simp add: set_zip) apply(clarsimp)
apply(subgoal_tac "map_of (map (\<lambda>(y, cl, var, var', v). (x_var var', v)) y_cl_var_var'_v_list) (x_var var'_k) = Some v_k")
apply(clarsimp) apply(drule map_y) apply(simp) apply(drule_tac x = "(y_k, ty_k)" in bspec, assumption) apply(clarsimp)
apply(erule sty_option.cases) apply(clarsimp) apply(rename_tac ty_y_k ty_k P) apply(drule_tac x = y_k in bspec, simp add: domI)
apply(clarsimp) apply(case_tac v_k) apply(clarify, rule wf_nullI, simp) apply(clarsimp) apply(rename_tac oid_y_k)
apply(erule_tac ?a4.0 = "Some ty_y_k" in wf_objectE) apply(simp) apply(erule sty_option.cases)
apply(clarsimp split: option.splits) apply(rule wf_objectI) apply(clarsimp) apply(rule sty_optionI, simp+)
apply(rule_tac ty' = ty' in sty_transitiveI, assumption+)
apply(rule map_of_is_SomeI) apply(simp add: map_fst_var'[simplified]) apply(rule set_zip_var'_v, simp)
apply(clarify) apply(rename_tac x_G ty_x_G) apply(frule_tac x = x_G in bspec, simp add: domI)
apply(case_tac "map_of (map (\<lambda>((y, cl, var, var', v), y', y). (x_var var', y)) (zip y_cl_var_var'_v_list y_ty_list)) x_G")
apply(frule map_of_map_zip_none, simp+) apply(simp add: map_add_def)
apply(frule map_of_map_zip_some, simp+) apply(clarsimp)
apply(drule map_of_SomeD) apply(drule map_of_SomeD) apply(clarsimp simp add: set_zip)
apply(frule same_el, simp+) apply(drule_tac s = "(aa, ab, ac, ai, b)" in sym)
apply(clarsimp) apply(rename_tac y_k cl_k var_k v_k y'_k ty_k var'_k i)
apply(frule_tac y_k = y_k and cl_k = cl_k and var_k = var_k and var'_k = var'_k and v_k = v_k in same_y, simp+) apply(clarsimp)
apply(drule_tac x = "(y_k, ty_k)" in bspec, simp) apply(clarsimp) apply(rename_tac y_k ty_k)
apply(erule sty_option.cases) apply(clarsimp) apply(drule_tac x = y_k in bspec, simp add: domI) apply(clarsimp)
apply(drule_tac x = "(y_k, cl_k, var_k, var'_k, v_k)" in bspec) apply(drule_tac t = "(y_k, cl_k, var_k, var'_k, v_k)" in sym) apply(simp)
apply(clarsimp) apply(erule_tac ?a4.0 = "Some ty" in wf_objectE) apply(clarsimp) apply(rule wf_nullI, simp)
apply(erule sty_option.cases) apply(clarsimp split: option.splits) apply(rule wf_objectI) apply(clarsimp)
apply(rule sty_optionI, simp+) apply(rule_tac ty' = ty'a in sty_transitiveI, assumption+)
(* assignment *)
apply(clarsimp split del: if_split) apply(rule)
apply(rule wf_var_assignI) apply(frule wf_method_from_find) apply(simp) apply(erule exE) apply(erule wf_methE)
apply(case_tac "ya = x_this")
apply(clarsimp split del: if_split)
apply(rule_tac ty' = ty_var in sty_optionI) apply(simp) apply(simp) apply(case_tac ctx, clarify)
apply(frule_tac ty_r_d = ty in mty_preservation, assumption+) apply(clarify)
apply(erule sty_option.cases) apply(clarsimp) apply(rule_tac ty' = ty' in sty_transitiveI, assumption+)
(* y != this *)
apply(erule sty_option.cases) apply(clarsimp split del: if_split)
apply(frule_tac var_k = ya and ty_k = tya in exists_var') apply(drule map_of_vd) apply(assumption+) apply(erule exE) apply(clarsimp)
apply(drule_tac y = "x_var var'_k" in map_of_SomeD) apply(clarsimp split del: if_split) apply(rename_tac y_k cl_k var_k var'_k v_k)
apply(frule x'_not_in_list, assumption+) apply(clarsimp) apply(drule map_of_SomeD)
apply(clarsimp split del: if_split) apply(rename_tac cl_k' var_k ty_k) apply(case_tac ctx, clarify) apply(frule mty_preservation, assumption+)
apply(clarsimp split del: if_split) apply(drule map_of_vd) apply(frule map_of_ty_k, assumption+)
apply(rule_tac ty = ty_k and ty' = ty_var in sty_optionI) apply(simp add: map_add_def) apply(simp) apply(simp)
apply(rule_tac ty' = ty_r in sty_transitiveI) apply(assumption+)
(* wf of translated statements *)
apply(clarsimp) apply(erule disjE) apply(erule wf_methE) apply(clarsimp)
apply(frule map_of_vd[rule_format]) apply(case_tac ctx, clarify) apply(frule mty_preservation, simp+) apply(clarsimp)
apply(rename_tac P dcl cl y meth s'' s')
apply(drule_tac x = "(s'', s')" in bspec, simp)+ apply(clarsimp)
apply(cut_tac cl_var_ty_list = cl_var_ty_list and y_cl_var_var'_v_list = y_cl_var_var'_v_list and y_ty_list = y_ty_list and
P = P and ty_x_d = ty_x_d and ty_x_m = "ty_def ctx_def dcl" and x' = x' and s'' = s'' and \<Gamma> = \<Gamma> in wf_tr_stmt_in_G')
apply(clarsimp)
(* wf of non-translated statements *)
apply(cut_tac L = L and x' = x' and y_cl_var_var'_v_list = y_cl_var_var'_v_list and \<Gamma> = \<Gamma> and P = P and H = H and
y_ty_list = y_ty_list and ty = "ty_def ctx dcl" and ss = ss in wf_stmt_in_G')
apply(clarsimp)
apply(erule wf_objectE) apply(force) apply(force)
done
end
diff --git a/thys/List_Interleaving/ListInterleaving.thy b/thys/List_Interleaving/ListInterleaving.thy
--- a/thys/List_Interleaving/ListInterleaving.thy
+++ b/thys/List_Interleaving/ListInterleaving.thy
@@ -1,1136 +1,1140 @@
(* Title: Reasoning about Lists via List Interleaving
Author: Pasquale Noce
Security Certification Specialist at Arjo Systems - Gep S.p.A.
pasquale dot noce dot lavoro at gmail dot com
pasquale dot noce at arjowiggins-it dot com
*)
section "List interleaving"
theory ListInterleaving
imports Main
begin
text \<open>
\null
Among the various mathematical tools introduced in his outstanding work on Communicating Sequential
Processes \cite{R2}, Hoare has defined \emph{interleaves} as the predicate satisfied by any three
lists \emph{s}, \emph{t}, emph{u} such that \emph{s} may be split into sublists alternately
extracted from \emph{t} and \emph{u}, whatever is the criterion for extracting an item from either
\emph{t} or \emph{u} in each step.
This paper enriches Hoare's definition by identifying such criterion with the truth value of a
predicate taking as inputs the head and the tail of \emph{s}. This enhanced \emph{interleaves}
predicate turns out to permit the proof of equalities between lists without the need of an
induction. Some rules that allow to infer \emph{interleaves} statements without induction,
particularly applying to the addition of a prefix to the input lists, are also proven. Finally, a
stronger version of the predicate, named \emph{Interleaves}, is shown to fulfil further rules
applying to the addition of a suffix to the input lists.
Throughout this paper, the salient points of definitions and proofs are commented; for additional
information, cf. Isabelle documentation, particularly \cite{R3}, \cite{R4}, \cite{R5}, and
\cite{R6}. For a sample nontrivial application of the mathematical machinery developed in this
paper, cf. \cite{R1}.
\<close>
subsection "A first version of interleaving"
text \<open>
Here below is the definition of predicate \<open>interleaves\<close>, as well as of a convenient symbolic
notation for it. As in the definition of predicate \emph{interleaves} formulated in \cite{R2}, the
recursive decomposition of the input lists is performed by item prepending. In the case of a list
\<open>ws\<close> constructed recursively by item appending rather than prepending, the statement that it
satisfies predicate \<open>interleaves\<close> with two further lists can nevertheless be proven by
induction using as input @{term "rev ws"}, rather than \<open>ws\<close> itself.
With respect to Hoare's homonymous predicate, \<open>interleaves\<close> takes as an additional input a
predicate \<open>P\<close>, which is a function of a single item and a list. Then, for statement
@{term "interleaves P (x # xs) (y # ys) (z # zs)"} to hold, the item picked for being \<open>x\<close> must
be \<open>y\<close> if \<open>P x xs\<close>, \<open>z\<close> otherwise. On the contrary, in case either the second or
the third list is empty, the truth value of \<open>P x xs\<close> does not matter and list
@{term "x # xs"} just has to match the other nonempty one, if any.
\null
\<close>
fun interleaves ::
"('a \<Rightarrow> 'a list \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"interleaves P (x # xs) (y # ys) (z # zs) = (if P x xs
then x = y \<and> interleaves P xs ys (z # zs)
else x = z \<and> interleaves P xs (y # ys) zs)" |
"interleaves P (x # xs) (y # ys) [] =
(x = y \<and> interleaves P xs ys [])" |
"interleaves P (x # xs) [] (z # zs) =
(x = z \<and> interleaves P xs [] zs)" |
"interleaves _ (_ # _) [] [] = False" |
"interleaves _ [] (_ # _) _ = False" |
"interleaves _ [] _ (_ # _) = False" |
"interleaves _ [] [] [] = True"
abbreviation interleaves_syntax ::
"'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> ('a \<Rightarrow> 'a list \<Rightarrow> bool) \<Rightarrow> bool"
("(_ \<simeq> {_, _, _})" [60, 60, 60] 51)
where "xs \<simeq> {ys, zs, P} \<equiv> interleaves P xs ys zs"
text \<open>
\null
The advantage provided by this enhanced \emph{interleaves} predicate is that in case
@{term "xs \<simeq> {ys, zs, P}"}, fixing the values of \<open>xs\<close> and either \<open>ys\<close> or \<open>zs\<close> has
the effect of fixing the value of the remaining list, too. Namely, if @{term "xs \<simeq> {ys', zs, P}"}
also holds, then @{term "ys = ys'"}, and likewise, if @{term "xs \<simeq> {ys, zs', P}"} also holds, then
@{term "zs = zs'"}. Therefore, once two \emph{interleaves} statements @{term "xs \<simeq> {ys, zs, P}"},
@{term "xs' \<simeq> {ys', zs', P'}"} have been proven along with equalities @{term "xs = xs'"},
@{term "P = P'"}, and either @{term "zs = zs'"} or @{term "ys = ys'"}, possibly by induction, the
remaining equality, i.e. respectively @{term "ys = ys'"} and @{term "zs = zs'"}, can be inferred
without the need of a further induction.
Here below is the proof of this property as well as of other ones. Particularly, it is also proven
that in case @{term "xs \<simeq> {ys, zs, P}"}, lists \<open>ys\<close> and \<open>zs\<close> can be swapped by
replacing predicate \<open>P\<close> with its negation.
It is worth noting that fixing the values of \<open>ys\<close> and \<open>zs\<close> does not fix the value of
\<open>xs\<close> instead. A counterexample is @{term "ys = [y]"}, @{term "zs = [z]"} with @{term "y \<noteq> z"},
@{term "P y [z] = True"}, @{term "P z [y] = False"}, in which case both of the \emph{interleaves}
statements @{term "[y, z] \<simeq> {ys, zs, P}"} and @{term "[z, y] \<simeq> {ys, zs, P}"} hold.
\null
\<close>
lemma interleaves_length [rule_format]:
"xs \<simeq> {ys, zs, P} \<longrightarrow> length xs = length ys + length zs"
proof (induction P xs ys zs rule: interleaves.induct, simp_all)
qed (rule conjI, (rule_tac [!] impI)+, simp_all)
lemma interleaves_nil:
"[] \<simeq> {ys, zs, P} \<Longrightarrow> ys = [] \<and> zs = []"
by (rule interleaves.cases [of "(P, [], ys, zs)"], simp_all)
lemma interleaves_swap:
"xs \<simeq> {ys, zs, P} = xs \<simeq> {zs, ys, \<lambda>w ws. \<not> P w ws}"
proof (induction P xs ys zs rule: interleaves.induct, simp_all)
fix y' :: 'a and ys' zs' P'
show "\<not> [] \<simeq> {zs', y' # ys', \<lambda>w ws. \<not> P' w ws}" by (cases zs', simp_all)
qed
lemma interleaves_equal_fst [rule_format]:
"xs \<simeq> {ys, zs, P} \<longrightarrow> xs \<simeq> {ys', zs, P} \<longrightarrow> ys = ys'"
proof (induction xs arbitrary: ys ys' zs, (rule_tac [!] impI)+)
fix ys ys' zs
assume "[] \<simeq> {ys, zs, P}"
hence "ys = [] \<and> zs = []" by (rule interleaves_nil)
moreover assume "[] \<simeq> {ys', zs, P}"
hence "ys' = [] \<and> zs = []" by (rule interleaves_nil)
ultimately show "ys = ys'" by simp
next
fix x xs ys ys' zs
assume
A: "\<And>ys ys' zs. xs \<simeq> {ys, zs, P} \<longrightarrow> xs \<simeq> {ys', zs, P} \<longrightarrow> ys = ys'" and
B: "x # xs \<simeq> {ys, zs, P}" and
B': "x # xs \<simeq> {ys', zs, P}"
show "ys = ys'"
proof (cases zs, case_tac [2] ys, case_tac [2-3] ys', simp_all)
assume C: "zs = []"
hence "\<exists>w ws. ys = w # ws" using B by (cases ys, simp_all)
then obtain w ws where Y: "ys = w # ws" by blast
hence D: "w = x" using B and C by simp
have "xs \<simeq> {ws, [], P}" using B and C and Y by simp
moreover have "\<exists>w' ws'. ys' = w' # ws'"
using B' and C by (cases ys', simp_all)
then obtain w' ws' where Y': "ys' = w' # ws'" by blast
hence D': "w' = x" using B' and C by simp
have "xs \<simeq> {ws', [], P}" using B' and C and Y' by simp
moreover have "xs \<simeq> {ws, [], P} \<longrightarrow> xs \<simeq> {ws', [], P} \<longrightarrow> ws = ws'"
using A .
ultimately have "ws = ws'" by simp
with Y and Y' and D and D' show ?thesis by simp
next
fix v vs w' ws'
assume C: "zs = v # vs" and "ys = []"
hence D: "xs \<simeq> {[], vs, P}" using B by simp
assume E: "ys' = w' # ws'"
have "P x xs \<or> \<not> P x xs" by simp
moreover {
assume "P x xs"
hence "xs \<simeq> {ws', v # vs, P}" using B' and C and E by simp
hence "length xs = Suc (length vs) + length ws'"
by (simp add: interleaves_length)
moreover have "length xs = length vs"
using D by (simp add: interleaves_length)
ultimately have False by simp
}
moreover {
assume "\<not> P x xs"
hence "xs \<simeq> {w' # ws', vs, P}" using B' and C and E by simp
moreover have "xs \<simeq> {[], vs, P} \<longrightarrow> xs \<simeq> {w' # ws', vs, P} \<longrightarrow>
[] = w' # ws'"
using A .
ultimately have "[] = w' # ws'" using D by simp
hence False by simp
}
ultimately show False ..
next
fix v vs w ws
assume C: "zs = v # vs" and "ys' = []"
hence D: "xs \<simeq> {[], vs, P}" using B' by simp
assume E: "ys = w # ws"
have "P x xs \<or> \<not> P x xs" by simp
moreover {
assume "P x xs"
hence "xs \<simeq> {ws, v # vs, P}" using B and C and E by simp
hence "length xs = Suc (length vs) + length ws"
by (simp add: interleaves_length)
moreover have "length xs = length vs"
using D by (simp add: interleaves_length)
ultimately have False by simp
}
moreover {
assume "\<not> P x xs"
hence "xs \<simeq> {w # ws, vs, P}" using B and C and E by simp
moreover have "xs \<simeq> {[], vs, P} \<longrightarrow> xs \<simeq> {w # ws, vs, P} \<longrightarrow> [] = w # ws"
using A .
ultimately have "[] = w # ws" using D by simp
hence False by simp
}
ultimately show False ..
next
fix v vs w ws w' ws'
assume C: "zs = v # vs" and D: "ys = w # ws" and D': "ys' = w' # ws'"
have "P x xs \<or> \<not> P x xs" by simp
moreover {
assume E: "P x xs"
hence F: "w = x" using B and C and D by simp
have "xs \<simeq> {ws, v # vs, P}" using B and C and D and E by simp
moreover have F': "w' = x" using B' and C and D' and E by simp
have "xs \<simeq> {ws', v # vs, P}" using B' and C and D' and E by simp
moreover have "xs \<simeq> {ws, v # vs, P} \<longrightarrow> xs \<simeq> {ws', v # vs, P} \<longrightarrow>
ws = ws'"
using A .
ultimately have "ws = ws'" by simp
hence "w = w' \<and> ws = ws'" using F and F' by simp
}
moreover {
assume E: "\<not> P x xs"
hence "xs \<simeq> {w # ws, vs, P}" using B and C and D by simp
moreover have "xs \<simeq> {w' # ws', vs, P}"
using B' and C and D' and E by simp
moreover have "xs \<simeq> {w # ws, vs, P} \<longrightarrow> xs \<simeq> {w' # ws', vs, P} \<longrightarrow>
w # ws = w' # ws'"
using A .
ultimately have "w # ws = w' # ws'" by simp
hence "w = w' \<and> ws = ws'" by simp
}
ultimately show "w = w' \<and> ws = ws'" ..
qed
qed
lemma interleaves_equal_snd:
"xs \<simeq> {ys, zs, P} \<Longrightarrow> xs \<simeq> {ys, zs', P} \<Longrightarrow> zs = zs'"
by (subst (asm) (1 2) interleaves_swap, rule interleaves_equal_fst)
text \<open>
\null
Since \emph{interleaves} statements permit to prove equalities between lists without the need of an
induction, it is useful to search for rules that allow to infer such statements themselves without
induction, which is precisely what is done here below. Particularly, it is proven that under proper
assumptions, predicate @{term interleaves} keeps being satisfied by applying a filter, a mapping, or
the addition or removal of a prefix to the input lists.
\null
\<close>
lemma interleaves_all_nil:
"xs \<simeq> {xs, [], P}"
by (induction xs, simp_all)
lemma interleaves_nil_all:
"xs \<simeq> {[], xs, P}"
by (induction xs, simp_all)
lemma interleaves_equal_all_nil:
"xs \<simeq> {ys, [], P} \<Longrightarrow> xs = ys"
by (insert interleaves_all_nil, rule interleaves_equal_fst)
lemma interleaves_equal_nil_all:
"xs \<simeq> {[], zs, P} \<Longrightarrow> xs = zs"
by (insert interleaves_nil_all, rule interleaves_equal_snd)
lemma interleaves_filter [rule_format]:
assumes A: "\<forall>x xs. P x (filter Q xs) = P x xs"
shows "xs \<simeq> {ys, zs, P} \<longrightarrow> filter Q xs \<simeq> {filter Q ys, filter Q zs, P}"
proof (induction xs arbitrary: ys zs, rule_tac [!] impI, simp)
fix ys zs
assume "[] \<simeq> {ys, zs, P}"
hence "ys = [] \<and> zs = []" by (rule interleaves_nil)
thus "[] \<simeq> {filter Q ys, filter Q zs, P}" by simp
next
fix x xs ys zs
assume
B: "\<And>ys' zs'. xs \<simeq> {ys', zs', P} \<longrightarrow>
filter Q xs \<simeq> {filter Q ys', filter Q zs', P}" and
C: "x # xs \<simeq> {ys, zs, P}"
show "filter Q (x # xs) \<simeq> {filter Q ys, filter Q zs, P}"
proof (cases ys, case_tac [!] zs, simp_all del: filter.simps, rule ccontr)
assume "ys = []" and "zs = []"
thus False using C by simp
next
fix z zs'
assume "ys = []" and "zs = z # zs'"
hence D: "x = z \<and> xs \<simeq> {[], zs', P}" using C by simp
moreover have "xs \<simeq> {[], zs', P} \<longrightarrow>
filter Q xs \<simeq> {filter Q [], filter Q zs', P}"
using B .
ultimately have "filter Q xs \<simeq> {[], filter Q zs', P}" by simp
thus "filter Q (x # xs) \<simeq> {[], filter Q (z # zs'), P}" using D by simp
next
fix y ys'
assume "ys = y # ys'" and "zs = []"
hence D: "x = y \<and> xs \<simeq> {ys', [], P}" using C by simp
moreover have "xs \<simeq> {ys', [], P} \<longrightarrow>
filter Q xs \<simeq> {filter Q ys', filter Q [], P}"
using B .
ultimately have "filter Q xs \<simeq> {filter Q ys', [], P}" by simp
thus "filter Q (x # xs) \<simeq> {filter Q (y # ys'), [], P}" using D by simp
next
fix y ys' z zs'
assume "ys = y # ys'" and "zs = z # zs'"
hence D: "x # xs \<simeq> {y # ys', z # zs', P}" using C by simp
show "filter Q (x # xs) \<simeq> {filter Q (y # ys'), filter Q (z # zs'), P}"
proof (cases "P x xs")
case True
hence E: "P x (filter Q xs)" using A by simp
have F: "x = y \<and> xs \<simeq> {ys', z # zs', P}" using D and True by simp
moreover have "xs \<simeq> {ys', z # zs', P} \<longrightarrow>
filter Q xs \<simeq> {filter Q ys', filter Q (z # zs'), P}"
using B .
ultimately have G: "filter Q xs \<simeq> {filter Q ys', filter Q (z # zs'), P}"
by simp
show ?thesis
proof (cases "Q x")
assume "Q x"
hence "filter Q (x # xs) = x # filter Q xs" by simp
moreover have "filter Q (y # ys') = x # filter Q ys'"
using \<open>Q x\<close> and F by simp
ultimately show ?thesis using E and G
by (cases "filter Q (z # zs')", simp_all)
next
assume "\<not> Q x"
hence "filter Q (x # xs) = filter Q xs" by simp
moreover have "filter Q (y # ys') = filter Q ys'"
using \<open>\<not> Q x\<close> and F by simp
ultimately show ?thesis using E and G
by (cases "filter Q (z # zs')", simp_all)
qed
next
case False
hence E: "\<not> P x (filter Q xs)" using A by simp
have F: "x = z \<and> xs \<simeq> {y # ys', zs', P}" using D and False by simp
moreover have "xs \<simeq> {y # ys', zs', P} \<longrightarrow>
filter Q xs \<simeq> {filter Q (y # ys'), filter Q zs', P}"
using B .
ultimately have G: "filter Q xs \<simeq> {filter Q (y # ys'), filter Q zs', P}"
by simp
show ?thesis
proof (cases "Q x")
assume "Q x"
hence "filter Q (x # xs) = x # filter Q xs" by simp
moreover have "filter Q (z # zs') = x # filter Q zs'"
using \<open>Q x\<close> and F by simp
ultimately show ?thesis using E and G
by (cases "filter Q (y # ys')", simp_all)
next
assume "\<not> Q x"
hence "filter Q (x # xs) = filter Q xs" by simp
moreover have "filter Q (z # zs') = filter Q zs'"
using \<open>\<not> Q x\<close> and F by simp
ultimately show ?thesis using E and G
by (cases "filter Q (z # zs')", simp_all)
qed
qed
qed
qed
lemma interleaves_map [rule_format]:
assumes A: "inj f"
shows "xs \<simeq> {ys, zs, P} \<longrightarrow>
map f xs \<simeq> {map f ys, map f zs, \<lambda>w ws. P (inv f w) (map (inv f) ws)}"
(is "_ \<longrightarrow> _ \<simeq> {_, _, ?P'}")
proof (induction xs arbitrary: ys zs, rule_tac [!] impI, simp_all)
fix ys zs
assume "[] \<simeq> {ys, zs, P}"
hence "ys = [] \<and> zs = []" by (rule interleaves_nil)
thus "[] \<simeq> {map f ys, map f zs, ?P'}" by simp
next
fix x xs ys zs
assume
B: "\<And>ys zs. xs \<simeq> {ys, zs, P} \<longrightarrow> map f xs \<simeq> {map f ys, map f zs, ?P'}" and
C: "x # xs \<simeq> {ys, zs, P}"
show "f x # map f xs \<simeq> {map f ys, map f zs, ?P'}"
proof (cases ys, case_tac [!] zs, simp_all del: interleaves.simps(1))
assume "ys = []" and "zs = []"
thus False using C by simp
next
fix z zs'
assume "ys = []" and "zs = z # zs'"
hence "x = z \<and> xs \<simeq> {[], zs', P}" using C by simp
moreover have "xs \<simeq> {[], zs', P} \<longrightarrow> map f xs \<simeq> {map f [], map f zs', ?P'}"
using B .
ultimately show "f x = f z \<and> map f xs \<simeq> {[], map f zs', ?P'}" by simp
next
fix y ys'
assume "ys = y # ys'" and "zs = []"
hence "x = y \<and> xs \<simeq> {ys', [], P}" using C by simp
moreover have "xs \<simeq> {ys', [], P} \<longrightarrow> map f xs \<simeq> {map f ys', map f [], ?P'}"
using B .
ultimately show "f x = f y \<and> map f xs \<simeq> {map f ys', [], ?P'}" by simp
next
fix y ys' z zs'
assume "ys = y # ys'" and "zs = z # zs'"
hence D: "x # xs \<simeq> {y # ys', z # zs', P}" using C by simp
show "f x # map f xs \<simeq> {f y # map f ys', f z # map f zs', ?P'}"
proof (cases "P x xs")
case True
hence E: "?P' (f x) (map f xs)" using A by simp
have "x = y \<and> xs \<simeq> {ys', z # zs', P}" using D and True by simp
moreover have "xs \<simeq> {ys', z # zs', P} \<longrightarrow>
map f xs \<simeq> {map f ys', map f (z # zs'), ?P'}"
using B .
ultimately have "f x = f y \<and> map f xs \<simeq> {map f ys', map f (z # zs'), ?P'}"
by simp
thus ?thesis using E by simp
next
case False
hence E: "\<not> ?P' (f x) (map f xs)" using A by simp
have "x = z \<and> xs \<simeq> {y # ys', zs', P}" using D and False by simp
moreover have "xs \<simeq> {y # ys', zs', P} \<longrightarrow>
map f xs \<simeq> {map f (y # ys'), map f zs', ?P'}"
using B .
ultimately have "f x = f z \<and> map f xs \<simeq> {map f (y # ys'), map f zs', ?P'}"
by simp
thus ?thesis using E by simp
qed
qed
qed
lemma interleaves_prefix_fst_1 [rule_format]:
assumes A: "xs \<simeq> {ys, zs, P}"
shows "(\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)) \<longrightarrow>
ws @ xs \<simeq> {ws @ ys, zs, P}"
proof (induction ws, simp_all add: A, rule impI)
fix w ws
assume B: "\<forall>n < Suc (length ws). P ((w # ws) ! n) (drop n ws @ xs)"
assume "(\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)) \<longrightarrow>
ws @ xs \<simeq> {ws @ ys, zs, P}"
moreover have "\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)"
proof (rule allI, rule impI)
fix n
assume "n < length ws"
moreover have "Suc n < Suc (length ws) \<longrightarrow>
P ((w # ws) ! (Suc n)) (drop (Suc n) ws @ xs)"
using B ..
ultimately show "P (ws ! n) (drop (Suc n) ws @ xs)" by simp
qed
ultimately have "ws @ xs \<simeq> {ws @ ys, zs, P}" ..
moreover have "0 < Suc (length ws) \<longrightarrow> P ((w # ws) ! 0) (drop 0 ws @ xs)"
using B ..
hence "P w (ws @ xs)" by simp
ultimately show "w # ws @ xs \<simeq> {w # ws @ ys, zs, P}" by (cases zs, simp_all)
qed
lemma interleaves_prefix_fst_2 [rule_format]:
"ws @ xs \<simeq> {ws @ ys, zs, P} \<longrightarrow>
(\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)) \<longrightarrow>
xs \<simeq> {ys, zs, P}"
proof (induction ws, simp_all, (rule impI)+)
fix w ws
assume A: "\<forall>n < Suc (length ws). P ((w # ws) ! n) (drop n ws @ xs)"
hence "0 < Suc (length ws) \<longrightarrow> P ((w # ws) ! 0) (drop 0 ws @ xs)" ..
hence "P w (ws @ xs)" by simp
moreover assume "w # ws @ xs \<simeq> {w # ws @ ys, zs, P}"
ultimately have "ws @ xs \<simeq> {ws @ ys, zs, P}" by (cases zs, simp_all)
moreover assume "ws @ xs \<simeq> {ws @ ys, zs, P} \<longrightarrow>
(\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)) \<longrightarrow>
xs \<simeq> {ys, zs, P}"
ultimately have "(\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)) \<longrightarrow>
xs \<simeq> {ys, zs, P}"
by simp
moreover have "\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)"
proof (rule allI, rule impI)
fix n
assume "n < length ws"
moreover have "Suc n < Suc (length ws) \<longrightarrow>
P ((w # ws) ! (Suc n)) (drop (Suc n) ws @ xs)"
using A ..
ultimately show "P (ws ! n) (drop (Suc n) ws @ xs)" by simp
qed
ultimately show "xs \<simeq> {ys, zs, P}" ..
qed
lemma interleaves_prefix_fst [rule_format]:
"\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs) \<Longrightarrow>
xs \<simeq> {ys, zs, P} = ws @ xs \<simeq> {ws @ ys, zs, P}"
proof (rule iffI, erule interleaves_prefix_fst_1, simp)
qed (erule interleaves_prefix_fst_2, simp)
lemma interleaves_prefix_snd [rule_format]:
"\<forall>n < length ws. \<not> P (ws ! n) (drop (Suc n) ws @ xs) \<Longrightarrow>
xs \<simeq> {ys, zs, P} = ws @ xs \<simeq> {ys, ws @ zs, P}"
proof (subst (1 2) interleaves_swap)
qed (rule interleaves_prefix_fst, simp)
subsection "A second, stronger version of interleaving"
text \<open>
Simple counterexamples show that unlike prefixes, the addition or removal of suffixes to the input
lists does not generally preserve the validity of predicate @{term interleaves}. In fact, if
@{term "P y [x] = True"} with @{term "x \<noteq> y"}, then @{term "[y, x] \<simeq> {[x], [y], P}"} does not hold
although @{term "[y] \<simeq> {[], [y], \<lambda>w ws. P w (ws @ [x])}"} does, by virtue of lemma
@{thm interleaves_nil_all}. Similarly, @{term "[x, y] \<simeq> {[], [y, x], \<lambda>w ws. P w (ws @ [x])}"} does
not hold for @{term "x \<noteq> y"} even though @{term "[x, y, x] \<simeq> {[x], [y, x], P}"} does.
Both counterexamples would not work any longer if the truth value of the input predicate were
significant even if either the second or the third list is empty. In fact, in the former case,
condition @{term "P y [x] = True"} would entail the falseness of statement
@{term "[y] \<simeq> {[], [y], \<lambda>w ws. P w (ws @ [x])}"}, so that the validity of rule
@{term "[y] \<simeq> {[], [y], \<lambda>w ws. P w (ws @ [x])} \<Longrightarrow> [y, x] \<simeq> {[x], [y], P}"} would be preserved. In
the latter case, statement @{term "[x, y, x] \<simeq> {[x], [y, x], P}"} may only hold provided the last
item \<open>x\<close> of the first list is extracted from the third one, which would require that
@{term "\<not> P x []"}; thus, subordinating rule
@{term "[x, y, x] \<simeq> {[x], [y, x], P} \<Longrightarrow> [x, y] \<simeq> {[], [y, x], \<lambda>w ws. P w (ws @ [x])}"} to
condition @{term "P x []"} would preserve its validity.
This argument suggests that in order to obtain an \emph{interleaves} predicate whose validity is
also preserved upon the addition or removal of a suffix to the input lists, the truth value of the
input predicate must matter until both the second and the third list are empty. In what follows,
such a stronger version of the predicate, named \<open>Interleaves\<close>, is defined along with a
convenient symbolic notation for it.
\null
\<close>
fun Interleaves ::
"('a \<Rightarrow> 'a list \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"Interleaves P (x # xs) (y # ys) (z # zs) = (if P x xs
then x = y \<and> Interleaves P xs ys (z # zs)
else x = z \<and> Interleaves P xs (y # ys) zs)" |
"Interleaves P (x # xs) (y # ys) [] =
(P x xs \<and> x = y \<and> Interleaves P xs ys [])" |
"Interleaves P (x # xs) [] (z # zs) =
(\<not> P x xs \<and> x = z \<and> Interleaves P xs [] zs)" |
"Interleaves _ (_ # _) [] [] = False" |
"Interleaves _ [] (_ # _) _ = False" |
"Interleaves _ [] _ (_ # _) = False" |
"Interleaves _ [] [] [] = True"
abbreviation Interleaves_syntax ::
"'a list \<Rightarrow> 'a list \<Rightarrow> 'a list \<Rightarrow> ('a \<Rightarrow> 'a list \<Rightarrow> bool) \<Rightarrow> bool"
("(_ \<cong> {_, _, _})" [60, 60, 60] 51)
where "xs \<cong> {ys, zs, P} \<equiv> Interleaves P xs ys zs"
text \<open>
\null
In what follows, it is proven that predicate @{term Interleaves} is actually not weaker than, viz.
is a sufficient condition for, predicate @{term interleaves}. Moreover, the former predicate is
shown to fulfil the same rules as the latter, although sometimes under more stringent assumptions
(cf. lemmas \<open>Interleaves_all_nil\<close>, \<open>Interleaves_nil_all\<close> with lemmas
@{thm interleaves_all_nil}, @{thm interleaves_nil_all}), and to have the further property that under
proper assumptions, its validity is preserved upon the addition or removal of a suffix to the input
lists.
\null
\<close>
lemma Interleaves_interleaves [rule_format]:
"xs \<cong> {ys, zs, P} \<longrightarrow> xs \<simeq> {ys, zs, P}"
proof (induction P xs ys zs rule: interleaves.induct, simp_all)
qed (rule conjI, (rule_tac [!] impI)+, simp_all)
lemma Interleaves_length:
"xs \<cong> {ys, zs, P} \<Longrightarrow> length xs = length ys + length zs"
by (drule Interleaves_interleaves, rule interleaves_length)
lemma Interleaves_nil:
"[] \<cong> {ys, zs, P} \<Longrightarrow> ys = [] \<and> zs = []"
by (drule Interleaves_interleaves, rule interleaves_nil)
lemma Interleaves_swap:
"xs \<cong> {ys, zs, P} = xs \<cong> {zs, ys, \<lambda>w ws. \<not> P w ws}"
proof (induction P xs ys zs rule: Interleaves.induct, simp_all)
fix y' :: 'a and ys' zs' P'
show "\<not> [] \<cong> {zs', y' # ys', \<lambda>w ws. \<not> P' w ws}" by (cases zs', simp_all)
qed
lemma Interleaves_equal_fst:
"xs \<cong> {ys, zs, P} \<Longrightarrow> xs \<cong> {ys', zs, P} \<Longrightarrow> ys = ys'"
by ((drule Interleaves_interleaves)+, rule interleaves_equal_fst)
lemma Interleaves_equal_snd:
"xs \<cong> {ys, zs, P} \<Longrightarrow> xs \<cong> {ys, zs', P} \<Longrightarrow> zs = zs'"
by ((drule Interleaves_interleaves)+, rule interleaves_equal_snd)
lemma Interleaves_equal_all_nil:
"xs \<cong> {ys, [], P} \<Longrightarrow> xs = ys"
by (drule Interleaves_interleaves, rule interleaves_equal_all_nil)
lemma Interleaves_equal_nil_all:
"xs \<cong> {[], zs, P} \<Longrightarrow> xs = zs"
by (drule Interleaves_interleaves, rule interleaves_equal_nil_all)
lemma Interleaves_filter [rule_format]:
assumes A: "\<forall>x xs. P x (filter Q xs) = P x xs"
shows "xs \<cong> {ys, zs, P} \<longrightarrow> filter Q xs \<cong> {filter Q ys, filter Q zs, P}"
proof (induction xs arbitrary: ys zs, rule_tac [!] impI, simp)
fix ys zs
assume "[] \<cong> {ys, zs, P}"
hence "ys = [] \<and> zs = []" by (rule Interleaves_nil)
thus "[] \<cong> {filter Q ys, filter Q zs, P}" by simp
next
fix x xs ys zs
assume
B: "\<And>ys' zs'. xs \<cong> {ys', zs', P} \<longrightarrow>
filter Q xs \<cong> {filter Q ys', filter Q zs', P}" and
C: "x # xs \<cong> {ys, zs, P}"
show "filter Q (x # xs) \<cong> {filter Q ys, filter Q zs, P}"
proof (cases ys, case_tac [!] zs, simp_all del: filter.simps, rule ccontr)
assume "ys = []" and "zs = []"
thus False using C by simp
next
fix z zs'
assume "ys = []" and "zs = z # zs'"
hence D: "\<not> P x xs \<and> x = z \<and> xs \<cong> {[], zs', P}" using C by simp+
moreover have "xs \<cong> {[], zs', P} \<longrightarrow>
filter Q xs \<cong> {filter Q [], filter Q zs', P}"
using B .
ultimately have "filter Q xs \<cong> {[], filter Q zs', P}" by simp
moreover have "\<not> P x (filter Q xs)" using A and D by simp+
ultimately show "filter Q (x # xs) \<cong> {[], filter Q (z # zs'), P}"
using D by simp
next
fix y ys'
assume "ys = y # ys'" and "zs = []"
hence D: "P x xs \<and> x = y \<and> xs \<cong> {ys', [], P}" using C by simp+
moreover have "xs \<cong> {ys', [], P} \<longrightarrow>
filter Q xs \<cong> {filter Q ys', filter Q [], P}"
using B .
ultimately have "filter Q xs \<cong> {filter Q ys', [], P}" by simp
moreover have "P x (filter Q xs)" using A and D by simp+
ultimately show "filter Q (x # xs) \<cong> {filter Q (y # ys'), [], P}"
using D by simp
next
fix y ys' z zs'
assume "ys = y # ys'" and "zs = z # zs'"
hence D: "x # xs \<cong> {y # ys', z # zs', P}" using C by simp
show "filter Q (x # xs) \<cong> {filter Q (y # ys'), filter Q (z # zs'), P}"
proof (cases "P x xs")
case True
hence E: "P x (filter Q xs)" using A by simp
have F: "x = y \<and> xs \<cong> {ys', z # zs', P}" using D and True by simp
moreover have "xs \<cong> {ys', z # zs', P} \<longrightarrow>
filter Q xs \<cong> {filter Q ys', filter Q (z # zs'), P}"
using B .
ultimately have G: "filter Q xs \<cong> {filter Q ys', filter Q (z # zs'), P}"
by simp
show ?thesis
proof (cases "Q x")
assume "Q x"
hence "filter Q (x # xs) = x # filter Q xs" by simp
moreover have "filter Q (y # ys') = x # filter Q ys'"
using \<open>Q x\<close> and F by simp
ultimately show ?thesis using E and G
by (cases "filter Q (z # zs')", simp_all)
next
assume "\<not> Q x"
hence "filter Q (x # xs) = filter Q xs" by simp
moreover have "filter Q (y # ys') = filter Q ys'"
using \<open>\<not> Q x\<close> and F by simp
ultimately show ?thesis using E and G
by (cases "filter Q (z # zs')", simp_all)
qed
next
case False
hence E: "\<not> P x (filter Q xs)" using A by simp
have F: "x = z \<and> xs \<cong> {y # ys', zs', P}" using D and False by simp
moreover have "xs \<cong> {y # ys', zs', P} \<longrightarrow>
filter Q xs \<cong> {filter Q (y # ys'), filter Q zs', P}"
using B .
ultimately have G: "filter Q xs \<cong> {filter Q (y # ys'), filter Q zs', P}"
by simp
show ?thesis
proof (cases "Q x")
assume "Q x"
hence "filter Q (x # xs) = x # filter Q xs" by simp
moreover have "filter Q (z # zs') = x # filter Q zs'"
using \<open>Q x\<close> and F by simp
ultimately show ?thesis using E and G
by (cases "filter Q (y # ys')", simp_all)
next
assume "\<not> Q x"
hence "filter Q (x # xs) = filter Q xs" by simp
moreover have "filter Q (z # zs') = filter Q zs'"
using \<open>\<not> Q x\<close> and F by simp
ultimately show ?thesis using E and G
by (cases "filter Q (z # zs')", simp_all)
qed
qed
qed
qed
lemma Interleaves_map [rule_format]:
assumes A: "inj f"
shows "xs \<cong> {ys, zs, P} \<longrightarrow>
map f xs \<cong> {map f ys, map f zs, \<lambda>w ws. P (inv f w) (map (inv f) ws)}"
(is "_ \<longrightarrow> _ \<cong> {_, _, ?P'}")
proof (induction xs arbitrary: ys zs, rule_tac [!] impI, simp_all)
fix ys zs
assume "[] \<cong> {ys, zs, P}"
hence "ys = [] \<and> zs = []" by (rule Interleaves_nil)
thus "[] \<cong> {map f ys, map f zs, ?P'}" by simp
next
fix x xs ys zs
assume
B: "\<And>ys zs. xs \<cong> {ys, zs, P} \<longrightarrow> map f xs \<cong> {map f ys, map f zs, ?P'}" and
C: "x # xs \<cong> {ys, zs, P}"
show "f x # map f xs \<cong> {map f ys, map f zs, ?P'}"
proof (cases ys, case_tac [!] zs, simp_all del: Interleaves.simps(1-3))
assume "ys = []" and "zs = []"
thus False using C by simp
next
fix z zs'
assume "ys = []" and "zs = z # zs'"
hence D: "\<not> P x xs \<and> x = z \<and> xs \<cong> {[], zs', P}" using C by simp+
moreover have "xs \<cong> {[], zs', P} \<longrightarrow> map f xs \<cong> {map f [], map f zs', ?P'}"
using B .
ultimately have "map f xs \<cong> {[], map f zs', ?P'}" by simp
moreover have "\<not> ?P' (f x) (map f xs)" using A and D by simp+
ultimately show "f x # map f xs \<cong> {[], f z # map f zs', ?P'}"
using D by simp
next
fix y ys'
assume "ys = y # ys'" and "zs = []"
hence D: "P x xs \<and> x = y \<and> xs \<cong> {ys', [], P}" using C by simp+
moreover have "xs \<cong> {ys', [], P} \<longrightarrow> map f xs \<cong> {map f ys', map f [], ?P'}"
using B .
ultimately have "map f xs \<cong> {map f ys', [], ?P'}" by simp
moreover have "?P' (f x) (map f xs)" using A and D by simp+
ultimately show "f x # map f xs \<cong> {f y # map f ys', [], ?P'}"
using D by simp
next
fix y ys' z zs'
assume "ys = y # ys'" and "zs = z # zs'"
hence D: "x # xs \<cong> {y # ys', z # zs', P}" using C by simp
show "f x # map f xs \<cong> {f y # map f ys', f z # map f zs', ?P'}"
proof (cases "P x xs")
case True
hence E: "?P' (f x) (map f xs)" using A by simp
have "x = y \<and> xs \<cong> {ys', z # zs', P}" using D and True by simp
moreover have "xs \<cong> {ys', z # zs', P} \<longrightarrow>
map f xs \<cong> {map f ys', map f (z # zs'), ?P'}"
using B .
ultimately have "f x = f y \<and> map f xs \<cong> {map f ys', map f (z # zs'), ?P'}"
by simp
thus ?thesis using E by simp
next
case False
hence E: "\<not> ?P' (f x) (map f xs)" using A by simp
have "x = z \<and> xs \<cong> {y # ys', zs', P}" using D and False by simp
moreover have "xs \<cong> {y # ys', zs', P} \<longrightarrow>
map f xs \<cong> {map f (y # ys'), map f zs', ?P'}"
using B .
ultimately have "f x = f z \<and> map f xs \<cong> {map f (y # ys'), map f zs', ?P'}"
by simp
thus ?thesis using E by simp
qed
qed
qed
lemma Interleaves_prefix_fst_1 [rule_format]:
assumes A: "xs \<cong> {ys, zs, P}"
shows "(\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)) \<longrightarrow>
ws @ xs \<cong> {ws @ ys, zs, P}"
proof (induction ws, simp_all add: A, rule impI)
fix w ws
assume B: "\<forall>n < Suc (length ws). P ((w # ws) ! n) (drop n ws @ xs)"
assume "(\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)) \<longrightarrow>
ws @ xs \<cong> {ws @ ys, zs, P}"
moreover have "\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)"
proof (rule allI, rule impI)
fix n
assume "n < length ws"
moreover have "Suc n < Suc (length ws) \<longrightarrow>
P ((w # ws) ! (Suc n)) (drop (Suc n) ws @ xs)"
using B ..
ultimately show "P (ws ! n) (drop (Suc n) ws @ xs)" by simp
qed
ultimately have "ws @ xs \<cong> {ws @ ys, zs, P}" ..
moreover have "0 < Suc (length ws) \<longrightarrow> P ((w # ws) ! 0) (drop 0 ws @ xs)"
using B ..
hence "P w (ws @ xs)" by simp
ultimately show "w # ws @ xs \<cong> {w # ws @ ys, zs, P}" by (cases zs, simp_all)
qed
lemma Interleaves_prefix_fst_2 [rule_format]:
"ws @ xs \<cong> {ws @ ys, zs, P} \<longrightarrow>
(\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)) \<longrightarrow>
xs \<cong> {ys, zs, P}"
proof (induction ws, simp_all, (rule impI)+)
fix w ws
assume A: "\<forall>n < Suc (length ws). P ((w # ws) ! n) (drop n ws @ xs)"
hence "0 < Suc (length ws) \<longrightarrow> P ((w # ws) ! 0) (drop 0 ws @ xs)" ..
hence "P w (ws @ xs)" by simp
moreover assume "w # ws @ xs \<cong> {w # ws @ ys, zs, P}"
ultimately have "ws @ xs \<cong> {ws @ ys, zs, P}" by (cases zs, simp_all)
moreover assume "ws @ xs \<cong> {ws @ ys, zs, P} \<longrightarrow>
(\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)) \<longrightarrow>
xs \<cong> {ys, zs, P}"
ultimately have "(\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)) \<longrightarrow>
xs \<cong> {ys, zs, P}"
by simp
moreover have "\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs)"
proof (rule allI, rule impI)
fix n
assume "n < length ws"
moreover have "Suc n < Suc (length ws) \<longrightarrow>
P ((w # ws) ! (Suc n)) (drop (Suc n) ws @ xs)"
using A ..
ultimately show "P (ws ! n) (drop (Suc n) ws @ xs)" by simp
qed
ultimately show "xs \<cong> {ys, zs, P}" ..
qed
lemma Interleaves_prefix_fst [rule_format]:
"\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws @ xs) \<Longrightarrow>
xs \<cong> {ys, zs, P} = ws @ xs \<cong> {ws @ ys, zs, P}"
proof (rule iffI, erule Interleaves_prefix_fst_1, simp)
qed (erule Interleaves_prefix_fst_2, simp)
lemma Interleaves_prefix_snd [rule_format]:
"\<forall>n < length ws. \<not> P (ws ! n) (drop (Suc n) ws @ xs) \<Longrightarrow>
xs \<cong> {ys, zs, P} = ws @ xs \<cong> {ys, ws @ zs, P}"
proof (subst (1 2) Interleaves_swap)
qed (rule Interleaves_prefix_fst, simp)
lemma Interleaves_all_nil_1 [rule_format]:
"xs \<cong> {xs, [], P} \<longrightarrow> (\<forall>n < length xs. P (xs ! n) (drop (Suc n) xs))"
proof (induction xs, simp_all, rule impI, erule conjE, rule allI, rule impI)
fix x xs n
assume
"xs \<cong> {xs, [], P} \<longrightarrow> (\<forall>n < length xs. P (xs ! n) (drop (Suc n) xs))" and
"xs \<cong> {xs, [], P}"
hence A: "\<forall>n < length xs. P (xs ! n) (drop (Suc n) xs)" ..
assume
B: "P x xs" and
C: "n < Suc (length xs)"
show "P ((x # xs) ! n) (drop n xs)"
proof (cases n, simp_all add: B)
case (Suc m)
have "m < length xs \<longrightarrow> P (xs ! m) (drop (Suc m) xs)" using A ..
moreover have "m < length xs" using C and Suc by simp
ultimately show "P (xs ! m) (drop (Suc m) xs)" ..
qed
qed
lemma Interleaves_all_nil_2 [rule_format]:
"\<forall>n < length xs. P (xs ! n) (drop (Suc n) xs) \<Longrightarrow> xs \<cong> {xs, [], P}"
by (insert Interleaves_prefix_fst [of xs P "[]" "[]" "[]"], simp)
lemma Interleaves_all_nil:
"xs \<cong> {xs, [], P} = (\<forall>n < length xs. P (xs ! n) (drop (Suc n) xs))"
proof (rule iffI, rule allI, rule impI, rule Interleaves_all_nil_1, assumption+)
qed (rule Interleaves_all_nil_2, simp)
lemma Interleaves_nil_all:
"xs \<cong> {[], xs, P} = (\<forall>n < length xs. \<not> P (xs ! n) (drop (Suc n) xs))"
by (subst Interleaves_swap, simp add: Interleaves_all_nil)
lemma Interleaves_suffix_one_aux:
assumes A: "P x []"
shows "\<not> xs @ [x] \<cong> {[], zs, P}"
+using [[simproc del: defined_all]]
proof (induction xs arbitrary: zs, simp_all, rule_tac [!] notI)
fix zs
assume "[x] \<cong> {[], zs, P}"
thus False by (cases zs, simp_all add: A)
next
fix w xs zs
assume B: "\<And>zs. \<not> xs @ [x] \<cong> {[], zs, P}"
assume "w # xs @ [x] \<cong> {[], zs, P}"
thus False proof (cases zs, simp_all, (erule_tac conjE)+)
fix zs'
assume "xs @ [x] \<cong> {[], zs', P}"
moreover have "\<not> xs @ [x] \<cong> {[], zs', P}" using B .
ultimately show False by contradiction
qed
qed
lemma Interleaves_suffix_one_fst_2 [rule_format]:
assumes A: "P x []"
shows "xs @ [x] \<cong> {ys @ [x], zs, P} \<longrightarrow> xs \<cong> {ys, zs, \<lambda>w ws. P w (ws @ [x])}"
(is "_ \<longrightarrow> _ \<cong> {_, _, ?P'}")
+using [[simproc del: defined_all]]
proof (induction xs arbitrary: ys zs, rule_tac [!] impI, simp_all)
fix ys zs
assume "[x] \<cong> {ys @ [x], zs, P}"
hence B: "length [x] = length (ys @ [x]) + length zs"
by (rule Interleaves_length)
have ys: "ys = []" by (cases ys, simp, insert B, simp)
then have "zs = []" by (cases zs, simp, insert B, simp)
with ys show "[] \<cong> {ys, zs, ?P'}" by simp
next
fix w xs ys zs
assume B: "\<And>ys zs. xs @ [x] \<cong> {ys @ [x], zs, P} \<longrightarrow> xs \<cong> {ys, zs, ?P'}"
assume "w # xs @ [x] \<cong> {ys @ [x], zs, P}"
thus "w # xs \<cong> {ys, zs, ?P'}"
proof (cases zs, case_tac [!] ys, simp_all del: Interleaves.simps(1,3),
(erule_tac [1-2] conjE)+)
assume "xs @ [x] \<cong> {[], [], P}"
thus False by (cases xs, simp_all)
next
fix ys'
have "xs @ [x] \<cong> {ys' @ [x], [], P} \<longrightarrow> xs \<cong> {ys', [], ?P'}" using B .
moreover assume "xs @ [x] \<cong> {ys' @ [x], [], P}"
ultimately show "xs \<cong> {ys', [], ?P'}" ..
next
fix z' zs'
assume "w # xs @ [x] \<cong> {[x], z' # zs', P}"
thus "w # xs \<cong> {[], z' # zs', ?P'}"
proof (cases "P w (xs @ [x])", simp_all, erule_tac [!] conjE)
assume "xs @ [x] \<cong> {[], z' # zs', P}"
moreover have "\<not> xs @ [x] \<cong> {[], z' # zs', P}"
using A by (rule Interleaves_suffix_one_aux)
ultimately show False by contradiction
next
have "xs @ [x] \<cong> {[x], zs', P} \<longrightarrow> xs \<cong> {[], zs', ?P'}" using B by simp
moreover assume "xs @ [x] \<cong> {[x], zs', P}"
ultimately show "xs \<cong> {[], zs', ?P'}" ..
qed
next
fix y' ys' z' zs'
assume "w # xs @ [x] \<cong> {y' # ys' @ [x], z' # zs', P}"
thus "w # xs \<cong> {y' # ys', z' # zs', ?P'}"
proof (cases "P w (xs @ [x])", simp_all, erule_tac [!] conjE)
have "xs @ [x] \<cong> {ys' @ [x], z' # zs', P} \<longrightarrow> xs \<cong> {ys', z' # zs', ?P'}"
using B .
moreover assume "xs @ [x] \<cong> {ys' @ [x], z' # zs', P}"
ultimately show "xs \<cong> {ys', z' # zs', ?P'}" ..
next
have "xs @ [x] \<cong> {y' # ys' @ [x], zs', P} \<longrightarrow> xs \<cong> {y' # ys', zs', ?P'}"
using B by simp
moreover assume "xs @ [x] \<cong> {y' # ys' @ [x], zs', P}"
ultimately show "xs \<cong> {y' # ys', zs', ?P'}" ..
qed
qed
qed
lemma Interleaves_suffix_fst_1 [rule_format]:
assumes A: "\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws)"
shows "xs \<cong> {ys, zs, \<lambda>v vs. P v (vs @ ws)} \<longrightarrow> xs @ ws \<cong> {ys @ ws, zs, P}"
(is "_ \<cong> {_, _, ?P'} \<longrightarrow> _")
+using [[simproc del: defined_all]]
proof (induction xs arbitrary: ys zs, rule_tac [!] impI, simp_all)
fix ys zs
assume "[] \<cong> {ys, zs, ?P'}"
hence "ys = [] \<and> zs = []" by (rule Interleaves_nil)
thus "ws \<cong> {ys @ ws, zs, P}" using A by (simp add: Interleaves_all_nil)
next
fix x xs ys zs
assume A: "\<And>ys zs. xs \<cong> {ys, zs, ?P'} \<longrightarrow> xs @ ws \<cong> {ys @ ws, zs, P}"
assume "x # xs \<cong> {ys, zs, ?P'}"
thus "x # xs @ ws \<cong> {ys @ ws, zs, P}"
proof (rule_tac Interleaves.cases [of "(?P', x # xs, ys, zs)"],
simp_all del: Interleaves.simps(1),
(erule_tac conjE)+, (erule_tac [2] conjE)+, (erule_tac [3] conjE)+)
fix P' x' xs' y' ys' z' zs'
assume
B: "x' # xs' \<cong> {y' # ys', z' # zs', P'}" and
C: "?P' = P'" and
D: "xs = xs'"
show "x' # xs' @ ws \<cong> {y' # ys' @ ws, z' # zs', P}"
proof (cases "P' x' xs'")
have "xs \<cong> {ys', z' # zs', ?P'} \<longrightarrow> xs @ ws \<cong> {ys' @ ws, z' # zs', P}"
using A .
moreover case True
hence "xs \<cong> {ys', z' # zs', ?P'}" using B and C and D by simp
ultimately have "xs @ ws \<cong> {ys' @ ws, z' # zs', P}" ..
moreover have "P x' (xs' @ ws)" using C [symmetric] and True by simp
moreover have "x' = y'" using B and True by simp
ultimately show ?thesis using D by simp
next
have "xs \<cong> {y' # ys', zs', ?P'} \<longrightarrow> xs @ ws \<cong> {(y' # ys') @ ws, zs', P}"
using A .
moreover case False
hence "xs \<cong> {y' # ys', zs', ?P'}" using B and C and D by simp
ultimately have "xs @ ws \<cong> {(y' # ys') @ ws, zs', P}" ..
moreover have "\<not> P x' (xs' @ ws)" using C [symmetric] and False by simp
moreover have "x' = z'" using B and False by simp
ultimately show ?thesis using D by simp
qed
next
fix P' x' xs' y' ys'
have "xs \<cong> {ys', [], ?P'} \<longrightarrow> xs @ ws \<cong> {ys' @ ws, [], P}" using A .
moreover assume
"xs' \<cong> {ys', [], P'}" and
B: "?P' = P'" and
C: "xs = xs'"
hence "xs \<cong> {ys', [], ?P'}" by simp
ultimately have "xs' @ ws \<cong> {ys' @ ws, [], P}" using C by simp
moreover assume
"P' x' xs'" and
"x' = y'"
hence "P y' (xs' @ ws)" using B [symmetric] by simp
ultimately show "P y' (xs' @ ws) \<and> xs' @ ws \<cong> {ys' @ ws, [], P}" by simp
next
fix P' x' xs' z' zs'
have "xs \<cong> {[], zs', ?P'} \<longrightarrow> xs @ ws \<cong> {[] @ ws, zs', P}" using A .
moreover assume
"xs' \<cong> {[], zs', P'}" and
B: "?P' = P'" and
C: "xs = xs'"
hence "xs \<cong> {[], zs', ?P'}" by simp
ultimately have "xs' @ ws \<cong> {ws, zs', P}" using C by simp
moreover assume
"\<not> P' x' xs'" and
"x' = z'"
hence "\<not> P z' (xs' @ ws)" using B [symmetric] by simp
ultimately show "z' # xs' @ ws \<cong> {ws, z' # zs', P}" by (cases ws, simp_all)
qed
qed
lemma Interleaves_suffix_one_fst_1 [rule_format]:
"P x [] \<Longrightarrow>
xs \<cong> {ys, zs, \<lambda>w ws. P w (ws @ [x])} \<Longrightarrow> xs @ [x] \<cong> {ys @ [x], zs, P}"
by (rule Interleaves_suffix_fst_1, simp)
lemma Interleaves_suffix_one_fst:
"P x [] \<Longrightarrow>
xs \<cong> {ys, zs, \<lambda>w ws. P w (ws @ [x])} = xs @ [x] \<cong> {ys @ [x], zs, P}"
proof (rule iffI, rule Interleaves_suffix_one_fst_1, assumption+)
qed (rule Interleaves_suffix_one_fst_2)
lemma Interleaves_suffix_one_snd:
"\<not> P x [] \<Longrightarrow>
xs \<cong> {ys, zs, \<lambda>w ws. P w (ws @ [x])} = xs @ [x] \<cong> {ys, zs @ [x], P}"
by (subst (1 2) Interleaves_swap, rule Interleaves_suffix_one_fst)
lemma Interleaves_suffix_aux [rule_format]:
"(\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws)) \<longrightarrow>
x # xs @ ws \<cong> {ws, zs, P} \<longrightarrow>
\<not> P x (xs @ ws)"
proof (induction ws arbitrary: P rule: rev_induct, simp_all,
rule impI, (rule_tac [2] impI)+)
fix P
assume "x # xs \<cong> {[], zs, P}"
thus "\<not> P x xs" by (cases zs, simp_all)
next
fix w ws P
assume
A: "\<And>P'. (\<forall>n < length ws. P' (ws ! n) (drop (Suc n) ws)) \<longrightarrow>
x # xs @ ws \<cong> {ws, zs, P'} \<longrightarrow> \<not> P' x (xs @ ws)" and
B: "\<forall>n < Suc (length ws). P ((ws @ [w]) ! n)
(drop (Suc n) ws @ drop (Suc n - length ws) [w])"
assume "x # xs @ ws @ [w] \<cong> {ws @ [w], zs, P}"
hence C: "(x # xs @ ws) @ [w] \<cong> {ws @ [w], zs, P}" by simp
let ?P' = "\<lambda>v vs. P v (vs @ [w])"
have "(\<forall>n < length ws. ?P' (ws ! n) (drop (Suc n) ws)) \<longrightarrow>
x # xs @ ws \<cong> {ws, zs, ?P'} \<longrightarrow> \<not> ?P' x (xs @ ws)"
using A .
moreover have "\<forall>n < length ws. ?P' (ws ! n) (drop (Suc n) ws)"
proof (rule allI, rule impI)
fix n
assume D: "n < length ws"
moreover have "n < Suc (length ws) \<longrightarrow> P ((ws @ [w]) ! n)
(drop (Suc n) ws @ drop (Suc n - length ws) [w])"
using B ..
ultimately have "P ((ws @ [w]) ! n) (drop (Suc n) ws @ [w])" by simp
moreover have "n < length (butlast (ws @ [w]))" using D by simp
hence "butlast (ws @ [w]) ! n = (ws @ [w]) ! n" by (rule nth_butlast)
ultimately show "P (ws ! n) (drop (Suc n) ws @ [w])" by simp
qed
ultimately have "x # xs @ ws \<cong> {ws, zs, ?P'} \<longrightarrow> \<not> ?P' x (xs @ ws)" ..
moreover have "length ws < Suc (length ws) \<longrightarrow> P ((ws @ [w]) ! length ws)
(drop (Suc (length ws)) ws @ drop (Suc (length ws) - length ws) [w])"
using B ..
hence "P w []" by simp
hence "x # xs @ ws \<cong> {ws, zs, ?P'}"
using C by (rule Interleaves_suffix_one_fst_2)
ultimately have "\<not> ?P' x (xs @ ws)" ..
thus "\<not> P x (xs @ ws @ [w])" by simp
qed
lemma Interleaves_suffix_fst_2 [rule_format]:
assumes A: "\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws)"
shows "xs @ ws \<cong> {ys @ ws, zs, P} \<longrightarrow> xs \<cong> {ys, zs, \<lambda>v vs. P v (vs @ ws)}"
(is "_ \<longrightarrow> _ \<cong> {_, _, ?P'}")
+using [[simproc del: defined_all]]
proof (induction xs arbitrary: ys zs, rule_tac [!] impI, simp_all)
fix ys zs
assume "ws \<cong> {ys @ ws, zs, P}"
hence B: "length ws = length (ys @ ws) + length zs"
by (rule Interleaves_length)
have ys: "ys = []" by (cases ys, simp, insert B, simp)
then have "zs = []" by (cases zs, simp, insert B, simp)
with ys show "[] \<cong> {ys, zs, ?P'}" by simp
next
fix x xs ys zs
assume B: "\<And>ys zs. xs @ ws \<cong> {ys @ ws, zs, P} \<longrightarrow> xs \<cong> {ys, zs, ?P'}"
assume "x # xs @ ws \<cong> {ys @ ws, zs, P}"
thus "x # xs \<cong> {ys, zs, ?P'}"
proof (cases zs, case_tac [!] ys, simp_all del: Interleaves.simps(1,3),
(erule_tac [2] conjE)+)
assume C: "x # xs @ ws \<cong> {ws, [], P}"
have "length (x # xs @ ws) = length ws + length []"
by (insert Interleaves_length [OF C], simp)
thus False by simp
next
fix ys'
have "xs @ ws \<cong> {ys' @ ws, [], P} \<longrightarrow> xs \<cong> {ys', [], ?P'}" using B .
moreover assume "xs @ ws \<cong> {ys' @ ws, [], P}"
ultimately show "xs \<cong> {ys', [], ?P'}" ..
next
fix z' zs'
assume "x # xs @ ws \<cong> {ws, z' # zs', P}"
thus "x # xs \<cong> {[], z' # zs', ?P'}"
proof (cases "P x (xs @ ws)", simp_all)
case True
moreover assume "x # xs @ ws \<cong> {ws, z' # zs', P}"
with A [rule_format] have "\<not> P x (xs @ ws)"
by (rule Interleaves_suffix_aux)
ultimately show False by contradiction
next
case False
moreover assume "x # xs @ ws \<cong> {ws, z' # zs', P}"
ultimately have "x = z' \<and> xs @ ws \<cong> {ws, zs', P}" by (cases ws, simp_all)
moreover have "xs @ ws \<cong> {[] @ ws, zs', P} \<longrightarrow> xs \<cong> {[], zs', ?P'}"
using B .
ultimately show "x = z' \<and> xs \<cong> {[], zs', ?P'}" by simp
qed
next
fix y' ys' z' zs'
assume "x # xs @ ws \<cong> {y' # ys' @ ws, z' # zs', P}"
thus "x # xs \<cong> {y' # ys', z' # zs', ?P'}"
proof (cases "P x (xs @ ws)", simp_all, erule_tac [!] conjE)
have "xs @ ws \<cong> {ys' @ ws, z' # zs', P} \<longrightarrow> xs \<cong> {ys', z' # zs', ?P'}"
using B .
moreover assume "xs @ ws \<cong> {ys' @ ws, z' # zs', P}"
ultimately show "xs \<cong> {ys', z' # zs', ?P'}" ..
next
have "xs @ ws \<cong> {y' # ys' @ ws, zs', P} \<longrightarrow> xs \<cong> {y' # ys', zs', ?P'}"
using B by simp
moreover assume "xs @ ws \<cong> {y' # ys' @ ws, zs', P}"
ultimately show "xs \<cong> {y' # ys', zs', ?P'}" ..
qed
qed
qed
lemma Interleaves_suffix_fst [rule_format]:
"\<forall>n < length ws. P (ws ! n) (drop (Suc n) ws) \<Longrightarrow>
xs \<cong> {ys, zs, \<lambda>v vs. P v (vs @ ws)} = xs @ ws \<cong> {ys @ ws, zs, P}"
proof (rule iffI, rule Interleaves_suffix_fst_1, simp_all)
qed (rule Interleaves_suffix_fst_2, simp)
lemma Interleaves_suffix_snd [rule_format]:
"\<forall>n < length ws. \<not> P (ws ! n) (drop (Suc n) ws) \<Longrightarrow>
xs \<cong> {ys, zs, \<lambda>v vs. P v (vs @ ws)} = xs @ ws \<cong> {ys, zs @ ws, P}"
by (subst (1 2) Interleaves_swap, rule Interleaves_suffix_fst, simp)
end
diff --git a/thys/Markov_Models/ex/MDP_RP.thy b/thys/Markov_Models/ex/MDP_RP.thy
--- a/thys/Markov_Models/ex/MDP_RP.thy
+++ b/thys/Markov_Models/ex/MDP_RP.thy
@@ -1,895 +1,895 @@
section \<open>Value Iteration for Reachability Probabilities of MDPs\<close>
theory MDP_RP
imports "../Markov_Models"
begin
subsection \<open>Auxiliary Theorems\<close>
lemma INF_Union_eq: "(INF x\<in>\<Union>A. f x) = (INF a\<in>A. INF x\<in>a. f x)" for f :: "_ \<Rightarrow> 'a::complete_lattice"
by (auto intro!: antisym INF_greatest intro: INF_lower2)
lemma lift_option_eq_None: "lift_option f A B = None \<longleftrightarrow> (A \<noteq> None \<longrightarrow> B = None)"
by (cases A; cases B; auto)
lemma lift_option_eq_Some: "lift_option f A B = Some y \<longleftrightarrow> (\<exists>a b. A = Some a \<and> B = Some b \<and> y = f a b)"
by (cases A; cases B; auto)
lemma ord_option_Some1_iff: "ord_option R (Some a) y \<longleftrightarrow> (\<exists>b. y = Some b \<and> R a b)"
by (cases y; auto)
lemma ord_option_Some2_iff: "ord_option R x (Some b) \<longleftrightarrow> (\<forall>a. x = Some a \<longrightarrow> R a b)"
by (cases x; auto)
lemma sym_Restr: "sym A \<Longrightarrow> sym (Restr A S)"
by (auto simp: sym_def)
lemma trans_Restr: "trans A \<Longrightarrow> trans (Restr A S)"
by (auto simp: trans_def)
lemma image_eq_singleton_iff: "inj_on f S \<Longrightarrow> f ` S = {y} \<longleftrightarrow> (\<exists>x. S = {x} \<and> y = f x)"
by (auto elim: inj_img_insertE)
lemma quotient_eq_singleton: "equiv A r \<Longrightarrow> A // r = {B} \<Longrightarrow> B = A"
using Union_quotient[of A r] by auto
lemma UN_singleton_image: "(\<Union>x\<in>A. {f x}) = f ` A"
by auto
lemma image_eq_singeltonD: "f ` A = {x} \<Longrightarrow> \<forall>a\<in>A. f a = x"
by auto
lemma fun_ord_refl: "reflp ord \<Longrightarrow> reflp (fun_ord ord)"
by (auto simp: fun_ord_def reflp_def)
lemma fun_ord_trans: "transp ord \<Longrightarrow> transp (fun_ord ord)"
by (fastforce simp: fun_ord_def transp_def)
lemma fun_ord_antisym: "antisymp ord \<Longrightarrow> antisymp (fun_ord ord)"
by (fastforce simp: fun_ord_def antisymp_def)
lemma fun_ord_combine:
"fun_ord ord a b \<Longrightarrow> fun_ord ord c d \<Longrightarrow> (\<And>s. ord (a s) (b s) \<Longrightarrow> ord (c s) (d s) \<Longrightarrow> ord (e s) (f s)) \<Longrightarrow> fun_ord ord e f"
by (auto simp: fun_ord_def)
lemma not_all_eq: "~ (\<forall>y. x \<noteq> y)"
by auto
lemma ball_vimage_iff: "(\<forall>x\<in>f -` X. P x) \<longleftrightarrow> (\<forall>x. f x \<in> X \<longrightarrow> P x)"
by auto
lemma UN_If_cases: "(\<Union>x\<in>X. if P x then A x else B x) = (\<Union>x\<in>{x\<in>X. P x}. A x) \<union> (\<Union>x\<in>{x\<in>X. \<not> P x}. B x)"
by (auto split: if_splits)
lemma (in Reachability_Problem) n_eq_0_closed:
assumes s: "s \<in> S'" and S': "S' \<subseteq> S" "S' \<inter> S2 = {}" and closed: "\<And>s. s \<in> S' \<Longrightarrow> \<exists>D\<in>K s. D \<subseteq> S'"
shows "n s = 0"
proof -
from closed obtain ct where ct: "\<And>s. s \<in> S' \<Longrightarrow> ct s \<in> K s" "\<And>s. s \<in> S' \<Longrightarrow> ct s \<subseteq> S'"
by metis
define cfg where "cfg = memoryless_on (\<lambda>s. if s \<in> S' then ct s else arb_act s)"
have cfg_on: "cfg s \<in> cfg_on s" for s
unfolding cfg_def using ct by (intro memoryless_on_cfg_onI) auto
have state_cfg[simp]: "state (cfg s) = s" for s
unfolding cfg_def by (intro state_memoryless_on)
have action_cfg[simp]: "action (cfg s) = (if s \<in> S' then ct s else arb_act s)" for s
unfolding cfg_def by (intro action_memoryless_on)
have cont_cfg[simp]: "s \<in> S' \<Longrightarrow> t \<in> ct s \<Longrightarrow> cont (cfg s) t = cfg t" for s t
unfolding cfg_def by (intro cont_memoryless_on) auto
from s have "v (cfg s) = 0"
proof (coinduction arbitrary: s rule: v_eq_0_coinduct)
case (valid cfg') with cfg_on s S' show ?case
by (auto simp: valid_cfg_def)
next
case (nS2 cfg') with S' show ?case
by auto
next
case (cont cfg') with S' ct show ?case
by (force simp: set_K_cfg)
qed
show "n s = 0"
proof (rule n_eq_0)
show "s \<in> S" using s S' by auto
qed fact+
qed
lemma (in Reachability_Problem) n_lb_ennreal:
fixes x
assumes "s \<in> S"
assumes solution: "\<And>s D. s \<in> S1 \<Longrightarrow> D \<in> K s \<Longrightarrow> x s \<le> (\<Sum>t\<in>S. ennreal (pmf D t) * x t)"
assumes solution_n0: "\<And>s. s \<in> S \<Longrightarrow> n s = 0 \<Longrightarrow> x s = 0"
assumes solution_S2: "\<And>s. s \<in> S2 \<Longrightarrow> x s = 1"
and le_1: "\<And>s. s \<in> S \<Longrightarrow> x s \<le> 1"
shows "x s \<le> n s" (is "_ \<le> ?y s")
proof -
have x_less_top[simp]: "s \<in> S \<Longrightarrow> x s < top" for s
using le_1[of s] by (auto simp: less_top[symmetric] top_unique)
have "enn2real (x s) \<le> enn2real (n s)"
apply (rule n_lb[OF \<open>s\<in>S\<close>])
subgoal for s D
by (rule ennreal_le_iff[THEN iffD1])
(use S1 in \<open>auto intro!: sum_nonneg simp add: subset_eq solution sum_ennreal[symmetric] ennreal_mult simp del: sum_ennreal\<close>)
apply (auto simp: solution_n0 solution_S2)
done
with \<open>s\<in>S\<close> show ?thesis
by (subst (asm) ennreal_le_iff[symmetric]) (simp_all add: real_n)
qed
lifting_forget pmf_as_function.pmf.lifting
text \<open>
Type to describe MDP components. The support (i.e. elements which are not mapped to an empty
set) is the set of states of the component.
Most of this is from:
Formal verification of probabilistic systems
Luca de Alfaro (PhD thesis, 1997)
and
Reachability in MDPs: Refining Convergence of Value Iteration
Serge Haddad and Benjamin Monmege (2014)
\<close>
typedef 's mdpc = "UNIV :: ('s \<rightharpoonup> 's pmf set) set"
by auto
setup_lifting type_definition_mdpc
lift_definition states :: "'s mdpc \<Rightarrow> 's set"
is dom .
declare [[coercion states]]
lift_definition actions :: "'s mdpc \<Rightarrow> 's \<Rightarrow> 's pmf set"
is "\<lambda>f s. case f s of None \<Rightarrow> {} | Some a \<Rightarrow> a" .
lemma in_states: "actions \<phi> s \<noteq> {} \<Longrightarrow> s \<in> states \<phi>"
by transfer auto
lemma mdpc_eqI: "states \<phi> = states \<psi> \<Longrightarrow> (\<And>s. s \<in> states \<phi> \<Longrightarrow> actions \<phi> s = actions \<psi> s) \<Longrightarrow> \<phi> = \<psi>"
apply transfer
apply (rule ext)
subgoal premises prems for \<phi> \<psi> x
using prems(1) prems(2)[of x]
by (cases "x \<in> dom \<phi>") (auto simp: fun_eq_iff split: option.splits)
done
lift_definition map_mdpc :: "('s \<Rightarrow> 't) \<Rightarrow> 's mdpc \<Rightarrow> 't mdpc"
is "\<lambda>m f s. if f ` (m -` {s}) \<subseteq> {None} then None else Some {map_pmf m d | d A t. m t = s \<and> f t = Some A \<and> d \<in> A}" .
lemma states_map_mdpc: "states (map_mdpc f M) = f ` (states M)"
by (transfer fixing: f) (auto simp: subset_eq image_iff dom_def split: if_splits)
lemma actions_map_mdpc_eq_Collect: "actions (map_mdpc f M) s = {map_pmf f d | d t. f t = s \<and> d \<in> actions M t}"
by transfer (force simp: subset_eq split: option.splits)
lemma actions_map_mdpc: "actions (map_mdpc f M) s = map_pmf f ` (\<Union>t\<in>f -` {s}. actions M t)"
by (auto simp: actions_map_mdpc_eq_Collect)
lemma map_mdpc_compose: "map_mdpc f (map_mdpc g M) = map_mdpc (f \<circ> g) M"
by (intro mdpc_eqI)
(auto simp add: states_map_mdpc image_comp actions_map_mdpc image_UN map_pmf_compose[symmetric]
vimage_comp[symmetric])
lemma map_mdpc_id: "map_mdpc id = id"
by (auto simp: fun_eq_iff states_map_mdpc actions_map_mdpc intro!: mdpc_eqI)
lemma finite_states_map: "finite (states M) \<Longrightarrow> finite (map_mdpc f M)"
by (simp add: states_map_mdpc)
lemma finite_actions_map:
assumes "finite (states M)" "\<And>s. finite (actions M s)" shows "finite (actions (map_mdpc f M) s)"
proof -
have "(\<Union>x\<in>f -` {s}. actions M x) = (\<Union>x\<in>f -` {s} \<inter> states M. actions M x)"
using in_states[of M] by auto
with assms show ?thesis
by (auto simp add: actions_map_mdpc)
qed
lift_definition fix_loop :: "'s \<Rightarrow> 's mdpc \<Rightarrow> 's mdpc"
is "\<lambda>s M t. if s = t then Some {return_pmf s} else M t" .
lemma states_fix_loop[simp]: "states (fix_loop s M) = insert s (states M)"
by transfer (auto simp: subset_eq image_iff dom_def split: if_splits)
lemma actions_fix_loop[simp]: "actions (fix_loop s M) t = (if s = t then {return_pmf s} else actions M t)"
by transfer auto
lemma fix_loop_idem: "fix_loop s (fix_loop s M) = fix_loop s M"
by (auto intro!: mdpc_eqI)
lemma fix_loop_commute: "fix_loop s (fix_loop t M) = fix_loop t (fix_loop s M)"
by (auto intro!: mdpc_eqI)
lemma map_fix_loop:
assumes f_s: "\<And>t. f s = f t \<Longrightarrow> t = s"
shows "map_mdpc f (fix_loop s M) = fix_loop (f s) (map_mdpc f M)"
by (auto simp: states_map_mdpc actions_map_mdpc_eq_Collect split: if_splits intro!: mdpc_eqI dest!: f_s f_s[OF sym]) force+
lift_definition map_actions :: "('s \<Rightarrow> 's pmf set \<Rightarrow> 's pmf set) \<Rightarrow> 's mdpc \<Rightarrow> 's mdpc"
is "\<lambda>m f s. map_option (m s) (f s)" .
lemma state_map_actions[simp]: "states (map_actions f \<phi>) = states \<phi>"
by transfer auto
lemma actions_map_actions[simp]: "(s \<notin> states \<phi> \<Longrightarrow> f s {} = {}) \<Longrightarrow> actions (map_actions f \<phi>) s = f s (actions \<phi> s)"
by transfer (auto split: option.splits)
lift_definition restrict_states :: "'s set \<Rightarrow> 's mdpc \<Rightarrow> 's mdpc"
is "\<lambda>S f s. if s \<in> S then f s else None" .
lemma state_restrict_states[simp]: "states (restrict_states S \<phi>) = states \<phi> \<inter> S"
by transfer (auto split: if_splits)
lemma actions_restrict_states[simp]: "actions (restrict_states S \<phi>) s = (if s \<in> S then actions \<phi> s else {})"
by transfer (auto split: if_splits)
lemma restrict_states_idem: "states \<phi> \<subseteq> A \<Longrightarrow> restrict_states A \<phi> = \<phi>"
by transfer (force simp: fun_eq_iff subset_eq dom_def)
instantiation mdpc :: (type) lattice
begin
lift_definition less_eq_mdpc :: "'s mdpc \<Rightarrow> 's mdpc \<Rightarrow> bool"
is "fun_ord (ord_option (\<subseteq>))" .
definition less_mdpc :: "'s mdpc \<Rightarrow> 's mdpc \<Rightarrow> bool"
where "less_mdpc f g \<longleftrightarrow> (f \<le> g \<and> \<not> g \<le> f)"
lift_definition inf_mdpc :: "'s mdpc \<Rightarrow> 's mdpc \<Rightarrow> 's mdpc"
is "\<lambda>f g s. lift_option (\<inter>) (f s) (g s)" .
lift_definition sup_mdpc :: "'s mdpc \<Rightarrow> 's mdpc \<Rightarrow> 's mdpc"
is "\<lambda>f g s. combine_options (\<union>) (f s) (g s)" .
instance
proof
fix x y z :: "'s mdpc"
show "(x < y) = (x \<le> y \<and> \<not> y \<le> x)"
by (rule less_mdpc_def)
note ord =
fun_ord_refl[where 'b="'s", OF reflp_ord_option[where 'a="'s pmf set"], of "(\<subseteq>)"]
fun_ord_trans[where 'b="'s", OF transp_ord_option[where 'a="'s pmf set"], of "(\<subseteq>)"]
fun_ord_antisym[where 'b="'s", OF antisymp_ord_option[where 'a="'s pmf set"], of "(\<subseteq>)"]
show "x \<le> x" "x \<le> y \<Longrightarrow> y \<le> z \<Longrightarrow> x \<le> z" "x \<le> y \<Longrightarrow> y \<le> x \<Longrightarrow> x = y"
by (transfer; insert ord; auto simp: transp_def antisymp_def reflp_def)+
show "x \<sqinter> y \<le> x" "x \<sqinter> y \<le> y"
by (transfer; auto simp: fun_ord_def ord_option.simps lift_option_def split: Option.bind_split)+
show "x \<le> y \<Longrightarrow> x \<le> z \<Longrightarrow> x \<le> y \<sqinter> z"
apply transfer
subgoal premises prems for a b c
using prems by (rule fun_ord_combine) (auto simp: ord_option.simps)
done
show "x \<le> x \<squnion> y" "y \<le> x \<squnion> y"
by (transfer; auto simp: fun_ord_def ord_option.simps combine_options_def not_all_eq split: option.splits)+
show "y \<le> x \<Longrightarrow> z \<le> x \<Longrightarrow> y \<squnion> z \<le> x"
apply transfer
subgoal premises prems for a b c
using prems by (rule fun_ord_combine) (auto simp: ord_option.simps)
done
qed
end
instantiation mdpc :: (type) complete_lattice
begin
lift_definition bot_mdpc :: "'a mdpc" is "\<lambda>_. None" .
lift_definition top_mdpc :: "'a mdpc" is "\<lambda>_. Some UNIV" .
lift_definition Sup_mdpc :: "'a mdpc set \<Rightarrow> 'a mdpc"
is "\<lambda>M s. if \<exists>m\<in>M. m s \<noteq> None then Some (\<Union>{ d | m d. m \<in> M \<and> m s = Some d}) else None" .
lift_definition Inf_mdpc :: "'a mdpc set \<Rightarrow> 'a mdpc"
is "\<lambda>M s. if \<exists>m\<in>M. m s = None then None else Some (\<Inter>{ d | m d. m \<in> M \<and> m s = Some d})" .
instance
proof
fix x :: "'a mdpc" and X :: "'a mdpc set"
show "x \<in> X \<Longrightarrow> \<Sqinter>X \<le> x" "x \<in> X \<Longrightarrow> x \<le> \<Squnion>X"
by (transfer; force simp: fun_ord_def ord_option_Some1_iff ord_option_Some2_iff)+
show "(\<And>y. y \<in> X \<Longrightarrow> x \<le> y) \<Longrightarrow> x \<le> \<Sqinter>X"
apply transfer
apply (clarsimp simp: fun_ord_def ord_option.simps)
subgoal premises P for X m x
using P[rule_format, of _ x]
by (cases "m x") fastforce+
done
show "(\<And>y. y \<in> X \<Longrightarrow> y \<le> x) \<Longrightarrow> \<Squnion>X \<le> x"
apply transfer
apply (clarsimp simp: fun_ord_def ord_option.simps)
subgoal premises P for X m x y z
using P(1)[rule_format, of _ x] P(1)[rule_format, of y x] P(2,3)
by auto force
done
qed (transfer; auto)+
end
lemma states_sup[simp]: "states (\<phi> \<squnion> \<psi>) = states \<phi> \<union> states \<psi>"
by transfer (auto simp: combine_options_def split: option.splits)
lemma states_SUP[simp]: "states (\<Squnion>A) = (\<Union>a\<in>A. states a)"
by transfer (auto simp: dom_def split: option.splits if_splits)
lemma states_inf[simp]: "states (\<phi> \<sqinter> \<psi>) = states \<phi> \<inter> states \<psi>"
by transfer (auto simp: lift_option_eq_Some split: option.splits)
lemma states_mono: "\<phi> \<le> \<psi> \<Longrightarrow> states \<phi> \<subseteq> states \<psi>"
using states_sup[of \<phi> \<psi>] by (auto simp del: states_sup simp add: sup_absorb2)
lemma actions_sup[simp]: "actions (\<phi> \<squnion> \<psi>) = actions \<phi> \<squnion> actions \<psi>"
by transfer (auto simp: combine_options_def split: option.splits)
lemma actions_SUP[simp]: "actions (\<Squnion>A) s = (\<Union>a\<in>A. actions a s)"
by transfer (auto simp: dom_def split: option.splits if_splits, blast)
lemma actions_inf[simp]: "actions (\<phi> \<sqinter> \<psi>) = actions \<phi> \<sqinter> actions \<psi>"
by transfer (auto simp: fun_eq_iff split: option.splits)
lemma actions_mono: assumes *: "\<phi> \<le> \<psi>" shows "actions \<phi> \<le> actions \<psi>"
proof -
have "actions \<phi> \<le> actions \<phi> \<squnion> actions \<psi>"
by auto
also have "\<dots> = actions \<psi>"
using * actions_sup[of \<phi> \<psi>] by (auto simp add: sup_absorb2)
finally show ?thesis .
qed
lemma le_mdpcI: "states M \<subseteq> states N \<Longrightarrow> (\<And>s. s \<in> states M \<Longrightarrow> actions M s \<subseteq> actions N s) \<Longrightarrow> M \<le> N"
by transfer
(force simp: fun_ord_def ord_option.simps subset_eq split: option.splits)
lemma le_mdpc_iff: "M \<le> N \<longleftrightarrow> states M \<subseteq> states N \<and> (\<forall>s. actions M s \<subseteq> actions N s)"
using states_mono[of M N] actions_mono[of M N] by (auto simp: le_fun_def intro!: le_mdpcI)
lemma map_actions_le: "(\<And>s A. s \<in> states \<phi> \<Longrightarrow> f s A \<subseteq> A) \<Longrightarrow> map_actions f \<phi> \<le> \<phi>"
apply (intro le_mdpcI)
subgoal by auto
subgoal premises p for s using p(1)[of s] p(1)[of s "{}"] p(2) actions_map_actions by auto
done
lemma restrict_states_mono: "A \<subseteq> B \<Longrightarrow> \<phi> \<le> \<psi> \<Longrightarrow> restrict_states A \<phi> \<le> restrict_states B \<psi>"
using states_mono[of \<phi> \<psi>] actions_mono[of \<phi> \<psi>] by (intro le_mdpcI) (auto simp: le_fun_def)
lemma restrict_states_le: "restrict_states A M \<le> M"
by (intro le_mdpcI) auto
lemma eq_bot_iff_states: "\<phi> = bot \<longleftrightarrow> states \<phi> = {}"
by transfer auto
lemma fix_loop_neq_bot: "fix_loop s N \<noteq> bot"
unfolding eq_bot_iff_states by simp
lemma
shows states_bot[simp]: "states bot = {}"
and actions_bot[simp]: "actions bot = (\<lambda>s. {})"
unfolding fun_eq_iff by (transfer; auto)+
lemma inf_eq_bot_eq_disjnt_states: "A \<sqinter> B = bot \<longleftrightarrow> disjnt (states A) (states B)"
unfolding disjnt_def by transfer (auto simp: fun_eq_iff lift_option_eq_None)
text \<open>Enabled States\<close>
definition en :: "'s mdpc \<Rightarrow> 's rel"
where "en \<phi> = {(s, t) | d s t. d \<in> actions \<phi> s \<and> t \<in> set_pmf d}"
lemma en_sup[simp]: "en (\<phi> \<squnion> \<psi>) = en \<phi> \<union> en \<psi>"
by (auto simp: en_def)
lemma en_SUP[simp]: "en (Sup A) = (\<Union>a\<in>A. en a)"
by (auto simp: en_def)
lemma en_mono: "\<phi> \<le> \<psi> \<Longrightarrow> en \<phi> \<subseteq> en \<psi>"
unfolding en_def
apply transfer
apply (auto simp: fun_ord_def split: option.splits)
- subgoal for \<phi> \<psi> a b x y
- by (auto elim!: allE[of _ x] simp: ord_option.simps intro!: exI[of _ b] exI[of _ x])
+ apply (auto simp add: ord_option.simps subset_iff)
+ apply force
done
lemma en_states: "(s, t) \<in> en M \<Longrightarrow> s \<in> states M"
using in_states[of M s] by (auto simp: en_def)
lemma en_bot[simp]: "en bot = {}"
by (simp add: en_def)
lemma en_fix_loop[simp]: "en (fix_loop s M) = insert (s, s) (en M - {s} \<times> UNIV)"
by (force simp add: en_def )
lift_definition trivial :: "'s \<Rightarrow> 's mdpc" is "\<lambda>s. (\<lambda>_. None)(s := Some {})" .
lemma states_trivial[simp]: "states (trivial s) = {s}"
by transfer auto
lemma actions_trivial[simp]: "actions (trivial s) = (\<lambda>_. {})"
by transfer (auto simp: fun_eq_iff)
lemma en_trivial[simp]: "en (trivial s) = {}"
by (simp add: en_def)
lemma trivial_le_iff: "trivial x \<le> \<phi> \<longleftrightarrow> x \<in> states \<phi>"
by transfer (auto simp: ord_option.simps fun_ord_def)
lemma trivial_le: "x \<in> states \<phi> \<Longrightarrow> trivial x \<le> \<phi>"
unfolding trivial_le_iff .
lemma trivial_neq_bot: "trivial x \<noteq> bot"
by transfer auto
lift_definition loop :: "'s \<Rightarrow> 's mdpc"
is "\<lambda>s. (\<lambda>_. None)(s := Some {return_pmf s})" .
lemma states_loop[simp]: "states (loop s) = {s}"
by transfer auto
lemma actions_loop: "actions (loop s) = ((\<lambda>_. {})(s := {return_pmf s}))"
by transfer (auto simp: fun_eq_iff)
lemma
shows actions_loop_self[simp]: "actions (loop s) s = {return_pmf s}"
and actions_loop_neq[simp]: "s \<noteq> t \<Longrightarrow> actions (loop s) t = {}"
by (simp_all add: actions_loop)
lemma en_loop[simp]: "en (loop s) = {(s, s)}"
by (auto simp: en_def actions_loop)
lemma loop_neq_bot: "loop s \<noteq> bot"
unfolding eq_bot_iff_states by simp
lemma loop_le: "loop x \<le> M \<longleftrightarrow> (x \<in> states M \<and> return_pmf x \<in> actions M x)"
by (auto simp: le_mdpc_iff actions_loop)
lemma le_loop: "M \<le> loop x \<longleftrightarrow> (states M \<subseteq> {x} \<and> actions M x \<subseteq> {return_pmf x})"
using in_states[of M] by (auto simp: le_mdpc_iff actions_loop)
text \<open>Strongly Connected (SC)\<close>
definition sc :: "'s mdpc \<Rightarrow> bool"
where "sc \<phi> \<longleftrightarrow> states \<phi> \<times> states \<phi> \<subseteq> (en \<phi>)\<^sup>*"
lemma scD: "sc \<phi> \<Longrightarrow> x \<in> states \<phi> \<Longrightarrow> y \<in> states \<phi> \<Longrightarrow> (x, y) \<in> (en \<phi>)\<^sup>*"
by (auto simp: sc_def)
lemma scI: "(\<And>x y. x \<in> states \<phi> \<Longrightarrow> y \<in> states \<phi> \<Longrightarrow> (x, y) \<in> (en \<phi>)\<^sup>*) \<Longrightarrow> sc \<phi>"
by (auto simp: sc_def)
lemma sc_trivial[simp]: "sc (trivial s)"
by (simp add: sc_def)
lemma sc_loop[simp]: "sc (loop s)"
by (auto simp: sc_def)
lemma sc_bot[simp]: "sc bot"
by (simp add: sc_def)
lemma sc_SupI_directed:
assumes A: "\<And>a. a \<in> A \<Longrightarrow> sc a"
and directed: "\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> \<exists>c\<in>A. a \<le> c \<and> b \<le> c"
shows "sc (Sup A)"
unfolding sc_def
proof clarsimp
fix x y a b assume "a \<in> A" "b \<in> A" and xy: "x \<in> states a" "y \<in> states b"
with directed obtain c where "c \<in> A" "a \<le> c" "b \<le> c"
by auto
with xy have "x \<in> states c" "y \<in> states c"
using states_mono[of a c] states_mono[of b c] by auto
with A[OF \<open>c \<in> A\<close>] \<open>c \<in> A\<close>
have "(x, y) \<in> (en c)\<^sup>*"
by (auto simp: sc_def subset_eq)
then show "(x, y) \<in> (\<Union>x\<in>A. en x)\<^sup>*"
using rtrancl_mono[of "en c" "\<Union>a\<in>A. en a"] \<open>c\<in>A\<close> by auto
qed
lemma sc_supI:
assumes \<phi>: "sc \<phi>" and \<psi>: "sc \<psi>" and not_disjoint: "\<phi> \<sqinter> \<psi> \<noteq> bot"
shows "sc (\<phi> \<squnion> \<psi>)"
unfolding sc_def
proof safe
fix x y assume "x \<in> states (\<phi> \<squnion> \<psi>)" "y \<in> states (\<phi> \<squnion> \<psi>)"
moreover obtain z where "z \<in> states \<phi>" "z \<in> states \<psi>"
using not_disjoint by (auto simp: inf_eq_bot_eq_disjnt_states disjnt_def)
moreover have "(en \<phi>)\<^sup>* \<union> (en \<psi>)\<^sup>* \<subseteq> (en (\<phi> \<squnion> \<psi>))\<^sup>*"
by (metis rtrancl_Un_subset en_sup)
ultimately have "(x, z) \<in> (en (\<phi> \<squnion> \<psi>))\<^sup>*" "(z, y) \<in> (en (\<phi> \<squnion> \<psi>))\<^sup>*"
using \<phi> \<psi> by (auto dest: scD)
then show "(x, y) \<in> (en (\<phi> \<squnion> \<psi>))\<^sup>*"
by auto
qed
lemma sc_eq_loop:
assumes M: "sc M" and s: "s \<in> M" "actions M s = {return_pmf s}" shows "M = loop s"
proof -
{ fix t assume "t \<in> M"
then have "(s, t) \<in> (en M)\<^sup>*"
using M[THEN scD, OF \<open>s \<in> M\<close> \<open>t \<in> M\<close>] by simp
from this have "t = s"
by (induction rule: rtrancl_induct) (auto simp: en_def \<open>actions M s = {return_pmf s}\<close>) }
then have "states M = {s}"
using \<open>s \<in> M\<close> by blast
then show ?thesis
by (intro mdpc_eqI) (auto simp: \<open>actions M s = {return_pmf s}\<close>)
qed
lemma sc_eq_trivial:
assumes M: "sc M" and s: "s \<in> M" "actions M s = {}" shows "M = trivial s"
proof -
{ fix t assume "t \<in> M" "t \<noteq> s"
then have "(s, t) \<in> (en M)\<^sup>+"
using M[THEN scD, OF \<open>s \<in> M\<close> \<open>t \<in> M\<close>] by (simp add: rtrancl_eq_or_trancl)
from tranclD[OF this] \<open>actions M s = {}\<close> have False
by (auto simp: en_def) }
then have "states M = {s}"
using \<open>s \<in> M\<close> by auto
then show ?thesis
by (intro mdpc_eqI) (auto simp: \<open>actions M s = {}\<close>)
qed
definition closed_mdpc :: "'s mdpc \<Rightarrow> bool"
where "closed_mdpc \<phi> \<longleftrightarrow> en \<phi> \<subseteq> states \<phi> \<times> states \<phi>"
lemma closed_mdpcD: "closed_mdpc \<phi> \<Longrightarrow> D \<in> actions \<phi> x \<Longrightarrow> y \<in> D \<Longrightarrow> y \<in> states \<phi>"
by (auto simp: closed_mdpc_def en_def)
lemma closed_mdpc_supI: "closed_mdpc \<phi> \<Longrightarrow> closed_mdpc \<psi> \<Longrightarrow> closed_mdpc (\<phi> \<squnion> \<psi>)"
by (auto simp: closed_mdpc_def)
lemma closed_mdpc_SupI: "(\<And>a. a \<in> A \<Longrightarrow> closed_mdpc a) \<Longrightarrow> closed_mdpc (\<Squnion>A)"
by (auto simp: closed_mdpc_def)
lemma closed_mdpc_infI: "closed_mdpc \<phi> \<Longrightarrow> closed_mdpc \<psi> \<Longrightarrow> closed_mdpc (\<phi> \<sqinter> \<psi>)"
using en_mono[of "\<phi> \<sqinter> \<psi>" \<phi>] en_mono[of "\<phi> \<sqinter> \<psi>" \<psi>]
by (auto simp: closed_mdpc_def lift_option_eq_Some)
lemma closed_mdpc_trivial[simp]: "closed_mdpc (trivial s)"
by (simp add: closed_mdpc_def)
lemma closed_mdpc_bot[simp]: "closed_mdpc bot"
by (simp add: closed_mdpc_def)
lemma closed_mdpc_loop[simp]: "closed_mdpc (loop s)"
by (auto simp: closed_mdpc_def)
lemma closed_mdpc_fix_loop: "closed_mdpc M \<Longrightarrow> closed_mdpc (fix_loop s M)"
by (auto simp: closed_mdpc_def)
lemma closed_mdpc_map: assumes M: "closed_mdpc M" shows "closed_mdpc (map_mdpc f M)"
using closed_mdpcD[OF M]
by (auto simp: closed_mdpc_def en_def actions_map_mdpc states_map_mdpc intro!: imageI intro: in_states)
definition close :: "'s mdpc \<Rightarrow> 's mdpc"
where "close \<phi> = map_actions (\<lambda>s A. {a\<in>A. set_pmf a \<subseteq> states \<phi>}) \<phi>"
lemma
shows states_close[simp]: "states (close \<phi>) = states \<phi>"
and actions_close[simp]: "actions (close \<phi>) s = {a\<in>actions \<phi> s. a \<subseteq> states \<phi>}"
by (auto simp: close_def)
lemma closed_close: "closed_mdpc (close \<phi>)"
by (auto simp add: closed_mdpc_def en_def intro: in_states)
lemma close_closed: "closed_mdpc \<phi> \<Longrightarrow> close \<phi> = \<phi>"
unfolding closed_mdpc_def by (intro mdpc_eqI) (auto simp: en_def)
lemma close_close: "close (close \<phi>) = close \<phi>"
by (simp add: closed_close close_closed)
lemma close_le: "close M \<le> M"
unfolding close_def by (intro map_actions_le) auto
lemma close_mono: "\<phi> \<le> \<psi> \<Longrightarrow> close \<phi> \<le> close \<psi>"
using states_mono[of \<phi> \<psi>] actions_mono[of \<phi> \<psi>]
unfolding close_def by (intro le_mdpcI) (auto simp: le_fun_def)
text \<open>End Component (EC)\<close>
definition ec :: "'s mdpc \<Rightarrow> bool"
where "ec \<phi> \<longleftrightarrow> sc \<phi> \<and> closed_mdpc \<phi>"
lemma ec_trivial[simp]: "ec (trivial s)"
by (auto simp: ec_def)
lemma ec_bot[simp]: "ec bot"
by (auto simp: ec_def)
lemma ec_loop[simp]: "ec (loop s)"
by (auto simp: ec_def)
lemma ec_sup: "ec \<phi> \<Longrightarrow> ec \<psi> \<Longrightarrow> \<phi> \<sqinter> \<psi> \<noteq> bot \<Longrightarrow> ec (\<phi> \<squnion> \<psi>)"
by (simp add: ec_def sc_supI closed_mdpc_supI)
lemma ec_Sup_directed:
"(\<And>a. a \<in> A \<Longrightarrow> ec a) \<Longrightarrow> (\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> \<exists>c\<in>A. a \<le> c \<and> b \<le> c) \<Longrightarrow> ec (\<Squnion>A)"
by (auto simp: ec_def closed_mdpc_SupI sc_SupI_directed)
text \<open>Maximal End Component (MEC) relative to @{term M}\<close>
definition mec :: "'s mdpc \<Rightarrow> 's mdpc \<Rightarrow> bool"
where "mec M \<phi> \<longleftrightarrow> ec \<phi> \<and> \<phi> \<le> M \<and> (\<forall>\<psi>\<le>M. ec \<psi> \<longrightarrow> \<phi> \<le> \<psi> \<longrightarrow> \<phi> = \<psi>)"
lemma mec_refl: "ec M \<Longrightarrow> mec M M"
by (auto simp: mec_def)
lemma mec_le: "mec M \<phi> \<Longrightarrow> \<phi> \<le> M"
by (auto simp: mec_def)
lemma mec_ec: "mec M \<phi> \<Longrightarrow> ec \<phi>"
by (auto simp: mec_def)
lemma mec_least: "mec M \<phi> \<Longrightarrow> \<psi> \<le> M \<Longrightarrow> \<phi> \<le> \<psi> \<Longrightarrow> ec \<psi> \<Longrightarrow> \<phi> \<ge> \<psi>"
by (auto simp: mec_def)
lemma mec_bot_imp_bot: assumes "mec \<phi> bot" shows "\<phi> = bot"
proof (rule ccontr)
assume "\<phi> \<noteq> bot"
then obtain x where "x \<in> states \<phi>"
unfolding eq_bot_iff_states by auto
then have "ec (trivial x)" "trivial x \<le> \<phi>"
by (auto intro: trivial_le)
then have "trivial x = bot"
using \<open>mec \<phi> bot\<close> by (auto simp: mec_def)
then show False
by (simp add: trivial_neq_bot)
qed
lemma mec_imp_bot_eq_bot: "mec \<phi> \<psi> \<Longrightarrow> \<phi> = bot \<longleftrightarrow> \<psi> = bot"
using mec_bot_imp_bot[of \<phi>] by (auto simp: mec_def)
lemma mec_unique: assumes \<phi>: "mec M \<phi>" and \<psi>: "mec M \<psi>" and "\<phi> \<sqinter> \<psi> \<noteq> bot" shows "\<phi> = \<psi>"
proof -
have "mec M (\<phi> \<squnion> \<psi>)"
using assms
by (auto intro!: mec_def[THEN iffD2] ec_sup antisym dest: mec_ec mec_le)
(blast intro: le_supI1 mec_least[of M])
with mec_least[OF \<phi>, of "\<phi> \<squnion> \<psi>"] mec_least[OF \<psi>, of "\<phi> \<squnion> \<psi>"] mec_le[OF this] mec_ec[OF this]
show "\<phi> = \<psi>"
by auto
qed
lemma mec_exists: assumes \<phi>: "\<phi> \<noteq> bot" "ec \<phi>" and M: "\<phi> \<le> M" shows "\<exists>\<psi>\<ge>\<phi>. mec M \<psi>"
proof (intro exI conjI)
show "\<phi> \<le> \<Squnion>{\<psi>. \<phi> \<le> \<psi> \<and> \<psi> \<le> M \<and> ec \<psi>}"
using \<phi> M by (intro Sup_upper) auto
show "mec M (\<Squnion>{\<psi>. \<phi> \<le> \<psi> \<and> \<psi> \<le> M \<and> ec \<psi>})"
unfolding mec_def
proof safe
show "ec (\<Squnion>{\<psi>. \<phi> \<le> \<psi> \<and> \<psi> \<le> M \<and> ec \<psi>})"
proof (safe intro!: ec_Sup_directed)
fix a b assume *: "\<phi> \<le> a" "\<phi> \<le> b" and "a \<le> M" "b \<le> M" "ec a" "ec b"
moreover have "a \<sqinter> b \<noteq> bot"
using * \<phi> bot_unique[of "\<phi>"] le_inf_iff[of \<phi> a b] by (auto simp del: inf.bounded_iff)
ultimately show "\<exists>c\<in>{\<psi>. \<phi> \<le> \<psi> \<and> \<psi> \<le> M \<and> ec \<psi>}. a \<le> c \<and> b \<le> c"
by (intro bexI[of _ "sup a b"]) (auto intro: le_supI1 intro!: ec_sup)
qed
fix \<psi> assume \<psi>: "\<psi> \<le> M" "ec \<psi>" "\<Squnion>{\<psi>. \<phi> \<le> \<psi> \<and> \<psi> \<le> M \<and> ec \<psi>} \<le> \<psi>"
have "\<phi> \<le> \<Squnion>{\<psi>. \<phi> \<le> \<psi> \<and> \<psi> \<le> M \<and> ec \<psi>}"
using assms by (auto intro!: Sup_upper)
also have "\<dots> \<le> \<psi>" by fact
finally show "\<Squnion>{\<psi>. \<phi> \<le> \<psi> \<and> \<psi> \<le> M \<and> ec \<psi>} = \<psi>"
using \<psi> by (intro antisym Sup_upper) auto
qed (auto intro!: Sup_least)
qed
lemma mec_exists': "x \<in> states M \<Longrightarrow> \<exists>\<psi>. x \<in> states \<psi> \<and> mec M \<psi>"
using mec_exists[of "trivial x"] by (auto simp: trivial_neq_bot trivial_le_iff)
lemma mec_loop: "x \<in> states M \<Longrightarrow> actions M x = {return_pmf x} \<Longrightarrow> mec M (loop x)"
apply (auto simp: mec_def loop_le ec_def)
subgoal for \<phi>
using sc_eq_loop[of \<phi> x] actions_mono[of \<phi> M, THEN le_funD, of x] by auto
done
lemma mec_fix_loop: "mec (fix_loop s M) (loop s)"
by (intro mec_loop) auto
definition trivials :: "'s mdpc \<Rightarrow> 's set"
where "trivials M = {x. mec M (trivial x)}"
lemma trivials_subset_states: "trivials M \<subseteq> states M"
by (auto simp: trivials_def mec_def trivial_le_iff)
text \<open>Bottom MEC (BEMC) in @{term M}\<close>
definition bmec :: "'s mdpc \<Rightarrow> 's mdpc \<Rightarrow> bool"
where "bmec M \<phi> \<longleftrightarrow> mec \<phi> M \<and> (\<forall>s\<in>states \<phi>. actions \<phi> s = actions M s)"
definition actions' :: "'s mdpc \<Rightarrow> 's \<Rightarrow> 's pmf set"
where "actions' M s = (if s \<in> states M then actions M s else {return_pmf s})"
lemma closed_mdpcD':
"closed_mdpc M \<Longrightarrow> s \<in> states M \<Longrightarrow> (\<Union>D\<in>actions' M s. set_pmf D) \<subseteq> states M"
by (auto simp: actions'_def dest: closed_mdpcD)
locale Finite_MDP =
fixes M :: "'s mdpc"
assumes closed_M: "closed_mdpc M" and M_neq_bot: "M \<noteq> bot"
and actions_neq_empty_M: "\<And>s. s \<in> states M \<Longrightarrow> actions M s \<noteq> {}"
and finite_states_M: "finite M"
and finite_actions_M: "\<And>s. finite (actions M s)"
begin
sublocale Finite_Markov_Decision_Process "actions' M" "states M"
proof
show "actions' M s \<noteq> {}" for s
using actions_neq_empty_M by (auto simp: actions'_def )
show "states M \<noteq> {}" "finite M" "\<And>s. finite (actions' M s)"
using M_neq_bot finite_states_M finite_actions_M by (auto simp: eq_bot_iff_states actions'_def)
show "s \<in> states M \<Longrightarrow> (\<Union>D\<in>actions' M s. set_pmf D) \<subseteq> states M" for s
using closed_M by (rule closed_mdpcD')
qed
lemma Finite_MDP_map_loop: "Finite_MDP (map_mdpc f M \<squnion> loop s)"
proof
show "closed_mdpc (map_mdpc f M \<squnion> loop s)"
by (intro closed_mdpc_supI closed_mdpc_map closed_M closed_mdpc_loop)
show "finite (actions (map_mdpc f M \<squnion> loop s) t)" for t
by (auto simp: actions_loop intro!: finite_actions_map finite_states_M finite_actions_M)
show "finite (map_mdpc f M \<squnion> loop s)"
by (auto intro!: finite_states_M finite_states_map)
qed (auto simp: loop_neq_bot states_map_mdpc actions_loop actions_map_mdpc dest: actions_neq_empty_M)
lemma Finite_MDP_map_fix_loop: "Finite_MDP (fix_loop s (map_mdpc f M))"
proof
show "closed_mdpc (fix_loop s (map_mdpc f M))"
by (intro closed_mdpc_supI closed_mdpc_map closed_M closed_mdpc_fix_loop)
show "finite (actions (fix_loop s (map_mdpc f M)) t)" for t
by (auto simp: actions_loop intro!: finite_actions_map finite_states_M finite_actions_M)
show "finite (fix_loop s (map_mdpc f M))"
by (auto intro!: finite_states_M finite_states_map)
qed (auto simp: fix_loop_neq_bot states_map_mdpc actions_map_mdpc dest: actions_neq_empty_M)
end
context
fixes M :: "'s mdpc"
and F :: "'s set"
assumes M_neq_bot: "M \<noteq> bot"
and closed_M: "closed_mdpc M"
and actions_neq_empty_M: "\<And>s. s \<in> states M \<Longrightarrow> actions M s \<noteq> {}"
and finite_states_M: "finite M"
and finite_actions_M: "\<And>s. finite (actions M s)"
and F_subset: "F \<subseteq> states M"
begin
lemma finite_F[simp]: "finite F"
using F_subset finite_states_M by (auto dest: finite_subset)
interpretation M: Finite_MDP M
proof qed fact+
interpretation M: Reachability_Problem "actions' M" "states M" "states M - F" F
proof qed (insert F_subset, auto)
definition r :: "'s \<Rightarrow> 's option"
where "r s = (if s \<in> F then None else Some s)"
lemma r_eq_None[simp]: "r s = None \<longleftrightarrow> s \<in> F"
by (simp add: r_def)
lemma r_eq_Some[simp]: "r s = Some t \<longleftrightarrow> (s \<notin> F \<and> s = t)"
by (simp add: r_def)
lemma r_in_Some_image: "r s \<in> Some ` X \<longleftrightarrow> (s \<notin> F \<and> s \<in> X)"
by (auto simp: r_def)
lemma r_inj: "s \<notin> F \<or> t \<notin> F \<Longrightarrow> r s = r t \<longleftrightarrow> s = t"
by (auto simp: r_def)
lemma shows r_F: "s \<in> F \<Longrightarrow> r s = None" and r_nF: "s \<notin> F \<Longrightarrow> r s = Some s"
by auto
definition R :: "'s option mdpc"
where "R = fix_loop None (map_mdpc r M)"
lemma closed_R: "closed_mdpc R"
unfolding R_def by (intro closed_mdpc_map closed_M closed_mdpc_fix_loop)
lemma states_R[simp]: "states R = Some ` (states M - F) \<union> {None}"
by (auto simp add: R_def r_def[abs_def] states_map_mdpc)
lemma actions_R_None[simp]:
"actions R None = {return_pmf None}"
by (auto simp add: R_def)
lemma actions_R_Some[simp]:
"actions R (Some s) = (if s \<in> F then {} else map_pmf r ` actions M s)"
by (auto simp add: R_def actions_map_mdpc split: if_splits intro!: imageI)
lemma mec_R_loop: "mec R (loop None)"
unfolding R_def by (intro mec_fix_loop)
interpretation R: Finite_MDP R
unfolding R_def by (rule M.Finite_MDP_map_fix_loop)
interpretation R: Reachability_Problem "actions' R" "states R" "{None}" "{}"
proof qed auto
lemma F_not_trivial: "s \<in> F \<Longrightarrow> Some s \<notin> trivials R"
by (auto simp: trivials_def mec_def trivial_le_iff)
primrec min_state :: "'s option \<Rightarrow> 's + bool"
where
"min_state None = Inr True"
| "min_state (Some s) = (if Some s \<in> trivials R then Inl s else Inr False)"
lemma min_state_eq_Inl: "min_state s = Inl t \<longleftrightarrow> (Some t \<in> trivials R \<and> s = Some t)"
by (cases s) auto
lemma min_state_eq_Inr: "min_state s = Inr b \<longleftrightarrow> (if b then s = None else s \<noteq> None \<and> s \<notin> trivials R)"
by (cases s) auto
lemma map_min_state_R: "map_mdpc min_state R = fix_loop (Inr True) (map_mdpc (min_state \<circ> r) M)"
unfolding R_def
by (subst map_fix_loop)
(auto simp: map_mdpc_compose min_state_eq_Inr eq_commute[of "Inr True"])
definition min_mdpc :: "('s + bool) mdpc"
where "min_mdpc = fix_loop (Inr False) (map_mdpc min_state R)"
lemma states_min_mdpc: "states min_mdpc = {Inl t | t. Some t \<in> trivials R} \<union> {Inr True, Inr False}"
using trivials_subset_states[of R] by (auto simp add: min_mdpc_def states_map_mdpc image_comp split: if_splits)
lemma actions_min_mdpc_Inl:
"actions min_mdpc (Inl t) = (if Some t \<in> trivials R then map_pmf (min_state \<circ> r) ` actions M t else {})"
proof -
have eq: "min_state -` {Inl t} = (if Some t \<in> trivials R then {Some t} else {})"
by (auto simp: min_state_eq_Inl)
show ?thesis using F_not_trivial[of t]
by (simp add: min_mdpc_def actions_map_mdpc eq image_comp map_pmf_compose[symmetric])
qed
lemma actions_min_mdpc_Inr: "actions min_mdpc (Inr b) = {return_pmf (Inr b)}"
by (simp add: min_mdpc_def map_min_state_R)
interpretation min: Finite_MDP min_mdpc
unfolding min_mdpc_def by (rule R.Finite_MDP_map_fix_loop)
interpretation min: Reachability_Problem "actions' min_mdpc" "states min_mdpc" "states min_mdpc - {Inr True}" "{Inr True}"
proof qed (auto simp: states_min_mdpc)
lemma M_n_eq_0_not_trivials:
assumes "s \<in> states M" "s \<notin> F" "Some s \<notin> trivials R"
shows "M.n s = 0"
proof -
have "Some s \<in> states R"
using assms by auto
obtain \<phi> where "mec R \<phi>" "s \<in> Some -` \<phi>"
using mec_exists'[OF \<open>Some s \<in> states R\<close>] by auto
then have action_\<phi>: "Some t \<in> \<phi> \<Longrightarrow> actions \<phi> (Some t) \<noteq> {}" for t
using mec_ec[OF \<open>mec R \<phi>\<close>] \<open>Some s \<notin> trivials R\<close> sc_eq_trivial[of \<phi> "Some t"]
by (auto simp: ec_def trivials_def)
have None_notin_states: "None \<notin> states \<phi>"
using mec_R_loop \<open>mec R \<phi>\<close> \<open>s \<in> Some -` \<phi>\<close> mec_unique[of R "loop None" \<phi>]
by (auto simp: inf_eq_bot_eq_disjnt_states disjnt_def)
from \<open>s \<in> Some -` \<phi>\<close> show "M.n s = 0"
proof (rule M.n_eq_0_closed)
show "Some -` states \<phi> \<subseteq> states M" "Some -` states \<phi> \<inter> F = {}"
using mec_le[OF \<open>mec R \<phi>\<close>] by (auto simp: r_def le_mdpc_iff)
fix s assume "s \<in> Some -` \<phi>"
then have s: "s \<in> states M" "s \<notin> F" "actions \<phi> (Some s) \<noteq> {}"
using mec_le[OF \<open>mec R \<phi>\<close>] by (auto simp: le_mdpc_iff action_\<phi>)
then obtain D where D: "D \<in> actions \<phi> (Some s)"
by auto
then have "D \<in> actions R (Some s)"
using mec_le[OF \<open>mec R \<phi>\<close>, THEN actions_mono] s by (auto simp add: le_fun_def simp del: actions_R_Some)
with s obtain D' where D_eq: "D = map_pmf r D'" and D': "D' \<in> actions M s"
by auto
have "set_pmf D \<subseteq> states \<phi>"
using closed_mdpcD[OF _ D] mec_ec[OF \<open>mec R \<phi>\<close>] by (auto simp: ec_def)
then have "set_pmf D = Some ` set_pmf D'"
using closed_mdpcD[of \<phi>, OF _ \<open>D \<in> actions \<phi> (Some s)\<close>] None_notin_states
mec_ec[OF \<open>mec R \<phi>\<close>]
unfolding D_eq by (auto intro!: image_cong simp: r_def ec_def)
then show "\<exists>x\<in>actions' M s. set_pmf x \<subseteq> Some -` states \<phi>"
using \<open>s \<in> states M\<close> \<open>set_pmf D \<subseteq> states \<phi>\<close> D'
by (intro bexI[of _ D']) (auto simp: actions'_def)
qed
qed
lemma min_state_r_in_min_mdpc[simp]: "s \<in> M \<Longrightarrow> min_state (r s) \<in> min_mdpc"
by (auto simp add: states_min_mdpc min_state_eq_Inr min_state_eq_Inl r_def)
end
end
diff --git a/thys/Noninterference_Concurrent_Composition/ConcurrentComposition.thy b/thys/Noninterference_Concurrent_Composition/ConcurrentComposition.thy
--- a/thys/Noninterference_Concurrent_Composition/ConcurrentComposition.thy
+++ b/thys/Noninterference_Concurrent_Composition/ConcurrentComposition.thy
@@ -1,3404 +1,3404 @@
(* Title: Conservation of CSP Noninterference Security under Concurrent Composition
Author: Pasquale Noce
Security Certification Specialist at Arjo Systems, Italy
pasquale dot noce dot lavoro at gmail dot com
pasquale dot noce at arjosystems dot com
*)
section "Concurrent composition and noninterference security"
theory ConcurrentComposition
imports Noninterference_Sequential_Composition.Propaedeutics
begin
text \<open>
\null
In his outstanding work on Communicating Sequential Processes \cite{R6}, Hoare has defined two
fundamental binary operations allowing to compose the input processes into another, typically more
complex, process: sequential composition and concurrent composition. Particularly, the output of the
latter operation is a process in which any event not shared by both operands can occur whenever the
operand that admits the event can engage in it, whereas any event shared by both operands can occur
just in case both can engage in it. In other words, shared events are those that synchronize the
concurrent processes, which on the contrary can engage asynchronously in the respective non-shared
events.
This paper formalizes Hoare's definition of concurrent composition and proves, in the general case
of a possibly intransitive policy, that CSP noninterference security \cite{R1} is conserved under
this operation, viz. the security of both of the input processes implies that of the output process.
This result, along with the analogous one concerning sequential composition attained in \cite{R5},
enables the construction of more and more complex processes enforcing noninterference security by
composing, sequentially or concurrently, simpler secure processes, whose security can in turn be
proven using either the definition of security formulated in \cite{R1}, or the unwinding theorems
demonstrated in \cite{R2}, \cite{R3}, and \cite{R4}.
Throughout this paper, the salient points of definitions and proofs are commented; for additional
information, cf. Isabelle documentation, particularly \cite{R7}, \cite{R8}, \cite{R9}, and
\cite{R10}.
\<close>
subsection "Propaedeutic definitions and lemmas"
text \<open>
The starting point is comprised of some definitions and lemmas propaedeutic to the proof of the
target security conservation theorem.
Particularly, the definition of operator \emph{after} given in \cite{R6} is formalized, and it is
proven that for any secure process @{term P} and any trace @{term xs} of @{term P}, @{term P} after
@{term xs} is still a secure process. Then, this result is used to generalize the lemma stating the
closure of the failures of a secure process @{term P} under intransitive purge, proven in \cite{R5},
to the futures of @{term P} associated to any one of its traces. This is a generalization of the
former result since @{term "futures P xs = failures P"} for @{term "xs = []"}.
\null
\<close>
lemma sinks_aux_elem [rule_format]:
"u \<in> sinks_aux I D U xs \<longrightarrow> u \<in> U \<or> (\<exists>x \<in> set xs. u = D x)"
by (induction xs rule: rev_induct, simp_all, blast)
lemma ipurge_ref_aux_cons:
"ipurge_ref_aux I D U (x # xs) X = ipurge_ref_aux I D (sinks_aux I D U [x]) xs X"
by (subgoal_tac "x # xs = [x] @ xs", simp only: ipurge_ref_aux_append, simp)
lemma process_rule_1_futures:
"xs \<in> traces P \<Longrightarrow> ([], {}) \<in> futures P xs"
by (simp add: futures_def, rule traces_failures)
lemma process_rule_3_futures:
"(ys, Y) \<in> futures P xs \<Longrightarrow> Y' \<subseteq> Y \<Longrightarrow> (ys, Y') \<in> futures P xs"
by (simp add: futures_def, rule process_rule_3)
lemma process_rule_4_futures:
"(ys, Y) \<in> futures P xs \<Longrightarrow>
(ys @ [x], {}) \<in> futures P xs \<or> (ys, insert x Y) \<in> futures P xs"
by (simp add: futures_def, subst append_assoc [symmetric], rule process_rule_4)
lemma process_rule_5_general [rule_format]:
"xs \<in> divergences P \<longrightarrow> xs @ ys \<in> divergences P"
proof (induction ys rule: rev_induct, simp, rule impI, simp)
qed (subst append_assoc [symmetric], rule process_rule_5)
text \<open>
\null
Here below is the definition of operator \emph{after}, for which a symbolic notation similar to the
one used in \cite{R6} is introduced. Then, it is proven that for any process @{term P} and any trace
@{term xs} of @{term P}, the failures set and the divergences set of @{term P} after @{term xs}
indeed enjoy their respective characteristic properties as defined in \cite{R1}.
\null
\<close>
definition future_divergences :: "'a process \<Rightarrow> 'a list \<Rightarrow> 'a list set" where
"future_divergences P xs \<equiv> {ys. xs @ ys \<in> divergences P}"
definition after :: "'a process \<Rightarrow> 'a list \<Rightarrow> 'a process" (infixl "\<setminus>" 64) where
"P \<setminus> xs \<equiv> Abs_process (futures P xs, future_divergences P xs)"
lemma process_rule_5_futures:
"ys \<in> future_divergences P xs \<Longrightarrow> ys @ [x] \<in> future_divergences P xs"
by (simp add: future_divergences_def, subst append_assoc [symmetric],
rule process_rule_5)
lemma process_rule_6_futures:
"ys \<in> future_divergences P xs \<Longrightarrow> (ys, Y) \<in> futures P xs"
by (simp add: futures_def future_divergences_def, rule process_rule_6)
lemma after_rep:
assumes A: "xs \<in> traces P"
shows "Rep_process (P \<setminus> xs) = (futures P xs, future_divergences P xs)"
(is "_ = ?X")
proof (subst after_def, rule Abs_process_inverse, simp add: process_set_def,
(subst conj_assoc [symmetric])+, (rule conjI)+)
show "process_prop_1 ?X"
proof (simp add: process_prop_1_def)
qed (rule process_rule_1_futures [OF A])
next
show "process_prop_2 ?X"
proof (simp add: process_prop_2_def del: all_simps, (rule allI)+, rule impI)
qed (rule process_rule_2_futures)
next
show "process_prop_3 ?X"
proof (simp add: process_prop_3_def del: all_simps, (rule allI)+, rule impI,
erule conjE)
qed (rule process_rule_3_futures)
next
show "process_prop_4 ?X"
proof (simp add: process_prop_4_def, (rule allI)+, rule impI)
qed (rule process_rule_4_futures)
next
show "process_prop_5 ?X"
proof (simp add: process_prop_5_def, rule allI, rule impI, rule allI)
qed (rule process_rule_5_futures)
next
show "process_prop_6 ?X"
proof (simp add: process_prop_6_def, rule allI, rule impI, rule allI)
qed (rule process_rule_6_futures)
qed
lemma after_failures:
assumes A: "xs \<in> traces P"
shows "failures (P \<setminus> xs) = futures P xs"
by (simp add: failures_def after_rep [OF A])
lemma after_futures:
assumes A: "xs \<in> traces P"
shows "futures (P \<setminus> xs) ys = futures P (xs @ ys)"
by (simp add: futures_def after_failures [OF A])
text \<open>
\null
Finally, the closure of the futures of a secure process under intransitive purge is proven.
\null
\<close>
lemma after_secure:
assumes A: "xs \<in> traces P"
shows "secure P I D \<Longrightarrow> secure (P \<setminus> xs) I D"
by (simp add: secure_def after_futures [OF A], blast)
lemma ipurge_tr_ref_aux_futures:
"\<lbrakk>secure P I D; (ys, Y) \<in> futures P xs\<rbrakk> \<Longrightarrow>
(ipurge_tr_aux I D U ys, ipurge_ref_aux I D U ys Y) \<in> futures P xs"
proof (subgoal_tac "xs \<in> traces P", simp add: after_failures [symmetric],
rule ipurge_tr_ref_aux_failures, rule after_secure, assumption+)
qed (simp add: futures_def, drule failures_traces, rule process_rule_2_traces)
lemma ipurge_tr_ref_aux_failures_general:
"\<lbrakk>secure P I D; (xs @ ys, Y) \<in> failures P\<rbrakk> \<Longrightarrow>
(xs @ ipurge_tr_aux I D U ys, ipurge_ref_aux I D U ys Y) \<in> failures P"
by (drule ipurge_tr_ref_aux_futures, simp_all add: futures_def)
subsection "Concurrent composition"
text \<open>
In \cite{R6}, the concurrent composition of two processes @{term P}, @{term Q}, expressed using
notation \<open>P \<parallel> Q\<close>, is defined as a process whose alphabet is the union of the alphabets of
@{term P} and @{term Q}, so that the shared events requiring the synchronous participation of both
processes are those in the intersection of their alphabets.
In the formalization of Communicating Sequential Processes developed in \cite{R1}, the alphabets of
@{term P} and @{term Q} are the data types @{typ 'a} and @{typ 'b} nested in their respective types
@{typ "'a process"} and @{typ "'b process"}. Therefore, for any two maps @{term "p :: 'a \<Rightarrow> 'c"},
@{term "q :: 'b \<Rightarrow> 'c"}, the concurrent composition of @{term P} and @{term Q} with respect to
@{term p} and @{term q}, expressed using notation \<open>P \<parallel> Q <p, q>\<close>, is defined in what follows
as a process of type @{typ "'c process"}, where meaningful events are those in
@{term "range p \<union> range q"} and shared events are those in @{term "range p \<inter> range q"}.
The case where @{term "- (range p \<union> range q) \<noteq> {}"} constitutes a generalization of the definition
given in \cite{R6}, and the events in @{term "- (range p \<union> range q)"}, not being mapped to any event
in the alphabets of the input processes, shall be understood as fake events lacking any meaning.
Consistently with this interpretation, such events are allowed to occur in divergent traces only --
necessarily, since divergences are capable by definition of giving rise to any sort of event. As a
result, while in \cite{R6} the refusals associated to non-divergent traces are the union of two
sets, a refusal of @{term P} and a refusal of @{term Q}, in the following definition they are the
union of three sets instead, where the third set is any subset of @{term "- (range p \<union> range q)"}.
Since the definition given in \cite{R6} preserves the identity of the events of the input processes,
a further generalization resulting from the following definition corresponds to the case where
either map @{term p}, @{term q} is not injective. However, as shown below, these generalizations
turn out to compromise neither the compliance of the output of concurrent composition with the
characteristic properties of processes as defined in \cite{R1}, nor even the validity of the target
security conservation theorem.
Since divergences can contain fake events, whereas non-divergent traces cannot, it is necessary to
add divergent failures to the failures set explicitly. The following definition of the divergences
set restricts the definition given in \cite{R6}, as it identifies a divergence with an arbitrary
extension of an event sequence @{term xs} being a divergence of both @{term P} and @{term Q}, rather
than a divergence of either process and a trace of the other one. This is a reasonable restriction,
in that it requires the concurrent composition of @{term P} and @{term Q} to admit a shared event
@{term x} in a divergent trace just in case both @{term P} and @{term Q} diverge and can then accept
@{term x}, analogously to what is required for a non-divergent trace. Anyway, the definitions match
if the input processes do not diverge, which is the case for any process of practical significance
(cf. \cite{R6}).
\null
\<close>
definition con_comp_divergences ::
"'a process \<Rightarrow> 'b process \<Rightarrow> ('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'c list set" where
"con_comp_divergences P Q p q \<equiv>
{xs @ ys | xs ys.
set xs \<subseteq> range p \<union> range q \<and>
map (inv p) [x\<leftarrow>xs. x \<in> range p] \<in> divergences P \<and>
map (inv q) [x\<leftarrow>xs. x \<in> range q] \<in> divergences Q}"
definition con_comp_failures ::
"'a process \<Rightarrow> 'b process \<Rightarrow> ('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'c failure set" where
"con_comp_failures P Q p q \<equiv>
{(xs, X \<union> Y \<union> Z) | xs X Y Z.
set xs \<subseteq> range p \<union> range q \<and>
X \<subseteq> range p \<and> Y \<subseteq> range q \<and> Z \<subseteq> - (range p \<union> range q) \<and>
(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` X) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` Y) \<in> failures Q} \<union>
{(xs, X). xs \<in> con_comp_divergences P Q p q}"
definition con_comp ::
"'a process \<Rightarrow> 'b process \<Rightarrow> ('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'c process" where
"con_comp P Q p q \<equiv>
Abs_process (con_comp_failures P Q p q, con_comp_divergences P Q p q)"
abbreviation con_comp_syntax ::
"'a process \<Rightarrow> 'b process \<Rightarrow> ('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'c process"
("(_ \<parallel> _ <_, _>)" 55)
where
"P \<parallel> Q <p, q> \<equiv> con_comp P Q p q"
text \<open>
\null
Here below is the proof that, for any two processes @{term P}, @{term Q} and any two maps @{term p},
@{term q}, sets @{term "con_comp_failures P Q p q"} and @{term "con_comp_divergences P Q p q"} enjoy
the characteristic properties of the failures and the divergences sets of a process as defined in
\cite{R1}.
\null
\<close>
lemma con_comp_prop_1:
"([], {}) \<in> con_comp_failures P Q p q"
proof (simp add: con_comp_failures_def)
qed (rule disjI1, rule conjI, (rule process_rule_1)+)
lemma con_comp_prop_2:
"(xs @ [x], X) \<in> con_comp_failures P Q p q \<Longrightarrow>
(xs, {}) \<in> con_comp_failures P Q p q"
proof (simp add: con_comp_failures_def del: filter_append,
erule disjE, (erule exE)+, (erule conjE)+, rule disjI1)
fix X Y
assume
A: "set xs \<subseteq> range p \<union> range q" and
B: "(map (inv p) [x\<leftarrow>xs @ [x]. x \<in> range p], inv p ` X) \<in> failures P" and
C: "(map (inv q) [x\<leftarrow>xs @ [x]. x \<in> range q], inv q ` Y) \<in> failures Q"
show "set xs \<subseteq> range p \<union> range q \<and>
(map (inv p) [x\<leftarrow>xs. x \<in> range p], {}) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs. x \<in> range q], {}) \<in> failures Q"
proof (simp add: A, rule conjI, cases "x \<in> range p",
case_tac [3] "x \<in> range q")
assume "x \<in> range p"
hence "(map (inv p) [x\<leftarrow>xs. x \<in> range p] @ [inv p x], inv p ` X) \<in> failures P"
using B by simp
thus "(map (inv p) [x\<leftarrow>xs. x \<in> range p], {}) \<in> failures P"
by (rule process_rule_2)
next
assume "x \<notin> range p"
hence "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` X) \<in> failures P"
using B by simp
moreover have "{} \<subseteq> inv p ` X" ..
ultimately show "(map (inv p) [x\<leftarrow>xs. x \<in> range p], {}) \<in> failures P"
by (rule process_rule_3)
next
assume "x \<in> range q"
hence "(map (inv q) [x\<leftarrow>xs. x \<in> range q] @ [inv q x], inv q ` Y) \<in> failures Q"
using C by simp
thus "(map (inv q) [x\<leftarrow>xs. x \<in> range q], {}) \<in> failures Q"
by (rule process_rule_2)
next
assume "x \<notin> range q"
hence "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` Y) \<in> failures Q"
using C by simp
moreover have "{} \<subseteq> inv q ` Y" ..
ultimately show "(map (inv q) [x\<leftarrow>xs. x \<in> range q], {}) \<in> failures Q"
by (rule process_rule_3)
qed
next
assume A: "xs @ [x] \<in> con_comp_divergences P Q p q"
show
"set xs \<subseteq> range p \<union> range q \<and>
(map (inv p) [x\<leftarrow>xs. x \<in> range p], {}) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs. x \<in> range q], {}) \<in> failures Q \<or>
xs \<in> con_comp_divergences P Q p q"
(is "?A \<or> _")
proof (insert A, simp add: con_comp_divergences_def,
((erule exE)?, erule conjE)+)
fix ws ys
assume
B: "xs @ [x] = ws @ ys" and
C: "set ws \<subseteq> range p \<union> range q" and
D: "map (inv p) [x\<leftarrow>ws. x \<in> range p] \<in> divergences P" and
E: "map (inv q) [x\<leftarrow>ws. x \<in> range q] \<in> divergences Q"
show "?A \<or> (\<exists>ws'.
(\<exists>ys'. xs = ws' @ ys') \<and>
set ws' \<subseteq> range p \<union> range q \<and>
map (inv p) [x\<leftarrow>ws'. x \<in> range p] \<in> divergences P \<and>
map (inv q) [x\<leftarrow>ws'. x \<in> range q] \<in> divergences Q)"
(is "_ \<or> (\<exists>ws'. ?B ws')")
proof (cases ys, rule disjI1, rule_tac [2] disjI2)
case Nil
hence "set (xs @ [x]) \<subseteq> range p \<union> range q"
using B and C by simp
hence "insert x (set xs) \<subseteq> range p \<union> range q"
by simp
moreover have "set xs \<subseteq> insert x (set xs)"
by (rule subset_insertI)
ultimately have "set xs \<subseteq> range p \<union> range q"
by simp
moreover have "map (inv p) [x\<leftarrow>xs @ [x]. x \<in> range p] \<in> divergences P"
using Nil and B and D by simp
hence "(map (inv p) [x\<leftarrow>xs. x \<in> range p], {}) \<in> failures P"
proof (cases "x \<in> range p", simp_all)
assume "map (inv p) [x\<leftarrow>xs. x \<in> range p] @ [inv p x] \<in> divergences P"
hence "(map (inv p) [x\<leftarrow>xs. x \<in> range p] @ [inv p x], {}) \<in> failures P"
by (rule process_rule_6)
thus ?thesis
by (rule process_rule_2)
next
assume "map (inv p) [x\<leftarrow>xs. x \<in> range p] \<in> divergences P"
thus ?thesis
by (rule process_rule_6)
qed
moreover have "map (inv q) [x\<leftarrow>xs @ [x]. x \<in> range q] \<in> divergences Q"
using Nil and B and E by simp
hence "(map (inv q) [x\<leftarrow>xs. x \<in> range q], {}) \<in> failures Q"
proof (cases "x \<in> range q", simp_all)
assume "map (inv q) [x\<leftarrow>xs. x \<in> range q] @ [inv q x] \<in> divergences Q"
hence "(map (inv q) [x\<leftarrow>xs. x \<in> range q] @ [inv q x], {}) \<in> failures Q"
by (rule process_rule_6)
thus ?thesis
by (rule process_rule_2)
next
assume "map (inv q) [x\<leftarrow>xs. x \<in> range q] \<in> divergences Q"
thus ?thesis
by (rule process_rule_6)
qed
ultimately show ?A
by blast
next
case Cons
moreover have "butlast (xs @ [x]) = butlast (ws @ ys)"
using B by simp
ultimately have "xs = ws @ butlast ys"
by (simp add: butlast_append)
hence "\<exists>ys'. xs = ws @ ys'" ..
hence "?B ws"
using C and D and E by simp
thus "\<exists>ws'. ?B ws'" ..
qed
qed
qed
lemma con_comp_prop_3:
"\<lbrakk>(xs, Y) \<in> con_comp_failures P Q p q; X \<subseteq> Y\<rbrakk> \<Longrightarrow>
(xs, X) \<in> con_comp_failures P Q p q"
proof (simp add: con_comp_failures_def, erule disjE, simp_all,
(erule exE)+, (erule conjE)+, rule disjI1, simp)
fix X' Y' Z'
assume
A: "X \<subseteq> X' \<union> Y' \<union> Z'" and
B: "X' \<subseteq> range p" and
C: "Y' \<subseteq> range q" and
D: "Z' \<subseteq> - range p" and
E: "Z' \<subseteq> - range q" and
F: "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` X') \<in> failures P" and
G: "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` Y') \<in> failures Q"
show "\<exists>X' Y' Z'.
X = X' \<union> Y' \<union> Z' \<and>
X' \<subseteq> range p \<and>
Y' \<subseteq> range q \<and>
Z' \<subseteq> - range p \<and>
Z' \<subseteq> - range q \<and>
(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` X') \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` Y') \<in> failures Q"
proof (rule_tac x = "X' \<inter> X" in exI, rule_tac x = "Y' \<inter> X" in exI,
rule_tac x = "Z' \<inter> X" in exI, (subst conj_assoc [symmetric])+, (rule conjI)+)
show "X = X' \<inter> X \<union> Y' \<inter> X \<union> Z' \<inter> X"
using A by blast
next
show "X' \<inter> X \<subseteq> range p"
using B by blast
next
show "Y' \<inter> X \<subseteq> range q"
using C by blast
next
show "Z' \<inter> X \<subseteq> - range p"
using D by blast
next
show "Z' \<inter> X \<subseteq> - range q"
using E by blast
next
have "inv p ` (X' \<inter> X) \<subseteq> inv p ` X'"
by blast
with F show "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` (X' \<inter> X))
\<in> failures P"
by (rule process_rule_3)
next
have "inv q ` (Y' \<inter> X) \<subseteq> inv q ` Y'"
by blast
with G show "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` (Y' \<inter> X))
\<in> failures Q"
by (rule process_rule_3)
qed
qed
lemma con_comp_prop_4:
"(xs, X) \<in> con_comp_failures P Q p q \<Longrightarrow>
(xs @ [x], {}) \<in> con_comp_failures P Q p q \<or>
(xs, insert x X) \<in> con_comp_failures P Q p q"
proof (simp add: con_comp_failures_def del: filter_append,
erule disjE, (erule exE)+, (erule conjE)+, simp_all del: filter_append)
fix X Y Z
assume
A: "X \<subseteq> range p" and
B: "Y \<subseteq> range q" and
C: "Z \<subseteq> - range p" and
D: "Z \<subseteq> - range q" and
E: "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` X) \<in> failures P" and
F: "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` Y) \<in> failures Q"
show
"(x \<in> range p \<or> x \<in> range q) \<and>
(map (inv p) [x\<leftarrow>xs @ [x]. x \<in> range p], {}) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs @ [x]. x \<in> range q], {}) \<in> failures Q \<or>
xs @ [x] \<in> con_comp_divergences P Q p q \<or>
(\<exists>X' Y' Z'.
insert x (X \<union> Y \<union> Z) = X' \<union> Y' \<union> Z' \<and>
X' \<subseteq> range p \<and>
Y' \<subseteq> range q \<and>
Z' \<subseteq> - range p \<and>
Z' \<subseteq> - range q \<and>
(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` X') \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` Y') \<in> failures Q) \<or>
xs \<in> con_comp_divergences P Q p q"
(is "_ \<or> _ \<or> ?A \<or> _")
proof (cases "x \<in> range p", case_tac [!] "x \<in> range q", simp_all)
assume
G: "x \<in> range p" and
H: "x \<in> range q"
show
"(map (inv p) [x\<leftarrow>xs. x \<in> range p] @ [inv p x], {}) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs. x \<in> range q] @ [inv q x], {}) \<in> failures Q \<or>
xs @ [x] \<in> con_comp_divergences P Q p q \<or>
?A \<or>
xs \<in> con_comp_divergences P Q p q"
(is "?B \<or> _")
proof (cases ?B, simp_all del: disj_not1, erule disjE)
assume
I: "(map (inv p) [x\<leftarrow>xs. x \<in> range p] @ [inv p x], {}) \<notin> failures P"
have ?A
proof (rule_tac x = "insert x X" in exI, rule_tac x = "Y" in exI,
rule_tac x = "Z" in exI, (subst conj_assoc [symmetric])+, (rule conjI)+)
show "insert x (X \<union> Y \<union> Z) = insert x X \<union> Y \<union> Z"
by simp
next
show "insert x X \<subseteq> range p"
using A and G by simp
next
show "Y \<subseteq> range q"
using B .
next
show "Z \<subseteq> - range p"
using C .
next
show "Z \<subseteq> - range q"
using D .
next
have
"(map (inv p) [x\<leftarrow>xs. x \<in> range p] @ [inv p x], {})
\<in> failures P \<or>
(map (inv p) [x\<leftarrow>xs. x \<in> range p], insert (inv p x) (inv p ` X))
\<in> failures P"
using E by (rule process_rule_4)
thus "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` insert x X) \<in> failures P"
using I by simp
next
show "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` Y) \<in> failures Q"
using F .
qed
thus ?thesis
by simp
next
assume
I: "(map (inv q) [x\<leftarrow>xs. x \<in> range q] @ [inv q x], {}) \<notin> failures Q"
have ?A
proof (rule_tac x = "X" in exI, rule_tac x = "insert x Y" in exI,
rule_tac x = "Z" in exI, (subst conj_assoc [symmetric])+, (rule conjI)+)
show "insert x (X \<union> Y \<union> Z) = X \<union> insert x Y \<union> Z"
by simp
next
show "X \<subseteq> range p"
using A .
next
show "insert x Y \<subseteq> range q"
using B and H by simp
next
show "Z \<subseteq> - range p"
using C .
next
show "Z \<subseteq> - range q"
using D .
next
show "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` X) \<in> failures P"
using E .
next
have
"(map (inv q) [x\<leftarrow>xs. x \<in> range q] @ [inv q x], {})
\<in> failures Q \<or>
(map (inv q) [x\<leftarrow>xs. x \<in> range q], insert (inv q x) (inv q ` Y))
\<in> failures Q"
using F by (rule process_rule_4)
thus "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` insert x Y) \<in> failures Q"
using I by simp
qed
thus ?thesis
by simp
qed
next
assume G: "x \<in> range p"
show
"(map (inv p) [x\<leftarrow>xs. x \<in> range p] @ [inv p x], {}) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs. x \<in> range q], {}) \<in> failures Q \<or>
xs @ [x] \<in> con_comp_divergences P Q p q \<or>
?A \<or>
xs \<in> con_comp_divergences P Q p q"
proof (cases "(map (inv p) [x\<leftarrow>xs. x \<in> range p] @ [inv p x], {})
\<in> failures P")
case True
moreover have "{} \<subseteq> inv q ` Y" ..
with F have "(map (inv q) [x\<leftarrow>xs. x \<in> range q], {}) \<in> failures Q"
by (rule process_rule_3)
ultimately show ?thesis
by simp
next
case False
have ?A
proof (rule_tac x = "insert x X" in exI, rule_tac x = "Y" in exI,
rule_tac x = "Z" in exI, (subst conj_assoc [symmetric])+, (rule conjI)+)
show "insert x (X \<union> Y \<union> Z) = insert x X \<union> Y \<union> Z"
by simp
next
show "insert x X \<subseteq> range p"
using A and G by simp
next
show "Y \<subseteq> range q"
using B .
next
show "Z \<subseteq> - range p"
using C .
next
show "Z \<subseteq> - range q"
using D .
next
have
"(map (inv p) [x\<leftarrow>xs. x \<in> range p] @ [inv p x], {})
\<in> failures P \<or>
(map (inv p) [x\<leftarrow>xs. x \<in> range p], insert (inv p x) (inv p ` X))
\<in> failures P"
using E by (rule process_rule_4)
thus "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` insert x X) \<in> failures P"
using False by simp
next
show "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` Y) \<in> failures Q"
using F .
qed
thus ?thesis
by simp
qed
next
assume G: "x \<in> range q"
show
"(map (inv p) [x\<leftarrow>xs. x \<in> range p], {}) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs. x \<in> range q] @ [inv q x], {}) \<in> failures Q \<or>
xs @ [x] \<in> con_comp_divergences P Q p q \<or>
?A \<or>
xs \<in> con_comp_divergences P Q p q"
proof (cases "(map (inv q) [x\<leftarrow>xs. x \<in> range q] @ [inv q x], {})
\<in> failures Q")
case True
moreover have "{} \<subseteq> inv p ` X" ..
with E have "(map (inv p) [x\<leftarrow>xs. x \<in> range p], {}) \<in> failures P"
by (rule process_rule_3)
ultimately show ?thesis
by simp
next
case False
have ?A
proof (rule_tac x = "X" in exI, rule_tac x = "insert x Y" in exI,
rule_tac x = "Z" in exI, (subst conj_assoc [symmetric])+, (rule conjI)+)
show "insert x (X \<union> Y \<union> Z) = X \<union> insert x Y \<union> Z"
by simp
next
show "X \<subseteq> range p"
using A .
next
show "insert x Y \<subseteq> range q"
using B and G by simp
next
show "Z \<subseteq> - range p"
using C .
next
show "Z \<subseteq> - range q"
using D .
next
show "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` X) \<in> failures P"
using E .
next
have
"(map (inv q) [x\<leftarrow>xs. x \<in> range q] @ [inv q x], {})
\<in> failures Q \<or>
(map (inv q) [x\<leftarrow>xs. x \<in> range q], insert (inv q x) (inv q ` Y))
\<in> failures Q"
using F by (rule process_rule_4)
thus "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` insert x Y) \<in> failures Q"
using False by simp
qed
thus ?thesis
by simp
qed
next
assume
G: "x \<notin> range p" and
H: "x \<notin> range q"
have ?A
proof (rule_tac x = "X" in exI, rule_tac x = "Y" in exI,
rule_tac x = "insert x Z" in exI, (subst conj_assoc [symmetric])+,
(rule conjI)+)
show "insert x (X \<union> Y \<union> Z) = X \<union> Y \<union> insert x Z"
by simp
next
show "X \<subseteq> range p"
using A .
next
show "Y \<subseteq> range q"
using B .
next
show "insert x Z \<subseteq> - range p"
using C and G by simp
next
show "insert x Z \<subseteq> - range q"
using D and H by simp
next
show "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` X) \<in> failures P"
using E .
next
show "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` Y) \<in> failures Q"
using F .
qed
thus
"xs @ [x] \<in> con_comp_divergences P Q p q \<or>
?A \<or>
xs \<in> con_comp_divergences P Q p q"
by simp
qed
qed
lemma con_comp_prop_5:
"xs \<in> con_comp_divergences P Q p q \<Longrightarrow>
xs @ [x] \<in> con_comp_divergences P Q p q"
proof (simp add: con_comp_divergences_def, erule exE, (erule conjE)+, erule exE)
fix xs' ys'
assume
A: "set xs' \<subseteq> range p \<union> range q" and
B: "map (inv p) [x\<leftarrow>xs'. x \<in> range p] \<in> divergences P" and
C: "map (inv q) [x\<leftarrow>xs'. x \<in> range q] \<in> divergences Q" and
D: "xs = xs' @ ys'"
show "\<exists>xs'.
(\<exists>ys'. xs @ [x] = xs' @ ys') \<and>
set xs' \<subseteq> range p \<union> range q \<and>
map (inv p) [x\<leftarrow>xs'. x \<in> range p] \<in> divergences P \<and>
map (inv q) [x\<leftarrow>xs'. x \<in> range q] \<in> divergences Q"
proof (rule_tac x = xs' in exI, simp_all add: A B C)
qed (rule_tac x = "ys' @ [x]" in exI, simp add: D)
qed
lemma con_comp_prop_6:
"xs \<in> con_comp_divergences P Q p q \<Longrightarrow>
(xs, X) \<in> con_comp_failures P Q p q"
by (simp add: con_comp_failures_def)
lemma con_comp_rep:
"Rep_process (P \<parallel> Q <p, q>) =
(con_comp_failures P Q p q, con_comp_divergences P Q p q)"
(is "_ = ?X")
proof (subst con_comp_def, rule Abs_process_inverse, simp add: process_set_def,
(subst conj_assoc [symmetric])+, (rule conjI)+)
show "process_prop_1 ?X"
proof (simp add: process_prop_1_def)
qed (rule con_comp_prop_1)
next
show "process_prop_2 ?X"
proof (simp add: process_prop_2_def del: all_simps, (rule allI)+, rule impI)
qed (rule con_comp_prop_2)
next
show "process_prop_3 ?X"
proof (simp add: process_prop_3_def del: all_simps, (rule allI)+, rule impI,
erule conjE)
qed (rule con_comp_prop_3)
next
show "process_prop_4 ?X"
proof (simp add: process_prop_4_def, (rule allI)+, rule impI)
qed (rule con_comp_prop_4)
next
show "process_prop_5 ?X"
proof (simp add: process_prop_5_def, rule allI, rule impI, rule allI)
qed (rule con_comp_prop_5)
next
show "process_prop_6 ?X"
proof (simp add: process_prop_6_def, rule allI, rule impI, rule allI)
qed (rule con_comp_prop_6)
qed
text \<open>
\null
Here below, the previous result is applied to derive useful expressions for the outputs of the
functions returning the elements of a process, as defined in \cite{R1} and \cite{R2}, when acting on
the concurrent composition of a pair of processes.
\null
\<close>
lemma con_comp_failures:
"failures (P \<parallel> Q <p, q>) = con_comp_failures P Q p q"
by (simp add: failures_def con_comp_rep)
lemma con_comp_divergences:
"divergences (P \<parallel> Q <p, q>) = con_comp_divergences P Q p q"
by (simp add: divergences_def con_comp_rep)
lemma con_comp_futures:
"futures (P \<parallel> Q <p, q>) xs =
{(ys, Y). (xs @ ys, Y) \<in> con_comp_failures P Q p q}"
by (simp add: futures_def con_comp_failures)
lemma con_comp_traces:
"traces (P \<parallel> Q <p, q>) = Domain (con_comp_failures P Q p q)"
by (simp add: traces_def con_comp_failures)
lemma con_comp_refusals:
"refusals (P \<parallel> Q <p, q>) xs \<equiv> con_comp_failures P Q p q `` {xs}"
by (simp add: refusals_def con_comp_failures)
lemma con_comp_next_events:
"next_events (P \<parallel> Q <p, q>) xs =
{x. xs @ [x] \<in> Domain (con_comp_failures P Q p q)}"
by (simp add: next_events_def con_comp_traces)
text \<open>
\null
In what follows, three lemmas are proven. The first one, whose proof makes use of the axiom of
choice, establishes an additional property required for the above definition of concurrent
composition to be correct, namely that for any two processes whose refusals are closed under set
union, their concurrent composition still be such, which is what is expected for any process of
practical significance (cf. \cite{R2}). The other two lemmas are auxiliary properties of concurrent
composition used in the proof of the target security conservation theorem.
\null
\<close>
lemma con_comp_ref_union_closed:
assumes
A: "ref_union_closed P" and
B: "ref_union_closed Q"
shows "ref_union_closed (P \<parallel> Q <p, q>)"
proof (simp add: ref_union_closed_def con_comp_failures con_comp_failures_def
con_comp_divergences_def del: SUP_identity_eq cong: SUP_cong_simp, (rule allI)+, (rule impI)+,
erule exE, rule disjI1)
fix xs A X
assume "\<forall>X \<in> A. \<exists>R S T.
X = R \<union> S \<union> T \<and>
set xs \<subseteq> range p \<union> range q \<and>
R \<subseteq> range p \<and>
S \<subseteq> range q \<and>
T \<subseteq> - range p \<and>
T \<subseteq> - range q \<and>
(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` R) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` S) \<in> failures Q"
(is "\<forall>X \<in> A. \<exists>R S T. ?F X R S T")
hence "\<exists>r. \<forall>X \<in> A. \<exists>S T. ?F X (r X) S T"
by (rule bchoice)
then obtain r where "\<forall>X \<in> A. \<exists>S T. ?F X (r X) S T" ..
hence "\<exists>s. \<forall>X \<in> A. \<exists>T. ?F X (r X) (s X) T"
by (rule bchoice)
then obtain s where "\<forall>X \<in> A. \<exists>T. ?F X (r X) (s X) T" ..
hence "\<exists>t. \<forall>X \<in> A. ?F X (r X) (s X) (t X)"
by (rule bchoice)
then obtain t where C: "\<forall>X \<in> A. ?F X (r X) (s X) (t X)" ..
assume D: "X \<in> A"
show "\<exists>R S T. ?F (\<Union>X \<in> A. X) R S T"
proof (rule_tac x = "\<Union>X \<in> A. r X" in exI, rule_tac x = "\<Union>X \<in> A. s X" in exI,
rule_tac x = "\<Union>X \<in> A. t X" in exI, (subst conj_assoc [symmetric])+,
(rule conjI)+)
show "(\<Union>X \<in> A. X) = (\<Union>X \<in> A. r X) \<union> (\<Union>X \<in> A. s X) \<union> (\<Union>X \<in> A. t X)"
proof (simp add: set_eq_iff, rule allI, rule iffI, erule_tac [2] disjE,
erule_tac [3] disjE, erule_tac [!] bexE)
fix x X
have "\<forall>X \<in> A. X = r X \<union> s X \<union> t X"
using C by simp
moreover assume E: "X \<in> A"
ultimately have "X = r X \<union> s X \<union> t X" ..
moreover assume "x \<in> X"
ultimately have "x \<in> r X \<or> x \<in> s X \<or> x \<in> t X"
by blast
hence "\<exists>X \<in> A. x \<in> r X \<or> x \<in> s X \<or> x \<in> t X"
using E ..
thus "(\<exists>X \<in> A. x \<in> r X) \<or> (\<exists>X \<in> A. x \<in> s X) \<or> (\<exists>X \<in> A. x \<in> t X)"
by blast
next
fix x X
have "\<forall>X \<in> A. X = r X \<union> s X \<union> t X"
using C by simp
moreover assume E: "X \<in> A"
ultimately have "X = r X \<union> s X \<union> t X" ..
moreover assume "x \<in> r X"
ultimately have "x \<in> X"
by blast
thus "\<exists>X \<in> A. x \<in> X"
using E ..
next
fix x X
have "\<forall>X \<in> A. X = r X \<union> s X \<union> t X"
using C by simp
moreover assume E: "X \<in> A"
ultimately have "X = r X \<union> s X \<union> t X" ..
moreover assume "x \<in> s X"
ultimately have "x \<in> X"
by blast
thus "\<exists>X \<in> A. x \<in> X"
using E ..
next
fix x X
have "\<forall>X \<in> A. X = r X \<union> s X \<union> t X"
using C by simp
moreover assume E: "X \<in> A"
ultimately have "X = r X \<union> s X \<union> t X" ..
moreover assume "x \<in> t X"
ultimately have "x \<in> X"
by blast
thus "\<exists>X \<in> A. x \<in> X"
using E ..
qed
next
have "\<forall>X \<in> A. set xs \<subseteq> range p \<union> range q"
using C by simp
thus "set xs \<subseteq> range p \<union> range q"
using D ..
next
show "(\<Union>X \<in> A. r X) \<subseteq> range p"
proof (rule subsetI, erule UN_E)
fix x X
have "\<forall>X \<in> A. r X \<subseteq> range p"
using C by simp
moreover assume "X \<in> A"
ultimately have "r X \<subseteq> range p" ..
moreover assume "x \<in> r X"
ultimately show "x \<in> range p" ..
qed
next
show "(\<Union>X \<in> A. s X) \<subseteq> range q"
proof (rule subsetI, erule UN_E)
fix x X
have "\<forall>X \<in> A. s X \<subseteq> range q"
using C by simp
moreover assume "X \<in> A"
ultimately have "s X \<subseteq> range q" ..
moreover assume "x \<in> s X"
ultimately show "x \<in> range q" ..
qed
next
show "(\<Union>X \<in> A. t X) \<subseteq> - range p"
proof (rule subsetI, erule UN_E)
fix x X
have "\<forall>X \<in> A. t X \<subseteq> - range p"
using C by simp
moreover assume "X \<in> A"
ultimately have "t X \<subseteq> - range p" ..
moreover assume "x \<in> t X"
ultimately show "x \<in> - range p" ..
qed
next
show "(\<Union>X \<in> A. t X) \<subseteq> - range q"
proof (rule subsetI, erule UN_E)
fix x X
have "\<forall>X \<in> A. t X \<subseteq> - range q"
using C by simp
moreover assume "X \<in> A"
ultimately have "t X \<subseteq> - range q" ..
moreover assume "x \<in> t X"
ultimately show "x \<in> - range q" ..
qed
next
let ?A' = "{inv p ` X | X. X \<in> r ` A}"
have
"(\<exists>X. X \<in> ?A') \<longrightarrow>
(\<forall>X \<in> ?A'. (map (inv p) [x\<leftarrow>xs. x \<in> range p], X) \<in> failures P) \<longrightarrow>
(map (inv p) [x\<leftarrow>xs. x \<in> range p], \<Union>X \<in> ?A'. X) \<in> failures P"
using A by (simp add: ref_union_closed_def)
moreover have "\<exists>X. X \<in> ?A'"
using D by blast
ultimately have
"(\<forall>X \<in> ?A'. (map (inv p) [x\<leftarrow>xs. x \<in> range p], X) \<in> failures P) \<longrightarrow>
(map (inv p) [x\<leftarrow>xs. x \<in> range p], \<Union>X \<in> ?A'. X) \<in> failures P" ..
moreover have
"\<forall>X \<in> ?A'. (map (inv p) [x\<leftarrow>xs. x \<in> range p], X) \<in> failures P"
proof (rule ballI, simp, erule exE, erule conjE)
fix R R'
assume "R \<in> r ` A"
hence "\<exists>X \<in> A. R = r X"
by (simp add: image_iff)
then obtain X where E: "X \<in> A" and F: "R = r X" ..
have "\<forall>X \<in> A. (map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` r X) \<in> failures P"
using C by simp
hence "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` r X) \<in> failures P"
using E ..
moreover assume "R' = inv p ` R"
ultimately show "(map (inv p) [x\<leftarrow>xs. x \<in> range p], R') \<in> failures P"
using F by simp
qed
ultimately have "(map (inv p) [x\<leftarrow>xs. x \<in> range p], \<Union>X \<in> ?A'. X)
\<in> failures P" ..
moreover have "(\<Union>X \<in> ?A'. X) = inv p ` (\<Union>X \<in> A. r X)"
proof (subst set_eq_iff, simp, rule allI, rule iffI, (erule exE, erule conjE)+)
fix a R R'
assume "R \<in> r ` A"
hence "\<exists>X \<in> A. R = r X"
by (simp add: image_iff)
then obtain X where E: "X \<in> A" and F: "R = r X" ..
assume "a \<in> R'" and "R' = inv p ` R"
hence "a \<in> inv p ` r X"
using F by simp
hence "\<exists>x \<in> r X. a = inv p x"
by (simp add: image_iff)
then obtain x where G: "x \<in> r X" and H: "a = inv p x" ..
have "x \<in> (\<Union>X \<in> A. r X)"
using E and G by (rule UN_I)
with H have "\<exists>x \<in> (\<Union>X \<in> A. r X). a = inv p x" ..
thus "a \<in> inv p ` (\<Union>X \<in> A. r X)"
by (simp add: image_iff)
next
fix a
assume "a \<in> inv p ` (\<Union>X \<in> A. r X)"
hence "\<exists>x \<in> (\<Union>X \<in> A. r X). a = inv p x"
by (simp add: image_iff)
then obtain x where E: "x \<in> (\<Union>X \<in> A. r X)" and F: "a = inv p x" ..
obtain X where G: "X \<in> A" and H: "x \<in> r X" using E ..
show "\<exists>R'. (\<exists>R. R' = inv p ` R \<and> R \<in> r ` A) \<and> a \<in> R'"
proof (rule_tac x = "inv p ` r X" in exI, rule conjI,
rule_tac x = "r X" in exI)
qed (rule_tac [2] image_eqI, simp add: G, simp add: F, simp add: H)
qed
ultimately show "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` (\<Union>X \<in> A. r X))
\<in> failures P"
by simp
next
let ?A' = "{inv q ` X | X. X \<in> s ` A}"
have
"(\<exists>X. X \<in> ?A') \<longrightarrow>
(\<forall>X \<in> ?A'. (map (inv q) [x\<leftarrow>xs. x \<in> range q], X) \<in> failures Q) \<longrightarrow>
(map (inv q) [x\<leftarrow>xs. x \<in> range q], \<Union>X \<in> ?A'. X) \<in> failures Q"
using B by (simp add: ref_union_closed_def)
moreover have "\<exists>X. X \<in> ?A'"
using D by blast
ultimately have
"(\<forall>X \<in> ?A'. (map (inv q) [x\<leftarrow>xs. x \<in> range q], X) \<in> failures Q) \<longrightarrow>
(map (inv q) [x\<leftarrow>xs. x \<in> range q], \<Union>X \<in> ?A'. X) \<in> failures Q" ..
moreover have
"\<forall>X \<in> ?A'. (map (inv q) [x\<leftarrow>xs. x \<in> range q], X) \<in> failures Q"
proof (rule ballI, simp, erule exE, erule conjE)
fix S S'
assume "S \<in> s ` A"
hence "\<exists>X \<in> A. S = s X"
by (simp add: image_iff)
then obtain X where E: "X \<in> A" and F: "S = s X" ..
have "\<forall>X \<in> A. (map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` s X) \<in> failures Q"
using C by simp
hence "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` s X) \<in> failures Q"
using E ..
moreover assume "S' = inv q ` S"
ultimately show "(map (inv q) [x\<leftarrow>xs. x \<in> range q], S') \<in> failures Q"
using F by simp
qed
ultimately have "(map (inv q) [x\<leftarrow>xs. x \<in> range q], \<Union>X \<in> ?A'. X)
\<in> failures Q" ..
moreover have "(\<Union>X \<in> ?A'. X) = inv q ` (\<Union>X \<in> A. s X)"
proof (subst set_eq_iff, simp, rule allI, rule iffI, (erule exE, erule conjE)+)
fix b S S'
assume "S \<in> s ` A"
hence "\<exists>X \<in> A. S = s X"
by (simp add: image_iff)
then obtain X where E: "X \<in> A" and F: "S = s X" ..
assume "b \<in> S'" and "S' = inv q ` S"
hence "b \<in> inv q ` s X"
using F by simp
hence "\<exists>x \<in> s X. b = inv q x"
by (simp add: image_iff)
then obtain x where G: "x \<in> s X" and H: "b = inv q x" ..
have "x \<in> (\<Union>X \<in> A. s X)"
using E and G by (rule UN_I)
with H have "\<exists>x \<in> (\<Union>X \<in> A. s X). b = inv q x" ..
thus "b \<in> inv q ` (\<Union>X \<in> A. s X)"
by (simp add: image_iff)
next
fix b
assume "b \<in> inv q ` (\<Union>X \<in> A. s X)"
hence "\<exists>x \<in> (\<Union>X \<in> A. s X). b = inv q x"
by (simp add: image_iff)
then obtain x where E: "x \<in> (\<Union>X \<in> A. s X)" and F: "b = inv q x" ..
obtain X where G: "X \<in> A" and H: "x \<in> s X" using E ..
show "\<exists>S'. (\<exists>S. S' = inv q ` S \<and> S \<in> s ` A) \<and> b \<in> S'"
proof (rule_tac x = "inv q ` s X" in exI, rule conjI,
rule_tac x = "s X" in exI)
qed (rule_tac [2] image_eqI, simp add: G, simp add: F, simp add: H)
qed
ultimately show "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` (\<Union>X \<in> A. s X))
\<in> failures Q"
by simp
qed
qed
lemma con_comp_failures_traces:
"(xs, X) \<in> con_comp_failures P Q p q \<Longrightarrow>
map (inv p) [x\<leftarrow>xs. x \<in> range p] \<in> traces P \<and>
map (inv q) [x\<leftarrow>xs. x \<in> range q] \<in> traces Q"
proof (simp add: con_comp_failures_def con_comp_divergences_def, erule disjE,
(erule exE)+, (erule conjE)+, erule_tac [2] exE, (erule_tac [2] conjE)+,
erule_tac [2] exE)
fix X Y
assume "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` X) \<in> failures P"
hence "map (inv p) [x\<leftarrow>xs. x \<in> range p] \<in> traces P"
by (rule failures_traces)
moreover assume "(map (inv q) [x\<leftarrow>xs. x \<in> range q], inv q ` Y) \<in> failures Q"
hence "map (inv q) [x\<leftarrow>xs. x \<in> range q] \<in> traces Q"
by (rule failures_traces)
ultimately show ?thesis ..
next
fix vs ws
assume A: "xs = vs @ ws"
assume "map (inv p) [x\<leftarrow>vs. x \<in> range p] \<in> divergences P"
hence "map (inv p) [x\<leftarrow>vs. x \<in> range p] @ map (inv p) [x\<leftarrow>ws. x \<in> range p]
\<in> divergences P"
by (rule process_rule_5_general)
hence "map (inv p) [x\<leftarrow>xs. x \<in> range p] \<in> divergences P"
using A by simp
hence "(map (inv p) [x\<leftarrow>xs. x \<in> range p], {}) \<in> failures P"
by (rule process_rule_6)
hence "map (inv p) [x\<leftarrow>xs. x \<in> range p] \<in> traces P"
by (rule failures_traces)
moreover assume "map (inv q) [x\<leftarrow>vs. x \<in> range q] \<in> divergences Q"
hence "map (inv q) [x\<leftarrow>vs. x \<in> range q] @ map (inv q) [x\<leftarrow>ws. x \<in> range q]
\<in> divergences Q"
by (rule process_rule_5_general)
hence "map (inv q) [x\<leftarrow>xs. x \<in> range q] \<in> divergences Q"
using A by simp
hence "(map (inv q) [x\<leftarrow>xs. x \<in> range q], {}) \<in> failures Q"
by (rule process_rule_6)
hence "map (inv q) [x\<leftarrow>xs. x \<in> range q] \<in> traces Q"
by (rule failures_traces)
ultimately show ?thesis ..
qed
lemma con_comp_failures_divergences:
"(xs @ y # ys, Y) \<in> con_comp_failures P Q p q \<Longrightarrow>
y \<notin> range p \<Longrightarrow>
y \<notin> range q \<Longrightarrow>
\<exists>xs'.
(\<exists>ys'. xs @ zs = xs' @ ys') \<and>
set xs' \<subseteq> range p \<union> range q \<and>
map (inv p) [x\<leftarrow>xs'. x \<in> range p] \<in> divergences P \<and>
map (inv q) [x\<leftarrow>xs'. x \<in> range q] \<in> divergences Q"
proof (simp add: con_comp_failures_def con_comp_divergences_def,
erule exE, (erule conjE)+, erule exE)
fix xs' ys'
assume
A: "y \<notin> range p" and
B: "y \<notin> range q" and
C: "set xs' \<subseteq> range p \<union> range q" and
D: "map (inv p) [x\<leftarrow>xs'. x \<in> range p] \<in> divergences P" and
E: "map (inv q) [x\<leftarrow>xs'. x \<in> range q] \<in> divergences Q" and
F: "xs @ y # ys = xs' @ ys'"
have "length xs' \<le> length xs"
proof (rule ccontr)
assume "\<not> length xs' \<le> length xs"
moreover have "take (length xs') (xs @ [y] @ ys) =
take (length xs') (xs @ [y]) @ take (length xs' - Suc (length xs)) ys"
(is "_ = _ @ ?vs")
by (simp only: take_append, simp)
ultimately have "take (length xs') (xs @ y # ys) = xs @ y # ?vs"
by simp
moreover have "take (length xs') (xs @ y # ys) =
take (length xs') (xs' @ ys')"
using F by simp
ultimately have "xs' = xs @ y # ?vs"
by simp
hence "set (xs @ y # ?vs) \<subseteq> range p \<union> range q"
using C by simp
hence "y \<in> range p \<union> range q"
by simp
thus False
using A and B by simp
qed
moreover have "xs @ zs =
take (length xs') (xs @ zs) @ drop (length xs') (xs @ zs)"
(is "_ = _ @ ?vs")
by (simp only: append_take_drop_id)
ultimately have "xs @ zs = take (length xs') (xs @ y # ys) @ ?vs"
by simp
moreover have "take (length xs') (xs @ y # ys) =
take (length xs') (xs' @ ys')"
using F by simp
ultimately have G: "xs @ zs = xs' @ ?vs"
by (simp del: take_append, simp)
show ?thesis
proof (rule_tac x = xs' in exI, rule conjI, rule_tac x = ?vs in exI)
qed (subst G, simp_all add: C D E)
qed
text \<open>
\null
In order to prove that CSP noninterference security is conserved under concurrent composition, the
first issue to be solved is to identify the noninterference policy @{term I'} and the event-domain
map @{term D'} with respect to which the output process is secure.
If the events of the input processes corresponding to those of the output process contained in
@{term "range p \<inter> range q"} were mapped by the respective event-domain maps @{term D}, @{term E}
into distinct security domains, there would be no criterion for determining the domains of the
aforesaid events of the output process, due to the equivalence of the input processes ensuing from
the commutative property of concurrent composition. Therefore, @{term D} and @{term E} must map the
events of the input processes into security domains of the same type @{typ 'd}, and for each
@{term x} in @{term "range p \<inter> range q"}, @{term D} and @{term E} must map the events of the input
processes corresponding to @{term x} into the same domain. This requirement is formalized here below
by means of predicate \<open>consistent_maps\<close>.
Similarly, if distinct noninterference policies applied to the input processes, there would exist
some ordered pair of security domains included in one of the policies, but not in the other one.
Thus, again, there would be no criterion for determining the inclusion of such a pair of domains in
the policy @{term I'} applying to the output process. As a result, the input processes are required
to enforce the same noninterference policy @{term I}, so that for any two domains @{term d},
@{term e} of type @{typ 'd}, the ordered pair comprised of the corresponding security domains for
the output process will be included in @{term I'} just in case @{term "(d, e) \<in> I"}.
However, in case @{term "- (range p \<union> range q) \<noteq> {}"}, the event-domain map @{term D'} for the
output process must assign a security domain to the fake events in @{term "- (range p \<union> range q)"}
as well. Since such events lack any meaning, they may all be mapped to the same security domain,
distinct from the domains of the meaningful events in @{term "range p \<union> range q"}. A simple way to
do this is to identify the type of the security domains for the output process with
@{typ "'d option"}. Then, for any meaningful event @{term x}, @{term D'} will assign @{term x} to
domain @{term "Some d"}, where @{term d} is the domain of the events of the input processes mapped
to @{term x}, whereas @{term "D' y = None"} for any fake event @{term y}. Such an event-domain map,
denoted using notation @{term "con_comp_map D E p q"}, is defined here below.
Therefore, for any two security domains @{term "Some d"}, @{term "Some e"} for the output process,
the above considerations about policy @{term I'} entail that @{term "(Some d, Some e) \<in> I'"} just in
case @{term "(d, e) \<in> I"}. Furthermore, since fake events may only occur in divergent traces, which
are extensions of divergences of the input processes comprised of meaningful events, @{term I'} must
allow the security domain @{term None} of fake events to be affected by any meaningful domain
matching pattern \<open>Some _\<close>. Such a noninterference policy, denoted using notation
@{term "con_comp_pol I"}, is defined here below. Observe that @{term "con_comp_pol I"} keeps being
reflexive or transitive if @{term I} is.
\null
\<close>
definition con_comp_pol ::
"('d \<times> 'd) set \<Rightarrow> ('d option \<times> 'd option) set" where
"con_comp_pol I \<equiv>
{(Some d, Some e) | d e. (d, e) \<in> I} \<union> {(u, v). v = None}"
function con_comp_map ::
"('a \<Rightarrow> 'd) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'c \<Rightarrow> 'd option" where
"x \<in> range p \<Longrightarrow>
con_comp_map D E p q x = Some (D (inv p x))" |
"x \<notin> range p \<Longrightarrow> x \<in> range q \<Longrightarrow>
con_comp_map D E p q x = Some (E (inv q x))" |
"x \<notin> range p \<Longrightarrow> x \<notin> range q \<Longrightarrow>
con_comp_map D E p q x = None"
by (atomize_elim, simp_all add: split_paired_all, blast)
termination by lexicographic_order
definition consistent_maps ::
"('a \<Rightarrow> 'd) \<Rightarrow> ('b \<Rightarrow> 'd) \<Rightarrow> ('a \<Rightarrow> 'c) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> bool" where
"consistent_maps D E p q \<equiv>
\<forall>x \<in> range p \<inter> range q. D (inv p x) = E (inv q x)"
subsection "Auxiliary intransitive purge functions"
text \<open>
Let @{term I} be a noninterference policy, @{term D} an event-domain map, @{term U} a domain set,
and @{term "xs = x # xs'"} an event list. Suppose to take event @{term x} just in case it satisfies
predicate @{term P}, to append @{term xs'} to the resulting list (matching either @{term "[x]"} or
@{term "[]"}), and then to compute the intransitive purge of the resulting list with domain set
@{term U}. If recursion with respect to the input list is added, replacing @{term xs'} with the list
produced by the same algorithm using @{term xs'} as input list and @{term "sinks_aux I D U [x]"} as
domain set, the final result matches that obtained by applying filter @{term P} to the intransitive
purge of @{term xs} with domain set @{term U}. In fact, in each recursive step, the processed item
of the input list is retained in the output list just in case it passes filter @{term P} and may be
affected neither by the domains in @{term U}, nor by the domains of the previous items affected by
some domain in @{term U}.
Here below is the formal definition of such purge function, named \<open>ipurge_tr_aux_foldr\<close> as its
action resembles that of function @{term foldr}.
\null
\<close>
primrec ipurge_tr_aux_foldr ::
"('d \<times> 'd) set \<Rightarrow> ('a \<Rightarrow> 'd) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'd set \<Rightarrow> 'a list \<Rightarrow> 'a list"
where
"ipurge_tr_aux_foldr I D P U [] = []" |
"ipurge_tr_aux_foldr I D P U (x # xs) = ipurge_tr_aux I D U
((if P x then [x] else []) @
ipurge_tr_aux_foldr I D P (sinks_aux I D U [x]) xs)"
text \<open>
\null
Likewise, given @{term I}, @{term D}, @{term U}, @{term "xs = x # xs'"}, and an event set @{term X},
suppose to take @{term x} just in case it satisfies predicate @{term P}, to append
@{term "ipurge_tr_aux_foldr I D P (sinks_aux I D U [x]) xs'"} to the resulting list (matching either
@{term "[x]"} or @{term "[]"}), and then to compute the intransitive purge of @{term X} using the
resulting list as input list and @{term U} as domain set. If recursion with respect to the input
list is added, replacing @{term X} with the set produced by the same algorithm using @{term xs'} as
input list, @{term X} as input set, and @{term "sinks_aux I D U [x]"} as domain set, the final
result matches the intransitive purge of @{term X} with input list @{term xs} and domain set
@{term U}. In fact, each recursive step is such as to remove from @{term X} any event that may be
affected either by the domains in @{term U}, or by the domains of the items of @{term xs} preceding
the processed one which are affected by some domain in @{term U}.
From the above considerations on function @{term ipurge_tr_aux_foldr}, it follows that the presence
of list @{term "ipurge_tr_aux_foldr I D P (sinks_aux I D U [x]) xs'"} has no impact on the final
result, because none of its items may be affected by the domains in @{term U}.
Here below is the formal definition of such purge function, named \<open>ipurge_ref_aux_foldr\<close>,
which at first glance just seems a uselessly complicate and inefficient way to compute the
intransitive purge of an event set.
\null
\<close>
primrec ipurge_ref_aux_foldr ::
"('d \<times> 'd) set \<Rightarrow> ('a \<Rightarrow> 'd) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'd set \<Rightarrow> 'a list \<Rightarrow> 'a set \<Rightarrow> 'a set"
where
"ipurge_ref_aux_foldr I D P U [] X = ipurge_ref_aux I D U [] X" |
"ipurge_ref_aux_foldr I D P U (x # xs) X = ipurge_ref_aux I D U
((if P x then [x] else []) @
ipurge_tr_aux_foldr I D P (sinks_aux I D U [x]) xs)
(ipurge_ref_aux_foldr I D P (sinks_aux I D U [x]) xs X)"
text \<open>
\null
The reason for the introduction of such intransitive purge functions is that the recursive equations
contained in their definitions, along with lemma @{thm [source] ipurge_tr_ref_aux_failures_general},
enable to prove by induction on list @{term ys}, assuming that process @{term P} be secure in
addition to further, minor premises, the following implication:
\null
@{term "(map (inv p) [x\<leftarrow>xs @ ys. x \<in> range p], inv p ` Y) \<in> failures P \<longrightarrow>
(map (inv p) [x\<leftarrow>xs. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr (con_comp_pol I) (con_comp_map D E p q)
(\<lambda>x. x \<in> range p) U ys),
inv p ` ipurge_ref_aux_foldr (con_comp_pol I) (con_comp_map D E p q)
(\<lambda>x. x \<in> range p) U ys Y) \<in> failures P"}
\null
In fact, for @{term "ys = y # ys'"}, the induction hypothesis entails that the consequent holds if
@{term xs}, @{term ys}, and @{term U} are replaced with @{term "xs @ [y]"}, @{term ys'}, and
@{term "sinks_aux (con_comp_pol I) (con_comp_map D E p q) U [y]"}, respectively. The proof can then
be accomplished by applying lemma @{thm [source] ipurge_tr_ref_aux_failures_general} to the
resulting future of trace @{term "map (inv p) [x\<leftarrow>xs. x \<in> range p]"}, moving functions
@{term ipurge_tr_aux} and @{term "ipurge_ref_aux"} into the arguments of @{term "map (inv p)"} and
@{term "(`) (inv p)"}, and using the recursive equations contained in the definitions of functions
@{term ipurge_tr_aux_foldr} and @{term ipurge_ref_aux_foldr}.
This property, along with the match of the outputs of functions @{term ipurge_tr_aux_foldr} and
@{term ipurge_ref_aux_foldr} with the filtered intransitive purge of the input event list and the
intransitive purge of the input event set, respectively, permits to solve the main proof obligations
arising from the demonstration of the target security conservation theorem.
Here below is the proof of the equivalence between function @{term ipurge_tr_aux_foldr} and the
filtered intransitive purge of an event list.
\null
\<close>
lemma ipurge_tr_aux_foldr_subset:
"U \<subseteq> V \<Longrightarrow>
ipurge_tr_aux I D U (ipurge_tr_aux_foldr I D P V xs) =
ipurge_tr_aux_foldr I D P V xs"
proof (induction xs, simp_all add: ipurge_tr_aux_union [symmetric])
qed (drule Un_absorb2, simp)
lemma ipurge_tr_aux_foldr_eq:
"[x\<leftarrow>ipurge_tr_aux I D U xs. P x] = ipurge_tr_aux_foldr I D P U xs"
proof (induction xs arbitrary: U, simp)
fix x xs U
assume
A: "\<And>U. [x\<leftarrow>ipurge_tr_aux I D U xs. P x] = ipurge_tr_aux_foldr I D P U xs"
show "[x\<leftarrow>ipurge_tr_aux I D U (x # xs). P x] =
ipurge_tr_aux_foldr I D P U (x # xs)"
proof (cases "\<exists>u \<in> U. (u, D x) \<in> I",
simp_all only: ipurge_tr_aux_foldr.simps ipurge_tr_aux_cons
sinks_aux_single_event if_True if_False)
case True
have B: "[x\<leftarrow>ipurge_tr_aux I D (insert (D x) U) xs. P x] =
ipurge_tr_aux_foldr I D P (insert (D x) U) xs"
using A .
show "[x\<leftarrow>ipurge_tr_aux I D (insert (D x) U) xs. P x] = ipurge_tr_aux I D U
((if P x then [x] else []) @ ipurge_tr_aux_foldr I D P (insert (D x) U) xs)"
proof (cases "P x", simp_all add: ipurge_tr_aux_cons True
del: con_comp_map.simps)
have "insert (D x) U \<subseteq> insert (D x) U" ..
hence "ipurge_tr_aux I D (insert (D x) U)
(ipurge_tr_aux_foldr I D P (insert (D x) U) xs) =
ipurge_tr_aux_foldr I D P (insert (D x) U) xs"
by (rule ipurge_tr_aux_foldr_subset)
thus "[x\<leftarrow>ipurge_tr_aux I D (insert (D x) U) xs. P x] =
ipurge_tr_aux I D (insert (D x) U)
(ipurge_tr_aux_foldr I D P (insert (D x) U) xs)"
using B by simp
next
have "U \<subseteq> insert (D x) U"
by (rule subset_insertI)
hence "ipurge_tr_aux I D U
(ipurge_tr_aux_foldr I D P (insert (D x) U) xs) =
ipurge_tr_aux_foldr I D P (insert (D x) U) xs"
by (rule ipurge_tr_aux_foldr_subset)
thus "[x\<leftarrow>ipurge_tr_aux I D (insert (D x) U) xs. P x] =
ipurge_tr_aux I D U
(ipurge_tr_aux_foldr I D P (insert (D x) U) xs)"
using B by simp
qed
next
case False
have B: "[x\<leftarrow>ipurge_tr_aux I D U xs. P x] = ipurge_tr_aux_foldr I D P U xs"
using A .
show "[x\<leftarrow>x # ipurge_tr_aux I D U xs. P x] = ipurge_tr_aux I D U
((if P x then [x] else []) @ ipurge_tr_aux_foldr I D P U xs)"
proof (cases "P x", simp_all add: ipurge_tr_aux_cons False
del: con_comp_map.simps)
have "U \<subseteq> U" ..
hence "ipurge_tr_aux I D U (ipurge_tr_aux_foldr I D P U xs) =
ipurge_tr_aux_foldr I D P U xs"
by (rule ipurge_tr_aux_foldr_subset)
thus "[x\<leftarrow>ipurge_tr_aux I D U xs. P x] =
ipurge_tr_aux I D U (ipurge_tr_aux_foldr I D P U xs)"
using B by simp
next
have "U \<subseteq> U" ..
hence "ipurge_tr_aux I D U (ipurge_tr_aux_foldr I D P U xs) =
ipurge_tr_aux_foldr I D P U xs"
by (rule ipurge_tr_aux_foldr_subset)
thus "[x\<leftarrow>ipurge_tr_aux I D U xs. P x] =
ipurge_tr_aux I D U (ipurge_tr_aux_foldr I D P U xs)"
using B by simp
qed
qed
qed
text \<open>
\null
Here below is the proof of the equivalence between function @{term ipurge_ref_aux_foldr} and the
intransitive purge of an event set.
\null
\<close>
lemma ipurge_tr_aux_foldr_sinks_aux [rule_format]:
"U \<subseteq> V \<longrightarrow> sinks_aux I D U (ipurge_tr_aux_foldr I D P V xs) = U"
proof (induction xs arbitrary: V, simp, rule impI)
fix x xs V
assume
A: "\<And>V. U \<subseteq> V \<longrightarrow> sinks_aux I D U (ipurge_tr_aux_foldr I D P V xs) = U" and
B: "U \<subseteq> V"
show "sinks_aux I D U (ipurge_tr_aux_foldr I D P V (x # xs)) = U"
proof (cases "P x", case_tac [!] "\<exists>v \<in> V. (v, D x) \<in> I",
simp_all (no_asm_simp) add: sinks_aux_cons ipurge_tr_aux_cons)
have "U \<subseteq> insert (D x) V \<longrightarrow>
sinks_aux I D U (ipurge_tr_aux_foldr I D P (insert (D x) V) xs) = U"
(is "_ \<longrightarrow> sinks_aux I D U ?ys = U")
using A .
moreover have "U \<subseteq> insert (D x) V"
using B by (rule subset_insertI2)
ultimately have "sinks_aux I D U ?ys = U" ..
moreover have "insert (D x) V \<subseteq> insert (D x) V" ..
hence "ipurge_tr_aux I D (insert (D x) V)
(ipurge_tr_aux_foldr I D P (insert (D x) V) xs) = ?ys"
(is "?zs = _")
by (rule ipurge_tr_aux_foldr_subset)
ultimately show "sinks_aux I D U ?zs = U"
by simp
next
assume C: "\<not> (\<exists>v \<in> V. (v, D x) \<in> I)"
have "\<not> (\<exists>u \<in> U. (u, D x) \<in> I)"
proof
assume "\<exists>u \<in> U. (u, D x) \<in> I"
then obtain u where D: "u \<in> U" and E: "(u, D x) \<in> I" ..
have "u \<in> V"
using B and D ..
with E have "\<exists>v \<in> V. (v, D x) \<in> I" ..
thus False
using C by contradiction
qed
thus
"((\<exists>v \<in> U. (v, D x) \<in> I) \<longrightarrow> sinks_aux I D (insert (D x) U)
(ipurge_tr_aux I D V (ipurge_tr_aux_foldr I D P V xs)) = U) \<and>
((\<forall>v \<in> U. (v, D x) \<notin> I) \<longrightarrow> sinks_aux I D U
(ipurge_tr_aux I D V (ipurge_tr_aux_foldr I D P V xs)) = U)"
proof simp
have "U \<subseteq> V \<longrightarrow> sinks_aux I D U (ipurge_tr_aux_foldr I D P V xs) = U"
(is "_ \<longrightarrow> sinks_aux I D U ?ys = U")
using A .
hence "sinks_aux I D U ?ys = U" using B ..
moreover have "V \<subseteq> V" ..
hence "ipurge_tr_aux I D V (ipurge_tr_aux_foldr I D P V xs) = ?ys"
(is "?zs = _")
by (rule ipurge_tr_aux_foldr_subset)
ultimately show "sinks_aux I D U ?zs = U"
by simp
qed
next
have "U \<subseteq> insert (D x) V \<longrightarrow>
sinks_aux I D U (ipurge_tr_aux_foldr I D P (insert (D x) V) xs) = U"
(is "_ \<longrightarrow> sinks_aux I D U ?ys = U")
using A .
moreover have "U \<subseteq> insert (D x) V"
using B by (rule subset_insertI2)
ultimately have "sinks_aux I D U ?ys = U" ..
moreover have "V \<subseteq> insert (D x) V"
by (rule subset_insertI)
hence "ipurge_tr_aux I D V
(ipurge_tr_aux_foldr I D P (insert (D x) V) xs) = ?ys"
(is "?zs = _")
by (rule ipurge_tr_aux_foldr_subset)
ultimately show "sinks_aux I D U ?zs = U"
by simp
next
have "U \<subseteq> V \<longrightarrow> sinks_aux I D U (ipurge_tr_aux_foldr I D P V xs) = U"
(is "_ \<longrightarrow> sinks_aux I D U ?ys = U")
using A .
hence "sinks_aux I D U ?ys = U" using B ..
moreover have "V \<subseteq> V" ..
hence "ipurge_tr_aux I D V (ipurge_tr_aux_foldr I D P V xs) = ?ys"
(is "?zs = _")
by (rule ipurge_tr_aux_foldr_subset)
ultimately show "sinks_aux I D U ?zs = U"
by simp
qed
qed
lemma ipurge_tr_aux_foldr_ref_aux:
assumes A: "U \<subseteq> V"
shows "ipurge_ref_aux I D U (ipurge_tr_aux_foldr I D P V xs) X =
ipurge_ref_aux I D U [] X"
by (simp add: ipurge_ref_aux_def ipurge_tr_aux_foldr_sinks_aux [OF A])
lemma ipurge_ref_aux_foldr_subset [rule_format]:
"sinks_aux I D U ys \<subseteq> V \<longrightarrow>
ipurge_ref_aux I D U ys (ipurge_ref_aux_foldr I D P V xs X) =
ipurge_ref_aux_foldr I D P V xs X"
proof (induction xs arbitrary: ys U V, rule_tac [!] impI,
simp add: ipurge_ref_aux_def, blast)
fix x xs ys U V
assume
A: "\<And>ys U V.
sinks_aux I D U ys \<subseteq> V \<longrightarrow>
ipurge_ref_aux I D U ys (ipurge_ref_aux_foldr I D P V xs X) =
ipurge_ref_aux_foldr I D P V xs X" and
B: "sinks_aux I D U ys \<subseteq> V"
show "ipurge_ref_aux I D U ys (ipurge_ref_aux_foldr I D P V (x # xs) X) =
ipurge_ref_aux_foldr I D P V (x # xs) X"
proof (cases "P x", simp_all add: ipurge_ref_aux_cons)
have C: "sinks_aux I D V [x] \<subseteq> sinks_aux I D V [x]" ..
show
"ipurge_ref_aux I D U ys (ipurge_ref_aux I D (sinks_aux I D V [x])
(ipurge_tr_aux_foldr I D P (sinks_aux I D V [x]) xs)
(ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X)) =
ipurge_ref_aux I D (sinks_aux I D V [x])
(ipurge_tr_aux_foldr I D P (sinks_aux I D V [x]) xs)
(ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X)"
proof (simp add: ipurge_tr_aux_foldr_ref_aux [OF C])
have "sinks_aux I D (sinks_aux I D V [x]) [] \<subseteq> sinks_aux I D V [x] \<longrightarrow>
ipurge_ref_aux I D (sinks_aux I D V [x]) []
(ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X) =
ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X"
(is "?A \<longrightarrow> ?us = ?vs")
using A .
moreover have ?A
by simp
ultimately have "?us = ?vs" ..
thus "ipurge_ref_aux I D U ys ?us = ?us"
proof simp
have "sinks_aux I D U ys \<subseteq> sinks_aux I D V [x] \<longrightarrow>
ipurge_ref_aux I D U ys
(ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X) =
ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X"
(is "_ \<longrightarrow> ?T")
using A .
moreover have "V \<subseteq> sinks_aux I D V [x]"
by (rule sinks_aux_subset)
hence "sinks_aux I D U ys \<subseteq> sinks_aux I D V [x]"
using B by simp
ultimately show ?T ..
qed
qed
next
have C: "V \<subseteq> sinks_aux I D V [x]"
by (rule sinks_aux_subset)
show
"ipurge_ref_aux I D U ys (ipurge_ref_aux I D V
(ipurge_tr_aux_foldr I D P (sinks_aux I D V [x]) xs)
(ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X)) =
ipurge_ref_aux I D V
(ipurge_tr_aux_foldr I D P (sinks_aux I D V [x]) xs)
(ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X)"
proof (simp add: ipurge_tr_aux_foldr_ref_aux [OF C])
have "sinks_aux I D V [] \<subseteq> sinks_aux I D V [x] \<longrightarrow>
ipurge_ref_aux I D V []
(ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X) =
ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X"
(is "?A \<longrightarrow> ?us = ?vs")
using A .
moreover have ?A
using C by simp
ultimately have "?us = ?vs" ..
thus "ipurge_ref_aux I D U ys ?us = ?us"
proof simp
have "sinks_aux I D U ys \<subseteq> sinks_aux I D V [x] \<longrightarrow>
ipurge_ref_aux I D U ys
(ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X) =
ipurge_ref_aux_foldr I D P (sinks_aux I D V [x]) xs X"
(is "_ \<longrightarrow> ?T")
using A .
moreover have "sinks_aux I D U ys \<subseteq> sinks_aux I D V [x]"
using B and C by simp
ultimately show ?T ..
qed
qed
qed
qed
lemma ipurge_ref_aux_foldr_eq:
"ipurge_ref_aux I D U xs X = ipurge_ref_aux_foldr I D P U xs X"
proof (induction xs arbitrary: U, simp)
fix x xs U
assume A: "\<And>U. ipurge_ref_aux I D U xs X = ipurge_ref_aux_foldr I D P U xs X"
show "ipurge_ref_aux I D U (x # xs) X =
ipurge_ref_aux_foldr I D P U (x # xs) X"
proof (cases "P x", simp_all add: ipurge_ref_aux_cons)
have "sinks_aux I D U [x] \<subseteq> sinks_aux I D U [x]" ..
hence
"ipurge_ref_aux I D (sinks_aux I D U [x])
(ipurge_tr_aux_foldr I D P (sinks_aux I D U [x]) xs)
(ipurge_ref_aux_foldr I D P (sinks_aux I D U [x]) xs X) =
ipurge_ref_aux I D (sinks_aux I D U [x]) []
(ipurge_ref_aux_foldr I D P (sinks_aux I D U [x]) xs X)"
(is "ipurge_ref_aux _ _ _ ?xs' ?X' = _")
by (rule ipurge_tr_aux_foldr_ref_aux)
also have "sinks_aux I D (sinks_aux I D U [x]) [] \<subseteq> sinks_aux I D U [x]"
by simp
hence "ipurge_ref_aux I D (sinks_aux I D U [x]) [] ?X' = ?X'"
by (rule ipurge_ref_aux_foldr_subset)
finally have "ipurge_ref_aux I D (sinks_aux I D U [x]) ?xs' ?X' = ?X'" .
thus "ipurge_ref_aux I D (sinks_aux I D U [x]) xs X =
ipurge_ref_aux I D (sinks_aux I D U [x]) ?xs' ?X'"
proof simp
show "ipurge_ref_aux I D (sinks_aux I D U [x]) xs X =
ipurge_ref_aux_foldr I D P (sinks_aux I D U [x]) xs X"
using A .
qed
next
have "U \<subseteq> sinks_aux I D U [x]"
by (rule sinks_aux_subset)
hence
"ipurge_ref_aux I D U
(ipurge_tr_aux_foldr I D P (sinks_aux I D U [x]) xs)
(ipurge_ref_aux_foldr I D P (sinks_aux I D U [x]) xs X) =
ipurge_ref_aux I D U []
(ipurge_ref_aux_foldr I D P (sinks_aux I D U [x]) xs X)"
(is "ipurge_ref_aux _ _ _ ?xs' ?X' = _")
by (rule ipurge_tr_aux_foldr_ref_aux)
also have "sinks_aux I D U [] \<subseteq> sinks_aux I D U [x]"
by (simp, rule sinks_aux_subset)
hence "ipurge_ref_aux I D U [] ?X' = ?X'"
by (rule ipurge_ref_aux_foldr_subset)
finally have "ipurge_ref_aux I D U ?xs' ?X' = ?X'" .
thus "ipurge_ref_aux I D (sinks_aux I D U [x]) xs X =
ipurge_ref_aux I D U ?xs' ?X'"
proof simp
show "ipurge_ref_aux I D (sinks_aux I D U [x]) xs X =
ipurge_ref_aux_foldr I D P (sinks_aux I D U [x]) xs X"
using A .
qed
qed
qed
text \<open>
\null
Finally, here below is the proof of the implication involving functions @{term ipurge_tr_aux_foldr}
and @{term ipurge_ref_aux_foldr} discussed above.
\null
\<close>
lemma con_comp_sinks_aux_range:
assumes
A: "U \<subseteq> range Some" and
B: "set xs \<subseteq> range p \<union> range q"
shows "sinks_aux (con_comp_pol I) (con_comp_map D E p q) U xs \<subseteq> range Some"
(is "sinks_aux _ ?D' _ _ \<subseteq> _")
proof (rule subsetI, drule sinks_aux_elem, erule disjE, erule_tac [2] bexE)
fix u
assume "u \<in> U"
with A show "u \<in> range Some" ..
next
fix u x
assume "x \<in> set xs"
with B have "x \<in> range p \<union> range q" ..
hence "?D' x \<in> range Some"
by (cases "x \<in> range p", simp_all)
moreover assume "u = ?D' x"
ultimately show "u \<in> range Some"
by simp
qed
lemma con_comp_sinks_aux [rule_format]:
assumes A: "U \<subseteq> range Some"
shows "set xs \<subseteq> range p \<longrightarrow>
sinks_aux I D (the ` U) (map (inv p) xs) =
the ` sinks_aux (con_comp_pol I) (con_comp_map D E p q) U xs"
(is "_ \<longrightarrow> _ = the ` sinks_aux ?I' ?D' _ _")
proof (induction xs rule: rev_induct, simp, rule impI)
fix x xs
assume "set xs \<subseteq> range p \<longrightarrow>
sinks_aux I D (the ` U) (map (inv p) xs) =
the ` sinks_aux ?I' ?D' U xs"
moreover assume B: "set (xs @ [x]) \<subseteq> range p"
ultimately have C: "sinks_aux I D (the ` U) (map (inv p) xs) =
the ` sinks_aux ?I' ?D' U xs"
by simp
show "sinks_aux I D (the ` U) (map (inv p) (xs @ [x])) =
the ` sinks_aux ?I' ?D' U (xs @ [x])"
proof (cases "\<exists>u \<in> sinks_aux ?I' ?D' U xs. (u, ?D' x) \<in> ?I'",
simp_all (no_asm_simp) del: map_append)
case True
then obtain u where
D: "u \<in> sinks_aux ?I' ?D' U xs" and E: "(u, ?D' x) \<in> ?I'" ..
have "(the u, D (inv p x)) \<in> I"
using B and E by (simp add: con_comp_pol_def, erule_tac exE, simp)
moreover have "the u \<in> the ` sinks_aux ?I' ?D' U xs"
using D by simp
hence "the u \<in> sinks_aux I D (the ` U) (map (inv p) xs)"
using C by simp
ultimately have "\<exists>d \<in> sinks_aux I D (the ` U) (map (inv p) xs).
(d, D (inv p x)) \<in> I" ..
hence "sinks_aux I D (the ` U) (map (inv p) (xs @ [x])) =
insert (D (inv p x)) (sinks_aux I D (the ` U) (map (inv p) xs))"
by simp
thus "sinks_aux I D (the ` U) (map (inv p) (xs @ [x])) =
insert (the (?D' x)) (the ` sinks_aux ?I' ?D' U xs)"
using B and C by simp
next
case False
have "\<not> (\<exists>d \<in> sinks_aux I D (the ` U) (map (inv p) xs).
(d, D (inv p x)) \<in> I)"
proof (rule notI, erule bexE)
fix d
assume "d \<in> sinks_aux I D (the ` U) (map (inv p) xs)"
hence "d \<in> the ` sinks_aux ?I' ?D' U xs"
using C by simp
hence "\<exists>u \<in> sinks_aux ?I' ?D' U xs. d = the u"
by (simp add: image_iff)
then obtain u where
D: "u \<in> sinks_aux ?I' ?D' U xs" and E: "d = the u" ..
have "set xs \<subseteq> range p \<union> range q"
using B by (simp, blast)
with A have "sinks_aux ?I' ?D' U xs \<subseteq> range Some"
by (rule con_comp_sinks_aux_range)
hence "u \<in> range Some"
using D ..
hence "u = Some d"
using E by (simp add: image_iff)
moreover assume "(d, D (inv p x)) \<in> I"
hence "(Some d, Some (D (inv p x))) \<in> ?I'"
by (simp add: con_comp_pol_def)
ultimately have "(u, ?D' x) \<in> ?I'"
using B by simp
hence "\<exists>u \<in> sinks_aux ?I' ?D' U xs. (u, ?D' x) \<in> ?I'"
using D ..
thus False
using False by contradiction
qed
thus "sinks_aux I D (the ` U) (map (inv p) (xs @ [x])) =
the ` sinks_aux ?I' ?D' U xs"
using C by simp
qed
qed
lemma con_comp_ipurge_tr_aux [rule_format]:
assumes A: "U \<subseteq> range Some"
shows "set xs \<subseteq> range p \<longrightarrow>
ipurge_tr_aux I D (the ` U) (map (inv p) xs) =
map (inv p) (ipurge_tr_aux (con_comp_pol I) (con_comp_map D E p q) U xs)"
(is "_ \<longrightarrow> _ = map (inv p) (ipurge_tr_aux ?I' ?D' _ _)")
proof (induction xs rule: rev_induct, simp, rule impI)
fix x xs
assume "set xs \<subseteq> range p \<longrightarrow>
ipurge_tr_aux I D (the ` U) (map (inv p) xs) =
map (inv p) (ipurge_tr_aux ?I' ?D' U xs)"
moreover assume B: "set (xs @ [x]) \<subseteq> range p"
ultimately have C: "ipurge_tr_aux I D (the ` U) (map (inv p) xs) =
map (inv p) (ipurge_tr_aux ?I' ?D' U xs)"
by simp
show "ipurge_tr_aux I D (the ` U) (map (inv p) (xs @ [x])) =
map (inv p) (ipurge_tr_aux ?I' ?D' U (xs @ [x]))"
proof (cases "\<exists>u \<in> sinks_aux ?I' ?D' U xs. (u, ?D' x) \<in> ?I'")
case True
then obtain u where
D: "u \<in> sinks_aux ?I' ?D' U xs" and E: "(u, ?D' x) \<in> ?I'" ..
have "(the u, D (inv p x)) \<in> I"
using B and E by (simp add: con_comp_pol_def, erule_tac exE, simp)
moreover have F: "the u \<in> the ` sinks_aux ?I' ?D' U xs"
using D by simp
have "set xs \<subseteq> range p"
using B by simp
with A have "sinks_aux I D (the ` U) (map (inv p) xs) =
the ` sinks_aux ?I' ?D' U xs"
by (rule con_comp_sinks_aux)
hence "the u \<in> sinks_aux I D (the ` U) (map (inv p) xs)"
using F by simp
ultimately have "\<exists>d \<in> sinks_aux I D (the ` U) (map (inv p) xs).
(d, D (inv p x)) \<in> I" ..
hence "ipurge_tr_aux I D (the ` U) (map (inv p) (xs @ [x])) =
ipurge_tr_aux I D (the ` U) (map (inv p) xs)"
by simp
moreover have "map (inv p) (ipurge_tr_aux ?I' ?D' U (xs @ [x])) =
map (inv p) (ipurge_tr_aux ?I' ?D' U xs)"
using True by simp
ultimately show ?thesis
using C by simp
next
case False
have "\<not> (\<exists>d \<in> sinks_aux I D (the ` U) (map (inv p) xs).
(d, D (inv p x)) \<in> I)"
proof (rule notI, erule bexE)
fix d
assume "d \<in> sinks_aux I D (the ` U) (map (inv p) xs)"
moreover have "set xs \<subseteq> range p"
using B by simp
with A have "sinks_aux I D (the ` U) (map (inv p) xs) =
the ` sinks_aux ?I' ?D' U xs"
by (rule con_comp_sinks_aux)
ultimately have "d \<in> the ` sinks_aux ?I' ?D' U xs"
by simp
hence "\<exists>u \<in> sinks_aux ?I' ?D' U xs. d = the u"
by (simp add: image_iff)
then obtain u where
D: "u \<in> sinks_aux ?I' ?D' U xs" and E: "d = the u" ..
have "set xs \<subseteq> range p \<union> range q"
using B by (simp, blast)
with A have "sinks_aux ?I' ?D' U xs \<subseteq> range Some"
by (rule con_comp_sinks_aux_range)
hence "u \<in> range Some"
using D ..
hence "u = Some d"
using E by (simp add: image_iff)
moreover assume "(d, D (inv p x)) \<in> I"
hence "(Some d, Some (D (inv p x))) \<in> ?I'"
by (simp add: con_comp_pol_def)
ultimately have "(u, ?D' x) \<in> ?I'"
using B by simp
hence "\<exists>u \<in> sinks_aux ?I' ?D' U xs. (u, ?D' x) \<in> ?I'"
using D ..
thus False
using False by contradiction
qed
hence "ipurge_tr_aux I D (the ` U) (map (inv p) (xs @ [x])) =
ipurge_tr_aux I D (the ` U) (map (inv p) xs) @ [inv p x]"
by simp
moreover have "map (inv p) (ipurge_tr_aux ?I' ?D' U (xs @ [x])) =
map (inv p) (ipurge_tr_aux ?I' ?D' U xs) @ [inv p x]"
using False by simp
ultimately show ?thesis
using C by simp
qed
qed
lemma con_comp_ipurge_ref_aux:
assumes
A: "U \<subseteq> range Some" and
B: "set xs \<subseteq> range p" and
C: "X \<subseteq> range p"
shows "ipurge_ref_aux I D (the ` U) (map (inv p) xs) (inv p ` X) =
inv p ` ipurge_ref_aux (con_comp_pol I) (con_comp_map D E p q) U xs X"
(is "_ = inv p ` ipurge_ref_aux ?I' ?D' _ _ _")
proof (simp add: ipurge_ref_aux_def set_eq_iff image_iff, rule allI, rule iffI,
erule conjE, erule bexE, erule_tac [2] exE, (erule_tac [2] conjE)+)
fix a x
assume
D: "x \<in> X" and
E: "a = inv p x" and
F: "\<forall>d \<in> sinks_aux I D (the ` U) (map (inv p) xs). (d, D a) \<notin> I"
show "\<exists>x. x \<in> X \<and> (\<forall>u \<in> sinks_aux ?I' ?D' U xs. (u, ?D' x) \<notin> ?I') \<and>
a = inv p x"
proof (rule_tac x = x in exI, simp add: D E, rule ballI)
fix u
assume G: "u \<in> sinks_aux ?I' ?D' U xs"
moreover have "sinks_aux I D (the ` U) (map (inv p) xs) =
the ` sinks_aux ?I' ?D' U xs"
using A and B by (rule con_comp_sinks_aux)
ultimately have "the u \<in> sinks_aux I D (the ` U) (map (inv p) xs)"
by simp
with F have "(the u, D a) \<notin> I" ..
moreover have "set xs \<subseteq> range p \<union> range q"
using B by blast
with A have "sinks_aux ?I' ?D' U xs \<subseteq> range Some"
by (rule con_comp_sinks_aux_range)
hence "u \<in> range Some"
using G ..
hence "\<exists>d. u = Some d"
by (simp add: image_iff)
then obtain d where H: "u = Some d" ..
ultimately have "(d, D (inv p x)) \<notin> I"
using E by simp
hence "(u, Some (D (inv p x))) \<notin> ?I'"
using H by (simp add: con_comp_pol_def)
moreover have "x \<in> range p"
using C and D ..
ultimately show "(u, ?D' x) \<notin> ?I'"
by simp
qed
next
fix a x
assume
D: "x \<in> X" and
E: "a = inv p x" and
F: "\<forall>u \<in> sinks_aux ?I' ?D' U xs. (u, ?D' x) \<notin> ?I'"
show "(\<exists>x \<in> X. a = inv p x) \<and>
(\<forall>u \<in> sinks_aux I D (the ` U) (map (inv p) xs). (u, D a) \<notin> I)"
proof (rule conjI, rule_tac [2] ballI)
show "\<exists>x \<in> X. a = inv p x"
using E and D ..
next
fix d
assume "d \<in> sinks_aux I D (the ` U) (map (inv p) xs)"
moreover have "sinks_aux I D (the ` U) (map (inv p) xs) =
the ` sinks_aux ?I' ?D' U xs"
using A and B by (rule con_comp_sinks_aux)
ultimately have "d \<in> the ` sinks_aux ?I' ?D' U xs"
by simp
hence "\<exists>u \<in> sinks_aux ?I' ?D' U xs. d = the u"
by (simp add: image_iff)
then obtain u where G: "u \<in> sinks_aux ?I' ?D' U xs" and H: "d = the u" ..
have "(u, ?D' x) \<notin> ?I'"
using F and G ..
moreover have "set xs \<subseteq> range p \<union> range q"
using B by blast
with A have "sinks_aux ?I' ?D' U xs \<subseteq> range Some"
by (rule con_comp_sinks_aux_range)
hence "u \<in> range Some"
using G ..
hence "u = Some d"
using H by (simp add: image_iff)
moreover have "x \<in> range p"
using C and D ..
ultimately have "(d, D (inv p x)) \<notin> I"
by (simp add: con_comp_pol_def)
thus "(d, D a) \<notin> I"
using E by simp
qed
qed
lemma con_comp_sinks_filter:
"sinks (con_comp_pol I) (con_comp_map D E p q) u
[x\<leftarrow>xs. x \<in> range p \<union> range q] =
sinks (con_comp_pol I) (con_comp_map D E p q) u xs \<inter> range Some"
(is "sinks ?I' ?D' _ _ = _")
proof (induction xs rule: rev_induct, simp)
fix x xs
assume A: "sinks ?I' ?D' u [x\<leftarrow>xs. x \<in> range p \<union> range q] =
sinks ?I' ?D' u xs \<inter> range Some"
(is "sinks _ _ _ ?xs' = _")
show "sinks ?I' ?D' u [x\<leftarrow>xs @ [x]. x \<in> range p \<union> range q] =
sinks ?I' ?D' u (xs @ [x]) \<inter> range Some"
proof (cases "x \<in> range p \<union> range q", simp_all del: Un_iff sinks.simps,
cases "(u, ?D' x) \<in> ?I' \<or> (\<exists>v \<in> sinks ?I' ?D' u ?xs'. (v, ?D' x) \<in> ?I')")
assume
B: "x \<in> range p \<union> range q" and
C: "(u, ?D' x) \<in> ?I' \<or> (\<exists>v \<in> sinks ?I' ?D' u ?xs'. (v, ?D' x) \<in> ?I')"
have "sinks ?I' ?D' u (?xs' @ [x]) =
insert (?D' x) (sinks ?I' ?D' u ?xs')"
using C by simp
also have "\<dots> =
insert (?D' x) (sinks ?I' ?D' u xs \<inter> range Some)"
using A by simp
also have "\<dots> =
insert (?D' x) (sinks ?I' ?D' u xs) \<inter> insert (?D' x) (range Some)"
by simp
finally have "sinks ?I' ?D' u (?xs' @ [x]) =
insert (?D' x) (sinks ?I' ?D' u xs) \<inter> insert (?D' x) (range Some)" .
moreover have "insert (?D' x) (range Some) = range Some"
using B by (rule_tac insert_absorb, cases "x \<in> range p", simp_all)
ultimately have "sinks ?I' ?D' u (?xs' @ [x]) =
insert (?D' x) (sinks ?I' ?D' u xs) \<inter> range Some"
by simp
moreover have "(u, ?D' x) \<in> ?I' \<or>
(\<exists>v \<in> sinks ?I' ?D' u xs. (v, ?D' x) \<in> ?I')"
using A and C by (simp, blast)
ultimately show "sinks ?I' ?D' u (?xs' @ [x]) =
sinks ?I' ?D' u (xs @ [x]) \<inter> range Some"
by simp
next
assume
B: "x \<in> range p \<union> range q" and
C: "\<not> ((u, ?D' x) \<in> ?I' \<or> (\<exists>v \<in> sinks ?I' ?D' u ?xs'. (v, ?D' x) \<in> ?I'))"
have "sinks ?I' ?D' u (?xs' @ [x]) = sinks ?I' ?D' u ?xs'"
using C by simp
hence "sinks ?I' ?D' u (?xs' @ [x]) = sinks ?I' ?D' u xs \<inter> range Some"
using A by simp
moreover from C have
"\<not> ((u, ?D' x) \<in> ?I' \<or> (\<exists>v \<in> sinks ?I' ?D' u xs. (v, ?D' x) \<in> ?I'))"
proof (rule_tac notI, simp del: bex_simps)
assume "\<exists>v \<in> sinks ?I' ?D' u xs. (v, ?D' x) \<in> ?I'"
then obtain v where E: "v \<in> sinks ?I' ?D' u xs" and F: "(v, ?D' x) \<in> ?I'" ..
have "\<exists>d. ?D' x = Some d"
using B by (cases "x \<in> range p", simp_all)
then obtain d where "?D' x = Some d" ..
hence "(v, Some d) \<in> ?I'"
using F by simp
hence "v \<in> range Some"
by (cases v, simp_all add: con_comp_pol_def)
with E have "v \<in> sinks ?I' ?D' u xs \<inter> range Some" ..
hence "v \<in> sinks ?I' ?D' u ?xs'"
using A by simp
with F have "\<exists>v \<in> sinks ?I' ?D' u ?xs'. (v, ?D' x) \<in> ?I'" ..
thus False
using C by simp
qed
ultimately show "sinks ?I' ?D' u (?xs' @ [x]) =
sinks ?I' ?D' u (xs @ [x]) \<inter> range Some"
by simp
next
assume B: "x \<notin> range p \<union> range q"
hence "(u, ?D' x) \<in> ?I'"
by (simp add: con_comp_pol_def)
hence "sinks ?I' ?D' u (xs @ [x]) = insert (?D' x) (sinks ?I' ?D' u xs)"
by simp
moreover have "insert (?D' x) (sinks ?I' ?D' u xs) \<inter> range Some =
sinks ?I' ?D' u xs \<inter> range Some"
using B by simp
ultimately have "sinks ?I' ?D' u (xs @ [x]) \<inter> range Some =
sinks ?I' ?D' u xs \<inter> range Some"
by simp
thus "sinks ?I' ?D' u ?xs' = sinks ?I' ?D' u (xs @ [x]) \<inter> range Some"
using A by simp
qed
qed
lemma con_comp_ipurge_tr_filter:
"ipurge_tr (con_comp_pol I) (con_comp_map D E p q) u
[x\<leftarrow>xs. x \<in> range p \<union> range q] =
ipurge_tr (con_comp_pol I) (con_comp_map D E p q) u xs"
(is "ipurge_tr ?I' ?D' _ _ = _")
proof (induction xs rule: rev_induct, simp)
fix x xs
assume A: "ipurge_tr ?I' ?D' u [x\<leftarrow>xs. x \<in> range p \<union> range q] =
ipurge_tr ?I' ?D' u xs"
(is "ipurge_tr _ _ _ ?xs' = _")
show "ipurge_tr ?I' ?D' u [x\<leftarrow>xs @ [x]. x \<in> range p \<union> range q] =
ipurge_tr ?I' ?D' u (xs @ [x])"
proof (cases "x \<in> range p \<union> range q", simp_all del: Un_iff ipurge_tr.simps,
cases "?D' x \<in> sinks ?I' ?D' u (?xs' @ [x])")
assume
B: "x \<in> range p \<union> range q" and
C: "?D' x \<in> sinks ?I' ?D' u (?xs' @ [x])"
have "ipurge_tr ?I' ?D' u (?xs' @ [x]) = ipurge_tr ?I' ?D' u ?xs'"
using C by simp
hence "ipurge_tr ?I' ?D' u (?xs' @ [x]) = ipurge_tr ?I' ?D' u xs"
using A by simp
moreover have "?D' x \<in> sinks ?I' ?D' u [x\<leftarrow>xs @ [x]. x \<in> range p \<union> range q]"
using B and C by simp
hence "?D' x \<in> sinks ?I' ?D' u (xs @ [x])"
by (simp only: con_comp_sinks_filter, blast)
ultimately show "ipurge_tr ?I' ?D' u (?xs' @ [x]) =
ipurge_tr ?I' ?D' u (xs @ [x])"
by simp
next
assume
B: "x \<in> range p \<union> range q" and
C: "\<not> (?D' x \<in> sinks ?I' ?D' u (?xs' @ [x]))"
have "ipurge_tr ?I' ?D' u (?xs' @ [x]) = ipurge_tr ?I' ?D' u ?xs' @ [x]"
using C by simp
hence "ipurge_tr ?I' ?D' u (?xs' @ [x]) = ipurge_tr ?I' ?D' u xs @ [x]"
using A by simp
moreover have "?D' x \<notin> sinks ?I' ?D' u [x\<leftarrow>xs @ [x]. x \<in> range p \<union> range q]"
using B and C by simp
hence "?D' x \<notin> sinks ?I' ?D' u (xs @ [x]) \<inter> range Some"
by (simp only: con_comp_sinks_filter, simp)
hence "?D' x \<notin> sinks ?I' ?D' u (xs @ [x])"
using B by (cases "x \<in> range p", simp_all)
ultimately show "ipurge_tr ?I' ?D' u (?xs' @ [x]) =
ipurge_tr ?I' ?D' u (xs @ [x])"
by simp
next
assume "x \<notin> range p \<union> range q"
hence "(u, ?D' x) \<in> ?I'"
by (simp add: con_comp_pol_def)
hence "?D' x \<in> sinks ?I' ?D' u (xs @ [x])"
by simp
thus "ipurge_tr ?I' ?D' u ?xs' = ipurge_tr ?I' ?D' u (xs @ [x])"
using A by simp
qed
qed
lemma con_comp_ipurge_ref_filter:
"ipurge_ref (con_comp_pol I) (con_comp_map D E p q) u
[x\<leftarrow>xs. x \<in> range p \<union> range q] X =
ipurge_ref (con_comp_pol I) (con_comp_map D E p q) u xs X"
(is "ipurge_ref ?I' ?D' _ _ _ = _")
proof (simp add: ipurge_ref_def con_comp_sinks_filter set_eq_iff del: Un_iff,
rule allI, rule iffI, simp_all, (erule conjE)+, rule ballI)
fix x v
assume
A: "(u, ?D' x) \<notin> ?I'" and
B: "\<forall>v \<in> sinks ?I' ?D' u xs \<inter> range Some. (v, ?D' x) \<notin> ?I'" and
C: "v \<in> sinks ?I' ?D' u xs"
show "(v, ?D' x) \<notin> ?I'"
proof (cases v, simp)
have "?D' x \<in> range Some"
using A by (cases "?D' x", simp_all add: con_comp_pol_def)
thus "(None, ?D' x) \<notin> ?I'"
by (simp add: image_iff con_comp_pol_def)
next
fix d
assume "v = Some d"
hence "v \<in> range Some"
by simp
with C have "v \<in> sinks ?I' ?D' u xs \<inter> range Some" ..
with B show "(v, ?D' x) \<notin> ?I'" ..
qed
qed
lemma con_comp_secure_aux [rule_format]:
assumes
A: "secure P I D" and
B: "Y \<subseteq> range p"
shows "set ys \<subseteq> range p \<union> range q \<longrightarrow> U \<subseteq> range Some \<longrightarrow>
(map (inv p) [x\<leftarrow>xs @ ys. x \<in> range p], inv p ` Y) \<in> failures P \<longrightarrow>
(map (inv p) [x\<leftarrow>xs. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr (con_comp_pol I) (con_comp_map D E p q)
(\<lambda>x. x \<in> range p) U ys),
inv p ` ipurge_ref_aux_foldr (con_comp_pol I) (con_comp_map D E p q)
(\<lambda>x. x \<in> range p) U ys Y) \<in> failures P"
proof (induction ys arbitrary: xs U, (rule_tac [!] impI)+, simp)
fix xs U
assume "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` Y) \<in> failures P"
moreover have
"ipurge_ref_aux (con_comp_pol I) (con_comp_map D E p q) U [] Y \<subseteq> Y"
(is "?Y' \<subseteq> _")
by (rule ipurge_ref_aux_subset)
hence "inv p ` ?Y' \<subseteq> inv p ` Y"
by (rule image_mono)
ultimately show "(map (inv p) [x\<leftarrow>xs. x \<in> range p], inv p ` ?Y') \<in> failures P"
by (rule process_rule_3)
next
fix y ys xs U
assume "\<And>xs U. set ys \<subseteq> range p \<union> range q \<longrightarrow> U \<subseteq> range Some \<longrightarrow>
(map (inv p) [x\<leftarrow>xs @ ys. x \<in> range p], inv p ` Y) \<in> failures P \<longrightarrow>
(map (inv p) [x\<leftarrow>xs. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr (con_comp_pol I) (con_comp_map D E p q)
(\<lambda>x. x \<in> range p) U ys),
inv p ` ipurge_ref_aux_foldr (con_comp_pol I) (con_comp_map D E p q)
(\<lambda>x. x \<in> range p) U ys Y) \<in> failures P"
hence "set ys \<subseteq> range p \<union> range q \<longrightarrow>
sinks_aux (con_comp_pol I) (con_comp_map D E p q) U [y] \<subseteq> range Some \<longrightarrow>
(map (inv p) [x\<leftarrow>(xs @ [y]) @ ys. x \<in> range p], inv p ` Y) \<in> failures P \<longrightarrow>
(map (inv p) [x\<leftarrow>xs @ [y]. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr (con_comp_pol I) (con_comp_map D E p q)
(\<lambda>x. x \<in> range p) (sinks_aux (con_comp_pol I) (con_comp_map D E p q)
U [y]) ys),
inv p ` ipurge_ref_aux_foldr (con_comp_pol I) (con_comp_map D E p q)
(\<lambda>x. x \<in> range p) (sinks_aux (con_comp_pol I) (con_comp_map D E p q)
U [y]) ys Y) \<in> failures P"
(is "_ \<longrightarrow> _ \<longrightarrow> _ \<longrightarrow>
(_ @ map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F ?U' _), _) \<in> _") .
moreover assume C: "set (y # ys) \<subseteq> range p \<union> range q"
hence "set ys \<subseteq> range p \<union> range q"
by simp
ultimately have "?U' \<subseteq> range Some \<longrightarrow>
(map (inv p) [x\<leftarrow>(xs @ [y]) @ ys. x \<in> range p], inv p ` Y) \<in> failures P \<longrightarrow>
(map (inv p) [x\<leftarrow>xs @ [y]. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F ?U' ys),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F ?U' ys Y) \<in> failures P" ..
moreover assume D: "U \<subseteq> range Some"
hence "?U' \<subseteq> range Some"
proof (cases "\<exists>u \<in> U. (u, ?D' y) \<in> ?I'", simp_all add: sinks_aux_single_event)
have "y \<in> range p \<union> range q"
using C by simp
thus "?D' y \<in> range Some"
by (cases "y \<in> range p", simp_all)
qed
ultimately have
"(map (inv p) [x\<leftarrow>(xs @ [y]) @ ys. x \<in> range p], inv p ` Y) \<in> failures P \<longrightarrow>
(map (inv p) [x\<leftarrow>xs @ [y]. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F ?U' ys),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F ?U' ys Y) \<in> failures P" ..
moreover assume
"(map (inv p) [x\<leftarrow>xs @ y # ys. x \<in> range p], inv p ` Y) \<in> failures P"
ultimately have
"(map (inv p) [x\<leftarrow>xs @ [y]. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F ?U' ys),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F ?U' ys Y) \<in> failures P"
by simp
hence
"(map (inv p) [x\<leftarrow>xs. x \<in> range p] @
map (inv p) ((if y \<in> range p then [y] else []) @
ipurge_tr_aux_foldr ?I' ?D' ?F ?U' ys),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F ?U' ys Y) \<in> failures P"
by (cases "y \<in> range p", simp_all)
with A have
"(map (inv p) [x\<leftarrow>xs. x \<in> range p] @
ipurge_tr_aux I D (the ` U) (map (inv p) ((if y \<in> range p then [y] else []) @
ipurge_tr_aux_foldr ?I' ?D' ?F ?U' ys)),
ipurge_ref_aux I D (the ` U) (map (inv p) ((if y \<in> range p then [y] else []) @
ipurge_tr_aux_foldr ?I' ?D' ?F ?U' ys))
(inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F ?U' ys Y)) \<in> failures P"
by (rule ipurge_tr_ref_aux_failures_general)
moreover have
"ipurge_tr_aux I D (the ` U) (map (inv p) ((if y \<in> range p then [y] else []) @
ipurge_tr_aux_foldr ?I' ?D' ?F ?U' ys)) =
map (inv p) (ipurge_tr_aux ?I' ?D' U ((if y \<in> range p then [y] else []) @
ipurge_tr_aux_foldr ?I' ?D' ?F ?U' ys))"
by (rule con_comp_ipurge_tr_aux, simp_all add:
D ipurge_tr_aux_foldr_eq [symmetric], blast)
moreover have
"ipurge_ref_aux I D (the ` U) (map (inv p) ((if y \<in> range p then [y] else []) @
ipurge_tr_aux_foldr ?I' ?D' ?F ?U' ys))
(inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F ?U' ys Y) =
inv p ` ipurge_ref_aux ?I' ?D' U ((if y \<in> range p then [y] else []) @
ipurge_tr_aux_foldr ?I' ?D' ?F ?U' ys)
(ipurge_ref_aux_foldr ?I' ?D' ?F ?U' ys Y)"
proof (rule con_comp_ipurge_ref_aux, simp_all add:
D ipurge_tr_aux_foldr_eq [symmetric] ipurge_ref_aux_foldr_eq [symmetric], blast)
have "ipurge_ref_aux ?I' ?D' ?U' ys Y \<subseteq> Y"
by (rule ipurge_ref_aux_subset)
thus "ipurge_ref_aux ?I' ?D' ?U' ys Y \<subseteq> range p"
using B by simp
qed
ultimately show
"(map (inv p) [x\<leftarrow>xs. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F U (y # ys)),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F U (y # ys) Y) \<in> failures P"
by simp
qed
subsection "Conservation of noninterference security under concurrent composition"
text \<open>
Everything is now ready for proving the target security conservation theorem. It states that for any
two processes @{term P}, @{term Q} being secure with respect to the noninterference policy @{term I}
and the event-domain maps @{term D}, @{term E}, their concurrent composition @{term "P \<parallel> Q <p, q>"}
is secure with respect to the noninterference policy @{term "con_comp_pol I"} and the event-domain
map @{term "con_comp_map D E p q"}, provided that condition @{term "consistent_maps D E p q"} is
satisfied.
The only assumption, in addition to the security of the input processes, is the consistency of the
respective event-domain maps. Particularly, this assumption permits to solve the proof obligations
concerning the latter input process by just swapping @{term D} for @{term E} and @{term p} for
@{term q} in the term @{term "con_comp_map D E p q"} and then applying the corresponding lemmas
proven for the former input process.
\null
\<close>
lemma con_comp_secure_del_aux_1:
assumes
A: "secure P I D" and
B: "y \<in> range p \<or> y \<in> range q" and
C: "set ys \<subseteq> range p \<union> range q" and
D: "Y \<subseteq> range p" and
E: "(map (inv p) [x\<leftarrow>xs @ y # ys. x \<in> range p], inv p ` Y) \<in> failures P"
shows
"(map (inv p) [x\<leftarrow>xs @ ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys. x \<in> range p],
inv p ` ipurge_ref (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys Y) \<in> failures P"
(is "(map (inv p) [x\<leftarrow>xs @ ipurge_tr ?I' ?D' _ _. _], _) \<in> _")
proof (simp add:
ipurge_tr_aux_single_dom [symmetric] ipurge_ref_aux_single_dom [symmetric]
ipurge_tr_aux_foldr_eq ipurge_ref_aux_foldr_eq [where P = "\<lambda>x. x \<in> range p"])
have "(map (inv p) [x\<leftarrow>xs @ [y]. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' (\<lambda>x. x \<in> range p) {?D' y} ys),
inv p ` ipurge_ref_aux_foldr ?I' ?D' (\<lambda>x. x \<in> range p) {?D' y} ys Y)
\<in> failures P"
(is "(_ @ map (inv p) (ipurge_tr_aux_foldr _ _ ?F _ _), _) \<in> _")
proof (rule con_comp_secure_aux [OF A D C])
show "{?D' y} \<subseteq> range Some"
using B by (cases "y \<in> range p", simp_all)
next
show "(map (inv p) [x\<leftarrow>(xs @ [y]) @ ys. x \<in> range p], inv p ` Y) \<in> failures P"
using E by simp
qed
thus "(map (inv p) [x\<leftarrow>xs. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {?D' y} ys),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {?D' y} ys Y)
\<in> failures P"
proof (cases "y \<in> range p", simp_all)
assume "(map (inv p) [x\<leftarrow>xs. x \<in> range p] @ inv p y #
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys Y)
\<in> failures P"
hence "(inv p y #
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys Y)
\<in> futures P (map (inv p) [x\<leftarrow>xs. x \<in> range p])"
by (simp add: futures_def)
hence
"(ipurge_tr I D (D (inv p y))
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys)),
ipurge_ref I D (D (inv p y))
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys))
(inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys Y))
\<in> futures P (map (inv p) [x\<leftarrow>xs. x \<in> range p])"
using A by (simp add: secure_def)
hence
"(ipurge_tr_aux I D (the ` {Some (D (inv p y))})
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys)),
ipurge_ref_aux I D (the ` {Some (D (inv p y))})
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys))
(inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys Y))
\<in> futures P (map (inv p) [x\<leftarrow>xs. x \<in> range p])"
by (simp add: ipurge_tr_aux_single_dom ipurge_ref_aux_single_dom)
moreover have
"ipurge_tr_aux I D (the ` {Some (D (inv p y))})
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys)) =
map (inv p) (ipurge_tr_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys))"
by (rule con_comp_ipurge_tr_aux, simp_all add:
ipurge_tr_aux_foldr_eq [symmetric], blast)
moreover have
"ipurge_ref_aux I D (the ` {Some (D (inv p y))})
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys))
(inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys Y) =
inv p ` ipurge_ref_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys)
(ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys Y)"
proof (rule con_comp_ipurge_ref_aux, simp_all add:
ipurge_tr_aux_foldr_eq [symmetric] ipurge_ref_aux_foldr_eq [symmetric], blast)
have "ipurge_ref_aux ?I' ?D' {Some (D (inv p y))} ys Y \<subseteq> Y"
by (rule ipurge_ref_aux_subset)
thus "ipurge_ref_aux ?I' ?D' {Some (D (inv p y))} ys Y \<subseteq> range p"
using D by simp
qed
ultimately have
"(map (inv p) (ipurge_tr_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys)),
inv p ` ipurge_ref_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys)
(ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys Y))
\<in> futures P (map (inv p) [x\<leftarrow>xs. x \<in> range p])"
by simp
moreover have
"ipurge_tr_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys) =
ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys"
by (rule ipurge_tr_aux_foldr_subset, simp)
moreover have
"ipurge_ref_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys)
(ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys Y) =
ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys Y"
by (rule ipurge_ref_aux_foldr_subset, subst ipurge_tr_aux_foldr_sinks_aux, simp)
ultimately have
"(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys Y)
\<in> futures P (map (inv p) [x\<leftarrow>xs. x \<in> range p])"
by simp
thus "(map (inv p) [x\<leftarrow>xs. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} ys Y)
\<in> failures P"
by (simp add: futures_def)
qed
qed
lemma con_comp_secure_add_aux_1:
assumes
A: "secure P I D" and
B: "y \<in> range p \<or> y \<in> range q" and
C: "set zs \<subseteq> range p \<union> range q" and
D: "Z \<subseteq> range p" and
E: "(map (inv p) [x\<leftarrow>xs @ zs. x \<in> range p], inv p ` Z) \<in> failures P" and
F: "map (inv p) [x\<leftarrow>xs @ [y]. x \<in> range p] \<in> traces P"
shows
"(map (inv p) [x\<leftarrow>xs @ y # ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) zs. x \<in> range p],
inv p ` ipurge_ref (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) zs Z) \<in> failures P"
(is "(map (inv p) [x\<leftarrow>xs @ y # ipurge_tr ?I' ?D' _ _. _], _) \<in> _")
proof -
have
"(map (inv p) [x\<leftarrow>(xs @ [y]) @ ipurge_tr ?I' ?D' (?D' y) zs. x \<in> range p],
inv p ` ipurge_ref ?I' ?D' (?D' y) zs Z) \<in> failures P"
proof (subst filter_append, simp del: filter_append add:
ipurge_tr_aux_single_dom [symmetric] ipurge_ref_aux_single_dom [symmetric]
ipurge_tr_aux_foldr_eq ipurge_ref_aux_foldr_eq [where P = "\<lambda>x. x \<in> range p"])
have "(map (inv p) [x\<leftarrow>xs. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' (\<lambda>x. x \<in> range p) {?D' y} zs),
inv p ` ipurge_ref_aux_foldr ?I' ?D' (\<lambda>x. x \<in> range p) {?D' y} zs Z)
\<in> failures P"
(is "(_ @ map (inv p) (ipurge_tr_aux_foldr _ _ ?F _ _), _) \<in> _")
proof (rule con_comp_secure_aux [OF A D C])
show "{?D' y} \<subseteq> range Some"
using B by (cases "y \<in> range p", simp_all)
next
show "(map (inv p) [x\<leftarrow>xs @ zs. x \<in> range p], inv p ` Z) \<in> failures P"
using E .
qed
thus "(map (inv p) [x\<leftarrow>xs @ [y]. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {?D' y} zs),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {?D' y} zs Z)
\<in> failures P"
proof (cases "y \<in> range p", simp_all)
case True
assume "(map (inv p) [x\<leftarrow>xs. x \<in> range p] @
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs Z)
\<in> failures P"
hence
"(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs Z)
\<in> futures P (map (inv p) [x\<leftarrow>xs. x \<in> range p])"
by (simp add: futures_def)
moreover have "(map (inv p) [x\<leftarrow>xs @ [y]. x \<in> range p], {}) \<in> failures P"
using F by (rule traces_failures)
hence "([inv p y], {}) \<in> futures P (map (inv p) [x\<leftarrow>xs. x \<in> range p])"
using True by (simp add: futures_def)
ultimately have
"(inv p y # ipurge_tr I D (D (inv p y))
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs)),
ipurge_ref I D (D (inv p y))
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs))
(inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs Z))
\<in> futures P (map (inv p) [x\<leftarrow>xs. x \<in> range p])"
using A by (simp add: secure_def)
hence
"(inv p y # ipurge_tr_aux I D (the ` {Some (D (inv p y))})
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs)),
ipurge_ref_aux I D (the ` {Some (D (inv p y))})
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs))
(inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs Z))
\<in> futures P (map (inv p) [x\<leftarrow>xs. x \<in> range p])"
by (simp add: ipurge_tr_aux_single_dom ipurge_ref_aux_single_dom)
moreover have
"ipurge_tr_aux I D (the ` {Some (D (inv p y))})
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs)) =
map (inv p) (ipurge_tr_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs))"
by (rule con_comp_ipurge_tr_aux, simp_all add:
ipurge_tr_aux_foldr_eq [symmetric], blast)
moreover have
"ipurge_ref_aux I D (the ` {Some (D (inv p y))})
(map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs))
(inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs Z) =
inv p ` ipurge_ref_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs)
(ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs Z)"
proof (rule con_comp_ipurge_ref_aux, simp_all add:
ipurge_tr_aux_foldr_eq [symmetric] ipurge_ref_aux_foldr_eq [symmetric], blast)
have "ipurge_ref_aux ?I' ?D' {Some (D (inv p y))} zs Z \<subseteq> Z"
by (rule ipurge_ref_aux_subset)
thus "ipurge_ref_aux ?I' ?D' {Some (D (inv p y))} zs Z \<subseteq> range p"
using D by simp
qed
ultimately have
"(inv p y # map (inv p) (ipurge_tr_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs)),
inv p ` ipurge_ref_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs)
(ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs Z))
\<in> futures P (map (inv p) [x\<leftarrow>xs. x \<in> range p])"
by simp
moreover have
"ipurge_tr_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs) =
ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs"
by (rule ipurge_tr_aux_foldr_subset, simp)
moreover have
"ipurge_ref_aux ?I' ?D' {Some (D (inv p y))}
(ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs)
(ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs Z) =
ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs Z"
by (rule ipurge_ref_aux_foldr_subset, subst ipurge_tr_aux_foldr_sinks_aux, simp)
ultimately have
"(inv p y # map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F
{Some (D (inv p y))} zs),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F
{Some (D (inv p y))} zs Z)
\<in> futures P (map (inv p) [x\<leftarrow>xs. x \<in> range p])"
by simp
thus "(map (inv p) [x\<leftarrow>xs. x \<in> range p] @ inv p y #
map (inv p) (ipurge_tr_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs),
inv p ` ipurge_ref_aux_foldr ?I' ?D' ?F {Some (D (inv p y))} zs Z)
\<in> failures P"
by (simp add: futures_def)
qed
qed
thus ?thesis
by simp
qed
lemma con_comp_consistent_maps:
"consistent_maps D E p q \<Longrightarrow> con_comp_map D E p q = con_comp_map E D q p"
-proof (simp add: consistent_maps_def, rule ext)
+using [[simproc del: defined_all]] proof (simp add: consistent_maps_def, rule ext)
fix x
assume A: "\<forall>x \<in> range p \<inter> range q. D (inv p x) = E (inv q x)"
show "con_comp_map D E p q x = con_comp_map E D q p x"
proof (rule con_comp_map.cases [of "(D, E, p, q, x)"], simp_all, (erule conjE)+)
fix p' q' D' E' x'
assume
B: "p = p'" and
C: "q = q'" and
D: "D = D'" and
E: "E = E'" and
F: "x' \<in> range p'"
show "Some (D' (inv p' x')) = con_comp_map E' D' q' p' x'"
proof (cases "x' \<in> range q'", simp_all add: F)
case True
with F have "x' \<in> range p' \<inter> range q'" ..
hence "x' \<in> range p \<inter> range q"
using B and C by simp
with A have "D (inv p x') = E (inv q x')" ..
thus "D' (inv p' x') = E' (inv q' x')"
using B and C and D and E by simp
qed
qed
qed
lemma con_comp_secure_del_aux_2:
assumes A: "consistent_maps D E p q"
shows
"secure Q I E \<Longrightarrow>
y \<in> range p \<or> y \<in> range q \<Longrightarrow>
set ys \<subseteq> range p \<union> range q \<Longrightarrow>
Y \<subseteq> range q \<Longrightarrow>
(map (inv q) [x\<leftarrow>xs @ y # ys. x \<in> range q], inv q ` Y) \<in> failures Q \<Longrightarrow>
(map (inv q) [x\<leftarrow>xs @ ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys. x \<in> range q],
inv q ` ipurge_ref (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys Y) \<in> failures Q"
proof (simp only: con_comp_consistent_maps [OF A], rule con_comp_secure_del_aux_1)
qed (simp_all, blast+)
lemma con_comp_secure_add_aux_2:
assumes A: "consistent_maps D E p q"
shows
"secure Q I E \<Longrightarrow>
y \<in> range p \<or> y \<in> range q \<Longrightarrow>
set zs \<subseteq> range p \<union> range q \<Longrightarrow>
Z \<subseteq> range q \<Longrightarrow>
(map (inv q) [x\<leftarrow>xs @ zs. x \<in> range q], inv q ` Z) \<in> failures Q \<Longrightarrow>
map (inv q) [x\<leftarrow>xs @ [y]. x \<in> range q] \<in> traces Q \<Longrightarrow>
(map (inv q) [x\<leftarrow>xs @ y # ipurge_tr (con_comp_pol I)
(con_comp_map D E p q) (con_comp_map D E p q y) zs. x \<in> range q],
inv q ` ipurge_ref (con_comp_pol I)
(con_comp_map D E p q) (con_comp_map D E p q y) zs Z) \<in> failures Q"
proof (simp only: con_comp_consistent_maps [OF A], rule con_comp_secure_add_aux_1)
qed (simp_all, blast+)
lemma con_comp_secure_del_case_1:
assumes
A: "consistent_maps D E p q" and
B: "secure P I D" and
C: "secure Q I E"
shows
"\<exists>R S T.
Y = R \<union> S \<union> T \<and>
(y \<in> range p \<or> y \<in> range q) \<and>
set xs \<subseteq> range p \<union> range q \<and>
set ys \<subseteq> range p \<union> range q \<and>
R \<subseteq> range p \<and>
S \<subseteq> range q \<and>
T \<subseteq> - range p \<and>
T \<subseteq> - range q \<and>
(map (inv p) [x\<leftarrow>xs @ y # ys. x \<in> range p], inv p ` R) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs @ y # ys. x \<in> range q], inv q ` S) \<in> failures Q \<Longrightarrow>
\<exists>R S T.
ipurge_ref (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys Y = R \<union> S \<union> T \<and>
set xs \<subseteq> range p \<union> range q \<and>
set (ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys) \<subseteq> range p \<union> range q \<and>
R \<subseteq> range p \<and>
S \<subseteq> range q \<and>
T \<subseteq> - range p \<and>
T \<subseteq> - range q \<and>
(map (inv p) [x\<leftarrow>xs @ ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys. x \<in> range p], inv p ` R) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs @ ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys. x \<in> range q], inv q ` S) \<in> failures Q"
(is "_ \<Longrightarrow> \<exists>_ _ _. ipurge_ref ?I' ?D' _ _ _ = _ \<and> _")
proof ((erule exE)+, (erule conjE)+)
fix R S T
assume
D: "Y = R \<union> S \<union> T" and
E: "y \<in> range p \<or> y \<in> range q" and
F: "set xs \<subseteq> range p \<union> range q" and
G: "set ys \<subseteq> range p \<union> range q" and
H: "R \<subseteq> range p" and
I: "S \<subseteq> range q" and
J: "T \<subseteq> - range p" and
K: "T \<subseteq> - range q" and
L: "(map (inv p) [x\<leftarrow>xs @ y # ys. x \<in> range p], inv p ` R) \<in> failures P" and
M: "(map (inv q) [x\<leftarrow>xs @ y # ys. x \<in> range q], inv q ` S) \<in> failures Q"
show ?thesis
proof (rule_tac x = "ipurge_ref ?I' ?D' (?D' y) ys R" in exI,
rule_tac x = "ipurge_ref ?I' ?D' (?D' y) ys S" in exI, rule_tac x = "{}" in exI,
(subst conj_assoc [symmetric])+, (rule conjI)+, simp_all del: filter_append)
have "ipurge_ref ?I' ?D' (?D' y) ys Y =
ipurge_ref ?I' ?D' (?D' y) ys (R \<union> S \<union> T)"
using D by simp
hence "ipurge_ref ?I' ?D' (?D' y) ys Y =
ipurge_ref ?I' ?D' (?D' y) ys R \<union>
ipurge_ref ?I' ?D' (?D' y) ys S \<union>
ipurge_ref ?I' ?D' (?D' y) ys T"
by (simp add: ipurge_ref_distrib_union)
moreover have "ipurge_ref ?I' ?D' (?D' y) ys T = {}"
proof (rule ipurge_ref_empty [of "?D' y"], simp, insert E,
cases "y \<in> range p", simp_all)
fix x
assume N: "x \<in> T"
with J have "x \<in> - range p" ..
moreover have "x \<in> - range q"
using K and N ..
ultimately have "?D' x = None"
by simp
thus "(Some (D (inv p y)), ?D' x) \<in> ?I'"
by (simp add: con_comp_pol_def)
next
fix x
assume N: "x \<in> T"
with J have "x \<in> - range p" ..
moreover have "x \<in> - range q"
using K and N ..
ultimately have "?D' x = None"
by simp
thus "(Some (E (inv q y)), ?D' x) \<in> ?I'"
by (simp add: con_comp_pol_def)
qed
ultimately show "ipurge_ref ?I' ?D' (?D' y) ys Y =
ipurge_ref ?I' ?D' (?D' y) ys R \<union>
ipurge_ref ?I' ?D' (?D' y) ys S"
by simp
next
show "set xs \<subseteq> range p \<union> range q"
using F .
next
have "set (ipurge_tr ?I' ?D' (?D' y) ys) \<subseteq> set ys"
by (rule ipurge_tr_set)
thus "set (ipurge_tr ?I' ?D' (?D' y) ys) \<subseteq> range p \<union> range q"
using G by simp
next
have "ipurge_ref ?I' ?D' (?D' y) ys R \<subseteq> R"
by (rule ipurge_ref_subset)
thus "ipurge_ref ?I' ?D' (?D' y) ys R \<subseteq> range p"
using H by simp
next
have "ipurge_ref ?I' ?D' (?D' y) ys S \<subseteq> S"
by (rule ipurge_ref_subset)
thus "ipurge_ref ?I' ?D' (?D' y) ys S \<subseteq> range q"
using I by simp
next
show "(map (inv p) [x\<leftarrow>xs @ ipurge_tr ?I' ?D' (?D' y) ys. x \<in> range p],
inv p ` ipurge_ref ?I' ?D' (?D' y) ys R) \<in> failures P"
by (rule con_comp_secure_del_aux_1 [OF B E G H L])
next
show "(map (inv q) [x\<leftarrow>xs @ ipurge_tr ?I' ?D' (?D' y) ys. x \<in> range q],
inv q ` ipurge_ref ?I' ?D' (?D' y) ys S) \<in> failures Q"
by (rule con_comp_secure_del_aux_2 [OF A C E G I M])
qed
qed
lemma con_comp_secure_del_case_2:
assumes
A: "consistent_maps D E p q" and
B: "secure P I D" and
C: "secure Q I E"
shows
"\<exists>xs'.
(\<exists>ys'. xs @ y # ys = xs' @ ys') \<and>
set xs' \<subseteq> range p \<union> range q \<and>
map (inv p) [x\<leftarrow>xs'. x \<in> range p] \<in> divergences P \<and>
map (inv q) [x\<leftarrow>xs'. x \<in> range q] \<in> divergences Q \<Longrightarrow>
(\<exists>R S T.
ipurge_ref (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys Y = R \<union> S \<union> T \<and>
set xs \<subseteq> range p \<union> range q \<and>
set (ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys) \<subseteq> range p \<union> range q \<and>
R \<subseteq> range p \<and>
S \<subseteq> range q \<and>
T \<subseteq> - range p \<and>
T \<subseteq> - range q \<and>
(map (inv p) [x\<leftarrow>xs @ ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys. x \<in> range p], inv p ` R) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs @ ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys. x \<in> range q], inv q ` S) \<in> failures Q) \<or>
(\<exists>xs'.
(\<exists>ys'. xs @ ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys = xs' @ ys') \<and>
set xs' \<subseteq> range p \<union> range q \<and>
map (inv p) [x\<leftarrow>xs'. x \<in> range p] \<in> divergences P \<and>
map (inv q) [x\<leftarrow>xs'. x \<in> range q] \<in> divergences Q)"
(is "_ \<Longrightarrow> (\<exists>R S T. ?F R S T ys) \<or> ?G")
proof (erule exE, (erule conjE)+, erule exE)
fix xs' ys'
assume
D: "set xs' \<subseteq> range p \<union> range q" and
E: "map (inv p) [x\<leftarrow>xs'. x \<in> range p] \<in> divergences P" and
F: "map (inv q) [x\<leftarrow>xs'. x \<in> range q] \<in> divergences Q" and
G: "xs @ y # ys = xs' @ ys'"
show ?thesis
proof (cases "length xs < length xs'", rule disjI1, rule_tac [2] disjI2)
case True
moreover have "take (length xs') (xs @ [y] @ ys) =
take (length xs') (xs @ [y]) @ take (length xs' - Suc (length xs)) ys"
by (simp only: take_append, simp)
ultimately have "take (length xs') (xs @ [y] @ ys) =
xs @ y # take (length xs' - Suc (length xs)) ys"
(is "_ = _ @ _ # ?vs")
by simp
moreover have "take (length xs') (xs @ [y] @ ys) =
take (length xs') (xs' @ ys')"
using G by simp
ultimately have H: "xs @ y # ?vs = xs'"
by simp
moreover have "y \<in> set (xs @ y # ?vs)"
by simp
ultimately have "y \<in> set xs'"
by simp
with D have I: "y \<in> range p \<union> range q" ..
have "set xs \<subseteq> set (xs @ y # ?vs)"
by auto
hence "set xs \<subseteq> set xs'"
using H by simp
hence J: "set xs \<subseteq> range p \<union> range q"
using D by simp
have "\<exists>R S T. ?F R S T [x\<leftarrow>ys. x \<in> range p \<union> range q]"
(is "\<exists>_ _ _. ipurge_ref ?I' ?D' _ _ _ = _ \<and> _")
proof (rule con_comp_secure_del_case_1 [OF A B C],
rule_tac x = "range p \<inter> Y" in exI, rule_tac x = "range q \<inter> Y" in exI,
rule_tac x = "- range p \<inter> - range q \<inter> Y" in exI,
(subst conj_assoc [symmetric])+, (rule conjI)+, simp_all del: filter_append)
show "Y = range p \<inter> Y \<union> range q \<inter> Y \<union> - range p \<inter> - range q \<inter> Y"
by blast
next
show "y \<in> range p \<or> y \<in> range q"
using I by simp
next
show "set xs \<subseteq> range p \<union> range q"
using J .
next
show "{x \<in> set ys. x \<in> range p \<or> x \<in> range q} \<subseteq> range p \<union> range q"
by blast
next
show "- range p \<inter> - range q \<inter> Y \<subseteq> - range p"
by blast
next
show "- range p \<inter> - range q \<inter> Y \<subseteq> - range q"
by blast
next
have "map (inv p) [x\<leftarrow>xs @ y # ?vs. x \<in> range p] \<in> divergences P"
using E and H by simp
hence "map (inv p) [x\<leftarrow>xs @ y # ?vs. x \<in> range p] @
map (inv p) [x\<leftarrow>drop (length xs' - Suc (length xs)) ys. x \<in> range p]
\<in> divergences P"
(is "_ @ map (inv p) [x\<leftarrow>?ws. _] \<in> _")
by (rule process_rule_5_general)
hence "map (inv p) [x\<leftarrow>(xs @ y # ?vs) @ ?ws. x \<in> range p] \<in> divergences P"
by (subst filter_append, simp)
hence "map (inv p) [x\<leftarrow>(xs @ [y]) @ ys. x \<in> range p] \<in> divergences P"
by simp
hence "map (inv p) [x\<leftarrow>(xs @ [y]) @ [x\<leftarrow>ys. x \<in> range p \<or> x \<in> range q].
x \<in> range p] \<in> divergences P"
proof (subst (asm) filter_append, subst filter_append, subst filter_filter)
qed (subgoal_tac "(\<lambda>x. (x \<in> range p \<or> x \<in> range q) \<and> x \<in> range p) =
(\<lambda>x. x \<in> range p)", simp, blast)
hence "map (inv p) [x\<leftarrow>xs @ y # [x\<leftarrow>ys. x \<in> range p \<or> x \<in> range q].
x \<in> range p] \<in> divergences P"
by simp
thus "(map (inv p) [x\<leftarrow>xs @ y # [x\<leftarrow>ys. x \<in> range p \<or> x \<in> range q].
x \<in> range p], inv p ` (range p \<inter> Y)) \<in> failures P"
by (rule process_rule_6)
next
have "map (inv q) [x\<leftarrow>xs @ y # ?vs. x \<in> range q] \<in> divergences Q"
using F and H by simp
hence "map (inv q) [x\<leftarrow>xs @ y # ?vs. x \<in> range q] @
map (inv q) [x\<leftarrow>drop (length xs' - Suc (length xs)) ys. x \<in> range q]
\<in> divergences Q"
(is "_ @ map (inv q) [x\<leftarrow>?ws. _] \<in> _")
by (rule process_rule_5_general)
hence "map (inv q) [x\<leftarrow>(xs @ y # ?vs) @ ?ws. x \<in> range q] \<in> divergences Q"
by (subst filter_append, simp)
hence "map (inv q) [x\<leftarrow>(xs @ [y]) @ ys. x \<in> range q] \<in> divergences Q"
by simp
hence "map (inv q) [x\<leftarrow>(xs @ [y]) @ [x\<leftarrow>ys. x \<in> range p \<or> x \<in> range q].
x \<in> range q] \<in> divergences Q"
proof (subst (asm) filter_append, subst filter_append, subst filter_filter)
qed (subgoal_tac "(\<lambda>x. (x \<in> range p \<or> x \<in> range q) \<and> x \<in> range q) =
(\<lambda>x. x \<in> range q)", simp, blast)
hence "map (inv q) [x\<leftarrow>xs @ y # [x\<leftarrow>ys. x \<in> range p \<or> x \<in> range q].
x \<in> range q] \<in> divergences Q"
by simp
thus "(map (inv q) [x\<leftarrow>xs @ y # [x\<leftarrow>ys. x \<in> range p \<or> x \<in> range q].
x \<in> range q], inv q ` (range q \<inter> Y)) \<in> failures Q"
by (rule process_rule_6)
qed
then obtain R and S and T where
"?F R S T [x\<leftarrow>ys. x \<in> range p \<union> range q]"
by blast
thus "\<exists>R S T. ?F R S T ys"
proof (rule_tac x = R in exI, rule_tac x = S in exI, rule_tac x = T in exI)
qed (simp only: con_comp_ipurge_tr_filter con_comp_ipurge_ref_filter)
next
let
?I' = "con_comp_pol I" and
?D' = "con_comp_map D E p q"
case False
moreover have "xs @ ipurge_tr ?I' ?D' (?D' y) ys =
take (length xs') (xs @ ipurge_tr ?I' ?D' (?D' y) ys) @
drop (length xs') (xs @ ipurge_tr ?I' ?D' (?D' y) ys)"
(is "_ = _ @ ?vs")
by (simp only: append_take_drop_id)
ultimately have "xs @ ipurge_tr ?I' ?D' (?D' y) ys =
take (length xs') (xs @ y # ys) @ ?vs"
by simp
hence H: "xs @ ipurge_tr ?I' ?D' (?D' y) ys = xs' @ ?vs"
using G by simp
show ?G
proof (rule_tac x = xs' in exI, rule conjI, rule_tac x = ?vs in exI)
qed (subst H, simp_all add: D E F)
qed
qed
lemma con_comp_secure_add_case_1:
assumes
A: "consistent_maps D E p q" and
B: "secure P I D" and
C: "secure Q I E" and
D: "(xs @ y # ys, Y) \<in> con_comp_failures P Q p q" and
E: "y \<in> range p \<or> y \<in> range q"
shows
"\<exists>R S T.
Z = R \<union> S \<union> T \<and>
set xs \<subseteq> range p \<union> range q \<and>
set zs \<subseteq> range p \<union> range q \<and>
R \<subseteq> range p \<and>
S \<subseteq> range q \<and>
T \<subseteq> - range p \<and>
T \<subseteq> - range q \<and>
(map (inv p) [x\<leftarrow>xs @ zs. x \<in> range p], inv p ` R) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs @ zs. x \<in> range q], inv q ` S) \<in> failures Q \<Longrightarrow>
\<exists>R S T.
ipurge_ref (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) zs Z = R \<union> S \<union> T \<and>
set xs \<subseteq> range p \<union> range q \<and>
set (ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) zs) \<subseteq> range p \<union> range q \<and>
R \<subseteq> range p \<and>
S \<subseteq> range q \<and>
T \<subseteq> - range p \<and>
T \<subseteq> - range q \<and>
(map (inv p) [x\<leftarrow>xs @ y # ipurge_tr (con_comp_pol I)
(con_comp_map D E p q) (con_comp_map D E p q y) zs. x \<in> range p],
inv p ` R) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs @ y # ipurge_tr (con_comp_pol I)
(con_comp_map D E p q) (con_comp_map D E p q y) zs. x \<in> range q],
inv q ` S) \<in> failures Q"
(is "_ \<Longrightarrow> \<exists>_ _ _. ipurge_ref ?I' ?D' _ _ _ = _ \<and> _")
proof ((erule exE)+, (erule conjE)+)
fix R S T
assume
F: "Z = R \<union> S \<union> T" and
G: "set xs \<subseteq> range p \<union> range q" and
H: "set zs \<subseteq> range p \<union> range q" and
I: "R \<subseteq> range p" and
J: "S \<subseteq> range q" and
K: "T \<subseteq> - range p" and
L: "T \<subseteq> - range q" and
M: "(map (inv p) [x\<leftarrow>xs @ zs. x \<in> range p], inv p ` R) \<in> failures P" and
N: "(map (inv q) [x\<leftarrow>xs @ zs. x \<in> range q], inv q ` S) \<in> failures Q"
show ?thesis
proof (rule_tac x = "ipurge_ref ?I' ?D' (?D' y) zs R" in exI,
rule_tac x = "ipurge_ref ?I' ?D' (?D' y) zs S" in exI, rule_tac x = "{}" in exI,
(subst conj_assoc [symmetric])+, (rule conjI)+, simp_all del: filter_append)
have "ipurge_ref ?I' ?D' (?D' y) zs Z =
ipurge_ref ?I' ?D' (?D' y) zs (R \<union> S \<union> T)"
using F by simp
hence "ipurge_ref ?I' ?D' (?D' y) zs Z =
ipurge_ref ?I' ?D' (?D' y) zs R \<union>
ipurge_ref ?I' ?D' (?D' y) zs S \<union>
ipurge_ref ?I' ?D' (?D' y) zs T"
by (simp add: ipurge_ref_distrib_union)
moreover have "ipurge_ref ?I' ?D' (?D' y) zs T = {}"
proof (rule ipurge_ref_empty [of "?D' y"], simp, insert E,
cases "y \<in> range p", simp_all)
fix x
assume O: "x \<in> T"
with K have "x \<in> - range p" ..
moreover have "x \<in> - range q"
using L and O ..
ultimately have "?D' x = None"
by simp
thus "(Some (D (inv p y)), ?D' x) \<in> ?I'"
by (simp add: con_comp_pol_def)
next
fix x
assume O: "x \<in> T"
with K have "x \<in> - range p" ..
moreover have "x \<in> - range q"
using L and O ..
ultimately have "?D' x = None"
by simp
thus "(Some (E (inv q y)), ?D' x) \<in> ?I'"
by (simp add: con_comp_pol_def)
qed
ultimately show "ipurge_ref ?I' ?D' (?D' y) zs Z =
ipurge_ref ?I' ?D' (?D' y) zs R \<union>
ipurge_ref ?I' ?D' (?D' y) zs S"
by simp
next
show "set xs \<subseteq> range p \<union> range q"
using G .
next
have "set (ipurge_tr ?I' ?D' (?D' y) zs) \<subseteq> set zs"
by (rule ipurge_tr_set)
thus "set (ipurge_tr ?I' ?D' (?D' y) zs) \<subseteq> range p \<union> range q"
using H by simp
next
have "ipurge_ref ?I' ?D' (?D' y) zs R \<subseteq> R"
by (rule ipurge_ref_subset)
thus "ipurge_ref ?I' ?D' (?D' y) zs R \<subseteq> range p"
using I by simp
next
have "ipurge_ref ?I' ?D' (?D' y) zs S \<subseteq> S"
by (rule ipurge_ref_subset)
thus "ipurge_ref ?I' ?D' (?D' y) zs S \<subseteq> range q"
using J by simp
next
have "map (inv p) [x\<leftarrow>xs @ y # ys. x \<in> range p] \<in> traces P \<and>
map (inv q) [x\<leftarrow>xs @ y # ys. x \<in> range q] \<in> traces Q"
using D by (rule con_comp_failures_traces)
hence "map (inv p) [x\<leftarrow>(xs @ [y]) @ ys. x \<in> range p] \<in> traces P"
by simp
hence "map (inv p) [x\<leftarrow>xs @ [y]. x \<in> range p] @
map (inv p) [x\<leftarrow>ys. x \<in> range p] \<in> traces P"
by (subst (asm) filter_append, simp)
hence "map (inv p) [x\<leftarrow>xs @ [y]. x \<in> range p] \<in> traces P"
by (rule process_rule_2_traces)
thus "(map (inv p) [x\<leftarrow>xs @ y # ipurge_tr ?I' ?D' (?D' y) zs. x \<in> range p],
inv p ` ipurge_ref ?I' ?D' (?D' y) zs R) \<in> failures P"
by (rule con_comp_secure_add_aux_1 [OF B E H I M])
next
have "map (inv p) [x\<leftarrow>xs @ y # ys. x \<in> range p] \<in> traces P \<and>
map (inv q) [x\<leftarrow>xs @ y # ys. x \<in> range q] \<in> traces Q"
using D by (rule con_comp_failures_traces)
hence "map (inv q) [x\<leftarrow>(xs @ [y]) @ ys. x \<in> range q] \<in> traces Q"
by simp
hence "map (inv q) [x\<leftarrow>xs @ [y]. x \<in> range q] @
map (inv q) [x\<leftarrow>ys. x \<in> range q] \<in> traces Q"
by (subst (asm) filter_append, simp)
hence "map (inv q) [x\<leftarrow>xs @ [y]. x \<in> range q] \<in> traces Q"
by (rule process_rule_2_traces)
thus "(map (inv q) [x\<leftarrow>xs @ y # ipurge_tr ?I' ?D' (?D' y) zs. x \<in> range q],
inv q ` ipurge_ref ?I' ?D' (?D' y) zs S) \<in> failures Q"
by (rule con_comp_secure_add_aux_2 [OF A C E H J N])
qed
qed
lemma con_comp_secure_add_case_2:
assumes
A: "consistent_maps D E p q" and
B: "secure P I D" and
C: "secure Q I E" and
D: "(xs @ y # ys, Y) \<in> con_comp_failures P Q p q" and
E: "y \<in> range p \<or> y \<in> range q"
shows
"\<exists>xs'.
(\<exists>ys'. xs @ zs = xs' @ ys') \<and>
set xs' \<subseteq> range p \<union> range q \<and>
map (inv p) [x\<leftarrow>xs'. x \<in> range p] \<in> divergences P \<and>
map (inv q) [x\<leftarrow>xs'. x \<in> range q] \<in> divergences Q \<Longrightarrow>
(\<exists>R S T.
ipurge_ref (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) zs Z = R \<union> S \<union> T \<and>
set xs \<subseteq> range p \<union> range q \<and>
set (ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) zs) \<subseteq> range p \<union> range q \<and>
R \<subseteq> range p \<and>
S \<subseteq> range q \<and>
T \<subseteq> - range p \<and>
T \<subseteq> - range q \<and>
(map (inv p) [x\<leftarrow>xs @ y # ipurge_tr (con_comp_pol I)
(con_comp_map D E p q) (con_comp_map D E p q y) zs. x \<in> range p],
inv p ` R) \<in> failures P \<and>
(map (inv q) [x\<leftarrow>xs @ y # ipurge_tr (con_comp_pol I)
(con_comp_map D E p q) (con_comp_map D E p q y) zs. x \<in> range q],
inv q ` S) \<in> failures Q) \<or>
(\<exists>xs'.
(\<exists>ys'. xs @ y # ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) zs = xs' @ ys') \<and>
set xs' \<subseteq> range p \<union> range q \<and>
map (inv p) [x\<leftarrow>xs'. x \<in> range p] \<in> divergences P \<and>
map (inv q) [x\<leftarrow>xs'. x \<in> range q] \<in> divergences Q)"
(is "_ \<Longrightarrow> (\<exists>R S T. ?F R S T zs) \<or> ?G")
proof (erule exE, (erule conjE)+, erule exE)
fix xs' ys'
assume
F: "set xs' \<subseteq> range p \<union> range q" and
G: "map (inv p) [x\<leftarrow>xs'. x \<in> range p] \<in> divergences P" and
H: "map (inv q) [x\<leftarrow>xs'. x \<in> range q] \<in> divergences Q" and
I: "xs @ zs = xs' @ ys'"
show ?thesis
proof (cases "length xs < length xs'", rule disjI1, rule_tac [2] disjI2)
case True
moreover have "take (length xs') (xs @ zs) =
take (length xs') xs @ take (length xs' - length xs) zs"
by simp
ultimately have "take (length xs') (xs @ zs) =
xs @ take (length xs' - length xs) zs"
(is "_ = _ @ ?vs")
by simp
moreover have "take (length xs') (xs @ zs) =
take (length xs') (xs' @ ys')"
using I by simp
ultimately have J: "xs @ ?vs = xs'"
by simp
moreover have "set xs \<subseteq> set (xs @ ?vs)"
by simp
ultimately have "set xs \<subseteq> set xs'"
by simp
hence K: "set xs \<subseteq> range p \<union> range q"
using F by simp
have "\<exists>R S T. ?F R S T [x\<leftarrow>zs. x \<in> range p \<union> range q]"
(is "\<exists>_ _ _. ipurge_ref ?I' ?D' _ _ _ = _ \<and> _")
proof (rule con_comp_secure_add_case_1 [OF A B C D E],
rule_tac x = "range p \<inter> Z" in exI, rule_tac x = "range q \<inter> Z" in exI,
rule_tac x = "- range p \<inter> - range q \<inter> Z" in exI,
(subst conj_assoc [symmetric])+, (rule conjI)+, simp_all del: filter_append)
show "Z = range p \<inter> Z \<union> range q \<inter> Z \<union> - range p \<inter> - range q \<inter> Z"
by blast
next
show "set xs \<subseteq> range p \<union> range q"
using K .
next
show "{x \<in> set zs. x \<in> range p \<or> x \<in> range q} \<subseteq> range p \<union> range q"
by blast
next
show "- range p \<inter> - range q \<inter> Z \<subseteq> - range p"
by blast
next
show "- range p \<inter> - range q \<inter> Z \<subseteq> - range q"
by blast
next
have "map (inv p) [x\<leftarrow>xs @ ?vs. x \<in> range p] \<in> divergences P"
using G and J by simp
hence "map (inv p) [x\<leftarrow>xs @ ?vs. x \<in> range p] @
map (inv p) [x\<leftarrow>drop (length xs' - length xs) zs. x \<in> range p]
\<in> divergences P"
(is "_ @ map (inv p) [x\<leftarrow>?ws. _] \<in> _")
by (rule process_rule_5_general)
hence "map (inv p) [x\<leftarrow>(xs @ ?vs) @ ?ws. x \<in> range p] \<in> divergences P"
by (subst filter_append, simp)
hence "map (inv p) [x\<leftarrow>xs @ zs. x \<in> range p] \<in> divergences P"
by simp
hence "map (inv p) [x\<leftarrow>xs @ [x\<leftarrow>zs. x \<in> range p \<or> x \<in> range q].
x \<in> range p] \<in> divergences P"
proof (subst (asm) filter_append, subst filter_append, subst filter_filter)
qed (subgoal_tac "(\<lambda>x. (x \<in> range p \<or> x \<in> range q) \<and> x \<in> range p) =
(\<lambda>x. x \<in> range p)", simp, blast)
thus "(map (inv p) [x\<leftarrow>xs @ [x\<leftarrow>zs. x \<in> range p \<or> x \<in> range q].
x \<in> range p], inv p ` (range p \<inter> Z)) \<in> failures P"
by (rule process_rule_6)
next
have "map (inv q) [x\<leftarrow>xs @ ?vs. x \<in> range q] \<in> divergences Q"
using H and J by simp
hence "map (inv q) [x\<leftarrow>xs @ ?vs. x \<in> range q] @
map (inv q) [x\<leftarrow>drop (length xs' - length xs) zs. x \<in> range q]
\<in> divergences Q"
(is "_ @ map (inv q) [x\<leftarrow>?ws. _] \<in> _")
by (rule process_rule_5_general)
hence "map (inv q) [x\<leftarrow>(xs @ ?vs) @ ?ws. x \<in> range q] \<in> divergences Q"
by (subst filter_append, simp)
hence "map (inv q) [x\<leftarrow>xs @ zs. x \<in> range q] \<in> divergences Q"
by simp
hence "map (inv q) [x\<leftarrow>xs @ [x\<leftarrow>zs. x \<in> range p \<or> x \<in> range q].
x \<in> range q] \<in> divergences Q"
proof (subst (asm) filter_append, subst filter_append, subst filter_filter)
qed (subgoal_tac "(\<lambda>x. (x \<in> range p \<or> x \<in> range q) \<and> x \<in> range q) =
(\<lambda>x. x \<in> range q)", simp, blast)
thus "(map (inv q) [x\<leftarrow>xs @ [x\<leftarrow>zs. x \<in> range p \<or> x \<in> range q].
x \<in> range q], inv q ` (range q \<inter> Z)) \<in> failures Q"
by (rule process_rule_6)
qed
then obtain R and S and T where
"?F R S T [x\<leftarrow>zs. x \<in> range p \<union> range q]"
by blast
thus "\<exists>R S T. ?F R S T zs"
proof (rule_tac x = R in exI, rule_tac x = S in exI, rule_tac x = T in exI)
qed (simp only: con_comp_ipurge_tr_filter con_comp_ipurge_ref_filter)
next
let
?I' = "con_comp_pol I" and
?D' = "con_comp_map D E p q"
case False
moreover have "xs @ y # ipurge_tr ?I' ?D' (?D' y) zs =
take (length xs') (xs @ y # ipurge_tr ?I' ?D' (?D' y) zs) @
drop (length xs') (xs @ y # ipurge_tr ?I' ?D' (?D' y) zs)"
(is "_ = _ @ ?vs")
by (simp only: append_take_drop_id)
ultimately have "xs @ y # ipurge_tr ?I' ?D' (?D' y) zs =
take (length xs') (xs @ zs) @ ?vs"
by simp
hence J: "xs @ y # ipurge_tr ?I' ?D' (?D' y) zs = xs' @ ?vs"
using I by simp
show ?G
proof (rule_tac x = xs' in exI, rule conjI, rule_tac x = ?vs in exI)
qed (subst J, simp_all add: F G H)
qed
qed
theorem con_comp_secure:
assumes
A: "consistent_maps D E p q" and
B: "secure P I D" and
C: "secure Q I E"
shows "secure (P \<parallel> Q <p, q>) (con_comp_pol I) (con_comp_map D E p q)"
proof (simp add: secure_def con_comp_futures, (rule allI)+, rule impI,
erule conjE, rule conjI, (erule rev_mp)+, rotate_tac [2], erule_tac [2] rev_mp)
fix xs y ys Y zs Z
show
"(xs @ zs, Z) \<in> con_comp_failures P Q p q \<longrightarrow>
(xs @ y # ys, Y) \<in> con_comp_failures P Q p q \<longrightarrow>
(xs @ ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys,
ipurge_ref (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) ys Y)
\<in> con_comp_failures P Q p q"
(is "_ \<longrightarrow> _ \<longrightarrow> (_ @ ipurge_tr ?I' ?D' _ _, _) \<in> _")
proof ((rule impI)+, thin_tac "(xs @ zs, Z) \<in> con_comp_failures P Q p q",
simp_all add: con_comp_failures_def con_comp_divergences_def
del: filter_append, erule disjE, rule disjI1)
qed (erule con_comp_secure_del_case_1 [OF A B C],
rule con_comp_secure_del_case_2 [OF A B C])
next
fix xs y ys Y zs Z
assume D: "(xs @ y # ys, Y) \<in> con_comp_failures P Q p q"
show
"(xs @ zs, Z) \<in> con_comp_failures P Q p q \<longrightarrow>
(xs @ y # ipurge_tr (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) zs,
ipurge_ref (con_comp_pol I) (con_comp_map D E p q)
(con_comp_map D E p q y) zs Z)
\<in> con_comp_failures P Q p q"
(is "_ \<longrightarrow> (_ @ _ # ipurge_tr ?I' ?D' _ _, _) \<in> _")
proof (rule impI, simp_all add: con_comp_failures_def con_comp_divergences_def
del: filter_append, cases "y \<in> range p \<or> y \<in> range q", simp del: filter_append,
erule disjE, rule disjI1, rule_tac [3] disjI2)
qed (erule con_comp_secure_add_case_1 [OF A B C D], assumption,
erule con_comp_secure_add_case_2 [OF A B C D], assumption,
rule con_comp_failures_divergences [OF D], simp_all)
qed
subsection "Conservation of noninterference security in the absence of fake events"
text \<open>
In what follows, it is proven that in the absence of fake events, namely if
@{term "range p \<union> range q = UNIV"}, the output of the concurrent composition of two secure processes
is secure with respect to the same noninterference policy enforced by the input processes, and to
the event-domain map that simply associates each event to the same security domain as the
corresponding events of the input processes.
More formally, for any two processes @{term P}, @{term Q} being secure with respect to the
noninterference policy @{term I} and the event-domain maps @{term D}, @{term E}, their concurrent
composition @{term "P \<parallel> Q <p, q>"} is secure with respect to the same noninterference policy
@{term I} and the event-domain map @{term "the \<circ> con_comp_map D E p q"}, provided that conditions
@{term "range p \<union> range q = UNIV"} and @{term "consistent_maps D E p q"} are satisfied.
\null
\<close>
lemma con_comp_sinks_range:
"u \<in> range Some \<Longrightarrow>
set xs \<subseteq> range p \<union> range q \<Longrightarrow>
sinks (con_comp_pol I) (con_comp_map D E p q) u xs \<subseteq> range Some"
by (insert con_comp_sinks_aux_range [of "{u}" xs p q I D E],
simp add: sinks_aux_single_dom)
lemma con_comp_sinks_no_fake:
assumes
A: "range p \<union> range q = UNIV" and
B: "u \<in> range Some"
shows "sinks I (the \<circ> con_comp_map D E p q) (the u) xs =
the ` sinks (con_comp_pol I) (con_comp_map D E p q) u xs"
(is "_ = the ` sinks ?I' ?D' _ _")
proof (induction xs rule: rev_induct, simp)
fix x xs
assume C: "sinks I (the \<circ> ?D') (the u) xs = the ` sinks ?I' ?D' u xs"
have "x \<in> range p \<union> range q"
using A by simp
hence D: "?D' x = Some (the (?D' x))"
by (cases "x \<in> range p", simp_all)
have E: "u = Some (the u)"
using B by (simp add: image_iff)
show "sinks I (the \<circ> ?D') (the u) (xs @ [x]) = the ` sinks ?I' ?D' u (xs @ [x])"
proof (cases "(u, ?D' x) \<in> ?I' \<or> (\<exists>v \<in> sinks ?I' ?D' u xs. (v, ?D' x) \<in> ?I')")
case True
hence "sinks ?I' ?D' u (xs @ [x]) = insert (?D' x) (sinks ?I' ?D' u xs)"
by simp
moreover have "(the u, the (?D' x)) \<in> I \<or>
(\<exists>d \<in> sinks I (the \<circ> ?D') (the u) xs. (d, the (?D' x)) \<in> I)"
proof (rule disjE [OF True], rule disjI1, rule_tac [2] disjI2)
assume "(u, ?D' x) \<in> ?I'"
hence "(Some (the u), Some (the (?D' x))) \<in> ?I'"
using D and E by simp
thus "(the u, the (?D' x)) \<in> I"
by (simp add: con_comp_pol_def)
next
assume "\<exists>v \<in> sinks ?I' ?D' u xs. (v, ?D' x) \<in> ?I'"
then obtain v where F: "v \<in> sinks ?I' ?D' u xs" and G: "(v, ?D' x) \<in> ?I'" ..
have "sinks ?I' ?D' u xs \<subseteq> range Some"
by (rule con_comp_sinks_range, simp_all add: A B)
hence "v \<in> range Some"
using F ..
hence "v = Some (the v)"
by (simp add: image_iff)
hence "(Some (the v), Some (the (?D' x))) \<in> ?I'"
using D and G by simp
hence "(the v, the (?D' x)) \<in> I"
by (simp add: con_comp_pol_def)
moreover have "the v \<in> sinks I (the \<circ> ?D') (the u) xs"
using C and F by simp
ultimately show "\<exists>d \<in> sinks I (the \<circ> ?D') (the u) xs.
(d, the (?D' x)) \<in> I" ..
qed
hence "sinks I (the \<circ> ?D') (the u) (xs @ [x]) =
insert (the (?D' x)) (sinks I (the \<circ> ?D') (the u) xs)"
by simp
ultimately show ?thesis
using C by simp
next
case False
hence "sinks ?I' ?D' u (xs @ [x]) = sinks ?I' ?D' u xs"
by simp
moreover have "\<not> ((the u, the (?D' x)) \<in> I \<or>
(\<exists>v \<in> sinks I (the \<circ> ?D') (the u) xs. (v, the (?D' x)) \<in> I))"
proof (insert False, simp, erule conjE, rule conjI, rule_tac [2] ballI)
assume "(u, ?D' x) \<notin> ?I'"
hence "(Some (the u), Some (the (?D' x))) \<notin> ?I'"
using D and E by simp
thus "(the u, the (?D' x)) \<notin> I"
by (simp add: con_comp_pol_def)
next
fix d
assume "d \<in> sinks I (the \<circ> ?D') (the u) xs"
hence "d \<in> the ` sinks ?I' ?D' u xs"
using C by simp
hence "\<exists>v \<in> sinks ?I' ?D' u xs. d = the v"
by (simp add: image_iff)
then obtain v where F: "v \<in> sinks ?I' ?D' u xs" and G: "d = the v" ..
have "sinks ?I' ?D' u xs \<subseteq> range Some"
by (rule con_comp_sinks_range, simp_all add: A B)
hence "v \<in> range Some"
using F ..
hence H: "v = Some d"
using G by (simp add: image_iff)
assume "\<forall>v \<in> sinks ?I' ?D' u xs. (v, ?D' x) \<notin> ?I'"
hence "(v, ?D' x) \<notin> ?I'"
using F ..
hence "(Some d, Some (the (?D' x))) \<notin> ?I'"
using D and H by simp
thus "(d, the (?D' x)) \<notin> I"
by (simp add: con_comp_pol_def)
qed
hence "sinks I (the \<circ> ?D') (the u) (xs @ [x]) = sinks I (the \<circ> ?D') (the u) xs"
by simp
ultimately show ?thesis
using C by simp
qed
qed
lemma con_comp_ipurge_tr_no_fake:
assumes
A: "range p \<union> range q = UNIV" and
B: "u \<in> range Some"
shows "ipurge_tr (con_comp_pol I) (con_comp_map D E p q) u xs =
ipurge_tr I (the \<circ> con_comp_map D E p q) (the u) xs"
(is "ipurge_tr ?I' ?D' _ _ = _")
proof (induction xs rule: rev_induct, simp)
fix x xs
assume C: "ipurge_tr ?I' ?D' u xs = ipurge_tr I (the \<circ> ?D') (the u) xs"
show "ipurge_tr ?I' ?D' u (xs @ [x]) = ipurge_tr I (the \<circ> ?D') (the u) (xs @ [x])"
proof (cases "?D' x \<in> sinks ?I' ?D' u (xs @ [x])")
case True
hence "ipurge_tr ?I' ?D' u (xs @ [x]) = ipurge_tr ?I' ?D' u xs"
by simp
moreover have "the (?D' x) \<in> the ` sinks ?I' ?D' u (xs @ [x])"
using True by simp
hence "the (?D' x) \<in> sinks I (the \<circ> ?D') (the u) (xs @ [x])"
by (subst con_comp_sinks_no_fake [OF A B])
hence "ipurge_tr I (the \<circ> ?D') (the u) (xs @ [x]) =
ipurge_tr I (the \<circ> ?D') (the u) xs"
by simp
ultimately show ?thesis
using C by simp
next
case False
hence "ipurge_tr ?I' ?D' u (xs @ [x]) = ipurge_tr ?I' ?D' u xs @ [x]"
by simp
moreover have "the (?D' x) \<notin> the ` sinks ?I' ?D' u (xs @ [x])"
proof
assume "the (?D' x) \<in> the ` sinks ?I' ?D' u (xs @ [x])"
hence "\<exists>v \<in> sinks ?I' ?D' u (xs @ [x]). the (?D' x) = the v"
by (simp add: image_iff)
then obtain v where
D: "v \<in> sinks ?I' ?D' u (xs @ [x])" and E: "the (?D' x) = the v" ..
have "x \<in> range p \<union> range q"
using A by simp
hence "\<exists>d. ?D' x = Some d"
by (cases "x \<in> range p", simp_all)
then obtain d where "?D' x = Some d" ..
moreover have "sinks ?I' ?D' u (xs @ [x]) \<subseteq> range Some"
by (rule con_comp_sinks_range, simp_all add: A B)
hence "v \<in> range Some"
using D ..
hence "\<exists>d'. v = Some d'"
by (simp add: image_iff)
then obtain d' where "v = Some d'" ..
ultimately have "?D' x = v"
using E by simp
hence "?D' x \<in> sinks ?I' ?D' u (xs @ [x])"
using D by simp
thus False
using False by contradiction
qed
hence "the (?D' x) \<notin> sinks I (the \<circ> ?D') (the u) (xs @ [x])"
by (subst con_comp_sinks_no_fake [OF A B])
hence "ipurge_tr I (the \<circ> ?D') (the u) (xs @ [x]) =
ipurge_tr I (the \<circ> ?D') (the u) xs @ [x]"
by simp
ultimately show ?thesis
using C by simp
qed
qed
lemma con_comp_ipurge_ref_no_fake:
assumes
A: "range p \<union> range q = UNIV" and
B: "u \<in> range Some"
shows "ipurge_ref (con_comp_pol I) (con_comp_map D E p q) u xs X =
ipurge_ref I (the \<circ> con_comp_map D E p q) (the u) xs X"
(is "ipurge_ref ?I' ?D' _ _ _ = _")
proof (simp add: ipurge_ref_def set_eq_iff, rule allI,
simp_all add: con_comp_sinks_no_fake [OF A B])
fix x
have "x \<in> range p \<union> range q"
using A by simp
hence C: "?D' x = Some (the (?D' x))"
by (cases "x \<in> range p", simp_all)
have D: "u = Some (the u)"
using B by (simp add: image_iff)
show
"(x \<in> X \<and> (u, ?D' x) \<notin> con_comp_pol I \<and>
(\<forall>v \<in> sinks ?I' ?D' u xs. (v, ?D' x) \<notin> con_comp_pol I)) =
(x \<in> X \<and> (the u, the (?D' x)) \<notin> I \<and>
(\<forall>v \<in> sinks ?I' ?D' u xs. (the v, the (?D' x)) \<notin> I))"
proof (rule iffI, (erule_tac [!] conjE)+, simp_all, rule_tac [!] conjI,
rule_tac [2] ballI, rule_tac [4] ballI)
assume "(u, ?D' x) \<notin> ?I'"
hence "(Some (the u), Some (the (?D' x))) \<notin> ?I'"
using C and D by simp
thus "(the u, the (?D' x)) \<notin> I"
by (simp add: con_comp_pol_def)
next
fix v
assume "\<forall>v \<in> sinks ?I' ?D' u xs. (v, ?D' x) \<notin> ?I'" and
E: "v \<in> sinks ?I' ?D' u xs"
hence "(v, ?D' x) \<notin> ?I'" ..
moreover have "sinks ?I' ?D' u xs \<subseteq> range Some"
by (rule con_comp_sinks_range, simp_all add: A B)
hence "v \<in> range Some"
using E ..
hence "v = Some (the v)"
by (simp add: image_iff)
ultimately have "(Some (the v), Some (the (?D' x))) \<notin> ?I'"
using C by simp
thus "(the v, the (?D' x)) \<notin> I"
by (simp add: con_comp_pol_def)
next
assume "(the u, the (?D' x)) \<notin> I"
hence "(Some (the u), Some (the (?D' x))) \<notin> ?I'"
by (simp add: con_comp_pol_def)
thus "(u, ?D' x) \<notin> ?I'"
using C and D by simp
next
fix v
assume "\<forall>v \<in> sinks ?I' ?D' u xs. (the v, the (?D' x)) \<notin> I" and
E: "v \<in> sinks ?I' ?D' u xs"
hence "(the v, the (?D' x)) \<notin> I" ..
hence "(Some (the v), Some (the (?D' x))) \<notin> ?I'"
by (simp add: con_comp_pol_def)
moreover have "sinks ?I' ?D' u xs \<subseteq> range Some"
by (rule con_comp_sinks_range, simp_all add: A B)
hence "v \<in> range Some"
using E ..
hence "v = Some (the v)"
by (simp add: image_iff)
ultimately show "(v, ?D' x) \<notin> ?I'"
using C by simp
qed
qed
theorem con_comp_secure_no_fake:
assumes
A: "range p \<union> range q = UNIV" and
B: "consistent_maps D E p q" and
C: "secure P I D" and
D: "secure Q I E"
shows "secure (P \<parallel> Q <p, q>) I (the \<circ> con_comp_map D E p q)"
proof (insert con_comp_secure [OF B C D], simp add: secure_def,
(rule allI)+, rule impI)
fix xs y ys Y zs Z
let
?I' = "con_comp_pol I" and
?D' = "con_comp_map D E p q"
have "y \<in> range p \<union> range q"
using A by simp
hence E: "?D' y \<in> range Some"
by (cases "y \<in> range p", simp_all)
assume "\<forall>xs y ys Y zs Z.
(y # ys, Y) \<in> futures (P \<parallel> Q <p, q>) xs \<and>
(zs, Z) \<in> futures (P \<parallel> Q <p, q>) xs \<longrightarrow>
(ipurge_tr ?I' ?D' (?D' y) ys, ipurge_ref ?I' ?D' (?D' y) ys Y)
\<in> futures (P \<parallel> Q <p, q>) xs \<and>
(y # ipurge_tr ?I' ?D' (?D' y) zs, ipurge_ref ?I' ?D' (?D' y) zs Z)
\<in> futures (P \<parallel> Q <p, q>) xs" and
"(y # ys, Y) \<in> futures (P \<parallel> Q <p, q>) xs \<and>
(zs, Z) \<in> futures (P \<parallel> Q <p, q>) xs"
hence
"(ipurge_tr ?I' ?D' (?D' y) ys, ipurge_ref ?I' ?D' (?D' y) ys Y)
\<in> futures (P \<parallel> Q <p, q>) xs \<and>
(y # ipurge_tr ?I' ?D' (?D' y) zs, ipurge_ref ?I' ?D' (?D' y) zs Z)
\<in> futures (P \<parallel> Q <p, q>) xs"
by blast
thus
"(ipurge_tr I (the \<circ> ?D') (the (?D' y)) ys,
ipurge_ref I (the \<circ> ?D') (the (?D' y)) ys Y) \<in> futures (P \<parallel> Q <p, q>) xs \<and>
(y # ipurge_tr I (the \<circ> ?D') (the (?D' y)) zs,
ipurge_ref I (the \<circ> ?D') (the (?D' y)) zs Z) \<in> futures (P \<parallel> Q <p, q>) xs"
by (simp add: con_comp_ipurge_tr_no_fake [OF A E]
con_comp_ipurge_ref_no_fake [OF A E])
qed
end
diff --git a/thys/NormByEval/NBE.thy b/thys/NormByEval/NBE.thy
--- a/thys/NormByEval/NBE.thy
+++ b/thys/NormByEval/NBE.thy
@@ -1,2491 +1,2490 @@
(* Author: Klaus Aehlig, Tobias Nipkow
Normalization by Evaluation.
*)
(*<*)
theory NBE imports Main begin
declare [[syntax_ambiguity_warning = false]]
declare Let_def[simp]
(*>*)
section "Terms"
type_synonym vname = nat
type_synonym ml_vname = nat
(* FIXME only for codegen*)
type_synonym cname = int
text\<open>ML terms:\<close>
datatype ml =
\<comment> \<open>ML\<close>
C_ML cname ("C\<^sub>M\<^sub>L") (* ref to compiled code *)
| V_ML ml_vname ("V\<^sub>M\<^sub>L")
| A_ML ml "(ml list)" ("A\<^sub>M\<^sub>L")
| Lam_ML ml ("Lam\<^sub>M\<^sub>L")
\<comment> \<open>the universal datatype\<close>
| C\<^sub>U cname "(ml list)"
| V\<^sub>U vname "(ml list)"
| Clo ml "(ml list)" nat
\<comment> \<open>ML function \emph{apply}\<close>
| "apply" ml ml
text\<open>Lambda-terms:\<close>
datatype tm = C cname | V vname | \<Lambda> tm | At tm tm (infix "\<bullet>" 100)
| "term" ml \<comment> \<open>ML function \texttt{term}\<close>
text \<open>The following locale captures type conventions for variables.
It is not actually used, merely a formal comment.\<close>
locale Vars =
fixes r s t:: tm
and rs ss ts :: "tm list"
and u v w :: ml
and us vs ws :: "ml list"
and nm :: cname
and x :: vname
and X :: ml_vname
text\<open>The subset of pure terms:\<close>
inductive pure :: "tm \<Rightarrow> bool" where
"pure(C nm)" |
"pure(V x)" |
Lam: "pure t \<Longrightarrow> pure(\<Lambda> t)" |
"pure s \<Longrightarrow> pure t \<Longrightarrow> pure(s\<bullet>t)"
declare pure.intros[simp]
declare Lam[simp del]
lemma pure_Lam[simp]: "pure(\<Lambda> t) = pure t"
proof
assume "pure(\<Lambda> t)" thus "pure t"
proof cases qed auto
next
assume "pure t" thus "pure(\<Lambda> t)" by(rule Lam)
qed
text\<open>Closed terms w.r.t.\ ML variables:\<close>
fun closed_ML :: "nat \<Rightarrow> ml \<Rightarrow> bool" ("closed\<^sub>M\<^sub>L") where
"closed\<^sub>M\<^sub>L i (C\<^sub>M\<^sub>L nm) = True" |
"closed\<^sub>M\<^sub>L i (V\<^sub>M\<^sub>L X) = (X<i)" |
"closed\<^sub>M\<^sub>L i (A\<^sub>M\<^sub>L v vs) = (closed\<^sub>M\<^sub>L i v \<and> (\<forall>v \<in> set vs. closed\<^sub>M\<^sub>L i v))" |
"closed\<^sub>M\<^sub>L i (Lam\<^sub>M\<^sub>L v) = closed\<^sub>M\<^sub>L (i+1) v" |
"closed\<^sub>M\<^sub>L i (C\<^sub>U nm vs) = (\<forall>v \<in> set vs. closed\<^sub>M\<^sub>L i v)" |
"closed\<^sub>M\<^sub>L i (V\<^sub>U nm vs) = (\<forall>v \<in> set vs. closed\<^sub>M\<^sub>L i v)" |
"closed\<^sub>M\<^sub>L i (Clo f vs n) = (closed\<^sub>M\<^sub>L i f \<and> (\<forall>v \<in> set vs. closed\<^sub>M\<^sub>L i v))" |
"closed\<^sub>M\<^sub>L i (apply v w) = (closed\<^sub>M\<^sub>L i v \<and> closed\<^sub>M\<^sub>L i w)"
fun closed_tm_ML :: "nat \<Rightarrow> tm \<Rightarrow> bool" ("closed\<^sub>M\<^sub>L") where
"closed_tm_ML i (r\<bullet>s) = (closed_tm_ML i r \<and> closed_tm_ML i s)" |
"closed_tm_ML i (\<Lambda> t) = (closed_tm_ML i t)" |
"closed_tm_ML i (term v) = closed_ML i v" |
"closed_tm_ML i v = True"
text\<open>Free variables:\<close>
fun fv_ML :: "ml \<Rightarrow> ml_vname set" ("fv\<^sub>M\<^sub>L") where
"fv\<^sub>M\<^sub>L (C\<^sub>M\<^sub>L nm) = {}" |
"fv\<^sub>M\<^sub>L (V\<^sub>M\<^sub>L X) = {X}" |
"fv\<^sub>M\<^sub>L (A\<^sub>M\<^sub>L v vs) = fv\<^sub>M\<^sub>L v \<union> (\<Union>v \<in> set vs. fv\<^sub>M\<^sub>L v)" |
"fv\<^sub>M\<^sub>L (Lam\<^sub>M\<^sub>L v) = {X. Suc X : fv\<^sub>M\<^sub>L v}" |
"fv\<^sub>M\<^sub>L (C\<^sub>U nm vs) = (\<Union>v \<in> set vs. fv\<^sub>M\<^sub>L v)" |
"fv\<^sub>M\<^sub>L (V\<^sub>U nm vs) = (\<Union>v \<in> set vs. fv\<^sub>M\<^sub>L v)" |
"fv\<^sub>M\<^sub>L (Clo f vs n) = fv\<^sub>M\<^sub>L f \<union> (\<Union>v \<in> set vs. fv\<^sub>M\<^sub>L v)" |
"fv\<^sub>M\<^sub>L (apply v w) = fv\<^sub>M\<^sub>L v \<union> fv\<^sub>M\<^sub>L w"
primrec fv :: "tm \<Rightarrow> vname set" where
"fv (C nm) = {}" |
"fv (V X) = {X}" |
"fv (s \<bullet> t) = fv s \<union> fv t" |
"fv (\<Lambda> t) = {X. Suc X : fv t}"
subsection "Iterated Term Application"
abbreviation foldl_At (infix "\<bullet>\<bullet>" 90) where
"t \<bullet>\<bullet> ts \<equiv> foldl (\<bullet>) t ts"
text\<open>Auxiliary measure function:\<close>
primrec depth_At :: "tm \<Rightarrow> nat"
where
"depth_At(C nm) = 0"
| "depth_At(V x) = 0"
| "depth_At(s \<bullet> t) = depth_At s + 1"
| "depth_At(\<Lambda> t) = 0"
| "depth_At(term v) = 0"
lemma depth_At_foldl:
"depth_At(s \<bullet>\<bullet> ts) = depth_At s + size ts"
by (induct ts arbitrary: s) simp_all
lemma foldl_At_eq_lemma: "size ts = size ts' \<Longrightarrow>
s \<bullet>\<bullet> ts = s' \<bullet>\<bullet> ts' \<longleftrightarrow> s = s' \<and> ts = ts'"
by (induct arbitrary: s s' rule:list_induct2) simp_all
lemma foldl_At_eq_length:
"s \<bullet>\<bullet> ts = s \<bullet>\<bullet> ts' \<Longrightarrow> length ts = length ts'"
apply(subgoal_tac "depth_At(s \<bullet>\<bullet> ts) = depth_At(s \<bullet>\<bullet> ts')")
apply(erule thin_rl)
apply (simp add:depth_At_foldl)
apply simp
done
lemma foldl_At_eq[simp]: "s \<bullet>\<bullet> ts = s \<bullet>\<bullet> ts' \<longleftrightarrow> ts = ts'"
apply(rule)
prefer 2 apply simp
apply(blast dest:foldl_At_eq_lemma foldl_At_eq_length)
done
lemma term_eq_foldl_At[simp]:
"term v = t \<bullet>\<bullet> ts \<longleftrightarrow> t = term v \<and> ts = []"
by (induct ts arbitrary:t) auto
lemma At_eq_foldl_At[simp]:
"r \<bullet> s = t \<bullet>\<bullet> ts \<longleftrightarrow>
(if ts=[] then t = r \<bullet> s else s = last ts \<and> r = t \<bullet>\<bullet> butlast ts)"
apply (induct ts arbitrary:t)
apply fastforce
apply rule
apply clarsimp
apply rule
apply clarsimp
apply clarsimp
apply(subgoal_tac "\<exists>ts' t'. ts = ts' @ [t']")
apply clarsimp
defer
apply (clarsimp split:list.split)
apply (metis append_butlast_last_id)
done
lemma foldl_At_eq_At[simp]:
"t \<bullet>\<bullet> ts = r \<bullet> s \<longleftrightarrow>
(if ts=[] then t = r \<bullet> s else s = last ts \<and> r = t \<bullet>\<bullet> butlast ts)"
by(metis At_eq_foldl_At)
lemma Lam_eq_foldl_At[simp]:
"\<Lambda> s = t \<bullet>\<bullet> ts \<longleftrightarrow> t = \<Lambda> s \<and> ts = []"
by (induct ts arbitrary:t) auto
lemma foldl_At_eq_Lam[simp]:
"t \<bullet>\<bullet> ts = \<Lambda> s \<longleftrightarrow> t = \<Lambda> s \<and> ts = []"
by (induct ts arbitrary:t) auto
lemma [simp]: "s \<bullet> t \<noteq> s"
apply(subgoal_tac "size(s \<bullet> t) \<noteq> size s")
apply metis
apply simp
done
(* Better: a simproc for disproving "s = t"
if s is a subterm of t or vice versa, by proving "size s ~= size t"
*)
fun atomic_tm :: "tm \<Rightarrow> bool" where
"atomic_tm(s \<bullet> t) = False" |
"atomic_tm(_) = True"
fun head_tm where
"head_tm(s \<bullet> t) = head_tm s" |
"head_tm(s) = s"
fun args_tm where
"args_tm(s \<bullet> t) = args_tm s @ [t]" |
"args_tm(_) = []"
lemma head_tm_foldl_At[simp]: "head_tm(s \<bullet>\<bullet> ts) = head_tm s"
by(induct ts arbitrary: s) auto
lemma args_tm_foldl_At[simp]: "args_tm(s \<bullet>\<bullet> ts) = args_tm s @ ts"
by(induct ts arbitrary: s) auto
lemma tm_eq_iff:
"atomic_tm(head_tm s) \<Longrightarrow> atomic_tm(head_tm t)
\<Longrightarrow> s = t \<longleftrightarrow> head_tm s = head_tm t \<and> args_tm s = args_tm t"
apply(induct s arbitrary: t)
apply(case_tac t, simp+)+
done
declare
tm_eq_iff[of "h \<bullet>\<bullet> ts", simp]
tm_eq_iff[of _ "h \<bullet>\<bullet> ts", simp]
for h ts
lemma atomic_tm_head_tm: "atomic_tm(head_tm t)"
by(induct t) auto
lemma head_tm_idem: "head_tm(head_tm t) = head_tm t"
by(induct t) auto
lemma args_tm_head_tm: "args_tm(head_tm t) = []"
by(induct t) auto
lemma eta_head_args: "t = head_tm t \<bullet>\<bullet> args_tm t"
by (subst tm_eq_iff) (auto simp: atomic_tm_head_tm head_tm_idem args_tm_head_tm)
lemma tm_vector_cases:
"(\<exists>n ts. t = V n \<bullet>\<bullet> ts) \<or>
(\<exists>nm ts. t = C nm \<bullet>\<bullet> ts) \<or>
(\<exists>t' ts. t = \<Lambda> t' \<bullet>\<bullet> ts) \<or>
(\<exists>v ts. t = term v \<bullet>\<bullet> ts)"
apply(induct t)
apply simp_all
by (metis snoc_eq_iff_butlast)
lemma fv_head_C[simp]: "fv (t \<bullet>\<bullet> ts) = fv t \<union> (\<Union>t\<in>set ts. fv t)"
by(induct ts arbitrary:t) auto
subsection "Lifting and Substitution"
fun lift_ml :: "nat \<Rightarrow> ml \<Rightarrow> ml" ("lift") where
"lift i (C\<^sub>M\<^sub>L nm) = C\<^sub>M\<^sub>L nm" |
"lift i (V\<^sub>M\<^sub>L X) = V\<^sub>M\<^sub>L X" |
"lift i (A\<^sub>M\<^sub>L v vs) = A\<^sub>M\<^sub>L (lift i v) (map (lift i) vs)" |
"lift i (Lam\<^sub>M\<^sub>L v) = Lam\<^sub>M\<^sub>L (lift i v)" |
"lift i (C\<^sub>U nm vs) = C\<^sub>U nm (map (lift i) vs)" |
"lift i (V\<^sub>U x vs) = V\<^sub>U (if x < i then x else x+1) (map (lift i) vs)" |
"lift i (Clo v vs n) = Clo (lift i v) (map (lift i) vs) n" |
"lift i (apply u v) = apply (lift i u) (lift i v)"
lemmas ml_induct = lift_ml.induct[of "\<lambda>i v. P v"] for P
fun lift_tm :: "nat \<Rightarrow> tm \<Rightarrow> tm" ("lift") where
"lift i (C nm) = C nm" |
"lift i (V x) = V(if x < i then x else x+1)" |
"lift i (s\<bullet>t) = (lift i s)\<bullet>(lift i t)" |
"lift i (\<Lambda> t) = \<Lambda>(lift (i+1) t)" |
"lift i (term v) = term (lift i v)"
fun lift_ML :: "nat \<Rightarrow> ml \<Rightarrow> ml" ("lift\<^sub>M\<^sub>L") where
"lift\<^sub>M\<^sub>L i (C\<^sub>M\<^sub>L nm) = C\<^sub>M\<^sub>L nm" |
"lift\<^sub>M\<^sub>L i (V\<^sub>M\<^sub>L X) = V\<^sub>M\<^sub>L (if X < i then X else X+1)" |
"lift\<^sub>M\<^sub>L i (A\<^sub>M\<^sub>L v vs) = A\<^sub>M\<^sub>L (lift\<^sub>M\<^sub>L i v) (map (lift\<^sub>M\<^sub>L i) vs)" |
"lift\<^sub>M\<^sub>L i (Lam\<^sub>M\<^sub>L v) = Lam\<^sub>M\<^sub>L (lift\<^sub>M\<^sub>L (i+1) v)" |
"lift\<^sub>M\<^sub>L i (C\<^sub>U nm vs) = C\<^sub>U nm (map (lift\<^sub>M\<^sub>L i) vs)" |
"lift\<^sub>M\<^sub>L i (V\<^sub>U x vs) = V\<^sub>U x (map (lift\<^sub>M\<^sub>L i) vs)" |
"lift\<^sub>M\<^sub>L i (Clo v vs n) = Clo (lift\<^sub>M\<^sub>L i v) (map (lift\<^sub>M\<^sub>L i) vs) n" |
"lift\<^sub>M\<^sub>L i (apply u v) = apply (lift\<^sub>M\<^sub>L i u) (lift\<^sub>M\<^sub>L i v)"
definition
cons :: "tm \<Rightarrow> (nat \<Rightarrow> tm) \<Rightarrow> (nat \<Rightarrow> tm)" (infix "##" 65) where
"t##\<sigma> \<equiv> \<lambda>i. case i of 0 \<Rightarrow> t | Suc j \<Rightarrow> lift 0 (\<sigma> j)"
definition
cons_ML :: "ml \<Rightarrow> (nat \<Rightarrow> ml) \<Rightarrow> (nat \<Rightarrow> ml)" (infix "##" 65) where
"v##\<sigma> \<equiv> \<lambda>i. case i of 0 \<Rightarrow> v::ml | Suc j \<Rightarrow> lift\<^sub>M\<^sub>L 0 (\<sigma> j)"
text\<open>Only for pure terms!\<close>
primrec subst :: "(nat \<Rightarrow> tm) \<Rightarrow> tm \<Rightarrow> tm"
where
"subst \<sigma> (C nm) = C nm"
| "subst \<sigma> (V x) = \<sigma> x"
| "subst \<sigma> (\<Lambda> t) = \<Lambda>(subst (V 0 ## \<sigma>) t)"
| "subst \<sigma> (s\<bullet>t) = (subst \<sigma> s) \<bullet> (subst \<sigma> t)"
fun subst_ML :: "(nat \<Rightarrow> ml) \<Rightarrow> ml \<Rightarrow> ml" ("subst\<^sub>M\<^sub>L") where
"subst\<^sub>M\<^sub>L \<sigma> (C\<^sub>M\<^sub>L nm) = C\<^sub>M\<^sub>L nm" |
"subst\<^sub>M\<^sub>L \<sigma> (V\<^sub>M\<^sub>L X) = \<sigma> X" |
"subst\<^sub>M\<^sub>L \<sigma> (A\<^sub>M\<^sub>L v vs) = A\<^sub>M\<^sub>L (subst\<^sub>M\<^sub>L \<sigma> v) (map (subst\<^sub>M\<^sub>L \<sigma>) vs)" |
"subst\<^sub>M\<^sub>L \<sigma> (Lam\<^sub>M\<^sub>L v) = Lam\<^sub>M\<^sub>L (subst\<^sub>M\<^sub>L (V\<^sub>M\<^sub>L 0 ## \<sigma>) v)" |
"subst\<^sub>M\<^sub>L \<sigma> (C\<^sub>U nm vs) = C\<^sub>U nm (map (subst\<^sub>M\<^sub>L \<sigma>) vs)" |
"subst\<^sub>M\<^sub>L \<sigma> (V\<^sub>U x vs) = V\<^sub>U x (map (subst\<^sub>M\<^sub>L \<sigma>) vs)" |
"subst\<^sub>M\<^sub>L \<sigma> (Clo v vs n) = Clo (subst\<^sub>M\<^sub>L \<sigma> v) (map (subst\<^sub>M\<^sub>L \<sigma>) vs) n" |
"subst\<^sub>M\<^sub>L \<sigma> (apply u v) = apply (subst\<^sub>M\<^sub>L \<sigma> u) (subst\<^sub>M\<^sub>L \<sigma> v)"
(* FIXME currrently needed for code generator
lemmas [code] = lift_tm.simps lift_ml.simps
lemmas [code] = subst_ML.simps *)
abbreviation
subst_decr :: "nat \<Rightarrow> tm \<Rightarrow> nat \<Rightarrow> tm" where
"subst_decr k t \<equiv> \<lambda>n. if n<k then V n else if n=k then t else V(n - 1)"
abbreviation
subst_decr_ML :: "nat \<Rightarrow> ml \<Rightarrow> nat \<Rightarrow> ml" where
"subst_decr_ML k v \<equiv> \<lambda>n. if n<k then V\<^sub>M\<^sub>L n else if n=k then v else V\<^sub>M\<^sub>L(n - 1)"
abbreviation
subst1 :: "tm \<Rightarrow> tm \<Rightarrow> nat \<Rightarrow> tm" ("(_/[_'/_])" [300, 0, 0] 300) where
"s[t/k] \<equiv> subst (subst_decr k t) s"
abbreviation
subst1_ML :: "ml \<Rightarrow> ml \<Rightarrow> nat \<Rightarrow> ml" ("(_/[_'/_])" [300, 0, 0] 300) where
"u[v/k] \<equiv> subst\<^sub>M\<^sub>L (subst_decr_ML k v) u"
lemma apply_cons[simp]:
"(t##\<sigma>) i = (if i=0 then t::tm else lift 0 (\<sigma>(i - 1)))"
by(simp add: cons_def split:nat.split)
lemma apply_cons_ML[simp]:
"(v##\<sigma>) i = (if i=0 then v::ml else lift\<^sub>M\<^sub>L 0 (\<sigma>(i - 1)))"
by(simp add: cons_ML_def split:nat.split)
lemma lift_foldl_At[simp]:
"lift k (s \<bullet>\<bullet> ts) = (lift k s) \<bullet>\<bullet> (map (lift k) ts)"
by(induct ts arbitrary:s) simp_all
lemma lift_lift_ml: fixes v :: ml shows
"i < k+1 \<Longrightarrow> lift (Suc k) (lift i v) = lift i (lift k v)"
by(induct i v rule:lift_ml.induct)
simp_all
lemma lift_lift_tm: fixes t :: tm shows
"i < k+1 \<Longrightarrow> lift (Suc k) (lift i t) = lift i (lift k t)"
by(induct t arbitrary: i rule:lift_tm.induct)(simp_all add:lift_lift_ml)
lemma lift_lift_ML:
"i < k+1 \<Longrightarrow> lift\<^sub>M\<^sub>L (Suc k) (lift\<^sub>M\<^sub>L i v) = lift\<^sub>M\<^sub>L i (lift\<^sub>M\<^sub>L k v)"
by(induct v arbitrary: i rule:lift_ML.induct)
simp_all
lemma lift_lift_ML_comm:
"lift j (lift\<^sub>M\<^sub>L i v) = lift\<^sub>M\<^sub>L i (lift j v)"
by(induct v arbitrary: i j rule:lift_ML.induct)
simp_all
lemma V_ML_cons_ML_subst_decr[simp]:
"V\<^sub>M\<^sub>L 0 ## subst_decr_ML k v = subst_decr_ML (Suc k) (lift\<^sub>M\<^sub>L 0 v)"
by(rule ext)(simp add:cons_ML_def split:nat.split)
lemma shift_subst_decr[simp]:
"V 0 ## subst_decr k t = subst_decr (Suc k) (lift 0 t)"
by(rule ext)(simp add:cons_def split:nat.split)
lemma lift_comp_subst_decr[simp]:
"lift 0 o subst_decr_ML k v = subst_decr_ML k (lift 0 v)"
by(rule ext) simp
lemma subst_ML_ext: "\<forall>i. \<sigma> i = \<sigma>' i \<Longrightarrow> subst\<^sub>M\<^sub>L \<sigma> v = subst\<^sub>M\<^sub>L \<sigma>' v"
by(metis ext)
lemma subst_ext: "\<forall>i. \<sigma> i = \<sigma>' i \<Longrightarrow> subst \<sigma> v = subst \<sigma>' v"
by(metis ext)
lemma lift_Pure_tms[simp]: "pure t \<Longrightarrow> pure(lift k t)"
by(induct arbitrary:k pred:pure) simp_all
lemma cons_ML_V_ML[simp]: "(V\<^sub>M\<^sub>L 0 ## V\<^sub>M\<^sub>L) = V\<^sub>M\<^sub>L"
by(rule ext) simp
lemma cons_V[simp]: "(V 0 ## V) = V"
by(rule ext) simp
lemma lift_o_shift: "lift k \<circ> (V\<^sub>M\<^sub>L 0 ## \<sigma>) = (V\<^sub>M\<^sub>L 0 ## (lift k \<circ> \<sigma>))"
by(rule ext)(simp add: lift_lift_ML_comm)
lemma lift_subst_ML:
"lift k (subst\<^sub>M\<^sub>L \<sigma> v) = subst\<^sub>M\<^sub>L (lift k \<circ> \<sigma>) (lift k v)"
apply(induct \<sigma> v rule:subst_ML.induct)
apply(simp_all add: o_assoc lift_o_shift del:apply_cons_ML)
apply(simp add:o_def)
done
corollary lift_subst_ML1:
"\<forall>v k. lift_ml 0 (u[v/k]) = (lift_ml 0 u)[lift 0 v/k]"
apply(induct u rule:ml_induct)
apply(simp_all add:lift_lift_ml lift_subst_ML)
apply(subst lift_lift_ML_comm)apply simp
done
lemma lift_ML_subst_ML:
"lift\<^sub>M\<^sub>L k (subst\<^sub>M\<^sub>L \<sigma> v) =
subst\<^sub>M\<^sub>L (\<lambda>i. if i<k then lift\<^sub>M\<^sub>L k (\<sigma> i) else if i=k then V\<^sub>M\<^sub>L k else lift\<^sub>M\<^sub>L k (\<sigma>(i - 1))) (lift\<^sub>M\<^sub>L k v)"
(is "_ = subst\<^sub>M\<^sub>L (?insrt k \<sigma>) (lift\<^sub>M\<^sub>L k v)")
apply (induct k v arbitrary: \<sigma> k rule: lift_ML.induct)
apply (simp_all add: o_assoc lift_o_shift)
apply(subgoal_tac "V\<^sub>M\<^sub>L 0 ## ?insrt k \<sigma> = ?insrt (Suc k) (V\<^sub>M\<^sub>L 0 ## \<sigma>)")
apply simp
apply (simp add:fun_eq_iff lift_lift_ML cons_ML_def split:nat.split)
done
corollary subst_cons_lift:
"subst\<^sub>M\<^sub>L (V\<^sub>M\<^sub>L 0 ## \<sigma>) o (lift\<^sub>M\<^sub>L 0) = lift\<^sub>M\<^sub>L 0 o (subst\<^sub>M\<^sub>L \<sigma>)"
apply(rule ext)
apply(simp add: lift_ML_subst_ML)
apply(subgoal_tac "(V\<^sub>M\<^sub>L 0 ## \<sigma>) = (\<lambda>i. if i = 0 then V\<^sub>M\<^sub>L 0 else lift\<^sub>M\<^sub>L 0 (\<sigma> (i - 1)))")
apply simp
apply(rule ext, simp)
done
lemma lift_ML_id[simp]: "closed\<^sub>M\<^sub>L k v \<Longrightarrow> lift\<^sub>M\<^sub>L k v = v"
by(induct k v rule: lift_ML.induct)(simp_all add:list_eq_iff_nth_eq)
lemma subst_ML_id:
"closed\<^sub>M\<^sub>L k v \<Longrightarrow> \<forall>i<k. \<sigma> i = V\<^sub>M\<^sub>L i \<Longrightarrow> subst\<^sub>M\<^sub>L \<sigma> v = v"
apply (induct \<sigma> v arbitrary: k rule: subst_ML.induct)
apply (auto simp add: list_eq_iff_nth_eq)
apply(simp add:Ball_def)
apply(erule_tac x="vs!i" in meta_allE)
apply(erule_tac x="k" in meta_allE)
apply(erule_tac x="k" in meta_allE)
apply simp
apply(erule_tac x="vs!i" in meta_allE)
apply(erule_tac x="k" in meta_allE)
apply simp
apply(erule_tac x="vs!i" in meta_allE)
apply(erule_tac x="k" in meta_allE)
apply simp
apply(erule_tac x="vs!i" in meta_allE)
apply(erule_tac x="k" in meta_allE)
apply(erule_tac x="k" in meta_allE)
apply simp
done
corollary subst_ML_id2[simp]: "closed\<^sub>M\<^sub>L 0 v \<Longrightarrow> subst\<^sub>M\<^sub>L \<sigma> v = v"
using subst_ML_id[where k=0] by simp
lemma subst_ML_coincidence:
"closed\<^sub>M\<^sub>L k v \<Longrightarrow> \<forall>i<k. \<sigma> i = \<sigma>' i \<Longrightarrow> subst\<^sub>M\<^sub>L \<sigma> v = subst\<^sub>M\<^sub>L \<sigma>' v"
by (induct \<sigma> v arbitrary: k \<sigma>' rule: subst_ML.induct) auto
lemma subst_ML_comp:
"subst\<^sub>M\<^sub>L \<sigma> (subst\<^sub>M\<^sub>L \<sigma>' v) = subst\<^sub>M\<^sub>L (subst\<^sub>M\<^sub>L \<sigma> \<circ> \<sigma>') v"
apply (induct \<sigma>' v arbitrary: \<sigma> rule: subst_ML.induct)
apply (simp_all add: list_eq_iff_nth_eq)
apply(rule subst_ML_ext)
apply simp
apply (metis o_apply subst_cons_lift)
done
lemma subst_ML_comp2:
"\<forall>i. \<sigma>'' i = subst\<^sub>M\<^sub>L \<sigma> (\<sigma>' i) \<Longrightarrow> subst\<^sub>M\<^sub>L \<sigma> (subst\<^sub>M\<^sub>L \<sigma>' v) = subst\<^sub>M\<^sub>L \<sigma>'' v"
by(simp add:subst_ML_comp subst_ML_ext)
lemma closed_tm_ML_foldl_At:
"closed\<^sub>M\<^sub>L k (t \<bullet>\<bullet> ts) \<longleftrightarrow> closed\<^sub>M\<^sub>L k t \<and> (\<forall>t \<in> set ts. closed\<^sub>M\<^sub>L k t)"
by(induct ts arbitrary: t) simp_all
lemma closed_ML_lift[simp]:
fixes v :: ml shows "closed\<^sub>M\<^sub>L k v \<Longrightarrow> closed\<^sub>M\<^sub>L k (lift m v)"
by(induct k v arbitrary: m rule: lift_ML.induct)
(simp_all add:list_eq_iff_nth_eq)
lemma closed_ML_Suc: "closed\<^sub>M\<^sub>L n v \<Longrightarrow> closed\<^sub>M\<^sub>L (Suc n) (lift\<^sub>M\<^sub>L k v)"
by (induct k v arbitrary: n rule: lift_ML.induct) simp_all
lemma closed_ML_subst_ML:
"\<forall>i. closed\<^sub>M\<^sub>L k (\<sigma> i) \<Longrightarrow> closed\<^sub>M\<^sub>L k (subst\<^sub>M\<^sub>L \<sigma> v)"
by(induct \<sigma> v arbitrary: k rule: subst_ML.induct) (auto simp: closed_ML_Suc)
lemma closed_ML_subst_ML2:
"closed\<^sub>M\<^sub>L k v \<Longrightarrow> \<forall>i<k. closed\<^sub>M\<^sub>L l (\<sigma> i) \<Longrightarrow> closed\<^sub>M\<^sub>L l (subst\<^sub>M\<^sub>L \<sigma> v)"
by(induct \<sigma> v arbitrary: k l rule: subst_ML.induct)(auto simp: closed_ML_Suc)
lemma subst_foldl[simp]:
"subst \<sigma> (s \<bullet>\<bullet> ts) = (subst \<sigma> s) \<bullet>\<bullet> (map (subst \<sigma>) ts)"
by (induct ts arbitrary: s) auto
lemma subst_V: "pure t \<Longrightarrow> subst V t = t"
by(induct pred:pure) simp_all
lemma lift_subst_aux:
"pure t \<Longrightarrow> \<forall>i<k. \<sigma>' i = lift k (\<sigma> i) \<Longrightarrow>
\<forall>i\<ge>k. \<sigma>'(Suc i) = lift k (\<sigma> i) \<Longrightarrow>
\<sigma>' k = V k \<Longrightarrow> lift k (subst \<sigma> t) = subst \<sigma>' (lift k t)"
apply(induct arbitrary:\<sigma> \<sigma>' k pred:pure)
apply (simp_all add: split:nat.split)
apply(erule meta_allE)+
apply(erule meta_impE)
defer
apply(erule meta_impE)
defer
apply(erule meta_mp)
apply (simp_all add: cons_def lift_lift_ml lift_lift_tm split:nat.split)
done
corollary lift_subst:
"pure t \<Longrightarrow> lift 0 (subst \<sigma> t) = subst (V 0 ## \<sigma>) (lift 0 t)"
by (simp add: lift_subst_aux lift_lift_ml)
lemma subst_comp:
"pure t \<Longrightarrow> \<forall>i. pure(\<sigma>' i) \<Longrightarrow>
\<sigma>'' = (\<lambda>i. subst \<sigma> (\<sigma>' i)) \<Longrightarrow> subst \<sigma> (subst \<sigma>' t) = subst \<sigma>'' t"
apply(induct arbitrary:\<sigma> \<sigma>' \<sigma>'' pred:pure)
apply simp
apply simp
defer
apply simp
apply (simp (no_asm))
apply(erule meta_allE)+
apply(erule meta_impE)
defer
apply(erule meta_mp)
prefer 2 apply simp
apply(rule ext)
apply(simp add:lift_subst)
done
section "Reduction"
subsection "Patterns"
inductive pattern :: "tm \<Rightarrow> bool"
and patterns :: "tm list \<Rightarrow> bool" where
"patterns ts \<equiv> \<forall>t\<in>set ts. pattern t" |
pat_V: "pattern(V X)" |
pat_C: "patterns ts \<Longrightarrow> pattern(C nm \<bullet>\<bullet> ts)"
lemma pattern_Lam[simp]: "\<not> pattern(\<Lambda> t)"
by(auto elim!: pattern.cases)
lemma pattern_At'D12: "pattern r \<Longrightarrow> r = (s \<bullet> t) \<Longrightarrow> pattern s \<and> pattern t"
proof(induct arbitrary: s t pred:pattern)
case pat_V thus ?case by simp
next
case pat_C thus ?case
by (simp add: atomic_tm_head_tm split:if_split_asm)
(metis eta_head_args in_set_butlastD pattern.pat_C)
qed
lemma pattern_AtD12: "pattern(s \<bullet> t) \<Longrightarrow> pattern s \<and> pattern t"
by(metis pattern_At'D12)
lemma pattern_At_vecD: "pattern(s \<bullet>\<bullet> ts) \<Longrightarrow> patterns ts"
apply(induct ts rule:rev_induct)
apply simp
apply (fastforce dest!:pattern_AtD12)
done
lemma pattern_At_decomp: "pattern(s \<bullet> t) \<Longrightarrow> \<exists>nm ss. s = C nm \<bullet>\<bullet> ss"
proof(induct s arbitrary: t)
case (At s1 s2) show ?case
using At by (metis foldl_Cons foldl_Nil foldl_append pattern_AtD12)
qed (auto elim!: pattern.cases split:if_split_asm)
subsection "Reduction of \<open>\<lambda>\<close>-terms"
text\<open>The source program:\<close>
axiomatization R :: "(cname * tm list * tm)set" where
pure_R: "(nm,ts,t) : R \<Longrightarrow> (\<forall>t \<in> set ts. pure t) \<and> pure t" and
fv_R: "(nm,ts,t) : R \<Longrightarrow> X : fv t \<Longrightarrow> \<exists>t' \<in> set ts. X : fv t'" and
pattern_R: "(nm,ts,t') : R \<Longrightarrow> patterns ts"
inductive_set
Red_tm :: "(tm * tm)set"
and red_tm :: "[tm, tm] => bool" (infixl "\<rightarrow>" 50)
where
"s \<rightarrow> t \<equiv> (s, t) \<in> Red_tm"
\<comment> \<open>$\beta$-reduction\<close>
| "(\<Lambda> t) \<bullet> s \<rightarrow> t[s/0]"
\<comment> \<open>$\eta$-expansion\<close>
| "t \<rightarrow> \<Lambda> ((lift 0 t) \<bullet> (V 0))"
\<comment> \<open>Rewriting\<close>
| "(nm,ts,t) : R \<Longrightarrow> (C nm) \<bullet>\<bullet> (map (subst \<sigma>) ts) \<rightarrow> subst \<sigma> t"
| "t \<rightarrow> t' \<Longrightarrow> \<Lambda> t \<rightarrow> \<Lambda> t'"
| "s \<rightarrow> s' \<Longrightarrow> s \<bullet> t \<rightarrow> s' \<bullet> t"
| "t \<rightarrow> t' \<Longrightarrow> s \<bullet> t \<rightarrow> s \<bullet> t'"
abbreviation
reds_tm :: "[tm, tm] => bool" (infixl "\<rightarrow>*" 50) where
"s \<rightarrow>* t \<equiv> (s, t) \<in> Red_tm^*"
inductive_set
Reds_tm_list :: "(tm list * tm list) set"
and reds_tm_list :: "[tm list, tm list] \<Rightarrow> bool" (infixl "\<rightarrow>*" 50)
where
"ss \<rightarrow>* ts \<equiv> (ss, ts) \<in> Reds_tm_list"
| "[] \<rightarrow>* []"
| "ts \<rightarrow>* ts' \<Longrightarrow> t \<rightarrow>* t' \<Longrightarrow> t#ts \<rightarrow>* t'#ts'"
declare Reds_tm_list.intros[simp]
lemma Reds_tm_list_refl[simp]: fixes ts :: "tm list" shows "ts \<rightarrow>* ts"
by(induct ts) auto
lemma Red_tm_append: "rs \<rightarrow>* rs' \<Longrightarrow> ts \<rightarrow>* ts' \<Longrightarrow> rs @ ts \<rightarrow>* rs' @ ts'"
by(induct set: Reds_tm_list) auto
lemma Red_tm_rev: "ts \<rightarrow>* ts' \<Longrightarrow> rev ts \<rightarrow>* rev ts'"
by(induct set: Reds_tm_list) (auto simp:Red_tm_append)
lemma red_Lam[simp]: "t \<rightarrow>* t' \<Longrightarrow> \<Lambda> t \<rightarrow>* \<Lambda> t'"
apply(induct rule:rtrancl_induct)
apply(simp_all)
apply(blast intro: rtrancl_into_rtrancl Red_tm.intros)
done
lemma red_At1[simp]: "t \<rightarrow>* t' \<Longrightarrow> t \<bullet> s \<rightarrow>* t' \<bullet> s"
apply(induct rule:rtrancl_induct)
apply(simp_all)
apply(blast intro: rtrancl_into_rtrancl Red_tm.intros)
done
lemma red_At2[simp]: "t \<rightarrow>* t' \<Longrightarrow> s \<bullet> t \<rightarrow>* s \<bullet> t'"
apply(induct rule:rtrancl_induct)
apply(simp_all)
apply(blast intro:rtrancl_into_rtrancl Red_tm.intros)
done
lemma Reds_tm_list_foldl_At:
"ts \<rightarrow>* ts' \<Longrightarrow> s \<rightarrow>* s' \<Longrightarrow> s \<bullet>\<bullet> ts \<rightarrow>* s' \<bullet>\<bullet> ts'"
apply(induct arbitrary:s s' rule:Reds_tm_list.induct)
apply simp
apply simp
apply(blast dest: red_At1 red_At2 intro:rtrancl_trans)
done
subsection "Reduction of ML-terms"
text\<open>The compiled rule set:\<close>
consts compR :: "(cname * ml list * ml)set"
text\<open>\noindent
The actual definition is given in \S\ref{sec:Compiler} below.\<close>
text\<open>Now we characterize ML values that cannot possibly be rewritten by a
rule in @{const compR}.\<close>
lemma termination_no_match_ML:
"i < length ps \<Longrightarrow> rev ps ! i = C\<^sub>U nm vs
\<Longrightarrow> sum_list (map size vs) < sum_list (map size ps)"
apply(subgoal_tac "C\<^sub>U nm vs : set ps")
apply(drule sum_list_map_remove1[of _ _ size])
apply (simp add:size_list_conv_sum_list)
apply (metis in_set_conv_nth length_rev set_rev)
done
declare conj_cong[fundef_cong]
function no_match_ML ("no'_match\<^sub>M\<^sub>L") where
"no_match\<^sub>M\<^sub>L ps os =
(\<exists>i < min (size os) (size ps).
\<exists>nm nm' vs vs'. (rev ps)!i = C\<^sub>U nm vs \<and> (rev os)!i = C\<^sub>U nm' vs' \<and>
(nm=nm' \<longrightarrow> no_match\<^sub>M\<^sub>L vs vs'))"
by pat_completeness auto
termination
apply(relation "measure(%(vs::ml list,_). \<Sum>v\<leftarrow>vs. size v)")
apply (auto simp:termination_no_match_ML)
done
abbreviation
"no_match_compR nm os \<equiv>
\<forall>(nm',ps,v)\<in> compR. nm=nm' \<longrightarrow> no_match\<^sub>M\<^sub>L ps os"
declare no_match_ML.simps[simp del]
inductive_set
Red_ml :: "(ml * ml)set"
and Red_ml_list :: "(ml list * ml list)set"
and red_ml :: "[ml, ml] => bool" (infixl "\<Rightarrow>" 50)
and red_ml_list :: "[ml list, ml list] => bool" (infixl "\<Rightarrow>" 50)
and reds_ml :: "[ml, ml] => bool" (infixl "\<Rightarrow>*" 50)
where
"s \<Rightarrow> t \<equiv> (s, t) \<in> Red_ml"
| "ss \<Rightarrow> ts \<equiv> (ss, ts) \<in> Red_ml_list"
| "s \<Rightarrow>* t \<equiv> (s, t) \<in> Red_ml^*"
\<comment> \<open>ML $\beta$-reduction\<close>
| "A\<^sub>M\<^sub>L (Lam\<^sub>M\<^sub>L u) [v] \<Rightarrow> u[v/0]"
\<comment> \<open>Execution of a compiled rewrite rule\<close>
| "(nm,vs,v) : compR \<Longrightarrow> \<forall> i. closed\<^sub>M\<^sub>L 0 (\<sigma> i) \<Longrightarrow>
A\<^sub>M\<^sub>L (C\<^sub>M\<^sub>L nm) (map (subst\<^sub>M\<^sub>L \<sigma>) vs) \<Rightarrow> subst\<^sub>M\<^sub>L \<sigma> v"
\<comment> \<open>default rule:\<close>
| "\<forall>i. closed\<^sub>M\<^sub>L 0 (\<sigma> i)
\<Longrightarrow> vs = map V\<^sub>M\<^sub>L [0..<arity nm] \<Longrightarrow> vs' = map (subst\<^sub>M\<^sub>L \<sigma>) vs
\<Longrightarrow> no_match_compR nm vs'
\<Longrightarrow> A\<^sub>M\<^sub>L (C\<^sub>M\<^sub>L nm) vs' \<Rightarrow> subst\<^sub>M\<^sub>L \<sigma> (C\<^sub>U nm vs)"
\<comment> \<open>Equations for function \texttt{apply}\<close>
| apply_Clo1: "apply (Clo f vs (Suc 0)) v \<Rightarrow> A\<^sub>M\<^sub>L f (v # vs)"
| apply_Clo2: "n > 0 \<Longrightarrow>
apply (Clo f vs (Suc n)) v \<Rightarrow> Clo f (v # vs) n"
| apply_C: "apply (C\<^sub>U nm vs) v \<Rightarrow> C\<^sub>U nm (v # vs)"
| apply_V: "apply (V\<^sub>U x vs) v \<Rightarrow> V\<^sub>U x (v # vs)"
\<comment> \<open>Context rules\<close>
| ctxt_C: "vs \<Rightarrow> vs' \<Longrightarrow> C\<^sub>U nm vs \<Rightarrow> C\<^sub>U nm vs'"
| ctxt_V: "vs \<Rightarrow> vs' \<Longrightarrow> V\<^sub>U x vs \<Rightarrow> V\<^sub>U x vs'"
| ctxt_Clo1: "f \<Rightarrow> f' \<Longrightarrow> Clo f vs n \<Rightarrow> Clo f' vs n"
| ctxt_Clo3: "vs \<Rightarrow> vs' \<Longrightarrow> Clo f vs n \<Rightarrow> Clo f vs' n"
| ctxt_apply1: "s \<Rightarrow> s' \<Longrightarrow> apply s t \<Rightarrow> apply s' t"
| ctxt_apply2: "t \<Rightarrow> t' \<Longrightarrow> apply s t \<Rightarrow> apply s t'"
| ctxt_A_ML1: "f \<Rightarrow> f' \<Longrightarrow> A\<^sub>M\<^sub>L f vs \<Rightarrow> A\<^sub>M\<^sub>L f' vs"
| ctxt_A_ML2: "vs \<Rightarrow> vs' \<Longrightarrow> A\<^sub>M\<^sub>L f vs \<Rightarrow> A\<^sub>M\<^sub>L f vs'"
| ctxt_list1: "v \<Rightarrow> v' \<Longrightarrow> v#vs \<Rightarrow> v'#vs"
| ctxt_list2: "vs \<Rightarrow> vs' \<Longrightarrow> v#vs \<Rightarrow> v#vs'"
inductive_set
Red_term :: "(tm * tm)set"
and red_term :: "[tm, tm] => bool" (infixl "\<Rightarrow>" 50)
and reds_term :: "[tm, tm] => bool" (infixl "\<Rightarrow>*" 50)
where
"s \<Rightarrow> t \<equiv> (s, t) \<in> Red_term"
| "s \<Rightarrow>* t \<equiv> (s, t) \<in> Red_term^*"
\<comment> \<open>function \texttt{term}\<close>
| term_C: "term (C\<^sub>U nm vs) \<Rightarrow> (C nm) \<bullet>\<bullet> (map term (rev vs))"
| term_V: "term (V\<^sub>U x vs) \<Rightarrow> (V x) \<bullet>\<bullet> (map term (rev vs))"
| term_Clo: "term(Clo vf vs n) \<Rightarrow> \<Lambda> (term (apply (lift 0 (Clo vf vs n)) (V\<^sub>U 0 [])))"
\<comment> \<open>context rules\<close>
| ctxt_Lam: "t \<Rightarrow> t' \<Longrightarrow> \<Lambda> t \<Rightarrow> \<Lambda> t'"
| ctxt_At1: "s \<Rightarrow> s' \<Longrightarrow> s \<bullet> t \<Rightarrow> s' \<bullet> t"
| ctxt_At2: "t \<Rightarrow> t' \<Longrightarrow> s \<bullet> t \<Rightarrow> s \<bullet> t'"
| ctxt_term: "v \<Rightarrow> v' \<Longrightarrow> term v \<Rightarrow> term v'"
section "Kernel"
text\<open>First a special size function and some lemmas for the
termination proof of the kernel function.\<close>
fun size' :: "ml \<Rightarrow> nat" where
"size' (C\<^sub>M\<^sub>L nm) = 1" |
"size' (V\<^sub>M\<^sub>L X) = 1" |
"size' (A\<^sub>M\<^sub>L v vs) = (size' v + (\<Sum>v\<leftarrow>vs. size' v))+1" |
"size' (Lam\<^sub>M\<^sub>L v) = size' v + 1" |
"size' (C\<^sub>U nm vs) = (\<Sum>v\<leftarrow>vs. size' v)+1" |
"size' (V\<^sub>U nm vs) = (\<Sum>v\<leftarrow>vs. size' v)+1" |
"size' (Clo f vs n) = (size' f + (\<Sum>v\<leftarrow>vs. size' v))+1" |
"size' (apply v w) = (size' v + size' w)+1"
lemma sum_list_size'[simp]:
"v \<in> set vs \<Longrightarrow> size' v < Suc(sum_list (map size' vs))"
by(induct vs)(auto)
corollary cor_sum_list_size'[simp]:
"v \<in> set vs \<Longrightarrow> size' v < Suc(m + sum_list (map size' vs))"
using sum_list_size'[of v vs] by arith
lemma size'_lift_ML: "size' (lift\<^sub>M\<^sub>L k v) = size' v"
apply(induct v arbitrary:k rule:size'.induct)
apply simp_all
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
done
lemma size'_subst_ML[simp]:
"\<forall>i j. size'(\<sigma> i) = 1 \<Longrightarrow> size' (subst\<^sub>M\<^sub>L \<sigma> v) = size' v"
apply(induct v arbitrary:\<sigma> rule:size'.induct)
apply simp_all
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
apply(erule meta_allE)
apply(erule meta_mp)
apply(simp add: size'_lift_ML split:nat.split)
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
done
lemma size'_lift[simp]: "size' (lift i v) = size' v"
apply(induct v arbitrary:i rule:size'.induct)
apply simp_all
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
apply(rule arg_cong[where f = sum_list])
apply(rule map_ext)
apply simp
done
function kernel :: "ml \<Rightarrow> tm" ("_!" 300) where
"(C\<^sub>M\<^sub>L nm)! = C nm" |
"(A\<^sub>M\<^sub>L v vs)! = v! \<bullet>\<bullet> (map kernel (rev vs))" |
"(Lam\<^sub>M\<^sub>L v)! = \<Lambda> (((lift 0 v)[V\<^sub>U 0 []/0])!)" |
"(C\<^sub>U nm vs)! = (C nm) \<bullet>\<bullet> (map kernel (rev vs))" |
"(V\<^sub>U x vs)! = (V x) \<bullet>\<bullet> (map kernel (rev vs))" |
"(Clo f vs n)! = f! \<bullet>\<bullet> (map kernel (rev vs))" |
"(apply v w)! = v! \<bullet> (w!)" |
"(V\<^sub>M\<^sub>L X)! = undefined"
by pat_completeness auto
termination by(relation "measure size'") auto
primrec kernelt :: "tm \<Rightarrow> tm" ("_!" 300)
where
"(C nm)! = C nm"
| "(V x)! = V x"
| "(s \<bullet> t)! = (s!) \<bullet> (t!)"
| "(\<Lambda> t)! = \<Lambda>(t!)"
| "(term v)! = v!"
abbreviation
kernels :: "ml list \<Rightarrow> tm list" ("_!" 300) where
"vs! \<equiv> map kernel vs"
lemma kernel_pure: assumes "pure t" shows "t! = t"
using assms by (induct) simp_all
lemma kernel_foldl_At[simp]: "(s \<bullet>\<bullet> ts)! = (s!) \<bullet>\<bullet> (map kernelt ts)"
by (induct ts arbitrary: s) simp_all
lemma kernelt_o_term[simp]: "(kernelt \<circ> term) = kernel"
by(rule ext) simp
lemma pure_foldl:
"pure t \<Longrightarrow> \<forall>t\<in>set ts. pure t \<Longrightarrow>
(!!s t. pure s \<Longrightarrow> pure t \<Longrightarrow> pure(f s t)) \<Longrightarrow>
pure(foldl f t ts)"
by(induct ts arbitrary: t) simp_all
lemma pure_kernel: fixes v :: ml shows "closed\<^sub>M\<^sub>L 0 v \<Longrightarrow> pure(v!)"
proof(induct v rule:kernel.induct)
case (3 v)
hence "closed\<^sub>M\<^sub>L (Suc 0) (lift 0 v)" by simp
then have "subst\<^sub>M\<^sub>L (\<lambda>n. V\<^sub>U 0 []) (lift 0 v) = lift 0 v[V\<^sub>U 0 []/0]"
by(rule subst_ML_coincidence) simp
moreover have "closed\<^sub>M\<^sub>L 0 (subst\<^sub>M\<^sub>L (\<lambda>n. V\<^sub>U 0 []) (lift 0 v))"
by(simp add: closed_ML_subst_ML)
ultimately have "closed\<^sub>M\<^sub>L 0 (lift 0 v[V\<^sub>U 0 []/0])" by simp
thus ?case using 3(1) by (simp add:pure_foldl)
qed (simp_all add:pure_foldl)
corollary subst_V_kernel: fixes v :: ml shows
"closed\<^sub>M\<^sub>L 0 v \<Longrightarrow> subst V (v!) = v!"
by (metis pure_kernel subst_V)
lemma kernel_lift_tm: fixes v :: ml shows
"closed\<^sub>M\<^sub>L 0 v \<Longrightarrow> (lift i v)! = lift i (v!)"
apply(induct v arbitrary: i rule: kernel.induct)
apply (simp_all add:list_eq_iff_nth_eq)
apply(simp add: rev_nth)
defer
apply(simp add: rev_nth)
apply(simp add: rev_nth)
apply(simp add: rev_nth)
apply(erule_tac x="Suc i" in meta_allE)
apply(erule meta_impE)
defer
apply (simp add:lift_subst_ML)
apply(subgoal_tac "lift (Suc i) \<circ> (\<lambda>n. if n = 0 then V\<^sub>U 0 [] else V\<^sub>M\<^sub>L (n - 1)) = (\<lambda>n. if n = 0 then V\<^sub>U 0 [] else V\<^sub>M\<^sub>L (n - 1))")
apply (simp add:lift_lift_ml)
apply(rule ext)
apply(simp)
apply(subst closed_ML_subst_ML2[of "1"])
apply(simp)
apply(simp)
apply(simp)
done
subsection "An auxiliary substitution"
text\<open>This function is only introduced to prove the involved susbtitution
lemma \<open>kernel_subst1\<close> below.\<close>
fun subst_ml :: "(nat \<Rightarrow> nat) \<Rightarrow> ml \<Rightarrow> ml" where
"subst_ml \<sigma> (C\<^sub>M\<^sub>L nm) = C\<^sub>M\<^sub>L nm" |
"subst_ml \<sigma> (V\<^sub>M\<^sub>L X) = V\<^sub>M\<^sub>L X" |
"subst_ml \<sigma> (A\<^sub>M\<^sub>L v vs) = A\<^sub>M\<^sub>L (subst_ml \<sigma> v) (map (subst_ml \<sigma>) vs)" |
"subst_ml \<sigma> (Lam\<^sub>M\<^sub>L v) = Lam\<^sub>M\<^sub>L (subst_ml \<sigma> v)" |
"subst_ml \<sigma> (C\<^sub>U nm vs) = C\<^sub>U nm (map (subst_ml \<sigma>) vs)" |
"subst_ml \<sigma> (V\<^sub>U x vs) = V\<^sub>U (\<sigma> x) (map (subst_ml \<sigma>) vs)" |
"subst_ml \<sigma> (Clo v vs n) = Clo (subst_ml \<sigma> v) (map (subst_ml \<sigma>) vs) n" |
"subst_ml \<sigma> (apply u v) = apply (subst_ml \<sigma> u) (subst_ml \<sigma> v)"
lemma lift_ML_subst_ml:
"lift\<^sub>M\<^sub>L k (subst_ml \<sigma> v) = subst_ml \<sigma> (lift\<^sub>M\<^sub>L k v)"
apply (induct \<sigma> v arbitrary: k rule:subst_ml.induct)
apply (simp_all add:list_eq_iff_nth_eq)
done
lemma subst_ml_subst_ML:
"subst_ml \<sigma> (subst\<^sub>M\<^sub>L \<sigma>' v) = subst\<^sub>M\<^sub>L (subst_ml \<sigma> o \<sigma>') (subst_ml \<sigma> v)"
apply (induct \<sigma>' v arbitrary: \<sigma> rule: subst_ML.induct)
apply(simp_all add:list_eq_iff_nth_eq)
apply(subgoal_tac "(subst_ml \<sigma>' \<circ> V\<^sub>M\<^sub>L 0 ## \<sigma>) = V\<^sub>M\<^sub>L 0 ## (subst_ml \<sigma>' \<circ> \<sigma>)")
apply simp
apply(rule ext)
apply(simp add: lift_ML_subst_ml)
done
text\<open>Maybe this should be the def of lift:\<close>
lemma lift_is_subst_ml: "lift k v = subst_ml (\<lambda>n. if n<k then n else n+1) v"
by(induct k v rule:lift_ml.induct)(simp_all add:list_eq_iff_nth_eq)
lemma subst_ml_comp: "subst_ml \<sigma> (subst_ml \<sigma>' v) = subst_ml (\<sigma> o \<sigma>') v"
by(induct \<sigma>' v rule:subst_ml.induct)(simp_all add:list_eq_iff_nth_eq)
lemma subst_kernel:
"closed\<^sub>M\<^sub>L 0 v \<Longrightarrow> subst (\<lambda>n. V(\<sigma> n)) (v!) = (subst_ml \<sigma> v)!"
apply (induct v arbitrary: \<sigma> rule:kernel.induct)
apply (simp_all add:list_eq_iff_nth_eq)
apply(simp add: rev_nth)
defer
apply(simp add: rev_nth)
apply(simp add: rev_nth)
apply(simp add: rev_nth)
apply(erule_tac x="\<lambda>n. case n of 0 \<Rightarrow> 0 | Suc k \<Rightarrow> Suc(\<sigma> k)" in meta_allE)
apply(erule_tac meta_impE)
apply(rule closed_ML_subst_ML2[where k="Suc 0"])
apply (metis closed_ML_lift)
apply simp
apply(subgoal_tac "(\<lambda>n. V(case n of 0 \<Rightarrow> 0 | Suc k \<Rightarrow> Suc (\<sigma> k))) = (V 0 ## (\<lambda>n. V(\<sigma> n)))")
apply (simp add:subst_ml_subst_ML)
defer
apply(simp add:fun_eq_iff split:nat.split)
apply(simp add:lift_is_subst_ml subst_ml_comp)
apply(rule arg_cong[where f = kernel])
apply(subgoal_tac "(case_nat 0 (\<lambda>k. Suc (\<sigma> k)) \<circ> Suc) = Suc o \<sigma>")
prefer 2 apply(simp add:fun_eq_iff split:nat.split)
apply(subgoal_tac "(subst_ml (case_nat 0 (\<lambda>k. Suc (\<sigma> k))) \<circ>
(\<lambda>n. if n = 0 then V\<^sub>U 0 [] else V\<^sub>M\<^sub>L (n - 1)))
= (\<lambda>n. if n = 0 then V\<^sub>U 0 [] else V\<^sub>M\<^sub>L (n - 1))")
apply simp
apply(simp add: fun_eq_iff)
done
lemma if_cong0: "If x y z = If x y z"
by simp
lemma kernel_subst1:
"closed\<^sub>M\<^sub>L 0 v \<Longrightarrow> closed\<^sub>M\<^sub>L (Suc 0) u \<Longrightarrow>
kernel(u[v/0]) = (kernel((lift 0 u)[V\<^sub>U 0 []/0]))[v!/0]"
proof(induct u arbitrary:v rule:kernel.induct)
case (3 w)
show ?case (is "?L = ?R")
proof -
have "?L = \<Lambda>(lift 0 (w[lift\<^sub>M\<^sub>L 0 v/Suc 0])[V\<^sub>U 0 []/0] !)"
by (simp cong:if_cong0)
also have "\<dots> = \<Lambda>((lift 0 w)[lift\<^sub>M\<^sub>L 0 (lift 0 v)/Suc 0][V\<^sub>U 0 []/0]!)"
by(simp only: lift_subst_ML1 lift_lift_ML_comm)
also have "\<dots> = \<Lambda>(subst\<^sub>M\<^sub>L (\<lambda>n. if n=0 then V\<^sub>U 0 [] else
if n=Suc 0 then lift 0 v else V\<^sub>M\<^sub>L (n - 2)) (lift 0 w) !)"
apply simp
apply(rule arg_cong[where f = kernel])
apply(rule subst_ML_comp2)
using 3
apply auto
done
also have "\<dots> = \<Lambda>((lift 0 w)[V\<^sub>U 0 []/0][lift 0 v/0]!)"
apply simp
apply(rule arg_cong[where f = kernel])
apply(rule subst_ML_comp2[symmetric])
using 3
apply auto
done
also have "\<dots> = \<Lambda>((lift_ml 0 ((lift_ml 0 w)[V\<^sub>U 0 []/0]))[V\<^sub>U 0 []/0]![(lift 0 v)!/0])"
apply(rule arg_cong[where f = \<Lambda>])
apply(rule 3(1))
apply (metis closed_ML_lift 3(2))
apply(subgoal_tac "closed\<^sub>M\<^sub>L (Suc(Suc 0)) w")
defer
using 3
apply force
apply(subgoal_tac "closed\<^sub>M\<^sub>L (Suc (Suc 0)) (lift 0 w)")
defer
apply(erule closed_ML_lift)
apply(erule closed_ML_subst_ML2)
apply simp
done
also have "\<dots> = \<Lambda>((lift_ml 0 (lift_ml 0 w)[V\<^sub>U 1 []/0])[V\<^sub>U 0 []/0]![(lift 0 v)!/0])" (is "_ = ?M")
apply(subgoal_tac "lift_ml 0 (lift_ml 0 w[V\<^sub>U 0 []/0])[V\<^sub>U 0 []/0] =
lift_ml 0 (lift_ml 0 w)[V\<^sub>U 1 []/0][V\<^sub>U 0 []/0]")
apply simp
apply(subst lift_subst_ML)
apply(simp add:comp_def if_distrib[where f="lift_ml 0"] cong:if_cong)
done
finally have "?L = ?M" .
have "?R = \<Lambda> (subst (V 0 ## subst_decr 0 (v!))
(lift 0 (lift_ml 0 w[V\<^sub>U 0 []/Suc 0])[V\<^sub>U 0 []/0]!))"
apply(subgoal_tac "(V\<^sub>M\<^sub>L 0 ## (\<lambda>n. if n = 0 then V\<^sub>U 0 [] else V\<^sub>M\<^sub>L (n - Suc 0))) = subst_decr_ML (Suc 0) (V\<^sub>U 0 [])")
apply(simp cong:if_cong)
apply(simp add:fun_eq_iff cons_ML_def split:nat.splits)
done
also have "\<dots> = \<Lambda> (subst (V 0 ## subst_decr 0 (v!))
((lift 0 (lift_ml 0 w))[V\<^sub>U 1 []/Suc 0][V\<^sub>U 0 []/0]!))"
apply(subgoal_tac "lift 0 (lift 0 w[V\<^sub>U 0 []/Suc 0]) = lift 0 (lift 0 w)[V\<^sub>U 1 []/Suc 0]")
apply simp
apply(subst lift_subst_ML)
apply(simp add:comp_def if_distrib[where f="lift_ml 0"] cong:if_cong)
done
also have "(lift_ml 0 (lift_ml 0 w))[V\<^sub>U 1 []/Suc 0][V\<^sub>U 0 []/0] =
(lift 0 (lift_ml 0 w))[V\<^sub>U 0 []/0][V\<^sub>U 1 []/ 0]" (is "?l = ?r")
proof -
have "?l = subst\<^sub>M\<^sub>L (\<lambda>n. if n= 0 then V\<^sub>U 0 [] else if n = 1 then V\<^sub>U 1 [] else
V\<^sub>M\<^sub>L (n - 2))
(lift_ml 0 (lift_ml 0 w))"
by(auto intro!:subst_ML_comp2)
also have "\<dots> = ?r" by(auto intro!:subst_ML_comp2[symmetric])
finally show ?thesis .
qed
also have "\<Lambda> (subst (V 0 ## subst_decr 0 (v!)) (?r !)) = ?M"
proof-
have "subst (subst_decr (Suc 0) (lift_tm 0 (kernel v))) (lift_ml 0 (lift_ml 0 w)[V\<^sub>U 0 []/0][V\<^sub>U 1 []/0]!) =
subst (subst_decr 0 (kernel(lift_ml 0 v))) (lift_ml 0 (lift_ml 0 w)[V\<^sub>U 1 []/0][V\<^sub>U 0 []/0]!)" (is "?a = ?b")
proof-
define pi where "pi n = (if n = 0 then 1 else if n = 1 then 0 else n)" for n :: nat
have "(\<lambda>i. V (pi i)[lift 0 (v!)/0]) = subst_decr (Suc 0) (lift 0 (v!))"
by(rule ext)(simp add:pi_def)
hence "?a =
subst (subst_decr 0 (lift_tm 0 (kernel v))) (subst (\<lambda> n. V(pi n)) (lift_ml 0 (lift_ml 0 w)[V\<^sub>U 0 []/0][V\<^sub>U 1 []/0]!))"
apply(subst subst_comp[OF _ _ refl])
prefer 3 apply simp
using 3(3)
apply simp
apply(rule pure_kernel)
apply(rule closed_ML_subst_ML2[where k="Suc 0"])
apply(rule closed_ML_subst_ML2[where k="Suc(Suc 0)"])
apply simp
apply simp
apply simp
apply simp
done
also have "\<dots> =
(subst_ml pi (lift_ml 0 (lift_ml 0 w)[V\<^sub>U 0 []/0][V\<^sub>U 1 []/0]))![lift_tm 0 (v!)/0]"
apply(subst subst_kernel)
using 3 apply auto
apply(rule closed_ML_subst_ML2[where k="Suc 0"])
apply(rule closed_ML_subst_ML2[where k="Suc(Suc 0)"])
apply simp
apply simp
apply simp
done
also have "\<dots> = (subst_ml pi (lift_ml 0 (lift_ml 0 w)[V\<^sub>U 0 []/0][V\<^sub>U 1 []/0]))![lift 0 v!/0]"
proof -
have "lift 0 (v!) = lift 0 v!" by (metis 3(2) kernel_lift_tm)
thus ?thesis by (simp cong:if_cong)
qed
also have "\<dots> = ?b"
proof-
have 1: "subst_ml pi (lift 0 (lift 0 w)) = lift 0 (lift 0 w)"
apply(simp add:lift_is_subst_ml subst_ml_comp)
apply(subgoal_tac "pi \<circ> (Suc \<circ> Suc) = (Suc \<circ> Suc)")
apply(simp)
apply(simp add:pi_def fun_eq_iff)
done
have "subst_ml pi (lift_ml 0 (lift_ml 0 w)[V\<^sub>U 0 []/0][V\<^sub>U 1 []/0]) =
lift_ml 0 (lift_ml 0 w)[V\<^sub>U 1 []/0][V\<^sub>U 0 []/0]"
apply(subst subst_ml_subst_ML)
apply(subst subst_ml_subst_ML)
apply(subst 1)
apply(subst subst_ML_comp)
apply(rule subst_ML_comp2[symmetric])
apply(auto simp:pi_def)
done
thus ?thesis by simp
qed
finally show ?thesis .
qed
thus ?thesis by(simp cong:if_cong0 add:shift_subst_decr)
qed
finally have "?R = ?M" .
then show "?L = ?R" using \<open>?L = ?M\<close> by metis
qed
qed (simp_all add:list_eq_iff_nth_eq, (simp_all add:rev_nth)?)
section \<open>Compiler \label{sec:Compiler}\<close>
axiomatization arity :: "cname \<Rightarrow> nat"
primrec compile :: "tm \<Rightarrow> (nat \<Rightarrow> ml) \<Rightarrow> ml"
where
"compile (V x) \<sigma> = \<sigma> x"
| "compile (C nm) \<sigma> =
(if arity nm > 0 then Clo (C\<^sub>M\<^sub>L nm) [] (arity nm) else A\<^sub>M\<^sub>L (C\<^sub>M\<^sub>L nm) [])"
| "compile (s \<bullet> t) \<sigma> = apply (compile s \<sigma>) (compile t \<sigma>)"
| "compile (\<Lambda> t) \<sigma> = Clo (Lam\<^sub>M\<^sub>L (compile t (V\<^sub>M\<^sub>L 0 ## \<sigma>))) [] 1"
text\<open>Compiler for open terms and for terms with fixed free variables:\<close>
definition "comp_open t = compile t V\<^sub>M\<^sub>L"
abbreviation "comp_fixed t \<equiv> compile t (\<lambda>i. V\<^sub>U i [])"
text\<open>Compiled rules:\<close>
lemma size_args_less_size_tm[simp]: "s \<in> set (args_tm t) \<Longrightarrow> size s < size t"
by(induct t) auto
fun comp_pat where
"comp_pat t =
(case head_tm t of
C nm \<Rightarrow> C\<^sub>U nm (map comp_pat (rev (args_tm t)))
| V X \<Rightarrow> V\<^sub>M\<^sub>L X)"
declare comp_pat.simps[simp del] size_args_less_size_tm[simp del]
lemma comp_pat_V[simp]: "comp_pat(V X) = V\<^sub>M\<^sub>L X"
by(simp add:comp_pat.simps)
lemma comp_pat_C[simp]:
"comp_pat(C nm \<bullet>\<bullet> ts) = C\<^sub>U nm (map comp_pat (rev ts))"
by(simp add:comp_pat.simps)
lemma comp_pat_C_Nil[simp]: "comp_pat(C nm) = C\<^sub>U nm []"
by(simp add:comp_pat.simps)
overloading compR \<equiv> compR
begin
definition "compR \<equiv> (\<lambda>(nm,ts,t). (nm, map comp_pat (rev ts), comp_open t)) ` R"
end
lemma fv_ML_comp_open: "pure t \<Longrightarrow> fv\<^sub>M\<^sub>L(comp_open t) = fv t"
by(induct t pred:pure) (simp_all add:comp_open_def)
lemma fv_ML_comp_pat: "pattern t \<Longrightarrow> fv\<^sub>M\<^sub>L(comp_pat t) = fv t"
by(induct t pred:pattern)(simp_all add:comp_open_def)
lemma fv_compR_aux:
"(nm,ts,t') : R \<Longrightarrow> x \<in> fv\<^sub>M\<^sub>L (comp_open t')
\<Longrightarrow> \<exists>t\<in>set ts. x \<in> fv\<^sub>M\<^sub>L(comp_pat t)"
apply(frule pure_R)
apply(simp add:fv_ML_comp_open)
apply(frule (1) fv_R)
apply clarsimp
apply(rule bexI) prefer 2 apply assumption
apply(drule pattern_R)
apply(simp add:fv_ML_comp_pat)
done
lemma fv_compR:
"(nm,vs,v) : compR \<Longrightarrow> x \<in> fv\<^sub>M\<^sub>L v \<Longrightarrow> \<exists>u\<in>set vs. x \<in> fv\<^sub>M\<^sub>L u"
by(fastforce simp add:compR_def image_def dest: fv_compR_aux)
lemma lift_compile:
"pure t \<Longrightarrow> \<forall>\<sigma> k. lift k (compile t \<sigma>) = compile t (lift k \<circ> \<sigma>)"
apply(induct pred:pure)
apply simp_all
apply clarsimp
apply(rule_tac f = "compile t" in arg_cong)
apply(rule ext)
apply (clarsimp simp: lift_lift_ML_comm)
done
lemma subst_ML_compile:
"pure t \<Longrightarrow> subst\<^sub>M\<^sub>L \<sigma>' (compile t \<sigma>) = compile t (subst\<^sub>M\<^sub>L \<sigma>' o \<sigma>)"
apply(induct arbitrary: \<sigma> \<sigma>' pred:pure)
apply simp_all
apply(erule_tac x="V\<^sub>M\<^sub>L 0 ## \<sigma>'" in meta_allE)
apply(erule_tac x= "V\<^sub>M\<^sub>L 0 ## (lift\<^sub>M\<^sub>L 0 \<circ> \<sigma>)" in meta_allE)
apply(rule_tac f = "compile t" in arg_cong)
apply(rule ext)
apply (auto simp add:subst_ML_ext lift_ML_subst_ML)
done
theorem kernel_compile:
"pure t \<Longrightarrow> \<forall>i. \<sigma> i = V\<^sub>U i [] \<Longrightarrow> (compile t \<sigma>)! = t"
apply(induct arbitrary: \<sigma> pred:pure)
apply simp_all
apply(subst lift_compile) apply simp
apply(subst subst_ML_compile) apply simp
apply(subgoal_tac "(subst\<^sub>M\<^sub>L (\<lambda>n. if n = 0 then V\<^sub>U 0 [] else V\<^sub>M\<^sub>L (n - 1)) \<circ>
(lift 0 \<circ> V\<^sub>M\<^sub>L 0 ## \<sigma>)) = (\<lambda>a. V\<^sub>U a [])")
apply(simp)
apply(rule ext)
apply(simp)
done
lemma kernel_subst_ML_pat:
"pure t \<Longrightarrow> pattern t \<Longrightarrow> \<forall>i. closed\<^sub>M\<^sub>L 0 (\<sigma> i) \<Longrightarrow>
(subst\<^sub>M\<^sub>L \<sigma> (comp_pat t))! = subst (kernel \<circ> \<sigma>) t"
apply(induct arbitrary: \<sigma> pred:pure)
apply simp_all
apply(frule pattern_At_decomp)
apply(frule pattern_AtD12)
apply clarsimp
apply(subst comp_pat.simps)
apply(simp add: rev_map)
done
lemma kernel_subst_ML:
"pure t \<Longrightarrow> \<forall>i. closed\<^sub>M\<^sub>L 0 (\<sigma> i) \<Longrightarrow>
(subst\<^sub>M\<^sub>L \<sigma> (comp_open t))! = subst (kernel \<circ> \<sigma>) t"
proof(induct arbitrary: \<sigma> pred:pure)
case (Lam t)
have "lift 0 o V\<^sub>M\<^sub>L = V\<^sub>M\<^sub>L" by (simp add:fun_eq_iff)
hence "(subst\<^sub>M\<^sub>L \<sigma> (comp_open (\<Lambda> t)))! =
\<Lambda> (subst\<^sub>M\<^sub>L (lift 0 \<circ> V\<^sub>M\<^sub>L 0 ## \<sigma>) (comp_open t)[V\<^sub>U 0 []/0]!)"
using Lam by(simp add: lift_subst_ML comp_open_def lift_compile)
also have "\<dots> = \<Lambda> (subst (V 0 ## (kernel \<circ> \<sigma>)) t)" using Lam
by(simp add: subst_ML_comp subst_ext kernel_lift_tm)
also have "\<dots> = subst (kernel o \<sigma>) (\<Lambda> t)" by simp
finally show ?case .
qed (simp_all add:comp_open_def)
lemma kernel_subst_ML_pat_map:
"\<forall>t \<in> set ts. pure t \<Longrightarrow> patterns ts \<Longrightarrow> \<forall>i. closed\<^sub>M\<^sub>L 0 (\<sigma> i) \<Longrightarrow>
map kernel (map (subst\<^sub>M\<^sub>L \<sigma>) (map comp_pat ts)) =
map (subst (kernel \<circ> \<sigma>)) ts"
by(simp add:list_eq_iff_nth_eq kernel_subst_ML_pat)
lemma compR_Red_tm: "(nm, vs, v) : compR \<Longrightarrow> \<forall> i. closed\<^sub>M\<^sub>L 0 (\<sigma> i)
\<Longrightarrow> C nm \<bullet>\<bullet> (map (subst\<^sub>M\<^sub>L \<sigma>) (rev vs))! \<rightarrow>* (subst\<^sub>M\<^sub>L \<sigma> v)!"
apply(auto simp add:compR_def rev_map simp del: map_map)
apply(frule pure_R)
apply(subst kernel_subst_ML) apply fast+
apply(subst kernel_subst_ML_pat_map)
apply fast
apply(fast dest:pattern_R)
apply assumption
apply(rule r_into_rtrancl)
apply(erule Red_tm.intros)
done
section "Correctness"
(* Without this special rule one "also" in the next proof *diverges*,
probably because of HOU. *)
lemma eq_Red_tm_trans: "s = t \<Longrightarrow> t \<rightarrow> t' \<Longrightarrow> s \<rightarrow> t'"
by simp
text\<open>Soundness of reduction:\<close>
theorem fixes v :: ml shows Red_ml_sound:
"v \<Rightarrow> v' \<Longrightarrow> closed\<^sub>M\<^sub>L 0 v \<Longrightarrow> v! \<rightarrow>* v'! \<and> closed\<^sub>M\<^sub>L 0 v'" and
"vs \<Rightarrow> vs' \<Longrightarrow> \<forall>v\<in>set vs. closed\<^sub>M\<^sub>L 0 v \<Longrightarrow>
vs! \<rightarrow>* vs'! \<and> (\<forall>v'\<in>set vs'. closed\<^sub>M\<^sub>L 0 v')"
proof(induct rule:Red_ml_Red_ml_list.inducts)
fix u v
let ?v = "A\<^sub>M\<^sub>L (Lam\<^sub>M\<^sub>L u) [v]"
assume cl: "closed\<^sub>M\<^sub>L 0 (A\<^sub>M\<^sub>L (Lam\<^sub>M\<^sub>L u) [v])"
let ?u' = "(lift_ml 0 u)[V\<^sub>U 0 []/0]"
have "?v! = (\<Lambda>((?u')!)) \<bullet> (v !)" by simp
also have "\<dots> \<rightarrow> (?u' !)[v!/0]" (is "_ \<rightarrow> ?R") by(rule Red_tm.intros)
also(eq_Red_tm_trans) have "?R = u[v/0]!" using cl
apply(cut_tac u = "u" and v = "v" in kernel_subst1)
apply(simp_all)
done
finally have "kernel(A\<^sub>M\<^sub>L (Lam\<^sub>M\<^sub>L u) [v]) \<rightarrow>* kernel(u[v/0])" (is ?A)
by(rule r_into_rtrancl)
moreover have "closed\<^sub>M\<^sub>L 0 (u[v/0])" (is "?C")
proof -
let ?\<sigma> = "\<lambda>n. if n = 0 then v else V\<^sub>M\<^sub>L (n - 1)"
let ?\<sigma>' = "\<lambda>n. v"
have clu: "closed\<^sub>M\<^sub>L (Suc 0) u" and clv: "closed\<^sub>M\<^sub>L 0 v" using cl by simp+
have "closed\<^sub>M\<^sub>L 0 (subst\<^sub>M\<^sub>L ?\<sigma>' u)"
by (metis closed_ML_subst_ML clv)
hence "closed\<^sub>M\<^sub>L 0 (subst\<^sub>M\<^sub>L ?\<sigma> u)"
using subst_ML_coincidence[OF clu, of ?\<sigma> ?\<sigma>'] by auto
thus ?thesis by simp
qed
ultimately show "?A \<and> ?C" ..
next
fix \<sigma> :: "nat \<Rightarrow> ml" and nm vs v
assume \<sigma>: "\<forall>i. closed\<^sub>M\<^sub>L 0 (\<sigma> i)" and compR: "(nm, vs, v) \<in> compR"
have "map (subst V) (map (subst\<^sub>M\<^sub>L \<sigma>) (rev vs)!) = map (subst\<^sub>M\<^sub>L \<sigma>) (rev vs)!"
by(simp add:list_eq_iff_nth_eq subst_V_kernel closed_ML_subst_ML[OF \<sigma>])
with compR_Red_tm[OF compR \<sigma>]
have "(C nm) \<bullet>\<bullet> ((map (subst\<^sub>M\<^sub>L \<sigma>) (rev vs)) !) \<rightarrow>* (subst\<^sub>M\<^sub>L \<sigma> v)!"
by(simp add:subst_V_kernel closed_ML_subst_ML[OF \<sigma>])
hence "A\<^sub>M\<^sub>L (C\<^sub>M\<^sub>L nm) (map (subst\<^sub>M\<^sub>L \<sigma>) vs)! \<rightarrow>* subst\<^sub>M\<^sub>L \<sigma> v!" (is ?A)
by(simp add:rev_map)
moreover
have "closed\<^sub>M\<^sub>L 0 (subst\<^sub>M\<^sub>L \<sigma> v)" (is ?C) by(metis closed_ML_subst_ML \<sigma>)
ultimately show "?A \<and> ?C" ..
qed (auto simp:Reds_tm_list_foldl_At Red_tm_rev rev_map[symmetric])
theorem Red_term_sound:
"t \<Rightarrow> t' \<Longrightarrow> closed\<^sub>M\<^sub>L 0 t \<Longrightarrow> kernelt t \<rightarrow>* kernelt t' \<and> closed\<^sub>M\<^sub>L 0 t'"
proof(induct rule:Red_term.inducts)
case term_C thus ?case
by (auto simp:closed_tm_ML_foldl_At)
next
case term_V thus ?case
by (auto simp:closed_tm_ML_foldl_At)
next
case (term_Clo vf vs n)
hence "(lift 0 vf!) \<bullet>\<bullet> map kernel (rev (map (lift 0) vs))
= lift 0 (vf! \<bullet>\<bullet> (rev vs)!)"
apply (simp add:kernel_lift_tm list_eq_iff_nth_eq)
apply(simp add:rev_nth rev_map kernel_lift_tm)
done
hence "term (Clo vf vs n)! \<rightarrow>*
\<Lambda> (term (apply (lift 0 (Clo vf vs n)) (V\<^sub>U 0 [])))!"
using term_Clo
by(simp del:lift_foldl_At add: r_into_rtrancl Red_tm.intros(2))
moreover
have "closed\<^sub>M\<^sub>L 0 (\<Lambda> (term (apply (lift 0 (Clo vf vs n)) (V\<^sub>U 0 []))))"
using term_Clo by simp
ultimately show ?case ..
next
case ctxt_term thus ?case by simp (metis Red_ml_sound)
qed auto
corollary kernel_inv:
"(t :: tm) \<Rightarrow>* t' \<Longrightarrow> closed\<^sub>M\<^sub>L 0 t \<Longrightarrow> t! \<rightarrow>* t'! \<and> closed\<^sub>M\<^sub>L 0 t' "
apply(induct rule: rtrancl.induct)
apply (metis rtrancl_eq_or_trancl)
apply (metis Red_term_sound rtrancl_trans)
done
lemma closed_ML_compile:
"pure t \<Longrightarrow> \<forall>i. closed\<^sub>M\<^sub>L n (\<sigma> i) \<Longrightarrow> closed\<^sub>M\<^sub>L n (compile t \<sigma>)"
proof(induct arbitrary:n \<sigma> pred:pure)
case (Lam t)
have 1: "\<forall>i. closed\<^sub>M\<^sub>L (Suc n) ((V\<^sub>M\<^sub>L 0 ## \<sigma>) i)" using Lam(3-)
by (auto simp: closed_ML_Suc)
show ?case using Lam(2)[OF 1] by (simp del:apply_cons_ML)
qed simp_all
theorem nbe_correct: fixes t :: tm
assumes "pure t" and "term (comp_fixed t) \<Rightarrow>* t'" and "pure t'" shows "t \<rightarrow>* t'"
proof -
have ML_cl: "closed\<^sub>M\<^sub>L 0 (term (comp_fixed t))"
by (simp add: closed_ML_compile[OF \<open>pure t\<close>])
have "(term (comp_fixed t))! = t"
using kernel_compile[OF \<open>pure t\<close>] by simp
moreover have "term (comp_fixed t)! \<rightarrow>* t'!"
using kernel_inv[OF assms(2) ML_cl] by auto
ultimately have "t \<rightarrow>* t'!" by simp
thus ?thesis using kernel_pure[OF \<open>pure t'\<close>] by simp
qed
section "Normal Forms"
inductive normal :: "tm \<Rightarrow> bool" where
"\<forall>t\<in>set ts. normal t \<Longrightarrow> normal(V x \<bullet>\<bullet> ts)" |
"normal t \<Longrightarrow> normal(\<Lambda> t)" |
"\<forall>t\<in>set ts. normal t \<Longrightarrow>
\<forall>\<sigma>. \<forall>(nm',ls,r)\<in>R. \<not>(nm = nm' \<and> take (size ls) ts = map (subst \<sigma>) ls)
\<Longrightarrow> normal(C nm \<bullet>\<bullet> ts)"
fun C_normal_ML :: "ml \<Rightarrow> bool" ("C'_normal\<^sub>M\<^sub>L") where
"C_normal\<^sub>M\<^sub>L(C\<^sub>U nm vs) =
((\<forall>v\<in>set vs. C_normal\<^sub>M\<^sub>L v) \<and> no_match_compR nm vs)" |
"C_normal\<^sub>M\<^sub>L (C\<^sub>M\<^sub>L _) = True" |
"C_normal\<^sub>M\<^sub>L (V\<^sub>M\<^sub>L _) = True" |
"C_normal\<^sub>M\<^sub>L (A\<^sub>M\<^sub>L v vs) = (C_normal\<^sub>M\<^sub>L v \<and> (\<forall>v \<in> set vs. C_normal\<^sub>M\<^sub>L v))" |
"C_normal\<^sub>M\<^sub>L (Lam\<^sub>M\<^sub>L v) = C_normal\<^sub>M\<^sub>L v" |
"C_normal\<^sub>M\<^sub>L (V\<^sub>U x vs) = (\<forall>v \<in> set vs. C_normal\<^sub>M\<^sub>L v)" |
"C_normal\<^sub>M\<^sub>L (Clo v vs _) = (C_normal\<^sub>M\<^sub>L v \<and> (\<forall>v \<in> set vs. C_normal\<^sub>M\<^sub>L v))" |
"C_normal\<^sub>M\<^sub>L (apply u v) = (C_normal\<^sub>M\<^sub>L u \<and> C_normal\<^sub>M\<^sub>L v)"
fun size_tm :: "tm \<Rightarrow> nat" where
"size_tm (C _) = 1" |
"size_tm (At s t) = size_tm s + size_tm t + 1" |
"size_tm _ = 0"
lemma size_tm_foldl_At: "size_tm(t \<bullet>\<bullet> ts) = size_tm t + size_list size_tm ts"
by (induct ts arbitrary:t) auto
lemma termination_no_match:
"i < length ss \<Longrightarrow> ss ! i = C nm \<bullet>\<bullet> ts
\<Longrightarrow> sum_list (map size_tm ts) < sum_list (map size_tm ss)"
apply(subgoal_tac "C nm \<bullet>\<bullet> ts : set ss")
apply(drule sum_list_map_remove1[of _ _ size_tm])
apply(simp add:size_tm_foldl_At size_list_conv_sum_list)
apply (metis in_set_conv_nth)
done
declare conj_cong [fundef_cong]
function no_match :: "tm list \<Rightarrow> tm list \<Rightarrow> bool" where
"no_match ps ts =
(\<exists>i < min (size ts) (size ps).
\<exists>nm nm' rs rs'. ps!i = (C nm) \<bullet>\<bullet> rs \<and> ts!i = (C nm') \<bullet>\<bullet> rs' \<and>
(nm=nm' \<longrightarrow> no_match rs rs'))"
by pat_completeness auto
termination
apply(relation "measure(%(ts::tm list,_). \<Sum>t\<leftarrow>ts. size_tm t)")
apply (auto simp:termination_no_match)
done
declare no_match.simps[simp del]
abbreviation
"no_match_R nm ts \<equiv> \<forall>(nm',ps,t)\<in> R. nm=nm' \<longrightarrow> no_match ps ts"
lemma no_match: "no_match ps ts \<Longrightarrow> \<not>(\<exists>\<sigma>. map (subst \<sigma>) ps = ts)"
proof(induct ps ts rule:no_match.induct)
case (1 ps ts)
thus ?case
apply auto
apply(subst (asm) no_match.simps[of ps])
apply fastforce
done
qed
lemma no_match_take: "no_match ps ts \<Longrightarrow> no_match ps (take (size ps) ts)"
apply(subst (asm) no_match.simps)
apply(subst no_match.simps)
apply fastforce
done
fun dterm_ML :: "ml \<Rightarrow> tm" ("dterm\<^sub>M\<^sub>L") where
"dterm\<^sub>M\<^sub>L (C\<^sub>U nm vs) = C nm \<bullet>\<bullet> map dterm\<^sub>M\<^sub>L (rev vs)" |
"dterm\<^sub>M\<^sub>L _ = V 0"
fun dterm :: "tm \<Rightarrow> tm" where
"dterm (V n) = V n" |
"dterm (C nm) = C nm" |
"dterm (s \<bullet> t) = dterm s \<bullet> dterm t" |
"dterm (\<Lambda> t) = \<Lambda> (dterm t)" |
"dterm (term v) = dterm\<^sub>M\<^sub>L v"
lemma dterm_pure[simp]: "pure t \<Longrightarrow> dterm t = t"
by (induct pred:pure) auto
lemma map_dterm_pure[simp]: "\<forall>t\<in>set ts. pure t \<Longrightarrow> map dterm ts = ts"
by (induct ts) auto
lemma map_dterm_term[simp]: "map dterm (map term vs) = map dterm\<^sub>M\<^sub>L vs"
by (induct vs) auto
lemma dterm_foldl_At[simp]: "dterm(t \<bullet>\<bullet> ts) = dterm t \<bullet>\<bullet> map dterm ts"
by(induct ts arbitrary: t) auto
lemma no_match_coincide:
"no_match\<^sub>M\<^sub>L ps vs \<Longrightarrow>
no_match (map dterm\<^sub>M\<^sub>L (rev ps)) (map dterm\<^sub>M\<^sub>L (rev vs))"
apply(induct ps vs rule:no_match_ML.induct)
apply(rotate_tac 1)
apply(subst (asm) no_match_ML.simps)
apply (elim exE conjE)
apply(case_tac "nm=nm'")
prefer 2
apply(subst no_match.simps)
apply(rule_tac x=i in exI)
apply rule
apply (simp (no_asm))
apply (metis min_less_iff_conj)
apply(simp add:min_less_iff_conj nth_map)
apply safe
apply(erule_tac x=i in meta_allE)
apply(erule_tac x=nm' in meta_allE)
apply(erule_tac x=nm' in meta_allE)
apply(erule_tac x="vs" in meta_allE)
apply(erule_tac x="vs'" in meta_allE)
apply(subst no_match.simps)
apply(rule_tac x=i in exI)
apply rule
apply (simp (no_asm))
apply (metis min_less_iff_conj)
apply(rule_tac x=nm' in exI)
apply(rule_tac x=nm' in exI)
apply(rule_tac x="map dterm\<^sub>M\<^sub>L (rev vs)" in exI)
apply(rule_tac x="map dterm\<^sub>M\<^sub>L (rev vs')" in exI)
apply(simp)
done
lemma dterm_ML_comp_patD:
"pattern t \<Longrightarrow> dterm\<^sub>M\<^sub>L (comp_pat t) = C nm \<bullet>\<bullet> rs \<Longrightarrow> \<exists>ts. t = C nm \<bullet>\<bullet> ts"
by(induct pred:pattern) simp_all
lemma no_match_R_coincide_aux[rule_format]: "patterns ts \<Longrightarrow>
no_match (map (dterm\<^sub>M\<^sub>L \<circ> comp_pat) ts) rs \<longrightarrow> no_match ts rs"
apply(induct ts rs rule:no_match.induct)
apply(subst (1 2) no_match.simps)
apply clarsimp
apply(rule_tac x=i in exI)
apply simp
apply(rule_tac x=nm in exI)
apply(cut_tac t = "ps!i" in dterm_ML_comp_patD, simp, assumption)
apply(clarsimp)
apply(erule_tac x = i in meta_allE)
apply(erule_tac x = nm' in meta_allE)
apply(erule_tac x = nm' in meta_allE)
apply(erule_tac x = tsa in meta_allE)
apply(erule_tac x = rs' in meta_allE)
apply (simp add:rev_map)
apply (metis in_set_conv_nth pattern_At_vecD)
done
lemma no_match_R_coincide:
"no_match_compR nm (rev vs) \<Longrightarrow> no_match_R nm (map dterm\<^sub>M\<^sub>L vs)"
apply auto
apply(drule_tac x="(nm, map comp_pat (rev aa), comp_open b)" in bspec)
unfolding compR_def
apply (simp add:image_def)
apply (force)
apply (simp)
apply(drule no_match_coincide)
apply(frule pure_R)
apply(drule pattern_R)
apply(clarsimp simp add: rev_map no_match.simps[of _ "map dterm\<^sub>M\<^sub>L vs"])
apply(rule_tac x=i in exI)
apply simp
apply(cut_tac t = "aa!i" in dterm_ML_comp_patD, simp, assumption)
apply clarsimp
apply(auto simp: rev_map)
apply(rule no_match_R_coincide_aux)
prefer 2 apply assumption
apply (metis in_set_conv_nth pattern_At_vecD)
done
inductive C_normal :: "tm \<Rightarrow> bool" where
"\<forall>t\<in>set ts. C_normal t \<Longrightarrow> C_normal(V x \<bullet>\<bullet> ts)" |
"C_normal t \<Longrightarrow> C_normal(\<Lambda> t)" |
"C_normal\<^sub>M\<^sub>L v \<Longrightarrow> C_normal(term v)" |
"\<forall>t\<in>set ts. C_normal t \<Longrightarrow> no_match_R nm (map dterm ts)
\<Longrightarrow> C_normal(C nm \<bullet>\<bullet> ts)"
declare C_normal.intros[simp]
lemma C_normal_term[simp]: "C_normal(term v) = C_normal\<^sub>M\<^sub>L v"
apply (auto)
apply(erule C_normal.cases)
apply auto
done
lemma [simp]: "C_normal(\<Lambda> t) = C_normal t"
apply (auto)
apply(erule C_normal.cases)
apply auto
done
lemma [simp]: "C_normal(V x)"
using C_normal.intros(1)[of "[]" x]
by simp
lemma [simp]: "dterm (dterm\<^sub>M\<^sub>L v) = dterm\<^sub>M\<^sub>L v"
apply(induct v rule:dterm_ML.induct)
apply simp_all
done
lemma "u\<Rightarrow>(v::ml) \<Longrightarrow> True" and
Red_ml_list_length: "vs \<Rightarrow> vs' \<Longrightarrow> length vs = length vs'"
by(induct rule: Red_ml_Red_ml_list.inducts) simp_all
lemma "(v::ml) \<Rightarrow> v' \<Longrightarrow> True" and
Red_ml_list_nth: "(vs::ml list) \<Rightarrow> vs'
\<Longrightarrow> \<exists>v' k. k<size vs \<and> vs!k \<Rightarrow> v' \<and> vs' = vs[k := v']"
apply (induct rule: Red_ml_Red_ml_list.inducts)
apply (auto split:nat.splits)
done
lemma Red_ml_list_pres_no_match:
"no_match\<^sub>M\<^sub>L ps vs \<Longrightarrow> vs \<Rightarrow> vs' \<Longrightarrow> no_match\<^sub>M\<^sub>L ps vs'"
proof(induct ps vs arbitrary: vs' rule:no_match_ML.induct)
case (1 vs os)
show ?case using 1(2-3)
apply-
apply(frule Red_ml_list_length)
apply(rotate_tac -2)
apply(subst (asm) no_match_ML.simps)
apply clarify
apply(rename_tac i nm nm' us us')
apply(subst no_match_ML.simps)
apply(rule_tac x=i in exI)
apply (simp)
apply(drule Red_ml_list_nth)
apply clarify
apply(rename_tac k)
apply(case_tac "k = length os - Suc i")
prefer 2
apply(rule_tac x=nm' in exI)
apply(rule_tac x=us' in exI)
apply (simp add: rev_nth nth_list_update)
apply (simp add: rev_nth)
apply(erule Red_ml.cases)
apply simp_all
apply(fastforce intro: 1(1) simp add:rev_nth)
done
qed
lemma no_match_ML_subst_ML[rule_format]:
"\<forall>v\<in>set vs. \<forall>x\<in>fv\<^sub>M\<^sub>L v. C_normal\<^sub>M\<^sub>L (\<sigma> x) \<Longrightarrow>
no_match\<^sub>M\<^sub>L ps vs \<longrightarrow> no_match\<^sub>M\<^sub>L ps (map (subst\<^sub>M\<^sub>L \<sigma>) vs)"
apply(induct ps vs rule:no_match_ML.induct)
apply simp
apply(subst (1 2) no_match_ML.simps)
apply clarsimp
apply(rule_tac x=i in exI)
apply simp
apply(rule_tac x=nm' in exI)
apply(rule_tac x="map (subst\<^sub>M\<^sub>L \<sigma>) vs'" in exI)
apply (auto simp:rev_nth)
apply(erule_tac x = i in meta_allE)
apply(erule_tac x = nm' in meta_allE)
apply(erule_tac x = nm' in meta_allE)
apply(erule_tac x = vs in meta_allE)
apply(erule_tac x = vs' in meta_allE)
apply simp
apply (metis UN_I fv_ML.simps(5) in_set_conv_nth length_rev rev_nth set_rev)
done
lemma lift_is_CUD:
"lift\<^sub>M\<^sub>L k v = C\<^sub>U nm vs' \<Longrightarrow> \<exists>vs. v = C\<^sub>U nm vs \<and> vs' = map (lift\<^sub>M\<^sub>L k) vs"
by(cases v) auto
lemma no_match_ML_lift_ML:
"no_match\<^sub>M\<^sub>L ps (map (lift\<^sub>M\<^sub>L k) vs) = no_match\<^sub>M\<^sub>L ps vs"
apply(induct ps vs rule:no_match_ML.induct)
apply simp
apply(subst (1 2) no_match_ML.simps)
apply rule
apply clarsimp
apply(rule_tac x=i in exI)
apply (simp add:rev_nth)
apply(drule lift_is_CUD)
apply fastforce
apply clarsimp
apply(rule_tac x=i in exI)
apply simp
apply(rule_tac x=nm' in exI)
apply(rule_tac x="map (lift\<^sub>M\<^sub>L k) vs'" in exI)
apply (fastforce simp:rev_nth)
done
lemma C_normal_ML_lift_ML: "C_normal\<^sub>M\<^sub>L(lift\<^sub>M\<^sub>L k v) = C_normal\<^sub>M\<^sub>L v"
by(induct v arbitrary: k rule:C_normal_ML.induct)(auto simp:no_match_ML_lift_ML)
lemma no_match_compR_Cons:
"no_match_compR nm vs \<Longrightarrow> no_match_compR nm (v # vs)"
apply auto
apply(drule bspec, assumption)
apply simp
apply(subst (asm) no_match_ML.simps)
apply(subst no_match_ML.simps)
apply clarsimp
apply(rule_tac x=i in exI)
apply (simp add:nth_append)
done
lemma C_normal_ML_comp_open: "pure t \<Longrightarrow> C_normal\<^sub>M\<^sub>L(comp_open t)"
by (induct pred:pure) (auto simp:comp_open_def)
lemma C_normal_compR_rhs: "(nm, vs, v) \<in> compR \<Longrightarrow> C_normal\<^sub>M\<^sub>L v"
by(auto simp: compR_def image_def Bex_def pure_R C_normal_ML_comp_open)
lemma C_normal_ML_subst_ML:
"C_normal\<^sub>M\<^sub>L (subst\<^sub>M\<^sub>L \<sigma> v) \<Longrightarrow> (\<forall>x\<in>fv\<^sub>M\<^sub>L v. C_normal\<^sub>M\<^sub>L (\<sigma> x))"
proof(induct \<sigma> v rule:subst_ML.induct)
case 4 thus ?case
by(simp del:apply_cons_ML)(force simp add: C_normal_ML_lift_ML)
(* weird - force suffices in apply style *)
qed auto
lemma C_normal_ML_subst_ML_iff: "C_normal\<^sub>M\<^sub>L v \<Longrightarrow>
C_normal\<^sub>M\<^sub>L (subst\<^sub>M\<^sub>L \<sigma> v) \<longleftrightarrow> (\<forall>x\<in>fv\<^sub>M\<^sub>L v. C_normal\<^sub>M\<^sub>L (\<sigma> x))"
proof(induct \<sigma> v rule:subst_ML.induct)
case 4 thus ?case
by(simp del:apply_cons_ML)(force simp add: C_normal_ML_lift_ML)
(* weird - force suffices in apply style *)
next
case 5 thus ?case by simp (blast intro: no_match_ML_subst_ML)
qed auto
lemma C_normal_ML_inv: "v \<Rightarrow> v' \<Longrightarrow> C_normal\<^sub>M\<^sub>L v \<Longrightarrow> C_normal\<^sub>M\<^sub>L v'" and
"vs \<Rightarrow> vs' \<Longrightarrow> \<forall>v\<in>set vs. C_normal\<^sub>M\<^sub>L v \<Longrightarrow> \<forall>v'\<in>set vs'. C_normal\<^sub>M\<^sub>L v'"
apply(induct rule:Red_ml_Red_ml_list.inducts)
apply(simp_all add: C_normal_ML_subst_ML_iff)
apply(metis C_normal_ML_subst_ML C_normal_compR_rhs
fv_compR C_normal_ML_subst_ML_iff)
apply(blast intro!:no_match_compR_Cons)
apply(blast dest:Red_ml_list_pres_no_match)
done
lemma Red_term_hnf_induct[consumes 1]:
assumes "(t::tm) \<Rightarrow> t'"
"\<And>nm vs ts. P ((term (C\<^sub>U nm vs)) \<bullet>\<bullet> ts) ((C nm \<bullet>\<bullet> map term (rev vs)) \<bullet>\<bullet> ts)"
"\<And>x vs ts. P (term (V\<^sub>U x vs) \<bullet>\<bullet> ts) ((V x \<bullet>\<bullet> map term (rev vs)) \<bullet>\<bullet> ts)"
"\<And>vf vs n ts.
P (term (Clo vf vs n) \<bullet>\<bullet> ts)
((\<Lambda> (term (apply (lift 0 (Clo vf vs n)) (V\<^sub>U 0 [])))) \<bullet>\<bullet> ts)"
"\<And>t t' ts. \<lbrakk>t \<Rightarrow> t'; P t t'\<rbrakk> \<Longrightarrow> P (\<Lambda> t \<bullet>\<bullet> ts) (\<Lambda> t' \<bullet>\<bullet> ts)"
"\<And>v v' ts. v \<Rightarrow> v' \<Longrightarrow> P (term v \<bullet>\<bullet> ts) (term v' \<bullet>\<bullet> ts)"
"\<And>x i t' ts. i<size ts \<Longrightarrow> ts!i \<Rightarrow> t' \<Longrightarrow> P (ts!i) (t')
\<Longrightarrow> P (V x \<bullet>\<bullet> ts) (V x \<bullet>\<bullet> ts[i:=t'])"
"\<And>nm i t' ts. i<size ts \<Longrightarrow> ts!i \<Rightarrow> t' \<Longrightarrow> P (ts!i) (t')
\<Longrightarrow> P (C nm \<bullet>\<bullet> ts) (C nm \<bullet>\<bullet> ts[i:=t'])"
"\<And>t i t' ts. i<size ts \<Longrightarrow> ts!i \<Rightarrow> t' \<Longrightarrow> P (ts!i) (t')
\<Longrightarrow> P (\<Lambda> t \<bullet>\<bullet> ts) (\<Lambda> t \<bullet>\<bullet> ts[i:=t'])"
"\<And>v i t' ts. i<size ts \<Longrightarrow> ts!i \<Rightarrow> t' \<Longrightarrow> P (ts!i) (t')
\<Longrightarrow> P (term v \<bullet>\<bullet> ts) (term v \<bullet>\<bullet> (ts[i:=t']))"
shows "P t t'"
proof-
{ fix ts from assms have "P (t \<bullet>\<bullet> ts) (t' \<bullet>\<bullet> ts)"
proof(induct arbitrary: ts rule:Red_term.induct)
case term_C thus ?case by metis
next
case term_V thus ?case by metis
next
case term_Clo thus ?case by metis
next
case ctxt_Lam thus ?case by simp (metis foldl_Nil)
next
case (ctxt_At1 s s' t ts)
thus ?case using ctxt_At1(2)[of "t#ts"] by simp
next
case (ctxt_At2 t t' s ts)
{ fix n rs assume "s = V n \<bullet>\<bullet> rs"
hence ?case using ctxt_At2(8)[of "size rs" "rs @ t # ts" t' n] ctxt_At2
by simp (metis foldl_Nil)
} moreover
{ fix nm rs assume "s = C nm \<bullet>\<bullet> rs"
hence ?case using ctxt_At2(9)[of "size rs" "rs @ t # ts" t' nm] ctxt_At2
by simp (metis foldl_Nil)
} moreover
{ fix r rs assume "s = \<Lambda> r \<bullet>\<bullet> rs"
hence ?case using ctxt_At2(10)[of "size rs" "rs @ t # ts" t'] ctxt_At2
by simp (metis foldl_Nil)
} moreover
{ fix v rs assume "s = term v \<bullet>\<bullet> rs"
hence ?case using ctxt_At2(11)[of "size rs" "rs @ t # ts" t'] ctxt_At2
by simp (metis foldl_Nil)
} ultimately show ?case using tm_vector_cases[of s] by blast
qed
}
from this[of "[]"] show ?thesis by simp
qed
corollary Red_term_hnf_cases[consumes 1]:
assumes "(t::tm) \<Rightarrow> t'"
"\<And>nm vs ts.
t = term (C\<^sub>U nm vs) \<bullet>\<bullet> ts \<Longrightarrow> t' = (C nm \<bullet>\<bullet> map term (rev vs)) \<bullet>\<bullet> ts \<Longrightarrow> P"
"\<And>x vs ts.
t = term (V\<^sub>U x vs) \<bullet>\<bullet> ts \<Longrightarrow> t' = (V x \<bullet>\<bullet> map term (rev vs)) \<bullet>\<bullet> ts \<Longrightarrow> P"
"\<And>vf vs n ts. t = term (Clo vf vs n) \<bullet>\<bullet> ts \<Longrightarrow>
t' = \<Lambda> (term (apply (lift 0 (Clo vf vs n)) (V\<^sub>U 0 []))) \<bullet>\<bullet> ts \<Longrightarrow> P"
"\<And>s s' ts. t = \<Lambda> s \<bullet>\<bullet> ts \<Longrightarrow> t' = \<Lambda> s' \<bullet>\<bullet> ts \<Longrightarrow> s \<Rightarrow> s' \<Longrightarrow> P"
"\<And>v v' ts. t = term v \<bullet>\<bullet> ts \<Longrightarrow> t' = term v' \<bullet>\<bullet> ts \<Longrightarrow> v \<Rightarrow> v' \<Longrightarrow> P"
"\<And>x i r' ts. i<size ts \<Longrightarrow> ts!i \<Rightarrow> r'
\<Longrightarrow> t = V x \<bullet>\<bullet> ts \<Longrightarrow> t' = V x \<bullet>\<bullet> ts[i:=r'] \<Longrightarrow> P"
"\<And>nm i r' ts. i<size ts \<Longrightarrow> ts!i \<Rightarrow> r'
\<Longrightarrow> t = C nm \<bullet>\<bullet> ts \<Longrightarrow> t' = C nm \<bullet>\<bullet> ts[i:=r'] \<Longrightarrow> P"
"\<And>s i r' ts. i<size ts \<Longrightarrow> ts!i \<Rightarrow> r'
\<Longrightarrow> t = \<Lambda> s \<bullet>\<bullet> ts \<Longrightarrow> t' = \<Lambda> s \<bullet>\<bullet> ts[i:=r'] \<Longrightarrow> P"
"\<And>v i r' ts. i<size ts \<Longrightarrow> ts!i \<Rightarrow> r'
\<Longrightarrow> t = term v \<bullet>\<bullet> ts \<Longrightarrow> t' = term v \<bullet>\<bullet> (ts[i:=r']) \<Longrightarrow> P"
shows "P" using assms
apply -
apply(induct rule:Red_term_hnf_induct)
apply metis+
done
lemma [simp]: "C_normal(term v \<bullet>\<bullet> ts) \<longleftrightarrow> C_normal\<^sub>M\<^sub>L v \<and> ts = []"
by(fastforce elim: C_normal.cases)
lemma [simp]: "C_normal(\<Lambda> t \<bullet>\<bullet> ts) \<longleftrightarrow> C_normal t \<and> ts = []"
by(fastforce elim: C_normal.cases)
lemma [simp]: "C_normal(C nm \<bullet>\<bullet> ts) \<longleftrightarrow>
(\<forall>t\<in>set ts. C_normal t) \<and> no_match_R nm (map dterm ts)"
by(fastforce elim: C_normal.cases)
lemma [simp]: "C_normal(V x \<bullet>\<bullet> ts) \<longleftrightarrow> (\<forall>t \<in> set ts. C_normal t)"
by(fastforce elim: C_normal.cases)
lemma no_match_ML_lift:
"no_match\<^sub>M\<^sub>L ps vs \<longrightarrow> no_match\<^sub>M\<^sub>L ps (map (lift k) vs)"
apply(induct ps vs rule:no_match_ML.induct)
apply simp
apply(subst (1 2) no_match_ML.simps)
apply clarsimp
apply(rule_tac x=i in exI)
apply simp
apply(rule_tac x=nm' in exI)
apply(rule_tac x="map (lift k) vs'" in exI)
apply (fastforce simp:rev_nth)
done
lemma no_match_compR_lift:
"no_match_compR nm vs \<Longrightarrow> no_match_compR nm (map (lift k) vs)"
by (fastforce simp: no_match_ML_lift)
lemma [simp]: "C_normal\<^sub>M\<^sub>L v \<Longrightarrow> C_normal\<^sub>M\<^sub>L(lift k v)"
apply(induct v arbitrary:k rule:lift_ml.induct)
apply(simp_all add:no_match_compR_lift)
done
declare [[simp_depth_limit = 10]]
lemma Red_term_pres_no_match:
"\<lbrakk>i < length ts; ts ! i \<Rightarrow> t'; no_match ps dts; dts = (map dterm ts)\<rbrakk>
\<Longrightarrow> no_match ps (map dterm (ts[i := t']))"
proof(induct ps dts arbitrary: ts i t' rule:no_match.induct)
case (1 ps dts ts i t')
from \<open>no_match ps dts\<close> \<open>dts = map dterm ts\<close>
obtain j nm nm' rs rs' where ob: "j < size ts" "j < size ps"
"ps!j = C nm \<bullet>\<bullet> rs" "dterm (ts!j) = C nm' \<bullet>\<bullet> rs'"
"nm = nm' \<longrightarrow> no_match rs rs'"
by (subst (asm) no_match.simps) fastforce
show ?case
proof (subst no_match.simps)
show "\<exists>k<min (length (map dterm (ts[i := t']))) (length ps).
\<exists>nm nm' rs rs'. ps!k = C nm \<bullet>\<bullet> rs \<and>
map dterm (ts[i := t']) ! k = C nm' \<bullet>\<bullet> rs' \<and>
(nm = nm' \<longrightarrow> no_match rs rs')"
(is "\<exists>k < ?m. ?P k")
proof-
{ assume [simp]: "j=i"
have "\<exists>rs'. dterm t' = C nm' \<bullet>\<bullet> rs' \<and> (nm = nm' \<longrightarrow> no_match rs rs')"
using \<open>ts ! i \<Rightarrow> t'\<close>
proof(cases rule:Red_term_hnf_cases)
case (5 v v' ts'')
then obtain vs where [simp]:
"v = C\<^sub>U nm' vs" "rs' = map dterm\<^sub>M\<^sub>L (rev vs) @ map dterm ts''"
using ob by(cases v) auto
obtain vs' where [simp]: "v' = C\<^sub>U nm' vs'" "vs \<Rightarrow> vs'"
using \<open>v\<Rightarrow>v'\<close> by(rule Red_ml.cases) auto
obtain v' k where [arith]: "k<size vs" and "vs!k \<Rightarrow> v'"
and [simp]: "vs' = vs[k := v']"
using Red_ml_list_nth[OF \<open>vs\<Rightarrow>vs'\<close>] by fastforce
show ?thesis (is "\<exists>rs'. ?P rs' \<and> ?Q rs'")
proof
let ?rs' = "map dterm ((map term (rev vs) @ ts'')[(size vs - k - 1):=term v'])"
have "?P ?rs'" using ob 5
by(simp add: list_update_append map_update[symmetric] rev_update)
moreover have "?Q ?rs'"
apply rule
apply(rule "1.hyps"[OF _ ob(3)])
using "1.prems" 5 ob
apply (auto simp:nth_append rev_nth ctxt_term[OF \<open>vs!k \<Rightarrow> v'\<close>] simp del: map_map)
done
ultimately show "?P ?rs' \<and> ?Q ?rs'" ..
qed
next
case (7 nm'' k r' ts'')
show ?thesis (is "\<exists>rs'. ?P rs'")
proof
show "?P(map dterm (ts''[k := r']))"
using 7 ob
apply clarsimp
apply(rule "1.hyps"[OF _ ob(3)])
using 7 "1.prems" ob apply auto
done
qed
next
case (9 v k r' ts'')
then obtain vs where [simp]: "v = C\<^sub>U nm' vs" "rs' = map dterm\<^sub>M\<^sub>L (rev vs) @ map dterm ts''"
using ob by(cases v) auto
show ?thesis (is "\<exists>rs'. ?P rs' \<and> ?Q rs'")
proof
let ?rs' = "map dterm ((map term (rev vs) @ ts'')[k+size vs:=r'])"
have "?P ?rs'" using ob 9 by (auto simp: list_update_append)
moreover have "?Q ?rs'"
apply rule
apply(rule "1.hyps"[OF _ ob(3)])
using 9 "1.prems" ob by (auto simp:nth_append simp del: map_map)
ultimately show "?P ?rs' \<and> ?Q ?rs'" ..
qed
qed (insert ob, auto simp del: map_map)
}
hence "\<exists>rs'. dterm (ts[i := t'] ! j) = C nm' \<bullet>\<bullet> rs' \<and> (nm = nm' \<longrightarrow> no_match rs rs')"
using \<open>i < size ts\<close> ob by(simp add:nth_list_update)
hence "?P j" using ob by auto
moreover have "j < ?m" using \<open>j < length ts\<close> \<open>j < size ps\<close> by simp
ultimately show ?thesis by blast
qed
qed
qed
declare [[simp_depth_limit = 50]]
lemma Red_term_pres_no_match_it:
"\<lbrakk> \<forall> i < length ts. (ts ! i, ts' ! i) : Red_term ^^ (ns!i);
size ts' = size ts; size ns = size ts;
no_match ps (map dterm ts)\<rbrakk>
\<Longrightarrow> no_match ps (map dterm ts')"
proof(induct "sum_list ns" arbitrary: ts ns)
case 0
hence "\<forall>i < size ts. ns!i = 0" by simp
with 0 show ?case by simp (metis nth_equalityI)
next
case (Suc n)
then have "sum_list ns \<noteq> 0" by arith
then obtain k l where "k<size ts" and [simp]: "ns!k = Suc l"
by simp (metis \<open>length ns = length ts\<close> gr0_implies_Suc in_set_conv_nth)
let ?ns = "ns[k := l]"
have "n = sum_list ?ns" using \<open>Suc n = sum_list ns\<close> \<open>k<size ts\<close> \<open>size ns = size ts\<close>
by (simp add:sum_list_update)
obtain t' where "ts!k \<Rightarrow> t'" "(t', ts'!k) : Red_term^^l"
using Suc(3) \<open>k<size ts\<close> \<open>size ns = size ts\<close> \<open>ns!k = Suc l\<close>
by (metis relpow_Suc_E2)
then have 1: "\<forall>i<size(ts[k:=t']). (ts[k:=t']!i, ts'!i) : Red_term^^(?ns!i)"
using Suc(3) \<open>k<size ts\<close> \<open>size ns = size ts\<close>
by (auto simp add:nth_list_update)
note nm1 = Red_term_pres_no_match[OF \<open>k<size ts\<close> \<open>ts!k \<Rightarrow> t'\<close> \<open>no_match ps (map dterm ts)\<close>]
show ?case by(rule Suc(1)[OF \<open>n = sum_list ?ns\<close> 1 _ _ nm1])
(simp_all add: \<open>size ts' = size ts\<close> \<open>size ns = size ts\<close>)
qed
lemma Red_term_pres_no_match_star:
assumes "\<forall>i < length(ts::tm list). ts ! i \<Rightarrow>* ts' ! i" and "size ts' = size ts"
and "no_match ps (map dterm ts)"
shows "no_match ps (map dterm ts')"
proof-
let ?P = "%ns. size ns = size ts \<and>
(\<forall>i < length ts.(ts!i, ts'!i) : Red_term^^(ns!i))"
have "\<exists>ns. ?P ns" using assms(1)
by(subst Skolem_list_nth[symmetric])
(simp add:rtrancl_power)
from someI_ex[OF this] show ?thesis
by(fast intro: Red_term_pres_no_match_it[OF _ assms(2) _ assms(3)])
qed
lemma not_pure_term[simp]: "\<not> pure(term v)"
proof
assume "pure(term v)" thus False
by cases
qed
abbreviation RedMLs :: "tm list \<Rightarrow> tm list \<Rightarrow> bool" (infix "[\<Rightarrow>*]" 50) where
"ss [\<Rightarrow>*] ts \<equiv> size ss = size ts \<and> (\<forall>i<size ss. ss!i \<Rightarrow>* ts!i)"
fun C_U_args :: "tm \<Rightarrow> tm list" ("C\<^sub>U'_args") where
"C\<^sub>U_args(s \<bullet> t) = C\<^sub>U_args s @ [t]" |
"C\<^sub>U_args(term(C\<^sub>U nm vs)) = map term (rev vs)" |
"C\<^sub>U_args _ = []"
lemma [simp]: "C\<^sub>U_args(C nm \<bullet>\<bullet> ts) = ts"
by (induct ts rule:rev_induct) auto
lemma redts_term_cong: "v \<Rightarrow>* v' \<Longrightarrow> term v \<Rightarrow>* term v'"
apply(erule converse_rtrancl_induct)
apply(rule rtrancl_refl)
apply(fast intro: converse_rtrancl_into_rtrancl dest: ctxt_term)
done
lemma C_Red_term_ML:
"v \<Rightarrow> v' \<Longrightarrow> C_normal\<^sub>M\<^sub>L v \<Longrightarrow> dterm\<^sub>M\<^sub>L v = C nm \<bullet>\<bullet> ts
\<Longrightarrow> dterm\<^sub>M\<^sub>L v' = C nm \<bullet>\<bullet> map dterm (C\<^sub>U_args(term v')) \<and>
C\<^sub>U_args(term v) [\<Rightarrow>*] C\<^sub>U_args(term v') \<and>
ts = map dterm (C\<^sub>U_args(term v))" and
"(vs:: ml list) \<Rightarrow> vs' \<Longrightarrow> i < length vs \<Longrightarrow> vs ! i \<Rightarrow>* vs' ! i"
apply(induct arbitrary: nm ts and i rule:Red_ml_Red_ml_list.inducts)
apply(simp_all add:Red_ml_list_length del: map_map)
apply(frule Red_ml_list_length)
apply(simp add: redts_term_cong rev_nth del: map_map)
apply(simp add:nth_Cons' r_into_rtrancl del: map_map)
apply(simp add:nth_Cons')
done
lemma C_normal_subterm:
"C_normal t \<Longrightarrow> dterm t = C nm \<bullet>\<bullet> ts \<Longrightarrow> s \<in> set(C\<^sub>U_args t) \<Longrightarrow> C_normal s"
apply(induct rule: C_normal.induct)
apply auto
apply(case_tac v)
apply auto
done
lemma C_normal_subterms:
"C_normal t \<Longrightarrow> dterm t = C nm \<bullet>\<bullet> ts \<Longrightarrow> ts = map dterm (C\<^sub>U_args t)"
apply(induct rule: C_normal.induct)
apply auto
apply(case_tac v)
apply auto
done
lemma C_redt: "t \<Rightarrow> t' \<Longrightarrow> C_normal t \<Longrightarrow>
C_normal t' \<and> (dterm t = C nm \<bullet>\<bullet> ts \<longrightarrow>
(\<exists>ts'. ts' = map dterm (C\<^sub>U_args t') \<and> dterm t' = C nm \<bullet>\<bullet> ts' \<and>
C\<^sub>U_args t [\<Rightarrow>*] C\<^sub>U_args t'))"
apply(induct arbitrary: ts nm rule:Red_term_hnf_induct)
apply (simp_all del: map_map)
apply (metis no_match_R_coincide rev_rev_ident)
- apply clarsimp
apply rule
apply (metis C_normal_ML_inv)
apply clarify
apply(drule (2) C_Red_term_ML)
apply clarsimp
apply clarsimp
apply (metis insert_iff subsetD set_update_subset_insert)
apply clarsimp
apply(rule)
apply (metis insert_iff subsetD set_update_subset_insert)
apply rule
apply clarify
apply(drule bspec, assumption)
apply simp
apply(subst no_match.simps)
apply(subst (asm) no_match.simps)
apply clarsimp
apply(rename_tac j nm nm' rs rs')
apply(rule_tac x=j in exI)
apply simp
apply(case_tac "i=j")
apply(erule_tac x=rs' in meta_allE)
apply(erule_tac x=nm' in meta_allE)
apply (clarsimp simp: all_set_conv_all_nth)
apply(metis C_normal_subterms Red_term_pres_no_match_star)
apply (auto simp:nth_list_update)
done
lemma C_redts: "t \<Rightarrow>* t' \<Longrightarrow> C_normal t \<Longrightarrow>
C_normal t' \<and> (dterm t = C nm \<bullet>\<bullet> ts \<longrightarrow>
(\<exists>ts'. dterm t' = C nm \<bullet>\<bullet> ts' \<and> C\<^sub>U_args t [\<Rightarrow>*] C\<^sub>U_args t' \<and>
ts' = map dterm (C\<^sub>U_args t')))"
apply(induct arbitrary: nm ts rule:converse_rtrancl_induct)
apply simp
using tm_vector_cases[of t']
apply(elim disjE)
apply clarsimp
apply clarsimp
apply clarsimp
apply clarsimp
apply(case_tac v)
apply simp
apply simp
apply simp
apply simp
apply clarsimp
apply simp
apply simp
apply simp
apply simp
apply(frule_tac nm=nm and ts="ts" in C_redt)
apply assumption
apply clarify
apply rule
apply metis
apply clarify
apply simp
apply rule
apply (metis rtrancl_trans)
done
lemma no_match_preserved:
"\<forall>t\<in>set ts. C_normal t \<Longrightarrow> ts [\<Rightarrow>*] ts'
\<Longrightarrow> no_match ps os \<Longrightarrow> os = map dterm ts \<Longrightarrow> no_match ps (map dterm ts')"
proof(induct ps os arbitrary: ts ts' rule: no_match.induct)
case (1 ps os)
obtain i nm nm' ps' os' where a: "ps!i = C nm \<bullet>\<bullet> ps'" "i < size ps"
"i < size os" "os!i = C nm' \<bullet>\<bullet> os'" "nm=nm' \<longrightarrow> no_match ps' os'"
using 1(4) no_match.simps[of ps os] by fastforce
note 1(5)[simp]
have "C_normal (ts ! i)" using 1(2) \<open>i < size os\<close> by auto
have "ts!i \<Rightarrow>* ts'!i" using 1(3) \<open>i < size os\<close> by auto
have "dterm (ts ! i) = C nm' \<bullet>\<bullet> os'" using \<open>os!i = C nm' \<bullet>\<bullet> os'\<close> \<open>i < size os\<close>
by (simp add:nth_map)
with C_redts [OF \<open>ts!i \<Rightarrow>* ts'!i\<close> \<open>C_normal (ts!i)\<close>]
C_normal_subterm[OF \<open>C_normal (ts!i)\<close>]
C_normal_subterms[OF \<open>C_normal (ts!i)\<close>]
obtain ss' rs rs' :: "tm list" where b: "\<forall>t\<in>set rs. C_normal t"
"dterm (ts' ! i) = C nm' \<bullet>\<bullet> ss'" "length rs = length rs'"
"\<forall>i<length rs. rs ! i \<Rightarrow>* rs' ! i" "ss' = map dterm rs'" "os' = map dterm rs"
by fastforce
show ?case
apply(subst no_match.simps)
apply(rule_tac x=i in exI)
using 1(2-5) a b
apply clarsimp
apply(rule 1(1)[of i nm' _ nm' "map dterm rs" rs])
apply simp_all
done
qed
lemma Lam_Red_term_itE:
"(\<Lambda> t, t') : Red_term^^i \<Longrightarrow> \<exists>t''. t' = \<Lambda> t'' \<and> (t,t'') : Red_term^^i"
apply(induct i arbitrary: t')apply simp
apply(erule relpow_Suc_E)
apply(erule Red_term.cases)
apply (simp_all)
apply blast+
done
lemma Red_term_it: "(V x \<bullet>\<bullet> rs, r) : Red_term^^i
\<Longrightarrow> \<exists>ts is. r = V x \<bullet>\<bullet> ts \<and> size ts = size rs & size is = size rs \<and>
(\<forall>j<size ts. (rs!j, ts!j) : Red_term^^(is!j) \<and> is!j <= i)"
proof(induct i arbitrary:rs)
case 0
moreover
have "\<exists>is. length is = length rs \<and>
(\<forall>j<size rs. (rs!j, rs!j) \<in> Red_term ^^ is!j \<and> is!j = 0)" (is "\<exists>is. ?P is")
proof
show "?P(replicate (size rs) 0)" by simp
qed
ultimately show ?case by auto
next
case (Suc i rs)
from \<open>(V x \<bullet>\<bullet> rs, r) \<in> Red_term ^^ Suc i\<close>
obtain r' where r': "V x \<bullet>\<bullet> rs \<Rightarrow> r'" and "(r',r) \<in> Red_term ^^ i"
by (metis relpow_Suc_D2)
from r' have "\<exists>k<size rs. \<exists>s. rs!k \<Rightarrow> s \<and> r' = V x \<bullet>\<bullet> rs[k:=s]"
proof(induct rs arbitrary: r' rule:rev_induct)
case Nil thus ?case by(fastforce elim: Red_term.cases)
next
case (snoc r rs)
hence "(V x \<bullet>\<bullet> rs) \<bullet> r \<Rightarrow> r'" by simp
thus ?case
proof(cases rule:Red_term.cases)
case (ctxt_At1 s')
then obtain k s'' where aux: "k<length rs" "rs ! k \<Rightarrow> s''" "s' = V x \<bullet>\<bullet> rs[k := s'']"
using snoc(1) by force
show ?thesis (is "\<exists>k < ?n. \<exists>s. ?P k s")
proof-
have "k<?n \<and> ?P k s''" using ctxt_At1 aux
by (simp add:nth_append) (metis last_snoc butlast_snoc list_update_append1)
thus ?thesis by blast
qed
next
case (ctxt_At2 t')
show ?thesis (is "\<exists>k < ?n. \<exists>s. ?P k s")
proof-
have "size rs<?n \<and> ?P (size rs) t'" using ctxt_At2 by simp
thus ?thesis by blast
qed
qed
qed
then obtain k s where "k<size rs" "rs!k \<Rightarrow> s" and [simp]: "r' = V x \<bullet>\<bullet> rs[k:=s]" by metis
from Suc(1)[of "rs[k:=s]"] \<open>(r',r) \<in> Red_term ^^ i\<close>
show ?case using \<open>k<size rs\<close> \<open>rs!k \<Rightarrow> s\<close>
apply auto
apply(rule_tac x="is[k := Suc(is!k)]" in exI)
apply (auto simp:nth_list_update)
apply(erule_tac x=k in allE)
apply auto
apply (metis relpow_Suc_I2 relpow.simps(2))
done
qed
lemma C_Red_term_it: "(C nm \<bullet>\<bullet> rs, r) : Red_term^^i
\<Longrightarrow> \<exists>ts is. r = C nm \<bullet>\<bullet> ts \<and> size ts = size rs \<and> size is = size rs \<and>
(\<forall>j<size ts. (rs!j, ts!j) \<in> Red_term^^(is!j) \<and> is!j \<le> i)"
proof(induct i arbitrary:rs)
case 0
moreover
have "\<exists>is. length is = length rs \<and>
(\<forall>j<size rs. (rs!j, rs!j) \<in> Red_term ^^ is!j \<and> is!j = 0)" (is "\<exists>is. ?P is")
proof
show "?P(replicate (size rs) 0)" by simp
qed
ultimately show ?case by auto
next
case (Suc i rs)
from \<open>(C nm \<bullet>\<bullet> rs, r) \<in> Red_term ^^ Suc i\<close>
obtain r' where r': "C nm \<bullet>\<bullet> rs \<Rightarrow> r'" and "(r',r) \<in> Red_term ^^ i"
by (metis relpow_Suc_D2)
from r' have "\<exists>k<size rs. \<exists>s. rs!k \<Rightarrow> s \<and> r' = C nm \<bullet>\<bullet> rs[k:=s]"
proof(induct rs arbitrary: r' rule:rev_induct)
case Nil thus ?case by(fastforce elim: Red_term.cases)
next
case (snoc r rs)
hence "(C nm \<bullet>\<bullet> rs) \<bullet> r \<Rightarrow> r'" by simp
thus ?case
proof(cases rule:Red_term.cases)
case (ctxt_At1 s')
then obtain k s'' where aux: "k<length rs" "rs ! k \<Rightarrow> s''" "s' = C nm \<bullet>\<bullet> rs[k := s'']"
using snoc(1) by force
show ?thesis (is "\<exists>k < ?n. \<exists>s. ?P k s")
proof-
have "k<?n \<and> ?P k s''" using ctxt_At1 aux
by (simp add:nth_append) (metis last_snoc butlast_snoc list_update_append1)
thus ?thesis by blast
qed
next
case (ctxt_At2 t')
show ?thesis (is "\<exists>k < ?n. \<exists>s. ?P k s")
proof-
have "size rs<?n \<and> ?P (size rs) t'" using ctxt_At2 by simp
thus ?thesis by blast
qed
qed
qed
then obtain k s where "k<size rs" "rs!k \<Rightarrow> s" and [simp]: "r' = C nm \<bullet>\<bullet> rs[k:=s]" by metis
from Suc(1)[of "rs[k:=s]"] \<open>(r',r) \<in> Red_term ^^ i\<close>
show ?case using \<open>k<size rs\<close> \<open>rs!k \<Rightarrow> s\<close>
apply auto
apply(rule_tac x="is[k := Suc(is!k)]" in exI)
apply (auto simp:nth_list_update)
apply(erule_tac x=k in allE)
apply auto
apply (metis relpow_Suc_I2 relpow.simps(2))
done
qed
lemma pure_At[simp]: "pure(s \<bullet> t) \<longleftrightarrow> pure s \<and> pure t"
by(fastforce elim: pure.cases)
lemma pure_foldl_At[simp]: "pure(s \<bullet>\<bullet> ts) \<longleftrightarrow> pure s \<and> (\<forall>t\<in>set ts. pure t)"
by(induct ts arbitrary: s) auto
lemma nbe_C_normal_ML:
assumes "term v \<Rightarrow>* t'" "C_normal\<^sub>M\<^sub>L v" "pure t'" shows "normal t'"
proof -
{ fix t t' i v
assume "(t,t') : Red_term^^i"
hence "t = term v \<Longrightarrow> C_normal\<^sub>M\<^sub>L v \<Longrightarrow> pure t' \<Longrightarrow> normal t'"
proof(induct i arbitrary: t t' v rule:less_induct)
case (less k)
show ?case
proof (cases k)
case 0 thus ?thesis using less by auto
next
case (Suc i)
then obtain i' s where "t \<Rightarrow> s" and red: "(s,t') : Red_term^^i'" and [arith]: "i' <= i"
by (metis eq_imp_le less(5) Suc relpow_Suc_D2)
hence "term v \<Rightarrow> s" using Suc less by simp
thus ?thesis
proof cases
case (term_C nm vs)
with less have 0:"no_match_compR nm vs" by auto
let ?n = "size vs"
have 1: "(C nm \<bullet>\<bullet> map term (rev vs),t') : Red_term^^i'"
using term_C \<open>(s,t') : Red_term^^i'\<close> by simp
with C_Red_term_it[OF 1]
obtain ts ks where [simp]: "t' = C nm \<bullet>\<bullet> ts"
and sz: "size ts = ?n \<and> size ks = ?n \<and>
(\<forall>i<?n. (term((rev vs)!i), ts!i) : Red_term^^(ks!i) \<and> ks ! i \<le> i')"
by(auto cong:conj_cong)
have pure_ts: "\<forall>t\<in>set ts. pure t" using \<open>pure t'\<close> by simp
{ fix i assume "i<size vs"
moreover hence "(term((rev vs)!i), ts!i) : Red_term^^(ks!i)" by(metis sz)
ultimately have "normal (ts!i)"
apply -
apply(rule less(1))
prefer 5 apply assumption
using sz Suc apply fastforce
apply(rule refl)
using less term_C
apply(auto)
apply (metis in_set_conv_nth length_rev set_rev)
apply (metis in_set_conv_nth pure_ts sz)
done
} note 2 = this
have 3: "no_match_R nm (map dterm (map term (rev vs)))"
apply(subst map_dterm_term)
apply(rule no_match_R_coincide) using 0 by simp
have 4: "map term (rev vs) [\<Rightarrow>*] ts"
proof -
have "(C nm \<bullet>\<bullet> map term (rev vs),t'): Red_term^^i'"
using red term_C by auto
from C_Red_term_it[OF this] obtain ts' "is" where "t' = C nm \<bullet>\<bullet> ts'"
and "length ts' = ?n \<and> length is =?n \<and>
(\<forall>j< ?n. (map term (rev vs) ! j, ts' ! j) \<in> Red_term ^^ is ! j \<and> is ! j \<le> i')"
using sz by auto
from \<open>t' = C nm \<bullet>\<bullet> ts'\<close> \<open>t' = C nm \<bullet>\<bullet> ts\<close> have "ts = ts'" by simp
show ?thesis using sz by (auto simp: rtrancl_is_UN_relpow)
qed
have 5: "\<forall>t\<in>set(map term vs). C_normal t"
using less term_C by auto
have "no_match_R nm (map dterm ts)"
apply auto
apply(subgoal_tac "no_match aa (map dterm (map term (rev vs)))")
prefer 2
using 3 apply blast
using 4 5 no_match_preserved[OF _ _ _ refl, of "map term (rev vs)" "ts"] by simp
hence 6: "no_match_R nm ts" by(metis map_dterm_pure[OF pure_ts])
then show "normal t'"
apply(simp)
apply(rule normal.intros(3))
using 2 sz apply(fastforce simp:set_conv_nth)
apply auto
apply(subgoal_tac "no_match aa (take (size aa) ts)")
apply (metis no_match)
apply(fastforce intro:no_match_take)
done
next
case (term_V x vs)
let ?n = "size vs"
have 1: "(V x \<bullet>\<bullet> map term (rev vs),t') : Red_term^^i'"
using term_V \<open>(s,t') : Red_term^^i'\<close> by simp
with Red_term_it[OF 1] obtain ts "is" where [simp]: "t' = V x \<bullet>\<bullet> ts"
and 2: "length ts = ?n \<and>
length is = ?n \<and> (\<forall>j<?n. (term (rev vs ! j), ts ! j) \<in> Red_term ^^ is ! j \<and>
is ! j \<le> i')"
by (auto cong:conj_cong)
have "\<forall>j<?n. normal(ts!j)"
proof(clarify)
fix j assume 0: "j < ?n"
then have "is!j < k" using \<open>k=Suc i\<close> 2 by auto
have red: "(term (rev vs ! j), ts ! j) \<in> Red_term ^^ is ! j" using \<open>j < ?n\<close> 2 by auto
have pure: "pure (ts ! j)" using \<open>pure t'\<close> 0 2 by auto
have Cnm: "C_normal\<^sub>M\<^sub>L (rev vs ! j)" using less term_V
by simp (metis 0 in_set_conv_nth length_rev set_rev)
from less(1)[OF \<open>is!j < k\<close> refl Cnm pure red] show "normal(ts!j)" .
qed
note 3=this
show ?thesis by simp (metis normal.intros(1) in_set_conv_nth 2 3)
next
case (term_Clo f vs n)
let ?u = "apply (lift 0 (Clo f vs n)) (V\<^sub>U 0 [])"
from term_Clo \<open>(s,t') : Red_term^^i'\<close>
obtain t'' where [simp]: "t' = \<Lambda> t''" and 1: "(term ?u, t'') : Red_term^^i'"
by(metis Lam_Red_term_itE)
have "i' < k" using \<open>k = Suc i\<close> by arith
have "pure t''" using \<open>pure t'\<close> by simp
have "C_normal\<^sub>M\<^sub>L ?u" using less term_Clo by(simp)
from less(1)[OF \<open>i' < k\<close> refl \<open>C_normal\<^sub>M\<^sub>L ?u\<close> \<open>pure t''\<close> 1]
show ?thesis by(simp add:normal.intros)
next
case (ctxt_term u')
have "i' < k" using \<open>k = Suc i\<close> by arith
have "C_normal\<^sub>M\<^sub>L u'" by (rule C_normal_ML_inv) (insert less ctxt_term, simp_all)
have "(term u', t') \<in> Red_term ^^ i'" using red ctxt_term by auto
from less(1)[OF \<open>i' < k\<close> refl \<open>C_normal\<^sub>M\<^sub>L u'\<close> \<open>pure t'\<close> this] show ?thesis .
qed
qed
qed
}
thus ?thesis using assms(2-) rtrancl_imp_relpow[OF assms(1)] by blast
qed
lemma C_normal_ML_compile:
"pure t \<Longrightarrow> \<forall>i. C_normal\<^sub>M\<^sub>L(\<sigma> i) \<Longrightarrow> C_normal\<^sub>M\<^sub>L (compile t \<sigma>)"
by(induct t arbitrary: \<sigma>) (simp_all add: C_normal_ML_lift_ML)
corollary nbe_normal:
"pure t \<Longrightarrow> term(comp_fixed t) \<Rightarrow>* t' \<Longrightarrow> pure t' \<Longrightarrow> normal t'"
apply(erule nbe_C_normal_ML)
apply(simp add: C_normal_ML_compile)
apply assumption
done
section\<open>Refinements\<close>
text\<open>We ensure that all occurrences of @{term "C\<^sub>U nm vs"} satisfy
the invariant @{prop"size vs = arity nm"}.\<close>
text\<open>A constructor value:\<close>
fun C\<^sub>Us :: "ml \<Rightarrow> bool" where
"C\<^sub>Us(C\<^sub>U nm vs) = (size vs = arity nm \<and> (\<forall>v\<in>set vs. C\<^sub>Us v))" |
"C\<^sub>Us _ = False"
lemma size_foldl_At: "size(C nm \<bullet>\<bullet> ts) = size ts + sum_list(map size ts)"
by(induct ts rule:rev_induct) auto
lemma termination_linpats:
"i < length ts \<Longrightarrow> ts!i = C nm \<bullet>\<bullet> ts'
\<Longrightarrow> length ts' + sum_list (map size ts') < length ts + sum_list (map size ts)"
apply(subgoal_tac "C nm \<bullet>\<bullet> ts' : set ts")
prefer 2 apply (metis in_set_conv_nth)
apply(drule sum_list_map_remove1[of _ _ size])
apply(simp add:size_foldl_At)
apply (metis gr_implies_not0 length_0_conv)
done
text\<open>Linear patterns:\<close>
function linpats :: "tm list \<Rightarrow> bool" where
"linpats ts \<longleftrightarrow>
(\<forall>i<size ts. (\<exists>x. ts!i = V x) \<or>
(\<exists>nm ts'. ts!i = C nm \<bullet>\<bullet> ts' \<and> arity nm = size ts' \<and> linpats ts')) \<and>
(\<forall>i<size ts. \<forall>j<size ts. i\<noteq>j \<longrightarrow> fv(ts!i) \<inter> fv(ts!j) = {})"
by pat_completeness auto
termination
apply(relation "measure(%ts. size ts + (SUM t<-ts. size t))")
apply (auto simp:termination_linpats)
done
declare linpats.simps[simp del]
(* FIXME move *)
lemma eq_lists_iff_eq_nth:
"size xs = size ys \<Longrightarrow> (xs=ys) = (\<forall>i<size xs. xs!i = ys!i)"
by (metis nth_equalityI)
lemma pattern_subst_ML_coincidence:
"pattern t \<Longrightarrow> \<forall>i\<in>fv t. \<sigma> i = \<sigma>' i
\<Longrightarrow> subst_ML \<sigma> (comp_pat t) = subst_ML \<sigma>' (comp_pat t)"
by(induct pred:pattern) auto
lemma linpats_pattern: "linpats ts \<Longrightarrow> patterns ts"
proof(induct ts rule:linpats.induct)
case (1 ts)
show ?case
proof
fix t assume "t : set ts"
then obtain i where "i < size ts" and [simp]: "t = ts!i"
by (auto simp: in_set_conv_nth)
hence "(\<exists>x. t = V x) \<or> (\<exists>nm ts'. t = C nm \<bullet>\<bullet> ts' \<and> arity nm = size ts' & linpats ts')"
(is "?V | ?C")
using 1(2) by(simp add:linpats.simps[of ts])
thus "pattern t"
proof
assume "?V" thus ?thesis by(auto simp:pat_V)
next
assume "?C" thus ?thesis using 1(1) \<open>i < size ts\<close>
by auto (metis pat_C)
qed
qed
qed
lemma no_match_ML_swap_rev:
"length ps = length vs \<Longrightarrow> no_match\<^sub>M\<^sub>L ps (rev vs) \<Longrightarrow> no_match\<^sub>M\<^sub>L (rev ps) vs"
apply(clarsimp simp: no_match_ML.simps[of ps] no_match_ML.simps[of _ vs])
apply(rule_tac x="size ps - i - 1" in exI)
apply (fastforce simp:rev_nth)
done
lemma no_match_ML_aux:
"\<forall>v \<in> set cvs. C\<^sub>Us v \<Longrightarrow> linpats ps \<Longrightarrow> size ps = size cvs \<Longrightarrow>
\<forall>\<sigma>. map (subst\<^sub>M\<^sub>L \<sigma>) (map comp_pat ps) \<noteq> cvs \<Longrightarrow>
no_match\<^sub>M\<^sub>L (map comp_pat ps) cvs"
apply(induct ps arbitrary: cvs rule:linpats.induct)
apply(frule linpats_pattern)
apply(subst (asm) linpats.simps) back
apply auto
apply(case_tac "\<forall>i<size ts. \<exists>\<sigma>. subst\<^sub>M\<^sub>L \<sigma> (comp_pat (ts!i)) = cvs!i")
apply(clarsimp simp:Skolem_list_nth)
apply(rename_tac "\<sigma>s")
apply(erule_tac x="%x. (\<sigma>s!(THE i. i<size ts & x : fv(ts!i)))x" in allE)
apply(clarsimp simp:eq_lists_iff_eq_nth)
apply(rotate_tac -3)
apply(erule_tac x=i in allE)
apply simp
apply(rotate_tac -1)
apply(drule sym)
apply simp
apply(erule contrapos_np)
apply(rule pattern_subst_ML_coincidence)
apply (metis in_set_conv_nth)
apply clarsimp
apply(rule_tac a=i in theI2)
apply simp
apply (metis disjoint_iff_not_equal)
apply (metis disjoint_iff_not_equal)
apply clarsimp
apply(subst no_match_ML.simps)
apply(rule_tac x="size ts - i - 1" in exI)
apply simp
apply rule
apply simp
apply(subgoal_tac "\<not>(\<exists>x. ts!i = V x)")
prefer 2
apply fastforce
apply(subgoal_tac "\<exists>nm ts'. ts!i = C nm \<bullet>\<bullet> ts' & size ts' = arity nm & linpats ts'")
prefer 2
apply fastforce
apply clarsimp
apply(rule_tac x=nm in exI)
apply(subgoal_tac "\<exists>nm' vs'. cvs!i = C\<^sub>U nm' vs' & size vs' = arity nm' & (\<forall>v' \<in> set vs'. C\<^sub>Us v')")
prefer 2
apply(drule_tac x="cvs!i" in bspec)
apply simp
apply(case_tac "cvs!i")
apply simp_all
apply (clarsimp simp:rev_nth rev_map[symmetric])
apply(erule_tac x=i in meta_allE)
apply(erule_tac x=nm' in meta_allE)
apply(erule_tac x="ts'" in meta_allE)
apply(erule_tac x="rev vs'" in meta_allE)
apply simp
apply(subgoal_tac "no_match\<^sub>M\<^sub>L (map comp_pat ts') (rev vs')")
apply(rule no_match_ML_swap_rev)
apply simp
apply assumption
apply(erule_tac meta_mp)
apply (metis rev_rev_ident)
done
(*<*)
end
(*>*)
diff --git a/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy b/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy
--- a/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy
+++ b/thys/Ordinary_Differential_Equations/Ex/Lorenz/Lorenz_Approximation.thy
@@ -1,3044 +1,3045 @@
section"Example: Lorenz attractor"
theory Lorenz_Approximation
imports
"HOL-ODE-Numerics.ODE_Numerics"
Result_File_Coarse
begin
text \<open>\label{sec:lorenz}\<close>
text \<open>TODO: move to isabelle? \<close>
lifting_update blinfun.lifting
lifting_forget blinfun.lifting
lemma eventually_uniformly_on:
"(\<forall>\<^sub>F x in uniformly_on T l. P x) = (\<exists>e>0. \<forall>f. (\<forall>x\<in>T. dist (f x) (l x) < e) \<longrightarrow> P f)"
unfolding uniformly_on_def
apply (subst eventually_INF)
apply safe
subgoal for E
apply (cases "E = {}")
subgoal by (auto intro!: exI[where x=1])
subgoal
apply (auto simp: INF_principal_finite eventually_principal elim!: )
proof goal_cases
case (1 x)
from 1 have "0 < Min E"
apply (subst Min_gr_iff)
apply force apply force apply force done
have *: "(\<Inter>e\<in>E. {f. \<forall>t\<in>T. dist (f t) (l t) < e}) = {f. \<forall>t\<in>T. dist (f t) (l t) < Min E}"
using 1 apply (auto simp: )
apply (subst Min_gr_iff)
apply force apply force apply force
apply (drule bspec, assumption)
apply (rule less_le_trans, assumption)
apply auto
done
from 1 have "\<forall>f. (\<forall>x\<in>T. dist (f x) (l x) < Min E) \<longrightarrow> P f" unfolding * by simp
then show ?case
using 1(4)[rule_format, OF \<open>0 < Min E\<close>] by auto
qed
done
subgoal for e
apply (rule exI[where x="{e}"])
by (auto simp: eventually_principal)
done
lemma op_cast_image_impl[autoref_rules]:
"(\<lambda>x. x, op_cast_image::'a::executable_euclidean_space set \<Rightarrow>
'b::executable_euclidean_space set)
\<in> aform.appr_rel \<rightarrow> aform.appr_rel"
if "DIM('a) = DIM('b)"
using that
apply (auto simp: aform.appr_rel_def intro!: relcompI)
unfolding lv_rel_def set_rel_br
by (force simp: intro!: brI dest!: brD)
lemma cast_bl_blinfun_of_list[simp]:
"cast_bl (blinfun_of_list xs::'a \<Rightarrow>\<^sub>L 'a) =
(blinfun_of_list xs::'b\<Rightarrow>\<^sub>L'b)"
if "DIM('a::executable_euclidean_space) = DIM('b::executable_euclidean_space)"
using that
apply (auto simp: cast_bl_rep intro!: blinfun_eqI)
by (auto simp: blinfun_of_list_def blinfun_of_matrix_apply
linear_sum linear.scaleR sum_Basis_sum_nth_Basis_list
linear_cast)
lemma cast_idem[simp]: "cast x = x"
by (auto simp: cast_def)
lemma cast_bl_idem[simp]: "cast_bl x = x"
by (auto simp: cast_bl_rep intro!: blinfun_eqI)
lemma op_cast_eucl1_image_impl[autoref_rules]:
"(\<lambda>x. x, op_cast_eucl1_image::'a::executable_euclidean_space c1_info set \<Rightarrow>
'b::executable_euclidean_space c1_info set)
\<in> aform.appr1_rel \<rightarrow> aform.appr1_rel"
if "DIM_precond TYPE('a) D"
"DIM_precond TYPE('b) D"
using that
proof (auto, goal_cases)
case (1 a b a')
then show ?case
apply (auto simp: aform.appr1_rel_br set_rel_br br_def)
subgoal for w x y z
apply (auto simp: aform.c1_info_of_appr_def cast_eucl1_def aform.c1_info_invar_def
split: option.splits)
apply (rule image_eqI)
apply (rule cast_eucl_of_list, force, force simp: Joints_imp_length_eq, force)
subgoal for s t
apply (rule image_eqI[where x="t"])
supply [simp del] = eucl_of_list_take_DIM
apply (auto simp: flow1_of_list_def)
apply (subst cast_eucl_of_list)
subgoal by simp
subgoal
by (auto dest!: Joints_imp_length_eq
simp: power2_eq_square flow1_of_list_def[abs_def])
subgoal by simp
done
done
subgoal for w x
apply (rule image_eqI[where x="cast_eucl1 (w, x)"])
apply (auto simp: aform.c1_info_of_appr_def cast_eucl1_def
aform.c1_info_invar_def
split: option.splits)
apply (rule image_eqI)
apply (rule cast_eucl_of_list, force, force simp: Joints_imp_length_eq, force)
subgoal for s t
apply (rule image_eqI[where x="t"])
supply [simp del] = eucl_of_list_take_DIM
apply (auto simp: flow1_of_list_def)
apply (subst cast_eucl_of_list)
subgoal by simp
subgoal
by (auto dest!: Joints_imp_length_eq
simp: power2_eq_square flow1_of_list_def[abs_def])
subgoal by simp
done
done
done
qed
lemma less_3_iff: "i < (3::nat) \<longleftrightarrow> i = 0 \<or> i = 1 \<or> i = 2"
by arith
definition mat3_of_vec::"R3 \<Rightarrow> real^3^3" where
"mat3_of_vec x = (let xs = list_of_eucl x in eucl_of_list [xs!0,0,0, xs!1,0,0, xs!2,0,0])"
lemma ll3: "{..<3} = {0,1,2::nat}"
by (auto simp: less_3_iff)
lemma mat3_of_vec: "cast (mat3_of_vec x *v eucl_of_list [1, 0, 0]) = x"
by (auto simp: mat3_of_vec_def eucl_of_list_matrix_vector_mult_eq_sum_nth_Basis_list
linear_sum linear_cast linear.scaleR ll3 linear_add Basis_list_R3 inner_prod_def
prod_eq_iff)
primrec bisect_form where
"bisect_form p f xs l u 0 = (l, u)"
| "bisect_form p f xs l u (Suc n) =
(let m = (l + u)/2 in
if approx_form_aform p (f m) xs
then bisect_form p f xs l m n
else bisect_form p f xs m u n)"
text \<open>This should prove that the expansion estimates are sufficient.\<close>
lemma expansion_main: "expansion_main (coarse_results) = Some True"
by eval
context includes floatarith_notation begin
definition "matrix_of_degrees2\<^sub>e =
(let
e = Var 2;
ur = Rad_of (Var 0);
vr = Rad_of (Var 1);
x1 = Cos ur;
x2 = Cos vr;
y1 = Sin ur;
y2 = Sin vr
in
[x1 + (e * (x2 - x1)), 0, 0,
y1 + (e * (y2 - y1)), 0, 0,
0, 0, 0])"
definition "matrix_of_degrees2 u v =
approx_floatariths 30 matrix_of_degrees2\<^sub>e (aforms_of_ivls [u, v, 0] [u, v, 1])"
text \<open>following \<open>vector_field.h\<close> / \<open>vector_field.cc\<close>\<close>
abbreviation "S \<equiv> 10::real"
abbreviation "B \<equiv> 8/3::real"
abbreviation "TEMP \<equiv> sqrt((S + 1) * (S + 1) + 4 * S * (28 - 1))"
abbreviation "K1 \<equiv> S / TEMP"
abbreviation "K2 \<equiv> (S - 1 + TEMP) / (2 * S)"
abbreviation "K3 \<equiv> (S - 1 - TEMP) / (2 * S)"
abbreviation "E1 \<equiv> (- (S + 1) + TEMP) / 2"
abbreviation "E2 \<equiv> (- (S + 1) - TEMP) / 2"
abbreviation "E3 \<equiv> - B"
abbreviation "C1 \<equiv> \<lambda>X. X ! 0 + X ! 1"
abbreviation "C2 \<equiv> \<lambda>X. K1 * C1 X * X ! 2"
schematic_goal
lorenz_fas:
"[E1 * X!0 - C2 X,
E2 * X!1 + C2 X,
E3 * X!2 + C1 X * (K2 * X!0 + K3 * X!1)] =
interpret_floatariths ?fas X"
by (reify_floatariths)
concrete_definition lorenz_fas uses lorenz_fas
end
interpretation lorenz: ode_interpretation true_form UNIV lorenz_fas "\<lambda>(X0, X1, X2).
(E1 * X0 - K1 * (X0 + X1) * X2,
E2 * X1 + K1 * (X0 + X1) * X2,
E3 * X2 + (X0 + X1) * (K2 * X0 + K3 * X1))::real*real*real"
"d::3" for d
by standard
(auto simp: lorenz_fas_def less_Suc_eq_0_disj nth_Basis_list_prod Basis_list_real_def
mk_ode_ops_def eucl_of_list_prod power2_eq_square inverse_eq_divide intro!: isFDERIV_I)
value [code] "length (slp_of_fas lorenz_fas)"
definition "mig_aform p x = mig_componentwise (Inf_aform' p x) (Sup_aform' p x)"
context includes floatarith_notation begin
definition "mig_aforms p x = real_of_float ((lower o the) ((approx p (Norm (map (Num o float_of o (mig_aform p)) x))) []))"
definition
"column_of_c1_info x N j = (map (\<lambda>i. the (snd x) ! i) (map (\<lambda>i. i * N + j) [0..<N]))"
definition "rotate_x_fa a =
[1, 0, 0,
0, Cos a, - Sin a,
0, Sin a, Cos a]"
definition "rotate_y_fa a =
[Cos a, 0, Sin a,
0, 1, 0,
- Sin a, 0, Cos a]"
definition "rotate_z_fa a =
[Cos a, - Sin a, 0,
Sin a, Cos a, 0,
0, 0, 1]"
definition "rotate_zx_slp a b xs =
slp_of_fas (mvmult_fa 3 3 (mmult_fa 3 3 3 (rotate_x_fa (Rad_of (R\<^sub>e b))) (rotate_z_fa (Rad_of (R\<^sub>e a)))) xs)"
definition "perspective_projection_aforms xs =
the (approx_slp_outer 30 3 (rotate_zx_slp (-30) (-60) (map Var [0..<3])) xs)"
definition "print_lorenz_aform print_fun cx cy cz ci
cd1 cd2
= (\<lambda>a b.
let (s1, n) = ((-6), False);
_ = print_fun (String.implode (''# gen(''@ show a@''): ''@ shows_aforms_hr (b) '''' @ ''\<newline>''));
_ = print_fun (String.implode (''# box(''@ show a@''): ''@ shows_box_of_aforms_hr (b) '''' @ ''\<newline>''));
((x0, y0, z0), (x1, y1, z1)) = case (map (Inf_aform' 30) (take 3 b), map (Sup_aform' 30) (take 3 b)) of
([x0, y0, z0], [x1, y1, z1]) \<Rightarrow> ((x0, y0, z0), (x1, y1, z1));
_ = print_fun (String.implode (shows_segments_of_aform 0 1 b ((shows cx o shows_space o shows z0 o shows_space o shows z1)'''') ''\<newline>''));
_ = print_fun (String.implode (shows_segments_of_aform 0 2 b ((shows cy o shows_space o shows y0 o shows_space o shows y1)'''') ''\<newline>''));
_ = print_fun (String.implode (shows_segments_of_aform 1 2 b ((shows cz o shows_space o shows x0 o shows_space o shows x1)'''') ''\<newline>''));
PS = perspective_projection_aforms b;
_ = print_fun (String.implode (shows_segments_of_aform 0 1 PS
((shows ci o shows_space o shows (fst (PS ! 2)) o shows_space o shows (fst (b ! 2))) '''') ''\<newline>''))
in if \<not> a \<and> length b > 10 then
print_fun (String.implode (shows_aforms_vareq 3 [(0, 1), (0, 2), (1, 2)] [0..<3]
cd1
cd2
(FloatR 1 s1 * (if n then
real_divl 30 1 (max (mig_aforms 30 (map (\<lambda>i. b ! i) [3,6,9]))
(mig_aforms 30 (map (\<lambda>i. b ! i) [4,7,10])))
else 1)) \<comment> \<open>always length \<open>2^s!\<close>\<close>
''# no C1 info'' b ''''))
else ())"
definition "print_lorenz_aform_std print_fun =
print_lorenz_aform print_fun ''0x000001'' ''0x000002'' ''0x000012'' ''0x000003''
[''0xa66f00'', ''0x06266f'', ''0xc60000'']
[''0xffaa00'', ''0x1240ab'', ''0xc60000'']"
definition "lorenz_optns print_funo =
(let
pf = the_default (\<lambda>_ _. ()) (map_option print_lorenz_aform_std print_funo);
tf = the_default (\<lambda>_ _. ()) (map_option (\<lambda>print_fun a b.
let
_ = print_fun (String.implode (''# '' @ a @ ''\<newline>''))
in case b of Some b \<Rightarrow>
(print_fun (String.implode (''# '' @ shows_box_of_aforms_hr (b) '''' @ ''\<newline>'')))
| None \<Rightarrow> ()) print_funo)
in
\<lparr>
precision = 30,
adaptive_atol = FloatR 1 (-30),
adaptive_rtol = FloatR 1 (-30),
method_id = 2,
start_stepsize = FloatR 1 (- 8),
iterations = 40,
halve_stepsizes = 10,
widening_mod = 40,
rk2_param = FloatR 1 0,
default_reduce = correct_girard 30 50 25,
printing_fun = pf,
tracing_fun = tf
\<rparr>)"
definition lorenz_optns'
where "lorenz_optns' pf m N rk2p a = lorenz_optns pf \<lparr>
default_reduce := correct_girard 30 m N,
rk2_param := rk2p,
adaptive_atol := a,
adaptive_rtol := a
\<rparr>"
definition mirror_irects
where "mirror_irects =
map (\<lambda>irect. case irect of [i, j, k] \<Rightarrow> [if j < 0 then - i else i , abs j, k] | irect \<Rightarrow> irect)"
definition "print_irects irects =
(let _ = map (\<lambda>is.
let _ = map (\<lambda>j.
let _ = print (String.implode (show j)) in print (STR '' '')) is in print (STR ''\<newline>'')) irects
in ())"
abbreviation "aforms_of_ivl \<equiv> \<lambda>x. aforms_of_ivls (fst x) (snd x)"
definition "conefield_propagation\<^sub>e =
([Deg_of (Arctan (Var (6) / Var (3))),
Deg_of (Arctan (Var (7) / Var (4))),
Min (Norm [Var(3), Var (6), Var (9)]) (Norm [Var(4), Var (7), Var (10)])])"
definition "conefield_propagation DX = approx_floatariths 30 conefield_propagation\<^sub>e DX"
definition "conefield_propagation_slp = slp_of_fas conefield_propagation\<^sub>e"
lemma [simp]: "length conefield_propagation_slp = 51"
by eval
definition op_with_unit_matrix :: "'a::real_normed_vector \<Rightarrow> 'a \<times> 'a \<Rightarrow>\<^sub>L 'a" where
"op_with_unit_matrix X = (X, 1\<^sub>L)"
context includes blinfun.lifting begin
lemma matrix_vector_mult_blinfun_works[simp]: "matrix e *v g = e g" for e::"(real^'n) \<Rightarrow>\<^sub>L (real^'m)"
by transfer (simp add: bounded_linear_def matrix_works)
end
lemma length_conefield_propagation\<^sub>e[simp]: "length conefield_propagation\<^sub>e = 3"
by (simp add: conefield_propagation\<^sub>e_def)
lemma interpret_floatariths_conefield_propagation:
"interpret_floatariths conefield_propagation\<^sub>e
(list_of_eucl (vec1_of_flow1 (xDX::(real^3) \<times> ((real^3)\<Rightarrow>\<^sub>L(real^3)))))
=
(let
DX = snd xDX;
DXu = DX (eucl_of_list [1, 0, 0]);
DXv = DX (eucl_of_list [0, 1, 0])
in
[deg_of (arctan (vec_nth DXu 1 / vec_nth DXu 0)),
deg_of (arctan (vec_nth DXv 1 / vec_nth DXv 0)),
min (norm DXu) (norm DXv)]
)"
apply (auto simp: conefield_propagation\<^sub>e_def Let_def interpret_mvmult_nth[where 'n=3 and 'm=3]
inverse_eq_divide vec1_of_flow1_def nth_append)
apply (auto simp: matrix_inner_Basis_list )
apply (auto simp: interpret_floatarith_norm[where 'a="real ^ 3"]
einterpret_mvmult_fa[where 'n=3 and 'm=3] matrix_inner_Basis_list nth_append)
by (auto simp: matrix_def axis_eq_eucl_of_list eucl_of_list_012)
definition "conefield_bounds_form l u =
(fold Conj [
Less (-90) (N\<^sub>r l),
LessEqual (N\<^sub>r l) (N\<^sub>r u),
LessEqual (Var 9) (0),
LessEqual 0 (Var 9),
Less (N\<^sub>r u) (90),
Less 0 (Var 3),
AtLeastAtMost (Var 6) (Tan (Rad_of (N\<^sub>r l)) * Var 3) (Tan (Rad_of (N\<^sub>r u)) * Var 3)] true_form)"
definition "contract_angles X i = (snd (bisect_form 30 (\<lambda>x. conefield_bounds_form x (89)) X 89 (-89) i),
snd (bisect_form 30 (conefield_bounds_form (-89)) X (-89) 89 i))"
definition "approx_conefield_bounds (DX::(R3 \<times> (R3 \<Rightarrow>\<^sub>L R3)) set) l u =
do {
let DX = (cast_eucl1 ` DX::3 eucl1 set);
DXo \<leftarrow> aform.vec1rep DX;
DX \<leftarrow> (case DXo of None \<Rightarrow> do {
let _ = aform.print_msg (''# approx_conefield_bounds failed DXo...'');
SUCCEED
}
| Some DX \<Rightarrow> RETURN DX);
let _ = aform.trace_set (''# approx_conefield_bounds DX: '') (Some DX);
approx_form_spec (conefield_bounds_form l u) (list_of_eucl ` DX)
}"
lemma [autoref_rules]:
includes autoref_syntax
shows "(conefield_bounds_form, conefield_bounds_form) \<in> Id \<rightarrow> Id \<rightarrow> Id "
by auto
lemma [autoref_rules_raw]:
"DIM_precond TYPE((real, 3) vec \<times> ((real, 3) vec, 3) vec) 12"
"DIM_precond TYPE(R3) 3"
"DIM_precond TYPE((real, 3) vec) 3"
by auto
schematic_goal approx_conefield_bounds_impl:
includes autoref_syntax
fixes optns::"real aform numeric_options"
assumes [autoref_rules]: "(DXi, DX) \<in> aform.appr1_rel"
assumes [autoref_rules]: "(li, l) \<in> Id"
assumes [autoref_rules]: "(ui, u) \<in> Id"
notes [autoref_rules] =
aform.print_msg_impl[where optns = optns]
aform.ivl_rep_of_set_autoref[where optns = optns]
aform.transfer_operations(12)[where optns = optns]
aform.approx_euclarithform[where optns=optns]
aform.trace_set_impl[of optns]
shows "(nres_of ?r, approx_conefield_bounds $ DX $ l $ u) \<in> \<langle>bool_rel\<rangle>nres_rel"
unfolding autoref_tag_defs
unfolding approx_conefield_bounds_def
including art
by autoref_monadic
concrete_definition approx_conefield_bounds_impl for optns li ui DXi uses
approx_conefield_bounds_impl
lemmas [autoref_rules] = approx_conefield_bounds_impl.refine
context includes autoref_syntax begin
lemma [autoref_rules]:
"(real_of_ereal, real_of_ereal) \<in> ereal_rel \<rightarrow> rnv_rel"
"(\<infinity>, \<infinity>) \<in> ereal_rel"
by auto
end
lemma interpret_form_true_form[simp]: "interpret_form true_form \<equiv> \<lambda>_. True"
by (force simp: true_form_def intro!: eq_reflection)
lemma interpret_form_conefield_bounds_form_list:
"interpret_form (conefield_bounds_form L U)
[x, y, z, ux, vx, wx,
uy, vy, wy,
uz, vz, wz]
\<longleftrightarrow>
(0 < ux \<and> -90 < L \<and> L \<le> U \<and> U < 90 \<and> uz = 0 \<and>
uy \<le> tan (rad_of U) * ux \<and>
tan (rad_of L) * ux \<le> uy)"
if "U \<in> float" "L \<in> float" "e \<in> float" "em \<in> float"
using that
by (auto simp: conefield_bounds_form_def L2_set_def)
lemma list_of_eucl_eucl1_3:
includes vec_syntax
shows "(list_of_eucl (vec1_of_flow1 (xDX::(real^3) \<times> ((real^3)\<Rightarrow>\<^sub>L(real^3))))) =
(let
(x, DX) = xDX;
DXu = DX (eucl_of_list [1, 0, 0]);
DXv = DX (eucl_of_list [0, 1, 0]);
DXw = DX (eucl_of_list [0, 0, 1])
in [x $ 0, x $ 1, x $ 2, vec_nth DXu 0, vec_nth DXv 0, vec_nth DXw 0,
vec_nth DXu 1, vec_nth DXv 1, vec_nth DXw 1,
vec_nth DXu 2, vec_nth DXv 2, vec_nth DXw 2])"
apply (auto simp: matrix_inner_Basis_list Let_def vec1_of_flow1_def
concat_map_map_index less_Suc_eq_0_disj
list_of_eucl_matrix eval_nat_numeral aform.inner_Basis_eq_vec_nth
intro!: nth_equalityI
split: prod.splits)
by (auto simp: matrix_def axis_eq_eucl_of_list eucl_of_list_012)
lemma interpret_form_conefield_bounds_form:
"interpret_form (conefield_bounds_form L U)
(list_of_eucl (vec1_of_flow1 (xDX::(real^3) \<times> ((real^3)\<Rightarrow>\<^sub>L(real^3)))))
=
(let
DX = snd xDX;
DXu = DX (eucl_of_list [1, 0, 0]);
DXv = DX (eucl_of_list [0, 1, 0]);
uz = vec_nth DXu 2;
uy = vec_nth DXu 1;
ux = vec_nth DXu 0;
vy = vec_nth DXv 1;
vx = vec_nth DXv 0
in
ux > 0 \<and> -90 < L \<and> L \<le> U \<and> U < 90 \<and> uz = 0 \<and>
(uy / ux) \<in> {tan (rad_of L) .. tan (rad_of U)}
)"
if "L \<in> float" "U \<in> float"
using that
unfolding list_of_eucl_eucl1_3 Let_def
by (auto split: prod.splits simp: interpret_form_conefield_bounds_form_list divide_simps)
lemma approx_conefield_bounds_cast:
"approx_conefield_bounds DX L U \<le> SPEC (\<lambda>b. b \<longrightarrow>
(\<forall>(x, dx) \<in> cast_eucl1 ` DX::3 eucl1 set.
let
u' = dx (eucl_of_list [1, 0, 0])
in
vec_nth u' 1 / vec_nth u' 0 \<in> {tan (rad_of L) .. tan (rad_of U)}
\<and> vec_nth u' 2 = 0 \<and> vec_nth u' 0 > 0 \<and> -90 < L \<and> L \<le> U \<and> U < 90)
)"
if "L \<in> float" "U \<in> float"
unfolding approx_conefield_bounds_def
apply refine_vcg
apply (auto simp: env_len_def )
subgoal for a b c
apply (drule bspec, assumption)
unfolding interpret_form_conefield_bounds_form[OF that]
by (auto simp: Let_def divide_simps)
done
lemma approx_conefield_bounds[le, refine_vcg]:
"approx_conefield_bounds DX l u \<le> SPEC (\<lambda>b. b \<longrightarrow>
(\<forall>(x, dx) \<in> DX::R3 c1_info set.
let
(u'1, u'2, u'3) = dx ((1, 0, 0))
in u'2 / u'1 \<in> {tan (rad_of l) .. tan (rad_of u)} \<and> u'3 = 0 \<and> u'1 > 0 \<and> -90 < l \<and> l \<le> u \<and> u < 90)
)"
if "l \<in> float" "u \<in> float"
apply (rule approx_conefield_bounds_cast[le, OF that])
apply (auto dest!: bspec simp: Let_def split: prod.splits)
by (auto simp: cast_eucl1_def cast_def)
schematic_goal MU\<^sub>e: "-E2 / E1 = interpret_floatarith ?fas []"
by (reify_floatariths)
concrete_definition MU\<^sub>e uses MU\<^sub>e
schematic_goal NU\<^sub>e: "-E3 / E1 = interpret_floatarith ?fas []"
by (reify_floatariths)
concrete_definition NU\<^sub>e uses NU\<^sub>e
definition "approx_ivls p fas xs = do {
let xs = ivls_of_aforms p xs;
res \<leftarrow> those (map (\<lambda>f. approx p f xs) fas);
Some (map (real_of_float o lower) res, map (real_of_float o upper) res)
}"
definition
"deform p t exit XDX = (case XDX of (lu, (X, DX)) \<Rightarrow>
let
d = ldec 0.1;
d = if exit then (1 - real_of_float (lb_sqrt 30 (1 - 2 * float_of (d)))) else d;
sd = real_of_float ((float_of (d*d)));
C0Deform = aforms_of_ivls [-sd,-sd,-sd] [sd, sd, sd];
result = msum_aforms' X C0Deform
in (lu,
case DX of None \<Rightarrow> (result, None)
| Some DX \<Rightarrow> let
C1_norm = 2 * d;
C1_norm = if exit then real_divr 30 C1_norm (1 - C1_norm) else C1_norm;
l = -C1_norm;
u = C1_norm;
D_M = aforms_of_ivls [1 + l,0,0, 0,1 + l,0, 0,0,1 + l] [1 + u,0,0, 0,1 + u,0, 0,0,1 + u];
(ri, ru) = the (approx_ivls p
(mmult_fa 3 3 3 (map Var [0..<9]) (map Var [9..<18])) (D_M @ DX));
Dresult = aforms_of_ivls ri ru;
resultDresult = product_aforms result Dresult
in (take 3 resultDresult, Some (drop 3 resultDresult))))"
definition "ivls_of_aforms' p r = (map (Inf_aform' p) r, map (Sup_aform' p) r)"
definition "compute_half_exit p t XDX = (case XDX of ((l, u::ereal), (X, DX)) \<Rightarrow>
let
\<comment> \<open>ASSERTING that \<open>Y\<close> straddles zero\<close>
(x0, y0, _) = case map (Inf_aform' p) X of [x,y,z] \<Rightarrow> (x, y, z);
(x1, y1, _) = case map (Sup_aform' p) X of [x,y,z] \<Rightarrow> (x, y, z);
splitting = x0 = 0 \<or> x1 = 0;
sign_x = if (x0 + x1) / 2 > 0 then 1 else -1;
mag_x = max (abs x0) (abs x1);
sign_x\<^sub>e = N\<^sub>r sign_x;
exit_rad\<^sub>e = N\<^sub>r (ldec 0.1);
X\<^sub>e = Var (0);
Y\<^sub>e = Var (1);
Z\<^sub>e = Var (2);
max_x_over_r\<^sub>e = N\<^sub>r mag_x / exit_rad\<^sub>e;
abs_x_over_r\<^sub>e = (Abs X\<^sub>e) / exit_rad\<^sub>e;
result = (if splitting
then let result = (the (approx_floatariths p [sign_x\<^sub>e * exit_rad\<^sub>e,
Y\<^sub>e * Powr (max_x_over_r\<^sub>e) MU\<^sub>e,
Z\<^sub>e * Powr (max_x_over_r\<^sub>e) NU\<^sub>e] X));
(ir, sr) = ivls_of_aforms' p result
in aforms_of_ivls (ir[1:=min 0 (ir!1), 2:=min 0 (ir!2)])
(sr[1:=max 0 (sr!1), 2:=max 0 (sr!2)])
else the (approx_floatariths p [sign_x\<^sub>e * exit_rad\<^sub>e,
Y\<^sub>e * Powr (abs_x_over_r\<^sub>e) MU\<^sub>e,
Z\<^sub>e * Powr (abs_x_over_r\<^sub>e) NU\<^sub>e] X));
_ = ()
in ((l::ereal, \<infinity>::ereal), (case DX of None \<Rightarrow> (result, None)
| Some DX \<Rightarrow>
let
ux\<^sub>e = Var (3);
uy\<^sub>e = Var (6);
P21\<^sub>e = if splitting
then (MU\<^sub>e / exit_rad\<^sub>e) * Y\<^sub>e * sign_x\<^sub>e * Powr (max_x_over_r\<^sub>e) (MU\<^sub>e - 1) \<comment> \<open>No need for \<open>Hull(0)\<close> because \<open>y\<close> straddles zero\<close>
else (MU\<^sub>e / exit_rad\<^sub>e) * Y\<^sub>e * sign_x\<^sub>e * Powr (abs_x_over_r\<^sub>e) (MU\<^sub>e - 1);
P22\<^sub>e = if splitting
then Powr (max_x_over_r\<^sub>e) MU\<^sub>e
else Powr (abs_x_over_r\<^sub>e) MU\<^sub>e;
P31\<^sub>e = if splitting
then sign_x\<^sub>e * (NU\<^sub>e / exit_rad\<^sub>e) * Z\<^sub>e * Powr (max_x_over_r\<^sub>e) (NU\<^sub>e - 1) \<comment> \<open>No need for \<open>Hull(\<infinity>)\<close> because scaling afterwards\<close>
else sign_x\<^sub>e * (NU\<^sub>e / exit_rad\<^sub>e) * Z\<^sub>e * Powr (abs_x_over_r\<^sub>e) (NU\<^sub>e - 1);
ry = (P21\<^sub>e * ux\<^sub>e) + (P22\<^sub>e * uy\<^sub>e);
rz = P31\<^sub>e * ux\<^sub>e;
(iDr, sDr) = the (approx_ivls p ([0, 0, 0,
ry, 0, 0,
rz, 0, 0]) (X @ DX));
Dresult = aforms_of_ivls (iDr[3:=min 0 (iDr!3)])
(sDr[3:=max 0 (sDr!3)]);
resultDresult = product_aforms result Dresult
in (take 3 resultDresult, Some (drop 3 resultDresult))
)))"
fun list3 where "list3 [a,b,c] = (a, b, c)"
definition
"split_x n x0 y0 z0 x1 y1 z1 =
(let
elem = (\<lambda>(x0, x1). aforms_of_ivls [x0, y0, z0] [x1, y1, z1]);
coord = (\<lambda>x0 n i. i * x0 * FloatR 1 (-int n));
us = map (coord x0 n) (rev [0..<(2^n)]) @ map (coord x1 n) [Suc 0..<Suc (2^n)];
ls = map (coord x0 n) (rev [Suc 0..<Suc (2^n)]) @ map (coord x1 n) [0..<(2^n)]
in map elem (zip ls us))"
definition "compute_cube_exit p t XDX =
(let
((l, u), (X', DX')) = deform p t False XDX;
((x0, y0, z0), (x1, y1, z1)) = pairself list3 (ivls_of_aforms' p X');
X's = [aforms_of_ivls [x0, y0, z0] [0, y1, z1], aforms_of_ivls [0, y0, z0] [x1, y1, z1]];
XDX's = map (\<lambda>X'. ((l, u), (X', DX'))) X's;
Xes = map (compute_half_exit p t) XDX's;
Xlumpies = map (deform p t True) Xes
in
Xlumpies)"
definition "cube_enteri = (map ldec [-0.1, -0.00015, 0.1, 0.8,0,0, 0.0005,0,0, 0,0,0],
map udec [ 0.1, 0.00015, 0.1, 1.7,0,0, 0.002,0,0, 0,0,0])"
definition "cube_enter = set_of_ivl (pairself eucl_of_list cube_enteri)"
value [code] "println ((show) (map (ivls_of_aforms' 100 o list_of_appr1e_aform)
(compute_cube_exit 30 (FloatR 1 (-10)) ((ereal 1, ereal 1),
(aforms_of_ivls (take 3 (fst cube_enteri))
(take 3 (snd cube_enteri)),
Some (aforms_of_ivls (drop 3 (fst cube_enteri))
(drop 3 (snd cube_enteri))))))))"
definition "cube_exiti =
[(aforms_of_ivls (map ldec [-0.12, -0.024, -0.012])
(map udec [-0.088, 0.024, 0.13]),
Some (aforms_of_ivls (map ldec [0,0,0, -0.56,0,0, -0.6,0,0])
(map udec [0,0,0, 0.56,0,0, -0.08,0,0]))),
(aforms_of_ivls (map ldec [ 0.088, -0.024, -0.012])
(map udec [ 0.12, 0.024, 0.13]),
Some (aforms_of_ivls (map ldec [0,0,0, -0.53,0,0, 0.08,0,0])
(map udec [0,0,0, 0.56,0,0, 0.6,0,0])))]"
definition "cube_exitv = aform.c1_info_of_apprs cube_exiti"
lemma cube_enteri[autoref_rules]: "(cube_enteri, cube_enter::'a set) \<in> lvivl_rel"
if "DIM_precond TYPE('a::executable_euclidean_space) 12"
using that
by (auto simp: cube_enteri_def cube_enter_def set_of_ivl_def
intro!: brI lv_relivl_relI)
lemma cube_exiti[autoref_rules]: "(cube_exiti, cube_exitv::'n eucl1 set) \<in> clw_rel aform.appr1_rel"
if "DIM_precond TYPE('n::enum rvec) 3"
unfolding cube_exitv_def cube_exiti_def
apply (rule aform.clw_rel_appr1_relI)
using that
by (auto simp: aform.c1_info_invar_def power2_eq_square)
definition "lorenz_interrupt (optns::real aform numeric_options) b (eX::3 eucl1 set) =
do {
((el, eu), X) \<leftarrow> scaleR2_rep eX;
let fX = fst ` X;
fentry \<leftarrow> op_image_fst_ivl (cube_enter::3 vec1 set);
interrupt \<leftarrow> aform.op_subset (fX:::aform.appr_rel) fentry;
(ol, ou) \<leftarrow> ivl_rep fentry;
aform.CHECKs (ST ''asdf'') (0 < el \<and> ol \<le> ou);
let _ = (if b then aform.trace_set (ST ''Potential Interrupt: '') (Some fX) else ());
let _ = (if b then aform.trace_set (ST ''With: '') (Some ({ol .. ou::3 rvec}:::aform.appr_rel)) else ());
if \<not>b \<or> \<not>interrupt then RETURN (op_empty_coll, mk_coll eX)
else do {
vX \<leftarrow> aform.vec1rep X;
let _ = (if b then aform.trace_set1e (ST ''Actual Interrupt: '') (Some eX) else ());
let l = (eucl_of_list [-1/2/2,-1/2/2,-1/2/2]::3 rvec);
let u = eucl_of_list [1/2/2, 1/2/2, 1/2/2];
ASSERT (l \<le> u);
let CX = mk_coll ({l .. u}:::aform.appr_rel);
(C0::3 eucl1 set) \<leftarrow> scaleRe_ivl_coll_spec el eu (fst ` cube_exitv \<times> UNIV);
(C1::3 eucl1 set) \<leftarrow> scaleRe_ivl_coll_spec el eu (cube_exitv);
case vX of None \<Rightarrow> RETURN (CX, C0)
| Some vX \<Rightarrow> do {
b \<leftarrow> aform.op_subset vX cube_enter;
aform.CHECKs (ST ''FAILED: TANGENT VECTORs are not contained'') b;
RETURN (CX, C1)
}
}
}"
lemma [autoref_rules]:
includes autoref_syntax
shows
"(real_of_int, real_of_int) \<in> int_rel \<rightarrow> rnv_rel"
"(ldec, ldec) \<in> Id \<rightarrow> rnv_rel"
"(udec, udec) \<in> Id \<rightarrow> rnv_rel"
by auto
schematic_goal lorenz_interrupti:
includes autoref_syntax
assumes[autoref_rules]: "(bi, b) \<in> bool_rel" "(Xi, X::3 eucl1 set) \<in> aform.appr1e_rel"
"(optnsi, optns) \<in> Id"
shows
"(nres_of ?r, lorenz_interrupt optns b X) \<in> \<langle>clw_rel aform.appr_rel \<times>\<^sub>r clw_rel aform.appr1e_rel\<rangle>nres_rel"
unfolding lorenz_interrupt_def
including art
by autoref_monadic
concrete_definition lorenz_interrupti for optnsi1 bi Xi uses
lorenz_interrupti[where optnsi = "optnsi"
and optnsa = "\<lambda>_ _ _ _ _ _ _ _. optnsi"
and optnsb = "\<lambda>_ _ _ _ _ _ _ _ _. optnsi"
and optnsc = "\<lambda>_ _ _ _ _ _ _ _ _ _ _. optnsi"
and optnsd = "\<lambda>_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. optnsi"
and optnse = "\<lambda>_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. optnsi"
and optnsf = "\<lambda>_ _ _ _ _ _ _ _ _. optnsi"
and optns = "\<lambda>_ _ _ _ _. optnsi"
for optnsi]
lemma lorenz_interrupti_refine[autoref_rules]:
includes autoref_syntax
shows
"(\<lambda>optnsi bi Xi. (lorenz_interrupti optnsi bi Xi),
lorenz_interrupt)
\<in> num_optns_rel \<rightarrow> bool_rel \<rightarrow> aform.appr1e_rel \<rightarrow> \<langle>clw_rel aform.appr_rel \<times>\<^sub>r clw_rel aform.appr1e_rel\<rangle>dres_nres_rel"
using lorenz_interrupti.refine
by (auto simp: nres_rel_def dres_nres_rel_def)
definition "(large_cube::R3 set) = {-1/4 .. 1/4} \<times> {-1/4 .. 1/4} \<times> {-1/4 .. 1/4}"
definition "cube_entry = (cast_eucl1 ` (flow1_of_vec1 ` cube_enter::3 eucl1 set)::R3 c1_info set)"
definition "cube_exit = (cast_eucl1 ` (cube_exitv::3 eucl1 set)::R3 c1_info set)"
text \<open>protect locale parameters\<close>
lemma flow0_cong[cong]: "auto_ll_on_open.flow0 ode X = auto_ll_on_open.flow0 ode X"
by (auto simp:)
lemma existence_ivl0_cong[cong]:
"auto_ll_on_open.existence_ivl0 ode X = auto_ll_on_open.existence_ivl0 ode X"
by (auto simp:)
lemma Dflow_cong[cong]: "c1_on_open_euclidean.Dflow ode ode_d X = c1_on_open_euclidean.Dflow ode ode_d X"
by (auto simp:)
lemma flowsto_cong[cong]:
"c1_on_open_euclidean.flowsto ode ode_d D = c1_on_open_euclidean.flowsto ode ode_d D"
by (auto simp:)
lemma poincare_mapsto_cong[cong]:
"c1_on_open_euclidean.poincare_mapsto ode X = c1_on_open_euclidean.poincare_mapsto ode X"
by (auto simp:)
lemma returns_to_cong[cong]:
"auto_ll_on_open.returns_to ode X = auto_ll_on_open.returns_to ode X"
by (auto simp:)
lemma return_time_cong[cong]:
"auto_ll_on_open.return_time ode X = auto_ll_on_open.return_time ode X"
by (auto simp: )
lemma poincare_map_cong[cong]:
"auto_ll_on_open.poincare_map ode X = auto_ll_on_open.poincare_map ode X"
by (auto simp: )
lemma eq_nth_iff_index:
"distinct xs \<Longrightarrow> n < length xs \<Longrightarrow> i = xs ! n \<longleftrightarrow> index xs i = n"
using index_nth_id by fastforce
lemma cast_in_BasisI: "(cast i::'a) \<in> Basis"
if "(i::'c) \<in> Basis""DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)"
using that
by (auto simp: cast_def eucl_of_list_nth inner_Basis if_distrib if_distribR
eq_nth_iff_index
cong: if_cong)
lemma cast_le_iff: "(cast (x::'a)::'c) \<le> y \<longleftrightarrow> x \<le> cast y"
if "DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)"
apply (auto simp: eucl_le[where 'a='a] eucl_le[where 'a='c]
dest!: bspec intro!: )
apply (rule cast_in_BasisI, assumption)
apply (auto simp: that)
apply (metis cast_eqI2 cast_inner that)
apply (rule cast_in_BasisI, assumption)
apply (auto simp: that)
apply (metis cast_eqI2 cast_inner that)
done
lemma cast_le_cast_iff: "(cast (x::'a)::'c) \<le> cast y \<longleftrightarrow> x \<le> y"
if "DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)"
apply (auto simp: eucl_le[where 'a='a] eucl_le[where 'a='c]
dest!: bspec intro!: )
apply (rule cast_in_BasisI, assumption)
apply (auto simp: that)
apply (rule cast_in_BasisI, assumption)
apply (auto simp: that)
by (metis cast_eqI2 cast_inner that)
lemma cast_image_Icc[simp]: "cast ` {a .. b::'c} = {cast a .. cast b::'a}"
if "DIM('c::executable_euclidean_space) = DIM('a::executable_euclidean_space)"
using that
apply (auto simp: cast_le_iff dest!:)
subgoal for x
apply (rule image_eqI[where x="cast x"])
by (auto simp: cast_le_iff)
done
lemma cast_eucl1_image_scaleR2:
"cast_eucl1 ` scaleR2 l u X = scaleR2 l u (cast_eucl1 ` (X::'b c1_info set)::'a c1_info set)"
if "DIM('a::executable_euclidean_space) = DIM('b::executable_euclidean_space)"
using that
apply (auto simp: scaleR2_def image_def vimage_def cast_eucl1_def
linear.scaleR linear_cast_bl)
apply force+
apply (rule exI conjI)+
apply assumption
apply (rule exI conjI)+
apply assumption
apply (rule bexI) prefer 2 apply assumption
apply force
by (auto simp: linear.scaleR linear_cast_bl)
lemma scaleR2_diff_prod2: "scaleR2 d e (X) - Y \<times> UNIV = scaleR2 d e (X - Y \<times> UNIV)"
by (force simp: scaleR2_def vimage_def image_def)
end
lemma (in c1_on_open_euclidean) flowsto_scaleR2I:
"flowsto (scaleR2 d e X0) T (CX \<times> UNIV) (scaleR2 d e Y)"
if "flowsto (X0) T (CX \<times> UNIV) (Y)"
using that
apply (auto simp: flowsto_def scaleR2_def)
apply (drule bspec, assumption)
apply auto
apply (rule bexI) prefer 2 apply assumption
apply auto
subgoal for x a b h
by (auto intro!: image_eqI[where x="(x, (flow0 a h, Dflow a h o\<^sub>L b))"] blinfun_eqI
simp: blinfun.bilinear_simps)
done
definition "aforms_of_resultrect x0 x1 y0 y1 = aforms_of_ivl (ivl_of_resultrect x0 x1 y0 y1)"
definition "flatten_varveq x = fst x @ the_default [] (snd x)"
(*
LessEqual (N\<^sub>r e) (N\<^sub>r em *\<^sub>e floatarith.Min (norm\<^sub>e [(3)\<^sub>e, (6)\<^sub>e, (9)\<^sub>e])
(norm\<^sub>e [(4)\<^sub>e, (7)\<^sub>e, (10)\<^sub>e]))
*)
derive "show" ereal
definition \<Sigma>::"(real*real*real) set" where
"\<Sigma> = {(-6, -6, 27) .. (6, 6, 27)}"
definition \<Sigma>\<^sub>l\<^sub>e::"(real*real*real) set" where
"\<Sigma>\<^sub>l\<^sub>e = {(x, y, z). z \<le> 27}"
definition "results = symmetrize coarse_results"
definition "results_at x = {res \<in> set results. x \<in> source_of_res res}"
text \<open>a part of the stable manifold (up to the (backward) first intersection with \<open>\<Sigma>\<close>)\<close>
definition \<Gamma>::"(real*real*real) set" where
"\<Gamma> = {x.
{0..} \<subseteq> lorenz.existence_ivl0 x \<and>
(\<forall>t>0. lorenz.flow0 x t \<notin> \<Sigma>) \<and>
(lorenz.flow0 x \<longlongrightarrow> 0) at_top}"
definition "\<Gamma>\<^sub>i intr = (if intr then \<Gamma> else {})"
definition "\<Gamma>\<^sub>i\<^sub>v intr = cast ` (\<Gamma>\<^sub>i intr)"
definition "sourcei_of_res res = source_of_res res - (\<Gamma>\<^sub>i (invoke_nf res))"
definition "resultsi_at x = {res \<in> set results. x \<in> sourcei_of_res res}"
definition "N = \<Union>(source_of_res ` (set results))"
definition "\<CC> x = \<Union>(conefield_of_res ` (results_at x))"
definition "R = lorenz.poincare_map \<Sigma>"
definition "DR x = frechet_derivative (lorenz.poincare_map \<Sigma>) (at x within \<Sigma>\<^sub>l\<^sub>e)"
definition "\<E> x = Min (expansion ` results_at x)"
definition "\<E>\<^sub>p x = Min (preexpansion ` results_at x)"
abbreviation returns_to (infixl "returns'_to" 50) where
"(x returns_to P) \<equiv> lorenz.returns_to P x"
lemma closed_\<Sigma>[intro, simp]: "closed \<Sigma>"
by (auto simp: \<Sigma>_def)
lemma \<Gamma>_stable: "lorenz.stable_on (- \<Sigma>) \<Gamma>"
unfolding lorenz.stable_on_def
proof (intro allI impI)
fix t x0
assume outside: "\<forall>s\<in>{0<..t}. lorenz.flow0 x0 s \<in> - \<Sigma>"
assume assms: "lorenz.flow0 x0 t \<in> \<Gamma>" "t \<in> lorenz.existence_ivl0 x0" "0 < t"
from assms have *: "{0..} \<subseteq> lorenz.existence_ivl0 (lorenz.flow0 x0 t)"
"(lorenz.flow0 (lorenz.flow0 x0 t) \<longlongrightarrow> 0) at_top"
by (auto simp: \<Gamma>_def)
have nonneg_exivl: "s \<in> lorenz.existence_ivl0 x0" if "s \<ge> 0" for s
proof (cases "s \<le> t")
case True
then show ?thesis
using \<open>0 \<le> s\<close> assms(2) lorenz.ivl_subset_existence_ivl[of t x0] by auto
next
case False
define u where "u = s - t"
with False have "u > 0" "s = t + u" by auto
note this(2)
also have "t + u \<in> lorenz.existence_ivl0 x0"
apply (rule lorenz.existence_ivl_trans)
apply fact
using * \<open>u > 0\<close> by auto
finally show ?thesis .
qed
show "x0 \<in> \<Gamma>"
unfolding \<Gamma>_def
proof (safe intro!: nonneg_exivl)
have "\<forall>\<^sub>F s in at_top. (s::real) \<ge> 0"
using eventually_ge_at_top by blast
then have "\<forall>\<^sub>F s in at_top. lorenz.flow0 (lorenz.flow0 x0 t) s = lorenz.flow0 x0 (s + t)"
proof eventually_elim
case (elim s)
then have "s \<in> lorenz.existence_ivl0 x0"
using nonneg_exivl[OF \<open>0 \<le> s\<close>] by simp
then have "lorenz.flow0 (lorenz.flow0 x0 t) s = lorenz.flow0 x0 (t + s)"
apply (subst lorenz.flow_trans)
using assms * elim by auto
then show ?case by (simp add: ac_simps)
qed
then have "((\<lambda>s. (lorenz.flow0 x0 (s + t))) \<longlongrightarrow> 0) at_top"
by (blast intro: * Lim_transform_eventually)
then show "(lorenz.flow0 x0 \<longlongrightarrow> 0) at_top"
unfolding aform.tendsto_at_top_translate_iff .
next
fix s::real assume s: "0 < s" "lorenz.flow0 x0 s \<in> \<Sigma>"
show False
proof (cases "s \<le> t")
case True
then show ?thesis
using outside s
by (auto simp: \<Gamma>_def)
next
case False
then obtain u where u: "u = s - t" "u > 0"
by auto
then have "lorenz.flow0 x0 (s) = lorenz.flow0 x0 (t + u)" by (simp add: algebra_simps)
also have "\<dots> = lorenz.flow0 (lorenz.flow0 x0 t) u"
apply (subst lorenz.flow_trans)
subgoal by fact
subgoal unfolding u
apply (rule lorenz.diff_existence_ivl_trans) apply fact+
apply (rule nonneg_exivl) using \<open>0 < s\<close> by simp
subgoal by simp
done
also from assms(1) \<open>u > 0\<close> have "\<dots> \<notin> \<Sigma>"
by (auto simp: \<Gamma>_def)
finally show ?thesis using s
by auto
qed
qed
qed
lemma (in auto_ll_on_open) stable_on_empty[intro, simp]: "stable_on asdf {}"
by (auto simp: stable_on_def)
lemma \<Gamma>\<^sub>i_stable: "lorenz.stable_on (- \<Sigma>) (\<Gamma>\<^sub>i b)"
using \<Gamma>_stable
unfolding \<Gamma>\<^sub>i_def
apply (cases b)
subgoal by auto
subgoal using lorenz.stable_on_empty
by (auto simp: \<Gamma>\<^sub>i_def)
done
definition "\<Gamma>\<^sub>v = (cast ` \<Gamma>)"
definition "NF = lorenz.flowsto (cube_entry - \<Gamma> \<times> UNIV) {0..} (large_cube \<times> UNIV) cube_exit"
lemma NF0: "lorenz.flowsto ((fst ` cube_entry - \<Gamma>) \<times> UNIV) {0..} (large_cube \<times> UNIV)
(fst ` cube_exit \<times> UNIV)"
if NF
using that
unfolding NF_def lorenz.flowsto_def
apply (auto simp: NF_def)
apply (drule bspec, force)
by auto
lemma [autoref_rules]: includes autoref_syntax shows "(\<lambda>_. (), \<Gamma>\<^sub>i\<^sub>v) \<in> bool_rel \<rightarrow> ghost_rel"
by (auto simp: ghost_rel_def)
lemma lorenz_interrupt[le, refine_vcg]:
"lorenz_interrupt optns b X \<le> SPEC (\<lambda>(CX, R).
lorenz.flowsto ((cast_eucl1 ` X::R3 c1_info set) - (\<Gamma>\<^sub>i b \<times> UNIV)) {0..} (cast ` CX \<times> UNIV) (cast_eucl1 ` R))"
if NF
unfolding lorenz_interrupt_def
apply refine_vcg
subgoal
by (rule lorenz.flowsto_self) auto
subgoal
by (auto simp: eucl_le[where 'a="3 rvec"] eucl_of_list_inner Basis_list_vec_def Basis_list_real_def)
subgoal for a b c d e f g h i j k l m n p
apply (auto)
apply (simp add: cast_eucl1_image_scaleR2 scaleR2_diff_prod2)
apply (erule make_neg_goal)
apply (rule lorenz.flowsto_scaleR2I)
using NF0[OF that]
apply (rule lorenz.flowsto_subset)
subgoal for q
apply (auto simp: scaleR2_def cast_eucl1_def)
apply (auto simp: linear_cast_bl linear.scaleR cube_entry_def cast_eucl1_def image_image)
subgoal premises prems for r s t u v
proof -
from prems have "fst ` c \<subseteq> fst ` (cube_enter::3 vec1 set)"
by auto
with \<open>(u, v) \<in> c\<close> obtain w where "((u, w)::3 vec1) \<in> cube_enter"
by auto
from _ this have "cast u \<in> (\<lambda>x. cast (fst (x::3 vec1))) ` cube_enter"
by (rule image_eqI) auto
then show ?thesis using prems
by blast
qed
subgoal by (auto simp: \<Gamma>\<^sub>i_def)
done
subgoal by simp
subgoal premises _
by (auto simp: eucl_le[where 'a="3 rvec"] eucl_of_list_inner Basis_list_vec_def Basis_list_real_def
eucl_of_list_def
large_cube_def)
subgoal premises prems for x
+ supply [[simproc del: defined_all]]
apply (auto simp:)
subgoal for A B C D E
apply (rule image_eqI[where x="cast_eucl1 (((B, C, D), A))"])
apply (auto simp: cast_eucl1_def)
subgoal premises prems
using prems(1)
apply (auto simp: cube_exit_def aform.c1_info_of_apprs_def cube_exiti_def
cast_eucl1_def aform.c1_info_of_appr_def)
done
done
done
done
subgoal for a b c d e f g h i j k l m n p
apply (auto)
apply (erule make_neg_goal, thin_tac "\<not> _")
apply (simp add: cast_eucl1_image_scaleR2 scaleR2_diff_prod2)
apply (rule lorenz.flowsto_scaleR2I)
using that[unfolded NF_def]
apply (rule lorenz.flowsto_subset)
subgoal for q
apply (auto simp: scaleR2_def cast_eucl1_def )
apply (auto simp: linear_cast_bl linear.scaleR cube_entry_def cast_eucl1_def image_image)
subgoal premises prems for r s t u v
proof -
from prems have \<open>vec1_of_flow1 (u, v) \<in> cube_enter\<close>
by auto
from _ this have "(cast u, cast_bl v) \<in> (\<lambda>x. (cast (fst (x::3 vec1)), cast_bl (snd (flow1_of_vec1 x)))) ` cube_enter"
by (rule image_eqI) (auto simp: )
then show ?thesis
using prems by blast
qed
subgoal by (auto simp: \<Gamma>\<^sub>i_def)
done
subgoal by simp
subgoal premises _
by (auto simp: eucl_le[where 'a="3 rvec"] eucl_of_list_inner Basis_list_vec_def Basis_list_real_def
eucl_of_list_def
large_cube_def)
subgoal by (simp add: cube_exit_def)
done
done
definition "lorenz_S X = (case X of (x, y, z) \<Rightarrow> (-x, -y, z))"
lemma lorenz_symI: "((\<lambda>t. lorenz_S (f t)) has_vderiv_on lf') T"
if "(f has_vderiv_on f') T" "\<And>t. t \<in> T \<Longrightarrow> lf' t = lorenz_S (f' t)"
using that
by (auto simp: has_vderiv_on_def lorenz_S_def split_beta' has_vector_derivative_def
intro!: derivative_eq_intros)
lemma lorenz_S:
"t \<in> lorenz.existence_ivl0 (lorenz_S X)" (is ?th1)
"lorenz.flow0 (lorenz_S X) t = lorenz_S (lorenz.flow0 X t)" (is ?th2)
if "t \<in> lorenz.existence_ivl0 X"
proof -
have 1: "((\<lambda>t. lorenz_S (lorenz.flow0 X t)) solves_ode
(\<lambda>_ (X0, X1, X2).
(E1 * X0 - K1 * (X0 + X1) * X2,
E2 * X1 + K1 * (X0 + X1) * X2,
E3 * X2 + (X0 + X1) * (K2 * X0 + K3 * X1))))
{0--t} UNIV"
apply (rule solves_odeI)
apply (rule lorenz_symI)
apply (rule lorenz.flow_has_vderiv_on_compose)
apply simp
apply simp
apply (rule derivative_intros)
apply (rule refl)
using that
apply (rule lorenz.in_existence_between_zeroI)
apply assumption
apply (rule refl)
unfolding lorenz_S_def
apply (split prod.splits)+
apply (simp add: field_simps)
apply simp
done
have "lorenz.flow0 X 0 = X"
unfolding lorenz.flow_initial_time_if
by simp
then have 2:
"lorenz_S (lorenz.flow0 X 0) = lorenz_S X"
"is_interval {0--t}"
"0 \<in> {0--t}"
"{0--t} \<subseteq> UNIV"
by auto
from lorenz.maximal_existence_flow[OF 1 2]
show ?th1 ?th2 by fast+
qed
lemma \<Sigma>\<^sub>l\<^sub>e_impl[autoref_rules]: "(Sctn [0, 0, 1] 27, \<Sigma>\<^sub>l\<^sub>e) \<in> \<langle>lv_rel\<rangle>below_rel"
apply (auto simp: below_rel_def \<Sigma>\<^sub>l\<^sub>e_def below_halfspace_def sctn_rel_def
intro!: relcompI[where b="Sctn (0, 0, 1) 27"] brI lv_relI)
subgoal
unfolding lv_rel_def
by (auto intro!: brI)
unfolding le_halfspace_def
by (auto intro!: brI)
lemma [autoref_rules]: "((), \<Gamma>\<^sub>v) \<in> ghost_rel"
by (auto intro!: ghost_relI)
no_notation vec_nth (infixl "$" 90) and vec_lambda (binder "\<chi>" 10)
abbreviation "guards_rel \<equiv> \<langle>clw_rel (\<langle>\<langle>lv_rel\<rangle>ivl_rel, \<langle>lv_rel\<rangle>plane_rel\<rangle>inter_rel) \<times>\<^sub>r aform.reach_optns_rel\<rangle>list_rel"
definition "aform_poincare_onto_from optns = aform.poincare_onto_from"
lemma aform_poincare_onto_from[autoref_rules]:
includes autoref_syntax
shows
"DIM_precond TYPE('b rvec) E \<Longrightarrow>
(XSi, XS::'b::enum eucl1 set) \<in> clw_rel aform.appr1e_rel \<Longrightarrow>
(sctni, sctn) \<in> \<langle>lv_rel\<rangle>sctn_rel \<Longrightarrow>
(ivli, ivl) \<in> \<langle>lv_rel\<rangle>ivl_rel \<Longrightarrow>
(Si, Sa) \<in> \<langle>lv_rel\<rangle>halfspaces_rel \<Longrightarrow>
(guardsi, guards) \<in> guards_rel \<Longrightarrow>
(symstartd, symstart) \<in> aform.appr1e_rel \<rightarrow> \<langle>clw_rel aform.appr_rel \<times>\<^sub>r clw_rel aform.appr1e_rel\<rangle>dres_nres_rel \<Longrightarrow>
((), trap) \<in> ghost_rel \<Longrightarrow>
(roi, roptn) \<in> aform.reach_optns_rel \<Longrightarrow>
(odoi, odo) \<in> ode_ops_rel \<Longrightarrow>
(optnsi, optns) \<in> num_optns_rel \<Longrightarrow>
(nres_of
(solve_poincare_map_aform optnsi E odoi symstartd Si guardsi ivli
sctni roi XSi),
aform_poincare_onto_from $ optns $ odo $ symstart $ trap $ Sa $ guards $
ivl $
sctn $
roptn $
XS)
\<in> \<langle>clw_rel aform.appr1e_rel\<rangle>nres_rel"
unfolding autoref_tag_defs aform_poincare_onto_from_def
using aform.poincare_onto_from_impl.refine[OF _ aform_ncc aform_ncc,
where 'a='b, of E odoi odo XSi XS Si Sa guardsi guards ivli ivl sctni sctn roi roptn
"(\<lambda>x. nres_of (symstartd x))" symstart symstartd trap optnsi,
unfolded autoref_tag_defs, OF _ _ _ _ _ _ _ _ _ order_refl]
by (auto simp: dest!: aform.dres_nres_rel_nres_relD)
definition "lorenz_odo_impl = init_ode_ops True True lorenz.odo"
interpretation autoref_op_pat_def lorenz.odo .
lemma lorenz_odo_impl[autoref_rules]: "(lorenz_odo_impl, lorenz.odo) \<in> ode_ops_rel"
by (auto simp: ode_ops_rel_def lorenz_odo_impl_def)
definition lorenz_poincare where
"lorenz_poincare optns interrupt guards roptn XS0 =
aform_poincare_onto_from optns lorenz.odo
(lorenz_interrupt optns interrupt)
(\<Gamma>\<^sub>i\<^sub>v interrupt:::ghost_rel)
((below_halfspaces {Sctn (eucl_of_list [0, 0, 1]) 27}::(real^3) set):::\<langle>lv_rel\<rangle>halfspaces_rel)
guards
(op_atLeastAtMost_ivl (eucl_of_list [-6, -6, 27]:::lv_rel) (eucl_of_list [6, 6, 27]:::lv_rel):::lvivl_rel::(real^3) set)
(Sctn (eucl_of_list [0, 0, -1]) (- 27)::(real^3) sctn)
roptn
XS0"
lemma [autoref_rules_raw]:
includes autoref_syntax
shows "((),
(OP \<Gamma>\<^sub>i\<^sub>v ::: bool_rel \<rightarrow> ghost_rel) $
(OP intr ::: bool_rel))
\<in> ghost_rel" by (auto simp: ghost_rel_def)
schematic_goal lorenz_poincare_impl[autoref_rules]:
includes autoref_syntax
assumes [autoref_rules]: "(XSi, XS) \<in> clw_rel aform.appr1e_rel"
"(intri, intr) \<in> bool_rel"
"(guardsi, guards) \<in> guards_rel"
"(roi, roptn) \<in> aform.reach_optns_rel"
"(optnsi, optns) \<in> num_optns_rel"
shows "(nres_of ?r, lorenz_poincare $ optns $ intr $ guards $ roptn $ XS) \<in>
\<langle>clw_rel aform.appr1e_rel\<rangle>nres_rel"
unfolding autoref_tag_defs
unfolding lorenz_poincare_def
including art
supply [autoref_rules_raw] = ghost_relI
by autoref_monadic
lemma cast_image_eqI: "cast ` X = Y"
if "DIM('a) = DIM('b)"
"(X::'a::executable_euclidean_space set) = cast ` (Y::'b::executable_euclidean_space set)"
using that
by (auto simp: image_image)
lemma transfer_\<Gamma>[transfer_rule]: "(rel_set lorenz.rel_ve) \<Gamma>\<^sub>v \<Gamma>"
unfolding \<Gamma>\<^sub>v_def
by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI)
lemma transfer_\<Sigma>\<^sub>l\<^sub>e[transfer_rule]: "(rel_set lorenz.rel_ve) (cast ` \<Sigma>\<^sub>l\<^sub>e) \<Sigma>\<^sub>l\<^sub>e"
by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI)
lemma transfer_\<Gamma>\<^sub>i[transfer_rule]: "(rel_fun (=) (rel_set lorenz.rel_ve)) \<Gamma>\<^sub>i\<^sub>v \<Gamma>\<^sub>i"
unfolding \<Gamma>\<^sub>i\<^sub>v_def
by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI)
lemma transfer_\<Sigma>[transfer_rule]: "(rel_set lorenz.rel_ve) (cast ` \<Sigma>) \<Sigma>"
by (auto simp: lorenz.rel_ve_cast' intro!: rel_setI)
lemma len_fas: "length lorenz_fas = 3"
by (auto simp: lorenz_fas_def)
lemma lorenz_poincare[le, refine_vcg]:
"lorenz_poincare optns intr guards roptn XS \<le> SPEC (\<lambda>R.
aform.poincare_mapsto lorenz.odo (cast ` \<Sigma>) (XS - (\<Gamma>\<^sub>i\<^sub>v intr \<times> UNIV)) (cast ` \<Sigma>\<^sub>l\<^sub>e)
UNIV R)"
if [refine_vcg]: NF
unfolding lorenz_poincare_def aform_poincare_onto_from_def
apply (refine_vcg)
subgoal by (simp add: aform.wd_def aform.ode_e_def len_fas)
subgoal for a b c d
apply (auto simp: lorenz.flowsto_eq[symmetric])
proof goal_cases
case 1
from 1(2)[unfolded lorenz.flowsto_eq[symmetric]]
show ?case
by transfer (simp add: lorenz.avflowsto_eq)
qed
subgoal
unfolding aform.stable_on_def unfolding autoref_tag_defs
proof (intro allI impI, goal_cases)
case (1 t x0)
from 1 have t: "t \<in> lorenz.v.existence_ivl0 x0"
using lorenz.vex_ivl_eq by simp
from 1 have f: "lorenz.v.flow0 x0 t \<in> \<Gamma>\<^sub>i\<^sub>v intr"
using lorenz.vflow_eq[OF t] by simp
from 1 have "lorenz.v.flow0 x0 s \<in> - cast ` \<Sigma>"
if "s\<in>{0<..t}" for s
proof -
from that t have s: "s \<in> lorenz.v.existence_ivl0 x0"
by (auto dest!: lorenz.a.v.closed_segment_subset_existence_ivl
simp: closed_segment_eq_real_ivl)
have "lorenz.v.flow0 x0 s
\<in> aform.Csafe lorenz.odo -
op_atLeastAtMost_ivl (eucl_of_list [- 6, - 6, 27]) (eucl_of_list [6, 6, 27]) \<inter>
plane_of (Sctn (eucl_of_list [0, 0, - 1]) (- 27))"
using 1(4)[rule_format, OF that]
unfolding lorenz.vflow_eq[OF s]
by auto
also have "\<dots> \<subseteq> - cast ` \<Sigma>"
by (auto simp: eucl_le[where 'a="real^3"] eucl_of_list_inner axis_eq_axis
cast_def Basis_list_real_def Basis_list_vec3 \<Sigma>_def plane_of_def
eucl_of_list_inner_eq inner_lv_rel_def)
finally show "lorenz.v.flow0 x0 s \<in> - cast ` \<Sigma>" .
qed
show "x0 \<in> \<Gamma>\<^sub>i\<^sub>v intr"
by (rule \<Gamma>\<^sub>i_stable[Transfer.untransferred, unfolded lorenz.v.stable_on_def, rule_format])
fact+
qed
subgoal for R
proof (clarsimp, goal_cases)
case 1
note 1(2)
also have "({eucl_of_list [- 6, - 6, 27]..eucl_of_list [6, 6, 27]::3 rvec} \<inter>
plane_of (Sctn (eucl_of_list [0, 0, - 1]) (- 27))) = cast ` \<Sigma>"
apply auto
apply (auto simp: \<Sigma>_def o_def plane_of_def
eucl_of_list_def Basis_list_R3 Basis_list_vec3)
subgoal
by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3)
subgoal
by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3)
subgoal
apply (auto simp: cast_le_iff[symmetric])
by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3 less_eq_prod_def
list_of_eucl_def inner_simps inner_axis_axis)
subgoal
apply (auto simp: cast_le_iff)
by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3 less_eq_prod_def
list_of_eucl_def inner_simps inner_axis_axis)
subgoal
by (auto simp: cast_def eucl_of_list_def Basis_list_R3 Basis_list_vec3 less_eq_prod_def
list_of_eucl_def inner_simps inner_axis_axis)
done
also have "(below_halfspaces {Sctn (eucl_of_list [0, 0, 1]::3 rvec) 27}) = (cast ` \<Sigma>\<^sub>l\<^sub>e)"
apply (auto simp: \<Sigma>\<^sub>l\<^sub>e_def below_halfspaces_def below_halfspace_def le_halfspace_def
eucl_of_list_def Basis_list_R3 Basis_list_vec3)
apply (rule image_eqI)
apply (rule cast_cast[symmetric])
by (auto simp: cast_def list_of_eucl_def o_def plane_of_def inner_simps inner_axis_axis
eucl_of_list_def Basis_list_R3 Basis_list_vec3)
finally
show ?case
by (rule aform.poincare_mapsto_subset) (force simp: lorenz.vX_def intro: cast_cast[symmetric])+
qed
done
context includes floatarith_notation begin
definition "mat1\<^sub>e =
[Var 0, Var 1, Var 2, Var 3, 0, 0,
Var 4, 0, 0,
Var 5, 0, 0]"
definition mat1_nres::"3 rvec set \<Rightarrow> 3 rvec set \<Rightarrow> 3 eucl1 set nres"
where
"mat1_nres X v = do {
Xv \<leftarrow> aform.approx_slp_appr mat1\<^sub>e (slp_of_fas mat1\<^sub>e) (concat ` listset [list_of_eucl ` X, list_of_eucl ` v]);
RETURN (flow1_of_vec1 ` Xv)
}"
lemma [simp]: "(x, x') \<in> aform.appr_rel \<Longrightarrow> aform.ncc x'"
using aform_ncc[where 'a='a]
by (auto simp: aform.ncc_precond_def)
lemma mat1e_autoref[autoref_rules]: "(mat1\<^sub>e, mat1\<^sub>e) \<in> \<langle>Id\<rangle>list_rel"
by auto
schematic_goal mat1_impl:
includes autoref_syntax
assumes [autoref_rules]: "(Xi, X) \<in> aform.appr_rel" "(vi, v) \<in> aform.appr_rel"
shows "(nres_of ?r, mat1_nres $ X $ v) \<in> \<langle>aform.appr1_rel\<rangle>nres_rel"
unfolding mat1_nres_def
including art
by autoref_monadic
concrete_definition mat1_impl for Xi vi uses mat1_impl
lemmas [autoref_rules] = mat1_impl.refine
lemma mat_nres[le, refine_vcg]:
"mat1_nres X v \<le> SPEC (\<lambda>M. X \<times> v \<subseteq> (\<lambda>x. (fst x, blinfun_apply (snd x) (eucl_of_list [1, 0, 0]))) ` M)"
unfolding mat1_nres_def
apply refine_vcg
apply (auto simp: dest!: bspec)
apply (auto simp: mat1\<^sub>e_def image_image )
subgoal for x a b
apply (rule image_eqI[where x="eucl_of_list
[(list_of_eucl a @ list_of_eucl b) ! 0, (list_of_eucl a @ list_of_eucl b) ! Suc 0,
(list_of_eucl a @ list_of_eucl b) ! 2, (list_of_eucl a @ list_of_eucl b) ! 3, 0, 0,
(list_of_eucl a @ list_of_eucl b) ! 4, 0, 0, (list_of_eucl a @ list_of_eucl b) ! 5, 0, 0]"])
apply (auto simp: eucl_of_list_prod eucl_of_list_inner nth_append Basis_list_vec3
intro!: euclidean_eqI[where 'a="3 rvec"])
unfolding Basis_list[symmetric] Basis_list_vec3
by (auto simp: flow1_of_vec1_def blinfun_of_vmatrix.rep_eq Basis_list_vec3 inner_simps
matrix_vector_mult_eq_list_of_eucl_nth ll3 inner_axis_axis)
done
definition [simp]: "op_image_cast_eucl1e = (`) cast_eucl1"
definition [simp]: "op_image_cast_eucl1e_coll = (`) cast_eucl1"
lemma prod_relI'': "\<lbrakk>(fst ab, a')\<in>R1; (snd ab, b')\<in>R2\<rbrakk> \<Longrightarrow> (ab,(a', b'))\<in>\<langle>R1,R2\<rangle>prod_rel"
by (auto simp: prod_rel_def)
lemma strange_aux_lemma:
"(b, b') \<in> A \<Longrightarrow> (b, snd (a'a, b')) \<in> A"
by auto
lemma [autoref_rules]:
includes autoref_syntax
assumes
"DIM_precond TYPE('a::executable_euclidean_space) D"
"DIM_precond TYPE('b::executable_euclidean_space) D"
shows
"(\<lambda>x. x, (op_image_cast_eucl1e::('a::executable_euclidean_space c1_info set \<Rightarrow> 'b::executable_euclidean_space c1_info set))) \<in> aform.appr1e_rel \<rightarrow> aform.appr1e_rel"
(is ?th1)
and "(\<lambda>x. x, op_image_cast_eucl1e_coll::'a::executable_euclidean_space c1_info set \<Rightarrow> 'b::executable_euclidean_space c1_info set) \<in> clw_rel aform.appr1e_rel \<rightarrow> clw_rel aform.appr1e_rel"
(is ?th2)
proof -
show 1: ?th1
unfolding scaleR2_rel_def
apply (rule subsetD)
apply (rule fun_rel_comp_dist)
apply (rule relcompI)
apply (rule fun_relI)
apply (erule prod_relE)
apply simp
apply (rule prod_relI)
apply simp
apply (rule fst_conv[symmetric])
apply (rule op_cast_eucl1_image_impl[OF assms, param_fo])
apply (rule strange_aux_lemma)
apply (auto simp: br_def scaleR2_def image_def vimage_def
cast_eucl1_def)
subgoal for a b c d e f g
apply (rule exI[where x=e] conjI)+
apply assumption
apply (rule exI conjI)+
apply assumption
apply (rule exI conjI)+
apply (rule bexI) prefer 2 apply assumption
apply (rule conjI) apply force
apply (rule refl)
using assms
by (auto simp: blinfun.bilinear_simps linear_cast linear.scaleR
intro!: blinfun_eqI)
subgoal for a b c d e f g
apply (rule exI[where x=e] exI conjI)+
apply assumption
apply (rule exI conjI)+
apply assumption
apply (rule bexI) prefer 2 apply assumption
apply force
using assms
by (auto simp: blinfun.bilinear_simps linear_cast linear.scaleR
intro!: blinfun_eqI)
done
have id_map: "(\<lambda>x. x) = (map (\<lambda>x. x))" by simp
show ?th2
apply (subst id_map)
apply (rule lift_clw_rel_map)
apply (rule relator_props)+
subgoal using 1 by auto
subgoal by auto
done
qed
definition "lorenz_poincare_tangents optns interrupt guards roptn c1 (X0::R3 set) (tangents::R3 set) =
do {
X0tanmat \<leftarrow> (if c1 then
do {
R \<leftarrow> mat1_nres (cast ` X0) (cast ` tangents);
RETURN (R::3 eucl1 set)
} else RETURN (cast ` X0 \<times> UNIV));
XDX0 \<leftarrow> scaleRe_ivl_spec 1 1 (X0tanmat);
let _ = aform.trace_set1e ''START'' (Some XDX0);
let _ = aform.print_set1e False (XDX0);
P \<leftarrow> lorenz_poincare optns interrupt guards roptn ((mk_coll XDX0:::clw_rel aform.appr1e_rel));
RETURN (op_image_cast_eucl1e_coll P::R3 c1_info set)
}"
lemma [autoref_rules_raw]: "DIM(real \<times> real \<times> real) = DIM((real, 3) vec)"
by auto
schematic_goal lorenz_poincare_tangents_impl:
includes autoref_syntax
assumes [autoref_rules]:
"(optnsi, optns) \<in> Id"
"(intrri, intr) \<in> bool_rel"
"(guardsi, guards) \<in> guards_rel"
"(roi, roptn) \<in> aform.reach_optns_rel"
"(c1i, c1) \<in> bool_rel"
"(X0i, X0) \<in> aform.appr_rel"
"(tangentsi, tangents) \<in> aform.appr_rel"
shows
"(nres_of ?r, lorenz_poincare_tangents $ optns $ intr $ guards $ roptn $ c1 $ (X0::R3 set) $ tangents) \<in>
\<langle>clw_rel aform.appr1e_rel\<rangle>nres_rel"
unfolding lorenz_poincare_tangents_def
including art
by (autoref_monadic)
concrete_definition lorenz_poincare_tangents_impl uses
lorenz_poincare_tangents_impl[where
optnsa = "\<lambda>_ _ _ _ _ _ _ _. optns"
and optnsb = "\<lambda>_ _ _ _ _ _ _ _ _. optns"
and optnsi = "optns"
and optnsc = "optns"
and optns = "\<lambda>_ _ _ _ _ _ _. optns"
for optns optnsc]
lemma lorenz_poincare_tangents_impl_refine[autoref_rules]:
includes autoref_syntax
shows
"(\<lambda>optnsi intrri guardsi roi c1i X0i tangentsi. nres_of
(lorenz_poincare_tangents_impl optnsi intrri guardsi roi c1i X0i tangentsi),
lorenz_poincare_tangents)
\<in> num_optns_rel \<rightarrow> bool_rel \<rightarrow> guards_rel \<rightarrow> aform.reach_optns_rel \<rightarrow> bool_rel \<rightarrow> aform.appr_rel \<rightarrow>
aform.appr_rel \<rightarrow>
\<langle>clw_rel aform.appr1e_rel\<rangle>nres_rel"
using lorenz_poincare_tangents_impl.refine by force
lemma transfer_UNIV_rel_blinfun[transfer_rule]:
"(rel_set (rel_blinfun lorenz.rel_ve lorenz.rel_ve)) UNIV UNIV"
apply (auto intro!: rel_setI simp: rel_blinfun_def)
subgoal for x
apply (rule exI[where x="cast_bl x"])
by (auto intro!: rel_funI simp: lorenz.rel_ve_cast)
subgoal for x
apply (rule exI[where x="cast_bl x"])
by (auto intro!: rel_funI simp: lorenz.rel_ve_cast)
done
lemma lorenz_vX[simp]: "lorenz.vX = (UNIV::3 rvec set)"
by (force simp: lorenz.vX_def intro!: cast_cast[symmetric])
lemma closed_cast_\<Sigma>[intro, simp]: "closed (cast ` \<Sigma>::3 rvec set)"
by (auto simp: \<Sigma>_def )
lemma blinfun_apply_transfer[transfer_rule]:
"(rel_fun (rel_blinfun lorenz.rel_ve lorenz.rel_ve)
(rel_fun (rel_prod (=) (rel_prod (=) (=))) lorenz.rel_ve))
(blinfun_apply o cast_bl) blinfun_apply"
by (auto intro!: rel_funI simp: rel_blinfun_def lorenz.rel_ve_cast
dest!: rel_funD)
lemma lorenz_poincare_tangents[le, refine_vcg]:
"lorenz_poincare_tangents optns intr guards roptn c1 (X0::R3 set) tangents \<le>
SPEC (\<lambda>x.
(if c1 then \<exists>tans. X0 \<times> tangents \<subseteq> (\<lambda>(x, y). (x, blinfun_apply y (1, 0, 0))) ` tans \<and> lorenz.poincare_mapsto \<Sigma> (tans - \<Gamma>\<^sub>i intr \<times> UNIV) (\<Sigma>\<^sub>l\<^sub>e) UNIV x
else lorenz.poincare_mapsto \<Sigma> ((X0 - \<Gamma>\<^sub>i intr) \<times> UNIV) (\<Sigma>\<^sub>l\<^sub>e) UNIV x))"
if [refine_vcg]: NF
unfolding lorenz_poincare_tangents_def
apply refine_vcg
apply auto
apply (subst lorenz.poincare_mapsto_eq[symmetric])
apply simp
proof goal_cases
case (2 R)
then show ?case
apply transfer
apply (subst lorenz.vpoincare_mapsto_eq[symmetric])
apply (auto simp: )
apply (rule aform.poincare_mapsto_subset, assumption)
by (force simp: scaleR2_def )+
next
case (1 tans R)
show ?case
apply (rule exI[where x="cast_eucl1 ` tans"])
apply (rule conjI)
subgoal including lifting_syntax
using 1(2)
by transfer (force simp: cast_def o_def)
subgoal
using 1(3) apply transfer
apply (subst lorenz.avpoincare_mapsto_eq[symmetric])
by (auto simp: )
done
qed
definition of_mat1_image::"R3 c1_info set \<Rightarrow> R3 set nres"
where [refine_vcg_def]: "of_mat1_image X = SPEC (\<lambda>R. R = (\<lambda>x. blinfun_apply (snd x) (1, 0, 0)) ` X)"
lemma of_mat1_image_impl[autoref_rules]:
"(\<lambda>x. (case x of (_, Some xs) \<Rightarrow> RETURN [xs ! 0, xs ! 3, xs ! 6]
| (_, None) \<Rightarrow> SUCCEED), of_mat1_image) \<in> aform.appr1_rel \<rightarrow> \<langle>aform.appr_rel\<rangle>nres_rel"
apply (auto simp: of_mat1_image_def RETURN_RES_refine_iff nres_rel_def
aform.appr1_rel_internal aform.appr_rel_def
intro!: relcompI
split: option.splits)
unfolding aforms_rel_def
apply (rule brI)
apply (auto simp: )
unfolding lv_rel_def set_rel_br
apply (rule brI)
prefer 2 apply (force simp: Joints_imp_length_eq)
apply (auto elim!: mem_Joints_appendE simp: flow1_of_list_def Joints_imp_length_eq)
subgoal for a b c d e f g h i j
apply (rule image_eqI[where x="[j ! 0, j ! 3, j! 6]"])
apply (auto simp: blinfun_of_list_def blinfun_of_matrix_apply
Basis_prod_def Basis_list_R3 Basis_list_vec3 eval_nat_numeral zero_prod_def)
apply (rule map_nth_Joints'[of _ _ "[0, Suc (Suc (Suc 0)), Suc (Suc (Suc (Suc (Suc (Suc 0)))))]",
simplified])
apply auto
done
subgoal for a b c d e f
unfolding image_image
apply (auto simp: Joints_def valuate_def)
subgoal for g
apply (rule image_eqI)
prefer 2
apply (rule image_eqI[where x=g])
apply (rule refl)
apply (auto simp: )
apply (auto simp: blinfun_of_list_def blinfun_of_matrix_apply flow1_of_list_def
Basis_prod_def Basis_list_R3 Basis_list_vec3 eval_nat_numeral zero_prod_def)
done
done
done
definition [refine_vcg_def]: "floatdegs res = SPEC (\<lambda>_::unit. min_deg res \<in> float \<and> max_deg res \<in> float)"
definition [simp]: "isinfloat x \<longleftrightarrow> x \<in> float"
lemma [code]: "isinfloat (real_of_float x) = True"
by (auto)
lemma floatdegs_impl[autoref_rules]:
includes autoref_syntax
shows
"(\<lambda>res. (if isinfloat (min_deg res) \<and> isinfloat (max_deg res) then RETURN () else SUCCEED), floatdegs)
\<in> Id \<rightarrow> \<langle>unit_rel\<rangle>nres_rel"
by (auto simp: nres_rel_def floatdegs_def)
definition "check_c1_entry optns em P (res0::result) (res::result) = do {
uv_ret \<leftarrow> of_mat1_image P;
nuv \<leftarrow> aform.mig_set 3 uv_ret;
floatdegs res0;
floatdegs res;
let e' = em * ereal nuv;
b1 \<leftarrow> approx_conefield_bounds P (min_deg res) (max_deg res);
let b2 = e' \<ge> preexpansion res;
let b3 = e' \<ge> expansion res0;
let _ = aform.print_msg ((shows ''# check_c1_entry: '' o shows_list [b1, b2, b3] o shows_space o
shows_list [e', preexpansion res, expansion res0]) '''');
RETURN (em \<ge> 0 \<and> b1 \<and> b2 \<and> b3)
}"
lemma [autoref_itype]:
"shows_prec ::\<^sub>i i_nat \<rightarrow>\<^sub>i A \<rightarrow>\<^sub>i i_string \<rightarrow>\<^sub>i i_string"
by auto
lemma [autoref_rules]:
includes autoref_syntax
shows
"PREFER_id A \<Longrightarrow> (shows_list, shows_list) \<in> \<langle>A\<rangle>list_rel \<rightarrow> string_rel \<rightarrow> string_rel"
"(shows_prec, shows_prec) \<in> nat_rel \<rightarrow> string_rel \<rightarrow> string_rel \<rightarrow> string_rel"
"(shows_prec, shows_prec) \<in> nat_rel \<rightarrow> ereal_rel \<rightarrow> string_rel \<rightarrow> string_rel"
"(shows_prec, shows_prec::_\<Rightarrow>result \<Rightarrow>_) \<in> nat_rel \<rightarrow> Id \<rightarrow> string_rel \<rightarrow> string_rel"
"(shows_space, shows_space) \<in> string_rel \<rightarrow> string_rel"
by (auto simp: string_rel_def)
lemma [autoref_rules]:
includes autoref_syntax
shows
"(expansion, expansion) \<in> Id \<rightarrow> rnv_rel"
"(preexpansion, preexpansion) \<in> Id \<rightarrow> rnv_rel"
"(min_deg, min_deg) \<in> Id \<rightarrow> rnv_rel"
"(max_deg, max_deg) \<in> Id \<rightarrow> rnv_rel"
by auto
interpretation autoref_op_pat_def aform.mig_set .
lemma [autoref_rules_raw]: "DIM_precond TYPE(real \<times> real \<times> real) (OP 3 ::: nat_rel)"
by simp
schematic_goal check_c1_entry_impl:
includes autoref_syntax
assumes [autoref_rules]:
"(optnsi, optns) \<in> Id"
"(res0i, res0) \<in> Id"
"(resi, res) \<in> Id"
"(emi, em) \<in> ereal_rel"
"(Pei, P) \<in> aform.appr1_rel"
shows
"(nres_of ?r, check_c1_entry optns em P res0 res) \<in> \<langle>bool_rel\<rangle>nres_rel"
unfolding check_c1_entry_def
including art
by autoref_monadic
concrete_definition check_c1_entry_impl uses check_c1_entry_impl[
where optns = "\<lambda>_ . optnsi"
and optnsi="optnsi"
and optnsc=optns
and optnsa="\<lambda>_ _ _ _ _. optnsi"
and optnsb="\<lambda>_ _ _ _ _ _ _ _ . optnsi"
and optns="\<lambda>_. optnsi"
for optns optnsi]
lemmas check_c1_entry_impl_refine[autoref_rules] = check_c1_entry_impl.refine[autoref_higher_order_rule]
definition "c1_entry_correct (em::ereal) (P::R3 c1_info set) res0 res = (\<forall>(_, d)\<in>P. case (d (1, 0, 0)) of (dx, dy, dz) \<Rightarrow>
dz = 0 \<and> dx > 0 \<and> -90 < min_deg res \<and> min_deg res \<le> max_deg res \<and> max_deg res < 90 \<and>
ereal (preexpansion res) \<le> em * (norm (dx, dy, dz)) \<and>
ereal (expansion res0) \<le> em * (norm (dx, dy, dz)) \<and>
dy / dx \<in> {tan (rad_of (min_deg res)) .. tan (rad_of (max_deg res))})"
lemma check_c1_entry[le, refine_vcg]:
"check_c1_entry optns em P res0 res \<le> SPEC (\<lambda>b. b \<longrightarrow> c1_entry_correct em P res0 res)"
unfolding check_c1_entry_def c1_entry_correct_def
apply refine_vcg
apply (auto dest!: bspec simp:)
apply (rule order_trans, assumption, rule ereal_mult_left_mono, force, force)
apply (rule order_trans, assumption, rule ereal_mult_left_mono, force, force)
done
subsection \<open>options for the lorenz system\<close>
definition aform_numeric_optns::"_ \<Rightarrow> _ \<Rightarrow> _ \<Rightarrow> _ \<Rightarrow> _ \<Rightarrow> _ \<Rightarrow> _ \<Rightarrow> _ \<Rightarrow> _ \<Rightarrow> _ \<Rightarrow> _ \<Rightarrow> _ \<Rightarrow>
real aform numeric_options" where
"aform_numeric_optns = numeric_options.fields"
fun zbucket::"real \<Rightarrow> real \<times> real \<Rightarrow> real \<times> real \<Rightarrow> real \<times> real \<Rightarrow> ((real list \<times> real list) \<times> real list sctn) list"
where "zbucket d (x0,x1) (y0, y1) (z0, z1) =
[zsec' (x0 - d, x0 + d) (y0 - d, y1 + d) z0, \<comment> \<open>bottom\<close>
xsec' x0 (y0 - d, y1 + d) (z0 - d, z1), \<comment> \<open>left\<close>
xsec x1 (y0 - d, y1 + d) (z0 - d, z1), \<comment> \<open>right\<close>
ysec' (x0 - d, x1 + d) y0 (z0 - d, z1), \<comment> \<open>backno\<close>
ysec (x0 - d, x1 + d) y1 (z0 - d, z1)] \<comment> \<open>front\<close>"
subsubsection \<open>Hybridizations\<close>
definition "reduce_weak_params c1 = (if c1 then (12::nat, 0::nat) else (3, 0))"
definition "reduce_hard_params c1 = (if c1 then (0::nat, 100::nat) else (0, 100))"
definition "ro_split_weak c1 w = (case reduce_weak_params c1 of (m, n) \<Rightarrow> ro (w + 1) m n w w (-5))"
definition "ro_split_weak' c1 w = (case reduce_weak_params c1 of (m, n) \<Rightarrow> ro w m n w w (-5))"
definition "ro_split_weak'' c1 w = (case reduce_weak_params c1 of (m, n) \<Rightarrow> ro (w + 2) m n w w (-5))"
definition "ro_split_weak4' c1 w = (case reduce_weak_params c1 of (m, n) \<Rightarrow> ro (w + 4) m n w w (-5))"
definition "ro_split_weak2 c1 w w2 = (case reduce_weak_params c1 of (m, n) \<Rightarrow> ro (w + 1) m n w w2 (-5))"
definition "ro_split_weak2' c1 w w2 = (case reduce_weak_params c1 of (m, n) \<Rightarrow> ro (w) m n w w2 (-5))"
definition "ro_split_hard c1 w0 w1 = (case reduce_hard_params c1 of (m, n) \<Rightarrow> ro (w0 + 1) m n w0 w1 (-5))"
definition "ro_split_hard'' c1 w0 w1 = (case reduce_hard_params c1 of (m, n) \<Rightarrow> ro (w0 + 2) m n w0 w1 (-5))"
definition "ro_split_not c1 w = (case reduce_weak_params c1 of (m, n) \<Rightarrow> ro 0 m n w w (-5))"
definition "ro_split_not2 c1 w w2 = (case reduce_weak_params c1 of (m, n) \<Rightarrow> ro 0 m n w w2 (-5))"
definition "xsecs x y z = [xsec' (-x) (-y, y) (0, z), xsec x (-y, y) (0, z)]"
type_synonym run_options = "(nat \<times> nat) \<times> int \<times>
(((real list \<times> real list) \<times> real list sctn) list \<times> real aform reach_options) list \<times> real aform reach_options"
abbreviation "p1 \<equiv> ldec 0.1"
definition mode_middle::"_ \<Rightarrow> run_options" where "mode_middle c1 = (reduce_weak_params c1, -14,
[([zsec' (-2, 2) (-1, 1) 10], ro_split_weak' c1 (-3)),
(xsecs (5 * p1) 10 10 @
xsecs p1 10 (6) @
[zsec' (-p1, p1) (-p1, p1) p1],
ro_split_hard c1 (-5) (-2)),
(xsecs (3/2/2) 10 (10), ro_split_not2 c1 0 (-2)), \<comment> \<open>To collect after interrupt\<close>
([zsec (-30, -6) (-10, 10) 30, zsec (6, 30) (-10, 10) 30], ro_split_not2 c1 (-1) (-3))
], ro_split_weak4' c1 (-5))"
definition mode_inner3::"bool\<Rightarrow>bool\<Rightarrow>run_options" where "mode_inner3 c1 very_inner =
(reduce_weak_params c1, -15,
(if very_inner then [([zsec' (-2, 2) (-1, 1) 10], ro_split_weak' c1 (-2))] else [])@
[(xsecs (3/2) 15 27@xsecs 1 (10) (11/2), ro_split_weak2 c1 (-2) (-1)),
([ zsec (-30, -6) (-10, 10) 30, zsec (6, 30) (-10, 10) 30], ro_split_not2 c1 (-1) (-3))
],
if very_inner then ro_split_weak4' c1 (-5) else ro_split_weak'' c1 (-5))"
definition mode_inner2::"bool \<Rightarrow> real \<Rightarrow> run_options" where "mode_inner2 c1 x =
(reduce_weak_params c1, -14,
[(xsecs x 10 27, ro_split_weak2' c1 (-2) (-1)),
([zsec ( -30, -6) (-10, 10) 30, zsec (6, 30) (-10, 10) 30], ro_split_not2 c1 (-3) (-3))],
ro_split_not c1 (-6))"
definition "ro_outer c1 w = (case reduce_weak_params c1 of (m, n) \<Rightarrow> ro w m n (-6) (-6) (-5))"
definition mode_outer::"bool\<Rightarrow>_\<Rightarrow>_\<Rightarrow>run_options" where
"mode_outer c1 w i = (reduce_weak_params c1, (-i),
[([zsec (-30, -6) (-10, 10) 27, zsec (6, 30) (-10, 10) 27], ro_split_not2 c1 (-3) (-4))], ro_outer c1 w)"
definition lookup_mode::"bool \<Rightarrow> result \<Rightarrow> _" where
"lookup_mode c1 i =
(if gridx0 i \<le> - 1024 then mode_outer c1 (-3) 16
else if gridx0 i \<le> - 120 then mode_outer c1 (-3) 14
else if gridx0 i \<le> 107 then mode_inner2 c1 (4)
else if gridx0 i \<le> 169 then mode_inner3 c1 False
else if gridx0 i \<le> 196 then mode_inner3 c1 True
else if gridx0 i \<le> 201 then mode_middle c1
else if gridx0 i \<le> 235 then mode_inner3 c1 True
else if gridx0 i \<le> 290 then mode_inner3 c1 False
else if gridx0 i \<le> 450 then mode_inner2 c1 4
else mode_outer c1 (-3) 14)"
definition mode_ro_spec::"bool \<Rightarrow> result \<Rightarrow> ((nat \<times> nat) \<times>
int \<times> ((real, 3) vec set \<times> unit) list \<times>
unit) nres"
where [refine_vcg_def]: "mode_ro_spec c1 res = SPEC (\<lambda>_. True)"
lemma reach_options_rel_br: "reach_options_rel TYPE('ty) = br (\<lambda>_. ()) (\<lambda>_. True)"
by (auto simp: reach_options_rel_def br_def)
lemma mode_ro_spec_impl[autoref_rules]:
includes autoref_syntax
shows "(\<lambda>b x. RETURN (lookup_mode b x), mode_ro_spec) \<in> bool_rel \<rightarrow> Id \<rightarrow>
\<langle>(nat_rel \<times>\<^sub>r nat_rel) \<times>\<^sub>r int_rel \<times>\<^sub>r guards_rel \<times>\<^sub>r aform.reach_optns_rel\<rangle>nres_rel"
supply [simp del] = prod_rel_id_simp
apply (rule fun_relI)
apply (rule fun_relI)
apply (rule nres_relI)
unfolding mode_ro_spec_def
apply (rule RETURN_SPEC_refine)
apply (auto simp: mode_ro_spec_def nres_rel_def RETURN_RES_refine_iff)
apply (rule exI)+
apply (rule prod_relI'' IdI)+
unfolding lv_rel_def ivl_rel_def br_rel_prod br_chain plane_rel_br inter_rel_br
clw_rel_br br_list_rel Id_br prod_eq_iff reach_options_rel_br
apply (rule brI refl)+
defer apply (rule brI) apply (rule refl) apply auto
apply (auto simp: lookup_mode_def mode_outer_def mode_inner2_def mode_inner3_def xsecs_def
mode_middle_def)
done
lemma [autoref_rules]: includes autoref_syntax shows
"(ivl_of_res, ivl_of_res) \<in> Id \<rightarrow> \<langle>rnv_rel\<rangle>list_rel \<times>\<^sub>r \<langle>rnv_rel\<rangle>list_rel"
by auto
lemma [autoref_rules]: includes autoref_syntax shows
"(Polygon.pairself, Polygon.pairself) \<in> (A \<rightarrow> C) \<rightarrow> (A \<times>\<^sub>r A) \<rightarrow> (C \<times>\<^sub>r C)"
by (auto dest: fun_relD)
lemma set_of_ivl_impl[autoref_rules]: includes autoref_syntax shows
"(\<lambda>x. x, set_of_ivl) \<in> (A \<times>\<^sub>r A) \<rightarrow> \<langle>A\<rangle>ivl_rel"
by (auto simp: ivl_rel_def br_def)
lemma eucl_of_list_pad: includes autoref_syntax shows
"DIM_precond TYPE('a::executable_euclidean_space) D \<Longrightarrow>
(\<lambda>xs. take D xs @ replicate (D - length xs) 0, eucl_of_list::_\<Rightarrow>'a) \<in> rl_rel \<rightarrow> lv_rel"
unfolding lv_rel_def
by (auto simp: intro!: brI)
concrete_definition eucl_of_list_pad uses eucl_of_list_pad
lemmas [autoref_rules] = eucl_of_list_pad.refine
lemma source_of_res_impl[autoref_rules]: includes autoref_syntax shows
"(ivl_of_res, source_of_res) \<in> Id \<rightarrow> \<langle>lv_rel\<rangle>ivl_rel"
unfolding source_of_res_def
apply (auto simp: ivl_rel_def intro!: relcompI brI)
subgoal for a
apply (auto simp: ivl_of_res_def ivl_of_resultrect_def
intro!: lv_relI)
unfolding lv_rel_def
apply (auto intro!: brI)
done
done
definition tangent_seg_of_res :: "real aform numeric_options \<Rightarrow> result \<Rightarrow> R3 set nres" where
"tangent_seg_of_res optns res0 = do {
let fas = map (OP (nth matrix_of_degrees2\<^sub>e)) [0, 3, 6];
let u = min_deg res0;
let v = max_deg res0;
aform.approx_slp_appr fas (slp_of_fas fas) (lv_ivl [u, v, 0] [u, v, 1])
}"
lemmas [refine_vcg_def] = tangent_seg_of_res_spec_def
lemma tangent_seg_of_res[le, refine_vcg]:
"tangent_seg_of_res optns res \<le> tangent_seg_of_res_spec res"
unfolding tangent_seg_of_res_def tangent_seg_of_res_spec_def
apply refine_vcg
apply (auto simp: matrix_of_degrees2\<^sub>e_def Let_def in_segment)
subgoal for x a b c u
by (drule bspec[where x="[min_deg res, max_deg res, u]"])
(auto simp: tangent_of_deg_def lv_ivl_def algebra_simps intro!:)
done
lemma [autoref_rules]: includes autoref_syntax shows
"(nth matrix_of_degrees2\<^sub>e, nth matrix_of_degrees2\<^sub>e) \<in> nat_rel \<rightarrow> Id"
by auto
schematic_goal tangent_seg_of_res_impl:
includes autoref_syntax
assumes [autoref_rules]: "(resi, res) \<in> Id"
"(optnsi, optns) \<in> num_optns_rel"
shows
"(nres_of ?r, tangent_seg_of_res optns res) \<in> \<langle>aform.appr_rel\<rangle>nres_rel"
unfolding tangent_seg_of_res_def
including art
by autoref_monadic
concrete_definition tangent_seg_of_res_impl uses tangent_seg_of_res_impl
lemmas [autoref_rules] = tangent_seg_of_res_impl.refine[where
optnsi = optnsi and optnsa=optns and optns="\<lambda>_ _ _. optnsi" for optns optnsi, autoref_higher_order_rule]
lemma return_of_res_impl:
includes autoref_syntax shows
"(\<lambda>results res. (get_results (inf_retx res) (inf_rety res) (sup_retx res) (sup_rety res) results),
return_of_res) \<in> \<langle>Id\<rangle>list_rel \<rightarrow> Id \<rightarrow> \<langle>Id\<rangle>list_wset_rel"
by (auto simp: return_of_res_def list_wset_rel_def intro!: brI)
concrete_definition return_of_res_impl uses return_of_res_impl
lemmas [autoref_rules] = return_of_res_impl.refine
lemma lorenz_optns'_impl[autoref_rules]: includes autoref_syntax shows
"(lorenz_optns', lorenz_optns') \<in>
\<langle>Id \<rightarrow> unit_rel\<rangle>option_rel \<rightarrow> nat_rel \<rightarrow> nat_rel \<rightarrow> rnv_rel \<rightarrow> rnv_rel \<rightarrow> num_optns_rel"
by auto
lemma [autoref_rules]:
includes autoref_syntax shows
"(results, results) \<in> \<langle>Id\<rangle>list_rel"
"(invoke_nf, invoke_nf) \<in> Id \<rightarrow> bool_rel"
by auto
definition "check_line_nres print_fun m0 n0 c1 res0 = do {
let X0 = source_of_res res0;
(X0l, X0u) \<leftarrow> ivl_rep X0;
((m::nat, n::nat), a::int, modes, ro) \<leftarrow> mode_ro_spec c1 res0;
let interrupt = invoke_nf res0;
let optns = lorenz_optns' print_fun (the_default m m0) (the_default n n0) 1 (FloatR 1 a);
tangents \<leftarrow> tangent_seg_of_res optns res0;
aform.CHECKs (ST ''check_line_nres le'') (X0l \<le> X0u);
sp \<leftarrow> aform.subset_spec_plane X0 (Sctn (eucl_of_list [0, 0, 1]) 27);
aform.CHECKs (ST ''check_line_nres le'') sp;
ASSERT (X0l \<le> X0u);
Pe \<leftarrow> lorenz_poincare_tangents optns interrupt modes ro c1 ({X0l .. X0u}) tangents;
PeS \<leftarrow> sets_of_coll Pe;
let RETs = (return_of_res results res0);
let RET = \<Union>((mk_coll ` (source_of_res ` RETs:::\<langle>lvivl_rel\<rangle>list_wset_rel):::\<langle>clw_rel lvivl_rel\<rangle>list_wset_rel));
every \<leftarrow> WEAK_ALL\<^bsup>\<lambda>Pe. \<exists>P em eM Rivls. em > 0 \<and> Pe = scaleR2 em eM P \<and> fst ` P \<subseteq> \<Union>Rivls \<and> (\<forall>Rivl \<in> Rivls. (\<exists>res\<in>RETs. Rivl \<subseteq> source_of_res res \<and> (c1 \<longrightarrow> c1_entry_correct em P res0 res)))\<^esup>
PeS (\<lambda>Pe. do {
let _ = aform.trace_set1e (ST ''# Return Element: '') (Some Pe);
((em, eM), P) \<leftarrow> scaleR2_rep Pe;
aform.CHECKs (ST ''check_line_nres pos'') (0 < em);
let R = (fst ` P:::aform.appr_rel);
(Ri, Rs) \<leftarrow> op_ivl_rep_of_set R;
let Rivl = (op_atLeastAtMost_ivl Ri Rs);
Rivls \<leftarrow> aform.split_along_ivls2 3 (mk_coll Rivl) RET;
Rivlss \<leftarrow> sets_of_coll Rivls;
WEAK_ALL\<^bsup>\<lambda>Rivl. \<exists>res\<in>RETs. Rivl \<subseteq> source_of_res res \<and> (c1 \<longrightarrow> c1_entry_correct em P res0 res)\<^esup> Rivlss
(\<lambda>Rivl. do {
b \<leftarrow>
WEAK_EX\<^bsup>\<lambda>res. Rivl \<subseteq> source_of_res res \<and> (c1 \<longrightarrow> c1_entry_correct em P res0 res)\<^esup> RETs
(\<lambda>res. do {
let src = (source_of_res res:::lvivl_rel);
let subs = Rivl \<subseteq> src;
cones \<leftarrow> if \<not>(c1 \<and> subs) then RETURN True else check_c1_entry optns em P res0 res;
RETURN (subs \<and> cones)
});
let _ = aform.print_msg ((shows (ST ''# return of '') o shows res0 o
shows (if b then ST '' OK'' else ST '' FAILED''))'''');
RETURN b
})
});
RETURN (every, Pe, RET)
}"
definition "aform_subset_spec_plane optns = aform.subset_spec_plane"
lemma aform_subset_spec_plane_impl[autoref_rules]: includes autoref_syntax shows
"DIM_precond TYPE('a::executable_euclidean_space) D \<Longrightarrow>
(Xi, X::'a set) \<in> \<langle>lv_rel\<rangle>ivl_rel \<Longrightarrow>
(sctni, sctn) \<in> \<langle>lv_rel\<rangle>sctn_rel \<Longrightarrow>
(optnsi, optns) \<in> num_optns_rel \<Longrightarrow>
(nres_of (subset_spec_plane_impl_aform optnsi D Xi sctni),
aform_subset_spec_plane $ optns $ X $ sctn)
\<in> \<langle>bool_rel\<rangle>nres_rel"
using aform.subset_spec_plane_impl.refine[where 'a='a, of D Xi X sctni sctn optnsi]
by (force simp: aform_subset_spec_plane_def)
schematic_goal check_line_impl:
includes autoref_syntax
assumes [autoref_rules]: "(pfi, pf) \<in> \<langle>Id \<rightarrow> unit_rel\<rangle>option_rel"
"(c1i, c1) \<in> bool_rel" "(res0i, res0) \<in> Id"
"(m0i, m0) \<in> \<langle>nat_rel\<rangle>option_rel" "(n0i, n0) \<in> \<langle>nat_rel\<rangle>option_rel"
shows
"(nres_of ?r, check_line_nres $ pf $ m0 $ n0 $ c1 $ res0) \<in>
\<langle>bool_rel \<times>\<^sub>r clw_rel aform.appr1e_rel \<times>\<^sub>r clw_rel lvivl_rel\<rangle>nres_rel"
unfolding check_line_nres_def
including art
by autoref_monadic
concrete_definition check_line_impl uses check_line_impl[where
optns = "\<lambda>_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ . lorenz_optns pfi"
and optnsa = "\<lambda>_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi"
and optnsb = "\<lambda>_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi"
and optnsc = "\<lambda>_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi"
and optnsd = "\<lambda>_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi"
and optnse = "\<lambda>_ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi"
and optnsf = "\<lambda> _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _. lorenz_optns pfi"
and pfi = pfi
for pfi]
lemmas [autoref_rules] = check_line_impl.refine
lemma check_line_nres:
"check_line_nres pf m0 n0 c1 res0 \<le> SPEC (\<lambda>(every, Pe, RET). \<exists>P. Pe = \<Union>P \<and>
(if c1
then \<exists>tans. source_of_res res0 \<times> {tangent_of_deg (min_deg res0)--tangent_of_deg (max_deg res0)} \<subseteq> (\<lambda>(x, y). (x, blinfun_apply y (1, 0, 0))) ` tans \<and>
lorenz.poincare_mapsto \<Sigma> (tans - \<Gamma>\<^sub>i (invoke_nf res0) \<times> UNIV) \<Sigma>\<^sub>l\<^sub>e UNIV (Pe)
else lorenz.poincare_mapsto \<Sigma> ((sourcei_of_res res0) \<times> UNIV) \<Sigma>\<^sub>l\<^sub>e UNIV (\<Union>P)) \<and>
source_of_res res0 \<subseteq> plane_of (Sctn (0, 0, 1) 27) \<and>
(every \<longrightarrow>
(\<forall>x\<in>P. \<exists>P em. em > 0 \<and> (\<exists>eM. x = scaleR2 em eM P) \<and>
(\<exists>Rivls. fst ` P \<subseteq> \<Union>Rivls \<and>
(\<forall>Rivl\<in>Rivls.
\<exists>res\<in>return_of_res results res0.
Rivl \<subseteq> source_of_res res \<and> (c1 \<longrightarrow> c1_entry_correct em P res0 res))))))"
if [refine_vcg]: NF
unfolding check_line_nres_def sourcei_of_res_def
apply (refine_vcg, clarsimp_all)
using [[goals_limit=1]]
subgoal for a0 a b c d e f g h i j k l m n p q r s t
apply (rule exI[where x=i])
apply (rule exI[where x=j])
apply (rule conjI)
apply force
apply (rule conjI)
apply (rule exI[where x=k])
apply force
apply (rule exI[where x=s])
apply (rule conjI) defer apply force
apply blast
done
subgoal for x0 y0 z0 x1 y1 z1 tangents s R every
apply (rule exI[where x=R])
apply auto
subgoal for tans
apply (rule exI[where x=tans])
by auto
subgoal for tans
apply (rule exI[where x=tans])
by auto
done
done
definition "print_sets_color (print_fun::String.literal \<Rightarrow> unit) (c::string) (X::'a::executable_euclidean_space set) = ()"
definition "print_lorenz_color print_fun cx cy cz ci cd1 cd2 P = ()"
definition "print_aforms print_fun c aforms =
fold (\<lambda>a _. print_fun (String.implode (shows_segments_of_aform 0 1 a c ''\<newline>''))) aforms ()"
lemma print_sets_color_impl[autoref_rules]: includes autoref_syntax shows
"(\<lambda>print_fun c X. print_aforms print_fun c X, print_sets_color) \<in>
(Id \<rightarrow> unit_rel) \<rightarrow> string_rel \<rightarrow> clw_rel aform.appr_rel \<rightarrow> unit_rel"
by auto
lemma print_lorenz_color_impl[autoref_rules]: includes autoref_syntax shows
"(\<lambda>print_fun cx cy cz ci cd1 cd2 P.
fold (\<lambda>(_, x) b.
print_lorenz_aform print_fun
cx cy cz ci cd1 cd2
False
(fst x @ the_default [] (snd x))
) P (), print_lorenz_color) \<in>
(Id \<rightarrow> unit_rel) \<rightarrow> string_rel \<rightarrow> string_rel \<rightarrow> string_rel \<rightarrow> string_rel \<rightarrow>
\<langle>\<langle>string_rel\<rangle>list_rel, \<langle>\<langle>string_rel\<rangle>list_rel, (\<langle>clw_rel aform.appr1e_rel, unit_rel\<rangle>fun_rel)\<rangle>fun_rel\<rangle>fun_rel"
by auto
definition check_line_core where
"check_line_core print_funo m0 n0 c1 i =
do {
let print_fun = the_default (\<lambda>_. ()) print_funo;
CHECK (\<lambda>_. print_fun (STR ''Hey, out of bounds!'')) (i < length results);
let res = ((results:::\<langle>Id\<rangle>list_rel) ! (i:::nat_rel));
(r, P, B) \<leftarrow> check_line_nres print_funo m0 n0 c1 res;
let _ = print_sets_color print_fun (ST ''0x007f00'') (aform.sets_of_ivls B);
(_, Pu) \<leftarrow> scaleR2_rep_coll P;
let _ = print_sets_color print_fun (ST ''0x7f0000'')
(aform.op_image_fst_coll (Pu:::clw_rel aform.appr1_rel):::clw_rel aform.appr_rel);
let _ = print_lorenz_color print_fun
(ST ''0x7f0000'') (ST ''0x7f0001'') (ST ''0x7f0002'') (ST ''0x7f0003'')
[(ST ''0xc60000''), (ST ''0xc60000''), (ST ''0xc60000'')]
[(ST ''0xf60000''), (ST ''0xf60000''), (ST ''0xf60000'')]
P;
let _ = (if r
then print_fun
(String.implode ((show (ST ''# VERIFIED '') @ show i @ show (ST ''\<newline>''))))
else print_fun (String.implode ((show (ST ''# Failed to verify '') @ show i @ show (ST ''\<newline>'')) )));
RETURN r
}"
lemma [autoref_rules]: includes autoref_syntax shows
"(shows_prec, shows_prec) \<in> nat_rel \<rightarrow> nat_rel \<rightarrow> string_rel \<rightarrow> string_rel"
"(shows_prec, shows_prec) \<in> nat_rel \<rightarrow> string_rel \<rightarrow> string_rel \<rightarrow> string_rel"
"(String.implode, String.implode) \<in> string_rel \<rightarrow> Id"
by (auto simp: string_rel_def)
schematic_goal check_line_core_impl:
includes autoref_syntax
assumes [autoref_rules]: "(pfi, pf) \<in> \<langle>Id \<rightarrow> unit_rel\<rangle>option_rel"
"(c1i, c1) \<in> bool_rel" "(ii, i) \<in> nat_rel"
"(m0i, m0) \<in> \<langle>nat_rel\<rangle>option_rel" "(n0i, n0) \<in> \<langle>nat_rel\<rangle>option_rel"
shows "(nres_of ?f, check_line_core $ pf $ m0 $ n0 $ c1 $ i) \<in> \<langle>bool_rel\<rangle>nres_rel"
unfolding check_line_core_def
including art
by autoref_monadic
concrete_definition check_line_core_impl for pfi m0i n0i c1i ii uses check_line_core_impl
lemmas [autoref_rules] = check_line_core_impl.refine
definition "c1i_of_res res = sourcei_of_res res \<times> conefield_of_res res"
definition "correct_res res =
((\<forall>(x, dx) \<in> c1i_of_res res.
x \<in> plane_of (Sctn (0, 0, 1) 27) \<and>
dx \<in> plane_of (Sctn (0, 0, 1) 0) \<and>
((lorenz.returns_to \<Sigma> x \<and>
lorenz.return_time \<Sigma> differentiable at x within \<Sigma>\<^sub>l\<^sub>e \<and>
(\<exists>D. (lorenz.poincare_map \<Sigma> has_derivative D) (at x within \<Sigma>\<^sub>l\<^sub>e) \<and>
norm (D dx) \<ge> expansion res * norm dx \<and>
(\<exists>res2 \<in> return_of_res results res.
(lorenz.poincare_map \<Sigma> x, D dx) \<in> c1_of_res res2 \<and>
norm (D dx) \<ge> preexpansion res2 * norm dx))))))"
lemma check_line_nres_c0_correct:
"check_line_nres pf m0 n0 c1 res \<le> SPEC (\<lambda>(every, Pe, RET). every \<longrightarrow>
(\<forall>x \<in> sourcei_of_res res. lorenz.poincare_map \<Sigma> x \<in> \<Union>(source_of_res ` return_of_res results res)))"
if NF
apply (rule check_line_nres[OF \<open>NF\<close>, le])
apply (auto simp: c1i_of_res_def lorenz.poincare_mapsto_def)
subgoal
apply (drule bspec, force)
apply (auto dest!: spec[where x="1\<^sub>L"])
apply (drule bspec, force)
apply (force simp: scaleR2_def image_def vimage_def)
done
subgoal premises prems for a b c d e tans
proof -
obtain t where
"((c, d, e), t) \<in> tans - \<Gamma>\<^sub>i (invoke_nf res) \<times> UNIV"
"((c, d, e), tangent_of_deg (min_deg res)) = (\<lambda>(x, y). (x, blinfun_apply y (1, 0, 0))) ((c, d, e), t)"
using prems
by (auto simp: sourcei_of_res_def)
with prems(6)[rule_format, of "((c, d, e), t)"] prems(3)
obtain D x where "tangent_of_deg (min_deg res) = blinfun_apply t (1, 0, 0)"
"(c, d, e) returns_to \<Sigma>"
"fst ` (tans - \<Gamma>\<^sub>i (invoke_nf res) \<times> UNIV) \<subseteq> \<Sigma>\<^sub>l\<^sub>e"
"lorenz.return_time \<Sigma> differentiable at (c, d, e) within \<Sigma>\<^sub>l\<^sub>e"
"(lorenz.poincare_map \<Sigma> has_derivative blinfun_apply D) (at (c, d, e) within \<Sigma>\<^sub>l\<^sub>e)"
"x \<in> b" "(lorenz.poincare_map \<Sigma> (c, d, e), D o\<^sub>L t) \<in> x"
by auto
with prems
show ?thesis
subgoal
apply (auto dest!: bspec[OF _ \<open>x \<in> b\<close>])
apply (auto simp: scaleR2_def image_def vimage_def)
apply (auto simp: subset_iff)
by fastforce \<comment>\<open>slow\<close>
done
qed
subgoal for a b c d e f
apply (drule bspec[where A="sourcei_of_res res"]) apply force
apply (auto dest!: spec[where x="1\<^sub>L"])
apply (drule bspec, force)
apply auto
apply (auto simp: scaleR2_def image_def vimage_def)
apply (auto simp: subset_iff)
by fastforce\<comment> \<open>slow\<close>
done
lemma cone_conefield[intro, simp]: "cone (conefield a b)"
unfolding conefield_alt_def
by (rule cone_cone_hull)
lemma in_segment_norm_bound: "c \<in> {a -- b} \<Longrightarrow> norm c \<le> max (norm a) (norm b)"
apply (auto simp: in_segment max_def intro!: norm_triangle_le)
apply (auto simp: algebra_simps intro: add_mono mult_left_mono
mult_right_mono)
using affine_ineq by blast
lemma norm_tangent_of_deg[simp]: "norm (tangent_of_deg d) = 1"
by (auto simp: tangent_of_deg_def norm_prod_def)
lemma check_line_nres_c1_correct:
"check_line_nres pf m0 n0 True res \<le> SPEC (\<lambda>(correct, Pe, RET). correct \<longrightarrow> correct_res res)"
if NF
proof (rule check_line_nres[OF \<open>NF\<close>, le], clarsimp, goal_cases)
case P: (1 a P tans)
let ?tans = "{tangent_of_deg (min_deg res)--tangent_of_deg (max_deg res)}"
have tans_plane: "?tans \<subseteq> UNIV \<times> UNIV \<times> {0}"
by (auto simp: in_segment tangent_of_deg_def)
from P have *: "x \<in> plane_of (Sctn (0, 0, 1) 27)" if "x \<in> sourcei_of_res res" for x
using that
by (auto simp: that sourcei_of_res_def)
from tans_plane P have **: "dx \<in> plane_of (Sctn (0, 0, 1) 0)"
if "x \<in> sourcei_of_res res" "dx \<in> conefield_of_res res" for x dx
proof -
from tans_plane that obtain c dx' dy' where
c: "dx = c *\<^sub>R (dx', dy', 0)" "(dx', dy', 0) \<in> ?tans" "c \<ge> 0"
unfolding conefield_of_res_def conefield_alt_def cone_hull_expl
by auto
then show ?thesis by (auto simp: plane_of_def)
qed
show ?case
unfolding correct_res_def
proof (intro ballI conjI, clarsimp_all simp add:
* ** c1i_of_res_def c1_of_res_def sourcei_of_res_def, goal_cases)
case source: (1 x y z dx dy dz)
from tans_plane source obtain c dx' dy' where
c: "(dx, dy, dz) = c *\<^sub>R (dx', dy', 0)" "(dx', dy', 0) \<in> ?tans" "c \<ge> 0"
unfolding conefield_of_res_def conefield_alt_def cone_hull_expl
by auto
from c source P
obtain t where tans: "((x, y, z), t) \<in> tans" "blinfun_apply t (1, 0, 0) = (dx', dy', 0)"
by auto
from P(3) tans(1) source(3) obtain Re D where Re:
"(x, y, z) returns_to \<Sigma>" "fst ` (tans - \<Gamma>\<^sub>i (invoke_nf res) \<times> UNIV) \<subseteq> \<Sigma>\<^sub>l\<^sub>e"
"lorenz.return_time \<Sigma> differentiable at (x, y, z) within \<Sigma>\<^sub>l\<^sub>e"
"(lorenz.poincare_map \<Sigma> has_derivative blinfun_apply D) (at (x, y, z) within \<Sigma>\<^sub>l\<^sub>e)"
"Re \<in> P" "(lorenz.poincare_map \<Sigma> (x, y, z), D o\<^sub>L t) \<in> Re"
by (auto simp: lorenz.poincare_mapsto_def dest!: bspec)
from P(5)[rule_format, OF \<open>Re \<in> P\<close>]
obtain R em eM Rivls
where R: "Re = scaleR2 em eM R"
"em > 0"
"fst ` R \<subseteq> \<Union>Rivls"
"\<And>Rivl. Rivl\<in>Rivls \<Longrightarrow>
\<exists>resa\<in>return_of_res results res. Rivl \<subseteq> source_of_res resa \<and> c1_entry_correct em R res resa"
by auto
have "lorenz.poincare_map \<Sigma> (x, y, z) \<in> fst ` R"
and s2: "(lorenz.poincare_map \<Sigma> (x, y, z), D o\<^sub>L t) \<in> scaleR2 em eM R"
using Re R by (auto simp: scaleR2_def)
then obtain Rivl res' where Rivl:
"lorenz.poincare_map \<Sigma> (x, y, z) \<in> Rivl" "Rivl \<in> Rivls"
"res' \<in> return_of_res results res" "Rivl \<subseteq> source_of_res res'"
and c1: "c1_entry_correct em R res res'"
using R
by force
from s2 obtain ed Dt where Dt:
"em \<le> ereal ed" "ereal ed \<le> eM" "D o\<^sub>L t = ed *\<^sub>R Dt"
"(lorenz.poincare_map \<Sigma> (x, y, z), Dt) \<in> R"
by (force simp: scaleR2_def)
then have Dt_simp[simp]: "Dt = inverse ed *\<^sub>R (D o\<^sub>L t)"
using \<open>0 < em\<close>
by (cases em) (auto simp: intro!: simp: blinfun.bilinear_simps inverse_eq_divide)
from c1[unfolded c1_entry_correct_def, rule_format, OF Dt(4)]
obtain dxr dyr where dxrdyr: "blinfun_apply D (dx', dy', 0) /\<^sub>R ed = (dxr, dyr, 0)"
"ereal (preexpansion res') \<le> em * ereal (norm (dxr, dyr, 0::real))"
"ereal (expansion res) \<le> em * ereal (norm (dxr, dyr, 0::real))"
"-90 < min_deg res'" "min_deg res' \<le> max_deg res'"
"tan (rad_of (min_deg res')) \<le> (dyr / dxr)"
"(dyr / dxr) \<le> tan (rad_of (max_deg res'))"
"max_deg res' < 90"
"0 < dxr"
by (auto simp: blinfun.bilinear_simps tans)
then obtain emr where emr: "em = ereal emr" "0 < emr" "emr \<le> ed"
"(preexpansion res') \<le> emr * (norm (dxr, dyr, 0::real))"
"(expansion res) \<le> emr * (norm (dxr, dyr, 0::real))"
using \<open>0 < em\<close> Dt
by (cases em) (auto simp: simp: blinfun.bilinear_simps divide_simps prod_eq_iff)
from dxrdyr have Ddx'dy': "D (dx', dy', 0) = ed *\<^sub>R (dxr, dyr, 0)"
using \<open>0 < em\<close> Dt
by (cases em) (auto simp: simp: blinfun.bilinear_simps divide_simps prod_eq_iff)
note \<open>(x, y, z) returns_to \<Sigma>\<close>
moreover note \<open>lorenz.return_time \<Sigma> differentiable at (x, y, z) within \<Sigma>\<^sub>l\<^sub>e\<close>
moreover note \<open>(lorenz.poincare_map \<Sigma> has_derivative D) (at (x, y, z) within \<Sigma>\<^sub>l\<^sub>e)\<close>
moreover note \<open>res' \<in> return_of_res results res\<close>
moreover have "lorenz.poincare_map \<Sigma> (x, y, z) \<in> source_of_res res'"
using Rivl by force
moreover
have \<open>0 \<le> ed\<close> using Dt \<open>0 < em\<close> by (cases em) auto
have \<open>D (dx, dy, dz) \<in> conefield_of_res res'\<close>
unfolding c blinfun.bilinear_simps conefield_of_res_def Ddx'dy'
apply (intro mem_cone, simp_all add: \<open>0 \<le> ed\<close> \<open>0 \<le> c\<close> tangent_of_deg_def)
apply (rule conefield_prod3I)
unfolding fun_cong[OF tan_def, symmetric]
subgoal by fact
subgoal using dxrdyr
apply (intro cos_gt_zero_pi)
unfolding rad_of_lt_iff rad_of_gt_iff
by (auto simp: deg_of_def)
subgoal using dxrdyr
apply (intro cos_gt_zero_pi)
unfolding rad_of_lt_iff rad_of_gt_iff
by (auto simp: deg_of_def)
subgoal by fact
subgoal by fact
done
moreover
have norms_le: "emr * norm (dx', dy', 0::real) * (\<bar>c\<bar> * norm (dxr, dyr, 0::real)) \<le> \<bar>ed\<bar> * (\<bar>c\<bar> * norm (dxr, dyr, 0::real))"
proof -
from c(2)[THEN in_segment_norm_bound] have "norm (dx', dy', 0::real) \<le> 1"
by auto
also have "\<dots> \<le> ed / emr"
using dxrdyr emr
unfolding Ddx'dy'
by auto
finally show ?thesis
using emr
by (intro mult_right_mono) (auto simp: divide_simps ac_simps)
qed
then have "expansion res * norm (dx, dy, dz) \<le> norm (D (dx, dy, dz))"
unfolding c blinfun.bilinear_simps conefield_of_res_def Ddx'dy' norm_scaleR
apply -
apply (rule order_trans)
apply (rule mult_right_mono)
apply (rule emr)
by (auto simp: ac_simps)
moreover have "preexpansion res' * norm (dx, dy, dz) \<le> norm (D (dx, dy, dz))"
using norms_le
unfolding c blinfun.bilinear_simps conefield_of_res_def Ddx'dy' norm_scaleR
apply -
apply (rule order_trans)
apply (rule mult_right_mono)
apply (rule emr)
by (auto simp: ac_simps)
ultimately
show ?case
by blast
qed
qed
lemma conefield_ne_empyt[simp]: "conefield a b \<noteq> {}"
by (auto simp: conefield_def conesegment_def cone_hull_empty_iff[symmetric])
lemma in_return_of_resD: "res' \<in> return_of_res results res \<Longrightarrow> res' \<in> set results"
by (auto simp: return_of_res_def get_results_def)
lemma finite_results_at[intro, simp]: "finite (results_at x)"
by (auto simp: results_at_def)
lemma lorenz_bounds_lemma:
"x returns_to \<Sigma>"
"R x \<in> N"
"(R has_derivative DR x) (at x within \<Sigma>\<^sub>l\<^sub>e)"
"\<And>c. c \<in> \<CC> x \<Longrightarrow> DR x c \<in> \<CC> (R x)"
"\<And>c. c \<in> \<CC> x \<Longrightarrow> norm (DR x c) \<ge> \<E> x * norm c"
"\<And>c. c \<in> \<CC> x \<Longrightarrow> norm (DR x c) \<ge> \<E>\<^sub>p (R x) * norm c"
if "x \<in> N - \<Gamma>" NF "\<And>res. res \<in> set results \<Longrightarrow> correct_res res"
proof -
from \<open>x \<in> N - \<Gamma>\<close> obtain res where res: "res \<in> set results" "x \<in> sourcei_of_res res"
by (auto simp: N_def sourcei_of_res_def \<Gamma>\<^sub>i_def)
then have ne: "c1i_of_res res \<noteq> {}"
by (auto simp: c1i_of_res_def conefield_of_res_def)
from res this obtain dx where dx: "(x, dx) \<in> c1i_of_res res"
by (auto simp: c1i_of_res_def)
from that(3)[OF \<open>res \<in> set _\<close>] have "correct_res res" by simp
from this[unfolded correct_res_def, rule_format, OF dx] res
obtain res' D where res': "x returns_to \<Sigma>"
"lorenz.return_time \<Sigma> differentiable at x within \<Sigma>\<^sub>l\<^sub>e"
"(lorenz.poincare_map \<Sigma> has_derivative D) (at x within \<Sigma>\<^sub>l\<^sub>e)"
"expansion res * norm dx \<le> norm (D dx)"
"res' \<in> return_of_res results res"
"(lorenz.poincare_map \<Sigma> x, D dx) \<in> c1_of_res res'"
"preexpansion res' * norm dx \<le> norm (D dx)"
by auto
show "x returns_to \<Sigma>" by fact
show "R x \<in> N" using res'
by (auto simp: R_def N_def N_def c1i_of_res_def c1_of_res_def in_return_of_resD
sourcei_of_res_def)
show "(R has_derivative DR x) (at x within \<Sigma>\<^sub>l\<^sub>e)"
apply (auto simp: R_def DR_def N_def c1_of_res_def in_return_of_resD)
apply (subst frechet_derivative_works[symmetric])
apply (rule differentiableI)
by fact
next
fix dx assume "dx \<in> \<CC> x"
then obtain res where res: "res \<in> set results" and dx: "(x, dx) \<in> c1_of_res res"
by (auto simp: \<CC>_def results_at_def c1_of_res_def )
then have dx: "(x, dx) \<in> c1i_of_res res"
using \<open>x \<in> N - _\<close>
by (auto simp: c1i_of_res_def sourcei_of_res_def c1_of_res_def \<Gamma>\<^sub>i_def)
from res dx have ne: "c1i_of_res res \<noteq> {}"
by (auto simp: c1_of_res_def conefield_of_res_def)
from that(3)[OF \<open>res \<in> set _\<close>] have "correct_res res" by simp
from that this[unfolded correct_res_def, rule_format, OF dx] res
obtain res' D where res': "x returns_to \<Sigma>" "x \<in> plane_of (Sctn (0, 0, 1) 27)"
"lorenz.return_time \<Sigma> differentiable at x within \<Sigma>\<^sub>l\<^sub>e"
"(lorenz.poincare_map \<Sigma> has_derivative D) (at x within \<Sigma>\<^sub>l\<^sub>e)"
"expansion res * norm dx \<le> norm (D dx)"
"res' \<in> return_of_res results res"
"(lorenz.poincare_map \<Sigma> x, D dx) \<in> c1_of_res res'"
"preexpansion res' * norm dx \<le> norm (D dx)"
by auto
have DRD: "DR x = D"
unfolding DR_def
apply (rule frechet_derivative_unique_within)
apply (subst frechet_derivative_works[symmetric])
apply (rule differentiableI)
apply fact
apply fact using \<open>x \<in> plane_of _\<close>
apply safe
subgoal for _ _ _ e
by (auto simp: \<Sigma>\<^sub>l\<^sub>e_def Basis_prod_def prod_eq_iff plane_of_def prod_eq_iff
inner_prod_def intro!: exI[where x="-e/2"])
done
have [intro, simp]: "res' \<in> results_at (lorenz.poincare_map \<Sigma> x)"
using res'
by (auto simp: c1_of_res_def results_at_def in_return_of_resD R_def
intro!: exI[where x=res'])
have [intro, simp]: "res \<in> results_at x"
using res dx
by (auto simp: c1i_of_res_def results_at_def sourcei_of_res_def)
show "DR x dx \<in> \<CC> (R x)"
unfolding DRD \<CC>_def
using res'
by (auto simp: c1_of_res_def R_def)
have "\<E> x * norm dx \<le> expansion res * norm dx"
by (rule mult_right_mono) (auto simp: \<E>_def)
also have "\<dots> \<le> norm (DR x dx)" unfolding DRD by fact
finally show "\<E> x * norm dx \<le> norm (DR x dx)" .
have "\<E>\<^sub>p (R x) * norm dx \<le> preexpansion res' * norm dx"
by (rule mult_right_mono) (auto simp: \<E>\<^sub>p_def R_def)
also have "\<dots> \<le> norm (DR x dx)" unfolding DRD by fact
finally show "\<E>\<^sub>p (R x) * norm dx \<le> norm (DR x dx)" .
qed
lemma check_line_core_correct:
"check_line_core pf m0 n0 True i \<le> SPEC (\<lambda>correct. correct \<longrightarrow> correct_res (results ! i))"
if [refine_vcg]: NF
unfolding check_line_core_def
supply [refine_vcg] = check_line_nres_c1_correct[le]
by refine_vcg
text \<open>The symmetric reduction\<close>
lemma source_of_res_mirror: "(x, y, z) \<in> source_of_res (mirror_result res) \<longleftrightarrow>
(-x, -y, z) \<in> source_of_res res"
by (cases res)
(auto simp: source_of_res_def ivl_of_res_def ivl_of_resultrect_def set_of_ivl_def)
lemma conefield_of_res_mirror[simp]: "(x, y, z) \<in> conefield_of_res (mirror_result res) \<longleftrightarrow>
(x, y, z) \<in> conefield_of_res res"
by (cases res) (auto simp: conefield_of_res_def ivl_of_res_def)
lemma c1_of_res_mirror: "((x, y, z), dx, dy, dz) \<in> c1_of_res (mirror_result res) \<longleftrightarrow>
((-x, -y, z), dx, dy, dz) \<in> c1_of_res res"
by (auto simp: c1_of_res_def source_of_res_mirror)
lemmas [simp] = lorenz_S(2)
lemma lorenz_S_idem[simp]: "lorenz_S (lorenz_S x) = (x::R3)"
by (auto simp: lorenz_S_def split_beta')
lemma lorenz_S_exivl[simp]:
"lorenz.existence_ivl0 (lorenz_S X) = lorenz.existence_ivl0 X"
using lorenz_S(1)[of _ X]
using lorenz_S(1)[of _ "lorenz_S X"]
by auto
lemma lorenz_S_zero[simp]: "lorenz_S x = 0 \<longleftrightarrow> (x::R3) = 0"
by (auto simp: lorenz_S_def split_beta' prod_eq_iff)
lemma lorenz_S_returns_toI[simp]:
"x returns_to (lorenz_S ` P) \<Longrightarrow> lorenz_S x returns_to P"
apply (auto simp: lorenz.returns_to_def)
subgoal premises prems for t
proof -
have " \<forall>\<^sub>F s in at_right 0. s < t"
using tendsto_ident_at \<open>0 < t\<close>
by (rule order_tendstoD)
then have " \<forall>\<^sub>F s in at_right 0. s \<in> lorenz.existence_ivl0 x"
unfolding eventually_at_filter
apply eventually_elim
using \<open>0 < t\<close> lorenz.closed_segment_subset_existence_ivl[OF prems(3)]
by (auto simp: closed_segment_eq_real_ivl subset_iff)
then show ?thesis using prems(1)
by eventually_elim force
qed
done
lemma lorenz_S_returns_to[simp]:
"lorenz_S x returns_to P \<longleftrightarrow> x returns_to (lorenz_S ` P)"
using lorenz_S_returns_toI[of P x] lorenz_S_returns_toI[of "lorenz_S ` P" "lorenz_S x"]
by (auto simp: image_image)
lemma lorenz_S_image_Sigma[simp]: "lorenz_S ` \<Sigma> = \<Sigma>"
apply (auto simp: \<Sigma>_def lorenz_S_def)
apply (rule image_eqI)
apply (rule lorenz_S_idem[symmetric])
apply (auto simp: \<Sigma>_def lorenz_S_def)
done
lemma linear_lorenz_S: "linear lorenz_S"
by unfold_locales (auto simp: lorenz_S_def)
lemma inj_lorenz_S: "inj_on (lorenz_S::R3 \<Rightarrow> _) G"
by (rule inj_onI) (auto simp: lorenz_S_def prod_eq_iff)
lemma lorenz_S_return_time:
"lorenz.return_time P (lorenz_S x) = lorenz.return_time (lorenz_S ` P) x"
if "x returns_to (lorenz_S ` P)" "closed P"
proof -
from lorenz.returns_toE[OF that(1)] obtain t0 t1
where f: "0 < t0" "t0 \<le> t1" " t1 \<in> lorenz.existence_ivl0 x"
"lorenz.flow0 x t1 \<in> lorenz_S ` P"
"\<And>t. 0 < t \<Longrightarrow> t < t0 \<Longrightarrow> lorenz.flow0 x t \<notin> lorenz_S ` P"
by auto
have [simp]: "lorenz.return_time (lorenz_S ` P) x \<in> lorenz.existence_ivl0 x"
by (auto intro!: that closed_injective_linear_image linear_lorenz_S
lorenz.return_time_exivl
inj_lorenz_S)
have c': "closed (lorenz_S ` P)"
by (auto intro!: that closed_injective_linear_image linear_lorenz_S
lorenz.return_time_exivl lorenz.return_time_pos
inj_lorenz_S)
show ?thesis
using f(1-4)
using lorenz.return_time_returns[OF that(1) c']
apply (intro lorenz.return_time_eqI)
apply (auto intro!: that closed_injective_linear_image linear_lorenz_S
lorenz.return_time_exivl lorenz.return_time_pos c'
inj_lorenz_S)
subgoal premises prems for a b c d e f g
proof -
have [simp]: "a \<in> lorenz.existence_ivl0 x"
using _ that(1)
apply (rule lorenz.less_return_time_imp_exivl)
using prems that(2) c'
by auto
have "lorenz.return_time (lorenz_S ` P) x \<le> a"
apply (rule lorenz.return_time_le)
using prems
apply (auto intro!: that closed_injective_linear_image linear_lorenz_S
lorenz.return_time_exivl lorenz.return_time_pos c'
inj_lorenz_S)
apply (rule image_eqI)
apply (rule lorenz_S_idem[symmetric])
by auto
then show ?thesis using prems
by simp
qed
done
qed
lemma lorenz_S_poincare_map:
"lorenz.poincare_map P (lorenz_S x) = lorenz_S (lorenz.poincare_map (lorenz_S ` P) x)"
if "x returns_to (lorenz_S ` P)" "closed P"
using that
unfolding lorenz.poincare_map_def
apply (auto simp: lorenz_S_return_time)
apply (subst lorenz_S)
by (auto intro!: lorenz.return_time_exivl that
closed_injective_linear_image linear_lorenz_S inj_lorenz_S)
lemma [continuous_intros]: "isCont (lorenz_S::_\<Rightarrow>R3) x"
"continuous_on (G::R3 set) lorenz_S"
by (auto simp:lorenz_S_def[abs_def] split_beta' continuous_intros)
lemma filtermap_lorenz_S_le: "filtermap lorenz_S (at x within lorenz_S ` P) \<le>(at (lorenz_S x::R3) within P)"\<comment> \<open>TODO: generalize!\<close>
unfolding at_within_def
apply (auto simp: intro!: antisym filtermap_inf[le] filtermap_inf[ge])
apply (rule inf.coboundedI1)
apply (subst filtermap_nhds_open_map)
apply (auto simp: intro!: invariance_of_domain inj_lorenz_S continuous_intros)
apply (rule inf.coboundedI2)
apply (auto simp: image_image )
apply (auto simp: lorenz_S_def split_beta')[]
done
lemma filtermap_lorenz_S_eq: "filtermap lorenz_S (at (x::R3) within lorenz_S ` P) = (at (lorenz_S x::R3) within P)"
apply (rule antisym)
using filtermap_lorenz_S_le[of "x" P]
apply simp
subgoal
proof -
have "filtermap lorenz_S (at (lorenz_S x) within P) \<le>
filtermap lorenz_S (filtermap lorenz_S (at x within lorenz_S ` P))"
using filtermap_lorenz_S_le[of "lorenz_S x" "lorenz_S ` P"]
by (auto simp: image_image filtermap_filtermap)
then show ?thesis
apply (subst (asm) filtermap_mono_strong)
by (auto simp: inj_lorenz_S)
qed
done
lemma norm_lorenz_S[simp]: "norm (lorenz_S x) = norm x"
by (auto simp: lorenz_S_def norm_prod_def split_beta')
lemma bl_lorenz_S: "bounded_linear (lorenz_S)"
by unfold_locales (auto simp: lorenz_S_def norm_prod_def intro!: exI[where x=1])
lemma filtermap_lorenz_S_eq_bot[simp]:
"filtermap (lorenz_S::R3\<Rightarrow>_) F = bot \<longleftrightarrow> F = bot"
apply (auto simp: )
apply (subst (asm) filtermap_bot[symmetric])
apply (subst (asm) filtermap_eq_strong)
by (auto simp: inj_lorenz_S)
lemma netlimit_filtermap[simp]:
"at x within X \<noteq> bot \<Longrightarrow> netlimit (filtermap lorenz_S (at x within X)) = lorenz_S (x::R3)"
apply (rule tendsto_Lim)
unfolding filterlim_filtermap
apply simp
by (auto intro!: tendsto_eq_intros simp: split_beta' lorenz_S_def[abs_def])
lemma lorenz_S_halfspace [simp]: "lorenz_S ` \<Sigma>\<^sub>l\<^sub>e = \<Sigma>\<^sub>l\<^sub>e"
apply (auto simp: \<Sigma>\<^sub>l\<^sub>e_def lorenz_S_def[abs_def])
apply (rule image_eqI)
apply auto
apply (rule sym)
apply (rule minus_minus)
apply (rule minus_minus[symmetric])
done
lemma closure_Sigma_le_eq: "closure \<Sigma>\<^sub>l\<^sub>e = \<Sigma>\<^sub>l\<^sub>e"
proof (rule closure_closed)
have "\<Sigma>\<^sub>l\<^sub>e = {x. x \<bullet> (0, 0, 1) \<le> 27}"
by (auto simp: \<Sigma>\<^sub>l\<^sub>e_def )
also have "closed \<dots>"
by (rule closed_halfspace_component_le)
finally show "closed \<Sigma>\<^sub>l\<^sub>e" .
qed
lemma closure_Sigma_le[simp]: "closure (\<Sigma>\<^sub>l\<^sub>e - {x}) = \<Sigma>\<^sub>l\<^sub>e"
proof (cases "x \<in> \<Sigma>\<^sub>l\<^sub>e")
case that: True
have "closure \<Sigma>\<^sub>l\<^sub>e \<subseteq> closure (insert x (\<Sigma>\<^sub>l\<^sub>e - {x}))" by (rule closure_mono) auto
also have "\<dots> = insert x (closure (\<Sigma>\<^sub>l\<^sub>e - {x}))"
apply (subst closure_insert) by simp
also
have "x \<in> closure (\<Sigma>\<^sub>l\<^sub>e - {x})"
apply (rule closed_sequentially[where f="\<lambda>n. x - (0, 0, inverse (Suc n))"])
apply (rule closed_closure)
subgoal
apply (auto simp: ) apply (rule subsetD) apply (rule closure_subset)
using that
apply (auto simp: \<Sigma>\<^sub>l\<^sub>e_def prod_eq_iff)
apply (rule order_trans)
apply (rule diff_right_mono)
apply (assumption)
apply simp
done
subgoal
apply (rule tendsto_eq_intros)
apply (rule tendsto_intros)
apply (rule tendsto_intros)
apply (rule tendsto_intros)
apply (rule tendsto_intros)
apply (rule tendsto_intros)
apply (rule LIMSEQ_inverse_real_of_nat)
by (auto simp: prod_eq_iff)
done
then have "insert x (closure (\<Sigma>\<^sub>l\<^sub>e - {x})) \<subseteq> closure (\<Sigma>\<^sub>l\<^sub>e - {x})"
by auto
finally have "closure \<Sigma>\<^sub>l\<^sub>e \<subseteq> closure (\<Sigma>\<^sub>l\<^sub>e - {x})" .
moreover
have "closure (\<Sigma>\<^sub>l\<^sub>e - {x}) \<subseteq> closure (\<Sigma>\<^sub>l\<^sub>e)"
by (rule closure_mono) auto
ultimately have "closure (\<Sigma>\<^sub>l\<^sub>e - {x}) = closure (\<Sigma>\<^sub>l\<^sub>e)"
by simp
also have "\<dots> = \<Sigma>\<^sub>l\<^sub>e"
by (rule closure_Sigma_le_eq)
finally show ?thesis .
next
case False
then show ?thesis
apply simp
apply (rule closure_Sigma_le_eq)
done
qed
lemma lorenz_S_return_time_has_derivative:
assumes "(lorenz.return_time \<Sigma> has_derivative D) (at x within \<Sigma>\<^sub>l\<^sub>e)"
and "lorenz.returns_to \<Sigma> x" and "x \<in> \<Sigma>\<^sub>l\<^sub>e"
shows "(lorenz.return_time \<Sigma> has_derivative D o lorenz_S) (at (lorenz_S x) within \<Sigma>\<^sub>l\<^sub>e)"
proof -
have [simp]: "\<not>trivial_limit (at x within \<Sigma>\<^sub>l\<^sub>e)"
unfolding at_within_eq_bot_iff
using assms
by simp
interpret bounded_linear "lorenz_S::R3\<Rightarrow>_" by (rule bl_lorenz_S)
have "\<forall>\<^sub>F x in at x within \<Sigma>\<^sub>l\<^sub>e. (x::R3) returns_to \<Sigma>"
by (blast intro: lorenz.eventually_returns_to_continuousI has_derivative_continuous assms)
then have "\<forall>\<^sub>F y in at x within \<Sigma>\<^sub>l\<^sub>e.
inverse (norm (y - x)) * (lorenz.return_time \<Sigma> y - lorenz.return_time \<Sigma> x - D (y - x))
= inverse (norm (lorenz_S y - lorenz_S x)) *
(lorenz.return_time \<Sigma> (lorenz_S y) - lorenz.return_time \<Sigma> (lorenz_S x) - D (y - x))"
by eventually_elim (auto simp: lorenz_S_return_time assms diff[symmetric])
then show ?thesis
using assms
apply (subst filtermap_lorenz_S_eq[symmetric])
apply (auto simp: has_derivative_def filterlim_filtermap)
unfolding o_def
apply (rule bounded_linear_compose, assumption, rule bl_lorenz_S)
unfolding diff lorenz_S_idem
apply (auto simp: Lim_ident_at)
apply (blast intro: Lim_transform_eventually)
done
qed
lemma lorenz_S_return_time_differentiable:
"lorenz.return_time \<Sigma> differentiable at (lorenz_S x) within \<Sigma>\<^sub>l\<^sub>e"
if "lorenz.return_time \<Sigma> differentiable at x within \<Sigma>\<^sub>l\<^sub>e"
"lorenz.returns_to \<Sigma> x" "x \<in> \<Sigma>\<^sub>l\<^sub>e"
proof -
from that obtain D where "(lorenz.return_time \<Sigma> has_derivative D) (at x within \<Sigma>\<^sub>l\<^sub>e)"
by (auto simp: differentiable_def)
then have "(lorenz.return_time \<Sigma> has_derivative D o lorenz_S) (at (lorenz_S x) within \<Sigma>\<^sub>l\<^sub>e)"
by (rule lorenz_S_return_time_has_derivative) fact+
then show ?thesis
by (auto simp: differentiable_def)
qed
lemma lorenz_S_has_derivative:
"(lorenz_S has_derivative lorenz_S) (at (x::R3) within X)"
by (auto simp: lorenz_S_def[abs_def] split_beta' intro!: derivative_eq_intros)
lemma lorenz_S_poincare_map_has_derivative:
assumes "(lorenz.poincare_map \<Sigma> has_derivative D) (at x within \<Sigma>\<^sub>l\<^sub>e)"
"(lorenz.return_time \<Sigma> has_derivative Dr) (at x within \<Sigma>\<^sub>l\<^sub>e)"
"lorenz.returns_to \<Sigma> x" "x \<in> \<Sigma>\<^sub>l\<^sub>e"
shows "(lorenz.poincare_map \<Sigma> has_derivative lorenz_S o D o lorenz_S) (at (lorenz_S x) within \<Sigma>\<^sub>l\<^sub>e)"
proof -
have [simp]: "\<not>trivial_limit (at x within \<Sigma>\<^sub>l\<^sub>e)"
unfolding at_within_eq_bot_iff
using assms
by simp
interpret bounded_linear "lorenz_S::R3\<Rightarrow>_" by (rule bl_lorenz_S)
have "\<forall>\<^sub>F x in at x within \<Sigma>\<^sub>l\<^sub>e. (x::R3) returns_to \<Sigma>"
by (blast intro: lorenz.eventually_returns_to_continuousI has_derivative_continuous assms)
then have "\<forall>\<^sub>F y in at x within \<Sigma>\<^sub>l\<^sub>e.
(lorenz_S (lorenz.poincare_map \<Sigma> y) - lorenz_S (lorenz.poincare_map \<Sigma> x) - lorenz_S (D (y - x))) /\<^sub>R
norm (y - x)
= (lorenz.poincare_map \<Sigma> (lorenz_S y) - lorenz.poincare_map \<Sigma> (lorenz_S x) - lorenz_S (D (y - x))) /\<^sub>R
norm (lorenz_S y - lorenz_S x)"
by eventually_elim (auto simp: lorenz_S_return_time lorenz_S_poincare_map assms diff[symmetric])
then show ?thesis
using has_derivative_compose[OF assms(1) lorenz_S_has_derivative] assms
apply (subst filtermap_lorenz_S_eq[symmetric])
apply (auto simp: has_derivative_def filterlim_filtermap)
unfolding o_def
apply (rule bounded_linear_compose, rule bl_lorenz_S)
apply (rule bounded_linear_compose, assumption, rule bl_lorenz_S)
unfolding diff lorenz_S_idem
apply (auto simp: Lim_ident_at)
apply (blast intro: Lim_transform_eventually)
done
qed
lemma [simp]: "expansion (mirror_result res) = expansion res"
by (cases res) auto
lemma lorenz_S_on_plane: "lorenz_S (dx, dy, 0::real) = - (dx, dy, 0)"
by (auto simp: lorenz_S_def )
lemma mirror_result_idem[simp]: "mirror_result (mirror_result x) = x"
by (cases x) (auto simp: mirror_result_def)
lemma mirror_in_set: "x \<in> set results \<Longrightarrow> mirror_result x \<in> set results"
by (auto simp: results_def symmetrize_def)
lemma mirror_result_in:
"mirror_result res2 \<in> return_of_res results (mirror_result res)"
if "res2 \<in> return_of_res results res"
proof -
from that have "res2 \<in> set results" by (rule in_return_of_resD)
from mirror_in_set[OF this] have "mirror_result res2 \<in> set results" .
then show ?thesis
apply (cases res2; cases res)
using that
by (auto simp: return_of_res_def get_results_def)
qed
lemma in_source_of_res_mirrorI:
"(x::R3) \<in> source_of_res (mirror_result (r))" if "lorenz_S x \<in> source_of_res r"
using that
apply (cases r; cases x)
by (auto simp: source_of_res_def set_of_ivl_def ivl_of_res_def lorenz_S_def
less_eq_prod_def ivl_of_resultrect_def)
lemma conefield_of_res_mirror_simp[simp]: "conefield_of_res (mirror_result res2) = conefield_of_res res2"
by (cases res2) (auto simp: conefield_of_res_def)
lemma lorenz_minus_planeI: "lorenz_S (- x) = x" if "snd (snd (x::R3)) = 0"
using that
by (auto simp: lorenz_S_def split_beta' prod_eq_iff)
lemma preexpansion_mirror_result[simp]: "preexpansion (mirror_result res2) = preexpansion res2"
by (cases res2) (auto simp: )
lemma lorenz_S_tendsto_0I: "(lorenz.flow0 (lorenz_S x) \<longlongrightarrow> 0) at_top"
if "{0..} \<subseteq> lorenz.existence_ivl0 x" "(lorenz.flow0 x \<longlongrightarrow> 0) at_top"
proof (rule Lim_transform_eventually)
have "\<forall>\<^sub>F s in at_top. (s::real) \<ge> 0"
using eventually_ge_at_top by blast
then show "\<forall>\<^sub>F s in at_top. lorenz_S (lorenz.flow0 x s) = lorenz.flow0 (lorenz_S x) s"
by eventually_elim (use that in auto)
show "((\<lambda>s. lorenz_S (lorenz.flow0 x s)) \<longlongrightarrow> 0) at_top"
unfolding Zfun_def[symmetric] by (rule bounded_linear.tendsto_zero[OF bl_lorenz_S that(2)])
qed
lemma lorenz_S_tendsto_0_iff:
"(lorenz.flow0 (lorenz_S x) \<longlongrightarrow> 0) at_top \<longleftrightarrow> (lorenz.flow0 x \<longlongrightarrow> 0) at_top"
if "{0..} \<subseteq> lorenz.existence_ivl0 x"
using lorenz_S_tendsto_0I[of x, OF that] lorenz_S_tendsto_0I[of "lorenz_S x"] that
by auto
lemma lorenz_S_eq_iff[simp]: "lorenz_S y = lorenz_S x \<longleftrightarrow> y = x" for x y::"real*real*real"
by (auto simp: lorenz_S_def split: prod.splits)
lemma lorenz_S_\<Gamma>: "lorenz_S x \<in> \<Gamma> \<longleftrightarrow> x \<in> \<Gamma>"
apply (auto simp: \<Gamma>_def lorenz_S_tendsto_0_iff )
subgoal for t
apply (auto simp: dest!: spec[where x=t])
apply (subst (asm) lorenz_S) apply auto
apply (subst (asm) (2) lorenz_S_image_Sigma[symmetric])
by (simp del: lorenz_S_image_Sigma)
subgoal for t
apply (auto simp: dest!: spec[where x=t])
apply (subst (asm) lorenz_S) apply auto
apply (subst (asm) lorenz_S_image_Sigma[symmetric])
apply (auto simp del: lorenz_S_image_Sigma)
done
done
lemma sourcei_of_res_mirror: "(x, y, z) \<in> sourcei_of_res (mirror_result res) \<longleftrightarrow>
(-x, -y, z) \<in> sourcei_of_res res"
using lorenz_S_\<Gamma>[of "(x, y, z)"]
by (cases res)
(auto simp: source_of_res_def sourcei_of_res_def ivl_of_res_def ivl_of_resultrect_def
set_of_ivl_def \<Gamma>\<^sub>i_def lorenz_S_def)
lemma c1i_of_res_mirror: "((x, y, z), dx, dy, dz) \<in> c1i_of_res (mirror_result res) \<longleftrightarrow>
((-x, -y, z), dx, dy, dz) \<in> c1i_of_res res"
by (auto simp: c1i_of_res_def sourcei_of_res_mirror)
lemma correct_res_mirror_result:
"correct_res (mirror_result res)" if "correct_res res"
unfolding correct_res_def
proof (clarsimp simp add: c1i_of_res_mirror, goal_cases)
case (1 x y z dx dy dz)
then have 1: "(lorenz_S (x, y, z), dx, dy, dz) \<in> c1i_of_res res"
by (auto simp: lorenz_S_def)
from that[unfolded correct_res_def, rule_format, OF 1, simplified]
have "(lorenz_S (x, y, z)) \<in> plane_of (Sctn (0, 0, 1) 27)"
"(dx, dy, dz) \<in> plane_of (Sctn (0, 0, 1) 0)"
by auto
then have plane: "(x, y, z) \<in> plane_of (Sctn (0, 0, 1) 27)"
"(dx, dy, dz) \<in> plane_of (Sctn (0, 0, 1) 0)"
by (auto simp: plane_of_def lorenz_S_def)
then show ?case
proof (clarsimp, goal_cases)
case mem: 1
with that[unfolded correct_res_def, rule_format, OF 1, simplified]
obtain D res2 where D:
"lorenz_S (x, y, z) returns_to \<Sigma>"
"lorenz.return_time \<Sigma> differentiable at (lorenz_S (x, y, z)) within \<Sigma>\<^sub>l\<^sub>e"
"(lorenz.poincare_map \<Sigma> has_derivative D) (at (lorenz_S (x, y, z)) within \<Sigma>\<^sub>l\<^sub>e)"
"expansion res * norm (dx, dy, dz) \<le> norm (D (dx, dy, dz))"
"res2 \<in> return_of_res results res"
"(lorenz.poincare_map \<Sigma> (lorenz_S (x, y, z)), D (dx, dy, dz)) \<in> c1_of_res res2"
"preexpansion res2 * norm (dx, dy, dz) \<le> norm (D (dx, dy, dz))"
by auto
from plane have S_le: "lorenz_S (x, y, z) \<in> \<Sigma>\<^sub>l\<^sub>e"
by (auto simp: \<Sigma>\<^sub>l\<^sub>e_def plane_of_def lorenz_S_def)
interpret linear D by (rule has_derivative_linear; fact)
have ret: "(x, y, z) returns_to \<Sigma>" using D(1) lorenz_S_returns_to by simp
moreover have "lorenz.return_time \<Sigma> differentiable at (x, y, z) within \<Sigma>\<^sub>l\<^sub>e"
using lorenz_S_return_time_differentiable[OF D(2) D(1) S_le] by auto
moreover from D obtain Dr where Dr: "(lorenz.return_time \<Sigma> has_derivative Dr) (at (lorenz_S (x, y, z)) within \<Sigma>\<^sub>l\<^sub>e)"
by (auto simp: differentiable_def)
let ?D = "lorenz_S \<circ> D \<circ> lorenz_S"
have "(lorenz.poincare_map \<Sigma> has_derivative ?D) (at (x, y, z) within \<Sigma>\<^sub>l\<^sub>e)"
using lorenz_S_poincare_map_has_derivative[OF D(3) Dr D(1) S_le]
by auto
moreover
from plane have [simp]: "dz = 0" by (auto simp: plane_of_def)
have "expansion (mirror_result res) * norm (dx, dy, dz) \<le> norm (?D (dx, dy, dz))"
using D(4) apply (auto simp: )
unfolding lorenz_S_on_plane neg
by simp
moreover have \<open>mirror_result res2 \<in> return_of_res results (mirror_result res)\<close>
using D(5) by (rule mirror_result_in)
moreover have "(lorenz.poincare_map \<Sigma> (x, y, z), ?D (dx, dy, dz)) \<in> c1_of_res (mirror_result res2)"
using D(6) apply (subst (asm) lorenz_S_poincare_map)
apply auto apply fact
apply (auto simp: c1_of_res_def in_source_of_res_mirrorI)
unfolding lorenz_S_on_plane neg
apply (subst lorenz_minus_planeI)
apply (auto simp: conefield_of_res_def conefield_alt_def cone_hull_expl
in_segment tangent_of_deg_def)
done
moreover have "preexpansion (mirror_result res2) * norm (dx, dy, dz) \<le> norm (?D (dx, dy, dz))"
using D(7) apply (auto simp: )
unfolding lorenz_S_on_plane neg
by simp
ultimately show ?case
by (force intro!: exI[where x = ?D] bexI[where x="mirror_result res2"])
qed
qed
lemma reduce_lorenz_symmetry: "Ball (set results) correct_res"
if "Ball (set coarse_results) correct_res"
using that
by (auto simp: results_def symmetrize_def intro!: correct_res_mirror_result)
end
subsection \<open>Code Generation\<close>
definition [code_abbrev]: "my_divide_integer (i::integer) (j::integer) = i div j"
code_printing constant my_divide_integer \<rightharpoonup> (SML) "IntInf.div/ (_,/ _)"
subsection \<open>Tuning code equations\<close>
definition mult_twopow_int::"int \<Rightarrow> int \<Rightarrow> int" where "mult_twopow_int x n = x * (power_int 2 n)"
definition div_twopow_int :: "int \<Rightarrow> int \<Rightarrow> int" where "div_twopow_int x n = x div (power_int 2 n)"
context includes integer.lifting begin
lift_definition mult_twopow_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer" is mult_twopow_int .
lift_definition div_twopow_integer :: "integer \<Rightarrow> integer \<Rightarrow> integer" is div_twopow_int .
end
lemma compute_float_round_down[code]:
"float_round_down prec (Float m e) =
(let d = bitlen \<bar>m\<bar> - int prec - 1 in
if 0 < d then Float (div_twopow_int m d) (e + d)
else Float m e)"
including float.lifting
using Float.compute_float_down[of "Suc prec - bitlen \<bar>m\<bar> - e" m e, symmetric]
by transfer
(auto simp add: field_simps abs_mult log_mult bitlen_alt_def truncate_down_def
div_twopow_int_def power_int_def
cong del: if_weak_cong)
lemma compute_float_plus_down[code]:
fixes p::nat and m1 e1 m2 e2::int
shows "float_plus_down p (Float m1 e1) (Float m2 e2) =
(if m1 = 0 then float_round_down p (Float m2 e2)
else if m2 = 0 then float_round_down p (Float m1 e1)
else
(if e1 \<ge> e2 then
(let k1 = Suc p - nat (bitlen \<bar>m1\<bar>) in
if bitlen \<bar>m2\<bar> > e1 - e2 - k1 - 2
then float_round_down p ((Float m1 e1) + (Float m2 e2))
else float_round_down p (Float (mult_twopow_int m1 (int k1 + 2) + sgn m2) (e1 - int k1 - 2)))
else float_plus_down p (Float m2 e2) (Float m1 e1)))"
using Float.compute_float_plus_down[of p m1 e1 m2 e2]
by (auto simp: mult_twopow_int_def Let_def power_int_def nat_add_distrib)
subsection \<open>Codegen\<close>
definition "is_dRETURN_True x = (case x of dRETURN b \<Rightarrow> b | _ \<Rightarrow> False)"
definition "file_output_option s f =
(case s of None \<Rightarrow> f None
| Some s \<Rightarrow> file_output (String.implode s) (\<lambda>pf. f (Some pf)))"
definition "check_line_lookup_out s m0 n0 c1 i =
is_dRETURN_True (file_output_option s (\<lambda>pf. check_line_core_impl pf m0 n0 c1 i))"
fun alternating where "alternating [] xs = xs"
| "alternating xs [] = xs"
| "alternating (x#xs) (y#ys) = x#y#alternating xs ys"
definition "ordered_lines = alternating (rev [0..<222]) ([222..<400])"
\<comment> \<open>the hard ones ``first'', potentially useless due to nondeterministic \<open>Parallel.map\<close>\<close>
definition "parallel_check filenameo m n c1 ns =
Parallel.forall (\<lambda>i.
let
_ = print (String.implode (''# Starting '' @ show i @ ''\<newline>''));
b =
check_line_lookup_out (map_option (\<lambda>f. f @ show i) filenameo)
(Some m) (Some n) c1 i;
_ = if b
then print (String.implode (''# Success: '' @ show i @ ''\<newline>''))
else print (String.implode (''# Failed: '' @ show i @ ''\<newline>''))
in b
) ns"
ML \<open>val check_line = @{computation_check
terms:
Trueprop
parallel_check
ordered_lines
check_line_core_impl
check_line_lookup_out
(* bool *)
True False
(* num *)
Num.One Num.Bit0 Num.Bit1
(* nat *)
Suc "0::nat" "1::nat" "numeral::num\<Rightarrow>nat"
(* int / integer*)
"numeral::num\<Rightarrow>int"
"numeral::num\<Rightarrow>integer"
"uminus::_\<Rightarrow>int"
"uminus::_\<Rightarrow>integer"
int_of_integer integer_of_int
"0::int"
"1::int"
(* Pairs *)
"Pair::_ \<Rightarrow> _\<Rightarrow> (real list \<times> real list)"
"Pair::_\<Rightarrow>_\<Rightarrow>(real list \<times> real list) \<times> real list sctn"
"Pair::_\<Rightarrow>_\<Rightarrow>((real list \<times> real list) \<times> real list sctn) list \<times> real aform reach_options"
(* Option *)
"None::nat option"
"Some::_\<Rightarrow>nat option"
"None::string option"
"Some::_\<Rightarrow>string option"
(* Lists *)
"Nil::real list"
"Cons::_\<Rightarrow>_\<Rightarrow>real list"
"Nil::nat list"
"Cons::_\<Rightarrow>_\<Rightarrow>nat list"
"Nil::real aform list"
"Cons::_\<Rightarrow>_\<Rightarrow>real aform list"
"Nil::((real list \<times> real list) \<times> real list sctn) list"
"Cons::_\<Rightarrow>_\<Rightarrow>((real list \<times> real list) \<times> real list sctn) list"
"Nil::(((real list \<times> real list) \<times> real list sctn) list \<times> real aform reach_options)list"
"Cons::_\<Rightarrow>_\<Rightarrow>(((real list \<times> real list) \<times> real list sctn) list \<times> real aform reach_options)list"
(* String *)
String.Char
String.implode "Cons::char \<Rightarrow> char list \<Rightarrow> char list" "Nil::char list"
(* float *)
Float float_of_int float_of_nat
(* real *)
"numeral::num\<Rightarrow>real" "real_of_float" "(/)::real\<Rightarrow>real\<Rightarrow>real" "uminus::real\<Rightarrow>_"
real_divl real_divr
real_of_int
(* section *)
"Sctn::_\<Rightarrow>_\<Rightarrow>real list sctn"
(* aform *)
"aforms_of_ivls::_\<Rightarrow>_\<Rightarrow>real aform list"
(* input *)
coarse_results
(* modes *)
xsec xsec' ysec ysec' zsec zsec' zbucket
lookup_mode
ro
ro_outer
mode_outer
(* unit *)
"()"
}\<close>
lemma is_dRETURN_True_iff[simp]: "is_dRETURN_True x \<longleftrightarrow> (x = dRETURN True)"
by (auto simp: is_dRETURN_True_def split: dres.splits)
lemma check_line_core_impl_True:
"check_line_core_impl pfo m n True i = dRETURN True \<Longrightarrow> NF \<Longrightarrow> correct_res (results ! i)"
apply (cases "check_line_core_impl pfo m n True i")
using check_line_core_correct[of pfo m n i]
check_line_core_impl.refine[of pfo pfo True True i i m m n n]
apply (auto simp: nres_rel_def)
apply (drule order_trans[where y="check_line_core pfo m n True i"])
apply assumption
by auto
lemma check_line_lookup_out: "correct_res (results ! i)"
if "\<exists>s m n. check_line_lookup_out s m n True i" NF
using that
by (auto simp: check_line_lookup_out_def file_output_iff check_line_core_impl_True
file_output_option_def split: dres.splits option.splits)
definition "check_lines c1 ns = list_all (\<lambda>i. \<exists>s m n. check_line_lookup_out s m n c1 i) ns"
lemma check_linesI:
"check_lines c1 ns"
if "parallel_check s m n c1 ns"
using that
by (auto simp: parallel_check_def check_lines_def list_all_iff)
subsection \<open>Automate generation of lemmas\<close>
lemma length_coarse_results[simp]: "length coarse_results = 400"
by (simp add: coarse_results_def)
lemma correct_res_coarse_resultsI:
"correct_res (results ! i) \<Longrightarrow> i < 400 \<Longrightarrow> correct_res (coarse_results ! i)"
by (auto simp: results_def symmetrize_def nth_append)
lemma Ball_coarseI: "Ball (set coarse_results) correct_res"
if NF "check_lines True xs" "set xs = {..<400}"
using that
by (force simp: check_lines_def list_all_iff in_set_conv_nth
intro!: correct_res_coarse_resultsI check_line_lookup_out)
ML \<open>map_option (using_master_directory_term @{context}) (SOME "a")\<close>
ML \<open>
fun mk_optionT ty = Type (@{type_name "option"}, [ty])
fun mk_None ty = Const (@{const_name "None"}, mk_optionT ty)
fun mk_Some ty x = Const (@{const_name "Some"}, ty --> mk_optionT ty) $ x
fun mk_option ty _ NONE = mk_None ty
| mk_option ty f (SOME x) = mk_Some ty (f x)
fun check_lines_tac' s m n ctxt =
resolve_tac ctxt
[Thm.instantiate ([],
[("s", @{typ "string option"}, mk_option @{typ string} (using_master_directory_term ctxt) s),
("m", @{typ nat}, HOLogic.mk_nat m),
("n", @{typ nat}, HOLogic.mk_nat n)]
|> map (fn (s, ty, t) => (((s, 0), ty), Thm.cterm_of ctxt t)))
@{thm check_linesI}]
THEN' CONVERSION (check_line ctxt)
THEN' resolve_tac ctxt @{thms TrueI}
\<close>
method_setup parallel_check = \<open>
Scan.lift (Parse.maybe Parse.string) -- Scan.lift Parse.nat -- Scan.lift Parse.nat
>> (fn ((s, m), n) => fn ctxt => SIMPLE_METHOD' (check_lines_tac' s m n ctxt))
\<close>
lemma lorenz_bounds_lemma_asym:
"\<forall>x \<in> N - \<Gamma>. x returns_to \<Sigma>"
"R ` (N - \<Gamma>) \<subseteq> N"
"\<forall>x \<in> N - \<Gamma>. (R has_derivative DR x) (at x within \<Sigma>\<^sub>l\<^sub>e)"
"\<forall>x \<in> N - \<Gamma>. DR x ` \<CC> x \<subseteq> \<CC> (R x)"
"\<forall>x \<in> N - \<Gamma>. \<forall>c \<in> \<CC> x. norm (DR x c) \<ge> \<E> x * norm c"
"\<forall>x \<in> N - \<Gamma>. \<forall>c \<in> \<CC> x. norm (DR x c) \<ge> \<E>\<^sub>p (R x) * norm c"
if NF "Ball (set results) correct_res"
using that
by (auto intro!: lorenz_bounds_lemma)
end
diff --git a/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis.thy b/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis.thy
--- a/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis.thy
+++ b/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis.thy
@@ -1,1701 +1,1701 @@
theory Refine_Reachability_Analysis
imports
Abstract_Reachability_Analysis
Refine_Rigorous_Numerics
begin
lemma isDERIV_simps[simp]:
"isDERIV i 1 xs" "isDERIV i 0 xs"
"isDERIV i (a + b) xs \<longleftrightarrow> isDERIV i a xs \<and> isDERIV i b xs"
"isDERIV i (a - b) xs \<longleftrightarrow> isDERIV i a xs \<and> isDERIV i b xs"
"isDERIV i (a * b) xs \<longleftrightarrow> isDERIV i a xs \<and> isDERIV i b xs"
"isDERIV i (a / b) xs \<longleftrightarrow> isDERIV i a xs \<and> isDERIV i b xs \<and> interpret_floatarith b xs \<noteq> 0"
"isDERIV i (-a) xs = isDERIV i a xs"
by (auto simp: one_floatarith_def zero_floatarith_def plus_floatarith_def minus_floatarith_def
times_floatarith_def divide_floatarith_def uminus_floatarith_def)
lemma list_of_eucl_of_env:
assumes [simp]: "length xs = DIM('a)"
shows "(list_of_eucl (eucl_of_env xs vs::'a)) = (map (\<lambda>i. vs ! (xs ! i)) [0..<DIM('a::executable_euclidean_space)])"
by (auto intro!: nth_equalityI simp: eucl_of_env_def eucl_of_list_inner)
context approximate_sets_ode
begin
lemma interpret_ode_fa[simp]:
assumes [simp]: "length xs = DIM('a::executable_euclidean_space)" "length vs \<ge> DIM('a)" "length ode_e = DIM('a)"
and mV: "max_Var_floatariths ode_e \<le> DIM('a)"
shows "(eucl_of_list (interpret_floatariths (ode_fa xs) vs)::'a) = ode (eucl_of_list (interpret_floatariths xs vs))"
unfolding ode_fa_def
apply (auto intro!: euclidean_eqI[where 'a='a] simp: eucl_of_list_inner ode_def)
apply (subst interpret_floatarith_subst_floatarith[where D="length vs"])
apply (auto intro!: max_Var_floatarith_le_max_Var_floatariths_nth[le]
mV[le])
apply (rule interpret_floatarith_max_Var_cong)
apply (drule max_Var_floatariths_lessI) apply simp
apply (drule less_le_trans[OF _ mV])
apply auto
apply (subst nth_map)
apply simp using assms(2) order.strict_trans2 apply blast
apply (subst nth_upt) apply simp
apply (rule less_le_trans, assumption, simp)
apply auto
done
lemma length_ode_fa[simp]: "length (ode_fa xs) = length ode_e"
by (auto simp: ode_fa_def)
lemma max_Var_ode_fa[le]:
assumes "max_Var_floatariths ode_e \<le> length xs"
shows "max_Var_floatariths (ode_fa xs) \<le> max_Var_floatariths xs"
by (auto simp: ode_fa_def intro!: assms max_Var_floatariths_subst_floatarith_le)
lemma max_Var_floatariths_ode_d_expr[le]:
"max_Var_floatariths ode_e \<le> d \<Longrightarrow> d > 0 \<Longrightarrow>
max_Var_floatariths (ode_d_expr d m) \<le> 3 * d"
by (auto simp: ode_d_expr_def
intro!: max_Var_floatarith_FDERIV_n_floatariths[le]
max_Var_floatarith_FDERIV_floatariths[le])
lemma interpret_ode_d_fa:
assumes FDERIV: "(eucl_of_list (interpret_floatariths xs vs)::'a::executable_euclidean_space) \<in> Csafe"
assumes [simp]: "length ds = DIM('a)" "length xs = DIM('a)"
notes [simp] = safe_length[OF FDERIV]
shows "(eucl_of_list (interpret_floatariths (ode_d_fa n xs ds) vs)::'a) =
ode_d n (eucl_of_list (interpret_floatariths xs vs)) (eucl_of_list (interpret_floatariths ds vs))
(eucl_of_list (interpret_floatariths ds vs))"
apply (transfer fixing: xs ds vs n)
using FDERIV apply (auto simp del: isnFDERIV.simps simp: interpret_floatariths_append)
apply (auto simp add: list_of_eucl_of_env ode_def
ode_d_raw_def eucl_of_list_inner
intro!: euclidean_eqI[where 'a='a])
apply (auto simp: ode_d_fa_def )
apply (subst interpret_floatarith_subst_floatarith[OF max_Var_floatarith_le_max_Var_floatariths_nth], simp)
apply (rule interpret_floatarith_max_Var_cong)
subgoal premises prems for b i
proof -
from prems have i: "i < max_Var_floatariths (ode_d_expr DIM('a) n)"
by (auto dest!: max_Var_floatariths_lessI)
also have "\<dots> \<le> 3 * DIM('a)"
by (auto intro!: max_Var_floatariths_ode_d_expr safe_max_Var[OF FDERIV])
finally have "i < 3 * DIM('a)" .
then show ?thesis
using prems i
by (auto simp: nth_append)
qed
done
lemma safe_at_within: "x \<in> Csafe \<Longrightarrow> at x = at x within Csafe"
by (subst (2) at_within_open) (auto simp: open_safe)
lemma ivlflowsD:
assumes "ivlflows stops stopcont trap rsctn" "ivl \<subseteq> \<Union>(plane_of ` stops) \<times> UNIV "
shows "ivl \<subseteq> (snd (stopcont ivl))"
"(fst (stopcont ivl)) \<subseteq> sbelow_halfspace rsctn \<times> UNIV"
"fst (stopcont ivl) \<subseteq> snd (stopcont ivl)"
"(snd (stopcont ivl)) \<subseteq> sbelow_halfspace rsctn \<times> UNIV"
"flowsto (ivl) {0..} ((snd (stopcont ivl))) ((fst (stopcont ivl)) \<union> trap)"
using assms(1)[unfolded ivlflows_def, rule_format, OF assms(2)]
by auto
lemma ivlflows_flowsto:
assumes "ivlflows stops stopcont trap rsctn" "ivl \<subseteq> \<Union>(plane_of ` stops) \<times> UNIV"
assumes "stopcont ivl = (x, y)"
shows "flowsto (ivl) {0..} y (x \<union> trap)"
using ivlflowsD[OF assms(1,2)] assms(3)
by auto
lemma ivlflows_emptyI: "ivlflows {} (\<lambda>x. (x, x)) J K"
by (auto simp: ivlflows_def set_of_ivl_def)
lemma plane_of_neg[simp]: "plane_of (Sctn (- normal sctn) (- pstn sctn)) = plane_of sctn"
by (auto simp: plane_of_def)
lemmas safe_max_Var_le[intro] = safe_max_Var[le]
lemmas [simp] = safe_length
lemma continuous_on_ode_d2: "continuous_on (Csafe) ode_d2"
proof -
have isn: "isnFDERIV DIM('a) ode_e [0..<DIM('a)] [DIM('a)..<2 * DIM('a)] (list_of_eucl x @ list_of_eucl i)
(Suc (Suc 0))"
if "x \<in> Csafe"
for x i::'a
by (rule safe_isnFDERIV) fact
have "continuous_on (Csafe::'a set) (\<lambda>x. ode_d_raw (Suc 0) x i j)"
if "i \<in> Basis" "j \<in> Basis" for i j
apply (rule has_derivative_continuous_on)
apply (auto simp: ode_d_raw_def at_within_open[OF _ open_safe])
apply (rule interpret_floatarith_FDERIV_floatariths_append)
apply (auto simp: ode_d_expr_def
intro!: isDERIV_FDERIV_floatariths safe_isFDERIV_append that isFDERIV_map_Var
safe_max_Var_le
max_Var_floatarith_FDERIV_floatariths[le])
apply assumption
apply arith
done
then show ?thesis
apply (auto intro!: continuous_on_blinfun_componentwise)
subgoal for i j
apply (rule continuous_on_eq[where f="(\<lambda>x. ode_d_raw (Suc 0) x i j)"])
apply force
apply (subst ode_d2.rep_eq)
apply simp
apply (subst ode_d.rep_eq)
apply (split if_splits)
apply (rule conjI) apply simp
using isn apply force
done
done
qed
lemmas continuous_on_ode_d2_comp[continuous_intros] = continuous_on_compose2[OF continuous_on_ode_d2]
lemma map_ode_fa_nth[simp]:
"d \<le> length ode_e \<Longrightarrow> map (ode_fa_nth CX) [0..<d] = map ((!) (ode_fa CX)) [0..<d]"
by (auto simp: ode_fa_nth cong: map_cong)
lemma map_ode_d_fa_nth[simp]:
"d \<le> length ode_e \<Longrightarrow> map (ode_d_fa_nth i CX X) [0..<d] = map ((!) (ode_d_fa i CX X)) [0..<d]"
by (auto simp: ode_d_fa_nth cong: map_cong)
lemma einterpret_euler_incr_fas:
assumes "length ode_e = DIM('a)" "length X0 = DIM('a)" "length CX = DIM('a)"
"DIM('a) \<le> length vs" "max_Var_floatariths ode_e \<le> DIM('a)"
shows "(einterpret (euler_incr_fas X0 h CX) vs::'a::executable_euclidean_space) =
einterpret X0 vs + (interpret_floatarith h vs) *\<^sub>R ode (einterpret CX vs)"
by (simp add: euler_incr_fas_def euler_incr_fas_nth_def assms ode_fa_nth cong: map_cong)
lemma einterpret_euler_err_fas:
assumes safe: "(einterpret CX vs::'a) \<in> Csafe"
assumes [simp]: "length X0 = DIM('a)" "length CX = DIM('a)" "DIM('a) \<le> length vs"
shows "(einterpret (euler_err_fas X0 h CX) vs::'a::executable_euclidean_space) =
(((interpret_floatarith h vs))\<^sup>2/2) *\<^sub>R ode_d 0 (einterpret CX vs) (ode (einterpret CX vs)) (ode (einterpret CX vs))"
using safe_length[OF safe] safe_max_Var[OF safe]
apply (simp add: euler_err_fas_def euler_err_fas_nth_def[abs_def] euler_incr_fas_def)
apply (subst interpret_ode_d_fa)
by (auto simp: safe)
lemma einterpret_euler_fas1:
assumes safe[simp]: "(einterpret CX vs::'a) \<in> Csafe"
assumes [simp]: "length X0 = DIM('a)" "length CX = DIM('a)" "DIM('a) \<le> length vs"
shows "(einterpret (take DIM('a) (euler_fas X0 h CX)) vs::'a::executable_euclidean_space) =
einterpret X0 vs + (interpret_floatarith h vs) *\<^sub>R ode (einterpret X0 vs) +
(((interpret_floatarith h vs))\<^sup>2/2) *\<^sub>R ode_d 0 (einterpret CX vs) (ode (einterpret CX vs)) (ode (einterpret CX vs))"
using safe_length[OF safe] safe_max_Var[OF safe]
by (simp add: euler_fas_def euler_incr_fas_def euler_incr_fas_nth_def[abs_def]
einterpret_euler_err_fas euler_err_fas_nth_def[abs_def] interpret_ode_d_fa)
lemma einterpret_euler_fas2:
assumes [simp]: "(einterpret CX vs::'a) \<in> Csafe"
assumes [simp]: "length X0 = DIM('a)" "length CX = DIM('a)" "DIM('a) \<le> length vs"
shows "(einterpret (drop DIM('a) (euler_fas X0 h CX)) vs::'a::executable_euclidean_space) =
(((interpret_floatarith h vs))\<^sup>2/2) *\<^sub>R ode_d 0 (einterpret CX vs) (ode (einterpret CX vs)) (ode (einterpret CX vs))"
by (simp add: euler_fas_def euler_incr_fas_def einterpret_euler_err_fas)
lemma ode_d_Suc_0_eq_ode_d2: "x \<in> Csafe \<Longrightarrow> ode_d (Suc 0) x = ode_d2 x"
unfolding ode_d2.rep_eq by auto
lemma rk2_increment_rk2_fas_err:
fixes h s1 s2 rkp x0 cx vs
defines "h' \<equiv> interpret_floatarith h vs"
defines "s2' \<equiv> interpret_floatarith s2 vs"
defines "rkp' \<equiv> interpret_floatarith rkp vs"
defines "x0' \<equiv> einterpret x0 vs"
defines "cx' \<equiv> einterpret cx vs"
assumes cx_flow: "cx' = flow0 x0' (h' * s1')"
assumes [simp]: "length x0 = DIM('a)" "length cx = DIM('a)" "DIM('a) \<le> length vs"
assumes safes: "x0' \<in> Csafe" "cx' \<in> Csafe" "(x0' + (s2' * h' * rkp') *\<^sub>R ode x0')\<in> Csafe"
shows "(einterpret (rk2_fas_err rkp x0 h cx s2) vs::'a::executable_euclidean_space) =
heun_remainder1 (flow0 x0') ode_na ode_d_na ode_d2_na 0 h' s1' -
heun_remainder2 rkp' (flow0 x0') ode_na ode_d2_na 0 h' s2'"
using safes
using safe_length[OF safes(1)] safe_max_Var[OF safes(1)]
apply (auto simp: heun_remainder1_def heun_remainder2_def discrete_evolution_def
ode_na_def ode_d_na_def ode_d2_na_def rk2_increment x0'_def rkp'_def h'_def s2'_def
cx'_def euler_incr_fas_def rk2_fas_err_def rk2_fas_err_nth_def[abs_def]
euler_incr_fas_nth_def[abs_def]
interpret_ode_d_fa)
apply (simp add: ode_d1_eq[symmetric] ode_d_Suc_0_eq_ode_d2 inverse_eq_divide)
apply (simp add: algebra_simps field_simps divide_simps)
unfolding cx'_def[symmetric] cx_flow x0'_def h'_def
apply (simp add: algebra_simps)
done
lemma map_rk2_fas_err_nth[simp]:
"d = length ode_e \<Longrightarrow> length b = length ode_e \<Longrightarrow> map (rk2_fas_err_nth a b c e f) [0..<d] = map ((!) (rk2_fas_err a b c e f)) [0..<d]"
unfolding rk2_fas_err_nth_def rk2_fas_err_def
by (rule map_cong) auto
lemma rk2_increment_rk2_fas1:
fixes h s1 s2 rkp x0 cx vs
defines "h' \<equiv> interpret_floatarith h vs"
defines "s2' \<equiv> interpret_floatarith s2 vs"
defines "rkp' \<equiv> interpret_floatarith rkp vs"
defines "x0' \<equiv> einterpret x0 vs"
defines "cx' \<equiv> einterpret cx vs"
assumes cx_flow: "cx' = flow0 x0' (h' * s1')"
assumes [simp]: "length x0 = DIM('a)" "length cx = DIM('a)" "DIM('a) \<le> length vs"
assumes safes: "(x0'::'a)\<in> Csafe" "(cx'::'a)\<in> Csafe" "(x0' + (s2' * h' * rkp') *\<^sub>R ode x0'::'a)\<in> Csafe"
shows "(einterpret (take DIM('a) (rk2_fas rkp x0 h cx s2)) vs::'a::executable_euclidean_space) =
discrete_evolution (rk2_increment rkp' (\<lambda>_. ode)) h' 0 x0' + (heun_remainder1 (flow0 x0') ode_na ode_d_na ode_d2_na 0 h' s1' -
heun_remainder2 rkp' (flow0 x0') ode_na ode_d2_na 0 h' s2')"
using safes using safe_length[OF safes(1)] safe_max_Var[OF safes(1)]
apply (auto simp: discrete_evolution_def rk2_fas_def)
apply (subst rk2_increment_rk2_fas_err[OF cx_flow[unfolded cx'_def x0'_def h'_def]])
subgoal by simp
subgoal by simp
subgoal by simp
subgoal by (simp add: x0'_def)
subgoal by (simp add: cx'_def)
subgoal by (simp add: x0'_def s2'_def h'_def rkp'_def)
subgoal using [[show_consts, show_sorts, show_types]]
by (auto simp: x0'_def s2'_def h'_def rkp'_def rk2_increment euler_incr_fas_def
euler_incr_fas_nth_def[abs_def] inverse_eq_divide)
done
lemma rk2_increment_rk2_fas2:
fixes h s1 s2 rkp x0 cx vs
defines "h' \<equiv> interpret_floatarith h vs"
defines "s2' \<equiv> interpret_floatarith s2 vs"
defines "rkp' \<equiv> interpret_floatarith rkp vs"
defines "x0' \<equiv> einterpret x0 vs"
defines "cx' \<equiv> einterpret cx vs"
assumes cx_flow: "cx' = flow0 x0' (h' * s1')"
assumes [simp]: "length x0 = DIM('a)" "length cx = DIM('a)" "DIM('a) \<le> length vs"
assumes safes: "x0'\<in> Csafe" "cx'\<in> Csafe" "(x0' + (s2' * h' * rkp') *\<^sub>R ode x0')\<in> Csafe"
shows "(einterpret (drop DIM('a) (rk2_fas rkp x0 h cx s2)) vs::'a::executable_euclidean_space) =
(heun_remainder1 (flow0 x0') ode_na ode_d_na ode_d2_na 0 h' s1' -
heun_remainder2 rkp' (flow0 x0') ode_na ode_d2_na 0 h' s2')"
using safes
apply (auto simp: discrete_evolution_def rk2_fas_def)
apply (subst rk2_increment_rk2_fas_err[OF cx_flow[unfolded cx'_def x0'_def h'_def]])
subgoal by simp
subgoal by simp
subgoal by simp
subgoal by (simp add: x0'_def)
subgoal by (simp add: cx'_def)
subgoal by (simp add: x0'_def s2'_def h'_def rkp'_def)
subgoal by (auto simp: x0'_def s2'_def h'_def rkp'_def rk2_increment euler_incr_fas_def inverse_eq_divide)
done
subsubsection \<open>safe set relation\<close>
lemma mk_safe[le, refine_vcg]: "wd TYPE('a::executable_euclidean_space)\<Longrightarrow>
mk_safe X \<le> SPEC (\<lambda>R::'a set. R = X \<and> X \<subseteq> Csafe)"
unfolding mk_safe_def
by refine_vcg
lemma mk_safe_coll[le, refine_vcg]: "wd TYPE('a::executable_euclidean_space) \<Longrightarrow>
mk_safe_coll X \<le> SPEC (\<lambda>R::'a set. R = X \<and> X \<subseteq> Csafe)"
unfolding mk_safe_coll_def autoref_tag_defs
by (refine_vcg FORWEAK_invarI[where I="\<lambda>a b. X = \<Union>b \<union> a \<and> a \<subseteq> Csafe"]) auto
lemma ode_set_spec[THEN order.trans, refine_vcg]:
assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)"
shows "ode_set X \<le> SPEC (\<lambda>r. ode ` X \<subseteq> (r::'a set))"
using assms wdD[OF assms(1)]
unfolding ode_set_def
apply (refine_vcg)
subgoal by (auto simp: env_len_def ode_slp_def)
subgoal premises prems
using prems(1,2,4-)
by (auto simp: env_len_def eucl_of_list_prod ode_def)
done
lemmas fderiv[derivative_intros] = ode_has_derivative_safeI
lemma fderiv2[derivative_intros]:
"x \<in> Csafe \<Longrightarrow> (ode_d1 has_derivative ode_d2 x) (at x)"
by (frule ode_d1_has_derivative_safeI)
(simp add: ode_d_Suc_0_eq_ode_d2)
lemma derivative_within_safe[derivative_intros]:
"(g has_derivative h) (at x) \<Longrightarrow> (g has_derivative h) (at x within Csafe)"
by (rule has_derivative_at_withinI)
lemma cont_fderiv: "continuous_on (Csafe) ode_d1"
by (rule has_derivative_continuous_on) (auto intro!: derivative_intros)
lemmas cont_fderiv'[continuous_intros] = continuous_on_compose2[OF cont_fderiv]
lemma continuous_on_ode1:
"continuous_on (Csafe) (ode)"
using fderiv
by (auto intro!: has_derivative_continuous_on derivative_intros)
lemma continuous_on_ode[continuous_intros]:
"continuous_on s g \<Longrightarrow> (\<And>x. x \<in> s \<Longrightarrow> (g x) \<in> Csafe) \<Longrightarrow> continuous_on s (\<lambda>x. ode (g x))"
using continuous_on_ode1
by (rule continuous_on_compose2) auto
lemma fderiv'[derivative_intros]:
assumes "(g has_derivative g' y) (at y within X)"
assumes "(g y) \<in> Csafe"
shows "((\<lambda>y. ode (g y)) has_derivative
(blinfun_apply (ode_d1 (g y)) \<circ>\<circ> g') y) (at y within X)"
using diff_chain_within[OF assms(1) has_derivative_within_subset[OF fderiv]] assms(2)
by (simp add: o_def)
lemma fderiv2'[derivative_intros]:
assumes "(g has_derivative g' y) (at y within X)"
assumes "(g y) \<in> Csafe"
shows "((\<lambda>y. ode_d1 (g y)) has_derivative
(blinfun_apply (ode_d2 (g y)) \<circ>\<circ> g') y) (at y within X)"
using diff_chain_within[OF assms(1) has_derivative_within_subset[OF fderiv2]] assms(2)
by (simp add: o_def)
subsubsection \<open>step of Picard iteration\<close>
lemma ncc_interval: "ncc {a .. b::'a::executable_euclidean_space} \<longleftrightarrow> a \<le> b"
by (auto simp: ncc_def)
lemma nonempty_interval: "nonempty {a .. b::'a::executable_euclidean_space} \<longleftrightarrow> a \<le> b"
by (auto simp: nonempty_def)
lemmas [refine_vcg] = Picard_step_def[THEN eq_refl, THEN order.trans]
lemma Basis_list_zero_mem_Basis[simp]:
"Basis_list ! 0 \<in> Basis"
unfolding Basis_list[symmetric]
apply (rule nth_mem)
apply (rule length_Basis_list_pos)
done
lemma cfuncset_empty_iff:
fixes l u::"'d::ordered_euclidean_space"
shows "l \<le> u \<Longrightarrow> cfuncset l u X = {} \<longleftrightarrow> X = {}"
unfolding cfuncset_def Pi_def
proof (safe, goal_cases)
case hyps: (1 x)
from \<open>x \<in> X\<close>
have "(\<lambda>_. x) \<in> {f. \<forall>x. x \<in> {l..u} \<longrightarrow> f x \<in> X} \<inter> Collect (continuous_on {l..u})"
by (auto intro!: continuous_on_const)
then show ?case using hyps by auto
qed auto
lemma lv_ivl_sings: "lv_ivl [x] [y] = (\<lambda>x. [x]) ` {x .. y}"
apply (auto simp: lv_ivl_def)
subgoal for x by (cases x) auto
done
lemma Picard_step_ivl_refine[le, refine_vcg]:
assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)"
assumes "(X::'a set) \<subseteq> Csafe"
assumes "0 \<le> h"
shows "Picard_step_ivl X0 t0 h X \<le> Picard_step X0 t0 h X"
proof -
have "h' \<in> {t0..t0 + h} \<Longrightarrow> t0 \<le> h'" for h' by simp
then show ?thesis
unfolding Picard_step_ivl_def Picard_step_def
apply (refine_vcg, clarsimp_all simp del: atLeastAtMost_iff)
subgoal using \<open>0 \<le> h\<close> by simp
subgoal by (auto simp: euler_incr_slp_def wdD)
subgoal by (auto simp: euler_incr_fas'_def)
subgoal for XS l u
apply (auto simp: lv_ivl_sings nonempty_interval
simp del: atLeastAtMost_iff
intro!: add_integral_ivl_bound)
subgoal for x0 h' phi x h''
apply (drule bspec, assumption)
apply (drule bspec[where x="h'' - t0"], force)
proof goal_cases
case (1)
have *: "map ((!) (list_of_eucl b)) [0..<DIM('a) - Suc 0] @ [b \<bullet> Basis_list ! (DIM('a) - Suc 0)]
= list_of_eucl b" for b::'a
apply (auto intro!: nth_equalityI simp: nth_append not_less)
using Intersection.le_less_Suc_eq by blast
have "phi x \<in> X" if "x \<in> {t0 .. h''}" for x
using 1 that by (auto simp: cfuncset_iff)
have "x0 + (h'' - t0) *\<^sub>R ode b \<in> {l .. u}" if "b \<in> X" for b
proof -
from 1(17)[rule_format, OF that] assms(1)
have "einterpret (euler_incr_fas' D) (list_of_eucl x0 @ (h'' - t0) # list_of_eucl b) \<in> eucl_of_list ` XS"
by (auto simp: wdD)
also have "eucl_of_list ` XS \<subseteq> {l .. u}" by fact
finally show ?thesis
by (simp add: euler_incr_fas'_def einterpret_euler_incr_fas map_nth_append1 nth_append wdD[OF \<open>wd _\<close>] *)
qed
then have *: "(h'' - t0) *\<^sub>R ode b \<in> {l - x0..u - x0}" if "b \<in> X" for b using that
by (auto simp: algebra_simps)
show ?case
apply (rule *)
using 1 by (auto simp: cfuncset_iff)
qed
subgoal
using assms(2)
by (auto intro!: integrable_continuous_real continuous_intros
simp: cfuncset_iff)
done
done
qed
subsubsection \<open>Picard iteration\<close>
lemma inf_le_supI[simp]:
fixes a b c d::"'d::ordered_euclidean_space"
shows
"a \<le> c \<Longrightarrow> inf a b \<le> sup c d"
"a \<le> d \<Longrightarrow> inf a b \<le> sup c d"
"b \<le> c \<Longrightarrow> inf a b \<le> sup c d"
"b \<le> d \<Longrightarrow> inf a b \<le> sup c d"
by (auto simp: eucl_le[where 'a='d] eucl_inf[where 'a='d] eucl_sup[where 'a='d] inf_real_def sup_real_def
intro!: sum_mono scaleR_right_mono)
lemmas [refine_vcg_def] = do_widening_spec_def
lemma P_iter_spec[le, refine_vcg]:
assumes "PHI \<subseteq> Csafe"
assumes "0 \<le> h"
assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)"
shows "P_iter X0 h i PHI \<le>
SPEC (\<lambda>r. case r of
None \<Rightarrow> True
| Some (PHI'::'a set) \<Rightarrow> nonempty PHI' \<and> compact PHI' \<and> (\<exists>PHI'' \<subseteq> PHI'. RETURN (Some PHI'') \<le> Picard_step X0 0 h PHI'))"
using assms[unfolded autoref_tag_defs]
proof (induction i arbitrary: PHI)
case 0 then show ?case
unfolding P_iter.simps
by refine_vcg
next
case (Suc i)
show ?case
unfolding P_iter.simps
apply (refine_vcg Suc)
subgoal by (auto simp: cfuncset_iff Picard_step_def algebra_simps add_increasing2)
subgoal for lu l u b CX CX' lu' l' u' b'
apply (simp add: nonempty_interval Picard_step_def)
apply (safe intro!: exI[where x="{l .. u}"] compact_interval)
subgoal by (auto simp: nonempty_interval)
apply (rule subsetD[of CX' "{l .. u}"])
subgoal
apply (rule order_trans, assumption)
unfolding atLeastatMost_subset_iff
apply (rule disjI2)
apply (rule conjI)
subgoal
apply (rule order_trans[where y="inf l' l - (if b' then \<bar>l' - l\<bar> else 0)"])
by (simp_all split: if_split_asm add: algebra_simps add_increasing2)
subgoal
apply (split if_split_asm)
apply (rule order_trans[where y="sup u' u + \<bar>u' - u\<bar>"])
by (auto split: if_split_asm simp add: algebra_simps add_increasing2)
done
subgoal by force
done
done
qed
subsubsection \<open>iterate step size\<close>
lemma Ball_cfuncset_continuous_on:
"\<forall>f\<in>cfuncset a b X. continuous_on {a..b} f"
by (simp add: cfuncset_iff)
lemmas one_step_methodD = one_step_method_def[THEN iffD1, rule_format, le]
lemmas one_step_methodI = one_step_method_def[THEN iffD2, rule_format]
lemma cert_stepsize_lemma:
assumes prems: " 0 < h"
"X0 \<subseteq> {l..u}"
"Res \<subseteq> Csafe"
"Res_ivl \<subseteq> Csafe"
"{l..u} \<subseteq> Csafe"
"nonempty PHI'"
"nonempty Res"
"\<forall>x0\<in>X0.
x0 \<in> Csafe \<longrightarrow>
h \<in> existence_ivl0 x0 \<longrightarrow>
(\<forall>h'\<in>{0..h}. flow0 x0 h' \<in> PHI') \<longrightarrow>
x0 + h *\<^sub>R ode x0 \<in> PHI' \<longrightarrow> flow0 x0 h \<in> Res"
"nonempty Res_ivl"
"\<forall>x0\<in>X0.
x0 \<in> Csafe \<longrightarrow>
(\<forall>h\<in>{0..h}.
h \<in> existence_ivl0 x0 \<longrightarrow>
(\<forall>h'\<in>{0..h}. flow0 x0 h' \<in> PHI') \<longrightarrow>
x0 + h *\<^sub>R ode x0 \<in> PHI' \<longrightarrow> flow0 x0 h \<in> Res_ivl)"
"compact PHI'"
"PHI'' \<subseteq> PHI'"
"RETURN (Some PHI'') \<le> Picard_step X0 0 h PHI'"
shows "flowpipe0 X0 h h Res_ivl Res"
proof -
from prems
have Ps: "RETURN (Some PHI'') \<le> Picard_step X0 0 h PHI'"
by simp
from Ps have PHI':
"compact PHI''" "PHI'' \<subseteq> Csafe"
"\<forall>x\<in>PHI''. x \<in> Csafe"
"\<And>x0 h'' phi. x0 \<in> X0 \<Longrightarrow> 0 \<le> h'' \<Longrightarrow> h'' \<le> h \<Longrightarrow> phi \<in> cfuncset 0 h'' PHI' \<Longrightarrow>
x0 + integral {0..h''} (\<lambda>t. ode (phi t)) \<in> PHI''"
by (auto simp: Picard_step_def)
then obtain cx where cx: "(\<lambda>t::real. cx) \<in> cfuncset 0 0 PHI'"
using \<open>PHI'' \<subseteq> PHI'\<close> \<open>nonempty PHI'\<close> by (auto simp: cfuncset_def continuous_on_const nonempty_def)
show "flowpipe0 X0 h h Res_ivl Res"
unfolding flowpipe0_def atLeastAtMost_singleton
proof safe
show "0 \<le> h" using \<open>0 < h\<close> by simp
show safe_X0: "x \<in> Csafe" if "x \<in> X0" for x using that \<open>{l..u} \<subseteq> Csafe\<close> \<open>X0 \<subseteq> {l..u}\<close> by blast
show "x \<in> Csafe" if "x \<in> Res_ivl" for x using prems that
by (auto simp: )
show "x \<in> Csafe" if "x \<in> Res" for x using prems that by (auto simp:)
fix x0 assume "x0 \<in> X0"
from PHI'(4)[OF \<open>x0 \<in> X0\<close> order_refl \<open>0 \<le> h\<close> cx]
have "x0 \<in> PHI''" by simp
have *: "h \<in> existence_ivl0 x0" "s \<in> {0 .. h} \<Longrightarrow> flow0 x0 s \<in> PHI''" for s
using \<open>x0 \<in> X0\<close> \<open>PHI'' \<subseteq> PHI'\<close> \<open>0 \<le> h\<close> PHI'(3) \<open>x0 \<in> PHI''\<close>
by (auto
simp: cfuncset_def Pi_iff closed_segment_eq_real_ivl ivl_integral_def
intro!: Picard_iterate_mem_existence_ivlI[OF UNIV_I _ UNIV_I \<open>compact PHI''\<close>
\<open>x0 \<in> PHI''\<close> \<open>PHI'' \<subseteq> Csafe\<close>] PHI'(4)) force+
show h_ex: "h \<in> existence_ivl0 x0" by fact
have cf: "(\<lambda>t::real. x0) \<in> cfuncset 0 h PHI'" for h
using \<open>x0 \<in> PHI''\<close> \<open>PHI'' \<subseteq> PHI'\<close>
by (auto simp: cfuncset_def continuous_intros)
have mem_PHI': "x0 + h' *\<^sub>R ode x0 \<in> PHI'" if "0 \<le> h'" "h' \<le> h" for h'
using that \<open>PHI'' \<subseteq> PHI'\<close> PHI'(4)[OF \<open>x0 \<in> X0\<close> that cf]
by auto
from this prems safe_X0
show "flow0 x0 h \<in> Res"
using \<open>0 \<le> h\<close> h_ex * \<open>PHI'' \<subseteq> PHI'\<close> \<open>x0 \<in> X0\<close>
by (auto simp: subset_iff dest!: bspec[where x=x0])
fix h' assume h': "h' \<in> {0..h}"
then have "h' \<in> existence_ivl0 x0"
by (meson atLeastAtMost_iff existence_ivl_zero h_ex is_interval_1
local.is_interval_existence_ivl local.mem_existence_ivl_iv_defined(2))
from h' this prems safe_X0
show "flow0 x0 h' \<in> Res_ivl"
using \<open>0 < h\<close> h_ex * \<open>PHI'' \<subseteq> PHI'\<close> \<open>x0 \<in> X0\<close> mem_PHI' \<open>x0 \<in> PHI''\<close>
by (auto simp: subset_iff dest!: bspec[where x=x0])
qed
qed
lemma cert_stepsize_spec[le,refine_vcg]:
assumes "h > 0"
assumes "one_step_method m"
assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)"
shows "cert_stepsize m X0 h i n \<le> SPEC (\<lambda>(h', RES::'a set, RES_ivl, _). nonempty RES \<and> nonempty RES_ivl \<and> 0 < h' \<and> h' \<le> h \<and> flowpipe0 X0 h' h' RES_ivl RES)"
using assms(1)[unfolded autoref_tag_defs]
proof (induction n arbitrary: h)
case 0 then show ?case by simp
next
case (Suc n)
note [refine_vcg] = Suc.IH[of "h/2", le]
show ?case
unfolding cert_stepsize.simps
using Suc.prems
by (refine_vcg Ball_cfuncset_continuous_on one_step_methodD[OF assms(2)])
(clarsimp_all simp: cert_stepsize_lemma)
qed
subsubsection \<open>Euler step\<close>
lemma embed_real_ivl_iff[simp]:
"(\<forall>x \<in> {0 .. h *\<^sub>R (One::'a::executable_euclidean_space)}. P (x \<bullet> hd Basis_list)) \<longleftrightarrow> (\<forall>x \<in> {0 .. h}. P x)"
proof (auto simp: eucl_le[where 'a='a], goal_cases)
case hyps: (1 x)
have "x = x *\<^sub>R (One::'a) \<bullet> hd Basis_list"
by auto
also have "P \<dots>"
apply (rule hyps[rule_format])
using hyps
by (auto simp: eucl_le[where 'a='a])
finally show ?case .
qed
lemma convex_on_segmentI:
assumes mem_convex: "convex C" "a \<in> C" "a + j *\<^sub>R b \<in> C"
assumes "0 \<le> i" "i \<le> j"
shows "a + i *\<^sub>R b \<in> C"
proof -
have "a + i *\<^sub>R b = (1 - i / j) *\<^sub>R a + (i / j) *\<^sub>R (a + j *\<^sub>R b)"
using assms
by (auto simp: algebra_simps diff_divide_distrib)
also have "\<dots> \<in> C"
using assms
by (auto simp: divide_simps intro!: convexD[OF mem_convex])
finally show ?thesis .
qed
lemma one_step_flowpipe:
assumes [THEN one_step_methodD, refine_vcg]: "one_step_method m"
assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)"
shows "one_step X0 h m \<le> SPEC (\<lambda>(h', _, RES_ivl, RES::'a set). 0 < h' \<and> h' \<le> h \<and> flowpipe0 X0 h' h' RES_ivl RES)"
using assms
unfolding one_step_def
by refine_vcg
lemma ncc_imageD:
assumes "ncc ((\<lambda>x. x ! i) ` env)"
assumes "nth_image_precond env i"
shows "compact ((\<lambda>x. x ! i) ` env::real set)" "closed ((\<lambda>x. x ! i) ` env)" "bounded ((\<lambda>x. x ! i) ` env)"
"((\<lambda>x. x ! i) ` env) \<noteq> {}" "convex ((\<lambda>x. x ! i) ` env)"
using assms
by (auto simp: ncc_def nth_image_precond_def compact_eq_bounded_closed)
lemma max_Var_floatariths_ode_d_fa[le]:
assumes [simp]: "length ode_e > 0" "max_Var_floatariths ode_e \<le> length ode_e"
"length cxs = length ode_e" "length ys = length ode_e"
shows "max_Var_floatariths (ode_d_fa i cxs ys) \<le> max (max_Var_floatariths (cxs)) (max_Var_floatariths ys)"
apply (auto simp: ode_d_fa_def max_Var_floatariths_Max )
using assms apply auto[1]
apply (auto intro!: max_Var_floatarith_subst_floatarith_le max_Var_floatariths_ode_d_expr
max_Var_floatarith_le_max_Var_floatariths_nthI max_Var_ode_fa
simp: in_set_conv_nth)
apply (auto simp: max_Var_floatariths_Max in_set_conv_nth)
done
lemma max_Var_floatariths_euler_err_fas[le]:
assumes nz: "0 < length ode_e"
and [simp]: "max_Var_floatariths ode_e \<le> length ode_e"
"length xs = length ode_e"
"length cxs = length ode_e"
shows "max_Var_floatariths (euler_err_fas xs h cxs)
\<le> max (max_Var_floatariths xs) (max (max_Var_floatarith h) (max_Var_floatariths cxs))"
using nz
by (auto simp: euler_err_fas_def[abs_def] euler_err_fas_nth_def[abs_def] map_nth_eq_self simp del: length_0_conv
intro!: max_Var_floatariths_ode_d_fa max_Var_floatariths_map_times
max_Var_floatariths_map_const max_Var_ode_fa; arith)
lemma max_Var_floatariths_euler_incr_fas[le]:
assumes [simp]: "max_Var_floatariths ode_e \<le> length ode_e"
"length xs = length ode_e"
"length cxs = length ode_e"
shows "max_Var_floatariths (euler_incr_fas xs h cxs)
\<le> max (max_Var_floatariths xs) (max (max_Var_floatarith h) (max_Var_floatariths cxs))"
using length_ode_fa
by (auto simp: euler_incr_fas_def euler_incr_fas_nth_def[abs_def] simp del: length_ode_fa
intro!: max_Var_floatariths_ode_d_fa max_Var_floatariths_map_plus max_Var_floatariths_map_times
max_Var_floatariths_map_const max_Var_ode_fa)
lemma map_euler_incr_fas_nth: "length X0 = d \<Longrightarrow> map (euler_incr_fas_nth X0 h CX) [0..<d] = euler_incr_fas X0 h CX"
by (auto simp: euler_incr_fas_def)
lemma map_euler_err_fas_nth: "length X0 = d \<Longrightarrow> map (euler_err_fas_nth X0 h CX) [0..<d] = euler_err_fas X0 h CX"
by (auto simp: euler_err_fas_def)
lemma max_Var_floatariths_euler_fas[le]:
assumes [simp]: "max_Var_floatariths ode_e \<le> length ode_e"
"length xs = length ode_e"
"length cxs = length ode_e"
assumes nz: "0 < length ode_e"
shows "max_Var_floatariths (euler_fas xs h cxs) \<le> Max {max_Var_floatariths xs, max_Var_floatarith h, max_Var_floatariths cxs}"
using nz
by (auto simp: euler_fas_def map_euler_incr_fas_nth map_euler_err_fas_nth
intro!: max_Var_floatariths_map_plus max_Var_floatariths_euler_incr_fas
max_Var_floatariths_euler_err_fas)
lemma take_interpret_floatariths:
"d < length fas \<Longrightarrow> take d (interpret_floatariths fas vs) = interpret_floatariths (take d fas) vs"
by (auto intro!: nth_equalityI)
lemma length_euler_slp_le: "2 * D \<le> length euler_slp"
by (auto simp: euler_fas'_def euler_slp_def intro!: order_trans[OF _ length_slp_of_fas_le])
lemma ncc_nonempty[simp]: "ncc x \<Longrightarrow> nonempty x"
by (simp add: ncc_def nonempty_def)
lemma nccD:
assumes "ncc X"
shows "compact X" "closed X" "bounded X" "X \<noteq> {}" "convex X"
using assms
by (auto simp: ncc_def nth_image_precond_def compact_eq_bounded_closed)
lemma D_DIM_wdD[simp]: "wd TYPE('a::executable_euclidean_space) \<Longrightarrow> D = DIM('a)"
by (auto simp: wdD)
lemma euler_step_flowpipe:
includes floatarith_notation
assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)"
shows "euler_step X0 h \<le> SPEC (\<lambda>(h', _, RES_ivl, RES::'a set). 0 < h' \<and> h' \<le> h \<and> flowpipe0 X0 h' h' RES_ivl RES)"
unfolding euler_step_def THE_NRES_def
apply (intro SPEC_rule_conjI one_step_flowpipe one_step_methodI)
apply (refine_vcg, clarsimp_all)
subgoal using assms by (auto simp: euler_slp_def euler_fas'_def)
subgoal by (auto simp: euler_slp_def euler_fas'_def)
subgoal using length_euler_slp_le assms by (auto simp: env_len_def wdD[OF \<open>wd _\<close>])
subgoal using length_euler_slp_le assms by (auto simp: env_len_def wdD[OF \<open>wd _\<close>])
proof (goal_cases)
case hyps: (1 X0 CX hl hu env res b x0 enve h)
then interpret derivative_on_prod "{0 .. h}" CX "\<lambda>_. ode" "\<lambda>(t, x). ode_d1 x o\<^sub>L snd_blinfun"
by unfold_locales (auto intro!: continuous_intros derivative_eq_intros
simp: split_beta' subset_iff wdD[OF \<open>wd _\<close>])
from \<open>h \<in> existence_ivl0 x0\<close> have s_ex: "s \<in> existence_ivl0 x0" if "0 \<le> s" "s \<le> h" for s
by (metis (no_types, lifting) atLeastAtMost_iff ivl_subset_existence_ivl subset_iff that)
have "flow0 x0 (h) = flow0 x0 (0 + (h))" by simp
also have "\<dots> \<in> eucl_of_list ` take D ` env"
using hyps
apply (intro euler_consistent_traj_set[where x="flow0 x0" and u = "h"])
apply (auto intro!: \<open>0 \<le> h\<close> flow_has_vector_derivative[THEN has_vector_derivative_at_within]
simp: nccD discrete_evolution_def euler_increment subset_iff wdD[OF \<open>wd _\<close>]
Let_def s_ex min_def max_def lv_ivl_sings)
subgoal premises prems for s
proof -
have "interpret_floatariths (euler_fas' DIM('a)) (list_of_eucl x0 @ list_of_eucl (flow0 x0 s) @ [h]) \<in> env"
using prems
by (auto intro!: prems(1)[rule_format])
then have "eucl_of_list (take D (interpret_floatariths (euler_fas' DIM('a)) (list_of_eucl x0 @ list_of_eucl (flow0 x0 s) @ [h])))
\<in> eucl_of_list ` take D ` env"
(is "eucl_of_list (take _ (interpret_floatariths _ ?vs)) \<in> _")
by auto
also
have "take (2 * D) (interpret_floatariths (euler_fas' DIM('a)) ?vs) =
interpret_floatariths (map fold_const_fa (euler_fas (map floatarith.Var [0..<D]) (Var (2 * D)) (map floatarith.Var [D..<2 * D]))) ?vs"
unfolding euler_fas'_def
by (auto simp: euler_fas_def wdD[OF \<open>wd _\<close>] simp del: map_map
intro!: max_Var_floatariths_map_plus max_Var_floatariths_euler_incr_fas
max_Var_floatariths_euler_err_fas \<open>wd _\<close>
max_Var_floatariths_fold_const_fa[le])
then have "take D (take (2 * D) (interpret_floatariths (euler_fas' DIM('a)) ?vs)) =
take D (interpret_floatariths (euler_fas (map floatarith.Var [0..<D]) (Var(2 * D)) (map floatarith.Var [D..<2 * D])) ?vs)"
by simp
then have "take D (interpret_floatariths (euler_fas' DIM('a)) ?vs) =
take DIM('a) (interpret_floatariths (euler_fas (map floatarith.Var [0..<D]) (Var(2 * D)) (map floatarith.Var [D..<2 * D])) ?vs)"
by (simp add: wdD[OF \<open>wd _\<close>])
also have "eucl_of_list \<dots> =
x0 + h *\<^sub>R ode x0 + (h\<^sup>2 / 2) *\<^sub>R (ode_d 0 (flow0 x0 s) (ode (flow0 x0 s))) (ode (flow0 x0 s))"
by (auto simp: take_interpret_floatariths einterpret_euler_fas1 map_nth_append1 prems nth_append
wdD[OF \<open>wd _\<close>])
finally show ?thesis
by (simp add: prems(10) prems(13) prems(14) prems(5) ode_d1_eq[symmetric] wdD[OF \<open>wd _\<close>])
qed
done
also have "\<dots> \<subseteq> res" using assms hyps by auto
finally show ?case by simp
qed (auto simp: assms)
lemma length_rk2_slp_le: "2 * D \<le> length rk2_slp"
by (auto simp: rk2_slp_def rk2_fas'_def intro!: order_trans[OF _ length_slp_of_fas_le])
lemma max_Var_floatarith_R\<^sub>e[simp]: "max_Var_floatarith (R\<^sub>e x) = 0"
by (auto simp: R\<^sub>e_def split: prod.splits)
lemma max_Var_floatariths_rk2_fas_err[le]:
assumes nz: "0 < length ode_e"
and [simp]: "max_Var_floatariths ode_e \<le> length ode_e" "length x0 = length ode_e" "length cx = length ode_e"
shows "max_Var_floatariths (rk2_fas_err rkp x0 h cx s2) \<le>
Max {max_Var_floatarith rkp, max_Var_floatariths x0, max_Var_floatarith h, max_Var_floatariths cx,
max_Var_floatarith s2}"
using nz
unfolding rk2_fas_err_def rk2_fas_err_nth_def
by (auto simp: rk2_fas_err_def
intro!: max_Var_floatariths_append max_Var_floatariths_map_plus max_Var_floatariths_map_times
max_Var_floatariths_map_const max_Var_ode_fa max_Var_floatariths_euler_incr_fas
max_Var_floatariths_ode_d_fa; arith)
lemma max_Var_floatarith_one[simp]: "max_Var_floatarith 1 = 0"
and max_Var_floatarith_zero[simp]: "max_Var_floatarith 0 = 0"
by (auto simp: one_floatarith_def zero_floatarith_def)
lemma max_Var_floatariths_rk2_fas[le]:
assumes nz: "0 < length ode_e"
and [simp]: "max_Var_floatariths ode_e \<le> length ode_e" "length x0 = length ode_e" "length cx = length ode_e"
shows "max_Var_floatariths (rk2_fas rkp x0 h cx s2) \<le>
Max {max_Var_floatarith rkp, max_Var_floatariths x0, max_Var_floatarith h, max_Var_floatariths cx,
max_Var_floatarith s2}"
using nz
by (auto simp: rk2_fas_def
intro!: max_Var_floatariths_append max_Var_floatariths_map_plus max_Var_floatariths_map_times
max_Var_floatariths_map_const max_Var_ode_fa max_Var_floatariths_euler_incr_fas
max_Var_floatariths_rk2_fas_err)
lemma rk2_step_flowpipe:
includes floatarith_notation
assumes [refine_vcg]: "wd TYPE('a::executable_euclidean_space)"
shows "rk2_step X0 h \<le> SPEC (\<lambda>(h', _, RES_ivl, RES::'a set).
0 < h' \<and> h' \<le> h \<and> flowpipe0 X0 h' h' RES_ivl RES)"
unfolding rk2_step_def THE_NRES_def
apply (intro one_step_flowpipe assms one_step_methodI)
apply (refine_vcg, clarsimp_all)
subgoal using assms by (auto simp: rk2_slp_def rk2_fas'_def)
subgoal by (auto simp: rk2_slp_def rk2_fas'_def)
subgoal using length_rk2_slp_le by (auto simp: env_len_def wdD[OF \<open>wd _\<close>])
subgoal using length_rk2_slp_le by (auto simp: env_len_def wdD[OF \<open>wd _\<close>])
proof (goal_cases)
case hyps: (1 X0 CX hl hu rk2_param env res b x0 el h)
from assms have "D = DIM('a)" by simp
have aux: "ode (flow0 x0 s) = ode (snd (s, flow0 x0 s))" for s
by simp
from hyps interpret derivative_on_prod "{0 .. h}" CX "\<lambda>_ x. ode x" "\<lambda>(t, x). ode_d1 x o\<^sub>L snd_blinfun"
by unfold_locales
(auto intro!: continuous_intros derivative_eq_intros simp: split_beta' subset_iff)
have aux2: "blinfun_apply (ode_d1 (snd tx)) \<circ> snd = blinfun_apply (ode_d1 (snd tx) o\<^sub>L snd_blinfun)"
for tx::"real\<times>'a"
by (auto intro!: blinfun_eqI)
have aux3: "blinfun_apply (ode_d2 (snd tx)) (snd h) o\<^sub>L snd_blinfun =
(flip_blinfun (flip_blinfun (ode_d2 (snd tx) o\<^sub>L snd_blinfun) o\<^sub>L snd_blinfun)) h"
for tx h::"real\<times>'a"
by (auto intro!: blinfun_eqI)
have "flow0 x0 (h) = flow0 x0 (0 + (h))" by simp
also have "\<dots> \<in> eucl_of_list ` take D ` env"
using hyps assms
apply (intro rk2_consistent_traj_set[where
x="flow0 x0" and u = "h" and T="{0..h}" and X="CX" and p="rk2_param"
and f = "ode_na" and f' = ode_d_na and g' = ode_d_na and f'' = ode_d2_na and g'' = ode_d2_na])
subgoal by (simp add: \<open>0 \<le> h\<close>)
subgoal by simp
subgoal by simp
subgoal by auto
subgoal by (auto simp add: ncc_def nonempty_def)
subgoal
apply (rule flow_has_vector_derivative[THEN has_vector_derivative_at_within, THEN has_vector_derivative_eq_rhs])
subgoal by (metis (no_types, lifting) ivl_subset_existence_ivl subset_iff)
subgoal by (force simp: ode_na_def[abs_def] ode_d_na_def[abs_def] ode_d2_na_def[abs_def])
done
subgoal
unfolding ode_na_def ode_d_na_def ode_d2_na_def
apply (rule derivative_eq_intros)
apply (rule derivative_intros)
apply (rule derivative_intros)
subgoal by (auto simp: ncc_def nonempty_def)
subgoal by force
done
subgoal
unfolding ode_na_def ode_d_na_def ode_d2_na_def
apply (rule derivative_eq_intros)
apply (rule derivative_intros)
apply (rule derivative_intros)
apply (rule derivative_intros)
subgoal by (force simp: nonempty_def)
apply (rule derivative_intros)
subgoal by (auto intro!: aux3)
done
subgoal by (rule refl)
subgoal by (rule refl)
subgoal
apply (rule compact_imp_bounded)
apply (rule compact_continuous_image)
subgoal
by (auto intro!: continuous_intros simp: ode_na_def ode_d_na_def ode_d2_na_def)
subgoal by (auto simp: ncc_def intro!: compact_Times)
done
subgoal by auto
subgoal by simp
subgoal by simp
subgoal
apply (rule convex_on_segmentI[where j=h])
using mult_left_le_one_le[of h "rk2_param"]
by (auto simp: ncc_def mult_left_le_one_le mult_le_one ac_simps ode_na_def
ode_d_na_def ode_d2_na_def dest: bspec[where x=0])
subgoal by (simp add: ncc_def)
subgoal by (simp add: ncc_def compact_imp_closed)
subgoal for s1 s2
apply (clarsimp simp add: lv_ivl_sings)
subgoal premises prems
proof -
have "s2 * rk2_param * h \<le> h"
apply (rule mult_left_le_one_le)
using assms prems
by (auto intro!: mult_le_one)
then have s2: "(s2 * h * rk2_param) \<in> {0 .. h}"
using prems assms by (auto simp: ac_simps)
have s1: "h * s1 \<in> {0 .. h}" using prems
by (auto intro!: mult_right_le_one_le)
then have
"interpret_floatariths (rk2_fas' D)
(list_of_eucl x0 @ list_of_eucl (flow0 x0 (h * s1)) @ [rk2_param, h, s2]) \<in> env"
unfolding \<open>D = _\<close> using prems
by (intro prems(17)[rule_format]) auto
then have "take (2 * D) (interpret_floatariths (rk2_fas' D)
(list_of_eucl x0 @ list_of_eucl (flow0 x0 (h * s1)) @ [rk2_param, h, s2])) \<in> take (2 * D) ` env"
(is "?l \<in> _")
by auto
also have "?l = interpret_floatariths
(map fold_const_fa (rk2_fas (Var (2 * D)) (map floatarith.Var [0..<D]) (Var (2 * D + 1))
(map floatarith.Var [D..<2 * D])
(Var (2 * D + 2))))
(list_of_eucl x0 @ list_of_eucl (flow0 x0 (h * s1)) @ [rk2_param, h, s2])"
(is "_ = interpret_floatariths (map fold_const_fa ?fas) ?xs")
unfolding rk2_fas'_def
by (auto intro!: max_Var_floatariths_rk2_fas max_Var_floatariths_fold_const_fa[le] simp: wdD[OF \<open>wd _\<close>])
finally have "take D (interpret_floatariths ?fas ?xs) \<in> take D ` take (2 * D) ` env"
by auto
also have "\<dots> = take D ` env" by (auto simp: image_image wdD[OF \<open>wd _\<close>])
finally have "eucl_of_list (take D (interpret_floatariths ?fas ?xs)) \<in> eucl_of_list ` take D ` env"
by simp
then have "einterpret (take D ?fas) ?xs \<in> eucl_of_list ` take D ` env"
by (simp add: take_interpret_floatariths wdD[OF \<open>wd _\<close>])
also have "einterpret (take D ?fas) ?xs =
discrete_evolution (rk2_increment (rk2_param) (\<lambda>t x. ode_na (t, x))) h 0 x0 +
heun_remainder1 (flow0 x0) ode_na ode_d_na ode_d2_na 0 h s1 -
heun_remainder2 (rk2_param) (flow0 x0) ode_na ode_d2_na 0 h s2"
apply (simp add: wdD[OF \<open>wd _\<close>])
apply (subst rk2_increment_rk2_fas1[where ?s1'.0 = s1])
subgoal by (auto simp: nth_append map_nth_append1)
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by (auto simp: nth_append map_nth_append1 \<open>x0 \<in> Csafe\<close>)
subgoal
apply (auto simp: nth_append map_nth_append1 \<open>x0 \<in> Csafe\<close>)
by (meson connectedD_interval existence_ivl_zero flow0_defined hyps
mult_right_le_one_le mult_sign_intros(1) mvar.connected prems)
subgoal
proof -
have "x0 + ((rk2_param * s2) * h) *\<^sub>R ode x0 \<in> CX"
by (rule convex_on_segmentI[where j=h])
(use prems in \<open>auto simp: ncc_def mult_left_le_one_le mult_le_one
dest: bspec[where x=0]\<close>)
also note \<open>\<dots> \<subseteq> Csafe\<close>
finally show ?thesis
by (auto simp: nth_append map_nth_append1 \<open>x0 \<in> Csafe\<close> ac_simps)
qed
subgoal by (auto simp: nth_append map_nth_append1 ode_na_def)
done
finally show ?thesis by (simp add: \<open>D = _\<close>)
qed
done
done
- also have "\<dots> \<subseteq> res" using hyps(5) by (auto simp: \<open>D = _\<close>)
+ also have "\<dots> \<subseteq> res" using hyps(6) by (simp add: \<open>D = _\<close>)
finally show ?case by (simp add: \<open>D = _\<close>)
qed
lemma interpret_adapt_stepsize_fa:
"interpret_floatarith (adapt_stepsize_fa rtol m_id e h') []
= float_of h' * (float_of(rtol) / float_of e) powr (1 / (float_of (m_id) + 1))"
by (auto simp: inverse_eq_divide adapt_stepsize_fa_def)
lemma choose_step_flowpipe[le, refine_vcg]:
assumes "wd TYPE('a::executable_euclidean_space)"
shows "choose_step X0 h \<le> SPEC (\<lambda>(h', _, RES_ivl, (RES::'a set)). 0 < h' \<and> h' \<le> h \<and> flowpipe0 X0 h' h' RES_ivl RES)"
using assms
unfolding choose_step_def
by (refine_vcg rk2_step_flowpipe euler_step_flowpipe)
lemma CsafeI: "t \<in> existence_ivl0 x \<Longrightarrow> x \<in> Csafe"
using local.mem_existence_ivl_iv_defined(2) by blast
lemma apply_vareq: "blinfun_apply (vareq x t) = ode_d1 (flow0 x t)"
by (auto simp: vareq_def)
lemma Dflow_has_derivative:
"t \<in> existence_ivl0 x \<Longrightarrow> (Dflow x has_derivative blinfun_scaleR_left (ode_d1 (flow0 x t) o\<^sub>L Dflow x t)) (at t)"
by (auto simp: Dflow_def blinfun.bilinear_simps scaleR_blinfun_compose_left apply_vareq CsafeI
intro!: derivative_eq_intros mvar.flow_has_derivative[THEN has_derivative_eq_rhs] ext
blinfun_eqI)
lemma matrix_scaleR: "matrix (blinfun_apply (h *\<^sub>R X)) = h *\<^sub>R matrix X"
by (vector matrix_def blinfun.bilinear_simps)
lemma blinfun_of_vmatrix_matrix_matrix_mult[simp]:
"blinfun_of_vmatrix (A ** B) = blinfun_of_vmatrix A o\<^sub>L blinfun_of_vmatrix B"
including blinfun.lifting
by transfer (auto simp: o_def matrix_vector_mul_assoc)
lemma blinfun_of_vmatrix_mat_1[simp]: "blinfun_of_vmatrix (mat 1) = 1\<^sub>L"
including blinfun.lifting
by transfer (auto simp: matrix_vector_mul_lid)
lemma blinfun_of_vmatrix_matrix[simp]:
"blinfun_of_vmatrix (matrix (blinfun_apply A)) = A"
including blinfun.lifting
by transfer (auto simp: bounded_linear.linear matrix_works)
lemma inner_Basis_eq_vec_nth: "b \<in> Basis \<Longrightarrow> v \<bullet> b = vec_nth v (enum_class.enum ! index Basis_list b)"
by (auto simp: inner_vec_def vec_nth_Basis if_distrib Basis_vec_def axis_eq_axis
index_Basis_list_axis1
cong: if_cong)
lemma intersects_sctns_spec_nres[le, refine_vcg]:
"intersects_sctns X' sctns \<le> intersects_sctns_spec X' sctns"
unfolding intersects_sctns_spec_def intersects_sctns_def
by refine_vcg auto
lemma intersects_sections_spec_clw_ref[le, refine_vcg]:
"intersects_sctns_spec_clw R sctns \<le> intersects_sctns_spec R sctns"
unfolding intersects_sctns_spec_def intersects_sctns_spec_clw_def
by (refine_vcg FORWEAK_mono_rule[where I="\<lambda>S b. \<not>b \<longrightarrow> \<Union>S \<inter> \<Union>(plane_of ` sctns) = {}"]) auto
lemma eq_nth_iff_index:
"distinct xs \<Longrightarrow> n < length xs \<Longrightarrow> i = xs ! n \<longleftrightarrow> index xs i = n"
using index_nth_id by fastforce
lemma
max_Var_floatariths_ode_e_wd:
assumes wd: "wd (TYPE('n::enum rvec))"
assumes "CARD('n) \<le> K"
shows "max_Var_floatariths ode_e \<le> K"
using wdD[OF wd] assms by auto
lemma nonzero_component[le, refine_vcg]: "nonzero_component s X n \<le> SPEC (\<lambda>_. \<forall>b\<in>X. b \<bullet> n \<noteq> 0)"
unfolding nonzero_component_def
by refine_vcg auto
lemma
interpret_slp_env_lenD:
assumes "\<forall>cx\<in>CX. interpret_slp (slp_of_fas (fas)) (env cx) \<in> R"
assumes "cx \<in> CX"
shows "interpret_floatariths fas (env cx) \<in> take (length fas) ` R"
proof -
from slp_of_fas
have "take (length fas) (interpret_slp (slp_of_fas fas) (env cx)) = interpret_floatariths fas (env cx)"
by auto
moreover
from assms(1)[rule_format, OF \<open>cx \<in> CX\<close>]
have "interpret_slp (slp_of_fas fas) (env cx) \<in> R" by auto
ultimately show ?thesis by force
qed
lemma flowpipe0_imp_flowpipe:
assumes "flowpipe0 (fst ` X0) x1 x1 aba bba"
shows "flowpipe X0 x1 x1 (aba \<times> UNIV) (bba \<times> UNIV)"
using assms
by (auto simp: flowpipe0_def flowpipe_def)
lemma disjoints_spec[le, refine_vcg]:
"disjoints_spec X Y \<le> SPEC (\<lambda>b. b \<longrightarrow> (X \<inter> Y = {}))"
unfolding disjoints_spec_def autoref_tag_defs
by refine_vcg auto
lemma inner_eq_zero_abs_BasisI:
"\<bar>y\<bar> \<in> Basis \<Longrightarrow> b \<in> Basis \<Longrightarrow> \<bar>y\<bar> \<noteq> b \<Longrightarrow> y \<bullet> b = 0"
for b y::"'a::executable_euclidean_space"
by (metis abs_inner inner_Basis linorder_not_le order_refl zero_less_abs_iff)
lemma abs_in_Basis_absE:
fixes x y::"'a::executable_euclidean_space"
assumes "abs y \<in> Basis"
obtains "abs y = y" | "abs y = -y"
proof -
have "abs y = (\<Sum>i\<in>Basis. (abs (y \<bullet> i)) *\<^sub>R i)"
by (simp add: euclidean_representation abs_inner[symmetric] assms)
also have "Basis = insert (abs y) (Basis - {abs y})" using assms by auto
also have "(\<Sum>i\<in>insert \<bar>y\<bar> (Basis - {\<bar>y\<bar>}). \<bar>y \<bullet> i\<bar> *\<^sub>R i) = \<bar>y \<bullet> \<bar>y\<bar>\<bar> *\<^sub>R \<bar>y\<bar>"
apply (subst sum.insert)
using assms
by (auto simp: abs_inner[symmetric] inner_Basis if_distribR if_distrib
cong: if_cong)
finally have "\<bar>y\<bar> = \<bar>y \<bullet> \<bar>y\<bar>\<bar> *\<^sub>R \<bar>y\<bar>" by simp
moreover have "\<dots> = y \<or> \<dots> = - y"
using assms
by (auto simp: abs_real_def algebra_simps intro!: euclidean_eqI[where 'a='a]
simp: inner_Basis inner_eq_zero_abs_BasisI split: if_splits)
ultimately consider "\<bar>y\<bar> = y" | "\<bar>y\<bar> = - y" by auto
then show ?thesis
by (cases; rule that)
qed
lemma abs_in_BasisE:
fixes x y::"'a::executable_euclidean_space"
assumes "abs y \<in> Basis"
obtains i where "i \<in> Basis" "y = i" | i where "i \<in> Basis" "y = -i"
proof -
from abs_in_Basis_absE[OF assms]
consider "\<bar>y\<bar> = y" | "\<bar>y\<bar> = - y"
by auto
then show ?thesis
proof cases
case 1 with assms have "abs y \<in> Basis" "y = abs y" by auto
then show ?thesis ..
next
case 2
with assms have "abs y \<in> Basis" "y = - abs y" by auto
then show ?thesis ..
qed
qed
lemma subset_spec_plane[le, refine_vcg]:
"subset_spec_plane X sctn \<le> SPEC (\<lambda>b. b \<longrightarrow> X \<subseteq> plane_of sctn)"
unfolding subset_spec_plane_def
by (refine_vcg) (auto simp: plane_of_def eucl_le[where 'a='a] dest!: bspec elim!: abs_in_BasisE)
lemma subset_spec_coll_refine[le, refine_vcg]: "subset_spec_coll X Y \<le> subset_spec X Y"
unfolding subset_spec_coll_def autoref_tag_defs subset_spec_def
by (refine_vcg FORWEAK_mono_rule[where I="\<lambda>X b. b \<longrightarrow> \<Union>X \<subseteq> Y"]) auto
lemma
eventually_in_planerectI:
fixes n::"'a::executable_euclidean_space"
assumes "abs n \<in> Basis"
assumes "{l .. u} \<subseteq> plane n c" "l \<le> u"
assumes "\<And>i. i \<in> Basis \<Longrightarrow> i \<noteq> abs n \<Longrightarrow> l \<bullet> i < x \<bullet> i"
assumes "\<And>i. i \<in> Basis \<Longrightarrow> i \<noteq> abs n \<Longrightarrow> x \<bullet> i < u \<bullet> i"
shows "\<forall>\<^sub>F x in at x within plane n c. x \<in> {l .. u}"
proof -
have "\<forall>\<^sub>F x in at x within plane n c. x \<in> plane n c"
unfolding eventually_at_filter
by simp
then have "\<forall>\<^sub>F x in at x within plane n c. l \<bullet> abs n \<le> x \<bullet> abs n \<and> x \<bullet> abs n \<le> u \<bullet> abs n"
apply eventually_elim
using assms(1,2,3)
by (auto simp: elim!: abs_in_BasisE)
moreover
{
fix i assume that: "i \<in> Basis" "i \<noteq> abs n"
have "\<forall>\<^sub>F x in at x within plane n c. l \<bullet> i < x \<bullet> i" "\<forall>\<^sub>F x in at x within plane n c. x \<bullet> i < u \<bullet> i"
by (auto intro!: order_tendstoD assms tendsto_eq_intros that)
then have "\<forall>\<^sub>F x in at x within plane n c. l \<bullet> i < x \<bullet> i \<and> x \<bullet> i < u \<bullet> i"
by eventually_elim auto
} then have "\<forall>\<^sub>F x in at x within plane n c. \<forall>i \<in> Basis - {abs n}. l \<bullet> i < x \<bullet> i \<and> x \<bullet> i < u \<bullet> i"
by (auto intro!: eventually_ball_finite)
then have "\<forall>\<^sub>F x in at x within plane n c. \<forall>i \<in> Basis - {abs n}. l \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> u \<bullet> i"
by eventually_elim (auto intro!: less_imp_le)
ultimately
have "\<forall>\<^sub>F x in at x within plane n c. \<forall>i\<in>Basis. l \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> u \<bullet> i"
by eventually_elim auto
then show ?thesis
by eventually_elim (auto simp: eucl_le[where 'a='a])
qed
lemma mem_ivl_euclI: "k \<in> {c..d::'x::ordered_euclidean_space}" if "\<And>i. i \<in> Basis \<Longrightarrow> c \<bullet> i \<le> k \<bullet> i" "\<And>i. i \<in> Basis \<Longrightarrow> k \<bullet> i \<le> d \<bullet> i"
using that
by (auto simp: eucl_le[where 'a='x])
lemma op_eventually_within_sctn[le, refine_vcg]:
"op_eventually_within_sctn X sctn S \<le>
SPEC (\<lambda>b. b \<longrightarrow> (\<forall>x \<in> X. x \<in> S \<and> (\<forall>\<^sub>F x in at x within plane_of sctn. x \<in> S)))"
unfolding op_eventually_within_sctn_def
apply refine_vcg
unfolding plane_of_def autoref_tag_defs
apply (safe intro!: eventually_in_planerectI mem_ivl_euclI)
subgoal premises prems for a b c d e f g h i j k B
proof cases
assume "B = \<bar>normal sctn\<bar>"
moreover
have "c \<in> plane (normal sctn) (pstn sctn)" "k \<in> plane (normal sctn) (pstn sctn)"
using prems by auto
ultimately show "c \<bullet> B \<le> k \<bullet> B" using \<open>\<bar>normal sctn\<bar> \<in> set Basis_list\<close>
by (auto simp: elim!: abs_in_Basis_absE)
next
assume B: "B \<noteq> \<bar>normal sctn\<bar>"
have "k \<bullet> B \<in> {g \<bullet> B .. h \<bullet> B}"
using \<open>k \<in> X\<close> \<open>X \<subseteq> {g..h}\<close> \<open>B \<in> Basis\<close> by (auto simp: eucl_le[where 'a='a])
with B prems show ?thesis by (auto dest!: bspec elim!: abs_in_Basis_absE)
qed
subgoal premises prems for a b c d e f g h i j k B
proof cases
assume "B = \<bar>normal sctn\<bar>"
moreover
have "d \<in> plane (normal sctn) (pstn sctn)" "k \<in> plane (normal sctn) (pstn sctn)"
using prems by auto
ultimately show "d \<bullet> B \<ge> k \<bullet> B" using \<open>\<bar>normal sctn\<bar> \<in> set Basis_list\<close>
by (auto simp: elim!: abs_in_Basis_absE)
qed (use prems in \<open>auto elim!: abs_in_BasisE simp: eucl_le[where 'a='a] dest!: bspec subsetD\<close>)
subgoal by simp
- subgoal by (auto elim!: abs_in_BasisE simp: eucl_le[where 'a='a] dest!: bspec subsetD cong del: image_cong_simp)
- subgoal by (auto elim!: abs_in_BasisE simp: eucl_le[where 'a='a] dest!: bspec subsetD cong del: image_cong_simp)
+ subgoal using [[simproc del: defined_all]] by (auto elim!: abs_in_BasisE simp: eucl_le[where 'a='a] dest!: bspec subsetD cong del: image_cong_simp)
+ subgoal using [[simproc del: defined_all]] by (auto elim!: abs_in_BasisE simp: eucl_le[where 'a='a] dest!: bspec subsetD cong del: image_cong_simp)
done
lemma Let_unit: "Let (x::unit) f = f ()"
by auto
lemma CHECK_no_text: "CHECKs (x#ys) a = CHECKs [] a"
by auto
lemma frontier_above_halfspace:
"normal sctn \<noteq> 0 \<Longrightarrow> frontier (above_halfspace sctn) = plane_of sctn"
using frontier_halfspace_ge[of "normal sctn" "pstn sctn"]
by (auto simp: halfspace_simps plane_of_def inner_commute)
lemma
flowpipe_subset:
assumes "flowpipe X0 hl hu CX X1"
and subs: "Y0 \<subseteq> X0" "hl \<le> il" "il \<le> iu" "iu \<le> hu" "CX \<subseteq> CY" "X1 \<subseteq> Y1"
and safe: "fst ` CY \<union> fst ` Y1 \<subseteq> Csafe"
shows "flowpipe Y0 il iu CY Y1"
proof -
from assms(1) have fp: "0 \<le> hl" "hl \<le> hu" "fst ` X0 \<subseteq> Csafe" "fst ` CX \<subseteq> Csafe" "fst ` X1 \<subseteq> Csafe"
"\<forall>(x0, d0)\<in>X0. \<forall>h\<in>{hl..hu}. h \<in> existence_ivl0 x0 \<and> (flow0 x0 h, Dflow x0 h o\<^sub>L d0) \<in> X1 \<and> (\<forall>h'\<in>{0..h}. (flow0 x0 h', Dflow x0 h' o\<^sub>L d0) \<in> CX)"
by (auto simp: flowpipe_def)
then show ?thesis
unfolding flowpipe_def
apply safe
subgoal using subs by auto
subgoal using subs by auto
subgoal using subs by fastforce
subgoal using safe by auto
subgoal using safe by auto
subgoal using subs by fastforce
subgoal using subs fp by fastforce
subgoal for x0 d0 h h'
using subs fp
apply -
apply (rule subsetD, assumption)
apply (drule bspec)
apply (rule subsetD; assumption)
apply safe
apply (drule bspec[where x=h], force)
apply auto
done
done
qed
lemma poincare_mapsto_unionI:
assumes "poincare_mapsto P r U t d"
assumes "poincare_mapsto P s U u e"
shows "poincare_mapsto P (r \<union> s) U (t \<union> u) (d \<union> e)"
using assms
apply (auto simp: poincare_mapsto_def)
subgoal
apply (drule bspec, assumption)
by auto
subgoal by fastforce
done
lemma sabove_not_le_halfspace:
"x \<in> sabove_halfspace sctn \<longleftrightarrow> \<not> le_halfspace sctn x"
by (auto simp: sabove_halfspace_def le_halfspace_def gt_halfspace_def)
lemma (in c1_on_open_euclidean) flowsto_self:\<comment> \<open>TODO: move!\<close>
"0 \<in> T \<Longrightarrow> X0 \<subseteq> Z \<Longrightarrow> fst ` X0 \<subseteq> X \<Longrightarrow> flowsto X0 T Y Z"
by (force simp: flowsto_def intro!: bexI[where x=0])
lemma (in c1_on_open_euclidean) flowpipe_imp_flowsto:\<comment> \<open>TODO: move!\<close>
assumes "flowpipe X0 hl hu CX Y" "hl > 0"
shows "flowsto X0 {0<..hl} CX Y"
using assms
by (fastforce simp: flowsto_def flowpipe_def open_segment_eq_real_ivl
dest: bspec[where x=hl]
intro!: bexI[where x=hl])
lemma flowsto_source_unionI:
"flowsto X0 T A B \<Longrightarrow> flowsto Z T A B \<Longrightarrow> flowsto (X0 \<union> Z) T A B"
by (fastforce simp: flowsto_def dest: bspec)
lemma poincare_mapsto_subset:
"poincare_mapsto P X0 U CX1 X1 \<Longrightarrow> X0' \<subseteq> X0 \<Longrightarrow> X1 \<subseteq> X2 \<Longrightarrow> CX1 \<subseteq> CX2 \<Longrightarrow> fst ` X2 \<subseteq> Csafe
\<Longrightarrow> poincare_mapsto P X0' U CX2 X2"
by (force simp: poincare_mapsto_def)
lemma PDP_abs_lemma:
fixes n::"'a::executable_euclidean_space"
assumes "abs n \<in> Basis"
shows
"(x, d - (blinfun_scaleR_left (f (x)) o\<^sub>L (blinfun_scaleR_left (inverse (f x \<bullet> n)) o\<^sub>L (blinfun_inner_left n o\<^sub>L d)))) =
(x, d - (blinfun_scaleR_left (f (x)) o\<^sub>L (blinfun_scaleR_left (inverse (f x \<bullet> (abs n))) o\<^sub>L (blinfun_inner_left (abs n) o\<^sub>L d))))"
proof -
consider "n \<in> Basis" | "- n \<in> Basis"
using abs_in_Basis_absE[OF assms] assms by metis
then show ?thesis
proof cases
case 1
then show ?thesis by simp
next
case 2
define i where "i = -n"
with 2 have "i \<in> Basis" "n = -i"
by (auto simp: )
then show ?thesis
by (auto simp: inverse_eq_divide intro!: blinfun_eqI blinfun.bilinear_simps euclidean_eqI[where 'a='a])
qed
qed
lemma abs_in_BasisI:
"\<bar>n\<bar> \<in> Basis" if n: "n \<in> Basis \<or> - n \<in> Basis" for n::"'a::executable_euclidean_space"
proof -
consider "n \<in> Basis" | "- n \<in> Basis"
using n by auto
then show ?thesis
proof cases
case 1
then show ?thesis by simp
next
case 2
define i where "i = -n"
with 2 have "i \<in> Basis" "n = -i"
by (auto simp: )
then show ?thesis
by (auto simp: inverse_eq_divide intro!: blinfun_eqI blinfun.bilinear_simps euclidean_eqI[where 'a='a])
qed
qed
lemma flowsto_poincareD:
assumes f: "flowsto X0 T CX X1"
assumes X1: "fst ` X1 \<subseteq> P"
assumes P: "(P \<times> UNIV) \<inter> CX = {}" "closed P"
assumes pos: "\<And>t. t \<in> T \<Longrightarrow> t > 0"
assumes x0: "x0 \<in> fst ` X0"
assumes "fst ` X1 \<subseteq> K"
shows returns_to_flowstoI: "returns_to P x0"
and poincare_map_mem_flowstoI: "poincare_map P x0 \<in> K"
proof -
from x0 obtain d where x0d: "(x0, d) \<in> X0" by auto
from flowstoE[OF f x0d] obtain h
where h:
"h \<in> T"
"h \<in> existence_ivl0 x0"
"(flow0 x0 h, Dflow x0 h o\<^sub>L d) \<in> X1"
and CX: "\<And>h'. h' \<in> {0<--<h} \<Longrightarrow> (flow0 x0 h', Dflow x0 h' o\<^sub>L d) \<in> CX"
by auto
have "h > 0" by (auto intro!: pos h)
have "flow0 x0 h \<in> P" using X1 h by auto
have "\<forall>\<^sub>F t in at_right 0. t \<in> {0<..<h}"
using order_tendstoD(2)[OF tendsto_ident_at \<open>0 < h\<close>, of "{0<..}"]
by (auto simp: eventually_at_filter)
then have "\<forall>\<^sub>F t in at_right 0. flow0 x0 t \<in> fst ` CX"
by eventually_elim (use CX \<open>0 < h\<close> open_segment_eq_real_ivl in auto)
then have evnP: "\<forall>\<^sub>F t in at_right 0. flow0 x0 t \<notin> P"
by eventually_elim (use P in force)
from \<open>h > 0\<close> h(2) \<open>flow0 x0 h \<in> P\<close> evnP P(2) show "returns_to P x0"
by (rule returns_toI)
have nin_P: "0 < s \<Longrightarrow> s < h \<Longrightarrow> flow0 x0 s \<notin> P" for s
using CX[of s] P by (auto simp: open_segment_eq_real_ivl)
have "return_time P x0 = h"
using h X1
by (auto intro!: return_time_eqI \<open>0 < h\<close> h assms simp: nin_P)
then have "poincare_map P x0 = flow0 x0 h" by (auto simp: poincare_map_def)
also have "\<dots> \<in> fst ` X1" using h by auto
also note \<open>_ \<subseteq> K\<close>
finally show "poincare_map P x0 \<in> K" .
qed
lemma
inner_abs_Basis_eq_zero_iff:
"abs n \<in> Basis \<Longrightarrow> x \<bullet> \<bar>n\<bar> = 0 \<longleftrightarrow> x \<bullet> n = 0" for n::"'a::executable_euclidean_space"
by (auto simp: elim!: abs_in_BasisE)
lemmas [simp] = sbelow_halfspaces_insert
lemma Int_Un_eq_emptyI: "a \<inter> (b \<union> c) = {}" if "a \<inter> b = {}" "a \<inter> c = {}"
using that by auto
lemma cancel_times_UNIV_subset: "A \<times> UNIV \<subseteq> B \<times> UNIV \<longleftrightarrow> A \<subseteq> B"
by auto
lemma split_spec_coll_spec[le,refine_vcg]:
"split_spec_coll X \<le> SPEC (\<lambda>(A, B). X \<subseteq> A \<union> B)"
unfolding split_spec_coll_def
by (refine_vcg)
lemma Un_snd_sing_Pair_eq:
"(e, f) \<in> a \<Longrightarrow> f \<union> (\<Union>x\<in>a - {(e, f)}. snd x) = (\<Union>x\<in>a. snd x)"
by force
lemma let_unit: "Let X y = y ()" by simp
lemma (in c1_on_open_euclidean) flowpipe_imp_flowsto_nonneg:\<comment> \<open>TODO: move!\<close>
assumes "flowpipe X0 hl hu CX Y"
shows "flowsto X0 {0..} CX Y"
using assms
by (fastforce simp: flowsto_def flowpipe_def open_segment_eq_real_ivl
dest: bspec[where x=hl]
intro!: bexI[where x=hl])
lemma subset_DiffI: "A \<subseteq> B \<Longrightarrow> A \<inter> C = {} \<Longrightarrow> A \<subseteq> B - C"
by auto
lemma flowsto_source_Union: "flowsto (\<Union>x\<in>R. X0 x) T CX X1"
if "\<And>x. x \<in> R \<Longrightarrow> flowsto (X0 x) T CX X1"
using that
by (auto simp: flowsto_def)
lemma times_subset_iff: "A \<times> B \<subseteq> C \<times> E \<longleftrightarrow> A = {} \<or> B = {} \<or> A \<subseteq> C \<and> B \<subseteq> E"
by auto
lemma
flowsto_UnionE:
assumes "finite Gs"
assumes "flowsto X T CX (\<Union>Gs)"
obtains XGs where "\<And>X G. (X, G) \<in> XGs \<Longrightarrow> flowsto X T CX G" "Gs = snd ` XGs" "X = \<Union>(fst ` XGs)"
apply atomize_elim
using assms
proof (induction arbitrary: X)
case empty
then show ?case by auto
next
case (insert x F)
from insert.prems obtain X1 X2 where X: "X = X1 \<union> X2" and X1: "flowsto X1 T CX x" and X2: "flowsto X2 T CX (\<Union>F)"
by (auto elim!: flowsto_unionE)
from insert.IH[OF X2] obtain XGs where XGs:
"\<And>X G. (X, G) \<in> XGs \<Longrightarrow> flowsto X T CX G" "F = snd ` XGs" "X2 = (\<Union>a\<in>XGs. fst a)"
by auto
then show ?case
using X X1 X2
by (intro exI[where x="insert (X1, x) XGs"]) auto
qed
lemma flowsto_Union_funE:
assumes "finite Gs"
assumes "flowsto X T CX (\<Union>Gs)"
obtains f where "\<And>G. G \<in> Gs \<Longrightarrow> flowsto (f G) T CX G" "X = \<Union>(f ` Gs)"
apply atomize_elim
using assms
proof (induction arbitrary: X)
case empty
then show ?case by auto
next
case (insert x F)
from insert.prems obtain X1 X2 where X: "X = X1 \<union> X2" and X1: "flowsto X1 T CX x" and X2: "flowsto X2 T CX (\<Union>F)"
by (auto elim!: flowsto_unionE)
from insert.IH[OF X2] obtain f where f:
"\<And>G. G \<in> F \<Longrightarrow> flowsto (f G) T CX G" "X2 = (\<Union>a\<in>F. f a)"
by auto
then show ?case
using X X1 X2 insert.hyps
by (intro exI[where x="f (x:=X1)"]) (auto split: if_splits)
qed
lemma flowsto_Union_Un_funE:
assumes "flowsto X T CX (\<Union>Gs \<union> trap)"
assumes "finite Gs" "Gs \<noteq> {}"
obtains f where "\<And>G. G \<in> Gs \<Longrightarrow> flowsto (f G) T CX (G \<union> trap)" "X = \<Union>(f ` Gs)"
proof -
from assms have "flowsto X T CX (\<Union>g \<in> Gs. g \<union> trap)" by auto
from flowsto_Union_funE[OF finite_imageI[OF \<open>finite Gs\<close>] this]
obtain f where f: "\<And>G. G \<in> (\<lambda>g. g \<union> trap) ` Gs \<Longrightarrow> flowsto (f G) T CX G"
"X = \<Union> (f ` ((\<lambda>g. g \<union> trap) ` Gs))"
by auto
define f' where "f' g = f (g \<union> trap)" for g
have "G \<in> Gs \<Longrightarrow> flowsto (f' G) T CX (G \<union> trap)" for G
using f(1)[of "G \<union> trap"]
by (auto simp: f'_def)
moreover
have "X = \<Union>(f' ` Gs)"
using f by (auto simp: f'_def)
ultimately show ?thesis ..
qed
lemma flowsto_Diff_to_Union_funE:
assumes "flowsto (X - trap) T CX (\<Union>Gs)"
assumes "finite Gs"
obtains f where "\<And>G. G \<in> Gs \<Longrightarrow> flowsto (f G - trap) T CX G" "Gs \<noteq> {} \<Longrightarrow> X = \<Union>(f ` Gs)"
apply atomize_elim
using assms(2,1)
proof (induction arbitrary: X)
case empty then show ?case by simp
next
case (insert x F)
from insert.prems obtain X1 X2 where X: "X - trap = X1 \<union> X2" and X1: "flowsto (X1) T CX x" and X2: "flowsto X2 T CX (\<Union>F)"
by (auto elim!: flowsto_unionE)
then have "X1 = X1 - trap" "X2 = X2 - trap" by auto
then have X1': "flowsto (X1 - trap) T CX x" and X2': "flowsto (X2 - trap) T CX (\<Union>F)"
using X1 X2 by auto
from insert.IH[OF X2'] obtain f where f: "\<And>G. G \<in> F \<Longrightarrow> flowsto (f G - trap) T CX G" "F \<noteq> {} \<Longrightarrow> X2 = (\<Union>a\<in>F. f a)"
by auto
show ?case
apply (cases "F = {}")
subgoal using f X X1 X2 X1' X2' insert.hyps insert.prems by auto
subgoal
apply (rule exI[where x="f (x:=X1 \<union> (trap \<inter> X))"])
apply auto
subgoal
using X1
by (rule flowsto_subset) auto
subgoal using X X1 X2 insert.hyps by auto
subgoal using f X X1 X2 insert.hyps by auto
subgoal using f X X1 X2 insert.hyps by auto
subgoal using f X X1 X2 X1' X2' insert.hyps insert.prems by auto
subgoal using f X X1 X2 X1' X2' insert.hyps insert.prems insert by auto
subgoal using f X X1 X2 insert.hyps by (auto split: if_splits)
done
done
qed
lemma refine_case_list[refine_vcg]:
assumes "xs = [] \<Longrightarrow> f \<le> SPEC P"
assumes "\<And>y ys. xs = y # ys \<Longrightarrow> g y ys \<le> SPEC P"
shows "(case xs of [] \<Rightarrow> f | (x#xs) \<Rightarrow> g x xs) \<le> SPEC P"
using assms
by (auto split: list.splits)
lemma flowsto_stays_sbelow:
assumes "flowsto X0 {0<..} CXS X1"
assumes "fst ` X0 \<subseteq> below_halfspace sctn"
assumes "\<And>x d. (x, d) \<in> CXS \<Longrightarrow> ode x \<bullet> normal sctn < 0"
shows "flowsto X0 {0<..} (CXS \<inter> sbelow_halfspace sctn \<times> UNIV) X1"
unfolding flowsto_def
proof safe
fix x d assume "(x, d) \<in> X0"
with assms obtain t where
"t>0" "t \<in> existence_ivl0 x" "(\<forall>s\<in>{0<..<t}. (flow0 x s, Dflow x s o\<^sub>L d) \<in> CXS)"
"(flow0 x t, Dflow x t o\<^sub>L d) \<in> X1"
by (auto simp: flowsto_def subset_iff open_segment_eq_real_ivl)
moreover
have "\<forall>s\<in>{0<..<t}. flow0 x s \<in> sbelow_halfspace sctn"
proof (rule ccontr, clarsimp)
fix s assume s: "flow0 x s \<notin> sbelow_halfspace sctn" "0 < s" "s < t"
let ?f = "\<lambda>t. flow0 x t \<bullet> normal sctn - pstn sctn"
let ?f' = "\<lambda>t dx. dx * (ode (flow0 x t) \<bullet> normal sctn)"
have "\<exists>xa\<in>{0<..<s}. ?f s - ?f 0 = ?f' xa (s - 0)"
by (rule mvt[OF \<open>0 < s\<close>, of ?f ?f'])
(use ivl_subset_existence_ivl[OF \<open>t \<in> existence_ivl0 x\<close>] \<open>0 < s\<close> \<open>s < t\<close> in
\<open>auto intro!: continuous_intros derivative_eq_intros flow_has_derivative
simp: flowderiv_def blinfun.bilinear_simps\<close>)
then obtain s' where "?f s - ?f 0 = ?f' s' (s - 0)" "0 < s'" "s' < s"
by (auto simp: algebra_simps)
note this(1)
also
have "(flow0 x s', Dflow x s' o\<^sub>L d )\<in> CXS"
using \<open>0 < s'\<close> \<open>\<forall>s\<in>{0<..<t}. _ \<in> CXS\<close> \<open>s < t\<close> \<open>s' < s\<close> by auto
then have "?f' s' (s - 0) < 0"
using assms \<open>(x, d) \<in> X0\<close> \<open>0 < s\<close>
by (auto simp: flowsto_def halfspace_simps algebra_simps subset_iff intro!: mult_pos_neg)
finally have 1: "?f s < ?f 0"
by simp
also have "?f 0 \<le> 0"
using \<open>(x, d) \<in> X0\<close> assms mem_existence_ivl_iv_defined[OF \<open>t \<in> existence_ivl0 _\<close>]
by (auto simp: halfspace_simps subset_iff)
finally have "?f s < 0" .
moreover from s have "0 \<le> ?f s"
by (auto simp: halfspace_simps not_less)
ultimately show False by simp
qed
ultimately
show "\<exists>t\<in>{0<..}. t \<in> existence_ivl0 x \<and> (flow0 x t, Dflow x t o\<^sub>L d) \<in> X1 \<and>
(\<forall>s\<in>{0<--<t}. (flow0 x s, Dflow x s o\<^sub>L d) \<in> CXS \<inter> sbelow_halfspace sctn \<times> UNIV)"
by (auto intro!: simp: open_segment_eq_real_ivl)
qed
lemma poincare_mapsto_Union: "poincare_mapsto P (\<Union>xa) S CXS PS"
if "\<And>x. x \<in> xa \<Longrightarrow> poincare_mapsto P x S CXS PS"
by (force simp: poincare_mapsto_def dest!: that)
lemma diff_subset: "(\<Union>x\<in>xa. f x) - (\<Union>x\<in>xa. g x) \<subseteq> (\<Union>x\<in>xa. f x - g x)"
by auto
lemma poincare_mapsto_imp_flowsto:
assumes "poincare_mapsto P X0 S CX X1"
assumes "closed P"
shows "flowsto X0 {0<..} (CX \<times> UNIV) (fst ` X1 \<times> UNIV)"
unfolding flowsto_def
proof safe
fix x0 d0 assume "(x0, d0) \<in> X0"
with assms obtain D where D:
"returns_to P x0"
"fst ` X0 \<subseteq> S"
"return_time P differentiable at x0 within S"
"(poincare_map P has_derivative blinfun_apply D) (at x0 within S)"
"(flow0 x0 (return_time P x0), D o\<^sub>L d0) \<in> X1"
"\<And>t. 0 < t \<Longrightarrow> t < return_time P x0 \<Longrightarrow> flow0 x0 t \<in> CX"
by (auto simp: poincare_mapsto_def poincare_map_def)
show "\<exists>h\<in>{0<..}.
h \<in> existence_ivl0 x0 \<and> (flow0 x0 h, Dflow x0 h o\<^sub>L d0) \<in> fst ` X1 \<times> UNIV \<and>
(\<forall>h'\<in>{0<--<h}. (flow0 x0 h', Dflow x0 h' o\<^sub>L d0) \<in> CX \<times> UNIV)"
by (auto intro!: bexI[where x="return_time P x0"] return_time_exivl D assms return_time_pos
simp: open_segment_eq_real_ivl)
qed
lemma flowsto_poincare_mapsto_trans_flowsto:
assumes "flowsto X0 T CX X1'"
assumes "poincare_mapsto P X1 S CY X2"
assumes "closed P"
assumes "fst ` X1' \<subseteq> fst ` X1"
assumes "X1' \<union> CX \<union> CY \<times> UNIV \<subseteq> CZ"
assumes "\<And>t. t \<in> T \<Longrightarrow> t \<ge> 0"
shows "flowsto X0 {0<..} CZ (fst ` X2 \<times> UNIV)"
proof -
have X1D: "(a, b) \<in> X1' \<Longrightarrow> \<exists>c. (a, c) \<in> X1" for a b using assms(4) by force
from poincare_mapsto_imp_flowsto[OF assms(2,3)]
have "flowsto X1 {0<..} (CY \<times> UNIV) (fst ` X2 \<times> UNIV)" .
then have "flowsto X1' {0<..} (CY \<times> UNIV) (fst ` X2 \<times> UNIV)"
by (auto simp: flowsto_def dest!: X1D)
from flowsto_trans[OF assms(1) this]
show ?thesis
apply (rule flowsto_subset)
using assms
by (auto intro!: add_nonneg_pos)
qed
lemma eq_blinfun_inner_left[intro]:
"(\<lambda>x. x \<bullet> n) = blinfun_apply (blinfun_inner_left n)"
by force
lemma flowsto_union_DiffE:
assumes "flowsto X0 T CX (Y \<union> Z)"
obtains X1 where "X1 \<subseteq> X0" "flowsto X1 T CX Y" "flowsto (X0 - X1) T CX Z"
proof -
let ?X1 = "{x\<in>X0. flowsto {x} T CX Y}"
from assms have "?X1 \<subseteq> X0" "flowsto ?X1 T CX Y" "flowsto (X0 - ?X1) T CX Z"
by (auto simp: flowsto_def)
thus ?thesis ..
qed
lemma eucl_less_le_trans:
fixes a b::"'a::executable_euclidean_space"
shows "eucl_less a b \<Longrightarrow> b \<le> c \<Longrightarrow> eucl_less a c"
by (force simp: eucl_less_def[where 'a='a] eucl_le[where 'a='a])
lemma le_eucl_less_trans:
fixes a b::"'a::executable_euclidean_space"
shows "a \<le> b \<Longrightarrow> eucl_less b c \<Longrightarrow> eucl_less a c"
by (force simp: eucl_less_def[where 'a='a] eucl_le[where 'a='a])
lemma flowsto_source_UnionI:
assumes "\<And>i. i \<in> I \<Longrightarrow> flowsto i T CXS (f i)"
shows "flowsto (\<Union>I) T CXS (\<Union>(f ` I))"
apply (auto simp: flowsto_def)
subgoal premises prems for y a b
using assms[unfolded flowsto_def, OF \<open>y \<in> I\<close>, rule_format, OF \<open>_ \<in> y\<close>] prems
by auto
done
lemma poincare_mapsto_UnionI:
assumes pm[unfolded poincare_mapsto_def, rule_format]: "\<And>i. i \<in> I \<Longrightarrow> poincare_mapsto p (X0 i) S (CX i) (X1 i)"
assumes R: "\<And>i x. i \<in> I \<Longrightarrow> x \<in> X1 i \<Longrightarrow> x \<in> R"
shows "poincare_mapsto p (\<Union>x\<in>I. X0 x) S ((\<Union>x\<in>I. CX x)) R"
unfolding poincare_mapsto_def
proof (safe del: conjI, goal_cases)
case (1 x0 d0 i)
moreover
have "fst ` \<Union>(X0 ` I) \<subseteq> S"
proof (safe, goal_cases)
case (1 _ x0 d0 i)
from this pm[OF 1]
show ?case by auto
qed
ultimately show ?case using pm[OF 1]
by (auto intro!: R)
qed
lemma tendsto_at_top_translateI:
assumes "(f \<longlongrightarrow> l) (at_top::real filter)"
shows "((\<lambda>x. f (x + t)::'a::topological_space) \<longlongrightarrow> l) at_top"
proof (rule topological_tendstoI)
fix S::"'a set" assume "open S" "l \<in> S"
from topological_tendstoD[OF assms this]
obtain N where "\<And>n. n \<ge> N \<Longrightarrow> f n \<in> S" by (auto simp: eventually_at_top_linorder)
then have "\<And>n. n \<ge> N - t \<Longrightarrow> f (n + t) \<in> S" by auto
then show "\<forall>\<^sub>F x in at_top. f (x + t) \<in> S"
unfolding eventually_at_top_linorder
by blast
qed
lemma tendsto_at_top_translate_iff:
"((\<lambda>x. f (x + t)::'a::topological_space) \<longlongrightarrow> l) at_top \<longleftrightarrow> (f \<longlongrightarrow> l) (at_top::real filter)"
using tendsto_at_top_translateI[of f l t]
tendsto_at_top_translateI[of "\<lambda>x. f (x + t)" l "- t"]
by auto
lemma stable_on_mono:
"stable_on A B" if "stable_on C B" "A \<subseteq> C"
using that
unfolding stable_on_def
by fastforce
lemma
flowsto_mapsto_avoid_trap:
assumes "flowsto (X0 - trap \<times> UNIV) {0<..} CX P"
assumes trapprop: "stable_on (fst ` (CX \<union> P)) trap"
shows "flowsto (X0 - trap \<times> UNIV) {0<..} CX (P - trap \<times> UNIV)"
unfolding flowsto_def
proof (rule, goal_cases)
case (1 xd)
from assms(1)[unfolded flowsto_def, rule_format, OF this] obtain h x0 d0 where
"xd = (x0, d0)" "0 < h"
"h \<in> existence_ivl0 (x0)"
"(flow0 x0 h, Dflow x0 h o\<^sub>L d0) \<in> P"
"(\<forall>h'\<in>{0<--<h}. (flow0 x0 h', Dflow x0 h' o\<^sub>L d0) \<in> CX)"
by auto
then show ?case
using 1 trapprop
apply (auto intro!: bexI[where x=h] dest!: stable_onD simp: open_segment_eq_real_ivl image_Un)
subgoal for s by (cases "s = h") auto
done
qed
end
lemma map_prod_def': "map_prod f g x = (f (fst x), g (snd x))"
by (cases x) auto
lemmas rel_prod_br = br_rel_prod
lemmas lvivl_relI = lv_relivl_relI
end
diff --git a/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis_C1.thy b/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis_C1.thy
--- a/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis_C1.thy
+++ b/thys/Ordinary_Differential_Equations/Numerics/Refine_Reachability_Analysis_C1.thy
@@ -1,2893 +1,2896 @@
theory Refine_Reachability_Analysis_C1
imports
Abstract_Reachability_Analysis_C1
Refine_Reachability_Analysis
begin
lemma fst_flow1_of_vec1[simp]: "fst (flow1_of_vec1 x) = fst x"
by (auto simp: flow1_of_vec1_def)
lemma fst_vec1_of_flow[simp]: "fst (vec1_of_flow1 x) = fst x"
by (auto simp: vec1_of_flow1_def)
context approximate_sets_ode'
begin
lemma poincare_mapsto_scaleR2I:
"poincare_mapsto P (scaleR2 x1 x2 baa) UNIV x1b (scaleR2 x1 x2 aca)"
if "poincare_mapsto P (baa) UNIV x1b (aca)"
using that
apply (auto simp: poincare_mapsto_def scaleR2_def image_def vimage_def)
apply (drule bspec, assumption)
apply auto
apply (rule exI, rule conjI, assumption)
apply (rule exI, rule conjI, assumption, rule conjI, assumption)
apply (rule bexI) prefer 2 apply assumption
apply (auto simp: scaleR_blinfun_compose_right)
done
context includes ode_ops.lifting begin
lemma var_safe_form_eq[simp]: "var.safe_form = safe_form"
unfolding var.safe_form_def
by transfer (auto simp: var_ode_ops_def safe_form_def)
lemma var_ode_e: "var.ode_e = ode_e'"
unfolding var.ode_e_def
by transfer (auto simp: var_ode_ops_def)
end
lemma wd_imp_var_wd[refine_vcg, intro]: "wd (TYPE('n rvec)) \<Longrightarrow> var.wd (TYPE('n::enum vec1))"
unfolding var.wd_def
by (auto simp: wd_def length_concat o_def sum_list_distinct_conv_sum_set
concat_map_map_index var_ode_e D_def ode_e'_def
intro!: max_Var_floatariths_mmult_fa[le] max_Var_floatariths_mapI
max_Var_floatarith_FDERIV_floatarith[le]
max_Var_floatariths_fold_const_fa[le]
max_Var_floatarith_le_max_Var_floatariths_nthI
max_Var_floatariths_list_updateI max_Var_floatariths_replicateI)
lemma safe_eq:
assumes "wd TYPE('n::enum rvec)"
shows "var.Csafe = ((Csafe \<times> UNIV)::'n vec1 set)"
using assms var.wdD[OF wd_imp_var_wd[OF assms]] wdD[OF assms]
unfolding var.safe_def safe_def var.wd_def wd_def var.Csafe_def Csafe_def
unfolding ode_e'_def var_ode_e
apply (auto simp: D_def)
subgoal
apply (subst interpret_form_max_Var_cong) prefer 2 apply assumption
by (auto simp: nth_Basis_list_prod)
subgoal for a b
apply (drule isFDERIV_appendD1)
apply simp apply simp apply (auto intro!: max_Var_floatariths_fold_const_fa[le])[]
apply (rule isFDERIV_max_Var_congI, assumption)
by (auto simp: nth_Basis_list_prod)
subgoal
apply (subst interpret_form_max_Var_cong) prefer 2 apply assumption
by (auto simp: nth_Basis_list_prod)
subgoal for a b
apply (rule isFDERIV_appendI1)
apply (rule isFDERIV_max_Var_congI, assumption)
apply (auto simp: nth_Basis_list_prod)
apply (auto simp: isFDERIV_def FDERIV_floatariths_def in_set_conv_nth isDERIV_inner_iff
length_concat o_def sum_list_distinct_conv_sum_set concat_map_map_index
intro!: isDERIV_FDERIV_floatarith isDERIV_mmult_fa_nth)
apply (rule isDERIV_max_Var_floatarithI[where ys="list_of_eucl a"])
subgoal for i j k
apply (cases "i < CARD('n)")
subgoal by auto
subgoal apply (rule isDERIV_max_VarI)
apply (rule max_Var_floatarith_le_max_Var_floatariths_nthI)
apply force
apply auto
done
done
subgoal for i j k l by (auto dest!: max_Var_floatariths_lessI simp: nth_Basis_list_prod)
subgoal by (auto simp: nth_list_update)
done
done
lemma
var_ode_eq:
fixes x::"'n::enum vec1"
assumes "wd TYPE('n rvec)" and [simp]: "(fst x) \<in> Csafe"
shows "var.ode x = (ode (fst x), matrix (ode_d1 (fst x)) ** snd x)"
proof -
have "interpret_floatariths ode_e (list_of_eucl x) =
interpret_floatariths ode_e (list_of_eucl (fst x))"
apply (rule interpret_floatariths_max_Var_cong)
using wdD[OF \<open>wd _\<close>]
by (auto simp: list_of_eucl_nth_if nth_Basis_list_prod inner_prod_def)
moreover
have "eucl_of_list
(interpret_floatariths
(mmult_fa D D D
(concat (map (\<lambda>j. map (\<lambda>i. FDERIV_floatarith (ode_e ! j) [0..<D] ((replicate D 0)[i := 1])) [0..<D]) [0..<D]))
(map floatarith.Var [D..<D + D * D])) (list_of_eucl x)) =
matrix (blinfun_apply (ode_d 0 (fst x) 0)) ** snd x"
unfolding matrix_eq
apply auto
apply (subst matrix_vector_mul_assoc[symmetric])
apply (subst matrix_works)
subgoal by (auto simp: linear_matrix_vector_mul_eq
intro!: bounded_linear.linear blinfun.bounded_linear_right)
apply (subst einterpret_mmult_fa[where 'n='n and 'm = 'n and 'l='n])
subgoal by (simp add: wdD[OF \<open>wd _\<close>])
subgoal by (simp add: length_concat o_def sum_list_distinct_conv_sum_set wdD[OF \<open>wd _\<close>])
subgoal by (simp add: length_concat o_def sum_list_distinct_conv_sum_set wdD[OF \<open>wd _\<close>])
subgoal for v
proof -
have v: "einterpret (map floatarith.Var [D..<D + D * D]) (list_of_eucl x) *v v = snd x *v v"
apply (vector matrix_vector_mult_def)
apply (simp add: vec_nth_eq_list_of_eucl2 wdD[OF \<open>wd _\<close>])
apply (auto simp: vec_nth_eq_list_of_eucl1 sum_index_enum_eq)
apply (subst sum_index_enum_eq)+
apply (rule sum.cong)
by (auto simp: nth_Basis_list_prod prod_eq_iff inner_prod_def)
show ?thesis
unfolding matrix_vector_mul_assoc[symmetric]
apply (subst v)
apply (auto simp: concat_map_map_index vec_nth_eq_list_of_eucl2)
apply (subst eucl_of_list_list_of_eucl[of "snd x *v v", symmetric])
apply (subst (2) eucl_of_list_list_of_eucl[of "snd x *v v", symmetric])
apply (subst eucl_of_list_matrix_vector_mult_eq_sum_nth_Basis_list)
subgoal by (simp add: length_concat o_def sum_list_distinct_conv_sum_set wdD[OF \<open>wd _\<close>])
subgoal by simp
apply (subst blinfun_apply_eq_sum)
apply (auto simp: vec_nth_eq_list_of_eucl1 sum_index_enum_eq)
apply (auto simp: scaleR_sum_left ode_d.rep_eq intro!: sum.cong[OF refl])
apply (auto simp: ode_d_raw_def wdD[OF \<open>wd _\<close>] eucl_of_list_inner )
apply (auto simp: ode_d_expr_def FDERIV_floatariths_def wdD[OF \<open>wd _\<close>] )
apply (rule interpret_floatarith_FDERIV_floatarith_cong)
subgoal for x y i
using wdD[OF \<open>wd _\<close>]
by (auto simp add: nth_append inner_prod_def
nth_Basis_list_prod dest!: max_Var_floatariths_lessI)
subgoal by auto
subgoal by auto
subgoal
apply (auto simp: wdD[OF \<open>wd _\<close>] nth_list_update inner_Basis intro!: nth_equalityI)
by (metis \<open>length (list_of_eucl (snd x *v v)) = CARD('n)\<close> index_Basis_list_nth length_list_of_eucl)
done
qed
done
ultimately show ?thesis
unfolding var.ode_def ode_def
unfolding ode_e'_def var_ode_e
by (auto simp: wdD[OF \<open>wd _\<close>] ode_d1_def intro!: euclidean_eqI[where 'a="'n vec1"])
qed
lemma var_existence_ivl_imp_existence_ivl:
fixes x::"'n::enum vec1"
assumes wd: "wd TYPE('n rvec)"
assumes t: "t \<in> var.existence_ivl0 x"
shows "t \<in> existence_ivl0 (fst x)"
proof (rule existence_ivl_maximal_segment)
from var.flow_solves_ode[OF UNIV_I var.mem_existence_ivl_iv_defined(2), OF t]
have D: "(var.flow0 x solves_ode (\<lambda>_. var.ode)) {0--t} (var.Csafe)"
apply (rule solves_ode_on_subset)
apply (rule var.closed_segment_subset_existence_ivl)
apply (rule t)
apply simp
done
show "((\<lambda>t. fst (var.flow0 x t)) solves_ode (\<lambda>_. ode)) {0--t} (Csafe)"
using var.closed_segment_subset_existence_ivl[OF t]
apply (auto simp: has_vderiv_on_def has_vector_derivative_def subset_iff
intro!: solves_odeI derivative_eq_intros)
apply (rule refl)
apply (rule refl)
apply (rule refl)
apply (auto simp: var.flowderiv_def )
apply (subst var_ode_eq[OF wd(1)])
apply (auto simp: blinfun.bilinear_simps)
subgoal for s
using solves_odeD(2)[OF D, of s]
by (subst(asm) (3) safe_eq[OF wd]) (auto )
subgoal for s
using solves_odeD(2)[OF D, of s]
by (subst(asm) (3) safe_eq[OF wd]) (auto )
done
next
show "fst (var.flow0 x 0) = fst x"
apply (subst var.flow_initial_time)
apply simp
apply (rule var.mem_existence_ivl_iv_defined[OF t])
apply auto
done
qed simp
lemma existence_ivl_imp_var_existence_ivl:
fixes x::"'n::enum rvec"
assumes wd: "wd TYPE('n rvec)"
assumes t: "t \<in> existence_ivl0 x"
shows "t \<in> var.existence_ivl0 ((x, W)::'n vec1)"
proof (rule var.existence_ivl_maximal_segment)
from flow_solves_ode[OF UNIV_I mem_existence_ivl_iv_defined(2), OF t]
have D: "(flow0 x solves_ode (\<lambda>_. ode)) {0--t} (Csafe)"
apply (rule solves_ode_on_subset)
apply (rule closed_segment_subset_existence_ivl)
apply (rule t)
apply simp
done
show "((\<lambda>t. (flow0 x t, matrix (Dflow x t) ** W)) solves_ode (\<lambda>_. var.ode)) {0--t} (var.Csafe)"
using closed_segment_subset_existence_ivl[OF t]
apply (auto simp: has_vderiv_on_def has_vector_derivative_def subset_iff
intro!: solves_odeI derivative_eq_intros)
apply (rule refl)
apply (rule refl)
apply (rule refl)
apply (rule has_derivative_at_withinI)
apply (rule Dflow_has_derivative)
apply force
apply (rule refl)
apply (auto simp: flowderiv_def )
apply (subst var_ode_eq)
apply (auto simp: blinfun.bilinear_simps matrix_blinfun_compose wd
intro!: ext)
subgoal for s h
unfolding matrix_scaleR matrix_blinfun_compose matrix_mul_assoc matrix_scaleR_right ..
subgoal for s
using solves_odeD(2)[OF D, of s] safe_eq[OF wd]
by auto
done
next
have "x \<in> Csafe" by rule fact
then show "(flow0 x 0, matrix (blinfun_apply (Dflow x 0)) ** W) = (x, W)"
apply (auto )
apply (vector matrix_def matrix_matrix_mult_def axis_def)
by (auto simp: if_distrib if_distribR cong: if_cong)
qed auto
theorem var_existence_ivl0_eq_existence_ivl0:
fixes x::"'n::enum vec1"
assumes wd: "wd TYPE('n rvec)"
shows "var.existence_ivl0 (x::'n vec1) = existence_ivl0 (fst x)"
apply safe
subgoal by (rule var_existence_ivl_imp_existence_ivl[OF wd, of _ "x", simplified], simp)
subgoal
by (rule existence_ivl_imp_var_existence_ivl[OF wd, of _ "fst x" "snd x", unfolded prod.collapse])
done
theorem var_flow_eq_flow_Dflow:
fixes x::"'n::enum vec1"
assumes wd: "wd TYPE('n rvec)"
assumes t: "t \<in> var.existence_ivl0 x"
shows "var.flow0 x t = vec1_of_flow1 (flow0 (fst x) t, Dflow (fst x) t o\<^sub>L blinfun_of_vmatrix (snd x)) "
proof -
have x: "x \<in> var.Csafe"
by (rule var.mem_existence_ivl_iv_defined[OF t])
then have "fst x \<in> Csafe"
by (subst (asm) safe_eq[OF wd]) auto
then have sx[simp]: "(fst x) \<in> Csafe" by simp
show ?thesis
proof (rule var.flow_unique_on[OF t])
show "vec1_of_flow1 (flow0 (fst x) 0, Dflow (fst x) 0 o\<^sub>L blinfun_of_vmatrix (snd x)) = x"
by (auto simp: vec1_of_flow1_def x)
show "((\<lambda>a. vec1_of_flow1 (flow0 (fst x) a, Dflow (fst x) a o\<^sub>L blinfun_of_vmatrix (snd x))) has_vderiv_on
(\<lambda>t. var.ode (vec1_of_flow1 (flow0 (fst x) t, Dflow (fst x) t o\<^sub>L blinfun_of_vmatrix (snd x)))))
(var.existence_ivl0 x)"
apply (auto simp: has_vderiv_on_def has_vector_derivative_def vec1_of_flow1_def
at_within_open[OF _ var.open_existence_ivl] flowderiv_def
intro!: derivative_eq_intros var_existence_ivl_imp_existence_ivl[OF wd]
Dflow_has_derivative ext)
apply (subst var_ode_eq[OF wd(1)])
apply (auto simp: blinfun.bilinear_simps)
subgoal for t
using flow_in_domain[of t "fst x"]
by (simp add: var_existence_ivl_imp_existence_ivl[OF wd])
subgoal for t h
by (simp add: matrix_blinfun_compose matrix_scaleR matrix_mul_assoc matrix_scaleR_right)
done
fix t
assume "t \<in> var.existence_ivl0 x"
then show "vec1_of_flow1 (flow0 (fst x) t, Dflow (fst x) t o\<^sub>L blinfun_of_vmatrix (snd x)) \<in> var.Csafe"
by (subst safe_eq[OF wd])
(auto simp: vec1_of_flow1_def dest!: var_existence_ivl_imp_existence_ivl[OF wd]
flow_in_domain)
qed
qed
theorem flow_Dflow_eq_var_flow:
fixes x::"'n::enum rvec"
assumes wd: "wd TYPE('n rvec)"
assumes t: "t \<in> existence_ivl0 x"
shows "(flow0 x t, Dflow x t o\<^sub>L W) = flow1_of_vec1 (var.flow0 (x, matrix W) t::'n vec1)"
using var_flow_eq_flow_Dflow[OF wd existence_ivl_imp_var_existence_ivl[OF wd t]]
unfolding var_flow_eq_flow_Dflow[OF wd existence_ivl_imp_var_existence_ivl[OF wd t]]
by (auto simp: flow1_of_vec1_def vec1_of_flow1_def)
context includes blinfun.lifting begin
lemma flow1_of_vec1_vec1_of_flow1[simp]:
"flow1_of_vec1 (vec1_of_flow1 X) = X"
unfolding vec1_of_flow1_def flow1_of_vec1_def
by (transfer) auto
end
lemma
var_flowpipe0_flowpipe:
assumes wd: "wd TYPE('n::enum rvec)"
assumes "var.flowpipe0 X0 hl hu (CX) X1"
assumes "fst ` X0 \<subseteq> Csafe"
assumes "fst ` CX \<subseteq> Csafe"
assumes "fst ` X1 \<subseteq> Csafe"
shows "flowpipe (flow1_of_vec1 ` X0) hl hu (flow1_of_vec1 ` (CX::'n vec1 set)) (flow1_of_vec1 ` X1)"
using assms
unfolding flowpipe_def var.flowpipe0_def
apply safe
subgoal by (auto simp add: flow1_of_vec1_def vec1_of_flow1_def safe_eq[OF wd])
subgoal by (auto simp add: flow1_of_vec1_def vec1_of_flow1_def safe_eq[OF wd])
subgoal by (auto simp add: flow1_of_vec1_def vec1_of_flow1_def safe_eq[OF wd])
subgoal for x W y V h
apply (drule bspec[where x="(y, V)"], force)
apply (drule bspec, assumption)
by (simp add: var_existence_ivl0_eq_existence_ivl0[OF wd] flow1_of_vec1_def)
subgoal for x W y V h
apply (drule bspec[where x="(y, V)"], force)
apply (drule bspec, assumption)
apply (subst flow_Dflow_eq_var_flow[OF wd],
force simp: var_existence_ivl0_eq_existence_ivl0[OF wd] flow1_of_vec1_def)
apply (rule imageI)
by (simp add: vec1_of_flow1_def flow1_of_vec1_def)
subgoal for x W y V h h'
apply (drule bspec[where x="vec1_of_flow1 (x, W)"], force)
apply (drule bspec, assumption)
apply (subst flow_Dflow_eq_var_flow[OF wd])
apply (subst (asm) var_existence_ivl0_eq_existence_ivl0[OF wd])
apply (simp add: flow1_of_vec1_def)
subgoal
by (meson local.existence_ivl_initial_time local.mem_existence_ivl_iv_defined(1)
local.mem_existence_ivl_iv_defined(2) mem_is_interval_1_I mvar.interval)
subgoal
apply (rule imageI)
by (simp add: vec1_of_flow1_def flow1_of_vec1_def)
done
done
theorem einterpret_solve_poincare_fas:
assumes wd: "wd TYPE('n rvec)"
assumes "length CXs = D + D*D" "n < D"
assumes nz: "ode (fst (eucl_of_list CXs::'n vec1)) \<bullet> Basis_list ! n \<noteq> 0"
shows
"flow1_of_vec1 (einterpret (solve_poincare_fas n) CXs::'n::enum vec1) =
(let (x, d) = flow1_of_vec1 (eucl_of_list CXs::'n vec1) in (x,
d - (blinfun_scaleR_left (ode (x)) o\<^sub>L
(blinfun_scaleR_left (inverse (ode x \<bullet> Basis_list ! n)) o\<^sub>L (blinfun_inner_left (Basis_list ! n) o\<^sub>L d)))))"
using assms
apply (auto intro!: simp: flow1_of_vec1_def solve_poincare_fas_def)
subgoal
apply (auto intro!: euclidean_eqI[where 'a="'n rvec"])
apply (subst eucl_of_list_prod)
by (auto simp: eucl_of_list_prod length_concat o_def sum_list_distinct_conv_sum_set D_def Let_def
wdD[OF wd] take_eq_map_nth)
subgoal premises prems
proof -
have ode_e_eq: "interpret_floatarith (ode_e ! i) (map ((!) CXs) [0..<CARD('n)]) = interpret_floatarith (ode_e ! i) CXs"
if "i < D"
for i
apply (rule interpret_floatarith_max_Var_cong)
apply (drule max_Var_floatariths_lessI)
using that apply (simp add: wdD[OF wd])
apply (subst nth_map)
apply auto
using wdD[OF wd]
apply (simp add: )
using wdD[OF wd]
apply (simp add: )
done
define z where "z = (0::float)"
show ?thesis
supply [simp] = snd_eucl_of_list_prod fst_eucl_of_list_prod
supply [simp del] = eucl_of_list_take_DIM
using prems unfolding z_def[symmetric] D_def Let_def
including blinfun.lifting
apply (transfer fixing: CXs n z)
unfolding z_def
apply (auto simp: o_def ode_def intro!: ext)
apply (vector matrix_vector_mult_def )
apply (auto intro!: blinfun_euclidean_eqI simp: inner_Basis_eq_vec_nth wdD[OF wd])
apply (auto simp: length_concat o_def sum_list_distinct_conv_sum_set wdD[OF wd] take_eq_map_nth)
apply (auto simp: concat_map_map_index)
apply (vector )
apply (subst vec_nth_eq_list_of_eucl2 vec_nth_eq_list_of_eucl1)+
apply (subst (asm) vec_nth_eq_list_of_eucl2 vec_nth_eq_list_of_eucl1)+
apply (simp add: less_imp_le wdD[OF wd] index_nth_id )
apply (auto simp: algebra_simps ode_e_eq wdD[OF wd] divide_simps)
done
qed
done
lemma choose_step'_flowpipe:
assumes wd[refine_vcg]: "wd TYPE('n::enum rvec)"
assumes safe: "fst ` X0 \<subseteq> Csafe"
shows "var.choose_step (X0::'n vec1 set) h \<le> SPEC (\<lambda>(h', _, RES_ivl, RES::'n vec1 set).
0 < h' \<and> h' \<le> h \<and> flowpipe (flow1_of_vec1 ` X0) h' h' (flow1_of_vec1 ` RES_ivl) (flow1_of_vec1 ` RES))"
apply refine_vcg
apply (auto simp: )
apply (frule var.flowpipe0_safeD)
apply (drule var_flowpipe0_flowpipe[rotated])
by (auto simp: safe_eq wd)
lemma max_Var_floatariths_solve_poincare_fas[le]:
assumes wd: "wd (TYPE('n::enum rvec))"
shows "i < D \<Longrightarrow> max_Var_floatariths (solve_poincare_fas i) \<le> D + D * D"
by (auto simp: solve_poincare_fas_def concat_map_map_index Let_def
intro!: max_Var_floatariths_leI Suc_leI)
(auto intro!: max_Var_floatarith_le_max_Var_floatariths_nthI max_Var_floatariths_ode_e_wd[OF wd]
simp: wdD[OF wd])
lemma length_solve_poincare_fas[simp]: "length (solve_poincare_fas n) = D + D * D"
by (auto simp: solve_poincare_fas_def length_concat o_def sum_list_distinct_conv_sum_set D_def Let_def)
theorem interpret_floatariths_solve_poincare_fas:
assumes wd: "wd TYPE('n::enum rvec)"
assumes "length CXs = D + D*D" "n < D"
assumes nz: "ode (fst (eucl_of_list CXs::'n vec1)) \<bullet> Basis_list ! n \<noteq> 0"
shows
"interpret_floatariths (solve_poincare_fas n) CXs =
list_of_eucl (vec1_of_flow1 (let (x, d) = flow1_of_vec1 (eucl_of_list CXs::'n vec1) in (x,
d - (blinfun_scaleR_left (ode (x)) o\<^sub>L
(blinfun_scaleR_left (inverse (ode x \<bullet> Basis_list ! n)) o\<^sub>L (blinfun_inner_left (Basis_list ! n) o\<^sub>L d))))))"
using arg_cong[where f="list_of_eucl::'n vec1 \<Rightarrow> _", OF arg_cong[where f=vec1_of_flow1, OF einterpret_solve_poincare_fas[OF assms]]]
apply (auto simp: )
apply (subst (asm) list_of_eucl_eucl_of_list)
apply (auto simp: )
apply (auto simp: wdD[OF wd])
done
lemma length_solve_poincare_slp[simp]: "length solve_poincare_slp = D"
by (auto simp: solve_poincare_slp_def)
lemma ne_zero_lemma:
assumes
"ode ` fst ` CX \<subseteq> FC"
"\<forall>b\<in>FC. b \<bullet> n \<noteq> 0"
"(a, b) \<in> CX"
"ode a \<bullet> n = 0"
shows "False"
proof -
have "(a, b) \<in> CX" by fact
then have "ode (fst (a, b)) \<in> ode ` fst ` CX" by blast
also have "\<dots> \<subseteq> FC"
by fact
finally have "ode a \<in> FC" by simp
with assms show False
by auto
qed
lemma ne_zero_lemma2:
assumes
"ode ` fst ` flow1_of_vec1 ` env \<subseteq> F"
"\<forall>x\<in>F. x \<bullet> n \<noteq> 0"
"(a, b) \<in> env"
"flow1_of_vec1 (a, b) = (a', b')"
"ode a' \<bullet> n = 0"
shows False
proof -
have "(a', b') \<in> flow1_of_vec1 ` env"
apply (rule image_eqI)
using assms by auto
then have "ode (fst (a', b')) \<in> ode ` fst ` \<dots>" by blast
also from assms have "\<dots> \<subseteq> F" by simp
finally have "ode a' \<in> F" by simp
with assms have "ode a' \<bullet> n \<noteq> 0" by auto
with assms show False by simp
qed
lemma solve_poincare_plane[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
assumes "n \<in> Basis"
shows "solve_poincare_plane (n::'n::enum rvec) CX \<le> SPEC (\<lambda>PDP.
fst ` PDP \<subseteq> Csafe \<and>
(\<forall>(x, d) \<in> CX. (x, d - (blinfun_scaleR_left (ode x) o\<^sub>L
(blinfun_scaleR_left (inverse (ode x \<bullet> n)) o\<^sub>L (blinfun_inner_left n o\<^sub>L d)))) \<in> PDP) \<and>
(\<forall>(x, d) \<in> PDP. ode x \<bullet> n \<noteq> 0))"
unfolding solve_poincare_plane_def
apply (refine_vcg)
subgoal using assms by auto
subgoal using assms by auto
subgoal using assms by auto
subgoal using assms by (auto simp: solve_poincare_slp_def)
subgoal using assms by auto
subgoal for C1 FC _ CX' CX'' P P1 FP _
apply auto
apply (drule bspec, assumption)
apply (rule image_eqI)
prefer 2 apply assumption
apply (subst einterpret_solve_poincare_fas)
subgoal using wd by auto
subgoal using wd by auto
subgoal using wd by auto
subgoal using wd assms by (auto elim!: ne_zero_lemma)
subgoal using wd assms by (auto simp: )
done
subgoal by (auto elim!: ne_zero_lemma2)
done
lemma choose_step1_flowpipe[le, refine_vcg]:
assumes wd[refine_vcg]: "wd TYPE('n::enum rvec)"
shows "choose_step1 (X0::'n eucl1 set) h \<le> SPEC (\<lambda>(h', _, RES_ivl, RES::'n eucl1 set).
0 < h' \<and> h' \<le> h \<and> flowpipe X0 h' h' RES_ivl RES)"
using assms
unfolding choose_step1_def
by (refine_vcg choose_step'_flowpipe[le] wd)
(auto simp: image_image,
auto simp: safe_eq vec1_of_flow1_def flowpipe0_imp_flowpipe env_len_def)
lemma image_flow1_of_vec1I:
"vec1_of_flow1 x \<in> X \<Longrightarrow> x \<in> flow1_of_vec1 ` X"
by (rule image_eqI) (rule flow1_of_vec1_vec1_of_flow1[symmetric])
lemma inter_sctn1_spec[le, refine_vcg]:
"inter_sctn1_spec X sctn \<le> SPEC (\<lambda>(R, S). X \<inter> plane_of sctn \<times> UNIV \<subseteq> R \<and> fst ` R \<subseteq> plane_of sctn
\<and> X \<inter> plane_of sctn \<times> UNIV \<subseteq> S \<and> fst ` S \<subseteq> plane_of sctn)"
unfolding inter_sctn1_spec_def
apply (refine_vcg, auto)
subgoal by (rule image_flow1_of_vec1I) (auto simp: plane_of_def inner_prod_def)
subgoal by (auto simp: plane_of_def inner_prod_def)
subgoal by (rule image_flow1_of_vec1I)
(force simp: set_plus_def plane_of_def inner_prod_def vec1_of_flow1_def)
subgoal by (force simp: set_plus_def)
done
lemma fst_safe_coll[le, refine_vcg]:
"wd TYPE('a) \<Longrightarrow>
fst_safe_coll (X::('a::executable_euclidean_space*'c) set) \<le> SPEC (\<lambda>R. R = fst ` X \<and> fst ` X \<subseteq> Csafe)"
unfolding fst_safe_coll_def
by refine_vcg
lemma vec1reps[THEN order_trans, refine_vcg]: "vec1reps CX \<le> SPEC (\<lambda>R. case R of None \<Rightarrow> True | Some X \<Rightarrow> X = vec1_of_flow1 ` CX)"
unfolding vec1reps_def
apply (refine_vcg FORWEAK_mono_rule[where
I="\<lambda>XS R. case R of None \<Rightarrow> True | Some R \<Rightarrow> vec1_of_flow1 ` (\<Union>XS) \<subseteq> R \<and> R \<subseteq> vec1_of_flow1 ` CX"])
by (auto simp: split: option.splits) force+
lemma nonzero_component_within[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "nonzero_component_within ivl sctn (PDP::'n eucl1 set) \<le> SPEC (\<lambda>b.
(b \<longrightarrow> (\<forall>x\<in>PDP. fst x \<in> ivl \<and> (\<forall>\<^sub>F x in at (fst x) within plane_of sctn. x \<in> ivl))) \<and>
fst ` PDP \<subseteq> Csafe \<and>
(\<forall>x\<in>PDP. ode (fst x) \<bullet> normal sctn \<noteq> 0))"
unfolding nonzero_component_within_def
by refine_vcg auto
lemma do_intersection_invar_inside:
"do_intersection_invar guards b ivl sctn X (e, f, m, n, p, q, True) \<Longrightarrow>
fst ` e \<subseteq> sabove_halfspace sctn \<Longrightarrow>
fst ` mn \<subseteq> ivl \<Longrightarrow>
mn = m \<or> mn = n \<Longrightarrow>
do_intersection_spec UNIV guards ivl sctn X (mn, p)"
subgoal premises prems
proof -
from prems have e: "e \<inter> sbelow_halfspace sctn \<times> UNIV = {}"
by (auto simp: halfspace_simps plane_of_def)
with prems(1) have
"poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} X UNIV p m"
"poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} X UNIV p n"
"e \<inter> sbelow_halfspace sctn \<times> UNIV = {}"
"fst ` X \<inter> b = {}"
"fst ` X \<subseteq> sbelow_halfspace sctn"
"ivl \<subseteq> plane (normal sctn) (pstn sctn)"
"fst ` X \<subseteq> p"
"fst ` m \<subseteq> Csafe"
"fst ` n \<subseteq> Csafe"
"p \<subseteq> Csafe"
"fst ` e \<subseteq> Csafe"
"f \<subseteq> {0..}"
"p \<subseteq> sbelow_halfspace sctn - guards"
"e \<subseteq> (- guards) \<times> UNIV"
"fst ` (m \<union> n) \<inter> guards = {}"
"0 \<notin> (\<lambda>x. ode x \<bullet> normal sctn) ` fst ` (m \<union> n)"
"\<forall>x\<in>m \<union> n. \<forall>\<^sub>F x in at (fst x) within plane (normal sctn) (pstn sctn). x \<in> ivl"
by (auto simp: do_intersection_invar_def do_intersection_spec_def plane_of_def)
then show ?thesis
using prems(2-)
by (auto simp: do_intersection_spec_def plane_of_def halfspace_simps)
qed
done
lemma do_intersection_body_lemma:
assumes "flowsto A T (i \<times> UNIV) (X' \<inter> sbelow_halfspace sctn \<times> UNIV)"
"poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} B UNIV i PS "
"poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} B UNIV i PS2"
"T \<subseteq> {0..}"
"i \<subseteq> sbelow_halfspace sctn - guards"
"fst ` (A \<union> B) \<subseteq> sbelow_halfspace sctn"
"fst ` PS \<subseteq> Csafe "
"fst ` PS2 \<subseteq> Csafe "
\<open>X = A \<union> B\<close>
assumes ivl: "closed ivl" "ivl \<subseteq> plane_of sctn"
assumes normal_Basis: "\<bar>normal sctn\<bar> \<in> Basis"
and inter_empties: "fst ` Y \<inter> GUARDS = {}" "fst ` CX' \<inter> GUARDS = {}"
"fst ` PDP' \<inter> GUARDS = {}" "fst ` PDP'' \<inter> GUARDS = {}"
and h': "0 < h'" "h' \<le> h"
and safe: "fst ` PDP \<subseteq> Csafe" "fst ` CX' \<subseteq> Csafe"
"fst ` PDP' \<subseteq> Csafe"
"fst ` PDP'' \<subseteq> Csafe"
and PDP:
"\<forall>(x,d)\<in>CX'. (x,
d - (blinfun_scaleR_left (ode x) o\<^sub>L
(blinfun_scaleR_left (inverse (ode x \<bullet> \<bar>normal sctn\<bar>)) o\<^sub>L
(blinfun_inner_left \<bar>normal sctn\<bar> o\<^sub>L d))))
\<in> PDP"
and PDP': "PDP \<inter> plane_of sctn \<times> UNIV \<subseteq> PDP'"
and PDP'': "PDP \<inter> plane_of sctn \<times> UNIV \<subseteq> PDP''"
and evin:
"\<forall>x\<in>PDP'. fst x \<in> ivl \<and> (\<forall>\<^sub>F x in at (fst x) within plane_of sctn. x \<in> ivl)"
"\<forall>x\<in>PDP''. fst x \<in> ivl \<and> (\<forall>\<^sub>F x in at (fst x) within plane_of sctn. x \<in> ivl)"
and through: "\<forall>(x, d)\<in>PDP. ode x \<bullet> \<bar>normal sctn\<bar> \<noteq> 0"
"\<forall>x\<in>PDP'. ode (fst x) \<bullet> normal sctn \<noteq> 0"
"\<forall>x\<in>PDP''. ode (fst x) \<bullet> normal sctn \<noteq> 0"
and plane:
"fst ` PDP' \<subseteq> plane_of sctn"
"fst ` PDP'' \<subseteq> plane_of sctn"
and flowpipe: "flowpipe X' h' h' CX' Y"
shows "\<exists>A B. X = A \<union> B \<and>
flowsto A {0<..} ((fst ` CX' \<inter> sbelow_halfspace sctn \<union> i) \<times> UNIV) (Y \<inter> sbelow_halfspace sctn \<times> UNIV) \<and>
poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} B UNIV (fst ` CX' \<inter> sbelow_halfspace sctn \<union> i) (PDP' \<union> PS) \<and>
poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} B UNIV (fst ` CX' \<inter> sbelow_halfspace sctn \<union> i) (PDP'' \<union> PS2)"
proof -
from flowpipe
have 1: "flowpipe (X' \<inter> (sbelow_halfspace sctn) \<times> UNIV) h' h' CX' Y"
by (rule flowpipe_subset) (use flowpipe in \<open>auto dest!: flowpipe_safeD\<close>)
have 2: "fst ` (X' \<inter> (sbelow_halfspace sctn) \<times> UNIV) \<inter> {x. pstn sctn \<le> x \<bullet> normal sctn} = {}"
by (auto simp: halfspace_simps plane_of_def)
from normal_Basis have 3: "normal sctn \<noteq> 0"
by (auto simp: )
note 4 = \<open>closed ivl\<close>
from \<open>ivl \<subseteq> plane_of sctn\<close> have 5: "ivl \<subseteq> plane (normal sctn) (pstn sctn)"
by (auto simp: plane_of_def)
have 6: "(x, d) \<in> CX' \<Longrightarrow> x \<in> plane (normal sctn) (pstn sctn) \<Longrightarrow>
(x, d - (blinfun_scaleR_left (ode x) o\<^sub>L
(blinfun_scaleR_left (inverse (ode x \<bullet> normal sctn)) o\<^sub>L (blinfun_inner_left (normal sctn) o\<^sub>L d))))
\<in> PDP' \<inter> PDP''" for x d
unfolding PDP_abs_lemma[OF normal_Basis]
apply (drule PDP[rule_format, of "(x, d)", unfolded split_beta' fst_conv snd_conv])
using PDP' PDP''
by (auto simp: plane_of_def)
from normal_Basis through
have 7: "(x, d) \<in> PDP' \<Longrightarrow> ode x \<bullet> normal sctn \<noteq> 0" for x d
by (auto elim!: abs_in_BasisE)
have 8: "(x, d) \<in> PDP' \<Longrightarrow> x \<in> ivl" for x d
using evin by auto
have 9: "(x, d) \<in> PDP' \<Longrightarrow> \<forall>\<^sub>F x in at x within plane (normal sctn) (pstn sctn). x \<in> ivl" for x d
using evin by (auto simp add: plane_of_def)
obtain X1 X2
where X1X2: "X' \<inter> sbelow_halfspace sctn \<times> UNIV = X1 \<union> X2"
and X1: "flowsto X1 {0<..h'} (CX' \<inter> {x. x \<bullet> normal sctn < pstn sctn} \<times> UNIV)
(CX' \<inter> {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} \<times> UNIV)"
and X2: "flowsto X2 {h'..h'} (CX' \<inter> {x. x \<bullet> normal sctn < pstn sctn} \<times> UNIV)
(Y \<inter> {x. x \<bullet> normal sctn < pstn sctn} \<times> UNIV)"
and P: "poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} X1 UNIV
(fst ` CX' \<inter> {x. x \<bullet> normal sctn < pstn sctn}) (PDP' \<inter> PDP'')"
by (rule flowpipe_split_at_above_halfspace[OF 1 2 3 4 5 6 7 8 9]) (auto simp: Ball_def)
from \<open>flowsto A _ _ _\<close>[unfolded X1X2]
obtain p1 p2 where p1p2: "A = p1 \<union> p2" and p1: "flowsto p1 T (i \<times> UNIV) X1" and p2: "flowsto p2 T (i \<times> UNIV) X2"
by (rule flowsto_unionE)
have "A \<union> B = p2 \<union> (p1 \<union> B)" using \<open>A = p1 \<union> p2\<close>
by auto
moreover
from flowsto_trans[OF p2 X2]
have "flowsto p2 {0<..} ((fst ` CX' \<inter> (sbelow_halfspace sctn) \<union> i) \<times> UNIV)
(Y \<inter> (sbelow_halfspace sctn) \<times> UNIV)"
apply (rule flowsto_subset)
subgoal by (auto simp: halfspace_simps)
subgoal using h' \<open>T \<subseteq> _\<close> by (auto simp: halfspace_simps intro!: add_nonneg_pos)
subgoal
using flowpipe_source_subset[OF 1, unfolded X1X2] X1X2
apply auto
by (auto simp: halfspace_simps)
subgoal by (auto simp: halfspace_simps)
done
moreover
have cls: "closed {x \<in> ivl. x \<bullet> normal sctn = pstn sctn}"
by (rule closed_levelset_within continuous_intros \<open>closed ivl\<close>)+
from flowsto_trans[OF p1 X1]
have ftt: "flowsto p1 ({s + t |s t. s \<in> T \<and> t \<in> {0<..h'}})
(i \<times> UNIV \<union> CX' \<inter> {x. x \<bullet> normal sctn < pstn sctn} \<times> UNIV \<union> X1 \<inter> X1)
(X1 - X1 \<union> CX' \<inter> {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} \<times> UNIV)"
by auto
from X1X2 have X1_sb: "X1 \<subseteq> sbelow_halfspace sctn \<times> UNIV" by auto
have "{x \<in> ivl. x \<bullet> normal sctn = pstn sctn} \<times> UNIV \<inter> (i \<times> UNIV \<union> CX' \<inter> {x. x \<bullet> normal sctn < pstn sctn} \<times> UNIV \<union> X1) = {}"
apply (intro Int_Un_eq_emptyI)
subgoal using \<open>i \<subseteq> sbelow_halfspace sctn - guards\<close> by (auto simp: halfspace_simps)
subgoal by (auto simp: halfspace_simps)
subgoal using X1_sb by (auto simp: halfspace_simps)
done
then have inter_empty:
"{x \<in> ivl. x \<bullet> normal sctn = pstn sctn} \<times> UNIV \<inter> (i \<times> UNIV \<union> CX' \<inter> {x. x \<bullet> normal sctn < pstn sctn} \<times> UNIV \<union> X1 \<inter> X1) = {}"
by auto
have p1ret: "returns_to {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} x"
and p1pm: "poincare_map {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} x \<in> fst ` (PDP' \<inter> PDP'')"
if "(x, d) \<in> p1" for x d
apply (rule flowsto_poincareD[OF ftt _ inter_empty _ _ _ order_refl])
subgoal by auto
subgoal by fact
subgoal using \<open>T \<subseteq> _\<close> by auto
subgoal using that by auto
subgoal
apply (rule flowsto_poincareD[OF ftt _ inter_empty])
subgoal by auto
subgoal by fact
subgoal using \<open>T \<subseteq> _\<close> by auto
subgoal using that by auto
subgoal using 6 by force
done
done
have crt: "isCont (return_time {x \<in> ivl. x \<bullet> normal sctn - pstn sctn = 0}) x" if "(x, d) \<in> p1" for x d
apply (rule return_time_isCont_outside[where Ds="\<lambda>_. blinfun_inner_left (normal sctn)"])
subgoal by (simp add: p1ret[OF that])
subgoal by fact
subgoal by (auto intro!: derivative_eq_intros)
subgoal by simp
subgoal apply simp
using p1pm[OF that]
by (auto dest!: 7)
subgoal
using p1pm[OF that]
by (auto dest!: 9 simp: eventually_at_filter)
subgoal
using \<open>fst ` (A \<union> B) \<subseteq> sbelow_halfspace sctn\<close> that p1p2
by (auto simp: halfspace_simps)
done
have pmij: "poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} p1 UNIV
(fst ` (i \<times> UNIV \<union> X1) \<union> fst ` CX' \<inter> {x. x \<bullet> normal sctn < pstn sctn}) (PDP' \<inter> PDP'')"
apply (rule flowsto_poincare_trans[OF \<open>flowsto _ _ _ X1\<close> P])
subgoal using \<open>T \<subseteq> {0..}\<close> by auto
subgoal by auto
subgoal
using \<open>i \<subseteq> sbelow_halfspace sctn - guards\<close> X1X2
by (force simp: halfspace_simps)
subgoal by fact
subgoal for x d using crt by simp
subgoal by auto
done
from pmij have "poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} p1 UNIV (fst ` (i \<times> UNIV \<union> X1) \<union> fst ` CX' \<inter> {x. x \<bullet> normal sctn < pstn sctn}) PDP'"
apply (rule poincare_mapsto_subset)
using \<open>fst ` PDP' \<subseteq> Csafe\<close>
by auto
from this \<open>poincare_mapsto _ _ _ i PS\<close>
have "poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} (p1 \<union> B) UNIV
((fst ` (i \<times> UNIV \<union> X1) \<union> fst ` CX' \<inter> {x. x \<bullet> normal sctn < pstn sctn}) \<union> i) (PDP' \<union> PS)"
by (intro poincare_mapsto_unionI) (auto simp: plane_of_def)
then have "poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} (p1 \<union> B) UNIV (fst ` CX' \<inter> sbelow_halfspace sctn \<union> i) (PDP' \<union> PS)"
apply (rule poincare_mapsto_subset)
subgoal by auto
subgoal by auto
subgoal
using flowpipe_source_subset[OF 1, unfolded X1X2] X1X2
apply (auto simp: halfspace_simps subset_iff)
done
subgoal using safe \<open>fst ` PS \<subseteq> Csafe\<close> by auto
done
moreover
from pmij have "poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} p1 UNIV (fst ` (i \<times> UNIV \<union> X1) \<union> fst ` CX' \<inter> {x. x \<bullet> normal sctn < pstn sctn}) PDP''"
apply (rule poincare_mapsto_subset)
using \<open>fst ` PDP'' \<subseteq> Csafe\<close>
by auto
from this \<open>poincare_mapsto _ _ _ i PS2\<close>
have "poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} (p1 \<union> B) UNIV
((fst ` (i \<times> UNIV \<union> X1) \<union> fst ` CX' \<inter> {x. x \<bullet> normal sctn < pstn sctn}) \<union> i) (PDP'' \<union> PS2)"
by (intro poincare_mapsto_unionI) (auto simp: plane_of_def)
then have "poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} (p1 \<union> B) UNIV (fst ` CX' \<inter> sbelow_halfspace sctn \<union> i) (PDP'' \<union> PS2)"
apply (rule poincare_mapsto_subset)
subgoal by auto
subgoal by auto
subgoal
using flowpipe_source_subset[OF 1, unfolded X1X2] X1X2
apply (auto simp: halfspace_simps subset_iff)
done
subgoal using safe \<open>fst ` PS2 \<subseteq> Csafe\<close> by auto
done
ultimately
show ?thesis
unfolding \<open>X = A \<union> B\<close> by blast
qed
lemma do_intersection_body_spec:
fixes guards::"'n::enum rvec set"
assumes invar: "do_intersection_invar guards GUARDS ivl sctn X (X', T, PS, PS2, i, True, True)"
and wdp[refine_vcg]: "wd TYPE('n rvec)"
and X: "fst ` X \<subseteq> Csafe"
and ivl: "closed ivl" and GUARDS: "guards \<subseteq> GUARDS"
shows "do_intersection_body GUARDS ivl sctn h (X', T, PS, PS2, i, True, True) \<le>
SPEC (do_intersection_invar guards GUARDS ivl sctn X)"
proof -
from invar
obtain A B where AB: "fst ` (A \<union> B) \<inter> GUARDS = {} "
"fst ` (A \<union> B) \<subseteq> sbelow_halfspace sctn "
"ivl \<subseteq> plane_of sctn "
"fst ` (A \<union> B) \<subseteq> i "
"fst ` PS \<subseteq> Csafe "
"fst ` PS2 \<subseteq> Csafe "
"i \<subseteq> Csafe "
"fst ` X' \<subseteq> Csafe "
"T \<subseteq> {0..}"
"i \<subseteq> sbelow_halfspace sctn - guards "
"X' \<subseteq> (- guards) \<times> UNIV "
"fst ` (PS \<union> PS2) \<inter> guards = {} "
"0 \<notin> (\<lambda>x. ode x \<bullet> normal sctn) ` fst ` (PS \<union> PS2) "
"\<forall>x\<in>PS \<union> PS2. \<forall>\<^sub>F x in at (fst x) within plane_of sctn. x \<in> ivl "
"X = A \<union> B "
"flowsto A T (i \<times> UNIV) (X' \<inter> sbelow_halfspace sctn \<times> UNIV)"
"poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} B UNIV i PS "
"poincare_mapsto {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} B UNIV i PS2"
by (auto simp: do_intersection_invar_def)
have ev_in_ivl: "\<forall>\<^sub>F x in at p within plane_of sctn. x \<in> ivl" if
\<open>\<forall>x\<in>d. fst x \<in> ivl \<and> (\<forall>\<^sub>F x in at (fst x) within plane_of sctn. x \<in> ivl)\<close>
\<open>\<forall>x\<in>e. fst x \<in> ivl \<and> (\<forall>\<^sub>F x in at (fst x) within plane_of sctn. x \<in> ivl)\<close>
\<open>(p, q) \<in> d \<or> (p, q) \<in> PS \<or> (p, q) \<in> e \<or> (p, q) \<in> PS2\<close>
for p q d e
using \<open>\<forall>x\<in>PS \<union> PS2. \<forall>\<^sub>F x in at (fst x) within plane_of sctn. x \<in> ivl\<close>
using that
by (auto dest!: bspec[where x="(p, q)"])
show ?thesis
unfolding do_intersection_body_def do_intersection_invar_def
apply simp
apply (refine_vcg, clarsimp_all)
subgoal using AB by (auto simp: )
subgoal using AB by (auto simp: )
subgoal using AB by (auto simp: )
subgoal
apply (rule conjI)
subgoal using AB by auto\<comment> \<open>unnecessarily slow\<close>
subgoal using AB by fastforce
done
subgoal using AB by (auto simp: )
subgoal using AB by (auto simp: )
subgoal using AB by (auto simp: )
subgoal by (auto dest!: flowpipe_safeD)
subgoal
apply safe
subgoal using AB GUARDS by auto
subgoal using AB by auto
subgoal using AB by auto
subgoal using AB GUARDS by auto
subgoal using AB by auto
subgoal using AB by auto
done
subgoal using AB GUARDS by auto
subgoal using AB GUARDS by auto\<comment> \<open>unnecessarily slow\<close>
subgoal using AB GUARDS by auto
subgoal using AB assms by (auto intro: ev_in_ivl)
subgoal using AB assms apply - by (rule do_intersection_body_lemma)
done
qed
lemma
do_intersection_spec[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "do_intersection guards ivl sctn (X::'n eucl1 set) h \<le>
SPEC (\<lambda>(inside, P, P2, CX). (inside \<longrightarrow>
(do_intersection_spec UNIV guards ivl sctn X (P, CX) \<and>
do_intersection_spec UNIV guards ivl sctn X (P2, CX)) \<and> fst ` X \<subseteq> CX))"
using assms
unfolding do_intersection_def autoref_tag_defs
apply (refine_vcg, clarsimp_all)
subgoal
unfolding do_intersection_invar_def
apply clarsimp
apply (intro conjI)
apply force
apply force
apply force
apply (rule exI[where x=X])
apply (rule exI[where x="{}"])
by (auto intro!: flowsto_self)
subgoal by (rule do_intersection_body_spec)
subgoal by (rule do_intersection_invar_inside, assumption) auto
subgoal by (rule do_intersection_invar_inside, assumption) auto
subgoal by (auto simp: plane_of_def halfspace_simps do_intersection_invar_def)
done
lemma mem_flow1_of_vec1_image_iff[simp]:
"(c, d) \<in> flow1_of_vec1 ` a \<longleftrightarrow> vec1_of_flow1 (c, d) \<in> a"
by force
lemma mem_vec1_of_flow1_image_iff[simp]:
"(c, d) \<in> vec1_of_flow1 ` a \<longleftrightarrow> flow1_of_vec1 (c, d) \<in> a"
by force
lemma split_spec_param1[le, refine_vcg]: "split_spec_param1 X \<le> SPEC (\<lambda>(A, B). X \<subseteq> A \<union> B)"
unfolding split_spec_param1_def
apply (refine_vcg)
apply (auto simp add: subset_iff split: option.splits)
by (metis flow1_of_vec1_vec1_of_flow1 surjective_pairing)
lemma do_intersection_spec_empty:
"X = {} \<Longrightarrow> Y = {} \<Longrightarrow> do_intersection_spec S sctns ivl sctn X ({}, Y)"
by (auto simp: do_intersection_spec_def halfspaces_union)
lemma do_intersection_spec_subset:
"do_intersection_spec S osctns ivl csctns Y (a, b) \<Longrightarrow> X \<subseteq> Y \<Longrightarrow> do_intersection_spec S osctns ivl csctns X (a, b)"
by (auto simp: do_intersection_spec_def halfspaces_union intro: flowsto_subset poincare_mapsto_subset)
lemma do_intersection_spec_union:
"do_intersection_spec S osctns ivl csctns a (b, c) \<Longrightarrow>
do_intersection_spec S osctns ivl csctns f (g, h) \<Longrightarrow>
do_intersection_spec S osctns ivl csctns (a \<union> f) (b \<union> g, c \<union> h)"
by (auto simp: do_intersection_spec_def intro!: poincare_mapsto_unionI)
lemma scaleR2_rep_of_coll[le, refine_vcg]:
"scaleR2_rep_coll X \<le> SPEC (\<lambda>((l, u), Y). X \<subseteq> scaleR2 l u Y)"
unfolding scaleR2_rep_coll_def
apply (refine_vcg FORWEAK_mono_rule[where I="\<lambda>Xs ((l, u), Y). \<Union>Xs \<subseteq> scaleR2 l u Y"])
subgoal by (auto intro: scaleR2_subset)
subgoal
apply clarsimp
apply safe
subgoal by (auto elim: scaleR2_subset)
subgoal
apply (rule set_rev_mp, assumption)
apply (rule order_trans)
apply (rule Union_upper, assumption)
apply (rule order_trans, assumption)
apply (rule subsetI)
apply (erule scaleR2_subset)
by (auto )
done
done
lemma split_spec_param1e[le, refine_vcg]: "split_spec_param1e X \<le> SPEC (\<lambda>(A, B). X \<subseteq> A \<union> B)"
unfolding split_spec_param1e_def
apply (refine_vcg)
apply clarsimp
apply (thin_tac "_ \<noteq> {}")
apply (auto simp: scaleR2_def vimage_def image_def)
apply (rule exI, rule conjI, assumption, rule conjI, assumption)
apply (auto simp: split_beta')
apply (drule_tac x = x in spec)
apply auto
by (metis (no_types, lifting) UnE prod.sel(1) prod.sel(2) subset_eq)
lemma reduce_spec1[le, refine_vcg]: "reduce_spec1 ro X \<le> SPEC (\<lambda>R. X \<subseteq> R)"
unfolding reduce_spec1_def
by refine_vcg auto
lemma reduce_spec1e[le, refine_vcg]: "reduce_spec1e ro X \<le> SPEC (\<lambda>R. X \<subseteq> R)"
unfolding reduce_spec1e_def
by refine_vcg (auto simp: scaleR2_def image_def vimage_def, force)
lemma split_under_threshold[le, refine_vcg]:
"split_under_threshold ro th X \<le> SPEC (\<lambda>R. X \<subseteq> R)"
unfolding split_under_threshold_def autoref_tag_defs
by (refine_vcg) auto
lemma step_split[le, refine_vcg]:
"wd TYPE((real, 'n::enum) vec) \<Longrightarrow> step_split ro (X::'n eucl1 set) \<le> SPEC (\<lambda>Y. X \<subseteq> Y \<and> fst ` Y \<subseteq> Csafe)"
unfolding step_split_def
by (refine_vcg refine_vcg) auto
lemma tolerate_error_SPEC[THEN order_trans, refine_vcg]:
"tolerate_error Y E \<le> SPEC (\<lambda>b. True)"
unfolding tolerate_error_def
by refine_vcg
lemma flowpipe_scaleR2I: "flowpipe (scaleR2 x1 x2 bc) x1a x1a (fst ` aca \<times> UNIV) (scaleR2 x1 x2 bca)"
if "flowpipe (bc) x1a x1a (fst ` aca \<times> UNIV) (bca)"
using that
apply (auto simp: flowpipe_def scaleR2_def)
apply (drule bspec, assumption)
apply (auto simp: image_def vimage_def )
apply (rule exI, rule conjI, assumption, rule conjI, assumption)
apply (rule bexI) prefer 2 apply assumption
by (auto simp: scaleR_blinfun_compose_right)
lemma choose_step1e_flowpipe[le, refine_vcg]:
assumes vwd[refine_vcg]: "wd TYPE('n::enum rvec)"
shows "choose_step1e (X0::'n eucl1 set) h \<le> SPEC (\<lambda>(h', _, RES_ivl, RES::'n eucl1 set).
0 < h' \<and> h' \<le> h \<and> flowpipe X0 h' h' (RES_ivl \<times> UNIV) RES)"
unfolding choose_step1e_def
apply (refine_vcg)
apply (auto intro: flowpipe_scaleR2I)
apply (erule contrapos_np)
apply (auto intro!: flowpipe_scaleR2I)
apply (rule flowpipe_subset)
apply assumption
apply (auto dest!: flowpipe_safeD)
done
lemma width_spec_appr1[THEN order_trans, refine_vcg]: "width_spec_appr1 X \<le> SPEC (\<lambda>_. True)"
unfolding width_spec_appr1_def
by refine_vcg
lemma tolerate_error1_SPEC[THEN order_trans, refine_vcg]:
"tolerate_error1 Y E \<le> SPEC (\<lambda>b. True)"
unfolding tolerate_error1_def
by refine_vcg
lemma
step_adapt_time[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "step_adapt_time (X::'n eucl1 set) h \<le> SPEC (\<lambda>(t, CX, X1, h). flowpipe X t t (CX \<times> UNIV) X1)"
unfolding step_adapt_time_def autoref_tag_defs
apply (refine_vcg refine_vcg, clarsimp)
apply (auto simp: flowpipe_def)
apply force
done
lemma
resolve_step[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "resolve_step roptns (X::'n::enum eucl1 set) h \<le> SPEC (\<lambda>(_, CX, X1, _).
flowsto X {0..} (CX \<times> UNIV) X1 \<and> X \<union> X1 \<subseteq> CX \<times> UNIV \<and> X1 \<union> CX \<times> UNIV \<subseteq> Csafe \<times> UNIV)"
unfolding resolve_step_def autoref_tag_defs
apply (refine_vcg refine_vcg)
subgoal by (rule flowsto_self) auto
subgoal by auto
subgoal by auto
subgoal
apply clarsimp
apply (frule flowpipe_imp_flowsto_nonneg)
apply (rule flowsto_subset, assumption)
by auto
subgoal
by (auto dest: flowpipe_source_subset)
subgoal
by (auto dest!: flowpipe_safeD)
done
lemma pre_intersection_step[THEN order_trans, refine_vcg]:
"pre_intersection_step ro (X::'n eucl1 set) h \<le> SPEC (\<lambda>(X', CX, G). X \<subseteq> X' \<union> G \<and> X \<union> X' \<union> G \<subseteq> CX \<times> UNIV)"
if [refine_vcg]: "wd TYPE('n::enum rvec)"
unfolding pre_intersection_step_def autoref_tag_defs
by (refine_vcg) auto
lemma [THEN order_trans, refine_vcg]: "select_with_inter ci a \<le> SPEC (\<lambda>_. True)"
unfolding select_with_inter_def
by (refine_vcg FORWEAK_mono_rule[where I="\<lambda>_ _. True"])
lemmas [refine_vcg del] = scaleR2_rep_of_coll
lemma fst_scaleR2_image[simp]: "ad \<le> ereal r \<Longrightarrow> ereal r \<le> bd \<Longrightarrow> fst ` scaleR2 ad bd be = fst ` be"
by (cases ad; cases bd; force simp: scaleR2_def image_image split_beta' vimage_def)
lemma scaleR2_rep_of_coll2[le, refine_vcg]:
"scaleR2_rep_coll X \<le> SPEC (\<lambda>((l, u), Y). X \<subseteq> scaleR2 l u Y \<and> fst ` X = fst ` Y)"
unfolding scaleR2_rep_coll_def
supply [simp del] = mem_scaleR2_union
apply (refine_vcg FORWEAK_mono_rule[where I="\<lambda>Xs ((l, u), Y).
\<Union>Xs \<subseteq> scaleR2 l u Y \<and> fst ` \<Union>Xs \<subseteq> fst ` Y \<and> fst ` Y \<subseteq> fst ` X"])
apply (auto intro: scaleR2_subset)
subgoal by (auto simp: scaleR2_def)
subgoal by (auto simp: scaleR2_def image_def vimage_def, fastforce)
subgoal
apply (rule scaleR2_subset)
apply (rule subsetD)
apply assumption
apply auto
done
subgoal by force
subgoal for a b c d e f g h i j k l
apply (rule scaleR2_subset)
apply (rule subsetD)
apply assumption
by auto
subgoal by (auto simp: scaleR2_def)
subgoal by (auto simp: scaleR2_def)
subgoal by (auto simp: scaleR2_def image_def vimage_def, fastforce)
done
lemma reach_cont[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "reach_cont roptns guards (X::'n eucl1 set) \<le> SPEC (\<lambda>(CX, G).
G \<union> (CX \<times> UNIV) \<subseteq> (Csafe - guards) \<times> UNIV \<and>
X \<union> G \<subseteq> CX \<times> UNIV \<and>
flowsto X {0..} (CX \<times> UNIV) G)"
+ using [[simproc del: defined_all]]
unfolding reach_cont_def autoref_tag_defs
apply (refine_vcg, clarsimp_all simp add: cancel_times_UNIV_subset)
subgoal by (rule flowsto_self) (auto simp: )
subgoal by (force simp: scaleR2_def)
subgoal by (fastforce simp: scaleR2_def vimage_def image_def)
subgoal premises prems for _ _ _ _ _ _ _ g
using \<open>flowsto X _ _ (g \<union> _ \<union> _)\<close> \<open>flowsto g _ _ _\<close>
apply (rule flowsto_stepI)
using prems
by auto
subgoal
apply safe
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
done
subgoal by auto
subgoal
by (rule flowsto_subset, assumption) auto
subgoal
apply safe
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by fastforce
subgoal by auto
subgoal by auto
subgoal
by (metis (mono_tags, lifting) Diff_eq_empty_iff Diff_iff IntI)
done
subgoal
apply safe
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
subgoal by auto
done
subgoal by auto
done
lemma reach_cont_par[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "reach_cont_par roptns guards (X::'n eucl1 set) \<le> SPEC (\<lambda>(CX, G).
G \<union> (CX \<times> UNIV) \<subseteq> (Csafe - guards) \<times> UNIV \<and>
X \<union> G \<subseteq> CX \<times> UNIV \<and>
flowsto X {0..} (CX \<times> UNIV) G)"
unfolding reach_cont_par_def
apply refine_vcg
apply auto
apply force
apply force
apply force
apply force
subgoal
apply (rule bexI)
prefer 2 apply assumption
by auto
subgoal
apply (rule bexI)
prefer 2 apply assumption
by auto
subgoal for R
apply (rule flowsto_source_Union)
apply (drule bspec, assumption)
apply auto
apply (rule flowsto_subset, assumption)
apply auto
done
done
lemma subset_iplane_coll[THEN order_trans, refine_vcg]:
"subset_iplane_coll x ics \<le> SPEC (\<lambda>b. b \<longrightarrow> x \<subseteq> ics)"
unfolding subset_iplane_coll_def
apply refine_vcg
subgoal for X icss
by (refine_vcg FORWEAK_mono_rule[where I="\<lambda>ic b. b \<longrightarrow> X \<subseteq> \<Union>(icss)"]) auto
done
lemma subsets_iplane_coll[THEN order_trans, refine_vcg]:
"subsets_iplane_coll x ics \<le> SPEC (\<lambda>b. b \<longrightarrow> \<Union>x \<subseteq> ics)"
unfolding subsets_iplane_coll_def
by (refine_vcg FORWEAK_mono_rule[where I="\<lambda>x b. (b \<longrightarrow> \<Union>x \<subseteq> ics)"]) auto
lemma symstart_coll[THEN order_trans, refine_vcg]:
assumes [refine_vcg]: "wd (TYPE('n::enum rvec))"
assumes [le, refine_vcg]:
"\<And>X0. X0 \<subseteq> Csafe \<times> UNIV \<Longrightarrow> symstart X0 \<le> SPEC (\<lambda>(CX, X). flowsto (X0 - trap \<times> UNIV) {0..} (CX \<times> UNIV) (X))"
shows "symstart_coll symstart X0 \<le> SPEC (\<lambda>(CX, X). flowsto ((X0::'n eucl1 set) - trap \<times> UNIV) {0..} (CX \<times> UNIV) X)"
unfolding symstart_coll_def autoref_tag_defs
apply (refine_vcg FORWEAK_mono_rule[where I="\<lambda>X (CY, Y). flowsto (\<Union>X - trap \<times> UNIV) {0..} (CY \<times> UNIV) Y"], clarsimp_all)
subgoal by force
subgoal for a b c d e by (rule flowsto_subset, assumption) auto
subgoal by force
subgoal for a b c d e f g
unfolding Un_Diff
apply (rule flowsto_source_unionI)
subgoal by (rule flowsto_subset, assumption) auto
subgoal by (rule flowsto_subset, assumption) auto
done
done
lemma reach_cont_symstart[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
assumes [le, refine_vcg]: "\<And>X0. X0 \<subseteq> Csafe \<times> UNIV \<Longrightarrow> symstart X0 \<le> SPEC (\<lambda>(CX, X). flowsto (X0 - trap \<times> UNIV) {0..} (CX \<times> UNIV) (X))"
shows "reach_cont_symstart roptns symstart guards (X::'n eucl1 set) \<le> SPEC (\<lambda>(CX, G).
G \<union> (CX \<times> UNIV) \<subseteq> (Csafe - guards) \<times> UNIV \<and>
X \<subseteq> CX \<times> UNIV \<and>
G \<subseteq> CX \<times> UNIV \<and>
flowsto (X - trap \<times> UNIV) {0..} (CX \<times> UNIV) (G))"
unfolding reach_cont_symstart_def autoref_tag_defs
apply (refine_vcg, clarsimp_all)
subgoal by (auto simp: times_subset_iff)
subgoal by auto
subgoal by auto
subgoal for a b c d e f g
apply (rule flowsto_stepI[OF _ _ order_refl])
apply assumption
by assumption auto
done
lemma reach_conts[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
assumes [refine_vcg]: "\<And>X0. X0 \<subseteq> Csafe \<times> UNIV \<Longrightarrow> symstart X0 \<le> SPEC (\<lambda>(CX, X). flowsto (X0 - trap \<times> UNIV) {0..} (CX \<times> UNIV) X)"
shows "reach_conts roptns symstart trap guards (X::'n eucl1 set) \<le> SPEC (\<lambda>(CX, IGs, X0).
\<Union>(snd ` IGs) \<union> (CX \<times> UNIV) \<subseteq> (Csafe - guards) \<times> UNIV \<and>
X \<subseteq> CX \<times> UNIV \<and>
\<Union>(snd ` IGs) \<subseteq> CX \<times> UNIV \<and>
\<Union>(fst ` IGs) \<subseteq> guards \<and>
X = \<Union>(X0 ` (snd ` IGs)) \<and>
(\<forall>(I, G) \<in> IGs. flowsto (X0 G - trap \<times> UNIV) {0..} (CX \<times> UNIV) G))"
unfolding reach_conts_def autoref_tag_defs
apply (refine_vcg, clarsimp_all)
subgoal for a b
apply (erule flowsto_Diff_to_Union_funE)
apply (force simp: split_beta')
subgoal for f
apply (rule exI[where x=f])
by (auto simp: split_beta')
done
subgoal by (auto)
subgoal by (auto)
subgoal by (auto)
done
lemma leaves_halfspace[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "leaves_halfspace S (X::'n::enum rvec set) \<le>
SPEC (\<lambda>b. case b of None \<Rightarrow> S = UNIV
| Some sctn \<Rightarrow>
(S = below_halfspace sctn \<and> X \<subseteq> plane_of sctn \<and> (\<forall>x \<in> X. ode x \<bullet> normal sctn < 0)))"
unfolding leaves_halfspace_def autoref_tag_defs op_set_to_list_def
apply (refine_vcg, clarsimp_all)
subgoal by (force simp add: halfspace_simps plane_of_def)
done
lemma poincare_start_on[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "poincare_start_on guards sctn (X0::'n eucl1 set) \<le> SPEC (\<lambda>(X1S, CX1S).
fst ` (X1S \<union> (CX1S \<times> UNIV)) \<subseteq> Csafe \<and>
fst ` X1S \<subseteq> sbelow_halfspace sctn \<and>
fst ` (X1S \<union> (CX1S \<times> UNIV)) \<inter> guards = {} \<and>
(X0 \<subseteq> (CX1S \<times> UNIV)) \<and>
(\<forall>(x, d) \<in> CX1S \<times> UNIV. ode x \<bullet> normal sctn < 0) \<and>
flowsto X0 pos_reals ((CX1S \<times> UNIV) \<inter> (sbelow_halfspace sctn \<times> UNIV)) X1S)"
unfolding poincare_start_on_def autoref_tag_defs
apply refine_vcg
apply (rule FORWEAK_mono_rule[where I="\<lambda>X0S (X1S, CX1S).
flowsto (\<Union>X0S) pos_reals ((CX1S \<times> UNIV) \<inter> sbelow_halfspace sctn \<times> UNIV) X1S \<and>
fst ` (X1S \<union> (CX1S \<times> UNIV)) \<subseteq> Csafe \<and>
(\<Union>X0S) \<subseteq> X0 \<and>
(\<Union>X0S) \<subseteq> (CX1S \<times> UNIV) \<and>
fst ` (X1S \<union> (CX1S \<times> UNIV)) \<inter> guards = {} \<and>
(\<forall>(x, d) \<in> (CX1S \<times> UNIV). ode x \<bullet> normal sctn < 0) \<and>
fst ` X1S \<subseteq> sbelow_halfspace sctn"])
subgoal by (refine_vcg)
subgoal for A B
apply (refine_vcg)
subgoal
apply (auto simp: dest!: flowpipe_imp_flowsto)
apply (rule flowsto_subset)
apply (rule flowsto_stays_sbelow[where sctn=sctn])
apply (rule flowsto_subset) apply assumption
apply (rule order_refl)
apply force
apply (rule order_refl)
apply (rule order_refl)
apply (auto simp: halfspace_simps)
apply (rule le_less_trans)
prefer 2 apply assumption
apply (drule bspec)
apply (rule subsetD, assumption)
prefer 2 apply assumption
apply auto
done
subgoal by auto
subgoal by force
subgoal by (auto simp: dest!: flowpipe_source_subset)
subgoal by auto
subgoal
apply (auto simp: halfspace_simps subset_iff)
apply (rule le_less_trans[rotated], assumption)
by fastforce
done
subgoal by (auto intro: flowsto_subset) force
subgoal for a b c d
using assms
apply (refine_vcg, clarsimp_all)
subgoal for e f g h i j k l m n
apply (rule flowsto_source_unionI)
subgoal
apply (drule flowpipe_imp_flowsto, assumption)
apply (rule flowsto_subset[OF flowsto_stays_sbelow[where sctn=sctn] order_refl])
apply (rule flowsto_subset[OF _ order_refl], assumption)
apply force
apply (rule order_refl)
apply (rule order_refl)
apply (auto simp: halfspace_simps)
apply (rule le_less_trans)
prefer 2 apply assumption
apply (drule bspec)
apply (rule subsetD, assumption)
prefer 2 apply assumption
apply auto
done
by (auto intro!: flowsto_source_unionI dest!: flowpipe_imp_flowsto intro: flowsto_subset[OF _ order_refl])
subgoal
apply (auto simp: subset_iff)
apply (auto simp: image_Un)
done
subgoal by auto
subgoal by (auto dest!: flowpipe_source_subset)
subgoal by auto
subgoal
apply (auto simp: halfspace_simps subset_iff)
apply (rule le_less_trans[rotated], assumption)
by fastforce
subgoal by auto
done
subgoal by auto
done
lemma op_inter_fst_coll[le, refine_vcg]: "op_inter_fst_coll X Y \<le> SPEC (\<lambda>R. R = X \<inter> Y \<times> UNIV)"
unfolding op_inter_fst_coll_def
by (refine_vcg FORWEAK_mono_rule[where I="\<lambda>Xs R. \<Union>Xs \<inter> Y \<times> UNIV \<subseteq> R \<and> R \<subseteq> X \<inter> Y \<times> UNIV"])
auto
lemma scaleRe_ivl_coll_spec[le, refine_vcg]: "scaleRe_ivl_coll_spec l u X \<le> SPEC (\<lambda>Y. Y = scaleR2 l u X)"
unfolding scaleRe_ivl_coll_spec_def
apply (refine_vcg FORWEAK_mono_rule[where I="\<lambda>Xs R. scaleR2 l u (\<Union>Xs) \<subseteq> R \<and> R \<subseteq> scaleR2 l u X"])
apply (auto simp: intro: scaleR2_subset)
subgoal
by (force simp: intro: scaleR2_subset)
done
lemma do_intersection_spec_scaleR2I:
"do_intersection_spec UNIV sctns ivl sctn (scaleR2 x1 x2 baa) (scaleR2 x1 x2 aca, x1b)"
if "do_intersection_spec UNIV sctns ivl sctn (baa) (aca, x1b)"
using that
by (auto simp: do_intersection_spec_def intro!: poincare_mapsto_scaleR2I)
(auto simp: scaleR2_def image_def vimage_def)
lemma do_intersection_core[refine_vcg, le]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "do_intersection_core sctns ivl sctn (X::'n eucl1 set) \<le>
SPEC (\<lambda>(P1, P2, CX, X0s).
do_intersection_spec UNIV sctns ivl sctn (X - X0s) (P1, CX) \<and>
do_intersection_spec UNIV sctns ivl sctn (X - X0s) (P2, CX)
\<and> fst ` (X - X0s) \<subseteq> CX
\<and> X0s \<subseteq> X)"
unfolding do_intersection_core_def autoref_tag_defs
apply (refine_vcg assms, clarsimp_all)
subgoal by (rule do_intersection_spec_scaleR2I) (auto simp: do_intersection_spec_def intro!: )
subgoal by (rule do_intersection_spec_scaleR2I) (auto simp: do_intersection_spec_def intro!: )
subgoal by (fastforce simp: scaleR2_def)
subgoal by (auto simp: do_intersection_spec_def)
subgoal by (auto simp: do_intersection_spec_def)
done
lemma do_intersection_spec_Union:
"do_intersection_spec S sctns ivl sctn (\<Union>X) A"
if "\<And>x. x \<in> X \<Longrightarrow> do_intersection_spec S sctns ivl sctn x A"
"X \<noteq> {}"
using that(2)
unfolding do_intersection_spec_def
apply clarsimp
apply safe
subgoal by (rule poincare_mapsto_Union) (auto simp: do_intersection_spec_def dest!: that(1))
subgoal by (auto simp: do_intersection_spec_def dest!: that(1))
subgoal by (auto simp: do_intersection_spec_def dest!: that(1))
subgoal by (fastforce simp: do_intersection_spec_def dest!: that(1))
subgoal by (fastforce simp: do_intersection_spec_def dest!: that(1))
subgoal by (fastforce simp: do_intersection_spec_def dest!: that(1))
subgoal by (force simp: do_intersection_spec_def dest!: that(1))
subgoal by (auto simp: do_intersection_spec_def dest!: that(1))
subgoal by (fastforce simp: do_intersection_spec_def dest!: that(1))
subgoal by (fastforce simp: do_intersection_spec_def dest!: that(1))
done
lemma do_intersection_spec_subset2:
"do_intersection_spec S p ivl sctn X1 (ab, CY) \<Longrightarrow> CY \<subseteq> CX \<Longrightarrow> CX \<subseteq> Csafe \<Longrightarrow>
CX \<inter> p = {} \<Longrightarrow> CX \<inter> ivl \<inter> plane_of sctn = {} \<Longrightarrow> X0 \<subseteq> X1 \<Longrightarrow>
do_intersection_spec S p ivl sctn X0 (ab, CX)"
by (auto simp: do_intersection_spec_def intro: poincare_mapsto_subset)
lemma do_intersection_spec_Union3:
"do_intersection_spec S osctns ivl csctns (\<Union>x\<in>X. a x) ((\<Union>x\<in>X. b x), (\<Union>x\<in>X. c x))"
if "finite X" "X \<noteq> {}" "\<And>x. x \<in> X \<Longrightarrow> do_intersection_spec S osctns ivl csctns (a x) (b x, c x)"
using that
proof induction
case empty
then show ?case by (auto simp: )
next
case (insert x F)
show ?case
apply (cases "F = {}")
subgoal using insert by simp
subgoal
apply simp
apply (rule do_intersection_spec_union)
apply (rule insert.prems) apply simp
apply (rule insert.IH)
apply (assumption)
apply (rule insert.prems) apply simp
done
done
qed
lemma do_intersection_coll[le]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "do_intersection_coll sctns ivl sctn (X::'n eucl1 set) \<le>
SPEC (\<lambda>(P1, P2, CX, X0s).
do_intersection_spec UNIV sctns ivl sctn (X - X0s) (P1, CX) \<and>
do_intersection_spec UNIV sctns ivl sctn (X - X0s) (P2, CX)
\<and> fst ` (X - X0s) \<subseteq> CX
\<and> X0s \<subseteq> X)"
unfolding do_intersection_coll_def autoref_tag_defs
apply (refine_vcg, clarsimp_all)
subgoal
apply (rule do_intersection_spec_subset[OF _ diff_subset])
apply (rule do_intersection_spec_Union3)
subgoal by auto
subgoal by auto
subgoal by auto
done
subgoal
apply (rule do_intersection_spec_subset[OF _ diff_subset])
apply (rule do_intersection_spec_Union3)
subgoal by auto
subgoal by auto
subgoal by auto
done
subgoal by fastforce
subgoal by fastforce
done
lemma
do_intersection_flowsto_trans_outside:
assumes "flowsto XS0 {0..} (CX \<times> UNIV) X1"
assumes "do_intersection_spec UNIV guards ivl sctn X1 (P, CP)"
assumes "fst ` X1 \<subseteq> CP"
assumes "{x \<in> ivl. x \<in> plane_of sctn} \<inter> CX = {}"
assumes "guards \<inter> (CX \<union> CP) = {}"
assumes "XS0 \<subseteq> CX \<times> UNIV"
assumes "closed ivl"
assumes "CX \<subseteq> Csafe"
shows "do_intersection_spec UNIV guards ivl sctn XS0 (P, CX \<union> CP)"
using assms
apply (auto simp: do_intersection_spec_def)
subgoal
apply (rule flowsto_poincare_trans, assumption, assumption)
subgoal by simp
subgoal by auto
subgoal using assms(3) by auto
subgoal by (auto intro!: closed_levelset_within continuous_intros simp: plane_of_def)
subgoal premises prems for x d
proof -
have [intro, simp]: "closed {x \<in> ivl. x \<in> plane_of sctn} " "closed {x \<in> ivl. x \<bullet> normal sctn = pstn sctn}"
by (auto intro!: closed_levelset_within continuous_intros simp: plane_of_def assms)
from flowsto_poincare_mapsto_trans_flowsto[OF \<open>flowsto _ _ _ _\<close> \<open>poincare_mapsto _ _ _ _ _\<close> _ _ order_refl]
have ft: "flowsto XS0 {0<..} (X1 \<union> CX \<times> UNIV \<union> CP \<times> UNIV) (fst ` P \<times> UNIV)"
by (auto simp: )
then have ret: "returns_to {x \<in> ivl. x \<bullet> normal sctn - pstn sctn = 0} x"
apply (rule returns_to_flowstoI[OF _ _ _ _ _ _ order_refl])
using prems by (auto simp: plane_of_def)
have pm: "poincare_map {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} x \<in> fst ` P"
apply (rule poincare_map_mem_flowstoI[OF ft])
using prems by (auto simp: plane_of_def)
from pm prems have "\<forall>\<^sub>F x in at (poincare_map {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} x) within
plane_of sctn. x \<in> ivl"
by auto
from ret have "isCont (return_time {x \<in> ivl. x \<bullet> normal sctn - pstn sctn = 0}) x"
apply (rule return_time_isCont_outside)
using prems pm
by (auto simp: eventually_at_filter plane_of_def intro!: assms derivative_eq_intros)
then show "isCont (return_time {x \<in> ivl. x \<in> plane_of sctn}) x" by (simp add: plane_of_def)
qed
subgoal by simp
done
done
lemma do_intersection_coll_flowsto[le]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
assumes ft: "flowsto X0 {0..} (CX0 \<times> UNIV) X"
assumes X_subset: "X \<subseteq> CX0 \<times> UNIV"
assumes X0_subset: "X0 \<subseteq> CX0 \<times> UNIV" and CX0_safe: "CX0 \<subseteq> Csafe"
assumes ci: "closed ivl"
assumes disj: "ivl \<inter> plane_of sctn \<inter> CX0 = {}" "sctns \<inter> CX0 = {}"
shows "do_intersection_coll sctns ivl sctn (X::'n eucl1 set) \<le>
SPEC (\<lambda>(P1, P2, CX, X0s).
\<exists>A.
do_intersection_spec UNIV sctns ivl sctn A (P1, CX0 \<union> CX) \<and>
do_intersection_spec UNIV sctns ivl sctn A (P2, CX0 \<union> CX) \<and>
flowsto (X0 - A) {0..} (CX0 \<times> UNIV) X0s \<and>
A \<subseteq> X0 \<and>
P1 \<inter> X0s = {} \<and>
P2 \<inter> X0s = {})"
apply (rule do_intersection_coll)
apply (rule wd)
proof (clarsimp, goal_cases)
case (1 P1 P2 CX R)
from ft have "flowsto X0 {0..} (CX0 \<times> UNIV) (X - R \<union> R)"
by (rule flowsto_subset) auto
from flowsto_union_DiffE[OF this]
obtain A where AB: "A \<subseteq> X0"
and A: "flowsto A {0..} (CX0 \<times> UNIV) (X - R)"
and B: "flowsto (X0 - A) {0..} (CX0 \<times> UNIV) (R)"
by auto
have di: "do_intersection_spec UNIV sctns ivl sctn A (P1, CX0 \<union> CX)"
apply (rule do_intersection_flowsto_trans_outside[OF A 1(1)])
subgoal using 1 by simp
subgoal using disj by auto
subgoal using 1 disj by (auto simp: do_intersection_spec_def)
subgoal using X0_subset AB by (auto simp: do_intersection_spec_def)
subgoal using ci by simp
subgoal using CX0_safe .
done
then have "P1 \<subseteq> (ivl \<inter> plane_of sctn) \<times> UNIV"
by (auto simp: do_intersection_spec_def)
then have disjoint: "P1 \<inter> R = {}"
using \<open>R \<subseteq> X\<close> disj X_subset
apply (auto simp: subset_iff)
by (metis (no_types, lifting) Int_iff disjoint_iff_not_equal)
have di2: "do_intersection_spec UNIV sctns ivl sctn A (P2, CX0 \<union> CX)"
apply (rule do_intersection_flowsto_trans_outside[OF A 1(2)])
subgoal using 1 by simp
subgoal using disj by auto
subgoal using 1 disj by (auto simp: do_intersection_spec_def)
subgoal using X0_subset AB by (auto simp: do_intersection_spec_def)
subgoal using ci by simp
subgoal using CX0_safe .
done
then have "P2 \<subseteq> (ivl \<inter> plane_of sctn) \<times> UNIV"
by (auto simp: do_intersection_spec_def)
then have "P2 \<inter> R = {}"
using \<open>R \<subseteq> X\<close> disj X_subset
apply (auto simp: subset_iff)
by (metis (no_types, lifting) Int_iff disjoint_iff_not_equal)
from AB this disjoint di di2 B show ?case
by (auto simp:)
qed
lemma op_enlarge_ivl_sctn[le, refine_vcg]:
"op_enlarge_ivl_sctn ivl sctn d \<le> SPEC (\<lambda>ivl'. ivl \<subseteq> ivl')"
unfolding op_enlarge_ivl_sctn_def
apply refine_vcg
unfolding plane_of_def
apply (safe intro!: eventually_in_planerectI)
apply (auto intro!: simp: eucl_le[where 'a='a] inner_sum_left inner_Basis if_distrib
algebra_simps cong: if_cong)
done
lemma resolve_ivlplanes[le]:
assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)"
assumes
"\<forall>x\<in>Xg. case x of (I, G) \<Rightarrow> flowsto (XSf G) {0..} (CXS \<times> UNIV) G"
"(\<Union>x\<in>Xg. snd x) \<subseteq> (Csafe - (ivlplanes \<union> guards)) \<times> UNIV"
"CXS \<times> UNIV \<subseteq> (Csafe - (ivlplanes \<union> guards)) \<times> UNIV"
"(\<Union>a\<in>Xg. XSf (snd a)) \<subseteq> (CXS::'a rvec set) \<times> UNIV"
"(\<Union>x\<in>Xg. snd x) \<subseteq> CXS \<times> UNIV"
"(\<Union>x\<in>Xg. fst x) \<subseteq> ivlplanes \<union> guards"
shows "resolve_ivlplanes guards ivlplanes Xg \<le> SPEC (\<lambda>PS.
CXS \<inter> (guards \<union> ivlplanes) = {} \<and>
CXS \<subseteq> Csafe \<and>
(\<exists>R0 P0. (\<Union>x\<in>PS. P0 x) \<union> (\<Union>x\<in>PS. R0 x) = (\<Union>a\<in>Xg. XSf (snd a))\<and>
(\<forall>x\<in>PS. case x of (X, P1, P2, R, ivl, sctn, CX) \<Rightarrow>
ivl \<inter> plane_of sctn \<subseteq> ivlplanes \<and> closed ivl \<and>
P0 (X, P1, P2, R, ivl, sctn, CX) \<inter> R0 (X, P1, P2, R, ivl, sctn, CX) = {} \<and>
R0 (X, P1, P2, R, ivl, sctn, CX) \<subseteq> (CXS \<times> UNIV) \<and>
flowsto (R0 (X, P1, P2, R, ivl, sctn, CX)) {0..} (CXS \<times> UNIV) R \<and>
do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX)) (P1, CXS \<union> CX) \<and>
do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX)) (P2, CXS \<union> CX))))"
using assms
unfolding resolve_ivlplanes_def
apply clarsimp_all
apply (refine_vcg FORWEAK_mono_rule[where I="\<lambda>Xgs PS.
(\<exists>R0 P0.
snd ` Xgs \<subseteq> fst ` PS \<and> fst ` PS \<subseteq> snd ` Xg \<and>
(\<forall>(X, P1, P2, R, ivl, sctn, CX) \<in> PS.
P0 (X, P1, P2, R, ivl, sctn, CX) \<union> R0 (X, P1, P2, R, ivl, sctn, CX) = XSf X
\<and> ivl \<inter> plane_of sctn \<subseteq> ivlplanes \<and> closed ivl
\<and> P0 (X, P1, P2, R, ivl, sctn, CX) \<inter> R0 (X, P1, P2, R, ivl, sctn, CX) = {}
\<and> R0 (X, P1, P2, R, ivl, sctn, CX) \<subseteq> (CXS \<times> UNIV)
\<and> flowsto (R0 (X, P1, P2, R, ivl, sctn, CX)) {0..} (CXS \<times> UNIV) R
\<and> do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX)) (P1, CXS \<union> CX)
\<and> do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX)) (P2, CXS \<union> CX)))"],
clarsimp_all)
using [[goals_limit=1]]
subgoal by auto
subgoal by auto
subgoal for a b c
apply (frule bspec, assumption, clarsimp)
apply (rule do_intersection_coll_flowsto)
apply (rule wd)
apply assumption
apply force
apply force
apply blast
apply assumption
subgoal premises prems
proof -
have "(b \<inter> plane_of c, a) \<in> Xg" using prems by simp
with \<open>(\<Union>x\<in>Xg. fst x) \<subseteq> ivlplanes \<union> guards\<close>
have "b \<inter> plane_of c \<subseteq> ivlplanes \<union> guards"
by (force simp: subset_iff)
then show ?thesis
using \<open>CXS \<times> UNIV \<subseteq> (Csafe - (ivlplanes \<union> guards)) \<times> UNIV\<close>
by auto
qed
subgoal by (auto simp: subset_iff)
subgoal apply (refine_vcg, clarsimp_all) apply force
apply (intro exI conjI)defer defer defer apply assumption+
apply simp
apply force
apply force
apply force
done
done
subgoal by (auto simp: subset_iff) blast
subgoal for a b c d e f R0 P0
apply (frule bspec, assumption, clarsimp)
apply (rule do_intersection_coll_flowsto)
apply (rule wd)
apply assumption
subgoal
apply (rule order_trans[where y="(\<Union>x\<in>Xg. snd x)"])
by auto
subgoal
apply (rule order_trans) defer apply assumption
by auto
subgoal by blast
subgoal by simp
subgoal premises prems
proof -
have "(d \<inter> plane_of e, c) \<in> Xg" using prems by simp
with \<open>(\<Union>x\<in>Xg. fst x) \<subseteq> ivlplanes \<union> guards\<close>
have "d \<inter> plane_of e \<subseteq> ivlplanes \<union> guards"
by (force simp: subset_iff)
then show ?thesis
using \<open>CXS \<times> UNIV \<subseteq> (Csafe - (ivlplanes \<union> guards)) \<times> UNIV\<close>
by auto
qed
subgoal by (auto simp: subset_iff)
subgoal
apply (refine_vcg, clarsimp_all)
subgoal by (auto simp: subset_iff)
subgoal by (auto simp: )
subgoal for x1 x1' x2 x3 A
apply (rule exI[where x="R0((c, x1, x1', x3, d, e, x2):=(XSf c - A))"])
apply (rule exI[where x="P0((c, x1, x1', x3, d, e, x2):=A)"])
apply clarsimp
apply (rule conjI)
subgoal by auto
apply (rule conjI)
subgoal premises prems
using prems
apply (auto simp: subset_iff)
by fastforce
apply clarsimp
subgoal
apply (drule bspec, assumption)
apply (drule bspec, assumption)
by force
done
done
done
subgoal by (auto simp: subset_iff)
subgoal by (auto simp: subset_iff)
subgoal for a R0 P0
apply (rule exI[where x=R0])
apply (rule exI[where x=P0])
apply (rule conjI)
subgoal premises prems
proof -
note prems
show ?thesis
using prems(9,8)
by fastforce
qed
by auto
done
lemma poincare_onto[le, refine_vcg]:
assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)"
assumes [refine_vcg]: "\<And>X0. X0 \<subseteq> Csafe \<times> UNIV \<Longrightarrow> symstart X0 \<le>
SPEC (\<lambda>(CX, X). flowsto (X0 - trap \<times> UNIV) {0..} (CX \<times> UNIV) X)"
assumes CXS0: "CXS0 \<inter> (guards \<union> ivlplanes) = {}"
shows "poincare_onto ro symstart trap guards ivlplanes (XS0::'a eucl1 set) CXS0 \<le>
SPEC (\<lambda>PS.
(\<exists>R0 P0.
\<Union>(P0 ` PS \<union> R0 ` PS) = XS0 - trap \<times> UNIV \<and>
(\<forall>(X, P1, P2, R, ivl, sctn, CX, CXS) \<in> PS.
ivl \<inter> plane_of sctn \<subseteq> ivlplanes \<and> closed ivl
\<and> XS0 \<subseteq> CXS \<times> UNIV \<and> CXS0 \<subseteq> CXS \<and> CXS \<inter> (guards \<union> ivlplanes) = {}
\<and> P0 (X, P1, P2, R, ivl, sctn, CX, CXS) \<inter> R0 (X, P1, P2, R, ivl, sctn, CX, CXS) = {}
\<and> R0 (X, P1, P2, R, ivl, sctn, CX, CXS) \<subseteq> CXS \<times> UNIV
\<and> flowsto (R0 (X, P1, P2, R, ivl, sctn, CX, CXS)) {0..} (CXS \<times> UNIV) R
\<and> do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, CXS \<union> CX)
\<and> do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, CXS \<union> CX))
))"
unfolding poincare_onto_def autoref_tag_defs
using [[goals_limit=1]]
apply (refine_vcg, clarsimp_all)
apply (refine_vcg resolve_ivlplanes[OF wd])
subgoal by force
apply clarsimp
subgoal for a b c d R0 P0
apply (rule exI[where x="\<lambda>(X, P1, P2, R, ivl, sctn, CX, CXS). R0 (X, P1, P2, R, ivl, sctn, CX)"])
apply (rule exI[where x="\<lambda>(X, P1, P2, R, ivl, sctn, CX, CXS). P0 (X, P1, P2, R, ivl, sctn, CX)"])
apply (rule conjI)
subgoal premises prems
using \<open>(\<Union>x\<in>d. P0 x) \<union> (\<Union>x\<in>d. R0 x) = (\<Union>x\<in>b. c (snd x)) - trap \<times> UNIV\<close>
by auto
subgoal
apply clarsimp
apply (drule bspec, assumption)+
apply (rule conjI, force)
apply (rule conjI, force)
apply (rule conjI, force)
apply (rule conjI)
subgoal using CXS0 by (auto simp: )
apply (rule conjI, force)
apply (rule conjI, force)
apply (rule conjI)
subgoal by (auto intro: flowsto_subset)
subgoal
apply clarsimp
apply (rule conjI)
subgoal
apply (rule do_intersection_spec_subset2, assumption)
subgoal by force
subgoal by (force simp: do_intersection_spec_def)
subgoal using CXS0 by (auto simp: do_intersection_spec_def)
subgoal using CXS0 by (auto simp: do_intersection_spec_def)
subgoal by auto
done
subgoal
apply (rule do_intersection_spec_subset2, assumption)
subgoal by force
subgoal by (force simp: do_intersection_spec_def)
subgoal using CXS0 by (auto simp: do_intersection_spec_def)
subgoal using CXS0 by (auto simp: do_intersection_spec_def)
subgoal by auto
done
done
done
done
done
lemma empty_remainders[le, refine_vcg]:
"empty_remainders PS \<le> SPEC (\<lambda>b. b \<longrightarrow> (\<forall>(X, P1, P2, R, ivl, sctn, CX) \<in> PS. R = {}))"
unfolding empty_remainders_def
by (refine_vcg FORWEAK_mono_rule[where I="\<lambda>Xs b. b \<longrightarrow> (\<forall>(X, P1, P2, R, ivl, sctn, CX) \<in> Xs. R = {})"])
auto
lemma poincare_onto_empty[le, refine_vcg]:
assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)"
assumes CXS0: "CXS0 \<inter> (guards \<union> ivlplanes) = {}"
shows "poincare_onto_empty ro guards ivlplanes (XS0::'a eucl1 set) CXS0 \<le>
SPEC (\<lambda>(PS).
(\<exists>R0 P0.
\<Union>(P0 ` PS \<union> R0 ` PS) = XS0 \<and>
(\<forall>(X, P1, P2, R, ivl, sctn, CX, CXS) \<in> PS.
ivl \<inter> plane_of sctn \<subseteq> ivlplanes \<and> closed ivl
\<and> XS0 \<subseteq> CXS \<times> UNIV \<and> CXS0 \<subseteq> CXS \<and> CXS \<inter> (guards \<union> ivlplanes) = {}
\<and> P0 (X, P1, P2, R, ivl, sctn, CX, CXS) \<inter> R0 (X, P1, P2, R, ivl, sctn, CX, CXS) = {}
\<and> R0 (X, P1, P2, R, ivl, sctn, CX, CXS) \<subseteq> CXS \<times> UNIV
\<and> flowsto (R0 (X, P1, P2, R, ivl, sctn, CX, CXS)) {0..} (CXS \<times> UNIV) R
\<and> do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, CXS \<union> CX)
\<and> do_intersection_spec UNIV guards ivl sctn (P0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, CXS \<union> CX))
))"
using CXS0
unfolding poincare_onto_empty_def autoref_tag_defs
by (refine_vcg) (auto intro!: flowsto_self)
lemma do_intersection_spec_union2:
assumes "do_intersection_spec S osctns ivl csctns a (b, c)"
"do_intersection_spec S osctns ivl csctns f (b, c)"
shows "do_intersection_spec S osctns ivl csctns (a \<union> f) (b, c)"
using do_intersection_spec_union[OF assms]
by auto
lemma poincare_onto2[le, refine_vcg]:
assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)"
assumes [refine_vcg]: "\<And>X0. X0 \<subseteq> Csafe \<times> UNIV \<Longrightarrow> symstart X0 \<le>
SPEC (\<lambda>(CX, X). flowsto (X0 - trap \<times> UNIV) {0..} (CX \<times> UNIV) X)"
notes [refine_vcg_def] = op_set_ndelete_spec
shows "poincare_onto2 ro symstart trap guards ivlplanes (XS0::'a eucl1 set) \<le>
SPEC (\<lambda>(PS).
(\<exists>P0. \<Union>(P0 ` PS) = XS0 - trap \<times> UNIV \<and>
(\<forall>(s, X, P1, P2, R, ivl, sctn, CX, CXS) \<in> PS.
XS0 \<subseteq> CXS \<times> UNIV \<and>
do_intersection_spec UNIV guards ivl sctn (P0 (s, X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, CXS \<union> CX) \<and>
do_intersection_spec UNIV guards ivl sctn (P0 (s, X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, CXS \<union> CX))))"
unfolding poincare_onto2_def autoref_tag_defs
apply (refine_vcg, clarsimp_all)
subgoal for PS R0 P0
apply (rule FORWEAK_mono_rule_empty[where I="\<lambda>PS1 PS2.
(\<exists>X0.
\<Union>(R0 ` PS1) \<subseteq> \<Union>(X0 ` PS2) \<and>
(\<forall>(X, P1, P2, R, ivl, sctn, CX, CXS) \<in> PS2.
XS0 \<subseteq> CXS \<times> UNIV \<and>
do_intersection_spec UNIV guards ivl sctn (X0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, CXS \<union> CX) \<and>
do_intersection_spec UNIV guards ivl sctn (X0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, CXS \<union> CX)))"])
subgoal by refine_vcg
subgoal by auto
subgoal by auto
subgoal
apply clarsimp
subgoal for c
apply (rule exI[where x=c])
apply (rule conjI)
apply (rule order_trans) prefer 2 apply assumption
apply (rule UN_mono) apply assumption apply (rule order_refl) apply assumption
done
done
subgoal for \<sigma>
apply (clarsimp)
subgoal for X0
apply (rule exI[where x="\<lambda>(b, x). (if b then X0 x else P0 x) \<inter> XS0 - trap \<times> UNIV "])
apply (rule conjI)
subgoal premises prems
using \<open>(\<Union>x\<in>PS. P0 x) \<union> (\<Union>x\<in>PS. R0 x) = XS0 - trap \<times> UNIV\<close>
\<open>(\<Union>x\<in>PS. R0 x) \<subseteq> (\<Union>x\<in>\<sigma>. X0 x)\<close>
by auto
subgoal by (auto intro: do_intersection_spec_subset)
done
done
apply clarsimp
subgoal for a b b' c d e f g h i j
apply (cases "c = {}")
subgoal by (auto intro!: exI[where x="j"])
subgoal
using [[goals_limit=1]]
apply clarsimp
apply refine_vcg
subgoal premises prems for k l
proof -
note prems
then show ?thesis
apply -
apply (drule bspec, assumption)+
apply clarsimp
subgoal premises prems
using \<open>g \<inter> (guards \<union> \<Union>k) = {}\<close> \<open>l = k - {d \<inter> plane_of e} \<or> l = k\<close> \<open>d \<inter> plane_of e \<subseteq> \<Union>k\<close>
by auto
done
qed
apply simp
apply (drule bspec, assumption)
apply simp
apply (erule exE conjE)+
subgoal for k l m n p q
apply (subgoal_tac "\<And>x. x \<in> m \<Longrightarrow> p x = {}")
defer
subgoal for x
proof goal_cases
case 1
from 1(10,15,24)
show ?case
by (auto dest!: bspec[where x=x])
qed
apply simp
subgoal premises prems
proof -
note prems
from prems have "finite (q ` m)" "flowsto (R0 (a, b, b', c, d, e, f, g)) {0..} (g \<times> UNIV) (\<Union>(q ` m))"
by auto
from flowsto_Union_funE[OF this]
obtain XGs where
XGs: "\<And>G. G \<in> q ` m \<Longrightarrow> flowsto (XGs G) {0..} (g \<times> UNIV) G"
"R0 (a, b, b', c, d, e, f, g) = \<Union>(XGs ` (q ` m))"
by metis
define q0 where "q0 = XGs o q"
have "case x of (X, P1, P2, R, ivl, sctn, CX, CXS) \<Rightarrow>
do_intersection_spec UNIV guards ivl sctn (q0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, CXS \<union> CX) \<and>
do_intersection_spec UNIV guards ivl sctn (q0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, CXS \<union> CX)"
if "x \<in> m"
for x
proof (clarsimp, goal_cases)
case (1 X P1 P2 R ivl sctn CX CXS)
with prems(10)[rule_format, OF \<open>x \<in> m\<close>] prems(15)[rule_format, OF \<open>x \<in> m\<close>] \<open>_ = c\<close>
have *: "R = {}"
"x = (X, P1, P2, {}, ivl, sctn, CX, CXS)"
"ivl \<inter> plane_of sctn \<subseteq> \<Union>l"
"closed ivl"
"c \<subseteq> CXS \<times> UNIV"
"g \<subseteq> CXS"
"\<Union>(q ` m) \<subseteq> CXS \<times> UNIV"
"CXS \<inter> (guards \<union> \<Union>l) = {}"
"p (X, P1, P2, {}, ivl, sctn, CX, CXS) = {}"
"p (X, P1, P2, R, ivl, sctn, CX, CXS) \<subseteq> CXS \<times> UNIV"
"do_intersection_spec UNIV guards ivl sctn (q (X, P1, P2, {}, ivl, sctn, CX, CXS)) (P1, CXS \<union> CX)"
"do_intersection_spec UNIV guards ivl sctn (q (X, P1, P2, {}, ivl, sctn, CX, CXS)) (P2, CXS \<union> CX)"
by auto
have "do_intersection_spec UNIV guards ivl sctn (q0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P1, (CXS \<union> CX) \<union> (CXS \<union> CX))"
apply (rule do_intersection_flowsto_trans_outside)
apply (simp add: q0_def)
apply (rule flowsto_subset)
apply (rule XGs)
using \<open>x \<in> m\<close> apply (rule imageI)
using 1 apply force
apply force
using * apply force
apply (rule order_refl)
using * apply (auto intro!: *)[]
subgoal
using * \<open>x \<in> m\<close>
by (auto simp add: )
subgoal using * by (auto simp: do_intersection_spec_def)
subgoal using * by (auto simp: do_intersection_spec_def)
subgoal
proof -
have "q0 (X, P1, P2, R, ivl, sctn, CX, CXS) \<subseteq> XGs (q x)"
by (auto simp: q0_def 1)
also have "\<dots> \<subseteq> R0 (a, b, b', c, d, e, f, g)" using \<open>x \<in>m\<close> XGs by auto
also have "\<dots> \<subseteq> (CXS \<union> CX) \<times> UNIV"
using prems(20) \<open>g \<subseteq> CXS\<close> by auto
finally show ?thesis by simp
qed
subgoal by fact
subgoal using * by (auto simp: do_intersection_spec_def)
done
moreover have "do_intersection_spec UNIV guards ivl sctn (q0 (X, P1, P2, R, ivl, sctn, CX, CXS)) (P2, (CXS \<union> CX) \<union> (CXS \<union> CX))"
apply (rule do_intersection_flowsto_trans_outside)
apply (simp add: q0_def)
apply (rule flowsto_subset)
apply (rule XGs)
using \<open>x \<in> m\<close> apply (rule imageI)
using 1 apply force
apply force
using * apply force
apply (rule order_refl)
using * apply (auto intro!: *)[]
subgoal
using * \<open>x \<in> m\<close>
by (auto simp add: )
subgoal using * by (auto simp: do_intersection_spec_def)
subgoal using * by (auto simp: do_intersection_spec_def)
subgoal
proof -
have "q0 (X, P1, P2, R, ivl, sctn, CX, CXS) \<subseteq> XGs (q x)"
by (auto simp: q0_def 1)
also have "\<dots> \<subseteq> R0 (a, b, b', c, d, e, f, g)" using \<open>x \<in>m\<close> XGs by auto
also have "\<dots> \<subseteq> (CXS \<union> CX) \<times> UNIV"
using prems(20) \<open>g \<subseteq> CXS\<close> by auto
finally show ?thesis by simp
qed
subgoal by fact
subgoal using * by (auto simp: do_intersection_spec_def)
done
ultimately show ?case
by (simp add: )
qed note q0 = this
have q0': "(a, aa, aa', ab, ac, ad, ae, b) \<in> m \<Longrightarrow> XS0 \<subseteq> b \<times> UNIV" for a aa aa' ab ac ad ae b
apply (drule prems(15)[rule_format])
using \<open>XS0 \<subseteq> g \<times> UNIV\<close>
by auto
from prems
show ?thesis
apply (intro exI[where x="\<lambda>x. if x \<in> i \<inter> m then j x \<union> q0 x else if x \<in> i then j x else q0 x"] conjI)
subgoal 1 premises prems
unfolding XGs
apply simp
by (auto simp: q0_def)
subgoal premises _
by (rule order_trans[OF \<open>(\<Union>x\<in>h. R0 x) \<subseteq> (\<Union>x\<in>i. j x)\<close>]) auto
subgoal premises _ using prems(6)[rule_format] q0
apply auto
subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2)
subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2)
subgoal by (auto intro!: do_intersection_spec_union2)
subgoal by (auto dest!: prems(6)[rule_format] q0' intro!: do_intersection_spec_union2)
subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2)
subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2)
subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2)
subgoal by (auto dest!: prems(6)[rule_format] q0 intro!: do_intersection_spec_union2)
done
done
qed
done
done
done
done
done
lemma width_spec_ivl[THEN order_trans, refine_vcg]: "width_spec_ivl M X \<le> SPEC (\<lambda>x. True)"
unfolding width_spec_ivl_def
by (refine_vcg)
lemma partition_ivl_spec[le, refine_vcg]:
shows "partition_ivl cg XS \<le> SPEC (\<lambda>YS. XS \<subseteq> YS)"
+ using [[simproc del: defined_all]]
unfolding partition_ivl_def autoref_tag_defs
apply (refine_vcg, clarsimp_all)
subgoal by fastforce
subgoal by fastforce
subgoal by fastforce
subgoal by fastforce
subgoal premises prems for a b c d e f ws g h i j k l m n
proof -
note prems
have disj: "\<And>A Aa. n \<notin> A \<or> \<not> XS \<inter> A \<subseteq> Aa \<or> n \<in> Aa"
using prems by blast
then have "n \<in> g"
using prems by (metis (no_types) Un_iff atLeastAtMost_iff subset_iff)
then show ?thesis
using disj prems by (meson atLeastAtMost_iff)
qed
done
lemma op_inter_fst_ivl_scaleR2[le,refine_vcg]:
"op_inter_fst_ivl_scaleR2 X Y \<le> SPEC (\<lambda>R. X \<inter> (Y \<times> UNIV) = R)"
unfolding op_inter_fst_ivl_scaleR2_def
apply refine_vcg
apply (auto simp: scaleR2_def)
subgoal for a b c d e f g h i j k
by (rule image_eqI[where x="(i, (j, k))"]; fastforce)
subgoal for a b c d e f g h i j k
by (rule image_eqI[where x="(i, (j, k))"]; fastforce)
done
lemma op_inter_fst_ivl_coll_scaleR2[le,refine_vcg]:
"op_inter_fst_ivl_coll_scaleR2 X Y \<le> SPEC (\<lambda>R. X \<inter> (Y \<times> UNIV) = R)"
unfolding op_inter_fst_ivl_coll_scaleR2_def
by (refine_vcg FORWEAK_mono_rule[where I="\<lambda>Xs R. (\<Union>Xs) \<inter> (Y \<times> UNIV) \<subseteq> R \<and> R \<subseteq> X \<inter> (Y \<times> UNIV)"])
auto
lemma op_inter_ivl_co[le, refine_vcg]: "op_ivl_of_ivl_coll X \<le> SPEC (\<lambda>R. X \<subseteq> R)"
unfolding op_ivl_of_ivl_coll_def
apply (refine_vcg FORWEAK_mono_rule[where I="\<lambda>R (l, u). \<Union>R \<subseteq> {l .. u}"])
apply auto
apply (metis Set.basic_monos(7) Sup_le_iff atLeastAtMost_iff inf.coboundedI2 inf_sup_aci(1))
by (meson Set.basic_monos(7) UnionI atLeastAtMost_iff le_supI1)
lemma op_inter_ivl_coll_scaleR2[le,refine_vcg]:
"op_inter_ivl_coll_scaleR2 X Y \<le> SPEC (\<lambda>R. X \<inter> (Y \<times> UNIV) \<subseteq> R)"
unfolding op_inter_ivl_coll_scaleR2_def
apply refine_vcg
subgoal for _ _ _ A l u
by (auto, rule scaleR2_subset[where i'=l and j'=u and k'=A], auto)
done
lemma [le, refine_vcg]: "op_image_fst_ivl_coll X \<le> SPEC (\<lambda>R. R = fst ` X)"
unfolding op_image_fst_ivl_coll_def
apply (refine_vcg FORWEAK_mono_rule[where I="\<lambda>Xs R. fst ` (\<Union>Xs) \<subseteq> R \<and> R \<subseteq> fst ` X"])
apply auto
apply force+
done
lemma op_single_inter_ivl[le, refine_vcg]: "op_single_inter_ivl a fxs \<le> SPEC (\<lambda>R. a \<inter> fxs \<subseteq> R)"
unfolding op_single_inter_ivl_def
by refine_vcg auto
lemma partition_ivle_spec[le, refine_vcg]:
shows "partition_ivle cg XS \<le> SPEC (\<lambda>YS. XS \<subseteq> YS)"
unfolding partition_ivle_def autoref_tag_defs
supply [refine_vcg del] = scaleR2_rep_of_coll2
and [refine_vcg] = scaleR2_rep_of_coll
apply (refine_vcg)
subgoal by (fastforce simp: scaleR2_def)
subgoal by auto
apply clarsimp
subgoal by (fastforce simp: scaleR2_def)
done
lemma vec1repse[THEN order_trans, refine_vcg]:
"vec1repse CX \<le> SPEC (\<lambda>R. case R of None \<Rightarrow> True | Some X \<Rightarrow> X = vec1_of_flow1 ` CX)"
unfolding vec1repse_def
apply (refine_vcg FORWEAK_mono_rule[where
I="\<lambda>XS R. case R of None \<Rightarrow> True | Some R \<Rightarrow> vec1_of_flow1 ` (\<Union>XS) \<subseteq> R \<and> R \<subseteq> vec1_of_flow1 ` CX"])
apply (auto simp: scaleR2_def split: option.splits)
subgoal for a b c d e f g h i j
apply (auto simp: vimage_def image_def)
apply (rule exI[where x="h"])
apply auto
apply (rule exI[where x=f])
apply (rule exI[where x="matrix j"])
apply auto
apply (rule bexI)
by (auto simp: vec1_of_flow1_def matrix_scaleR)
subgoal for a b c d e f g h i j
apply (rule bexI)
defer apply assumption
apply (rule image_eqI[where x="(f, g, j)"])
by (auto simp: flow1_of_vec1_def vec1_of_flow1_def matrix_scaleR[symmetric])
subgoal by fastforce
subgoal for a b c d e f g h i j k l
apply (auto simp: vimage_def image_def)
apply (rule exI[where x="j"])
apply auto
apply (rule exI[where x=h])
apply (rule exI[where x="matrix l"])
apply auto
apply (rule bexI)
by (auto simp: vec1_of_flow1_def matrix_scaleR)
subgoal by fastforce
subgoal for a b c d e f g h i j k l
apply (rule bexI)
defer apply assumption
apply (rule image_eqI[where x="(h, i, l)"])
by (auto simp: flow1_of_vec1_def vec1_of_flow1_def matrix_scaleR[symmetric])
done
lemma scaleR2_rep1[le, refine_vcg]: "scaleR2_rep1 Y \<le> SPEC (\<lambda>R. Y \<subseteq> R)"
unfolding scaleR2_rep1_def
apply refine_vcg
subgoal by (auto simp: norm2_slp_def)
subgoal for a b c d e y z f g h i j prec k l m n p q r s
apply (auto simp: scaleR2_def image_def vimage_def)
subgoal premises prems for B C D E
proof -
define ij where "ij = (i + j) / 2"
from prems
have "ij > 0"
by (auto simp: ij_def)
show ?thesis
unfolding ij_def[symmetric]
apply (rule exI[where x="1 / ij * B"])
apply (intro conjI) prefer 3
apply (rule bexI[where x="(D, ij *\<^sub>R E)"])
subgoal using \<open>ij > 0\<close> by auto
subgoal
using prems
using \<open>(D, E) \<in> c\<close> \<open>c \<subseteq> {(n, p)..(q, r)}\<close> \<open>ij > 0\<close>
by (auto simp: ij_def[symmetric] intro!: scaleR_left_mono)
subgoal
using \<open>d \<le> ereal B\<close> \<open>0 < ij\<close> \<open>0 < d\<close>
apply (cases d)
apply (simp only: times_ereal.simps ereal_less_eq)
apply (rule mult_mono)
apply (rule real_divl)
by auto
subgoal
using \<open>0 < d\<close> \<open>d \<le> ereal B\<close> \<open>ereal B \<le> e\<close> \<open>0 < ij\<close> \<open>0 < e\<close>
\<open>0 < real_divr prec 1 ((i + j) / 2)\<close>
unfolding ij_def[symmetric]
apply (cases e; cases d)
apply (simp only: times_ereal.simps ereal_less_eq)
apply (rule mult_mono)
apply (rule real_divr)
by auto
done
qed
done
done
lemma reduce_ivl[le, refine_vcg]: "reduce_ivl Y b \<le> SPEC (\<lambda>R. Y \<subseteq> R)"
unfolding reduce_ivl_def
apply refine_vcg
apply (auto simp add: scaleR2_def image_def vimage_def plane_of_def )
prefer 2
subgoal using basic_trans_rules(23) by blast
prefer 3
subgoal using basic_trans_rules(23) by blast
proof goal_cases
case (1 i0 i1 s0 s1 y0 y1)
from 1 have le: "1 \<le> (y1 \<bullet> b) / (i1 \<bullet> b)"
by (auto simp: min_def dest!: inner_Basis_mono[OF _ \<open>b \<in> Basis\<close>])
show ?case
apply (rule exI[where x="(y1 \<bullet> b) / (i1 \<bullet> b)"])
apply (rule conjI) apply fact
apply (rule bexI[where x="(y0, ((i1 \<bullet> b) / (y1 \<bullet> b)) *\<^sub>R y1)"])
subgoal using 1 le by simp
subgoal using 1 le apply simp
apply (rule conjI)
subgoal
apply (auto simp: eucl_le[where 'a="'c"])
apply (auto simp: divide_simps)
apply (subst mult.commute)
subgoal for i
apply (cases " y1 \<bullet> b \<le> i1 \<bullet> b")
apply (rule order_trans)
apply (rule mult_left_mono[where b="y1 \<bullet> i"])
apply (auto simp: mult_le_cancel_right)
apply (cases "i1 \<bullet> i \<le> 0")
apply (rule order_trans)
apply (rule mult_right_mono_neg[where b="i1 \<bullet> b"])
apply auto
by (auto simp: not_le inner_Basis split: if_splits dest!: bspec[where x=i])
done
subgoal
apply (auto simp: eucl_le[where 'a="'c"])
subgoal for i
apply (cases "i = b")
apply (auto simp: divide_simps)
subgoal by (auto simp: divide_simps algebra_simps)
subgoal apply (auto simp: divide_simps algebra_simps inner_Basis)
apply (subst mult.commute)
apply (rule order_trans)
apply (rule mult_right_mono[where b="s1 \<bullet> i"]) apply simp
apply simp
apply (rule mult_left_mono)
by auto
done
done
done
done
next
case (2 i0 i1 s0 s1 y0 y1)
from 2 have le: "1 \<le> (y1 \<bullet> b) / (s1 \<bullet> b)"
by (auto simp: min_def abs_real_def divide_simps dest!: inner_Basis_mono[OF _ \<open>b \<in> Basis\<close>])
show ?case
apply (rule exI[where x="(y1 \<bullet> b) / (s1 \<bullet> b)"])
apply (rule conjI) apply fact
apply (rule bexI[where x="(y0, ((s1 \<bullet> b) / (y1 \<bullet> b)) *\<^sub>R y1)"])
subgoal using 2 le by simp
subgoal using 2 le apply simp
apply (rule conjI)
subgoal
apply (auto simp: eucl_le[where 'a="'c"])
subgoal for i
apply (cases "i = b")
apply (auto simp: divide_simps)
subgoal by (auto simp: divide_simps algebra_simps)
subgoal apply (auto simp: divide_simps algebra_simps inner_Basis)
apply (subst mult.commute)
apply (cases "y1 \<bullet> i \<le> 0")
apply (rule order_trans)
apply (rule mult_left_mono_neg[where b="y1 \<bullet> b"])
apply (auto simp: mult_le_cancel_right not_le)
apply (rule order_trans)
apply (rule mult_right_mono_neg[where b="i1 \<bullet> i"])
apply (auto intro!: mult_left_mono_neg)
done
done
done
subgoal
apply (auto simp: eucl_le[where 'a="'c"])
subgoal for i
apply (cases "i = b")
subgoal by (auto simp: divide_simps algebra_simps)
subgoal apply (auto simp: divide_simps algebra_simps inner_Basis)
apply (subst mult.commute)
apply (cases "y1 \<bullet> i \<ge> 0")
apply (rule order_trans)
apply (rule mult_left_mono_neg[where b="y1 \<bullet> i"]) apply simp
apply simp
apply (rule mult_right_mono) apply force
apply force
proof -
assume a1: "\<forall>i\<in>Basis. s1 \<bullet> b * (if b = i then 1 else 0) \<le> s1 \<bullet> i"
assume a2: "i \<in> Basis"
assume a3: "i \<noteq> b"
assume a4: "y1 \<bullet> b < 0"
assume a5: "s1 \<bullet> b < 0"
assume a6: "\<not> 0 \<le> y1 \<bullet> i"
have "s1 \<bullet> b * (if b = i then 1 else 0) \<le> s1 \<bullet> i"
using a2 a1 by metis
then have f7: "0 \<le> s1 \<bullet> i"
using a3 by (metis (full_types) mult_zero_right)
have f8: "y1 \<bullet> b \<le> 0"
using a4 by (metis eucl_less_le_not_le)
have "s1 \<bullet> b \<le> 0"
using a5 by (metis eucl_less_le_not_le)
then show "y1 \<bullet> b * (s1 \<bullet> i) \<le> s1 \<bullet> b * (y1 \<bullet> i)"
using f8 f7 a6 by (metis mult_right_mono_le mult_zero_left zero_le_mult_iff zero_le_square)
qed
done
done
done
done
qed
lemma reduce_ivle[le, refine_vcg]:
"reduce_ivle Y b \<le> SPEC (\<lambda>R. Y \<subseteq> R)"
+ using [[simproc del: defined_all]]
unfolding reduce_ivle_def
apply refine_vcg
apply (auto simp: scaleR2_def image_def vimage_def)
subgoal for a b c d e f g h i j k
apply (drule subsetD, assumption)
apply auto
subgoal for l m
apply (rule exI[where x="l * g"])
apply (intro conjI)
subgoal
unfolding times_ereal.simps[symmetric]
apply (rule ereal_mult_mono)
subgoal by (cases e) auto
subgoal by (cases b) auto
subgoal by (cases b) auto
subgoal by (cases e) auto
done
subgoal
unfolding times_ereal.simps[symmetric]
apply (rule ereal_mult_mono)
subgoal by (cases b) auto
subgoal by (cases b) auto
subgoal by (cases b) auto
subgoal by (cases e) auto
done
subgoal by force
done
done
done
lemma reduces_ivle[le, refine_vcg]:
"reduces_ivle X \<le> SPEC (\<lambda>R. X \<subseteq> R)"
unfolding reduces_ivle_def
by refine_vcg auto
lemma ivlse_of_setse[le, refine_vcg]: "ivlse_of_setse X \<le> SPEC (\<lambda>R. X \<subseteq> R)"
unfolding ivlse_of_setse_def
by (refine_vcg FORWEAK_mono_rule[where I="\<lambda>Xs R. \<Union>Xs \<subseteq> R"])
(auto simp: scaleR2_def image_def vimage_def)
lemma setse_of_ivlse[le, refine_vcg]:
"setse_of_ivlse X \<le> SPEC (\<lambda>R. R = X)"
unfolding setse_of_ivlse_def
apply (refine_vcg FORWEAK_mono_rule[where I="\<lambda>Xs R. \<Union>Xs \<subseteq> R \<and> R \<subseteq> X"])
apply clarsimp_all
subgoal by (rule bexI)
subgoal by auto
subgoal by auto
subgoal by auto
done
lemma partition_set_spec[le, refine_vcg]:
shows "partition_set ro XS \<le> SPEC (\<lambda>YS. XS \<subseteq> YS)"
unfolding partition_set_def autoref_tag_defs
apply (refine_vcg)
subgoal by (fastforce simp: scaleR2_def vimage_def image_def)
subgoal by fastforce
done
lemma partition_sets_spec[le, refine_vcg]:
shows "partition_sets ro XS \<le> SPEC (\<lambda>YS. (\<Union>(_, _, PS, _, _, _, _, _) \<in> XS. PS) \<subseteq> YS)"
unfolding partition_sets_def autoref_tag_defs
by (refine_vcg FORWEAK_mono_rule[where I="\<lambda>X Y. (\<Union>(_, _, PS, _, _, _, _, _) \<in> X. PS) \<subseteq> Y"]) auto
lemma
do_intersection_poincare_mapstos_trans:
assumes pm: "\<And>i. i \<in> I \<Longrightarrow> poincare_mapsto (p i) (X0 i) UNIV (CX i) (X1 i)"
assumes di: "do_intersection_spec UNIV guards ivl sctn (\<Union>i\<in>I. X1 i) (P, CP)"
assumes "\<And>i. i \<in> I \<Longrightarrow> fst ` (X1 i) \<subseteq> CP"
assumes "\<And>i. i \<in> I \<Longrightarrow> {x \<in> ivl. x \<in> plane_of sctn} \<inter> CX i = {}"
assumes "\<And>i. i \<in> I \<Longrightarrow> guards \<inter> (CX i \<union> CP) = {}"
assumes "\<And>i. i \<in> I \<Longrightarrow> X0 i \<subseteq> CX i \<times> UNIV"
assumes "\<And>i. i \<in> I \<Longrightarrow> closed (p i)"
assumes "closed ivl"
assumes "\<And>i. i \<in> I \<Longrightarrow> CX i \<subseteq> Csafe"
shows "do_intersection_spec UNIV guards ivl sctn (\<Union>i\<in>I. X0 i) (P, (\<Union>i\<in>I. CX i) \<union> CP)"
apply (auto simp: do_intersection_spec_def)
subgoal
apply (simp del: UN_simps add: UN_extend_simps)
apply (rule impI)
apply (thin_tac "I \<noteq> {}")
subgoal
proof -
from di have pmi: "poincare_mapsto {x \<in> ivl. x \<in> plane_of sctn} (X1 i) UNIV CP P" if "i \<in> I" for i
by (auto simp: do_intersection_spec_def intro: poincare_mapsto_subset that)
show ?thesis
apply (rule poincare_mapsto_UnionI)
apply (rule poincare_mapsto_trans[OF pm pmi])
apply clarsimp_all
subgoal s1 using assms by (auto simp: do_intersection_spec_def)
subgoal using assms apply (auto simp: do_intersection_spec_def)
apply blast
by (metis (mono_tags, lifting) s1 mem_Collect_eq mem_simps(2) mem_simps(4))
subgoal using assms by auto
subgoal using assms by auto
subgoal premises prems for i x d
proof -
note prems
have [intro, simp]: "closed {x \<in> ivl. x \<in> plane_of sctn} " "closed {x \<in> ivl. x \<bullet> normal sctn = pstn sctn}"
by (auto intro!: closed_levelset_within continuous_intros simp: plane_of_def assms)
have set_eq: "(CX i \<union> CP) \<times> UNIV = (fst ` X1 i \<times> UNIV \<union> CX i \<times> UNIV \<union> CP \<times> UNIV)"
using assms prems
by auto
have empty_inter: "{x \<in> ivl. x \<bullet> normal sctn - pstn sctn = 0} \<times> UNIV \<inter> (CX i \<union> CP) \<times> UNIV = {}"
apply safe
subgoal
using assms(4)[of i] \<open>i \<in> I\<close>
by (auto simp: plane_of_def )
subgoal
using assms(4)[of i]
using prems assms by (auto simp: plane_of_def do_intersection_spec_def)
done
have ft: "flowsto (X0 i) {0<..} ((CX i \<union> CP) \<times> UNIV) (fst ` P \<times> UNIV)"
unfolding set_eq
apply (rule flowsto_poincare_mapsto_trans_flowsto[OF poincare_mapsto_imp_flowsto[OF pm[OF \<open>i \<in> I\<close>]]
pmi[OF \<open>i \<in> I\<close>] _ _ order_refl])
using assms prems by (auto)
then have ret: "returns_to {x \<in> ivl. x \<bullet> normal sctn - pstn sctn = 0} x"
apply (rule returns_to_flowstoI[OF _ _ _ _ _ _ order_refl])
subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal by (rule empty_inter)
subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def)
done
have pm: "poincare_map {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} x \<in> fst ` P"
apply (rule poincare_map_mem_flowstoI[OF ft])
subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal using empty_inter by simp
subgoal by auto
subgoal by auto
subgoal using prems assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal by auto
done
from ret have "isCont (return_time {x \<in> ivl. x \<bullet> normal sctn - pstn sctn = 0}) x"
apply (rule return_time_isCont_outside)
subgoal by fact
apply (force intro!: derivative_eq_intros)
subgoal by (auto intro!: continuous_intros)
subgoal using prems pm assms by (auto simp: do_intersection_spec_def)
subgoal using prems pm assms
by (auto simp: eventually_at_filter plane_of_def do_intersection_spec_def)
subgoal
proof -
have "x \<in> CX i" using \<open>_ \<in> I \<Longrightarrow> X0 _ \<subseteq> CX _ \<times> UNIV\<close>[OF \<open>i \<in> I\<close>] \<open>(x, _) \<in> _\<close>
by auto
with assms(4)[OF \<open>i \<in> I\<close>] show ?thesis
by (auto simp: plane_of_def)
qed
done
then show "isCont (return_time {x \<in> ivl. x \<in> plane_of sctn}) x" by (simp add: plane_of_def)
qed
done
qed
done
subgoal using assms by (fastforce simp: plane_of_def do_intersection_spec_def)
subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal using assms by (fastforce simp: plane_of_def do_intersection_spec_def)
subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal using assms(9) by (fastforce simp: plane_of_def do_intersection_spec_def)
subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def)
subgoal using assms by (auto simp: plane_of_def do_intersection_spec_def)
done
lemma flow_in_stable_setD:
"flow0 x0 t \<in> stable_set trap \<Longrightarrow> t \<in> existence_ivl0 x0 \<Longrightarrow> x0 \<in> stable_set trap"
apply (auto simp: stable_set_def)
proof goal_cases
case (1 s)
then show ?case
apply (cases "s \<le> t")
apply (meson atLeastAtMost_iff contra_subsetD local.ivl_subset_existence_ivl)
using contra_subsetD local.existence_ivl_reverse local.existence_ivl_trans' local.flows_reverse by fastforce
next
case (2)
have "((\<lambda>s. flow0 x0 (t + s)) \<longlongrightarrow> trap) (at_top)"
proof (rule Lim_transform_eventually)
have "\<forall>\<^sub>F x in at_top. x > max t 0"
by (simp add: max_def)
then show "\<forall>\<^sub>F x in at_top. flow0 (flow0 x0 t) x = flow0 x0 (t + x)"
apply eventually_elim
apply (subst flow_trans)
using 2
by auto
qed (use 2 in auto)
then show ?case by (simp add: tendsto_at_top_translate_iff ac_simps)
qed
lemma
poincare_mapsto_avoid_trap:
assumes "poincare_mapsto p (X0 - trap \<times> UNIV) S CX P"
assumes "closed p"
assumes trapprop[THEN stable_onD]: "stable_on (CX \<union> fst ` P) trap"
shows "poincare_mapsto p (X0 - trap \<times> UNIV) S CX (P - trap \<times> UNIV)"
using assms(1,2)
apply (auto simp: poincare_mapsto_def)
apply (drule bspec, force)
apply auto
subgoal for x0 d0 D
apply (rule exI[where x=D])
apply (auto dest!: trapprop simp: poincare_map_def intro!: return_time_exivl assms(1,2) return_time_pos)
subgoal for s
by (cases "s = return_time p x0") (auto simp: )
done
done
lemma poincare_onto_series[le, refine_vcg]:
assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)"
assumes [refine_vcg]: "\<And>X0. X0 \<subseteq> Csafe \<times> UNIV \<Longrightarrow> symstart X0 \<le> SPEC (\<lambda>(CX, X). flowsto (X0 - trap \<times> UNIV) {0..} (CX \<times> UNIV) (X))"
assumes trapprop: "stable_on (Csafe - (ivl \<inter> plane_of sctn)) trap"
shows "poincare_onto_series symstart trap guards (X0::'a eucl1 set) ivl sctn ro \<le>
SPEC (\<lambda>XS. do_intersection_spec UNIV {} ivl sctn (X0 - trap \<times> UNIV)
(XS, Csafe - (ivl \<inter> plane_of sctn)) \<and>
fst ` X0 - trap \<subseteq> Csafe - (ivl \<inter> plane_of sctn))"
proof (induction guards arbitrary: X0)
case Nil
then show ?case
apply (simp add:)
apply refine_vcg
apply (clarsimp simp add: ivlsctn_to_set_def)
apply (rule do_intersection_spec_subset2, assumption)
subgoal by (auto simp: do_intersection_spec_def)
subgoal by (auto simp: do_intersection_spec_def)
subgoal by (auto simp: do_intersection_spec_def)
subgoal by (auto simp: do_intersection_spec_def)
subgoal by (auto simp: do_intersection_spec_def)
subgoal by (auto simp: do_intersection_spec_def)
done
next
case (Cons a guards)
note Cons.IH[simplified, le, refine_vcg]
show ?case
+ supply [[simproc del: defined_all]]
apply auto
apply refine_vcg
apply clarsimp_all
defer
subgoal premises prems for b c d e f g h
proof -
from prems have "(f, g) \<in> (\<Union>x\<in>c. h x)"
by auto
then obtain x where "x \<in> c" "(f, g) \<in> (h x)"
by auto
then show ?thesis
using prems(14)[rule_format, OF \<open>x \<in> c\<close>] prems(5-7)
by (cases x) (auto simp: do_intersection_spec_def)
qed
subgoal premises prems for c ro d e f
proof -
let ?s = "trap \<times> UNIV"
note prems
from \<open>do_intersection_spec _ _ _ _ _ _ \<close>
have disro: "do_intersection_spec UNIV {} ivl sctn ((\<Union>i\<in>ro. case i of (_, _, PS, _, _, _, _, _, _) \<Rightarrow> PS - ?s))
(e, Csafe - ivl \<inter> plane_of sctn)"
apply (rule do_intersection_spec_subset)
using prems by auto
have subset: "(Csafe - ivl \<inter> plane (normal sctn) (pstn sctn)) \<supseteq>
(snd (snd (snd (snd (snd (snd (snd (snd i))))))) \<union>
fst (snd (snd (snd (snd (snd (snd (snd i))))))) \<union> fst ` fst (snd (snd i)))" if "i \<in> ro" for i
using prems(12)[rule_format, unfolded do_intersection_spec_def, OF that]
apply (clarsimp )
subgoal for s X P1 P2 R ivla sctna CX CXS
apply (rule conjI)
subgoal by (auto simp: plane_of_def)
subgoal by (auto simp: plane_of_def)
done
done
have pmro: "poincare_mapsto
(case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \<Rightarrow> {x \<in> ivla. x \<in> plane_of sctna})
(f i - ?s) UNIV
(case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \<Rightarrow> CXS \<union> CX)
(case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \<Rightarrow> P1)"
if "i \<in> ro"
for i
using prems(12)[rule_format, unfolded do_intersection_spec_def, OF that]
by (auto intro: poincare_mapsto_subset)
then have pmro: "poincare_mapsto
(case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \<Rightarrow> {x \<in> ivla. x \<in> plane_of sctna})
(f i - ?s) UNIV
(case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \<Rightarrow> CXS \<union> CX)
(case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \<Rightarrow> P1 - ?s)"
if "i \<in> ro"
for i
unfolding split_beta'
apply (rule poincare_mapsto_avoid_trap)
using that prems assms
by (auto intro!: closed_levelset_within continuous_intros
stable_on_mono[OF _ subset]
simp: plane_of_def)
have "do_intersection_spec UNIV {} ivl sctn (\<Union>i\<in>ro. f i - ?s)
(e, (\<Union>i\<in>ro. case i of (s, X, P1, P2, R, ivla, sctna, CX, CXS) \<Rightarrow> CXS \<union> CX) \<union>
(Csafe - ivl \<inter> plane_of sctn))"
apply (rule do_intersection_poincare_mapstos_trans[OF pmro disro])
subgoal by auto
subgoal premises that for i
- using prems(12)[rule_format, unfolded do_intersection_spec_def, OF that]
+ using prems(12)[rule_format, unfolded do_intersection_spec_def, OF that] using [[simproc del: defined_all]]
by (auto simp: do_intersection_spec_def)
subgoal using assms(1,2) prems by (auto simp: do_intersection_spec_def)
subgoal by auto
subgoal premises that for i
using prems(12)[rule_format, unfolded do_intersection_spec_def, OF that]
prems(11) that
by (auto simp: do_intersection_spec_def)
subgoal using assms(1,2) prems by (auto simp: do_intersection_spec_def)
subgoal using assms(1,2) prems by (auto simp: do_intersection_spec_def)
subgoal using assms(1,2) prems by (auto simp: do_intersection_spec_def)
done
then show ?thesis
unfolding \<open>(\<Union>x\<in>ro. f x) = X0 - trap \<times> UNIV\<close>
apply (rule do_intersection_spec_subset2)
subgoal using assms(1,2) prems by (auto simp: do_intersection_spec_def)
using prems
by (auto simp: do_intersection_spec_def intro: poincare_mapsto_subset)
qed
done
qed
lemma
do_intersection_flowsto_trans_return:
assumes "flowsto XS0 {0<..} (CX \<times> UNIV) X1"
assumes "do_intersection_spec UNIV guards ivl sctn X1 (P, CP)"
assumes "fst ` X1 \<subseteq> CP"
assumes "{x \<in> ivl. x \<in> plane_of sctn} \<inter> CX = {}"
assumes "guards \<inter> (CX \<union> CP) = {}"
assumes "closed ivl"
assumes "CX \<subseteq> sbelow_halfspace sctn \<inter> Csafe"
assumes subset_plane: "fst ` XS0 \<subseteq> plane_of sctn \<inter> ivl"
assumes down: "\<And>x d. (x, d) \<in> XS0 \<Longrightarrow> ode x \<bullet> normal sctn < 0" "\<And>x. x \<in> CX \<Longrightarrow> ode x \<bullet> normal sctn < 0"
shows "do_intersection_spec (below_halfspace sctn) guards ivl sctn XS0 (P, CX \<union> CP)"
using assms
apply (auto simp: do_intersection_spec_def)
subgoal
apply (rule flowsto_poincare_trans, assumption, assumption)
subgoal by simp
subgoal by auto
subgoal using assms(3) by auto
subgoal by (auto intro!: closed_levelset_within continuous_intros simp: plane_of_def)
prefer 2
subgoal by (auto simp add: plane_of_def halfspace_simps)
subgoal premises prems for x d
proof -
have [intro, simp]: "closed {x \<in> ivl. x \<in> plane_of sctn} " "closed {x \<in> ivl. x \<bullet> normal sctn = pstn sctn}"
by (auto intro!: closed_levelset_within continuous_intros simp: plane_of_def assms)
from subset_plane have "fst ` XS0 \<subseteq> below_halfspace sctn" by (auto simp: )
from flowsto_stays_sbelow[OF \<open>flowsto _ _ _ _\<close> this down(2)]
have ft_below: "flowsto XS0 pos_reals (CX \<times> UNIV \<inter> sbelow_halfspace sctn \<times> UNIV) X1"
by auto
from flowsto_poincare_mapsto_trans_flowsto[OF ft_below \<open>poincare_mapsto _ _ _ _ _\<close> _ _ order_refl]
have ft: "flowsto XS0 {0<..} (X1 \<union> CX \<times> UNIV \<inter> sbelow_halfspace sctn \<times> UNIV \<union> CP \<times> UNIV) (fst ` P \<times> UNIV)"
by (auto simp: )
have ret: "returns_to {x \<in> ivl. x \<bullet> normal sctn - pstn sctn = 0} x"
apply (rule returns_to_flowstoI[OF ft])
using prems by (auto simp: plane_of_def halfspace_simps)
have pm: "poincare_map {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} x \<in> fst ` P"
apply (rule poincare_map_mem_flowstoI[OF ft])
using prems by (auto simp: plane_of_def halfspace_simps)
from pm prems have evmem: "\<forall>\<^sub>F x in at (poincare_map {x \<in> ivl. x \<bullet> normal sctn = pstn sctn} x) within
plane_of sctn. x \<in> ivl"
by auto
from ret have "continuous (at x within {x. x \<bullet> normal sctn - pstn sctn \<le> 0})
(return_time {x \<in> ivl. x \<bullet> normal sctn - pstn sctn = 0})"
apply (rule return_time_continuous_below)
apply (rule derivative_eq_intros refl)+
apply force
subgoal using \<open>closed ivl\<close> by auto
subgoal using prems pm by (auto simp: plane_of_def eventually_at_filter)
subgoal by (auto intro!: )
subgoal using prems pm by auto
subgoal using prems by auto
subgoal using prems pm by (auto intro!: assms simp: plane_of_def)
subgoal using prems pm by auto
done
then show "continuous (at x within below_halfspace sctn) (return_time {x \<in> ivl. x \<in> plane_of sctn})"
by (simp add: plane_of_def halfspace_simps)
qed
done
done
lemma do_intersection_spec_sctn_cong:
assumes "sctn = sctn' \<or> (normal sctn = - normal sctn' \<and> pstn sctn = - pstn sctn')"
shows "do_intersection_spec a b c sctn d e = do_intersection_spec a b c sctn' d e"
using assms
by (auto simp: do_intersection_spec_def plane_of_def set_eq_iff intro!: )
lemma poincare_onto_from[le, refine_vcg]:
assumes wd[refine_vcg]: "wd TYPE('a::enum rvec)"
assumes [refine_vcg]: "\<And>X0. X0 \<subseteq> Csafe \<times> UNIV \<Longrightarrow> symstart X0 \<le> SPEC (\<lambda>(CX, X). flowsto (X0 - trap \<times> UNIV) {0..} (CX \<times> UNIV) (X))"
assumes trapprop: "stable_on (Csafe - (ivl \<inter> plane_of sctn)) trap"
shows "poincare_onto_from symstart trap S guards ivl sctn ro (XS0::'a eucl1 set) \<le>
SPEC (poincare_mapsto (ivl \<inter> plane_of sctn) (XS0 - trap \<times> UNIV) S (Csafe - ivl \<inter> plane_of sctn))"
unfolding poincare_onto_from_def autoref_tag_defs
apply (refine_vcg, clarsimp_all simp: trapprop)
subgoal by (auto simp: do_intersection_spec_def Int_def intro: poincare_mapsto_subset)
subgoal premises prems for a b c d e f
proof -
note prems
from trapprop
have stable: "stable_on (fst ` (e \<times> UNIV \<inter> sbelow_halfspace a \<times> UNIV \<union> d)) trap"
apply (rule stable_on_mono)
using \<open>fst ` (d \<union> e \<times> UNIV) \<subseteq> Csafe\<close> \<open>a = sctn \<or> normal a = - normal sctn \<and> pstn a = - pstn sctn\<close>
\<open>fst ` d \<subseteq> sbelow_halfspace a\<close>
by (auto simp: halfspace_simps plane_of_def image_Un)
from prems(16) have "flowsto (XS0 - trap \<times> UNIV) {0<..} (e \<times> UNIV \<inter> sbelow_halfspace a \<times> UNIV) d"
by (rule flowsto_subset) auto
then have ft: "flowsto (XS0 - trap \<times> UNIV) {0<..} ((e \<inter> sbelow_halfspace a) \<times> UNIV) (d - trap \<times> UNIV)"
by (auto intro!: flowsto_mapsto_avoid_trap stable simp: Times_Int_distrib1)
from prems(8) have di: "do_intersection_spec UNIV {} ivl a (d - trap \<times> UNIV) (f, Csafe - ivl \<inter> plane_of sctn)"
apply (subst do_intersection_spec_sctn_cong)
defer apply assumption
using prems(2)
by auto
have "do_intersection_spec (below_halfspace a) {} ivl a (XS0 - trap \<times> UNIV)
(f, e \<inter> sbelow_halfspace a \<union> (Csafe - ivl \<inter> plane_of sctn))"
apply (rule do_intersection_flowsto_trans_return[OF ft di])
subgoal using prems by (auto simp: do_intersection_spec_def halfspace_simps plane_of_def)
subgoal by (auto simp: halfspace_simps plane_of_def)
subgoal using prems by (auto simp: halfspace_simps plane_of_def)
subgoal using prems by (auto simp: do_intersection_spec_def halfspace_simps plane_of_def)
subgoal using prems by (auto simp: image_Un)
subgoal using prems by (auto simp: do_intersection_spec_def halfspace_simps plane_of_def)
subgoal using prems by (auto simp: do_intersection_spec_def halfspace_simps plane_of_def)
subgoal using prems by (auto simp: do_intersection_spec_def halfspace_simps plane_of_def)
done
moreover have "plane_of a = plane_of sctn"
using prems(2) by (auto simp: plane_of_def)
ultimately show ?thesis
apply (auto simp add: do_intersection_spec_def Int_def)
apply (rule poincare_mapsto_subset, assumption)
by auto
qed
done
lemma subset_spec1[refine_vcg]: "subset_spec1 R P dP \<le> SPEC (\<lambda>b. b \<longrightarrow> R \<subseteq> flow1_of_vec1 ` (P \<times> dP))"
unfolding subset_spec1_def
by refine_vcg (auto simp: vec1_of_flow1_def)
lemma subset_spec1_coll[le, refine_vcg]:
"subset_spec1_coll R P dP \<le> subset_spec R (flow1_of_vec1 ` (P \<times> dP))"
unfolding autoref_tag_defs subset_spec_def subset_spec1_coll_def
by (refine_vcg) (auto simp: subset_iff set_of_ivl_def)
lemma one_step_until_time_spec[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "one_step_until_time (X0::'n eucl1 set) CX t1 \<le> SPEC (\<lambda>(R, CX).
(\<forall>(x0, d0) \<in> X0. t1 \<in> existence_ivl0 x0 \<and>
(flow0 x0 t1, Dflow x0 t1 o\<^sub>L d0) \<in> R \<and>
(\<forall>t \<in> {0 .. t1}. flow0 x0 t \<in> CX)) \<and>
fst ` R \<union> CX \<subseteq> Csafe)"
+ using [[simproc del: defined_all]]
unfolding one_step_until_time_def autoref_tag_defs
apply (refine_vcg WHILE_rule[where I="\<lambda>(t, h, X, CX). fst ` X \<subseteq> Csafe \<and> CX \<subseteq> Csafe \<and> 0 \<le> h \<and> 0 \<le> t \<and> t \<le> t1 \<and>
(\<forall>(x0, d0) \<in> X0. t \<in> existence_ivl0 x0 \<and>
(flow0 x0 t, Dflow x0 t o\<^sub>L d0) \<in> X \<and>
(\<forall>s \<in> {0 .. t}. flow0 x0 s \<in> CX))"])
subgoal by auto
subgoal by (force simp: flowpipe_def existence_ivl_trans flow_trans)
subgoal by (auto simp: flowpipe_def existence_ivl_trans flow_trans)
apply clarsimp subgoal for startstep rk2_param a b c d e f g h i j
apply (safe)
subgoal by (auto simp: flowpipe_def intro!: existence_ivl_trans flow_trans)
subgoal
apply (subst flow_trans, force)
subgoal by (auto simp: flowpipe_def intro!: existence_ivl_trans flow_trans)
apply (subst Dflow_trans, force)
subgoal by (auto simp: flowpipe_def intro!: existence_ivl_trans flow_trans)
by (auto simp: blinfun_compose_assoc flowpipe_def)
subgoal for s
apply (drule bspec[where x="(i, j)"], assumption)
apply auto
apply (cases "s \<le> a")
subgoal by auto
subgoal
apply (auto simp: blinfun_compose_assoc flowpipe_def)
apply (drule bspec, assumption)
apply auto
proof goal_cases
case 1
have a: "a \<in> existence_ivl0 i" using 1 by auto
have sa: "s - a \<in> existence_ivl0 (flow0 i a)"
using "1"(15) "1"(19) "1"(20) local.ivl_subset_existence_ivl by fastforce
have "flow0 i s = flow0 (flow0 i a) (s - a)"
by (auto simp: a sa flow_trans[symmetric])
also have "\<dots> \<in> f"
using 1 by auto
finally show ?case
using 1 by simp
qed
done
done
subgoal by auto
done
text \<open>solve ODE until the time interval \<open>{t1 .. t2}\<close>\<close>
lemma ivl_of_eucl1_coll[THEN order_trans, refine_vcg]: "ivl_of_eucl_coll X \<le> SPEC (\<lambda>R. X \<times> UNIV \<subseteq> R)"
unfolding ivl_of_eucl_coll_def
by refine_vcg auto
lemma one_step_until_time_ivl_spec[le, refine_vcg]:
assumes wd[refine_vcg]: "wd (TYPE('n::enum rvec))"
shows "one_step_until_time_ivl (X0::'n eucl1 set) CX t1 t2 \<le> SPEC (\<lambda>(R, CX).
(\<forall>(x0, d0) \<in> X0. {t1 .. t2} \<subseteq> existence_ivl0 x0 \<and>
(\<forall>t \<in> {t1 .. t2}. (flow0 x0 t, Dflow x0 t o\<^sub>L d0) \<in> R) \<and>
(\<forall>t \<in> {0 .. t1}. (flow0 x0 t) \<in> CX)) \<and>
fst ` R \<union> CX \<subseteq> Csafe)"
unfolding one_step_until_time_ivl_def
apply (refine_vcg, clarsimp_all)
subgoal for X CX Y CY CY' x0 d0
apply (drule bspec, assumption, clarsimp)
apply (drule bspec, assumption, clarsimp simp add: nonneg_interval_mem_existence_ivlI)
apply (rule subsetD, assumption)
subgoal for t
apply (drule bspec[where x=0], force)
apply (drule bspec[where x="t - t1"], force)
using interval_subset_existence_ivl[of t1 x0 t2]
by (auto simp: flow_trans')
done
- subgoal
- by (auto simp: scaleR2_def image_def vimage_def)
done
lemma empty_symstart_flowsto:
"X0 \<subseteq> Csafe \<times> UNIV \<Longrightarrow>
RETURN ({}, X0) \<le> SPEC (\<lambda>(CX, X). flowsto (X0 - {} \<times> UNIV) {0..} (CX \<times> UNIV) X)"
by (auto intro!: flowsto_self)
subsection \<open>Poincare map returning to\<close>
lemma poincare_onto_from_ivla[le, refine_vcg]:
assumes [refine_vcg]: "wd TYPE('n::enum rvec)"
assumes [refine_vcg]: "\<And>X0. X0 \<subseteq> Csafe \<times> UNIV \<Longrightarrow> symstart X0 \<le> SPEC (\<lambda>(CX, X). flowsto (X0 - trap \<times> UNIV) {0..} (CX \<times> UNIV) (X))"
assumes trapprop[refine_vcg]: "stable_on (Csafe - (ivl \<inter> plane_of sctn)) trap"
shows "poincare_onto_from symstart trap S guards ivl sctn ro (XS0::'n eucl1 set) \<le> SPEC
(\<lambda>P.
wd TYPE((real, 'n) vec) \<and>
poincare_mapsto (ivl \<inter> plane_of sctn) (XS0 - trap \<times> UNIV) S (Csafe - ivl \<inter> plane_of sctn) P)"
by (refine_vcg)
subsection \<open>Poincare map onto (from outside of target)\<close>
subsection \<open>One step method (reachability in time)\<close>
lemma c0_info_of_apprsI:
assumes "(b, a) \<in> clw_rel appr_rel"
assumes "x \<in> a"
shows "x \<in> c0_info_of_apprs b"
using assms
by (auto simp: appr_rel_br clw_rel_br c0_info_of_apprs_def c0_info_of_appr_def dest!: brD)
lemma c0_info_of_appr'I:
assumes "(b, a) \<in> \<langle>clw_rel appr_rel\<rangle>phantom_rel"
assumes "x \<in> a"
shows "x \<in> c0_info_of_appr' b"
using assms
by (auto simp add: c0_info_of_appr'_def intro!: c0_info_of_apprsI split: option.splits)
lemma poincare_onto_from_in_ivl[le, refine_vcg]:
assumes [refine_vcg]: "wd TYPE('n::enum rvec)"
assumes [refine_vcg]: "\<And>X0. X0 \<subseteq> Csafe \<times> UNIV \<Longrightarrow> symstart X0 \<le> SPEC (\<lambda>(CX, X). flowsto (X0 - trap \<times> UNIV) {0..} (CX \<times> UNIV) (X))"
assumes trapprop: "stable_on (Csafe - (ivl \<inter> plane_of sctn)) trap"
shows "poincare_onto_from_in_ivl symstart trap S guards ivl sctn ro (XS0::'n::enum eucl1 set) P dP \<le>
SPEC (\<lambda>b. b \<longrightarrow> poincare_mapsto (ivl \<inter> plane_of sctn) (XS0 - trap \<times> UNIV) S (Csafe - ivl \<inter> plane_of sctn) (flow1_of_vec1 ` (P \<times> dP)))"
unfolding poincare_onto_from_in_ivl_def
apply (refine_vcg, clarsimp_all)
apply (rule trapprop)
apply (rule poincare_mapsto_subset)
apply assumption
by (auto simp: )
lemma lvivl_default_relI:
"(dRi, set_of_lvivl' dRi::'e::executable_euclidean_space set) \<in> \<langle>lvivl_rel\<rangle>default_rel UNIV"
if "lvivl'_invar DIM('e) dRi"
using that
by (auto simp: set_of_lvivl'_def set_of_lvivl_def set_of_ivl_def lvivl'_invar_def
intro!: mem_default_relI lvivl_relI)
lemma stable_on_empty[simp]: "stable_on A {}"
by (auto simp: stable_on_def)
lemma poincare_onto_in_ivl[le, refine_vcg]:
assumes [simp]: "length (ode_e) = CARD('n::enum)"
shows "poincare_onto_in_ivl guards ivl sctn ro (XS0::'n::enum eucl1 set) P dP \<le>
SPEC (\<lambda>b. b \<longrightarrow> poincare_mapsto (ivl \<inter> plane_of sctn) (XS0) UNIV (Csafe - ivl \<inter> plane_of sctn) (flow1_of_vec1 ` (P \<times> dP)))"
proof -
have wd[refine_vcg]: "wd TYPE((real, 'n) vec)" by (simp add: wd_def)
show ?thesis
unfolding poincare_onto_in_ivl_def
apply (refine_vcg)
subgoal by (auto intro!: flowsto_self)
subgoal
apply (clarsimp simp add: do_intersection_spec_def Int_def[symmetric])
apply (rule poincare_mapsto_subset)
apply assumption
by auto
done
qed
end
end
\ No newline at end of file
diff --git a/thys/Password_Authentication_Protocol/Protocol.thy b/thys/Password_Authentication_Protocol/Protocol.thy
--- a/thys/Password_Authentication_Protocol/Protocol.thy
+++ b/thys/Password_Authentication_Protocol/Protocol.thy
@@ -1,5546 +1,5551 @@
(* Title: Verification of a Diffie-Hellman Password-based Authentication Protocol by Extending the Inductive Method
Author: Pasquale Noce
Security Certification Specialist at Arjo Systems, Italy
pasquale dot noce dot lavoro at gmail dot com
pasquale dot noce at arjosystems dot com
*)
section "Protocol modeling and verification"
theory Protocol
imports Propaedeutics
begin
subsection "Protocol modeling"
text \<open>
The protocol under consideration can be formalized by means of the following inductive definition.
\null
\<close>
inductive_set protocol :: "(event list \<times> state \<times> msg set \<times> msg set) set" where
Nil: "([], start_S, start_A, start_U) \<in> protocol" |
R1: "\<lbrakk>(evsR1, S, A, U) \<in> protocol; Pri_AgrK s \<notin> U; s \<noteq> 0;
NonceS (S (Card n, n, run)) = None\<rbrakk>
\<Longrightarrow> (Says n run 1 (Card n) (User m) (Crypt (symK n) (Pri_AgrK s)) # evsR1,
S ((Card n, n, run) := S (Card n, n, run) \<lparr>NonceS := Some s\<rparr>),
if n \<in> bad then insert (Pri_AgrK s) A else A,
insert (Pri_AgrK s) U) \<in> protocol" |
FR1: "\<lbrakk>(evsFR1, S, A, U) \<in> protocol; User m \<noteq> Spy; s \<noteq> 0;
Crypt (symK m) (Pri_AgrK s) \<in> synth (analz (A \<union> spies evsFR1))\<rbrakk>
\<Longrightarrow> (Says n run 1 Spy (User m) (Crypt (symK m) (Pri_AgrK s)) # evsFR1,
S, A, U) \<in> protocol" |
C2: "\<lbrakk>(evsC2, S, A, U) \<in> protocol; Pri_AgrK a \<notin> U;
NonceS (S (User m, n, run)) = None;
Says n run 1 X (User m) (Crypt (symK n') (Pri_AgrK s)) \<in> set evsC2;
s' = (if symK n' = symK m then s else 0)\<rbrakk>
\<Longrightarrow> (Says n run 2 (User m) (Card n) (pubAK a) # evsC2,
S ((User m, n, run) := S (User m, n, run)
\<lparr>NonceS := Some s', IntMapK := Some a\<rparr>),
if User m = Spy then insert (Pri_AgrK a) A else A,
insert (Pri_AgrK a) U) \<in> protocol" |
FC2: "\<lbrakk>(evsFC2, S, A, U) \<in> protocol;
Pri_AgrK a \<in> analz (A \<union> spies evsFC2)\<rbrakk>
\<Longrightarrow> (Says n run 2 Spy (Card n) (pubAK a) # evsFC2,
S, A, U) \<in> protocol" |
R2: "\<lbrakk>(evsR2, S, A, U) \<in> protocol; Pri_AgrK b \<notin> U;
NonceS (S (Card n, n, run)) \<noteq> None;
IntMapK (S (Card n, n, run)) = None;
Says n run 2 X (Card n) (pubAK a) \<in> set evsR2\<rbrakk>
\<Longrightarrow> (Says n run 2 (Card n) X (pubAK b) # evsR2,
S ((Card n, n, run) := S (Card n, n, run)
\<lparr>IntMapK := Some b, ExtMapK := Some a\<rparr>),
A, insert (Pri_AgrK b) U) \<in> protocol" |
FR2: "\<lbrakk>(evsFR2, S, A, U) \<in> protocol; User m \<noteq> Spy;
Pri_AgrK b \<in> analz (A \<union> spies evsFR2)\<rbrakk>
\<Longrightarrow> (Says n run 2 Spy (User m) (pubAK b) # evsFR2,
S, A, U) \<in> protocol" |
C3: "\<lbrakk>(evsC3, S, A, U) \<in> protocol; Pri_AgrK c \<notin> U;
NonceS (S (User m, n, run)) = Some s;
IntMapK (S (User m, n, run)) = Some a;
ExtMapK (S (User m, n, run)) = None;
Says n run 2 X (User m) (pubAK b) \<in> set evsC3;
c * (s + a * b) \<noteq> 0\<rbrakk>
\<Longrightarrow> (Says n run 3 (User m) (Card n) (pubAK (c * (s + a * b))) # evsC3,
S ((User m, n, run) := S (User m, n, run)
\<lparr>ExtMapK := Some b, IntAgrK := Some c\<rparr>),
if User m = Spy then insert (Pri_AgrK c) A else A,
insert (Pri_AgrK c) U) \<in> protocol" |
FC3: "\<lbrakk>(evsFC3, S, A, U) \<in> protocol;
NonceS (S (Card n, n, run)) = Some s;
IntMapK (S (Card n, n, run)) = Some b;
ExtMapK (S (Card n, n, run)) = Some a;
{Pri_AgrK s, Pri_AgrK a, Pri_AgrK c} \<subseteq> analz (A \<union> spies evsFC3)\<rbrakk>
\<Longrightarrow> (Says n run 3 Spy (Card n) (pubAK (c * (s + a * b))) # evsFC3,
S, A, U) \<in> protocol" |
R3: "\<lbrakk>(evsR3, S, A, U) \<in> protocol; Pri_AgrK d \<notin> U;
NonceS (S (Card n, n, run)) = Some s;
IntMapK (S (Card n, n, run)) = Some b;
ExtMapK (S (Card n, n, run)) = Some a;
IntAgrK (S (Card n, n, run)) = None;
Says n run 3 X (Card n) (pubAK (c * (s' + a * b))) \<in> set evsR3;
Key (sesK (c * d * (s' + a * b))) \<notin> U;
Key (sesK (c * d * (s + a * b))) \<notin> U;
d * (s + a * b) \<noteq> 0\<rbrakk>
\<Longrightarrow> (Says n run 3 (Card n) X (pubAK (d * (s + a * b))) # evsR3,
S ((Card n, n, run) := S (Card n, n, run)
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s' + a * b))\<rparr>),
if s' = s \<and> Pri_AgrK c \<in> analz (A \<union> spies evsR3)
then insert (Key (sesK (c * d * (s + a * b)))) A else A,
{Pri_AgrK d,
Key (sesK (c * d * (s' + a * b))), Key (sesK (c * d * (s + a * b))),
\<lbrace>Key (sesK (c * d * (s + a * b))), Agent X, Number n, Number run\<rbrace>} \<union> U) \<in> protocol" |
FR3: "\<lbrakk>(evsFR3, S, A, U) \<in> protocol; User m \<noteq> Spy;
NonceS (S (User m, n, run)) = Some s;
IntMapK (S (User m, n, run)) = Some a;
ExtMapK (S (User m, n, run)) = Some b;
IntAgrK (S (User m, n, run)) = Some c;
{Pri_AgrK s, Pri_AgrK b, Pri_AgrK d} \<subseteq> analz (A \<union> spies evsFR3);
Key (sesK (c * d * (s + a * b))) \<notin> U\<rbrakk>
\<Longrightarrow> (Says n run 3 Spy (User m) (pubAK (d * (s + a * b))) # evsFR3, S,
insert (Key (sesK (c * d * (s + a * b)))) A,
{Key (sesK (c * d * (s + a * b))),
\<lbrace>Key (sesK (c * d * (s + a * b))), Agent (User m), Number n, Number run\<rbrace>} \<union> U) \<in> protocol" |
C4: "\<lbrakk>(evsC4, S, A, U) \<in> protocol;
IntAgrK (S (User m, n, run)) = Some c;
ExtAgrK (S (User m, n, run)) = None;
Says n run 3 X (User m) (pubAK f) \<in> set evsC4;
\<lbrace>Key (sesK (c * f)), Agent (User m), Number n, Number run\<rbrace> \<in> U\<rbrakk>
\<Longrightarrow> (Says n run 4 (User m) (Card n) (Crypt (sesK (c * f)) (pubAK f)) # evsC4,
S ((User m, n, run) := S (User m, n, run) \<lparr>ExtAgrK := Some f\<rparr>),
A, U) \<in> protocol" |
FC4: "\<lbrakk>(evsFC4, S, A, U) \<in> protocol;
NonceS (S (Card n, n, run)) = Some s;
IntMapK (S (Card n, n, run)) = Some b;
ExtMapK (S (Card n, n, run)) = Some a;
IntAgrK (S (Card n, n, run)) = Some d;
ExtAgrK (S (Card n, n, run)) = Some e;
Crypt (sesK (d * e)) (pubAK (d * (s + a * b)))
\<in> synth (analz (A \<union> spies evsFC4))\<rbrakk>
\<Longrightarrow> (Says n run 4 Spy (Card n)
(Crypt (sesK (d * e)) (pubAK (d * (s + a * b)))) # evsFC4,
S, A, U) \<in> protocol" |
R4: "\<lbrakk>(evsR4, S, A, U) \<in> protocol;
NonceS (S (Card n, n, run)) = Some s;
IntMapK (S (Card n, n, run)) = Some b;
ExtMapK (S (Card n, n, run)) = Some a;
IntAgrK (S (Card n, n, run)) = Some d;
ExtAgrK (S (Card n, n, run)) = Some e;
Says n run 4 X (Card n) (Crypt (sesK (d * e))
(pubAK (d * (s + a * b)))) \<in> set evsR4\<rbrakk>
\<Longrightarrow> (Says n run 4 (Card n) X (Crypt (sesK (d * e))
\<lbrace>pubAK e, Auth_Data (priAK n) b, pubAK (priAK n),
Crypt (priSK CA) (Hash (pubAK (priAK n)))\<rbrace>) # evsR4,
S, A, U) \<in> protocol" |
FR4: "\<lbrakk>(evsFR4, S, A, U) \<in> protocol; User m \<noteq> Spy;
NonceS (S (User m, n, run)) = Some s;
IntMapK (S (User m, n, run)) = Some a;
ExtMapK (S (User m, n, run)) = Some b;
IntAgrK (S (User m, n, run)) = Some c;
ExtAgrK (S (User m, n, run)) = Some f;
Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace> \<in> synth (analz (A \<union> spies evsFR4))\<rbrakk>
\<Longrightarrow> (Says n run 4 Spy (User m) (Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>) # evsFR4,
S, A, U) \<in> protocol" |
C5: "\<lbrakk>(evsC5, S, A, U) \<in> protocol;
NonceS (S (User m, n, run)) = Some s;
IntMapK (S (User m, n, run)) = Some a;
ExtMapK (S (User m, n, run)) = Some b;
IntAgrK (S (User m, n, run)) = Some c;
ExtAgrK (S (User m, n, run)) = Some f;
Says n run 4 X (User m) (Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>) \<in> set evsC5\<rbrakk>
\<Longrightarrow> (Says n run 5 (User m) (Card n) (Crypt (sesK (c * f)) (Passwd m)) # evsC5,
S, A, U) \<in> protocol" |
FC5: "\<lbrakk>(evsFC5, S, A, U) \<in> protocol;
IntAgrK (S (Card n, n, run)) = Some d;
ExtAgrK (S (Card n, n, run)) = Some e;
Crypt (sesK (d * e)) (Passwd n) \<in> synth (analz (A \<union> spies evsFC5))\<rbrakk>
\<Longrightarrow> (Says n run 5 Spy (Card n) (Crypt (sesK (d * e)) (Passwd n)) # evsFC5,
S, A, U) \<in> protocol" |
R5: "\<lbrakk>(evsR5, S, A, U) \<in> protocol;
IntAgrK (S (Card n, n, run)) = Some d;
ExtAgrK (S (Card n, n, run)) = Some e;
Says n run 5 X (Card n) (Crypt (sesK (d * e)) (Passwd n)) \<in> set evsR5\<rbrakk>
\<Longrightarrow> (Says n run 5 (Card n) X (Crypt (sesK (d * e)) (Number 0)) # evsR5,
S, A, U) \<in> protocol" |
FR5: "\<lbrakk>(evsFR5, S, A, U) \<in> protocol; User m \<noteq> Spy;
IntAgrK (S (User m, n, run)) = Some c;
ExtAgrK (S (User m, n, run)) = Some f;
Crypt (sesK (c * f)) (Number 0) \<in> synth (analz (A \<union> spies evsFR5))\<rbrakk>
\<Longrightarrow> (Says n run 5 Spy (User m) (Crypt (sesK (c * f)) (Number 0)) # evsFR5,
S, A, U) \<in> protocol"
text \<open>
\null
Here below are some comments about the most significant points of this definition.
\begin{itemize}
\item
Rules @{thm [source] R1} and @{thm [source] FR1} constrain the values of nonce $s$ to be different
from 0. In this way, the state of affairs where an incorrect PACE authentication key has been used
to encrypt nonce $s$, so that a wrong value results from the decryption, can be modeled in rule
@{thm [source] C2} by identifying such value with 0.
\item
The spy can disclose session keys as soon as they are established, namely in correspondence with
rules @{thm [source] R3} and @{thm [source] FR3}.
\\In the former rule, condition @{term "s' = s"} identifies Diffie-Hellman private key \<open>c\<close> as
the terminal's ephemeral private key for key agreement, and then $[c \times d \times (s + a \times
b)]G$ as the terminal's value of the shared secret, which the spy can compute by multiplying the
card's ephemeral public key $[d \times (s + a \times b)]G$ by \<open>c\<close> provided she knows
\<open>c\<close>.
\\In the latter rule, the spy is required to know private keys \<open>s\<close>, \<open>b\<close>, and \<open>d\<close>
to be able to compute and send public key $[d \times (s + a \times b)]G$. This is the only way to
share with @{term "User m"} the same shared secret's value $[c \times d \times (s + a \times b)]G$,
which the spy can compute by multiplying the user's ephemeral public key $[c \times (s + a \times
b)]G$ by \<open>d\<close>.
\item
Rules @{thm [source] R3} and @{thm [source] FR3} record the user, the communication channel, and the
protocol run associated to the session key having been established by adding this information to the
set of the messages already used. In this way, rule @{thm [source] C4} can specify that the session
key computed by @{term "User m"} is fresh by assuming that a corresponding record be included in set
\<open>U\<close>. In fact, a simple check that the session key be not included in \<open>U\<close> would vacuously
fail, as session keys are added to the set of the messages already used in rules @{thm [source] R3}
and @{thm [source] FR3}.
\end{itemize}
\<close>
subsection "Secrecy theorems"
text \<open>
This section contains a series of lemmas culminating in the secrecy theorems \<open>pr_key_secrecy\<close>
and \<open>pr_passwd_secrecy\<close>. Structured Isar proofs are used, possibly preceded by
\emph{apply}-style scripts in case a substantial amount of backward reasoning steps is required at
the beginning.
\null
\<close>
lemma pr_state:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
(NonceS (S (X, n, run)) = None \<longrightarrow> IntMapK (S (X, n, run)) = None) \<and>
(IntMapK (S (X, n, run)) = None \<longrightarrow> ExtMapK (S (X, n, run)) = None) \<and>
(ExtMapK (S (X, n, run)) = None \<longrightarrow> IntAgrK (S (X, n, run)) = None) \<and>
(IntAgrK (S (X, n, run)) = None \<longrightarrow> ExtAgrK (S (X, n, run)) = None)"
proof (erule protocol.induct, simp_all)
qed (rule_tac [!] impI, simp_all)
lemma pr_state_1:
"\<lbrakk>(evs, S, A, U) \<in> protocol; NonceS (S (X, n, run)) = None\<rbrakk> \<Longrightarrow>
IntMapK (S (X, n, run)) = None"
by (simp add: pr_state)
lemma pr_state_2:
"\<lbrakk>(evs, S, A, U) \<in> protocol; IntMapK (S (X, n, run)) = None\<rbrakk> \<Longrightarrow>
ExtMapK (S (X, n, run)) = None"
by (simp add: pr_state)
lemma pr_state_3:
"\<lbrakk>(evs, S, A, U) \<in> protocol; ExtMapK (S (X, n, run)) = None\<rbrakk> \<Longrightarrow>
IntAgrK (S (X, n, run)) = None"
by (simp add: pr_state)
lemma pr_state_4:
"\<lbrakk>(evs, S, A, U) \<in> protocol; IntAgrK (S (X, n, run)) = None\<rbrakk> \<Longrightarrow>
ExtAgrK (S (X, n, run)) = None"
by (simp add: pr_state)
lemma pr_analz_used:
"(evs, S, A, U) \<in> protocol \<Longrightarrow> A \<subseteq> U"
by (erule protocol.induct, auto)
lemma pr_key_parts_intro [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Key K \<in> parts (A \<union> spies evs) \<longrightarrow>
Key K \<in> A"
proof (erule protocol.induct, subst parts_simp, simp, blast, simp)
qed (simp_all add: parts_simp_insert parts_auth_data parts_crypt parts_mpair)
lemma pr_key_analz:
"(evs, S, A, U) \<in> protocol \<Longrightarrow> (Key K \<in> analz (A \<union> spies evs)) = (Key K \<in> A)"
proof (rule iffI, drule subsetD [OF analz_parts_subset], erule pr_key_parts_intro,
assumption)
qed (rule subsetD [OF analz_subset], simp)
lemma pr_symk_used:
"(evs, S, A, U) \<in> protocol \<Longrightarrow> Key (symK n) \<in> U"
by (erule protocol.induct, simp_all)
lemma pr_symk_analz:
"(evs, S, A, U) \<in> protocol \<Longrightarrow> (Key (symK n) \<in> analz (A \<union> spies evs)) = (n \<in> bad)"
proof (simp add: pr_key_analz, erule protocol.induct, simp_all, rule_tac [2] impI,
rule_tac [!] iffI, simp_all, erule disjE, erule_tac [2-4] disjE, simp_all)
assume "Key (symK n) \<in> Key ` symK ` bad"
hence "\<exists>m \<in> bad. symK n = symK m"
by (simp add: image_iff)
then obtain m where "m \<in> bad" and "symK n = symK m" ..
thus "n \<in> bad"
by (rule symK_bad)
next
assume "Key (symK n) \<in> Key ` range pubSK"
thus "n \<in> bad"
by (auto, drule_tac sym, erule_tac notE [OF pubSK_symK])
next
assume "Key (symK n) \<in> pubAK ` range priAK"
thus "n \<in> bad"
by blast
next
fix evsR3 S A U s a b c d
assume "(evsR3, S, A, U) \<in> protocol"
hence "Key (symK n) \<in> U"
by (rule pr_symk_used)
moreover assume "symK n = sesK (c * d * (s + a * b))"
ultimately have "Key (sesK (c * d * (s + a * b))) \<in> U"
by simp
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show "n \<in> bad"
by contradiction
next
fix evsFR3 S A U s a b c d
assume "(evsFR3, S, A, U) \<in> protocol"
hence "Key (symK n) \<in> U"
by (rule pr_symk_used)
moreover assume "symK n = sesK (c * d * (s + a * b))"
ultimately have "Key (sesK (c * d * (s + a * b))) \<in> U"
by simp
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show "n \<in> bad"
by contradiction
qed
lemma pr_sesk_card [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
IntAgrK (S (Card n, n, run)) = Some d \<longrightarrow>
ExtAgrK (S (Card n, n, run)) = Some e \<longrightarrow>
Key (sesK (d * e)) \<in> U"
proof (erule protocol.induct, simp_all, (rule impI)+, simp)
qed (subst (2) mult.commute, subst mult.assoc, simp)
lemma pr_sesk_user_1 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
IntAgrK (S (User m, n, run)) = Some c \<longrightarrow>
ExtAgrK (S (User m, n, run)) = Some f \<longrightarrow>
\<lbrace>Key (sesK (c * f)), Agent (User m), Number n, Number run\<rbrace> \<in> U"
proof (erule protocol.induct, simp_all, (rule_tac [!] impI)+, simp_all)
fix evsC3 S A U m n run
assume A: "(evsC3, S, A, U) \<in> protocol" and
"ExtMapK (S (User m, n, run)) = None"
hence "IntAgrK (S (User m, n, run)) = None"
by (rule pr_state_3)
with A have "ExtAgrK (S (User m, n, run)) = None"
by (rule pr_state_4)
moreover assume "ExtAgrK (S (User m, n, run)) = Some f"
ultimately show "\<lbrace>Key (sesK (c * f)), Agent (User m), Number n, Number run\<rbrace> \<in> U"
by simp
qed
lemma pr_sesk_user_2 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
\<lbrace>Key (sesK K), Agent (User m), Number n, Number run\<rbrace> \<in> U \<longrightarrow>
Key (sesK K) \<in> U"
by (erule protocol.induct, blast, simp_all)
lemma pr_auth_key_used:
"(evs, S, A, U) \<in> protocol \<Longrightarrow> Pri_AgrK (priAK n) \<in> U"
by (erule protocol.induct, simp_all)
lemma pr_int_mapk_used [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<in> U"
by (erule protocol.induct, simp_all)
lemma pr_valid_key_analz:
"(evs, S, A, U) \<in> protocol \<Longrightarrow> Key (pubSK X) \<in> analz (A \<union> spies evs)"
by (simp add: pr_key_analz, erule protocol.induct, simp_all)
lemma pr_pri_agrk_parts [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Pri_AgrK x \<notin> U \<longrightarrow>
Pri_AgrK x \<notin> parts (A \<union> spies evs)"
proof (induction arbitrary: x rule: protocol.induct,
simp_all add: parts_simp_insert parts_auth_data parts_crypt parts_mpair,
subst parts_simp, blast, blast, rule_tac [!] impI)
fix evsFR1 A U m s x
assume
"\<And>x. Pri_AgrK x \<notin> U \<longrightarrow> Pri_AgrK x \<notin> parts (A \<union> spies evsFR1)" and
"Pri_AgrK x \<notin> U"
hence A: "Pri_AgrK x \<notin> parts (A \<union> spies evsFR1)" ..
assume B: "Crypt (symK m) (Pri_AgrK s) \<in> synth (analz (A \<union> spies evsFR1))"
show "x \<noteq> s"
proof
assume "x = s"
hence "Crypt (symK m) (Pri_AgrK x) \<in> synth (analz (A \<union> spies evsFR1))"
using B by simp
hence "Crypt (symK m) (Pri_AgrK x) \<in> analz (A \<union> spies evsFR1) \<or>
Pri_AgrK x \<in> synth (analz (A \<union> spies evsFR1)) \<and>
Key (symK m) \<in> analz (A \<union> spies evsFR1)"
(is "?P \<or> ?Q")
by (rule synth_crypt)
moreover {
assume ?P
hence "Crypt (symK m) (Pri_AgrK x) \<in> parts (A \<union> spies evsFR1)"
by (rule subsetD [OF analz_parts_subset])
hence "Pri_AgrK x \<in> parts (A \<union> spies evsFR1)"
by (rule parts.Body)
hence False
using A by contradiction
}
moreover {
assume ?Q
hence "Pri_AgrK x \<in> synth (analz (A \<union> spies evsFR1))" ..
hence "Pri_AgrK x \<in> analz (A \<union> spies evsFR1)"
by (rule synth_simp_intro, simp)
hence "Pri_AgrK x \<in> parts (A \<union> spies evsFR1)"
by (rule subsetD [OF analz_parts_subset])
hence False
using A by contradiction
}
ultimately show False ..
qed
next
fix evsR4 S A U b n run x
assume
A: "(evsR4, S, A, U) \<in> protocol" and
B: "IntMapK (S (Card n, n, run)) = Some b" and
C: "Pri_AgrK x \<notin> U"
show "x \<noteq> priAK n \<and> x \<noteq> b"
proof (rule conjI, rule_tac [!] notI)
assume "x = priAK n"
moreover have "Pri_AgrK (priAK n) \<in> U"
using A by (rule pr_auth_key_used)
ultimately have "Pri_AgrK x \<in> U"
by simp
thus False
using C by contradiction
next
assume "x = b"
moreover have "Pri_AgrK b \<in> U"
using A and B by (rule pr_int_mapk_used)
ultimately have "Pri_AgrK x \<in> U"
by simp
thus False
using C by contradiction
qed
next
fix evsFR4 S A U s a b c f g x
assume
A: "\<And>x. Pri_AgrK x \<notin> U \<longrightarrow> Pri_AgrK x \<notin> parts (A \<union> spies evsFR4)" and
B: "(evsFR4, S, A, U) \<in> protocol" and
C: "Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace> \<in> synth (analz (A \<union> spies evsFR4))"
(is "Crypt _ ?M \<in> synth (analz ?A)") and
D: "Pri_AgrK x \<notin> U"
show "x \<noteq> g \<and> x \<noteq> b"
proof -
have E: "Pri_AgrK b \<in> U \<and> Pri_AgrK g \<in> U"
proof -
have "Crypt (sesK (c * f)) ?M \<in> analz ?A \<or>
?M \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
using C by (rule synth_crypt)
moreover {
assume "Crypt (sesK (c * f)) ?M \<in> analz ?A"
hence "Crypt (sesK (c * f)) ?M \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
hence "?M \<in> parts ?A"
by (rule parts.Body)
hence "\<lbrace>Auth_Data g b, pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> parts ?A"
by (rule parts.Snd)
hence F: "Auth_Data g b \<in> parts ?A"
by (rule parts.Fst)
hence "Pri_AgrK b \<in> parts ?A"
by (rule parts.Auth_Snd)
moreover have "Pri_AgrK g \<in> parts ?A"
using F by (rule parts.Auth_Fst)
ultimately have "Pri_AgrK b \<in> parts ?A \<and> Pri_AgrK g \<in> parts ?A" ..
}
moreover {
assume "?M \<in> synth (analz ?A) \<and>
Key (sesK (c * f)) \<in> analz ?A"
hence "?M \<in> synth (analz ?A)" ..
hence "\<lbrace>Auth_Data g b, pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "Auth_Data g b \<in> synth (analz ?A)"
by (rule synth_analz_fst)
hence "Auth_Data g b \<in> analz ?A \<or>
Pri_AgrK g \<in> analz ?A \<and> Pri_AgrK b \<in> analz ?A"
by (rule synth_auth_data)
moreover {
assume "Auth_Data g b \<in> analz ?A"
hence F: "Auth_Data g b \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
hence "Pri_AgrK b \<in> parts ?A"
by (rule parts.Auth_Snd)
moreover have "Pri_AgrK g \<in> parts ?A"
using F by (rule parts.Auth_Fst)
ultimately have "Pri_AgrK b \<in> parts ?A \<and> Pri_AgrK g \<in> parts ?A" ..
}
moreover {
assume F: "Pri_AgrK g \<in> analz ?A \<and> Pri_AgrK b \<in> analz ?A"
hence "Pri_AgrK b \<in> analz ?A" ..
hence "Pri_AgrK b \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
moreover have "Pri_AgrK g \<in> analz ?A"
using F ..
hence "Pri_AgrK g \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
ultimately have "Pri_AgrK b \<in> parts ?A \<and> Pri_AgrK g \<in> parts ?A" ..
}
ultimately have "Pri_AgrK b \<in> parts ?A \<and> Pri_AgrK g \<in> parts ?A" ..
}
ultimately have F: "Pri_AgrK b \<in> parts ?A \<and> Pri_AgrK g \<in> parts ?A" ..
hence "Pri_AgrK b \<in> parts ?A" ..
hence "Pri_AgrK b \<in> U"
by (rule contrapos_pp, insert A, simp)
moreover have "Pri_AgrK g \<in> parts ?A"
using F ..
hence "Pri_AgrK g \<in> U"
by (rule contrapos_pp, insert A, simp)
ultimately show ?thesis ..
qed
show ?thesis
proof (rule conjI, rule_tac [!] notI)
assume "x = g"
hence "Pri_AgrK x \<in> U"
using E by simp
thus False
using D by contradiction
next
assume "x = b"
hence "Pri_AgrK x \<in> U"
using E by simp
thus False
using D by contradiction
qed
qed
qed
lemma pr_pri_agrk_items:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Pri_AgrK x \<notin> U \<Longrightarrow>
items (insert (Pri_AgrK x) (A \<union> spies evs)) =
insert (Pri_AgrK x) (items (A \<union> spies evs))"
by (rule items_pri_agrk_out, rule pr_pri_agrk_parts)
lemma pr_auth_data_items:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Pri_AgrK (priAK n) \<notin> items (A \<union> spies evs) \<and>
(IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<notin> items (A \<union> spies evs))"
proof (induction arbitrary: n run b rule: protocol.induct,
simp_all add: items_simp_insert_2 items_crypt items_mpair pr_pri_agrk_items,
subst items_simp, blast+)
fix evsR1 S A U n' run' s n run b
assume
A: "(evsR1, S, A, U) \<in> protocol" and
B: "Pri_AgrK s \<notin> U"
show
"(n = n' \<and> run = run' \<longrightarrow>
priAK n' \<noteq> s \<and> (IntMapK (S (Card n', n', run')) = Some b \<longrightarrow> b \<noteq> s)) \<and>
((n = n' \<longrightarrow> run \<noteq> run') \<longrightarrow>
priAK n \<noteq> s \<and> (IntMapK (S (Card n, n, run)) = Some b \<longrightarrow> b \<noteq> s))"
proof (rule conjI, rule_tac [!] impI, rule_tac [!] conjI, rule_tac [2] impI,
rule_tac [4] impI, rule_tac [!] notI)
have "Pri_AgrK (priAK n) \<in> U"
using A by (rule pr_auth_key_used)
moreover assume "priAK n = s"
ultimately have "Pri_AgrK s \<in> U"
by simp
thus False
using B by contradiction
next
assume "IntMapK (S (Card n, n, run)) = Some b"
with A have "Pri_AgrK b \<in> U"
by (rule pr_int_mapk_used)
moreover assume "b = s"
ultimately have "Pri_AgrK s \<in> U"
by simp
thus False
using B by contradiction
next
have "Pri_AgrK (priAK n') \<in> U"
using A by (rule pr_auth_key_used)
moreover assume "priAK n' = s"
ultimately have "Pri_AgrK s \<in> U"
by simp
thus False
using B by contradiction
next
assume "IntMapK (S (Card n', n', run')) = Some b"
with A have "Pri_AgrK b \<in> U"
by (rule pr_int_mapk_used)
moreover assume "b = s"
ultimately have "Pri_AgrK s \<in> U"
by simp
thus False
using B by contradiction
qed
next
fix evsFR1 A m s n run b and S :: state
assume A: "\<And>n run b. Pri_AgrK (priAK n) \<notin> items (A \<union> spies evsFR1) \<and>
(IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<notin> items (A \<union> spies evsFR1))"
assume "Crypt (symK m) (Pri_AgrK s) \<in> synth (analz (A \<union> spies evsFR1))"
hence "Crypt (symK m) (Pri_AgrK s) \<in> analz (A \<union> spies evsFR1) \<or>
Pri_AgrK s \<in> synth (analz (A \<union> spies evsFR1)) \<and>
Key (symK m) \<in> analz (A \<union> spies evsFR1)"
(is "?P \<or> ?Q")
by (rule synth_crypt)
moreover {
assume ?P
hence "Crypt (symK m) (Pri_AgrK s) \<in> items (A \<union> spies evsFR1)"
by (rule subsetD [OF analz_items_subset])
hence "Pri_AgrK s \<in> items (A \<union> spies evsFR1)"
by (rule items.Body)
}
moreover {
assume ?Q
hence "Pri_AgrK s \<in> synth (analz (A \<union> spies evsFR1))" ..
hence "Pri_AgrK s \<in> analz (A \<union> spies evsFR1)"
by (rule synth_simp_intro, simp)
hence "Pri_AgrK s \<in> items (A \<union> spies evsFR1)"
by (rule subsetD [OF analz_items_subset])
}
ultimately have B: "Pri_AgrK s \<in> items (A \<union> spies evsFR1)" ..
show "Pri_AgrK (priAK n) \<notin> items (insert (Pri_AgrK s) (A \<union> spies evsFR1)) \<and>
(IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<notin> items (insert (Pri_AgrK s) (A \<union> spies evsFR1)))"
by (simp add: items_simp_insert_1 [OF B] A)
next
fix evsC2 S A U a n run b and m :: nat
assume
A: "(evsC2, S, A, U) \<in> protocol" and
B: "Pri_AgrK a \<notin> U"
show "m = 0 \<longrightarrow> priAK n \<noteq> a \<and> (IntMapK (S (Card n, n, run)) = Some b \<longrightarrow> b \<noteq> a)"
proof (rule impI, rule conjI, rule_tac [2] impI, rule_tac [!] notI)
have "Pri_AgrK (priAK n) \<in> U"
using A by (rule pr_auth_key_used)
moreover assume "priAK n = a"
ultimately have "Pri_AgrK a \<in> U"
by simp
thus False
using B by contradiction
next
assume "IntMapK (S (Card n, n, run)) = Some b"
with A have "Pri_AgrK b \<in> U"
by (rule pr_int_mapk_used)
moreover assume "b = a"
ultimately have "Pri_AgrK a \<in> U"
by simp
thus False
using B by contradiction
qed
next
fix evsR2 S A U b' n' run' b and n :: nat and run :: nat
assume
A: "(evsR2, S, A, U) \<in> protocol" and
B: "Pri_AgrK b' \<notin> U"
show "n = n' \<and> run = run' \<longrightarrow> b' = b \<longrightarrow> Pri_AgrK b \<notin> items (A \<union> spies evsR2)"
proof ((rule impI)+, drule sym, simp)
qed (rule contra_subsetD [OF items_parts_subset], rule pr_pri_agrk_parts [OF A B])
next
fix evsC3 S A U c n run b and m :: nat
assume
A: "(evsC3, S, A, U) \<in> protocol" and
B: "Pri_AgrK c \<notin> U"
show "m = 0 \<longrightarrow> priAK n \<noteq> c \<and> (IntMapK (S (Card n, n, run)) = Some b \<longrightarrow> b \<noteq> c)"
proof (rule impI, rule conjI, rule_tac [2] impI, rule_tac [!] notI)
have "Pri_AgrK (priAK n) \<in> U"
using A by (rule pr_auth_key_used)
moreover assume "priAK n = c"
ultimately have "Pri_AgrK c \<in> U"
by simp
thus False
using B by contradiction
next
assume "IntMapK (S (Card n, n, run)) = Some b"
with A have "Pri_AgrK b \<in> U"
by (rule pr_int_mapk_used)
moreover assume "b = c"
ultimately have "Pri_AgrK c \<in> U"
by simp
thus False
using B by contradiction
qed
next
fix evsR3 A n' run' s b' c n run b and S :: state and s' :: pri_agrk
assume
A: "\<And>n run b. Pri_AgrK (priAK n) \<notin> items (A \<union> spies evsR3) \<and>
(IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<notin> items (A \<union> spies evsR3))" and
B: "IntMapK (S (Card n', n', run')) = Some b'"
show
"(s' = s \<and> Pri_AgrK c \<in> analz (A \<union> spies evsR3) \<longrightarrow>
n = n' \<and> run = run' \<longrightarrow> b' = b \<longrightarrow>
Pri_AgrK b \<notin> items (A \<union> spies evsR3)) \<and>
((s' = s \<longrightarrow> Pri_AgrK c \<notin> analz (A \<union> spies evsR3)) \<longrightarrow>
n = n' \<and> run = run' \<longrightarrow> b' = b \<longrightarrow>
Pri_AgrK b \<notin> items (A \<union> spies evsR3))"
proof (rule conjI, (rule_tac [!] impI)+)
have "Pri_AgrK (priAK n') \<notin> items (A \<union> spies evsR3) \<and>
(IntMapK (S (Card n', n', run')) = Some b' \<longrightarrow>
Pri_AgrK b' \<notin> items (A \<union> spies evsR3))"
using A .
hence "Pri_AgrK b' \<notin> items (A \<union> spies evsR3)"
using B by simp
moreover assume "b' = b"
ultimately show "Pri_AgrK b \<notin> items (A \<union> spies evsR3)"
by simp
next
have "Pri_AgrK (priAK n') \<notin> items (A \<union> spies evsR3) \<and>
(IntMapK (S (Card n', n', run')) = Some b' \<longrightarrow>
Pri_AgrK b' \<notin> items (A \<union> spies evsR3))"
using A .
hence "Pri_AgrK b' \<notin> items (A \<union> spies evsR3)"
using B by simp
moreover assume "b' = b"
ultimately show "Pri_AgrK b \<notin> items (A \<union> spies evsR3)"
by simp
qed
next
fix evsR4 A n' run' b' n run b and S :: state
let ?M = "\<lbrace>pubAK (priAK n'), Crypt (priSK CA) (Hash (pubAK (priAK n')))\<rbrace>"
assume
A: "\<And>n run b. Pri_AgrK (priAK n) \<notin> items (A \<union> spies evsR4) \<and>
(IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<notin> items (A \<union> spies evsR4))" and
B: "IntMapK (S (Card n', n', run')) = Some b'"
show
"Pri_AgrK (priAK n) \<notin> items (insert (Auth_Data (priAK n') b')
(insert ?M (A \<union> spies evsR4))) \<and>
(IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<notin> items (insert (Auth_Data (priAK n') b')
(insert ?M (A \<union> spies evsR4))))"
proof (subst (1 2) insert_commute, simp add: items_mpair,
subst (1 3) insert_commute, simp add: items_simp_insert_2,
subst (1 2) insert_commute, simp add: items_crypt items_simp_insert_2)
have C: "Pri_AgrK (priAK n') \<notin> items (A \<union> spies evsR4) \<and>
(IntMapK (S (Card n', n', run')) = Some b' \<longrightarrow>
Pri_AgrK b' \<notin> items (A \<union> spies evsR4))"
using A .
hence "Pri_AgrK (priAK n') \<notin> items (A \<union> spies evsR4)" ..
moreover have "Pri_AgrK b' \<notin> items (A \<union> spies evsR4)"
using B and C by simp
ultimately show
"Pri_AgrK (priAK n) \<notin> items (insert (Auth_Data (priAK n') b')
(A \<union> spies evsR4)) \<and>
(IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<notin> items (insert (Auth_Data (priAK n') b')
(A \<union> spies evsR4)))"
by (simp add: items_auth_data_out A)
qed
next
fix evsFR4 A s a b' c f g n run b and S :: state
let ?M = "\<lbrace>pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>"
assume
A: "\<And>n run b. Pri_AgrK (priAK n) \<notin> items (A \<union> spies evsFR4) \<and>
(IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<notin> items (A \<union> spies evsFR4))" and
B: "Crypt (sesK (c * f)) \<lbrace>pubAK (c * (s + a * b')), Auth_Data g b', ?M\<rbrace>
\<in> synth (analz (A \<union> spies evsFR4))"
(is "Crypt _ ?M' \<in> synth (analz ?A)")
have C: "Pri_AgrK b' \<in> items ?A \<or> Pri_AgrK g \<in> items ?A \<longrightarrow>
Pri_AgrK b' \<in> items ?A \<and> Pri_AgrK g \<in> items ?A"
(is "?P \<longrightarrow> ?Q")
proof
assume ?P
have "Crypt (sesK (c * f)) ?M' \<in> analz ?A \<or>
?M' \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
using B by (rule synth_crypt)
moreover {
assume "Crypt (sesK (c * f)) ?M' \<in> analz ?A"
hence "Crypt (sesK (c * f)) ?M' \<in> items ?A"
by (rule subsetD [OF analz_items_subset])
hence "?M' \<in> items ?A"
by (rule items.Body)
hence "\<lbrace>Auth_Data g b', pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> items ?A"
by (rule items.Snd)
hence D: "Auth_Data g b' \<in> items ?A"
by (rule items.Fst)
have ?Q
proof (rule disjE [OF \<open>?P\<close>])
assume "Pri_AgrK b' \<in> items ?A"
moreover from this have "Pri_AgrK g \<in> items ?A"
by (rule items.Auth_Fst [OF D])
ultimately show ?Q ..
next
assume "Pri_AgrK g \<in> items ?A"
moreover from this have "Pri_AgrK b' \<in> items ?A"
by (rule items.Auth_Snd [OF D])
ultimately show ?Q
by simp
qed
}
moreover {
assume "?M' \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
hence "?M' \<in> synth (analz ?A)" ..
hence "\<lbrace>Auth_Data g b', pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "Auth_Data g b' \<in> synth (analz ?A)"
by (rule synth_analz_fst)
hence "Auth_Data g b' \<in> analz ?A \<or>
Pri_AgrK g \<in> analz ?A \<and> Pri_AgrK b' \<in> analz ?A"
by (rule synth_auth_data)
moreover {
assume "Auth_Data g b' \<in> analz ?A"
hence D: "Auth_Data g b' \<in> items ?A"
by (rule subsetD [OF analz_items_subset])
have ?Q
proof (rule disjE [OF \<open>?P\<close>])
assume "Pri_AgrK b' \<in> items ?A"
moreover from this have "Pri_AgrK g \<in> items ?A"
by (rule items.Auth_Fst [OF D])
ultimately show ?Q ..
next
assume "Pri_AgrK g \<in> items ?A"
moreover from this have "Pri_AgrK b' \<in> items ?A"
by (rule items.Auth_Snd [OF D])
ultimately show ?Q
by simp
qed
}
moreover {
assume D: "Pri_AgrK g \<in> analz ?A \<and> Pri_AgrK b' \<in> analz ?A"
hence "Pri_AgrK b' \<in> analz ?A" ..
hence "Pri_AgrK b' \<in> items ?A"
by (rule subsetD [OF analz_items_subset])
moreover have "Pri_AgrK g \<in> analz ?A"
using D ..
hence "Pri_AgrK g \<in> items ?A"
by (rule subsetD [OF analz_items_subset])
ultimately have ?Q ..
}
ultimately have ?Q ..
}
ultimately show ?Q ..
qed
show
"Pri_AgrK (priAK n) \<notin> items (insert (Auth_Data g b')
(insert ?M (A \<union> spies evsFR4))) \<and>
(IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<notin> items (insert (Auth_Data g b')
(insert ?M (A \<union> spies evsFR4))))"
proof (subst (1 2) insert_commute, simp add: items_mpair,
subst (1 3) insert_commute, simp add: items_simp_insert_2,
subst (1 2) insert_commute, simp add: items_crypt items_simp_insert_2, cases ?P)
case True
with C have ?Q ..
thus
"Pri_AgrK (priAK n) \<notin> items (insert (Auth_Data g b')
(A \<union> spies evsFR4)) \<and>
(IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<notin> items (insert (Auth_Data g b')
(A \<union> spies evsFR4)))"
by (simp add: items_auth_data_in items_simp_insert_1 A)
next
case False
thus
"Pri_AgrK (priAK n) \<notin> items (insert (Auth_Data g b')
(A \<union> spies evsFR4)) \<and>
(IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
Pri_AgrK b \<notin> items (insert (Auth_Data g b')
(A \<union> spies evsFR4)))"
by (simp add: items_auth_data_out A)
qed
qed
lemma pr_auth_key_analz:
"(evs, S, A, U) \<in> protocol \<Longrightarrow> Pri_AgrK (priAK n) \<notin> analz (A \<union> spies evs)"
proof (rule contra_subsetD [OF analz_items_subset], drule pr_auth_data_items)
qed (erule conjE)
lemma pr_int_mapk_analz:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
IntMapK (S (Card n, n, run)) = Some b \<Longrightarrow>
Pri_AgrK b \<notin> analz (A \<union> spies evs)"
proof (rule contra_subsetD [OF analz_items_subset], drule pr_auth_data_items)
qed (erule conjE, rule mp)
lemma pr_key_set_unused [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evs) = H \<union> analz (A \<union> spies evs)"
proof (induction arbitrary: H rule: protocol.induct, simp_all add: analz_simp_insert_2,
rule impI, (subst analz_simp, blast)+, simp)
fix evsR1 S A U n s H
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsR1) = H \<union> analz (A \<union> spies evsR1)" and
B: "(evsR1, S, A, U) \<in> protocol" and
C: "Pri_AgrK s \<notin> U"
let
?B = "H \<union> A \<union> spies evsR1" and
?C = "A \<union> spies evsR1"
show
"(n \<in> bad \<longrightarrow>
H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK s) U \<longrightarrow>
analz (insert (Crypt (symK n) (Pri_AgrK s)) (insert (Pri_AgrK s) ?B)) =
H \<union> analz (insert (Crypt (symK n) (Pri_AgrK s)) (insert (Pri_AgrK s) ?C))) \<and>
(n \<notin> bad \<longrightarrow>
H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK s) U \<longrightarrow>
analz (insert (Crypt (symK n) (Pri_AgrK s)) ?B) =
H \<union> analz (insert (Crypt (symK n) (Pri_AgrK s)) ?C))"
(is "(_ \<longrightarrow> _ \<longrightarrow> ?T) \<and> (_ \<longrightarrow> _ \<longrightarrow> ?T')")
proof (rule conjI, (rule_tac [!] impI)+)
assume "H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK s) U"
hence "insert (Pri_AgrK s) H \<subseteq> range Key \<union> range Pri_AgrK - U"
using C by blast
with A have "analz (insert (Pri_AgrK s) H \<union> A \<union> spies evsR1) =
insert (Pri_AgrK s) H \<union> analz (A \<union> spies evsR1)" ..
hence "analz (insert (Pri_AgrK s) ?B) = H \<union> insert (Pri_AgrK s) (analz ?C)"
by simp
moreover have "{Pri_AgrK s} \<subseteq> range Key \<union> range Pri_AgrK - U"
using C by simp
with A have "analz ({Pri_AgrK s} \<union> A \<union> spies evsR1) =
{Pri_AgrK s} \<union> analz (A \<union> spies evsR1)" ..
hence "insert (Pri_AgrK s) (analz ?C) = analz (insert (Pri_AgrK s) ?C)"
by simp
ultimately have D: "analz (insert (Pri_AgrK s) ?B) =
H \<union> analz (insert (Pri_AgrK s) ?C)"
by simp
assume "n \<in> bad"
hence E: "Key (invK (symK n)) \<in> analz ?C"
using B by (simp add: pr_symk_analz invK_symK)
have "Key (invK (symK n)) \<in> analz (insert (Pri_AgrK s) ?B)"
by (rule subsetD [OF _ E], rule analz_mono, blast)
hence "analz (insert (Crypt (symK n) (Pri_AgrK s)) (insert (Pri_AgrK s) ?B)) =
insert (Crypt (symK n) (Pri_AgrK s)) (analz (insert (Pri_AgrK s) ?B))"
by (simp add: analz_crypt_in)
moreover have "Key (invK (symK n)) \<in> analz (insert (Pri_AgrK s) ?C)"
by (rule subsetD [OF _ E], rule analz_mono, blast)
hence "analz (insert (Crypt (symK n) (Pri_AgrK s)) (insert (Pri_AgrK s) ?C)) =
insert (Crypt (symK n) (Pri_AgrK s)) (analz (insert (Pri_AgrK s) ?C))"
by (simp add: analz_crypt_in)
ultimately show ?T
using D by simp
next
assume "H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK s) U"
hence D: "H \<subseteq> range Key \<union> range Pri_AgrK - U"
by blast
with A have E: "analz ?B = H \<union> analz ?C" ..
moreover assume "n \<notin> bad"
hence F: "Key (invK (symK n)) \<notin> analz ?C"
using B by (simp add: pr_symk_analz invK_symK)
ultimately have "Key (invK (symK n)) \<notin> analz ?B"
proof (simp add: invK_symK, insert pr_symk_used [OF B, of n])
qed (rule notI, drule subsetD [OF D], simp)
hence "analz (insert (Crypt (symK n) (Pri_AgrK s)) ?B) =
insert (Crypt (symK n) (Pri_AgrK s)) (analz ?B)"
by (simp add: analz_crypt_out)
moreover have "H \<union> analz (insert (Crypt (symK n) (Pri_AgrK s)) ?C) =
insert (Crypt (symK n) (Pri_AgrK s)) (H \<union> analz ?C)"
using F by (simp add: analz_crypt_out)
ultimately show ?T'
using E by simp
qed
next
fix evsFR1 S A U m s H
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsFR1) = H \<union> analz (A \<union> spies evsFR1)" and
B: "(evsFR1, S, A, U) \<in> protocol" and
C: "Crypt (symK m) (Pri_AgrK s) \<in> synth (analz (A \<union> spies evsFR1))"
let
?B = "H \<union> A \<union> spies evsFR1" and
?C = "A \<union> spies evsFR1"
show "H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (insert (Crypt (symK m) (Pri_AgrK s)) ?B) =
H \<union> analz (insert (Crypt (symK m) (Pri_AgrK s)) ?C)"
(is "_ \<longrightarrow> ?T")
proof (rule impI, cases "Key (invK (symK m)) \<in> analz ?C")
case True
assume "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have "analz ?B = H \<union> analz ?C" ..
moreover have "Pri_AgrK s \<in> analz ?C"
proof (insert synth_crypt [OF C], erule disjE, erule_tac [2] conjE)
assume "Crypt (symK m) (Pri_AgrK s) \<in> analz ?C"
thus ?thesis
using True by (rule analz.Decrypt)
next
assume "Pri_AgrK s \<in> synth (analz ?C)"
thus ?thesis
by (rule synth_simp_intro, simp)
qed
moreover from this have "Pri_AgrK s \<in> analz ?B"
by (rule rev_subsetD, rule_tac analz_mono, blast)
ultimately have D: "analz (insert (Pri_AgrK s) ?B) =
H \<union> analz (insert (Pri_AgrK s) ?C)"
by (simp add: analz_simp_insert_1)
have "Key (invK (symK m)) \<in> analz ?B"
by (rule subsetD [OF _ True], rule analz_mono, blast)
hence "analz (insert (Crypt (symK m) (Pri_AgrK s)) ?B) =
insert (Crypt (symK m) (Pri_AgrK s)) (analz (insert (Pri_AgrK s) ?B))"
by (simp add: analz_crypt_in)
moreover have "analz (insert (Crypt (symK m) (Pri_AgrK s)) ?C) =
insert (Crypt (symK m) (Pri_AgrK s)) (analz (insert (Pri_AgrK s) ?C))"
using True by (simp add: analz_crypt_in)
ultimately show ?T
using D by simp
next
case False
moreover assume D: "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have E: "analz ?B = H \<union> analz ?C" ..
ultimately have "Key (invK (symK m)) \<notin> analz ?B"
proof (simp add: invK_symK, insert pr_symk_used [OF B, of m])
qed (rule notI, drule subsetD [OF D], simp)
hence "analz (insert (Crypt (symK m) (Pri_AgrK s)) ?B) =
insert (Crypt (symK m) (Pri_AgrK s)) (analz ?B)"
by (simp add: analz_crypt_out)
moreover have "H \<union> analz (insert (Crypt (symK m) (Pri_AgrK s)) ?C) =
insert (Crypt (symK m) (Pri_AgrK s)) (H \<union> analz ?C)"
using False by (simp add: analz_crypt_out)
ultimately show ?T
using E by simp
qed
next
fix evsC2 S A U a H and m :: nat
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsC2) = H \<union> analz (A \<union> spies evsC2)" and
B: "Pri_AgrK a \<notin> U"
let
?B = "H \<union> A \<union> spies evsC2" and
?C = "A \<union> spies evsC2"
show
"(m = 0 \<longrightarrow>
H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK a) U \<longrightarrow>
insert (pubAK a) (analz (insert (Pri_AgrK a) ?B)) =
insert (pubAK a) (H \<union> analz (insert (Pri_AgrK a) ?C))) \<and>
(0 < m \<longrightarrow>
H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK a) U \<longrightarrow>
insert (pubAK a) (analz ?B) =
insert (pubAK a) (H \<union> analz ?C))"
(is "(_ \<longrightarrow> _ \<longrightarrow> ?T) \<and> (_ \<longrightarrow> _ \<longrightarrow> ?T')")
proof (rule conjI, (rule_tac [!] impI)+)
assume "H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK a) U"
hence "insert (Pri_AgrK a) H \<subseteq> range Key \<union> range Pri_AgrK - U"
using B by blast
with A have "analz (insert (Pri_AgrK a) H \<union> A \<union> spies evsC2) =
insert (Pri_AgrK a) H \<union> analz (A \<union> spies evsC2)" ..
hence "analz (insert (Pri_AgrK a) ?B) = H \<union> insert (Pri_AgrK a) (analz ?C)"
by simp
moreover have "{Pri_AgrK a} \<subseteq> range Key \<union> range Pri_AgrK - U"
using B by simp
with A have "analz ({Pri_AgrK a} \<union> A \<union> spies evsC2) =
{Pri_AgrK a} \<union> analz (A \<union> spies evsC2)" ..
hence "insert (Pri_AgrK a) (analz ?C) = analz (insert (Pri_AgrK a) ?C)"
by simp
ultimately have "analz (insert (Pri_AgrK a) ?B) =
H \<union> analz (insert (Pri_AgrK a) ?C)"
by simp
thus ?T
by simp
next
assume "H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK a) U"
hence "H \<subseteq> range Key \<union> range Pri_AgrK - U"
by blast
with A have "analz ?B = H \<union> analz ?C" ..
thus ?T'
by simp
qed
next
fix evsR2 S A U b H
assume A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsR2) = H \<union> analz (A \<union> spies evsR2)"
let
?B = "H \<union> A \<union> spies evsR2" and
?C = "A \<union> spies evsR2"
show "H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK b) U \<longrightarrow>
insert (pubAK b) (analz ?B) = insert (pubAK b) (H \<union> analz ?C)"
(is "_ \<longrightarrow> ?T")
proof
assume "H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK b) U"
hence "H \<subseteq> range Key \<union> range Pri_AgrK - U"
by blast
with A have "analz ?B = H \<union> analz ?C" ..
thus ?T
by simp
qed
next
fix evsC3 S A U s a b c H and m :: nat
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsC3) = H \<union> analz (A \<union> spies evsC3)" and
B: "Pri_AgrK c \<notin> U"
let
?B = "H \<union> A \<union> spies evsC3" and
?C = "A \<union> spies evsC3"
show
"(m = 0 \<longrightarrow>
H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK c) U \<longrightarrow>
insert (pubAK (c * (s + a * b))) (analz (insert (Pri_AgrK c) ?B)) =
insert (pubAK (c * (s + a * b))) (H \<union> analz (insert (Pri_AgrK c) ?C))) \<and>
(0 < m \<longrightarrow>
H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK c) U \<longrightarrow>
insert (pubAK (c * (s + a * b))) (analz ?B) =
insert (pubAK (c * (s + a * b))) (H \<union> analz ?C))"
(is "(_ \<longrightarrow> _ \<longrightarrow> ?T) \<and> (_ \<longrightarrow> _ \<longrightarrow> ?T')")
proof (rule conjI, (rule_tac [!] impI)+)
assume "H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK c) U"
hence "insert (Pri_AgrK c) H \<subseteq> range Key \<union> range Pri_AgrK - U"
using B by blast
with A have "analz (insert (Pri_AgrK c) H \<union> A \<union> spies evsC3) =
insert (Pri_AgrK c) H \<union> analz (A \<union> spies evsC3)" ..
hence "analz (insert (Pri_AgrK c) ?B) = H \<union> insert (Pri_AgrK c) (analz ?C)"
by simp
moreover have "{Pri_AgrK c} \<subseteq> range Key \<union> range Pri_AgrK - U"
using B by simp
with A have "analz ({Pri_AgrK c} \<union> A \<union> spies evsC3) =
{Pri_AgrK c} \<union> analz (A \<union> spies evsC3)" ..
hence "insert (Pri_AgrK c) (analz ?C) = analz (insert (Pri_AgrK c) ?C)"
by simp
ultimately have "analz (insert (Pri_AgrK c) ?B) =
H \<union> analz (insert (Pri_AgrK c) ?C)"
by simp
thus ?T
by simp
next
assume "H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK c) U"
hence "H \<subseteq> range Key \<union> range Pri_AgrK - U"
by blast
with A have "analz ?B = H \<union> analz ?C" ..
thus ?T'
by simp
qed
next
fix evsR3 S A U n run s s' a b c d X H
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsR3) = H \<union> analz (A \<union> spies evsR3)" and
B: "Key (sesK (c * d * (s + a * b))) \<notin> U"
(is "Key ?K \<notin> _")
let
?B = "H \<union> A \<union> spies evsR3" and
?C = "A \<union> spies evsR3" and
?K' = "sesK (c * d * (s' + a * b))"
show
"(s' = s \<and> Pri_AgrK c \<in> analz (A \<union> spies evsR3) \<longrightarrow>
H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK d)
(insert (Key ?K) (insert \<lbrace>Key ?K, Agent X, Number n, Number run\<rbrace> U)) \<longrightarrow>
insert (pubAK (d * (s + a * b))) (analz (insert (Key ?K) ?B)) =
insert (pubAK (d * (s + a * b))) (H \<union> analz (insert (Key ?K) ?C))) \<and>
((s' = s \<longrightarrow> Pri_AgrK c \<notin> analz (A \<union> spies evsR3)) \<longrightarrow>
H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK d) (insert (Key ?K')
(insert (Key ?K) (insert \<lbrace>Key ?K, Agent X, Number n, Number run\<rbrace> U))) \<longrightarrow>
insert (pubAK (d * (s + a * b))) (analz ?B) =
insert (pubAK (d * (s + a * b))) (H \<union> analz ?C))"
(is "(_ \<longrightarrow> _ \<longrightarrow> ?T) \<and> (_ \<longrightarrow> _ \<longrightarrow> ?T')")
proof (rule conjI, (rule_tac [!] impI)+)
assume "H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK d)
(insert (Key ?K) (insert \<lbrace>Key ?K, Agent X, Number n, Number run\<rbrace> U))"
hence "insert (Key ?K) H \<subseteq> range Key \<union> range Pri_AgrK - U"
using B by blast
with A have "analz (insert (Key ?K) H \<union> A \<union> spies evsR3) =
insert (Key ?K) H \<union> analz (A \<union> spies evsR3)" ..
hence "analz (insert (Key ?K) ?B) = H \<union> insert (Key ?K) (analz ?C)"
by simp
moreover have "{Key ?K} \<subseteq> range Key \<union> range Pri_AgrK - U"
using B by simp
with A have "analz ({Key ?K} \<union> A \<union> spies evsR3) =
{Key ?K} \<union> analz (A \<union> spies evsR3)" ..
hence "insert (Key ?K) (analz ?C) = analz (insert (Key ?K) ?C)"
by simp
ultimately have "analz (insert (Key ?K) ?B) = H \<union> analz (insert (Key ?K) ?C)"
by simp
thus ?T
by simp
next
assume "H \<subseteq> range Key \<union> range Pri_AgrK - insert (Pri_AgrK d) (insert (Key ?K')
(insert (Key ?K) (insert \<lbrace>Key ?K, Agent X, Number n, Number run\<rbrace> U)))"
hence "H \<subseteq> range Key \<union> range Pri_AgrK - U"
by blast
with A have "analz ?B = H \<union> analz ?C" ..
thus ?T'
by simp
qed
next
fix evsFR3 S A U m n run s a b c d H
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsFR3) = H \<union> analz (A \<union> spies evsFR3)" and
B: "Key (sesK (c * d * (s + a * b))) \<notin> U"
(is "Key ?K \<notin> _")
let
?B = "H \<union> A \<union> spies evsFR3" and
?C = "A \<union> spies evsFR3"
show
"H \<subseteq> range Key \<union> range Pri_AgrK - insert (Key ?K)
(insert \<lbrace>Key ?K, Agent (User m), Number n, Number run\<rbrace> U) \<longrightarrow>
insert (pubAK (d * (s + a * b))) (analz (insert (Key ?K) ?B)) =
insert (pubAK (d * (s + a * b))) (H \<union> analz (insert (Key ?K) ?C))"
(is "_ \<longrightarrow> ?T")
proof
assume "H \<subseteq> range Key \<union> range Pri_AgrK - insert (Key ?K)
(insert \<lbrace>Key ?K, Agent (User m), Number n, Number run\<rbrace> U)"
hence "insert (Key ?K) H \<subseteq> range Key \<union> range Pri_AgrK - U"
using B by blast
with A have "analz (insert (Key ?K) H \<union> A \<union> spies evsFR3) =
insert (Key ?K) H \<union> analz (A \<union> spies evsFR3)" ..
hence "analz (insert (Key ?K) ?B) = H \<union> insert (Key ?K) (analz ?C)"
by simp
moreover have "{Key ?K} \<subseteq> range Key \<union> range Pri_AgrK - U"
using B by simp
with A have "analz ({Key ?K} \<union> A \<union> spies evsFR3) =
{Key ?K} \<union> analz (A \<union> spies evsFR3)" ..
hence "insert (Key ?K) (analz ?C) = analz (insert (Key ?K) ?C)"
by simp
ultimately have "analz (insert (Key ?K) ?B) = H \<union> analz (insert (Key ?K) ?C)"
by simp
thus ?T
by simp
qed
next
fix evsC4 S A U m n run c f H
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsC4) = H \<union> analz (A \<union> spies evsC4)" and
B: "(evsC4, S, A, U) \<in> protocol" and
C: "\<lbrace>Key (sesK (c * f)), Agent (User m), Number n, Number run\<rbrace> \<in> U"
let
?B = "H \<union> A \<union> spies evsC4" and
?C = "A \<union> spies evsC4"
show "H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (insert (Crypt (sesK (c * f)) (pubAK f)) ?B) =
H \<union> analz (insert (Crypt (sesK (c * f)) (pubAK f)) ?C)"
(is "_ \<longrightarrow> ?T")
proof (rule impI, cases "Key (invK (sesK (c * f))) \<in> analz ?C")
case True
assume "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have D: "analz ?B = H \<union> analz ?C" ..
have "Key (invK (sesK (c * f))) \<in> analz ?B"
by (rule subsetD [OF _ True], rule analz_mono, blast)
hence "analz (insert (Crypt (sesK (c * f)) (pubAK f)) ?B) =
insert (Crypt (sesK (c * f)) (pubAK f)) (insert (pubAK f) (analz ?B))"
by (simp add: analz_crypt_in analz_simp_insert_2)
moreover have "H \<union> analz (insert (Crypt (sesK (c * f)) (pubAK f)) ?C) =
insert (Crypt (sesK (c * f)) (pubAK f)) (insert (pubAK f) (H \<union> analz ?C))"
using True by (simp add: analz_crypt_in analz_simp_insert_2)
ultimately show ?T
using D by simp
next
case False
moreover assume D: "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have E: "analz ?B = H \<union> analz ?C" ..
ultimately have "Key (invK (sesK (c * f))) \<notin> analz ?B"
proof (simp add: invK_sesK, insert pr_sesk_user_2 [OF B C])
qed (rule notI, drule subsetD [OF D], simp)
hence "analz (insert (Crypt (sesK (c * f)) (pubAK f)) ?B) =
insert (Crypt (sesK (c * f)) (pubAK f)) (analz ?B)"
by (simp add: analz_crypt_out)
moreover have "H \<union> analz (insert (Crypt (sesK (c * f)) (pubAK f)) ?C) =
insert (Crypt (sesK (c * f)) (pubAK f)) (H \<union> analz ?C)"
using False by (simp add: analz_crypt_out)
ultimately show ?T
using E by simp
qed
next
fix evsFC4 S A U n run s a b d e H
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsFC4) = H \<union> analz (A \<union> spies evsFC4)" and
B: "(evsFC4, S, A, U) \<in> protocol" and
C: "IntAgrK (S (Card n, n, run)) = Some d" and
D: "ExtAgrK (S (Card n, n, run)) = Some e"
let
?B = "H \<union> A \<union> spies evsFC4" and
?C = "A \<union> spies evsFC4" and
?f = "d * (s + a * b)"
show "H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (insert (Crypt (sesK (d * e)) (pubAK ?f)) ?B) =
H \<union> analz (insert (Crypt (sesK (d * e)) (pubAK ?f)) ?C)"
(is "_ \<longrightarrow> ?T")
proof (rule impI, cases "Key (invK (sesK (d * e))) \<in> analz ?C")
case True
assume "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have E: "analz ?B = H \<union> analz ?C" ..
have "Key (invK (sesK (d * e))) \<in> analz ?B"
by (rule subsetD [OF _ True], rule analz_mono, blast)
hence "analz (insert (Crypt (sesK (d * e)) (pubAK ?f)) ?B) =
insert (Crypt (sesK (d * e)) (pubAK ?f)) (insert (pubAK ?f) (analz ?B))"
by (simp add: analz_crypt_in analz_simp_insert_2)
moreover have "H \<union> analz (insert (Crypt (sesK (d * e)) (pubAK ?f)) ?C) =
insert (Crypt (sesK (d * e)) (pubAK ?f)) (insert (pubAK ?f) (H \<union> analz ?C))"
using True by (simp add: analz_crypt_in analz_simp_insert_2)
ultimately show ?T
using E by simp
next
case False
moreover assume E: "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have F: "analz ?B = H \<union> analz ?C" ..
ultimately have "Key (invK (sesK (d * e))) \<notin> analz ?B"
proof (simp add: invK_sesK, insert pr_sesk_card [OF B C D])
qed (rule notI, drule subsetD [OF E], simp)
hence "analz (insert (Crypt (sesK (d * e)) (pubAK ?f)) ?B) =
insert (Crypt (sesK (d * e)) (pubAK ?f)) (analz ?B)"
by (simp add: analz_crypt_out)
moreover have "H \<union> analz (insert (Crypt (sesK (d * e)) (pubAK ?f)) ?C) =
insert (Crypt (sesK (d * e)) (pubAK ?f)) (H \<union> analz ?C)"
using False by (simp add: analz_crypt_out)
ultimately show ?T
using F by simp
qed
next
fix evsR4 S A U n run b d e H
let
?B = "H \<union> A \<union> spies evsR4" and
?C = "A \<union> spies evsR4" and
?H = "Hash (pubAK (priAK n))" and
?M = "\<lbrace>pubAK (priAK n), Crypt (priSK CA) (Hash (pubAK (priAK n)))\<rbrace>" and
?M' = "\<lbrace>pubAK e, Auth_Data (priAK n) b, pubAK (priAK n),
Crypt (priSK CA) (Hash (pubAK (priAK n)))\<rbrace>"
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsR4) = H \<union> analz (A \<union> spies evsR4)" and
B: "(evsR4, S, A, U) \<in> protocol" and
C: "IntMapK (S (Card n, n, run)) = Some b" and
D: "IntAgrK (S (Card n, n, run)) = Some d" and
E: "ExtAgrK (S (Card n, n, run)) = Some e"
show "H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (insert (Crypt (sesK (d * e)) ?M') ?B) =
H \<union> analz (insert (Crypt (sesK (d * e)) ?M') ?C)"
(is "_ \<longrightarrow> ?T")
proof
assume F: "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have G: "analz ?B = H \<union> analz ?C" ..
have H: "Key (pubSK CA) \<in> analz ?C"
using B by (rule pr_valid_key_analz)
hence I: "analz (insert (Crypt (priSK CA) ?H) ?C) =
{Crypt (priSK CA) ?H, ?H} \<union> analz ?C"
by (simp add: analz_crypt_in analz_simp_insert_2)
have "Key (pubSK CA) \<in> analz ?B"
by (rule subsetD [OF _ H], rule analz_mono, blast)
hence J: "analz (insert (Crypt (priSK CA) ?H) ?B) =
{Crypt (priSK CA) ?H, ?H} \<union> analz ?B"
by (simp add: analz_crypt_in analz_simp_insert_2)
have K: "Pri_AgrK (priAK n) \<notin> analz ?C"
using B by (rule pr_auth_key_analz)
hence L: "Pri_AgrK (priAK n) \<notin> analz (insert (Crypt (priSK CA) ?H) ?C)"
using I by simp
have M: "Pri_AgrK b \<notin> analz ?C"
using B and C by (rule pr_int_mapk_analz)
hence N: "Pri_AgrK b \<notin> analz (insert (Crypt (priSK CA) ?H) ?C)"
using I by simp
have "Pri_AgrK (priAK n) \<in> U"
using B by (rule pr_auth_key_used)
hence "Pri_AgrK (priAK n) \<notin> H"
using F by blast
hence O: "Pri_AgrK (priAK n) \<notin> analz (insert (Crypt (priSK CA) ?H) ?B)"
using G and J and K by simp
have "Pri_AgrK b \<in> U"
using B and C by (rule pr_int_mapk_used)
hence "Pri_AgrK b \<notin> H"
using F by blast
hence P: "Pri_AgrK b \<notin> analz (insert (Crypt (priSK CA) ?H) ?B)"
using G and J and M by simp
show ?T
proof (cases "Key (invK (sesK (d * e))) \<in> analz ?C")
case True
have Q: "Key (invK (sesK (d * e))) \<in> analz ?B"
by (rule subsetD [OF _ True], rule analz_mono, blast)
show ?T
proof (simp add: analz_crypt_in analz_mpair analz_simp_insert_2 True Q,
rule equalityI, (rule_tac [!] insert_mono)+)
show "analz (insert (Auth_Data (priAK n) b) (insert ?M ?B)) \<subseteq>
H \<union> analz (insert (Auth_Data (priAK n) b) (insert ?M ?C))"
proof (subst (1 2) insert_commute, simp add: analz_mpair analz_simp_insert_2
del: Un_insert_right, subst (1 3) insert_commute,
subst analz_auth_data_out [OF O P], subst analz_auth_data_out [OF L N])
qed (auto simp add: G I J)
next
show "H \<union> analz (insert (Auth_Data (priAK n) b) (insert ?M ?C)) \<subseteq>
analz (insert (Auth_Data (priAK n) b) (insert ?M ?B))"
proof (subst (1 2) insert_commute, simp add: analz_mpair analz_simp_insert_2
del: Un_insert_right Un_subset_iff semilattice_sup_class.sup.bounded_iff,
subst (1 3) insert_commute, subst analz_auth_data_out [OF L N],
subst analz_auth_data_out [OF O P])
qed (auto simp add: G I J)
qed
next
case False
hence "Key (invK (sesK (d * e))) \<notin> analz ?B"
proof (simp add: invK_sesK G, insert pr_sesk_card [OF B D E])
qed (rule notI, drule subsetD [OF F], simp)
hence "analz (insert (Crypt (sesK (d * e)) ?M') ?B) =
insert (Crypt (sesK (d * e)) ?M') (analz ?B)"
by (simp add: analz_crypt_out)
moreover have "H \<union> analz (insert (Crypt (sesK (d * e)) ?M') ?C) =
insert (Crypt (sesK (d * e)) ?M') (H \<union> analz ?C)"
using False by (simp add: analz_crypt_out)
ultimately show ?T
using G by simp
qed
qed
next
fix evsFR4 S A U m n run s a b c f g H
let
?B = "H \<union> A \<union> spies evsFR4" and
?C = "A \<union> spies evsFR4" and
?H = "Hash (pubAK g)" and
?M = "\<lbrace>pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>" and
?M' = "\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>"
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsFR4) = H \<union> analz (A \<union> spies evsFR4)" and
B: "(evsFR4, S, A, U) \<in> protocol" and
C: "IntAgrK (S (User m, n, run)) = Some c" and
D: "ExtAgrK (S (User m, n, run)) = Some f" and
E: "Crypt (sesK (c * f)) ?M' \<in> synth (analz ?C)"
have F:
"Key (invK (sesK (c * f))) \<in> analz ?C \<longrightarrow>
Pri_AgrK b \<in> analz ?C \<or> Pri_AgrK g \<in> analz ?C \<longrightarrow>
Pri_AgrK b \<in> analz ?C \<and> Pri_AgrK g \<in> analz ?C"
(is "?P \<longrightarrow> ?Q \<longrightarrow> ?R")
proof (rule impI)+
assume ?P and ?Q
have "Crypt (sesK (c * f)) ?M' \<in> analz ?C \<or>
?M' \<in> synth (analz ?C) \<and> Key (sesK (c * f)) \<in> analz ?C"
using E by (rule synth_crypt)
moreover {
assume "Crypt (sesK (c * f)) ?M' \<in> analz ?C"
hence "?M' \<in> analz ?C"
using \<open>?P\<close> by (rule analz.Decrypt)
hence "\<lbrace>Auth_Data g b, pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> analz ?C"
by (rule analz.Snd)
hence G: "Auth_Data g b \<in> analz ?C"
by (rule analz.Fst)
have ?R
proof (rule disjE [OF \<open>?Q\<close>])
assume "Pri_AgrK b \<in> analz ?C"
moreover from this have "Pri_AgrK g \<in> analz ?C"
by (rule analz.Auth_Fst [OF G])
ultimately show ?R ..
next
assume "Pri_AgrK g \<in> analz ?C"
moreover from this have "Pri_AgrK b \<in> analz ?C"
by (rule analz.Auth_Snd [OF G])
ultimately show ?R
by simp
qed
}
moreover {
assume "?M' \<in> synth (analz ?C) \<and> Key (sesK (c * f)) \<in> analz ?C"
hence "?M' \<in> synth (analz ?C)" ..
hence "\<lbrace>Auth_Data g b, pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> synth (analz ?C)"
by (rule synth_analz_snd)
hence "Auth_Data g b \<in> synth (analz ?C)"
by (rule synth_analz_fst)
hence "Auth_Data g b \<in> analz ?C \<or>
Pri_AgrK g \<in> analz ?C \<and> Pri_AgrK b \<in> analz ?C"
by (rule synth_auth_data)
moreover {
assume G: "Auth_Data g b \<in> analz ?C"
have ?R
proof (rule disjE [OF \<open>?Q\<close>])
assume "Pri_AgrK b \<in> analz ?C"
moreover from this have "Pri_AgrK g \<in> analz ?C"
by (rule analz.Auth_Fst [OF G])
ultimately show ?R ..
next
assume "Pri_AgrK g \<in> analz ?C"
moreover from this have "Pri_AgrK b \<in> analz ?C"
by (rule analz.Auth_Snd [OF G])
ultimately show ?R
by simp
qed
}
moreover {
assume "Pri_AgrK g \<in> analz ?C \<and> Pri_AgrK b \<in> analz ?C"
hence ?R
by simp
}
ultimately have ?R ..
}
ultimately show ?R ..
qed
show "H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (insert (Crypt (sesK (c * f)) ?M') ?B) =
H \<union> analz (insert (Crypt (sesK (c * f)) ?M') ?C)"
(is "_ \<longrightarrow> ?T")
proof
assume G: "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have H: "analz ?B = H \<union> analz ?C" ..
have I: "Key (pubSK CA) \<in> analz ?C"
using B by (rule pr_valid_key_analz)
hence J: "analz (insert (Crypt (priSK CA) ?H) ?C) =
{Crypt (priSK CA) ?H, ?H} \<union> analz ?C"
by (simp add: analz_crypt_in analz_simp_insert_2)
have "Key (pubSK CA) \<in> analz ?B"
by (rule subsetD [OF _ I], rule analz_mono, blast)
hence K: "analz (insert (Crypt (priSK CA) ?H) ?B) =
{Crypt (priSK CA) ?H, ?H} \<union> analz ?B"
by (simp add: analz_crypt_in analz_simp_insert_2)
show ?T
proof (cases "Key (invK (sesK (c * f))) \<in> analz ?C",
cases "Pri_AgrK g \<in> analz ?C \<or> Pri_AgrK b \<in> analz ?C", simp_all)
assume L: "Key (invK (sesK (c * f))) \<in> analz ?C"
have M: "Key (invK (sesK (c * f))) \<in> analz ?B"
by (rule subsetD [OF _ L], rule analz_mono, blast)
assume N: "Pri_AgrK g \<in> analz ?C \<or> Pri_AgrK b \<in> analz ?C"
hence O: "Pri_AgrK g \<in> analz (insert (Crypt (priSK CA) ?H) ?C) \<or>
Pri_AgrK b \<in> analz (insert (Crypt (priSK CA) ?H) ?C)"
using J by simp
have "Pri_AgrK g \<in> analz ?B \<or> Pri_AgrK b \<in> analz ?B"
using H and N by blast
hence P: "Pri_AgrK g \<in> analz (insert (Crypt (priSK CA) ?H) ?B) \<or>
Pri_AgrK b \<in> analz (insert (Crypt (priSK CA) ?H) ?B)"
using K by simp
have Q: "Pri_AgrK b \<in> analz ?C \<and> Pri_AgrK g \<in> analz ?C"
using F and L and N by blast
hence "Pri_AgrK g \<in> analz (insert (Crypt (priSK CA) ?H) ?C)"
using J by simp
hence R: "Pri_AgrK g \<in> analz (insert (Pri_AgrK b)
(insert (Crypt (priSK CA) ?H) ?C))"
by (rule rev_subsetD, rule_tac analz_mono, blast)
have S: "Pri_AgrK b \<in> analz (insert (Crypt (priSK CA) ?H) ?C)"
using J and Q by simp
have T: "Pri_AgrK b \<in> analz ?B \<and> Pri_AgrK g \<in> analz ?B"
using H and Q by simp
hence "Pri_AgrK g \<in> analz (insert (Crypt (priSK CA) ?H) ?B)"
using K by simp
hence U: "Pri_AgrK g \<in> analz (insert (Pri_AgrK b)
(insert (Crypt (priSK CA) ?H) ?B))"
by (rule rev_subsetD, rule_tac analz_mono, blast)
have V: "Pri_AgrK b \<in> analz (insert (Crypt (priSK CA) ?H) ?B)"
using K and T by simp
show ?T
proof (simp add: analz_crypt_in analz_mpair analz_simp_insert_2 L M,
rule equalityI, (rule_tac [!] insert_mono)+)
show "analz (insert (Auth_Data g b) (insert ?M ?B)) \<subseteq>
H \<union> analz (insert (Auth_Data g b) (insert ?M ?C))"
proof (subst (1 2) insert_commute, simp add: analz_mpair analz_simp_insert_2
del: Un_insert_right, subst (1 3) insert_commute,
subst analz_auth_data_in [OF P], subst analz_auth_data_in [OF O],
simp del: Un_insert_right)
show
"analz (insert (Pri_AgrK g) (insert (Pri_AgrK b)
(insert (Crypt (priSK CA) ?H) ?B))) \<subseteq>
H \<union> insert ?M (insert (pubAK g) (insert (Auth_Data g b)
(analz (insert (Pri_AgrK g) (insert (Pri_AgrK b)
(insert (Crypt (priSK CA) ?H) ?C))))))"
proof (subst analz_simp_insert_1 [OF U], subst analz_simp_insert_1 [OF V],
subst analz_simp_insert_1 [OF R], subst analz_simp_insert_1 [OF S])
qed (auto simp add: H J K)
qed
next
show "H \<union> analz (insert (Auth_Data g b) (insert ?M ?C)) \<subseteq>
analz (insert (Auth_Data g b) (insert ?M ?B))"
proof (subst (1 2) insert_commute, simp add: analz_mpair analz_simp_insert_2
del: Un_insert_right Un_subset_iff semilattice_sup_class.sup.bounded_iff,
subst (2 4) insert_commute, subst analz_auth_data_in [OF O],
subst analz_auth_data_in [OF P], simp only: Un_insert_left Un_empty_left)
show
"H \<union> insert ?M (insert (pubAK g) (insert (Auth_Data g b)
(analz (insert (Pri_AgrK g) (insert (Pri_AgrK b)
(insert (Crypt (priSK CA) ?H) ?C)))))) \<subseteq>
insert ?M (insert (pubAK g) (insert (Auth_Data g b)
(analz (insert (Pri_AgrK g) (insert (Pri_AgrK b)
(insert (Crypt (priSK CA) ?H) ?B))))))"
proof (subst analz_simp_insert_1 [OF R], subst analz_simp_insert_1 [OF S],
subst analz_simp_insert_1 [OF U], subst analz_simp_insert_1 [OF V])
qed (auto simp add: H J K)
qed
qed
next
assume L: "Key (invK (sesK (c * f))) \<in> analz ?C"
have M: "Key (invK (sesK (c * f))) \<in> analz ?B"
by (rule subsetD [OF _ L], rule analz_mono, blast)
assume N: "Pri_AgrK g \<notin> analz ?C \<and> Pri_AgrK b \<notin> analz ?C"
hence O: "Pri_AgrK g \<notin> analz (insert (Crypt (priSK CA) ?H) ?C)"
using J by simp
have P: "Pri_AgrK b \<notin> analz (insert (Crypt (priSK CA) ?H) ?C)"
using J and N by simp
have Q: "Pri_AgrK g \<in> U \<and> Pri_AgrK b \<in> U"
proof -
have "Crypt (sesK (c * f)) ?M' \<in> analz ?C \<or>
?M' \<in> synth (analz ?C) \<and> Key (sesK (c * f)) \<in> analz ?C"
using E by (rule synth_crypt)
moreover {
assume "Crypt (sesK (c * f)) ?M' \<in> analz ?C"
hence "Crypt (sesK (c * f)) ?M' \<in> parts ?C"
by (rule subsetD [OF analz_parts_subset])
hence "?M' \<in> parts ?C"
by (rule parts.Body)
hence "\<lbrace>Auth_Data g b, pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> parts ?C"
by (rule parts.Snd)
hence R: "Auth_Data g b \<in> parts ?C"
by (rule parts.Fst)
hence "Pri_AgrK g \<in> parts ?C"
by (rule parts.Auth_Fst)
hence "Pri_AgrK g \<in> U"
by (rule contrapos_pp, rule_tac pr_pri_agrk_parts [OF B])
moreover have "Pri_AgrK b \<in> parts ?C"
using R by (rule parts.Auth_Snd)
hence "Pri_AgrK b \<in> U"
by (rule contrapos_pp, rule_tac pr_pri_agrk_parts [OF B])
ultimately have ?thesis ..
}
moreover {
assume "?M' \<in> synth (analz ?C) \<and>
Key (sesK (c * f)) \<in> analz ?C"
hence "?M' \<in> synth (analz ?C)" ..
hence "\<lbrace>Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace> \<in> synth (analz ?C)"
by (rule synth_analz_snd)
hence "Auth_Data g b \<in> synth (analz ?C)"
by (rule synth_analz_fst)
hence "Auth_Data g b \<in> analz ?C \<or>
Pri_AgrK g \<in> analz ?C \<and> Pri_AgrK b \<in> analz ?C"
by (rule synth_auth_data)
moreover {
assume "Auth_Data g b \<in> analz ?C"
hence R: "Auth_Data g b \<in> parts ?C"
by (rule subsetD [OF analz_parts_subset])
hence "Pri_AgrK g \<in> parts ?C"
by (rule parts.Auth_Fst)
hence "Pri_AgrK g \<in> U"
by (rule contrapos_pp, rule_tac pr_pri_agrk_parts [OF B])
moreover have "Pri_AgrK b \<in> parts ?C"
using R by (rule parts.Auth_Snd)
hence "Pri_AgrK b \<in> U"
by (rule contrapos_pp, rule_tac pr_pri_agrk_parts [OF B])
ultimately have ?thesis ..
}
moreover {
assume R: "Pri_AgrK g \<in> analz ?C \<and> Pri_AgrK b \<in> analz ?C"
hence "Pri_AgrK g \<in> analz ?C" ..
hence "Pri_AgrK g \<in> parts ?C"
by (rule subsetD [OF analz_parts_subset])
hence "Pri_AgrK g \<in> U"
by (rule contrapos_pp, rule_tac pr_pri_agrk_parts [OF B])
moreover have "Pri_AgrK b \<in> analz ?C"
using R ..
hence "Pri_AgrK b \<in> parts ?C"
by (rule subsetD [OF analz_parts_subset])
hence "Pri_AgrK b \<in> U"
by (rule contrapos_pp, rule_tac pr_pri_agrk_parts [OF B])
ultimately have ?thesis ..
}
ultimately have ?thesis ..
}
ultimately show ?thesis ..
qed
have R: "Pri_AgrK g \<notin> analz ?B \<and> Pri_AgrK b \<notin> analz ?B"
proof (simp add: H N, rule conjI, rule_tac [!] notI, drule_tac [!] subsetD [OF G])
qed (simp_all add: Q)
hence S: "Pri_AgrK g \<notin> analz (insert (Crypt (priSK CA) ?H) ?B)"
using K by simp
have T: "Pri_AgrK b \<notin> analz (insert (Crypt (priSK CA) ?H) ?B)"
using K and R by simp
show ?T
proof (simp add: analz_crypt_in analz_mpair analz_simp_insert_2 L M,
rule equalityI, (rule_tac [!] insert_mono)+)
show "analz (insert (Auth_Data g b) (insert ?M ?B)) \<subseteq>
H \<union> analz (insert (Auth_Data g b) (insert ?M ?C))"
proof (subst (1 2) insert_commute, simp add: analz_mpair analz_simp_insert_2
del: Un_insert_right, subst (1 3) insert_commute,
subst analz_auth_data_out [OF S T], subst analz_auth_data_out [OF O P])
qed (auto simp add: H J K)
next
show "H \<union> analz (insert (Auth_Data g b) (insert ?M ?C)) \<subseteq>
analz (insert (Auth_Data g b) (insert ?M ?B))"
proof (subst (1 2) insert_commute, simp add: analz_mpair analz_simp_insert_2
del: Un_insert_right Un_subset_iff semilattice_sup_class.sup.bounded_iff,
subst (2 4) insert_commute, subst analz_auth_data_out [OF O P],
subst analz_auth_data_out [OF S T])
qed (simp add: H J K)
qed
next
assume L: "Key (invK (sesK (c * f))) \<notin> analz ?C"
hence "Key (invK (sesK (c * f))) \<notin> analz ?B"
proof (simp add: invK_sesK, insert pr_sesk_user_1 [OF B C D,
THEN pr_sesk_user_2 [OF B]])
qed (rule notI, simp add: H, drule subsetD [OF G], simp)
hence "analz (insert (Crypt (sesK (c * f)) ?M') ?B) =
insert (Crypt (sesK (c * f)) ?M') (analz ?B)"
by (simp add: analz_crypt_out)
moreover have "H \<union> analz (insert (Crypt (sesK (c * f)) ?M') ?C) =
insert (Crypt (sesK (c * f)) ?M') (H \<union> analz ?C)"
using L by (simp add: analz_crypt_out)
ultimately show ?T
using H by simp
qed
qed
next
fix evsC5 S A U m n run c f H
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsC5) = H \<union> analz (A \<union> spies evsC5)" and
B: "(evsC5, S, A, U) \<in> protocol" and
C: "IntAgrK (S (User m, n, run)) = Some c" and
D: "ExtAgrK (S (User m, n, run)) = Some f"
let
?B = "H \<union> A \<union> spies evsC5" and
?C = "A \<union> spies evsC5"
show "H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (insert (Crypt (sesK (c * f)) (Passwd m)) ?B) =
H \<union> analz (insert (Crypt (sesK (c * f)) (Passwd m)) ?C)"
(is "_ \<longrightarrow> ?T")
proof (rule impI, cases "Key (invK (sesK (c * f))) \<in> analz ?C")
case True
assume "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have E: "analz ?B = H \<union> analz ?C" ..
have "Key (invK (sesK (c * f))) \<in> analz ?B"
by (rule subsetD [OF _ True], rule analz_mono, blast)
hence "analz (insert (Crypt (sesK (c * f)) (Passwd m)) ?B) =
insert (Crypt (sesK (c * f)) (Passwd m)) (insert (Passwd m) (analz ?B))"
by (simp add: analz_crypt_in analz_simp_insert_2)
moreover have "H \<union> analz (insert (Crypt (sesK (c * f)) (Passwd m)) ?C) =
insert (Crypt (sesK (c * f)) (Passwd m)) (insert (Passwd m) (H \<union> analz ?C))"
using True by (simp add: analz_crypt_in analz_simp_insert_2)
ultimately show ?T
using E by simp
next
case False
moreover assume E: "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have F: "analz ?B = H \<union> analz ?C" ..
ultimately have "Key (invK (sesK (c * f))) \<notin> analz ?B"
proof (simp add: invK_sesK, insert pr_sesk_user_1 [OF B C D,
THEN pr_sesk_user_2 [OF B]])
qed (rule notI, drule subsetD [OF E], simp)
hence "analz (insert (Crypt (sesK (c * f)) (Passwd m)) ?B) =
insert (Crypt (sesK (c * f)) (Passwd m)) (analz ?B)"
by (simp add: analz_crypt_out)
moreover have "H \<union> analz (insert (Crypt (sesK (c * f)) (Passwd m)) ?C) =
insert (Crypt (sesK (c * f)) (Passwd m)) (H \<union> analz ?C)"
using False by (simp add: analz_crypt_out)
ultimately show ?T
using F by simp
qed
next
fix evsFC5 S A U n run d e H
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsFC5) = H \<union> analz (A \<union> spies evsFC5)" and
B: "(evsFC5, S, A, U) \<in> protocol" and
C: "IntAgrK (S (Card n, n, run)) = Some d" and
D: "ExtAgrK (S (Card n, n, run)) = Some e"
let
?B = "H \<union> A \<union> spies evsFC5" and
?C = "A \<union> spies evsFC5"
show "H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (insert (Crypt (sesK (d * e)) (Passwd n)) ?B) =
H \<union> analz (insert (Crypt (sesK (d * e)) (Passwd n)) ?C)"
(is "_ \<longrightarrow> ?T")
proof (rule impI, cases "Key (invK (sesK (d * e))) \<in> analz ?C")
case True
assume "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have E: "analz ?B = H \<union> analz ?C" ..
have "Key (invK (sesK (d * e))) \<in> analz ?B"
by (rule subsetD [OF _ True], rule analz_mono, blast)
hence "analz (insert (Crypt (sesK (d * e)) (Passwd n)) ?B) =
insert (Crypt (sesK (d * e)) (Passwd n)) (insert (Passwd n) (analz ?B))"
by (simp add: analz_crypt_in analz_simp_insert_2)
moreover have "H \<union> analz (insert (Crypt (sesK (d * e)) (Passwd n)) ?C) =
insert (Crypt (sesK (d * e)) (Passwd n)) (insert (Passwd n) (H \<union> analz ?C))"
using True by (simp add: analz_crypt_in analz_simp_insert_2)
ultimately show ?T
using E by simp
next
case False
moreover assume E: "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have F: "analz ?B = H \<union> analz ?C" ..
ultimately have "Key (invK (sesK (d * e))) \<notin> analz ?B"
proof (simp add: invK_sesK, insert pr_sesk_card [OF B C D])
qed (rule notI, drule subsetD [OF E], simp)
hence "analz (insert (Crypt (sesK (d * e)) (Passwd n)) ?B) =
insert (Crypt (sesK (d * e)) (Passwd n)) (analz ?B)"
by (simp add: analz_crypt_out)
moreover have "H \<union> analz (insert (Crypt (sesK (d * e)) (Passwd n)) ?C) =
insert (Crypt (sesK (d * e)) (Passwd n)) (H \<union> analz ?C)"
using False by (simp add: analz_crypt_out)
ultimately show ?T
using F by simp
qed
next
fix evsR5 S A U n run d e H
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsR5) = H \<union> analz (A \<union> spies evsR5)" and
B: "(evsR5, S, A, U) \<in> protocol" and
C: "IntAgrK (S (Card n, n, run)) = Some d" and
D: "ExtAgrK (S (Card n, n, run)) = Some e"
let
?B = "H \<union> A \<union> spies evsR5" and
?C = "A \<union> spies evsR5"
show "H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (insert (Crypt (sesK (d * e)) (Number 0)) (H \<union> A \<union> spies evsR5)) =
H \<union> analz (insert (Crypt (sesK (d * e)) (Number 0)) (A \<union> spies evsR5))"
(is "_ \<longrightarrow> ?T")
proof (rule impI, cases "Key (invK (sesK (d * e))) \<in> analz ?C")
case True
assume "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have E: "analz ?B = H \<union> analz ?C" ..
have "Key (invK (sesK (d * e))) \<in> analz ?B"
by (rule subsetD [OF _ True], rule analz_mono, blast)
hence "analz (insert (Crypt (sesK (d * e)) (Number 0)) ?B) =
insert (Crypt (sesK (d * e)) (Number 0)) (insert (Number 0) (analz ?B))"
by (simp add: analz_crypt_in analz_simp_insert_2)
moreover have "H \<union> analz (insert (Crypt (sesK (d * e)) (Number 0)) ?C) =
insert (Crypt (sesK (d * e)) (Number 0)) (insert (Number 0) (H \<union> analz ?C))"
using True by (simp add: analz_crypt_in analz_simp_insert_2)
ultimately show ?T
using E by simp
next
case False
moreover assume E: "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have F: "analz ?B = H \<union> analz ?C" ..
ultimately have "Key (invK (sesK (d * e))) \<notin> analz ?B"
proof (simp add: invK_sesK, insert pr_sesk_card [OF B C D])
qed (rule notI, drule subsetD [OF E], simp)
hence "analz (insert (Crypt (sesK (d * e)) (Number 0)) ?B) =
insert (Crypt (sesK (d * e)) (Number 0)) (analz ?B)"
by (simp add: analz_crypt_out)
moreover have "H \<union> analz (insert (Crypt (sesK (d * e)) (Number 0)) ?C) =
insert (Crypt (sesK (d * e)) (Number 0)) (H \<union> analz ?C)"
using False by (simp add: analz_crypt_out)
ultimately show ?T
using F by simp
qed
next
fix evsFR5 S A U m n run c f H
assume
A: "\<And>H. H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (H \<union> A \<union> spies evsFR5) = H \<union> analz (A \<union> spies evsFR5)" and
B: "(evsFR5, S, A, U) \<in> protocol" and
C: "IntAgrK (S (User m, n, run)) = Some c" and
D: "ExtAgrK (S (User m, n, run)) = Some f"
let
?B = "H \<union> A \<union> spies evsFR5" and
?C = "A \<union> spies evsFR5"
show "H \<subseteq> range Key \<union> range Pri_AgrK - U \<longrightarrow>
analz (insert (Crypt (sesK (c * f)) (Number 0)) (H \<union> A \<union> spies evsFR5)) =
H \<union> analz (insert (Crypt (sesK (c * f)) (Number 0)) (A \<union> spies evsFR5))"
(is "_ \<longrightarrow> ?T")
proof (rule impI, cases "Key (invK (sesK (c * f))) \<in> analz ?C")
case True
assume "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have E: "analz ?B = H \<union> analz ?C" ..
have "Key (invK (sesK (c * f))) \<in> analz ?B"
by (rule subsetD [OF _ True], rule analz_mono, blast)
hence "analz (insert (Crypt (sesK (c * f)) (Number 0)) ?B) =
insert (Crypt (sesK (c * f)) (Number 0)) (insert (Number 0) (analz ?B))"
by (simp add: analz_crypt_in analz_simp_insert_2)
moreover have "H \<union> analz (insert (Crypt (sesK (c * f)) (Number 0)) ?C) =
insert (Crypt (sesK (c * f)) (Number 0)) (insert (Number 0) (H \<union> analz ?C))"
using True by (simp add: analz_crypt_in analz_simp_insert_2)
ultimately show ?T
using E by simp
next
case False
moreover assume E: "H \<subseteq> range Key \<union> range Pri_AgrK - U"
with A have F: "analz ?B = H \<union> analz ?C" ..
ultimately have "Key (invK (sesK (c * f))) \<notin> analz ?B"
proof (simp add: invK_sesK, insert pr_sesk_user_1 [OF B C D,
THEN pr_sesk_user_2 [OF B]])
qed (rule notI, drule subsetD [OF E], simp)
hence "analz (insert (Crypt (sesK (c * f)) (Number 0)) ?B) =
insert (Crypt (sesK (c * f)) (Number 0)) (analz ?B)"
by (simp add: analz_crypt_out)
moreover have "H \<union> analz (insert (Crypt (sesK (c * f)) (Number 0)) ?C) =
insert (Crypt (sesK (c * f)) (Number 0)) (H \<union> analz ?C)"
using False by (simp add: analz_crypt_out)
ultimately show ?T
using F by simp
qed
qed
lemma pr_key_unused:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Key K \<notin> U \<Longrightarrow>
analz (insert (Key K) (A \<union> spies evs)) =
insert (Key K) (analz (A \<union> spies evs))"
by (simp only: insert_def Un_assoc [symmetric], rule pr_key_set_unused, simp_all)
lemma pr_pri_agrk_unused:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Pri_AgrK x \<notin> U \<Longrightarrow>
analz (insert (Pri_AgrK x) (A \<union> spies evs)) =
insert (Pri_AgrK x) (analz (A \<union> spies evs))"
by (simp only: insert_def Un_assoc [symmetric], rule pr_key_set_unused, simp_all)
lemma pr_pri_agrk_analz_intro [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Pri_AgrK x \<in> analz (A \<union> spies evs) \<longrightarrow>
Pri_AgrK x \<in> A"
proof (erule protocol.induct, subst analz_simp, simp, blast,
simp_all add: analz_simp_insert_2 pr_key_unused pr_pri_agrk_unused,
rule conjI, rule_tac [1-2] impI, rule_tac [!] impI)
fix evsR1 S A U n s
assume
A: "Pri_AgrK x \<in> analz (A \<union> spies evsR1) \<longrightarrow> Pri_AgrK x \<in> A"
(is "_ \<in> analz ?A \<longrightarrow> _") and
B: "(evsR1, S, A, U) \<in> protocol" and
C: "n \<in> bad" and
D: "Pri_AgrK x \<in> analz (insert (Crypt (symK n) (Pri_AgrK s))
(insert (Pri_AgrK s) (A \<union> spies evsR1)))" and
E: "Pri_AgrK s \<notin> U"
have "Key (symK n) \<in> analz ?A"
using B and C by (simp add: pr_symk_analz)
hence "Key (invK (symK n)) \<in> analz ?A"
by (simp add: invK_symK)
hence "Key (invK (symK n)) \<in> analz (insert (Pri_AgrK s) ?A)"
using B and E by (simp add: pr_pri_agrk_unused)
hence "Pri_AgrK x \<in> analz (insert (Pri_AgrK s) ?A)"
using D by (simp add: analz_crypt_in)
hence "x = s \<or> Pri_AgrK x \<in> analz ?A"
using B and E by (simp add: pr_pri_agrk_unused)
thus "x = s \<or> Pri_AgrK x \<in> A"
using A by blast
next
fix evsR1 S A U n s
assume
A: "Pri_AgrK x \<in> analz (A \<union> spies evsR1) \<longrightarrow> Pri_AgrK x \<in> A"
(is "_ \<in> analz ?A \<longrightarrow> _") and
B: "(evsR1, S, A, U) \<in> protocol" and
C: "n \<notin> bad" and
D: "Pri_AgrK x \<in> analz (insert (Crypt (symK n) (Pri_AgrK s))
(A \<union> spies evsR1))"
have "Key (symK n) \<notin> analz ?A"
using B and C by (simp add: pr_symk_analz)
hence "Key (invK (symK n)) \<notin> analz ?A"
by (simp add: invK_symK)
hence "Pri_AgrK x \<in> analz ?A"
using D by (simp add: analz_crypt_out)
with A show "Pri_AgrK x \<in> A" ..
next
fix evsFR1 A m s
assume
A: "Pri_AgrK x \<in> analz (A \<union> spies evsFR1) \<longrightarrow> Pri_AgrK x \<in> A"
(is "_ \<in> analz ?A \<longrightarrow> _") and
B: "Crypt (symK m) (Pri_AgrK s) \<in> synth (analz (A \<union> spies evsFR1))" and
C: "Pri_AgrK x \<in> analz (insert (Crypt (symK m) (Pri_AgrK s))
(A \<union> spies evsFR1))"
show "Pri_AgrK x \<in> A"
proof (cases "Key (invK (symK m)) \<in> analz ?A")
case True
have "Crypt (symK m) (Pri_AgrK s) \<in> analz ?A \<or>
Pri_AgrK s \<in> synth (analz ?A) \<and> Key (symK m) \<in> analz ?A"
using B by (rule synth_crypt)
moreover {
assume "Crypt (symK m) (Pri_AgrK s) \<in> analz ?A"
hence "Pri_AgrK s \<in> analz ?A"
using True by (rule analz.Decrypt)
}
moreover {
assume "Pri_AgrK s \<in> synth (analz ?A) \<and> Key (symK m) \<in> analz ?A"
hence "Pri_AgrK s \<in> synth (analz ?A)" ..
hence "Pri_AgrK s \<in> analz ?A"
by (rule synth_simp_intro, simp)
}
ultimately have "Pri_AgrK s \<in> analz ?A" ..
moreover have "Pri_AgrK x \<in> analz (insert (Pri_AgrK s) ?A)"
using C and True by (simp add: analz_crypt_in)
ultimately have "Pri_AgrK x \<in> analz ?A"
by (simp add: analz_simp_insert_1)
with A show ?thesis ..
next
case False
hence "Pri_AgrK x \<in> analz ?A"
using C by (simp add: analz_crypt_out)
with A show ?thesis ..
qed
next
fix evsC4 A c f
assume
A: "Pri_AgrK x \<in> analz (A \<union> spies evsC4) \<longrightarrow> Pri_AgrK x \<in> A"
(is "_ \<in> analz ?A \<longrightarrow> _") and
B: "Pri_AgrK x \<in> analz (insert (Crypt (sesK (c * f)) (pubAK f))
(A \<union> spies evsC4))"
show "Pri_AgrK x \<in> A"
proof (cases "Key (invK (sesK (c * f))) \<in> analz ?A")
case True
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_in analz_simp_insert_2)
with A show ?thesis ..
next
case False
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_out)
with A show ?thesis ..
qed
next
fix evsFC4 A s a b d e
assume
A: "Pri_AgrK x \<in> analz (A \<union> spies evsFC4) \<longrightarrow> Pri_AgrK x \<in> A"
(is "_ \<in> analz ?A \<longrightarrow> _") and
B: "Pri_AgrK x \<in> analz (insert (Crypt (sesK (d * e)) (pubAK (d * (s + a * b))))
(A \<union> spies evsFC4))"
show "Pri_AgrK x \<in> A"
proof (cases "Key (invK (sesK (d * e))) \<in> analz ?A")
case True
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_in analz_simp_insert_2)
with A show ?thesis ..
next
case False
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_out)
with A show ?thesis ..
qed
next
fix evsR4 S A U n run b d e
let
?H = "Hash (pubAK (priAK n))" and
?M = "\<lbrace>pubAK (priAK n), Crypt (priSK CA) (Hash (pubAK (priAK n)))\<rbrace>" and
?M' = "\<lbrace>pubAK e, Auth_Data (priAK n) b, pubAK (priAK n),
Crypt (priSK CA) (Hash (pubAK (priAK n)))\<rbrace>"
assume
A: "Pri_AgrK x \<in> analz (A \<union> spies evsR4) \<longrightarrow> Pri_AgrK x \<in> A"
(is "_ \<in> analz ?A \<longrightarrow> _") and
B: "(evsR4, S, A, U) \<in> protocol" and
C: "IntMapK (S (Card n, n, run)) = Some b" and
D: "Pri_AgrK x \<in> analz (insert (Crypt (sesK (d * e)) ?M')
(A \<union> spies evsR4))"
show "Pri_AgrK x \<in> A"
proof (cases "Key (invK (sesK (d * e))) \<in> analz ?A")
case True
have "Key (pubSK CA) \<in> analz ?A"
using B by (rule pr_valid_key_analz)
hence E: "analz (insert (Crypt (priSK CA) ?H) ?A) =
{Crypt (priSK CA) ?H, ?H} \<union> analz ?A"
by (simp add: analz_crypt_in analz_simp_insert_2)
have "Pri_AgrK (priAK n) \<notin> analz ?A"
using B by (rule pr_auth_key_analz)
hence F: "Pri_AgrK (priAK n) \<notin> analz (insert (Crypt (priSK CA) ?H) ?A)"
using E by simp
have "Pri_AgrK b \<notin> analz ?A"
using B and C by (rule pr_int_mapk_analz)
hence G: "Pri_AgrK b \<notin> analz (insert (Crypt (priSK CA) ?H) ?A)"
using E by simp
have "Pri_AgrK x \<in> analz (insert ?M' ?A)"
using D and True by (simp add: analz_crypt_in)
hence "Pri_AgrK x \<in> analz (insert (Auth_Data (priAK n) b) (insert ?M ?A))"
by (simp add: analz_mpair analz_simp_insert_2)
hence "Pri_AgrK x \<in> analz ?A"
proof (subst (asm) insert_commute, simp add: analz_mpair analz_simp_insert_2
del: Un_insert_right, subst (asm) insert_commute,
subst (asm) analz_auth_data_out [OF F G])
qed (simp add: E)
with A show ?thesis ..
next
case False
hence "Pri_AgrK x \<in> analz ?A"
using D by (simp add: analz_crypt_out)
with A show ?thesis ..
qed
next
fix evsFR4 S A U m n run s a b c f g
let
?H = "Hash (pubAK g)" and
?M = "\<lbrace>pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>" and
?M' = "\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>"
assume
A: "Pri_AgrK x \<in> analz (A \<union> spies evsFR4) \<longrightarrow> Pri_AgrK x \<in> A"
(is "_ \<in> analz ?A \<longrightarrow> _") and
B: "(evsFR4, S, A, U) \<in> protocol" and
C: "Crypt (sesK (c * f)) ?M' \<in> synth (analz (A \<union> spies evsFR4))" and
D: "Pri_AgrK x \<in> analz (insert (Crypt (sesK (c * f)) ?M')
(A \<union> spies evsFR4))"
have E:
"Key (invK (sesK (c * f))) \<in> analz ?A \<longrightarrow>
Pri_AgrK b \<in> analz ?A \<or> Pri_AgrK g \<in> analz ?A \<longrightarrow>
Pri_AgrK b \<in> analz ?A \<and> Pri_AgrK g \<in> analz ?A"
(is "?P \<longrightarrow> ?Q \<longrightarrow> ?R")
proof (rule impI)+
assume ?P and ?Q
have "Crypt (sesK (c * f)) ?M' \<in> analz ?A \<or>
?M' \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
using C by (rule synth_crypt)
moreover {
assume "Crypt (sesK (c * f)) ?M' \<in> analz ?A"
hence "?M' \<in> analz ?A"
using \<open>?P\<close> by (rule analz.Decrypt)
hence "\<lbrace>Auth_Data g b, pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> analz ?A"
by (rule analz.Snd)
hence F: "Auth_Data g b \<in> analz ?A"
by (rule analz.Fst)
have ?R
proof (rule disjE [OF \<open>?Q\<close>])
assume "Pri_AgrK b \<in> analz ?A"
moreover from this have "Pri_AgrK g \<in> analz ?A"
by (rule analz.Auth_Fst [OF F])
ultimately show ?R ..
next
assume "Pri_AgrK g \<in> analz ?A"
moreover from this have "Pri_AgrK b \<in> analz ?A"
by (rule analz.Auth_Snd [OF F])
ultimately show ?R
by simp
qed
}
moreover {
assume "?M' \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
hence "?M' \<in> synth (analz ?A)" ..
hence "\<lbrace>Auth_Data g b, pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "Auth_Data g b \<in> synth (analz ?A)"
by (rule synth_analz_fst)
hence "Auth_Data g b \<in> analz ?A \<or>
Pri_AgrK g \<in> analz ?A \<and> Pri_AgrK b \<in> analz ?A"
by (rule synth_auth_data)
moreover {
assume F: "Auth_Data g b \<in> analz ?A"
have ?R
proof (rule disjE [OF \<open>?Q\<close>])
assume "Pri_AgrK b \<in> analz ?A"
moreover from this have "Pri_AgrK g \<in> analz ?A"
by (rule analz.Auth_Fst [OF F])
ultimately show ?R ..
next
assume "Pri_AgrK g \<in> analz ?A"
moreover from this have "Pri_AgrK b \<in> analz ?A"
by (rule analz.Auth_Snd [OF F])
ultimately show ?R
by simp
qed
}
moreover {
assume "Pri_AgrK g \<in> analz ?A \<and> Pri_AgrK b \<in> analz ?A"
hence ?R
by simp
}
ultimately have ?R ..
}
ultimately show ?R ..
qed
show "Pri_AgrK x \<in> A"
proof (cases "Key (invK (sesK (c * f))) \<in> analz ?A")
case True
have "Key (pubSK CA) \<in> analz ?A"
using B by (rule pr_valid_key_analz)
hence F: "analz (insert (Crypt (priSK CA) ?H) ?A) =
{Crypt (priSK CA) ?H, ?H} \<union> analz ?A"
by (simp add: analz_crypt_in analz_simp_insert_2)
show ?thesis
proof (cases "Pri_AgrK g \<in> analz ?A \<or> Pri_AgrK b \<in> analz ?A", simp_all)
assume G: "Pri_AgrK g \<in> analz ?A \<or> Pri_AgrK b \<in> analz ?A"
hence H: "Pri_AgrK g \<in> analz (insert (Crypt (priSK CA) ?H) ?A) \<or>
Pri_AgrK b \<in> analz (insert (Crypt (priSK CA) ?H) ?A)"
using F by simp
have I: "Pri_AgrK b \<in> analz ?A \<and> Pri_AgrK g \<in> analz ?A"
using E and G and True by blast
hence "Pri_AgrK g \<in> analz (insert (Crypt (priSK CA) ?H) ?A)"
using F by simp
hence J: "Pri_AgrK g \<in> analz (insert (Pri_AgrK b)
(insert (Crypt (priSK CA) ?H) ?A))"
by (rule rev_subsetD, rule_tac analz_mono, blast)
have K: "Pri_AgrK b \<in> analz (insert (Crypt (priSK CA) ?H) ?A)"
using F and I by simp
have "Pri_AgrK x \<in> analz (insert ?M' ?A)"
using D and True by (simp add: analz_crypt_in)
hence "Pri_AgrK x \<in> analz (insert (Auth_Data g b) (insert ?M ?A))"
by (simp add: analz_mpair analz_simp_insert_2)
hence "Pri_AgrK x \<in> analz ?A"
proof (subst (asm) insert_commute, simp add: analz_mpair analz_simp_insert_2
del: Un_insert_right, subst (asm) insert_commute,
subst (asm) analz_auth_data_in [OF H], simp del: Un_insert_right)
assume "Pri_AgrK x \<in> analz (insert (Pri_AgrK g) (insert (Pri_AgrK b)
(insert (Crypt (priSK CA) ?H) ?A)))"
thus ?thesis
proof (subst (asm) analz_simp_insert_1 [OF J],
subst (asm) analz_simp_insert_1 [OF K])
qed (simp add: F)
qed
with A show ?thesis ..
next
assume G: "Pri_AgrK g \<notin> analz ?A \<and> Pri_AgrK b \<notin> analz ?A"
hence H: "Pri_AgrK g \<notin> analz (insert (Crypt (priSK CA) ?H) ?A)"
using F by simp
have I: "Pri_AgrK b \<notin> analz (insert (Crypt (priSK CA) ?H) ?A)"
using F and G by simp
have "Pri_AgrK x \<in> analz (insert ?M' ?A)"
using D and True by (simp add: analz_crypt_in)
hence "Pri_AgrK x \<in> analz (insert (Auth_Data g b) (insert ?M ?A))"
by (simp add: analz_mpair analz_simp_insert_2)
hence "Pri_AgrK x \<in> analz ?A"
proof (subst (asm) insert_commute, simp add: analz_mpair analz_simp_insert_2
del: Un_insert_right, subst (asm) insert_commute,
subst (asm) analz_auth_data_out [OF H I])
qed (simp add: F)
with A show ?thesis ..
qed
next
case False
hence "Pri_AgrK x \<in> analz ?A"
using D by (simp add: analz_crypt_out)
with A show ?thesis ..
qed
next
fix evsC5 A m c f
assume
A: "Pri_AgrK x \<in> analz (A \<union> spies evsC5) \<longrightarrow> Pri_AgrK x \<in> A"
(is "_ \<in> analz ?A \<longrightarrow> _") and
B: "Pri_AgrK x \<in> analz (insert (Crypt (sesK (c * f)) (Passwd m))
(A \<union> spies evsC5))"
show "Pri_AgrK x \<in> A"
proof (cases "Key (invK (sesK (c * f))) \<in> analz ?A")
case True
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_in analz_simp_insert_2)
with A show ?thesis ..
next
case False
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_out)
with A show ?thesis ..
qed
next
fix evsFC5 A n d e
assume
A: "Pri_AgrK x \<in> analz (A \<union> spies evsFC5) \<longrightarrow> Pri_AgrK x \<in> A"
(is "_ \<in> analz ?A \<longrightarrow> _") and
B: "Pri_AgrK x \<in> analz (insert (Crypt (sesK (d * e)) (Passwd n))
(A \<union> spies evsFC5))"
show "Pri_AgrK x \<in> A"
proof (cases "Key (invK (sesK (d * e))) \<in> analz ?A")
case True
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_in analz_simp_insert_2)
with A show ?thesis ..
next
case False
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_out)
with A show ?thesis ..
qed
next
fix evsR5 A d e
assume
A: "Pri_AgrK x \<in> analz (A \<union> spies evsR5) \<longrightarrow> Pri_AgrK x \<in> A"
(is "_ \<in> analz ?A \<longrightarrow> _") and
B: "Pri_AgrK x \<in> analz (insert (Crypt (sesK (d * e)) (Number 0))
(A \<union> spies evsR5))"
show "Pri_AgrK x \<in> A"
proof (cases "Key (invK (sesK (d * e))) \<in> analz ?A")
case True
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_in analz_simp_insert_2)
with A show ?thesis ..
next
case False
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_out)
with A show ?thesis ..
qed
next
fix evsFR5 A c f
assume
A: "Pri_AgrK x \<in> analz (A \<union> spies evsFR5) \<longrightarrow> Pri_AgrK x \<in> A"
(is "_ \<in> analz ?A \<longrightarrow> _") and
B: "Pri_AgrK x \<in> analz (insert (Crypt (sesK (c * f)) (Number 0))
(A \<union> spies evsFR5))"
show "Pri_AgrK x \<in> A"
proof (cases "Key (invK (sesK (c * f))) \<in> analz ?A")
case True
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_in analz_simp_insert_2)
with A show ?thesis ..
next
case False
hence "Pri_AgrK x \<in> analz ?A"
using B by (simp add: analz_crypt_out)
with A show ?thesis ..
qed
qed
lemma pr_pri_agrk_analz:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
(Pri_AgrK x \<in> analz (A \<union> spies evs)) = (Pri_AgrK x \<in> A)"
proof (rule iffI, erule pr_pri_agrk_analz_intro, assumption)
qed (rule subsetD [OF analz_subset], simp)
lemma pr_ext_agrk_user_1 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
User m \<noteq> Spy \<longrightarrow>
Says n run 4 (User m) (Card n) (Crypt (sesK K) (pubAK e)) \<in> set evs \<longrightarrow>
ExtAgrK (S (User m, n, run)) \<noteq> None"
by (erule protocol.induct, simp_all, (rule_tac [!] impI)+, simp_all)
lemma pr_ext_agrk_user_2 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
User m \<noteq> Spy \<longrightarrow>
Says n run 4 X (User m) (Crypt (sesK K)
\<lbrace>pubAK e, Auth_Data x y, pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>)
\<in> set evs \<longrightarrow>
ExtAgrK (S (User m, n, run)) \<noteq> None"
-proof (erule protocol.induct, simp_all, (rule_tac [!] impI)+, simp_all,
+using [[simproc del: defined_all]] proof (erule protocol.induct, simp_all, (rule_tac [!] impI)+, simp_all,
(erule conjE)+)
fix evs S A U n run s a b d e X
assume "(evs, S, A, U) \<in> protocol"
moreover assume "0 < m"
hence "User m \<noteq> Spy"
by simp
moreover assume A: "User m = X" and
"Says n run 4 X (Card n) (Crypt (sesK (d * e))
(pubAK (d * (s + a * b)))) \<in> set evs"
hence "Says n run 4 (User m) (Card n) (Crypt (sesK (d * e))
(pubAK (d * (s + a * b)))) \<in> set evs"
by simp
ultimately have "ExtAgrK (S (User m, n, run)) \<noteq> None"
by (rule pr_ext_agrk_user_1)
thus "\<exists>e. ExtAgrK (S (X, n, run)) = Some e"
using A by simp
qed
lemma pr_ext_agrk_user_3 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
User m \<noteq> Spy \<longrightarrow>
ExtAgrK (S (User m, n, run)) = Some e \<longrightarrow>
Says n run 4 (User m) (Card n) (Crypt (sesK K) (pubAK e')) \<in> set evs \<longrightarrow>
e' = e"
proof (erule protocol.induct, simp_all, (rule conjI, (rule_tac [!] impI)+)+,
(erule conjE)+, simp_all)
assume "agrK e' = agrK e"
with agrK_inj show "e' = e"
by (rule injD)
next
fix evsC4 S A U n run and m :: nat
assume "(evsC4, S, A, U) \<in> protocol"
moreover assume "0 < m"
hence "User m \<noteq> Spy"
by simp
moreover assume
"Says n run 4 (User m) (Card n) (Crypt (sesK K) (pubAK e')) \<in> set evsC4"
ultimately have "ExtAgrK (S (User m, n, run)) \<noteq> None"
by (rule pr_ext_agrk_user_1)
moreover assume "ExtAgrK (S (User m, n, run)) = None"
ultimately show "e' = e"
by contradiction
qed
lemma pr_ext_agrk_user_4 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
ExtAgrK (S (User m, n, run)) = Some f \<longrightarrow>
(\<exists>X. Says n run 3 X (User m) (pubAK f) \<in> set evs)"
proof (erule protocol.induct, simp_all, rule_tac [!] impI, rule_tac [1-2] impI,
rule_tac [5] impI, simp_all)
qed blast+
declare fun_upd_apply [simp del]
lemma pr_ext_agrk_user_5 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Says n run 3 X (User m) (pubAK f) \<in> set evs \<longrightarrow>
(\<exists>s a b d. f = d * (s + a * b) \<and>
NonceS (S (Card n, n, run)) = Some s \<and>
IntMapK (S (Card n, n, run)) = Some b \<and>
ExtMapK (S (Card n, n, run)) = Some a \<and>
IntAgrK (S (Card n, n, run)) = Some d \<and>
d \<noteq> 0 \<and> s + a * b \<noteq> 0) \<or>
(\<exists>b. Pri_AgrK b \<in> A \<and>
ExtMapK (S (User m, n, run)) = Some b)"
(is "_ \<Longrightarrow> ?H evs \<longrightarrow> ?P S n run \<or> ?Q S A n run")
apply (erule protocol.induct, simp_all add: pr_pri_agrk_analz)
apply (rule conjI)
apply (rule_tac [1-2] impI)+
apply (rule_tac [3] conjI, (rule_tac [3] impI)+)+
apply (rule_tac [4] impI)+
apply ((rule_tac [5] impI)+, (rule_tac [5] conjI)?)+
apply (rule_tac [6] impI)+
apply ((rule_tac [7] impI)+, (rule_tac [7] conjI)?)+
apply (rule_tac [8] impI)+
apply ((rule_tac [9] impI)+, (rule_tac [9] conjI)?)+
apply (rule_tac [10] impI)+
apply (rule_tac [11] impI)+
apply (rule_tac [12] conjI, (rule_tac [12] impI)+)+
apply (rule_tac [13] impI)+
apply (rule_tac [14] conjI, (rule_tac [14] impI)+)+
apply (erule_tac [14] conjE)+
apply (rule_tac [15] impI)+
apply ((rule_tac [16] impI)+, (rule_tac [16] conjI)?)+
apply (erule_tac [16] conjE)+
apply (rule_tac [17-18] impI)
proof -
fix evsR1 S A U s n' run'
assume
"?H evsR1 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsR1"
hence A: "?P S n run \<or> ?Q S A n run" ..
assume B: "NonceS (S (Card n', n', run')) = None"
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>NonceS := Some s\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = s \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply, rule impI, simp add: B)
qed (rule disjI2, simp add: fun_upd_apply, blast)
next
fix evsR1 S A U s n' run'
assume
"?H evsR1 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsR1"
hence A: "?P S n run \<or> ?Q S A n run" ..
assume B: "NonceS (S (Card n', n', run')) = None"
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>NonceS := Some s\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply, rule impI, simp add: B)
qed (rule disjI2, simp add: fun_upd_apply)
next
fix evsC2 S A U s a n' run'
assume
"?H evsC2 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsC2"
hence A: "?P S n run \<or> ?Q S A n run" ..
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = a \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a m' n' run'
assume
"?H evsC2 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsC2"
hence A: "?P S n run \<or> ?Q S A n run" ..
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a n' run'
assume
"?H evsC2 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsC2"
hence A: "?P S n run \<or> ?Q S A n run" ..
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = a \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a m' n' run'
assume
"?H evsC2 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsC2"
hence A: "?P S n run \<or> ?Q S A n run" ..
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a n' run'
assume
"?H evsC2 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsC2"
hence A: "?P S n run \<or> ?Q S A n run" ..
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = a \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a m' n' run'
assume
"?H evsC2 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsC2"
hence A: "?P S n run \<or> ?Q S A n run" ..
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a n' run'
assume
"?H evsC2 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsC2"
hence A: "?P S n run \<or> ?Q S A n run" ..
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = a \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a m' n' run'
assume
"?H evsC2 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsC2"
hence A: "?P S n run \<or> ?Q S A n run" ..
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsR2 S A U a b n' run'
assume
"?H evsR2 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsR2"
hence A: "?P S n run \<or> ?Q S A n run" ..
assume B: "IntMapK (S (Card n', n', run')) = None"
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntMapK := Some b, ExtMapK := Some a\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply, rule impI, simp add: B)
qed (rule disjI2, simp add: fun_upd_apply)
next
fix evsC3 S A U b c m' n' run'
assume
"?H evsC3 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsC3"
hence A: "?P S n run \<or> ?Q S A n run" ..
assume
"ExtMapK (S (User m', n', run')) = None" and
"m' = 0"
hence B: "ExtMapK (S (Spy, n', run')) = None"
by simp
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>ExtMapK := Some b, IntAgrK := Some c\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = c \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
proof (rule disjE [OF A], simp add: fun_upd_apply)
qed (rule disjI2, simp add: fun_upd_apply, rule conjI, rule impI, simp add: B, blast)
next
fix evsC3 S A U b c m' n' run'
assume
"?H evsC3 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsC3"
hence A: "?P S n run \<or> ?Q S A n run" ..
assume B: "ExtMapK (S (User m', n', run')) = None"
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>ExtMapK := Some b, IntAgrK := Some c\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
proof (rule disjE [OF A], simp add: fun_upd_apply)
qed (rule disjI2, simp add: fun_upd_apply, rule impI, simp add: B)
next
fix evsR3 A U s a b c d n' run' X and S :: state
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s + a * b))\<rparr>)"
assume "agrK f = agrK (d * (s + a * b))"
with agrK_inj have "f = d * (s + a * b)"
by (rule injD)
moreover assume
"NonceS (S (Card n', n', run')) = Some s" and
"IntMapK (S (Card n', n', run')) = Some b" and
"ExtMapK (S (Card n', n', run')) = Some a" and
"d \<noteq> 0" and
"s + a * b \<noteq> 0"
ultimately show "?P ?S n' run' \<or>
(\<exists>b. Pri_AgrK b \<in> A \<and> ExtMapK (?S (X, n', run')) = Some b)"
by (simp add: fun_upd_apply)
next
fix evsR3 S A U s a b c d n' run'
assume
"?H evsR3 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsR3"
hence A: "?P S n run \<or> ?Q S A n run" ..
assume B: "IntAgrK (S (Card n', n', run')) = None"
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s + a * b))\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply, rule impI, simp add: B)
qed (rule disjI2, simp add: fun_upd_apply)
next
fix evsR3 A U s s' a b c d n' run' X and S :: state
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s' + a * b))\<rparr>)"
assume "agrK f = agrK (d * (s + a * b))"
with agrK_inj have "f = d * (s + a * b)"
by (rule injD)
moreover assume
"NonceS (S (Card n', n', run')) = Some s" and
"IntMapK (S (Card n', n', run')) = Some b" and
"ExtMapK (S (Card n', n', run')) = Some a" and
"d \<noteq> 0" and
"s + a * b \<noteq> 0"
ultimately show "?P ?S n' run' \<or>
(\<exists>b. Pri_AgrK b \<in> A \<and> ExtMapK (?S (X, n', run')) = Some b)"
by (simp add: fun_upd_apply)
next
fix evsR3 S A U s a b c d n' run'
assume
"?H evsR3 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsR3"
hence A: "?P S n run \<or> ?Q S A n run" ..
assume B: "IntAgrK (S (Card n', n', run')) = None"
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s + a * b))\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply, rule impI, simp add: B)
qed (rule disjI2, simp add: fun_upd_apply)
next
fix evsC4 S A U f m' n' run'
assume
"?H evsC4 \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H evsC4"
hence A: "?P S n run \<or> ?Q S A n run" ..
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>ExtAgrK := Some f\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
qed
declare fun_upd_apply [simp]
lemma pr_int_agrk_user_1 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
IntAgrK (S (User m, n, run)) = Some c \<longrightarrow>
Pri_AgrK c \<in> U"
by (erule protocol.induct, simp_all)
lemma pr_int_agrk_user_2 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
User m \<noteq> Spy \<longrightarrow>
IntAgrK (S (User m, n, run)) = Some c \<longrightarrow>
Pri_AgrK c \<notin> A"
proof (erule protocol.induct, simp_all, rule_tac [3] conjI, (rule_tac [!] impI)+,
rule_tac [2] impI, rule_tac [4] impI, rule_tac [!] notI, simp_all)
fix evsR1 S A U s
assume
"(evsR1, S, A, U) \<in> protocol" and
"IntAgrK (S (User m, n, run)) = Some s"
hence "Pri_AgrK s \<in> U"
by (rule pr_int_agrk_user_1)
moreover assume "Pri_AgrK s \<notin> U"
ultimately show False
by contradiction
next
fix evsC2 S A U a
assume
"(evsC2, S, A, U) \<in> protocol" and
"IntAgrK (S (User m, n, run)) = Some a"
hence "Pri_AgrK a \<in> U"
by (rule pr_int_agrk_user_1)
moreover assume "Pri_AgrK a \<notin> U"
ultimately show False
by contradiction
next
fix evsC3 S A U
assume "(evsC3, S, A, U) \<in> protocol"
hence "A \<subseteq> U"
by (rule pr_analz_used)
moreover assume "Pri_AgrK c \<in> A"
ultimately have "Pri_AgrK c \<in> U" ..
moreover assume "Pri_AgrK c \<notin> U"
ultimately show False
by contradiction
next
fix evsC3 S A U c'
assume
"(evsC3, S, A, U) \<in> protocol" and
"IntAgrK (S (User m, n, run)) = Some c'"
hence "Pri_AgrK c' \<in> U"
by (rule pr_int_agrk_user_1)
moreover assume "Pri_AgrK c' \<notin> U"
ultimately show False
by contradiction
qed
lemma pr_int_agrk_user_3 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
NonceS (S (User m, n, run)) = Some s \<longrightarrow>
IntMapK (S (User m, n, run)) = Some a \<longrightarrow>
ExtMapK (S (User m, n, run)) = Some b \<longrightarrow>
IntAgrK (S (User m, n, run)) = Some c \<longrightarrow>
c * (s + a * b) \<noteq> 0"
proof (erule protocol.induct, simp_all, rule conjI, (rule_tac [1-2] impI)+,
(rule_tac [3] impI)+, simp_all)
fix evsC2 S A U n run m
assume A: "(evsC2, S, A, U) \<in> protocol"
moreover assume "NonceS (S (User m, n, run)) = None"
ultimately have "IntMapK (S (User m, n, run)) = None"
by (rule pr_state_1)
with A have "ExtMapK (S (User m, n, run)) = None"
by (rule pr_state_2)
moreover assume "ExtMapK (S (User m, n, run)) = Some b"
ultimately show "c \<noteq> 0 \<and> s + a * b \<noteq> 0"
by simp
next
fix evsC2 S A U n run m
assume A: "(evsC2, S, A, U) \<in> protocol"
moreover assume "NonceS (S (User m, n, run)) = None"
ultimately have "IntMapK (S (User m, n, run)) = None"
by (rule pr_state_1)
with A have "ExtMapK (S (User m, n, run)) = None"
by (rule pr_state_2)
moreover assume "ExtMapK (S (User m, n, run)) = Some b"
ultimately show "c \<noteq> 0 \<and> a \<noteq> 0 \<and> b \<noteq> 0"
by simp
qed
lemma pr_int_agrk_card [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
NonceS (S (Card n, n, run)) = Some s \<longrightarrow>
IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
ExtMapK (S (Card n, n, run)) = Some a \<longrightarrow>
IntAgrK (S (Card n, n, run)) = Some d \<longrightarrow>
d * (s + a * b) \<noteq> 0"
proof (erule protocol.induct, simp_all, (rule_tac [!] impI)+, simp_all)
fix evsR1 S A U n run
assume
"(evsR1, S, A, U) \<in> protocol" and
"NonceS (S (Card n, n, run)) = None"
hence "IntMapK (S (Card n, n, run)) = None"
by (rule pr_state_1)
moreover assume "IntMapK (S (Card n, n, run)) = Some b"
ultimately show "d \<noteq> 0 \<and> s + a * b \<noteq> 0"
by simp
next
fix evsR2 S A U n run
assume A: "(evsR2, S, A, U) \<in> protocol" and
"IntMapK (S (Card n, n, run)) = None"
hence "ExtMapK (S (Card n, n, run)) = None"
by (rule pr_state_2)
with A have "IntAgrK (S (Card n, n, run)) = None"
by (rule pr_state_3)
moreover assume "IntAgrK (S (Card n, n, run)) = Some d"
ultimately show "d \<noteq> 0 \<and> s + a * b \<noteq> 0"
by simp
qed
lemma pr_ext_agrk_card [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
NonceS (S (Card n, n, run)) = Some s \<longrightarrow>
IntMapK (S (Card n, n, run)) = Some b \<longrightarrow>
ExtMapK (S (Card n, n, run)) = Some a \<longrightarrow>
IntAgrK (S (Card n, n, run)) = Some d \<longrightarrow>
ExtAgrK (S (Card n, n, run)) = Some (c * (s + a * b)) \<longrightarrow>
Pri_AgrK c \<notin> A \<longrightarrow>
Key (sesK (c * d * (s + a * b))) \<notin> A"
apply (erule protocol.induct, simp_all add: pr_pri_agrk_analz)
apply (rule conjI)
apply (rule_tac [1-2] impI)+
apply (rule_tac [3] impI)+
apply (rule_tac [4] conjI, (rule_tac [4] impI)+)+
apply (erule_tac [4] conjE)+
apply (rule_tac [5] impI)
apply (erule_tac [5] conjE)+
apply (rule_tac [6] impI)+
apply (erule_tac [6] conjE)+
apply (rule_tac [6] notI)
apply ((rule_tac [7] impI)+, (rule_tac [7] conjI)?)+
apply (erule_tac [7] conjE)+
apply (rule_tac [8] impI)
apply (erule_tac [8] conjE)+
apply (rule_tac [9] impI)+
apply (rule_tac [9] notI)
apply (rule_tac [10] impI)+
apply (rule_tac [11] impI)+
apply (rule_tac [11] notI)
proof simp_all
fix evsR1 S A U n run
assume
"(evsR1, S, A, U) \<in> protocol" and
"NonceS (S (Card n, n, run)) = None"
hence "IntMapK (S (Card n, n, run)) = None"
by (rule pr_state_1)
moreover assume "IntMapK (S (Card n, n, run)) = Some b"
ultimately show "Key (sesK (c * d * (s + a * b))) \<notin> A"
by simp
next
fix evsR1 S A U n run
assume
"(evsR1, S, A, U) \<in> protocol" and
"NonceS (S (Card n, n, run)) = None"
hence "IntMapK (S (Card n, n, run)) = None"
by (rule pr_state_1)
moreover assume "IntMapK (S (Card n, n, run)) = Some b"
ultimately show "Key (sesK (c * d * (s + a * b))) \<notin> A"
by simp
next
fix evsR2 S A U n run
assume A: "(evsR2, S, A, U) \<in> protocol" and
"IntMapK (S (Card n, n, run)) = None"
hence "ExtMapK (S (Card n, n, run)) = None"
by (rule pr_state_2)
with A have "IntAgrK (S (Card n, n, run)) = None"
by (rule pr_state_3)
moreover assume "IntAgrK (S (Card n, n, run)) = Some d"
ultimately show "Key (sesK (c * d * (s + a * b))) \<notin> A"
by simp
next
fix evsR3 S A U s a' b' c' d'
assume
"(evsR3, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some (c * (s + a * b))"
hence "Key (sesK (d * (c * (s + a * b)))) \<in> U"
by (rule pr_sesk_card)
moreover have "d * (c * (s + a * b)) = c * d * (s + a * b)"
by simp
ultimately have "Key (sesK (c * d * (s + a * b))) \<in> U"
by simp
moreover assume "sesK (c * d * (s + a * b)) = sesK (c' * d' * (s + a' * b'))"
ultimately have "Key (sesK (c' * d' * (s + a' * b'))) \<in> U"
by simp
moreover assume "Key (sesK (c' * d' * (s + a' * b'))) \<notin> U"
ultimately show False
by contradiction
next
fix evsR3 S A U s' a' b' c' d'
assume
"(evsR3, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some (c * (s + a * b))"
hence "Key (sesK (d * (c * (s + a * b)))) \<in> U"
by (rule pr_sesk_card)
moreover have "d * (c * (s + a * b)) = c * d * (s + a * b)"
by simp
ultimately have "Key (sesK (c * d * (s + a * b))) \<in> U"
by simp
moreover assume "sesK (c * d * (s + a * b)) = sesK (c' * d' * (s' + a' * b'))"
ultimately have "Key (sesK (c' * d' * (s' + a' * b'))) \<in> U"
by simp
moreover assume "Key (sesK (c' * d' * (s' + a' * b'))) \<notin> U"
ultimately show False
by contradiction
next
fix evsR3 S A U s' c'
assume "(evsR3, S, A, U) \<in> protocol"
hence "A \<subseteq> U"
by (rule pr_analz_used)
moreover assume "Key (sesK (c' * d * (s' + a * b))) \<notin> U"
ultimately have "Key (sesK (c' * d * (s' + a * b))) \<notin> A"
by (rule contra_subsetD)
moreover assume "c' * (s' + a * b) = c * (s + a * b)"
hence "c' * d * (s' + a * b) = c * d * (s + a * b)"
by simp
ultimately show "Key (sesK (c * d * (s + a * b))) \<notin> A"
by simp
next
fix evsFR3 S A U s' a' b' c' d'
assume
"(evsFR3, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some (c * (s + a * b))"
hence "Key (sesK (d * (c * (s + a * b)))) \<in> U"
by (rule pr_sesk_card)
moreover have "d * (c * (s + a * b)) = c * d * (s + a * b)"
by simp
ultimately have "Key (sesK (c * d * (s + a * b))) \<in> U"
by simp
moreover assume "sesK (c * d * (s + a * b)) = sesK (c' * d' * (s' + a' * b'))"
ultimately have "Key (sesK (c' * d' * (s' + a' * b'))) \<in> U"
by simp
moreover assume "Key (sesK (c' * d' * (s' + a' * b'))) \<notin> U"
ultimately show False
by contradiction
qed
declare fun_upd_apply [simp del]
lemma pr_sesk_user_3 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
\<lbrace>Key (sesK K), Agent (User m), Number n, Number run\<rbrace> \<in> U \<longrightarrow>
Key (sesK K) \<in> A \<longrightarrow>
(\<exists>d e. K = d * e \<and>
IntAgrK (S (Card n, n, run)) = Some d \<and>
ExtAgrK (S (Card n, n, run)) = Some e) \<or>
(\<exists>b. Pri_AgrK b \<in> A \<and>
ExtMapK (S (User m, n, run)) = Some b)"
(is "_ \<Longrightarrow> ?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run")
apply (erule protocol.induct, simp_all add: pr_pri_agrk_analz, blast)
apply (rule conjI)
apply (rule_tac [1-2] impI)+
apply (rule_tac [3] conjI, (rule_tac [3] impI)+)+
apply (rule_tac [4] impI)+
apply ((rule_tac [5] impI)+, (rule_tac [5] conjI)?)+
apply (rule_tac [6] impI)+
apply ((rule_tac [7] impI)+, (rule_tac [7] conjI)?)+
apply (rule_tac [8] impI)+
apply ((rule_tac [9] impI)+, (rule_tac [9] conjI)?)+
apply (rule_tac [10] impI)+
apply (rule_tac [11] impI)+
apply (rule_tac [12] conjI, (rule_tac [12] impI)+)+
apply (rule_tac [13] impI)+
apply (rule_tac [14] conjI, (rule_tac [14] impI)+)+
apply (erule_tac [14] conjE)+
apply ((rule_tac [15] impI)+, (rule_tac [15] conjI)?)+
apply (rule_tac [16] impI)+
apply ((rule_tac [17] impI)+, (rule_tac [17] conjI)?)+
apply (rule_tac [18-20] impI)+
proof -
fix evsR1 S A U s n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>NonceS := Some s\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = s \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply, rule impI, simp)
qed (rule disjI2, simp add: fun_upd_apply, blast)
next
fix evsR1 S A U s n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>NonceS := Some s\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply, rule impI, simp)
qed (rule disjI2, simp add: fun_upd_apply)
next
fix evsC2 S A U s a n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = a \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a m n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((User m, n', run') := S (User m, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = a \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a m n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((User m, n', run') := S (User m, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = a \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a m n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((User m, n', run') := S (User m, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = a \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsC2 S A U s a m n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((User m, n', run') := S (User m, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
next
fix evsR2 S A U a b n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntMapK := Some b, ExtMapK := Some a\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply, rule impI, simp)
qed (rule disjI2, simp add: fun_upd_apply)
next
fix evsC3 S A U b c m' n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
assume
"ExtMapK (S (User m', n', run')) = None" and
"m' = 0"
hence B: "ExtMapK (S (Spy, n', run')) = None"
by simp
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>ExtMapK := Some b, IntAgrK := Some c\<rparr>)"
show "?P ?S n run \<or>
(\<exists>b. (b = c \<or> Pri_AgrK b \<in> A) \<and> ExtMapK (?S (User m, n, run)) = Some b)"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply)
qed (rule disjI2, simp add: fun_upd_apply, rule conjI, rule impI, simp add: B, blast)
next
fix evsC3 S A U b c m' n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
assume B: "ExtMapK (S (User m', n', run')) = None"
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>ExtMapK := Some b, IntAgrK := Some c\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply)
qed (rule disjI2, simp add: fun_upd_apply, rule impI, simp add: B)
next
fix evsR3 A s a b c d n' run' X and S :: state
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s + a * b))\<rparr>)"
assume "sesK K = sesK (c * d * (s + a * b))"
with sesK_inj have "K = c * d * (s + a * b)"
by (rule injD)
thus "?P ?S n' run' \<or>
(\<exists>b. Pri_AgrK b \<in> A \<and> ExtMapK (?S (X, n', run')) = Some b)"
by (simp add: fun_upd_apply)
next
fix evsR3 A U s a b c d n' run' and S :: state
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s + a * b))\<rparr>)"
assume "sesK K = sesK (c * d * (s + a * b))"
with sesK_inj have "K = c * d * (s + a * b)"
by (rule injD)
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately have "Key (sesK K) \<notin> U"
by simp
moreover assume
"(evsR3, S, A, U) \<in> protocol" and
"\<lbrace>Key (sesK K), Agent (User m), Number n, Number run\<rbrace> \<in> U"
hence "Key (sesK K) \<in> U"
by (rule pr_sesk_user_2)
ultimately show "?P ?S n run \<or> ?Q ?S A n run"
by contradiction
next
fix evsR3 S A U s a b c d n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
assume B: "IntAgrK (S (Card n', n', run')) = None"
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s + a * b))\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply, rule impI, simp add: B)
qed (rule disjI2, simp add: fun_upd_apply)
next
fix evsR3 A U s s' a b c d n' run' X and S :: state
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s' + a * b))\<rparr>)"
assume "(evsR3, S, A, U) \<in> protocol"
hence "A \<subseteq> U"
by (rule pr_analz_used)
moreover assume "Key (sesK (c * d * (s + a * b))) \<in> A"
ultimately have "Key (sesK (c * d * (s + a * b))) \<in> U" ..
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show "?P ?S n' run' \<or>
(\<exists>b. Pri_AgrK b \<in> A \<and> ExtMapK (?S (X, n', run')) = Some b)"
by contradiction
next
fix evsR3 S A U s a b c d n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
assume B: "IntAgrK (S (Card n', n', run')) = None"
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s + a * b))\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
proof (rule disjE [OF A], rule disjI1, simp add: fun_upd_apply, rule impI, simp add: B)
qed (rule disjI2, simp add: fun_upd_apply)
next
fix evsFR3 A U s a b c d n' run' and S :: state
assume "sesK K = sesK (c * d * (s + a * b))"
with sesK_inj have "K = c * d * (s + a * b)"
by (rule injD)
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately have "Key (sesK K) \<notin> U"
by simp
moreover assume
"(evsFR3, S, A, U) \<in> protocol" and
"\<lbrace>Key (sesK K), Agent (User m), Number n, Number run\<rbrace> \<in> U"
hence "Key (sesK K) \<in> U"
by (rule pr_sesk_user_2)
ultimately show "?P S n run \<or> ?Q S A n run"
by contradiction
next
fix evsC4 S A U f m' n' run'
assume
"?H1 U \<longrightarrow> ?H2 A \<longrightarrow> ?P S n run \<or> ?Q S A n run" and
"?H1 U" and
"?H2 A"
hence A: "?P S n run \<or> ?Q S A n run"
by simp
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>ExtAgrK := Some f\<rparr>)"
show "?P ?S n run \<or> ?Q ?S A n run"
by (rule disjE [OF A], simp_all add: fun_upd_apply, blast)
qed
declare fun_upd_apply [simp]
lemma pr_sesk_auth [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Crypt (sesK K) \<lbrace>pubAK e, Auth_Data x y, pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> parts (A \<union> spies evs) \<longrightarrow>
Key (sesK K) \<in> U"
proof (erule protocol.induct, subst parts_simp, (simp, blast)+,
simp_all add: parts_simp_insert parts_auth_data parts_crypt parts_mpair,
rule_tac [!] impI)
fix evsR4 S A U n run d e
assume
"(evsR4, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some e"
thus "Key (sesK (d * e)) \<in> U"
by (rule pr_sesk_card)
next
fix evsFR4 S A U m n run c f
assume A: "(evsFR4, S, A, U) \<in> protocol" and
"IntAgrK (S (User m, n, run)) = Some c" and
"ExtAgrK (S (User m, n, run)) = Some f"
hence "\<lbrace>Key (sesK (c * f)), Agent (User m), Number n, Number run\<rbrace> \<in> U"
by (rule pr_sesk_user_1)
with A show "Key (sesK (c * f)) \<in> U"
by (rule pr_sesk_user_2)
qed
lemma pr_sesk_passwd [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Says n run 5 X (Card n) (Crypt (sesK K) (Passwd m)) \<in> set evs \<longrightarrow>
Key (sesK K) \<in> U"
proof (erule protocol.induct, simp_all, rule_tac [!] impI)
fix evsC5 S A U m n run s a b c f g X
assume "(evsC5, S, A, U) \<in> protocol"
moreover assume "Says n run 4 X (User m) (Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>) \<in> set evsC5"
(is "Says _ _ _ _ _ ?M \<in> _")
hence "?M \<in> spies evsC5"
by (rule set_spies)
hence "?M \<in> A \<union> spies evsC5" ..
hence "?M \<in> parts (A \<union> spies evsC5)"
by (rule parts.Inj)
ultimately show "Key (sesK (c * f)) \<in> U"
by (rule pr_sesk_auth)
next
fix evsFC5 S A U n run d e
assume
"(evsFC5, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some e"
thus "Key (sesK (d * e)) \<in> U"
by (rule pr_sesk_card)
qed
lemma pr_sesk_card_user [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
User m \<noteq> Spy \<longrightarrow>
NonceS (S (User m, n, run)) = Some s \<longrightarrow>
IntMapK (S (User m, n, run)) = Some a \<longrightarrow>
ExtMapK (S (User m, n, run)) = Some b \<longrightarrow>
IntAgrK (S (User m, n, run)) = Some c \<longrightarrow>
NonceS (S (Card n, n, run)) = Some s' \<longrightarrow>
IntMapK (S (Card n, n, run)) = Some b' \<longrightarrow>
ExtMapK (S (Card n, n, run)) = Some a' \<longrightarrow>
IntAgrK (S (Card n, n, run)) = Some d \<longrightarrow>
ExtAgrK (S (Card n, n, run)) = Some (c * (s + a * b)) \<longrightarrow>
s' + a' * b' = s + a * b \<longrightarrow>
Key (sesK (c * d * (s + a * b))) \<notin> A"
apply (erule protocol.induct, rule_tac [!] impI, simp_all add: pr_pri_agrk_analz)
apply (rule impI)+
apply (rule_tac [2] impI)
apply (rule_tac [2] conjI)
apply (rule_tac [2-3] impI)+
apply (rule_tac [4] impI)+
apply (rule_tac [5] impI)+
apply (rule_tac [6] conjI, (rule_tac [6] impI)+)+
apply (rule_tac [6] conjI)
apply (erule_tac [6] conjE)+
apply (rule_tac [8] impI)+
apply (rule_tac [8] notI)
apply (rule_tac [9] impI, rule_tac [9] conjI)+
apply (rule_tac [9] impI)+
apply (rule_tac [10] impI)+
apply (rule_tac [10] notI)
apply (rule_tac [11] impI)+
apply (rule_tac [12] impI)+
apply (rule_tac [12] notI)
proof simp_all
fix evsR1 S A U n run
assume "(evsR1, S, A, U) \<in> protocol" and
"NonceS (S (Card n, n, run)) = None"
hence "IntMapK (S (Card n, n, run)) = None"
by (rule pr_state_1)
moreover assume "IntMapK (S (Card n, n, run)) = Some b'"
ultimately show "Key (sesK (c * d * (s + a * b))) \<notin> A"
by simp
next
fix evsC2 S A U m n run
assume A: "(evsC2, S, A, U) \<in> protocol"
moreover assume "NonceS (S (User m, n, run)) = None"
ultimately have "IntMapK (S (User m, n, run)) = None"
by (rule pr_state_1)
with A have "ExtMapK (S (User m, n, run)) = None"
by (rule pr_state_2)
moreover assume "ExtMapK (S (User m, n, run)) = Some b"
ultimately show "Key (sesK (c * d * (a * b))) \<notin> A"
by simp
next
fix evsC2 S A U m n run
assume A: "(evsC2, S, A, U) \<in> protocol"
moreover assume "NonceS (S (User m, n, run)) = None"
ultimately have "IntMapK (S (User m, n, run)) = None"
by (rule pr_state_1)
with A have "ExtMapK (S (User m, n, run)) = None"
by (rule pr_state_2)
moreover assume "ExtMapK (S (User m, n, run)) = Some b"
ultimately show "Key (sesK (c * d * (s + a * b))) \<notin> A"
by simp
next
fix evsR2 S A U n run
assume A: "(evsR2, S, A, U) \<in> protocol" and
"IntMapK (S (Card n, n, run)) = None"
hence "ExtMapK (S (Card n, n, run)) = None"
by (rule pr_state_2)
with A have "IntAgrK (S (Card n, n, run)) = None"
by (rule pr_state_3)
moreover assume "IntAgrK (S (Card n, n, run)) = Some d"
ultimately show "Key (sesK (c * d * (s + a * b))) \<notin> A"
by simp
next
fix evsC3 S A U n run
assume A: "(evsC3, S, A, U) \<in> protocol" and
"NonceS (S (Card n, n, run)) = Some s'" and
"IntMapK (S (Card n, n, run)) = Some b'" and
"ExtMapK (S (Card n, n, run)) = Some a'" and
"IntAgrK (S (Card n, n, run)) = Some d"
moreover assume B: "s' + a' * b' = s + a * b" and
"ExtAgrK (S (Card n, n, run)) = Some (c * (s + a * b))"
hence "ExtAgrK (S (Card n, n, run)) = Some (c * (s' + a' * b'))"
by simp
moreover assume C: "Pri_AgrK c \<notin> U"
have "A \<subseteq> U"
using A by (rule pr_analz_used)
hence "Pri_AgrK c \<notin> A"
using C by (rule contra_subsetD)
ultimately have "Key (sesK (c * d * (s' + a' * b'))) \<notin> A"
by (rule pr_ext_agrk_card)
thus "Key (sesK (c * d * (s + a * b))) \<notin> A"
using B by simp
next
fix evsR3 S A U n run
assume "(evsR3, S, A, U) \<in> protocol"
moreover assume "0 < m"
hence "User m \<noteq> Spy"
by simp
moreover assume "IntAgrK (S (User m, n, run)) = Some c"
ultimately have "Pri_AgrK c \<notin> A"
by (rule pr_int_agrk_user_2)
moreover assume "Pri_AgrK c \<in> A"
ultimately show False
by contradiction
next
fix evsR3 S A U
assume "(evsR3, S, A, U) \<in> protocol"
hence "A \<subseteq> U"
by (rule pr_analz_used)
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show "Key (sesK (c * d * (s + a * b))) \<notin> A"
by (rule contra_subsetD)
next
fix evsR3 S A U s' a' b' c' d'
assume
"(evsR3, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some (c * (s + a * b))"
hence "Key (sesK (d * (c * (s + a * b)))) \<in> U"
by (rule pr_sesk_card)
moreover have "d * (c * (s + a * b)) = c * d * (s + a * b)"
by simp
ultimately have "Key (sesK (c * d * (s + a * b))) \<in> U"
by simp
moreover assume "sesK (c * d * (s + a * b)) = sesK (c' * d' * (s' + a' * b'))"
ultimately have "Key (sesK (c' * d' * (s' + a' * b'))) \<in> U"
by simp
moreover assume "Key (sesK (c' * d' * (s' + a' * b'))) \<notin> U"
ultimately show False
by simp
next
fix evsR3 S A U s' a' b' c' d'
assume
"(evsR3, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some (c * (s + a * b))"
hence "Key (sesK (d * (c * (s + a * b)))) \<in> U"
by (rule pr_sesk_card)
moreover have "d * (c * (s + a * b)) = c * d * (s + a * b)"
by simp
ultimately have "Key (sesK (c * d * (s + a * b))) \<in> U"
by simp
moreover assume "sesK (c * d * (s + a * b)) = sesK (c' * d' * (s' + a' * b'))"
ultimately have "Key (sesK (c' * d' * (s' + a' * b'))) \<in> U"
by simp
moreover assume "Key (sesK (c' * d' * (s' + a' * b'))) \<notin> U"
ultimately show False
by simp
next
fix evsR3 S A U s' c'
assume "(evsR3, S, A, U) \<in> protocol"
hence "A \<subseteq> U"
by (rule pr_analz_used)
moreover assume "Key (sesK (c' * d * (s' + a' * b'))) \<notin> U"
ultimately have "Key (sesK (c' * d * (s' + a' * b'))) \<notin> A"
by (rule contra_subsetD)
moreover assume "c' * (s' + a' * b') = c * (s + a * b)"
hence "c' * d * (s' + a' * b') = c * d * (s + a * b)"
by simp
ultimately show "Key (sesK (c * d * (s + a * b))) \<notin> A"
by simp
next
fix evsFR3 S A U s' a' b' c' d'
assume
"(evsFR3, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some (c * (s + a * b))"
hence "Key (sesK (d * (c * (s + a * b)))) \<in> U"
by (rule pr_sesk_card)
moreover have "d * (c * (s + a * b)) = c * d * (s + a * b)"
by simp
ultimately have "Key (sesK (c * d * (s + a * b))) \<in> U"
by simp
moreover assume "sesK (c * d * (s + a * b)) = sesK (c' * d' * (s' + a' * b'))"
ultimately have "Key (sesK (c' * d' * (s' + a' * b'))) \<in> U"
by simp
moreover assume "Key (sesK (c' * d' * (s' + a' * b'))) \<notin> U"
ultimately show False
by simp
qed
lemma pr_sign_key_used:
"(evs, S, A, U) \<in> protocol \<Longrightarrow> Key (priSK X) \<in> U"
by (erule protocol.induct, simp_all)
lemma pr_sign_key_analz:
"(evs, S, A, U) \<in> protocol \<Longrightarrow> Key (priSK X) \<notin> analz (A \<union> spies evs)"
proof (simp add: pr_key_analz, erule protocol.induct,
auto simp add: priSK_pubSK priSK_symK)
fix evsR3 S A U s a b c d
assume "(evsR3, S, A, U) \<in> protocol"
hence "Key (priSK X) \<in> U"
by (rule pr_sign_key_used)
moreover assume "priSK X = sesK (c * d * (s + a * b))"
ultimately have "Key (sesK (c * d * (s + a * b))) \<in> U"
by simp
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show False
by contradiction
next
fix evsFR3 S A U s a b c d
assume "(evsFR3, S, A, U) \<in> protocol"
hence "Key (priSK X) \<in> U"
by (rule pr_sign_key_used)
moreover assume "priSK X = sesK (c * d * (s + a * b))"
ultimately have "Key (sesK (c * d * (s + a * b))) \<in> U"
by simp
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show False
by contradiction
qed
lemma pr_auth_data_parts [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Auth_Data (priAK n) b \<in> parts (A \<union> spies evs) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)"
(is "_ \<Longrightarrow> ?M \<in> _ \<longrightarrow> _")
apply (erule protocol.induct, simp, subst parts_simp, simp, blast+, simp_all
add: parts_simp_insert parts_auth_data parts_crypt parts_mpair del: fun_upd_apply)
apply (rule impI)
apply ((rule_tac [2] conjI)?, rule_tac [2] impI)+
apply (rule_tac [3] impI)+
apply (rule_tac [4] impI, (rule_tac [4] conjI)?)+
apply (rule_tac [5] impI)+
apply (rule_tac [6] impI, (rule_tac [6] conjI)?)+
apply (rule_tac [7] impI)+
apply (rule_tac [8] impI, (rule_tac [8] conjI)?)+
apply (rule_tac [9] impI)+
apply (rule_tac [10] impI)
apply (rule_tac [11] conjI)
apply (rule_tac [11-12] impI)+
apply (rule_tac [13] conjI)
apply (rule_tac [13-14] impI)+
apply (rule_tac [15-17] impI)
apply (erule_tac [17] conjE)
proof -
fix evsR1 A n' run' s and S :: state
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>NonceS := Some s\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsR1) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsR1)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp, blast)
next
fix evsC2 A n' run' s a and S :: state
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsC2) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsC2)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp)
next
fix evsC2 A n' run' m' s a and S :: state
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsC2) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsC2)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp)
next
fix evsC2 A n' run' s a and S :: state
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsC2) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsC2)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp)
next
fix evsC2 A n' run' m' a and S :: state
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>NonceS := Some 0, IntMapK := Some a\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsC2) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsC2)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp)
next
fix evsC2 A n' run' s a and S :: state
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsC2) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsC2)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp)
next
fix evsC2 A n' run' m' s a and S :: state
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>NonceS := Some s, IntMapK := Some a\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsC2) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsC2)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp)
next
fix evsC2 A n' run' a and S :: state
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>NonceS := Some 0, IntMapK := Some a\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsC2) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsC2)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp)
next
fix evsC2 A n' run' m' a and S :: state
let ?S = "S((User m', n', run') := S (User m', n', run')
\<lparr>NonceS := Some 0, IntMapK := Some a\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsC2) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsC2)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp)
next
fix evsR2 A n' run' a b' and S :: state
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntMapK := Some b', ExtMapK := Some a\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsR2) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsR2)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
moreover assume "IntMapK (S (Card n', n', run')) = None"
ultimately show "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
proof (rule_tac x = m in exI, rule_tac x = run in exI, simp)
qed (rule impI, simp)
next
fix evsC3 A n' run' b' c and S :: state
let ?S = "S((Spy, n', run') := S (Spy, n', run')
\<lparr>ExtMapK := Some b', IntAgrK := Some c\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsC3) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsC3)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp)
next
fix evsC3 A n' run' b' c m and S :: state
let ?S = "S((User m, n', run') := S (User m, n', run')
\<lparr>ExtMapK := Some b', IntAgrK := Some c\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsC3) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsC3)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp)
next
fix evsR3 A n' run' s a b' c d and S :: state
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s + a * b'))\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsR3) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsR3)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp, blast)
next
fix evsR3 A n' run' s a b' c d and S :: state
let ?S = "S((Card n', n', run') := S (Card n', n', run')
\<lparr>IntAgrK := Some d, ExtAgrK := Some (c * (s + a * b'))\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsR3) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsR3)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp, blast)
next
fix evsC4 A n' run' c f m and S :: state
let ?S = "S((User m, n', run') := S (User m, n', run')\<lparr>ExtAgrK := Some f\<rparr>)"
assume
"?M \<in> parts (A \<union> spies evsC4) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
"?M \<in> parts (A \<union> spies evsC4)"
hence "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b" ..
then obtain m and run where "IntMapK (S (Card m, m, run)) = Some b"
by blast
thus "\<exists>m run. IntMapK (?S (Card m, m, run)) = Some b"
by (rule_tac x = m in exI, rule_tac x = run in exI, simp)
next
fix evsR4 n' run' b' and S :: state
assume "IntMapK (S (Card n', n', run')) = Some b'"
thus "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b'"
by blast
next
fix evsFR4 A U s a b' c f g and S :: state
assume
A: "?M \<in> parts (A \<union> spies evsFR4) \<longrightarrow>
(\<exists>m run. IntMapK (S (Card m, m, run)) = Some b)" and
B: "(evsFR4, S, A, U) \<in> protocol" and
C: "Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b')), Auth_Data g b', pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace> \<in> synth (analz (A \<union> spies evsFR4))"
(is "Crypt _ ?M' \<in> synth (analz ?A)") and
D: "priAK n = g" and
E: "b = b'"
show "\<exists>m run. IntMapK (S (Card m, m, run)) = Some b'"
proof -
have "Crypt (sesK (c * f)) ?M' \<in> analz ?A \<or>
?M' \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
using C by (rule synth_crypt)
moreover {
assume "Crypt (sesK (c * f)) ?M' \<in> analz ?A"
hence "Crypt (sesK (c * f)) ?M' \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
hence "?M' \<in> parts ?A"
by (rule parts.Body)
hence "\<lbrace>Auth_Data g b', pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> parts ?A"
by (rule parts.Snd)
hence "Auth_Data g b' \<in> parts ?A"
by (rule parts.Fst)
}
moreover {
assume "?M' \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
hence "?M' \<in> synth (analz ?A)" ..
hence "\<lbrace>Auth_Data g b', pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "Auth_Data g b' \<in> synth (analz ?A)"
by (rule synth_analz_fst)
hence "Auth_Data g b' \<in> analz ?A \<or>
Pri_AgrK g \<in> analz ?A \<and> Pri_AgrK b' \<in> analz ?A"
by (rule synth_auth_data)
moreover {
assume "Auth_Data g b' \<in> analz ?A"
hence "Auth_Data g b' \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
}
moreover {
assume "Pri_AgrK g \<in> analz ?A \<and> Pri_AgrK b' \<in> analz ?A"
hence "Pri_AgrK g \<in> analz ?A" ..
hence "Pri_AgrK (priAK n) \<in> analz ?A"
using D by simp
moreover have "Pri_AgrK (priAK n) \<notin> analz ?A"
using B by (rule pr_auth_key_analz)
ultimately have "Auth_Data g b' \<in> parts ?A"
by contradiction
}
ultimately have "Auth_Data g b' \<in> parts ?A" ..
}
ultimately have "Auth_Data g b' \<in> parts ?A" ..
thus ?thesis
using A and D and E by simp
qed
qed
lemma pr_sign_parts [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Crypt (priSK CA) (Hash (pubAK g)) \<in> parts (A \<union> spies evs) \<longrightarrow>
(\<exists>n. g = priAK n)"
(is "_ \<Longrightarrow> ?M g \<in> _ \<longrightarrow> _")
proof (erule protocol.induct, subst parts_simp, (simp, blast)+, simp_all add:
parts_simp_insert parts_auth_data parts_crypt parts_mpair, rule_tac [!] impI)
fix evsR4 n
assume "agrK g = agrK (priAK n)"
with agrK_inj have "g = priAK n"
by (rule injD)
thus "\<exists>n. g = priAK n" ..
next
fix evsFR4 S A U s a b c f g'
assume
A: "?M g \<in> parts (A \<union> spies evsFR4) \<longrightarrow> (\<exists>n. g = priAK n)" and
B: "(evsFR4, S, A, U) \<in> protocol" and
C: "Crypt (sesK (c * f)) \<lbrace>pubAK (c * (s + a * b)), Auth_Data g' b,
pubAK g', ?M g'\<rbrace> \<in> synth (analz (A \<union> spies evsFR4))"
(is "?M' \<in> synth (analz ?A)")
assume "agrK g = agrK g'"
with agrK_inj have D: "g = g'"
by (rule injD)
show "\<exists>n. g = priAK n"
proof -
have "?M' \<in> analz ?A \<or>
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g' b, pubAK g', ?M g'\<rbrace>
\<in> synth (analz ?A) \<and>
Key (sesK (c * f)) \<in> analz ?A"
using C by (rule synth_crypt)
moreover {
assume "?M' \<in> analz ?A"
hence "?M' \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
hence "\<lbrace>pubAK (c * (s + a * b)), Auth_Data g' b, pubAK g', ?M g'\<rbrace>
\<in> parts ?A"
by (rule parts.Body)
hence "?M g' \<in> parts ?A"
by (rule_tac parts.Snd, assumption?)+
hence "\<exists>n. g' = priAK n"
using A and D by simp
}
moreover {
assume
"\<lbrace>pubAK (c * (s + a * b)), Auth_Data g' b, pubAK g', ?M g'\<rbrace>
\<in> synth (analz ?A) \<and>
Key (sesK (c * f)) \<in> analz ?A"
hence
"\<lbrace>pubAK (c * (s + a * b)), Auth_Data g' b, pubAK g', ?M g'\<rbrace>
\<in> synth (analz ?A)" ..
hence "\<lbrace>Auth_Data g' b, pubAK g', ?M g'\<rbrace> \<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "\<lbrace>pubAK g', ?M g'\<rbrace> \<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "?M g' \<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "?M g' \<in> analz ?A \<or>
Hash (pubAK g') \<in> synth (analz ?A) \<and> Key (priSK CA) \<in> analz ?A"
by (rule synth_crypt)
moreover {
assume "?M g' \<in> analz ?A"
hence "?M g' \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
hence "\<exists>n. g' = priAK n"
using A and D by simp
}
moreover {
assume "Hash (pubAK g') \<in> synth (analz ?A) \<and>
Key (priSK CA) \<in> analz ?A"
hence "Key (priSK CA) \<in> analz ?A" ..
moreover have "Key (priSK CA) \<notin> analz ?A"
using B by (rule pr_sign_key_analz)
ultimately have "\<exists>n. g' = priAK n"
by contradiction
}
ultimately have "\<exists>n. g' = priAK n" ..
}
ultimately have "\<exists>n. g' = priAK n" ..
thus "\<exists>n. g = priAK n"
using D by simp
qed
qed
lemma pr_key_secrecy_aux [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
User m \<noteq> Spy \<longrightarrow>
NonceS (S (User m, n, run)) = Some s \<longrightarrow>
IntMapK (S (User m, n, run)) = Some a \<longrightarrow>
ExtMapK (S (User m, n, run)) = Some b \<longrightarrow>
IntAgrK (S (User m, n, run)) = Some c \<longrightarrow>
ExtAgrK (S (User m, n, run)) = Some f \<longrightarrow>
Says n run 4 X (User m) (Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>) \<in> set evs \<longrightarrow>
Key (sesK (c * f)) \<notin> A"
+supply [[simproc del: defined_all]]
apply (erule protocol.induct, (rule_tac [!] impI)+, simp_all split: if_split_asm)
apply (erule_tac [7] disjE)
apply simp_all
apply (erule_tac [7] conjE)+
apply (erule_tac [8] disjE)+
apply (erule_tac [8] conjE)+
apply simp_all
apply (rule_tac [8] notI)
proof -
fix evsC2 S A U m n run
assume A: "(evsC2, S, A, U) \<in> protocol"
moreover assume "NonceS (S (User m, n, run)) = None"
ultimately have "IntMapK (S (User m, n, run)) = None"
by (rule pr_state_1)
with A have "ExtMapK (S (User m, n, run)) = None"
by (rule pr_state_2)
moreover assume "ExtMapK (S (User m, n, run)) = Some b"
ultimately show "Key (sesK (c * f)) \<notin> A"
by simp
next
fix evsC2 S A U m n run
assume A: "(evsC2, S, A, U) \<in> protocol"
moreover assume "NonceS (S (User m, n, run)) = None"
ultimately have "IntMapK (S (User m, n, run)) = None"
by (rule pr_state_1)
with A have "ExtMapK (S (User m, n, run)) = None"
by (rule pr_state_2)
moreover assume "ExtMapK (S (User m, n, run)) = Some b"
ultimately show "Key (sesK (c * f)) \<notin> A"
by simp
next
fix evsC3 S A U m n run
assume A: "(evsC3, S, A, U) \<in> protocol" and
"ExtMapK (S (User m, n, run)) = None"
hence "IntAgrK (S (User m, n, run)) = None"
by (rule pr_state_3)
with A have "ExtAgrK (S (User m, n, run)) = None"
by (rule pr_state_4)
moreover assume "ExtAgrK (S (User m, n, run)) = Some f"
ultimately show "Key (sesK (c * f)) \<notin> A"
by simp
next
fix evsR3 S A U d s' s'' a' b' c'
assume
A: "(evsR3, S, A, U) \<in> protocol" and
B: "Key (sesK (c' * d * (s'' + a' * b'))) \<notin> U" and
C: "Says n run 4 X (User m) (Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>) \<in> set evsR3"
(is "Says _ _ _ _ _ ?M \<in> _")
show "s'' = s' \<and> Pri_AgrK c' \<in> analz (A \<union> spies evsR3) \<longrightarrow>
sesK (c * f) \<noteq> sesK (c' * d * (s' + a' * b'))"
proof (rule impI, rule notI, erule conjE)
have "?M \<in> spies evsR3"
using C by (rule set_spies)
hence "?M \<in> A \<union> spies evsR3"
by simp
hence "?M \<in> parts (A \<union> spies evsR3)"
by (rule parts.Inj)
with A have "Key (sesK (c * f)) \<in> U"
by (rule pr_sesk_auth)
moreover assume "sesK (c * f) = sesK (c' * d * (s' + a' * b'))" and "s'' = s'"
hence "Key (sesK (c * f)) \<notin> U"
using B by simp
ultimately show False
by contradiction
qed
next
fix evsFR3 S A U s' a' b' c' d
assume
A: "(evsFR3, S, A, U) \<in> protocol" and
B: "Key (sesK (c' * d * (s' + a' * b'))) \<notin> U" and
C: "Says n run 4 X (User m) (Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>) \<in> set evsFR3"
(is "Says _ _ _ _ _ ?M \<in> _")
show "sesK (c * f) \<noteq> sesK (c' * d * (s' + a' * b'))"
proof
have "?M \<in> spies evsFR3"
using C by (rule set_spies)
hence "?M \<in> A \<union> spies evsFR3"
by simp
hence "?M \<in> parts (A \<union> spies evsFR3)"
by (rule parts.Inj)
with A have "Key (sesK (c * f)) \<in> U"
by (rule pr_sesk_auth)
moreover assume "sesK (c * f) = sesK (c' * d * (s' + a' * b'))"
hence "Key (sesK (c * f)) \<notin> U"
using B by simp
ultimately show False
by contradiction
qed
next
fix evsC4 S A U n run and m :: nat
assume "(evsC4, S, A, U) \<in> protocol"
moreover assume "0 < m"
hence "User m \<noteq> Spy"
by simp
moreover assume "Says n run 4 X (User m) (Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>) \<in> set evsC4"
ultimately have "ExtAgrK (S (User m, n, run)) \<noteq> None"
by (rule pr_ext_agrk_user_2)
moreover assume "ExtAgrK (S (User m, n, run)) = None"
ultimately show "Key (sesK (c * f)) \<notin> A"
by contradiction
next
fix evsR4 A U n run X s' a' b' d e and S :: state
assume A: "User m = X"
assume "agrK (c * (s + a * b')) = agrK e"
with agrK_inj have B: "c * (s + a * b') = e"
by (rule injD)
assume "sesK (c * f) = sesK (d * e)"
with sesK_inj have "c * f = d * e"
by (rule injD)
hence C: "c * f = c * d * (s + a * b')"
using B by simp
assume "ExtAgrK (S (X, n, run)) = Some f"
hence D: "ExtAgrK (S (User m, n, run)) = Some f"
using A by simp
assume E: "(evsR4, S, A, U) \<in> protocol"
moreover assume "0 < m"
hence F: "User m \<noteq> Spy"
by simp
moreover assume "NonceS (S (X, n, run)) = Some s"
hence G: "NonceS (S (User m, n, run)) = Some s"
using A by simp
moreover assume "IntMapK (S (X, n, run)) = Some a"
hence H: "IntMapK (S (User m, n, run)) = Some a"
using A by simp
moreover assume "ExtMapK (S (X, n, run)) = Some b'"
hence I: "ExtMapK (S (User m, n, run)) = Some b'"
using A by simp
moreover assume "IntAgrK (S (X, n, run)) = Some c"
hence J: "IntAgrK (S (User m, n, run)) = Some c"
using A by simp
moreover assume K: "NonceS (S (Card n, n, run)) = Some s'"
moreover assume L: "IntMapK (S (Card n, n, run)) = Some b'"
moreover assume M: "ExtMapK (S (Card n, n, run)) = Some a'"
moreover assume N: "IntAgrK (S (Card n, n, run)) = Some d"
moreover assume "ExtAgrK (S (Card n, n, run)) = Some e"
hence "ExtAgrK (S (Card n, n, run)) = Some (c * (s + a * b'))"
using B by simp
moreover assume "Says n run 4 X (Card n)
(Crypt (sesK (d * e)) (pubAK (d * (s' + a' * b')))) \<in> set evsR4"
hence "Says n run 4 (User m) (Card n)
(Crypt (sesK (d * e)) (pubAK (d * (s' + a' * b')))) \<in> set evsR4"
using A by simp
with E and F and D have "d * (s' + a' * b') = f"
by (rule pr_ext_agrk_user_3)
hence "c * d * (s' + a' * b') = c * d * (s + a * b')"
using C by auto
hence "s' + a' * b' = s + a * b'"
proof auto
assume "c = 0"
moreover have "c * (s + a * b') \<noteq> 0"
using E and G and H and I and J by (rule pr_int_agrk_user_3)
ultimately show ?thesis
by simp
next
assume "d = 0"
moreover have "d * (s' + a' * b') \<noteq> 0"
using E and K and L and M and N by (rule pr_int_agrk_card)
ultimately show ?thesis
by simp
qed
ultimately have "Key (sesK (c * d * (s + a * b'))) \<notin> A"
by (rule pr_sesk_card_user)
moreover have "c * d * (s + a * b') = d * e"
using B by simp
ultimately show "Key (sesK (d * e)) \<notin> A"
by simp
next
fix evsFR4 S A U m n run b g
assume
A: "(evsFR4, S, A, U) \<in> protocol" and
B: "ExtMapK (S (User m, n, run)) = Some b" and
C: "IntAgrK (S (User m, n, run)) = Some c" and
D: "ExtAgrK (S (User m, n, run)) = Some f" and
E: "0 < m" and
F: "Key (sesK (c * f)) \<in> A"
assume G: "Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace> \<in> synth (analz (A \<union> spies evsFR4))"
(is "Crypt _ ?M \<in> synth (analz ?A)")
hence "Crypt (sesK (c * f)) ?M \<in> analz ?A \<or>
?M \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
by (rule synth_crypt)
moreover {
assume "Crypt (sesK (c * f)) ?M \<in> analz ?A"
hence "Crypt (sesK (c * f)) ?M \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
hence "?M \<in> parts ?A"
by (rule parts.Body)
hence "Crypt (priSK CA) (Hash (pubAK g)) \<in> parts ?A"
by (rule_tac parts.Snd, assumption?)+
with A have "\<exists>n. g = priAK n"
by (rule pr_sign_parts)
}
moreover {
assume "?M \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
hence "?M \<in> synth (analz ?A)" ..
hence "\<lbrace>Auth_Data g b, pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "\<lbrace>pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace> \<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "Crypt (priSK CA) (Hash (pubAK g)) \<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "Crypt (priSK CA) (Hash (pubAK g)) \<in> analz ?A \<or>
Hash (pubAK g) \<in> synth (analz ?A) \<and> Key (priSK CA) \<in> analz ?A"
by (rule synth_crypt)
moreover {
assume "Crypt (priSK CA) (Hash (pubAK g)) \<in> analz ?A"
hence "Crypt (priSK CA) (Hash (pubAK g)) \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
with A have "\<exists>n. g = priAK n"
by (rule pr_sign_parts)
}
moreover {
assume "Hash (pubAK g) \<in> synth (analz ?A) \<and> Key (priSK CA) \<in> analz ?A"
hence "Key (priSK CA) \<in> analz ?A" ..
moreover have "Key (priSK CA) \<notin> analz ?A"
using A by (rule pr_sign_key_analz)
ultimately have "\<exists>n. g = priAK n"
by contradiction
}
ultimately have "\<exists>n. g = priAK n" ..
}
ultimately have "\<exists>n. g = priAK n" ..
then obtain n' where "g = priAK n'" ..
hence "Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data (priAK n') b, pubAK (priAK n'),
Crypt (priSK CA) (Hash (pubAK (priAK n')))\<rbrace> \<in> synth (analz ?A)"
(is "Crypt _ ?M \<in> _")
using G by simp
hence "Crypt (sesK (c * f)) ?M \<in> analz ?A \<or>
?M \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
by (rule synth_crypt)
moreover {
assume "Crypt (sesK (c * f)) ?M \<in> analz ?A"
hence "Crypt (sesK (c * f)) ?M \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
hence "?M \<in> parts ?A"
by (rule parts.Body)
hence "\<lbrace>Auth_Data (priAK n') b, pubAK (priAK n'),
Crypt (priSK CA) (Hash (pubAK (priAK n')))\<rbrace> \<in> parts ?A"
by (rule parts.Snd)
hence "Auth_Data (priAK n') b \<in> parts ?A"
by (rule parts.Fst)
}
moreover {
assume "?M \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
hence "?M \<in> synth (analz ?A)" ..
hence "\<lbrace>Auth_Data (priAK n') b, pubAK (priAK n'),
Crypt (priSK CA) (Hash (pubAK (priAK n')))\<rbrace> \<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "Auth_Data (priAK n') b \<in> synth (analz ?A)"
by (rule synth_analz_fst)
hence "Auth_Data (priAK n') b \<in> analz ?A \<or>
Pri_AgrK (priAK n') \<in> analz ?A \<and> Pri_AgrK b \<in> analz ?A"
by (rule synth_auth_data)
moreover {
assume "Auth_Data (priAK n') b \<in> analz ?A"
hence "Auth_Data (priAK n') b \<in> parts ?A"
by (rule subsetD [OF analz_parts_subset])
}
moreover {
assume "Pri_AgrK (priAK n') \<in> analz ?A \<and> Pri_AgrK b \<in> analz ?A"
hence "Pri_AgrK (priAK n') \<in> analz ?A" ..
moreover have "Pri_AgrK (priAK n') \<notin> analz ?A"
using A by (rule pr_auth_key_analz)
ultimately have "Auth_Data (priAK n') b \<in> parts ?A"
by contradiction
}
ultimately have "Auth_Data (priAK n') b \<in> parts ?A" ..
}
ultimately have "Auth_Data (priAK n') b \<in> parts ?A" ..
with A have "\<exists>n run. IntMapK (S (Card n, n, run)) = Some b"
by (rule pr_auth_data_parts)
then obtain n' and run' where "IntMapK (S (Card n', n', run')) = Some b"
by blast
with A have "Pri_AgrK b \<notin> analz ?A"
by (rule pr_int_mapk_analz)
hence H: "Pri_AgrK b \<notin> A"
using A by (simp add: pr_pri_agrk_analz)
have "\<exists>X. Says n run 3 X (User m) (pubAK f) \<in> set evsFR4"
using A and D by (rule pr_ext_agrk_user_4)
then obtain X where "Says n run 3 X (User m) (pubAK f) \<in> set evsFR4" ..
with A have
"(\<exists>s a b d. f = d * (s + a * b) \<and>
NonceS (S (Card n, n, run)) = Some s \<and>
IntMapK (S (Card n, n, run)) = Some b \<and>
ExtMapK (S (Card n, n, run)) = Some a \<and>
IntAgrK (S (Card n, n, run)) = Some d \<and>
d \<noteq> 0 \<and> s + a * b \<noteq> 0) \<or>
(\<exists>b. Pri_AgrK b \<in> A \<and>
ExtMapK (S (User m, n, run)) = Some b)"
(is "(\<exists>s a b d. ?P s a b d) \<or> ?Q")
by (rule pr_ext_agrk_user_5)
moreover have I: "\<not> ?Q"
proof (rule notI, erule exE, erule conjE)
fix b'
assume "ExtMapK (S (User m, n, run)) = Some b'"
hence "b' = b"
using B by simp
moreover assume "Pri_AgrK b' \<in> A"
ultimately have "Pri_AgrK b \<in> A"
by simp
thus False
using H by contradiction
qed
ultimately have "\<exists>s a b d. ?P s a b d"
by simp
then obtain s' and a' and b' and d where J: "?P s' a' b' d"
by blast
hence "ExtAgrK (S (User m, n, run)) = Some (d * (s' + a' * b'))"
using D by simp
with A and C have
"\<lbrace>Key (sesK (c * (d * (s' + a' * b')))), Agent (User m), Number n, Number run\<rbrace> \<in> U"
by (rule pr_sesk_user_1)
moreover have K: "Key (sesK (c * (d * (s' + a' * b')))) \<in> A"
using F and J by simp
ultimately have
"(\<exists>d' e'. c * (d * (s' + a' * b')) = d' * e' \<and>
IntAgrK (S (Card n, n, run)) = Some d' \<and>
ExtAgrK (S (Card n, n, run)) = Some e') \<or>
(\<exists>b. Pri_AgrK b \<in> A \<and>
ExtMapK (S (User m, n, run)) = Some b)"
(is "(\<exists>d e. ?P d e) \<or> _")
by (rule pr_sesk_user_3 [OF A])
hence "\<exists>d e. ?P d e"
using I by simp
then obtain d' and e' where L: "?P d' e'"
by blast
hence "d' = d"
using J by simp
hence "d * c * (s' + a' * b') = d * e'"
using L by simp
hence "e' = c * (s' + a' * b')"
using J by simp
hence M: "ExtAgrK (S (Card n, n, run)) = Some (c * (s' + a' * b'))"
using L by simp
have "c * (d * (s' + a' * b')) = c * d * (s' + a' * b')"
by simp
hence "Key (sesK (c * (d * (s' + a' * b')))) = Key (sesK (c * d * (s' + a' * b')))"
by (rule arg_cong)
hence "Key (sesK (c * d * (s' + a' * b'))) \<in> A"
using K by simp
moreover have "Key (sesK (c * d * (s' + a' * b'))) \<notin> A"
proof (rule pr_ext_agrk_card [OF A, of n run], simp_all add: J M)
qed (rule pr_int_agrk_user_2 [OF A, of m n run], simp_all add: C E)
ultimately show False
by contradiction
qed
theorem pr_key_secrecy [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
User m \<noteq> Spy \<longrightarrow>
Says n run 5 (User m) (Card n) (Crypt (sesK K) (Passwd m)) \<in> set evs \<longrightarrow>
Key (sesK K) \<notin> analz (A \<union> spies evs)"
proof (simp add: pr_key_analz, erule protocol.induct, simp_all,
(rule_tac [1] impI)+, (rule_tac [2-3] impI)+, rule_tac [1-2] notI, simp_all)
fix evsR3 S A U s a b c d
assume
"(evsR3, S, A, U) \<in> protocol" and
"Says n run 5 (User m) (Card n)
(Crypt (sesK (c * d * (s + a * b))) (Passwd m)) \<in> set evsR3"
hence "Key (sesK (c * d * (s + a * b))) \<in> U"
by (rule pr_sesk_passwd)
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show False
by contradiction
next
fix evsFR3 S A U s a b c d
assume
"(evsFR3, S, A, U) \<in> protocol" and
"Says n run 5 (User m) (Card n)
(Crypt (sesK (c * d * (s + a * b))) (Passwd m)) \<in> set evsFR3"
hence "Key (sesK (c * d * (s + a * b))) \<in> U"
by (rule pr_sesk_passwd)
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show False
by contradiction
next
fix evsC5 S A U n run s a b c f g X and m :: nat
assume "(evsC5, S, A, U) \<in> protocol"
moreover assume "0 < m"
hence "User m \<noteq> Spy"
by simp
moreover assume
"NonceS (S (User m, n, run)) = Some s" and
"IntMapK (S (User m, n, run)) = Some a" and
"ExtMapK (S (User m, n, run)) = Some b" and
"IntAgrK (S (User m, n, run)) = Some c" and
"ExtAgrK (S (User m, n, run)) = Some f" and
"Says n run 4 X (User m) (Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>)
\<in> set evsC5"
ultimately show "Key (sesK (c * f)) \<notin> A"
by (rule pr_key_secrecy_aux)
qed
theorem pr_passwd_secrecy [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
User m \<noteq> Spy \<longrightarrow>
Passwd m \<notin> analz (A \<union> spies evs)"
proof (erule protocol.induct, rule_tac [!] impI, simp_all add: analz_simp_insert_2,
rule contra_subsetD [OF analz_parts_subset], subst parts_simp, simp, blast+,
rule_tac [3] impI)
fix evsR1 S A U n s
assume
A: "Passwd m \<notin> analz (A \<union> spies evsR1)" and
B: "(evsR1, S, A, U) \<in> protocol" and
C: "Pri_AgrK s \<notin> U"
show
"(n \<in> bad \<longrightarrow> Passwd m \<notin> analz (insert (Crypt (symK n) (Pri_AgrK s))
(insert (Pri_AgrK s) (A \<union> spies evsR1)))) \<and>
(n \<notin> bad \<longrightarrow> Passwd m \<notin> analz (insert (Crypt (symK n) (Pri_AgrK s))
(A \<union> spies evsR1)))"
(is "(_ \<longrightarrow> ?T) \<and> (_ \<longrightarrow> ?T')")
proof (rule conjI, rule_tac [!] impI)
assume "n \<in> bad"
hence "Key (invK (symK n)) \<in> analz (A \<union> spies evsR1)"
using B by (simp add: pr_symk_analz invK_symK)
hence "Key (invK (symK n)) \<in> analz (insert (Pri_AgrK s) (A \<union> spies evsR1))"
by (rule rev_subsetD, rule_tac analz_mono, blast)
moreover have "analz (insert (Pri_AgrK s) (A \<union> spies evsR1)) =
insert (Pri_AgrK s) (analz (A \<union> spies evsR1))"
using B and C by (rule pr_pri_agrk_unused)
ultimately show ?T
using A by (simp add: analz_crypt_in)
next
assume "n \<notin> bad"
hence "Key (invK (symK n)) \<notin> analz (A \<union> spies evsR1)"
using B by (simp add: pr_symk_analz invK_symK)
thus ?T'
using A by (simp add: analz_crypt_out)
qed
next
fix evsFR1 A m' s
assume
A: "Passwd m \<notin> analz (A \<union> spies evsFR1)" and
B: "Crypt (symK m') (Pri_AgrK s) \<in> synth (analz (A \<union> spies evsFR1))"
thus "Passwd m \<notin> analz (insert (Crypt (symK m') (Pri_AgrK s)) (A \<union> spies evsFR1))"
proof (cases "Key (invK (symK m')) \<in> analz (A \<union> spies evsFR1)",
simp_all add: analz_crypt_in analz_crypt_out)
case True
have "Crypt (symK m') (Pri_AgrK s) \<in> analz (A \<union> spies evsFR1) \<or>
Pri_AgrK s \<in> synth (analz (A \<union> spies evsFR1)) \<and>
Key (symK m') \<in> analz (A \<union> spies evsFR1)"
(is "?P \<or> ?Q")
using B by (rule synth_crypt)
moreover {
assume ?P
hence "Pri_AgrK s \<in> analz (A \<union> spies evsFR1)"
using True by (rule analz.Decrypt)
}
moreover {
assume ?Q
hence "Pri_AgrK s \<in> synth (analz (A \<union> spies evsFR1))" ..
hence "Pri_AgrK s \<in> analz (A \<union> spies evsFR1)"
by (rule synth_simp_intro, simp)
}
ultimately have "Pri_AgrK s \<in> analz (A \<union> spies evsFR1)" ..
thus "Passwd m \<notin> analz (insert (Pri_AgrK s) (A \<union> spies evsFR1))"
using A by (simp add: analz_simp_insert_1)
qed
next
fix evsC2 S A U a
assume
"Passwd m \<notin> analz (A \<union> spies evsC2)" and
"(evsC2, S, A, U) \<in> protocol" and
"Pri_AgrK a \<notin> U"
thus "Passwd m \<notin> analz (insert (Pri_AgrK a) (A \<union> spies evsC2))"
by (subst pr_pri_agrk_unused, simp_all)
next
fix evsC3 S A U c and m' :: nat
assume
"Passwd m \<notin> analz (A \<union> spies evsC3)" and
"(evsC3, S, A, U) \<in> protocol" and
"Pri_AgrK c \<notin> U"
thus "m' = 0 \<longrightarrow> Passwd m \<notin> analz (insert (Pri_AgrK c) (A \<union> spies evsC3))"
by (rule_tac impI, subst pr_pri_agrk_unused, simp_all)
next
fix evsR3 S A U s s' a b c d
assume "Passwd m \<notin> analz (A \<union> spies evsR3)"
moreover assume
"(evsR3, S, A, U) \<in> protocol" and
"Key (sesK (c * d * (s + a * b))) \<notin> U"
(is "Key ?K \<notin> _")
hence "analz (insert (Key ?K) (A \<union> spies evsR3)) =
insert (Key ?K) (analz (A \<union> spies evsR3))"
by (rule pr_key_unused)
ultimately show "s' = s \<and> Pri_AgrK c \<in> analz (A \<union> spies evsR3) \<longrightarrow>
Passwd m \<notin> analz (insert (Key ?K) (A \<union> spies evsR3))"
by simp
next
fix evsFR3 S A U s a b c d
assume "Passwd m \<notin> analz (A \<union> spies evsFR3)"
moreover assume
"(evsFR3, S, A, U) \<in> protocol" and
"Key (sesK (c * d * (s + a * b))) \<notin> U"
(is "Key ?K \<notin> _")
hence "analz (insert (Key ?K) (A \<union> spies evsFR3)) =
insert (Key ?K) (analz (A \<union> spies evsFR3))"
by (rule pr_key_unused)
ultimately show "Passwd m \<notin> analz (insert (Key ?K) (A \<union> spies evsFR3))"
by simp
next
fix evsC4 A c f
assume "Passwd m \<notin> analz (A \<union> spies evsC4)"
thus "Passwd m \<notin> analz (insert (Crypt (sesK (c * f)) (pubAK f)) (A \<union> spies evsC4))"
by (cases "Key (invK (sesK (c * f))) \<in> analz (A \<union> spies evsC4)",
simp_all add: analz_crypt_in analz_crypt_out analz_simp_insert_2)
next
fix evsFC4 A s a b d e
assume "Passwd m \<notin> analz (A \<union> spies evsFC4)"
thus "Passwd m \<notin> analz (insert (Crypt (sesK (d * e)) (pubAK (d * (s + a * b))))
(A \<union> spies evsFC4))"
by (cases "Key (invK (sesK (d * e))) \<in> analz (A \<union> spies evsFC4)",
simp_all add: analz_crypt_in analz_crypt_out analz_simp_insert_2)
next
fix evsR4 S A U n run b d e
let
?A = "A \<union> spies evsR4" and
?H = "Hash (pubAK (priAK n))" and
?M = "\<lbrace>pubAK (priAK n), Crypt (priSK CA) (Hash (pubAK (priAK n)))\<rbrace>"
assume
A: "Passwd m \<notin> analz ?A" and
B: "(evsR4, S, A, U) \<in> protocol" and
C: "IntMapK (S (Card n, n, run)) = Some b"
show "Passwd m \<notin> analz (insert (Crypt (sesK (d * e))
\<lbrace>pubAK e, Auth_Data (priAK n) b, ?M\<rbrace>) ?A)"
proof (cases "Key (invK (sesK (d * e))) \<in> analz ?A",
simp_all add: analz_crypt_in analz_crypt_out analz_mpair analz_simp_insert_2 A)
have "Key (pubSK CA) \<in> analz ?A"
using B by (rule pr_valid_key_analz)
hence D: "analz (insert (Crypt (priSK CA) ?H) ?A) =
{Crypt (priSK CA) ?H, ?H} \<union> analz ?A"
by (simp add: analz_crypt_in analz_simp_insert_2)
have "Pri_AgrK (priAK n) \<notin> analz ?A"
using B by (rule pr_auth_key_analz)
hence E: "Pri_AgrK (priAK n) \<notin> analz (insert (Crypt (priSK CA) ?H) ?A)"
using D by simp
have "Pri_AgrK b \<notin> analz ?A"
using B and C by (rule pr_int_mapk_analz)
hence F: "Pri_AgrK b \<notin> analz (insert (Crypt (priSK CA) ?H) ?A)"
using D by simp
show "Passwd m \<notin> analz (insert (Auth_Data (priAK n) b) (insert ?M ?A))"
proof ((subst insert_commute, simp add: analz_mpair analz_simp_insert_2)+,
subst analz_auth_data_out [OF E F])
qed (simp add: A D)
qed
next
fix evsFR4 S A U s a b c f g
let
?A = "A \<union> spies evsFR4" and
?H = "Hash (pubAK g)" and
?M = "\<lbrace>pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>" and
?M' = "\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>"
assume
A: "Passwd m \<notin> analz ?A" and
B: "(evsFR4, S, A, U) \<in> protocol" and
C: "Crypt (sesK (c * f)) ?M' \<in> synth (analz ?A)"
have D:
"Key (invK (sesK (c * f))) \<in> analz ?A \<longrightarrow>
Pri_AgrK b \<in> analz ?A \<or> Pri_AgrK g \<in> analz ?A \<longrightarrow>
Pri_AgrK b \<in> analz ?A \<and> Pri_AgrK g \<in> analz ?A"
(is "?P \<longrightarrow> ?Q \<longrightarrow> ?R")
proof (rule impI)+
assume ?P and ?Q
have "Crypt (sesK (c * f)) ?M' \<in> analz ?A \<or>
?M' \<in> synth (analz ?A) \<and> Key (sesK (c * f)) \<in> analz ?A"
using C by (rule synth_crypt)
moreover {
assume "Crypt (sesK (c * f)) ?M' \<in> analz ?A"
hence "?M' \<in> analz ?A"
using \<open>?P\<close> by (rule analz.Decrypt)
hence "\<lbrace>Auth_Data g b, pubAK g, Crypt (priSK CA) (Hash (pubAK g))\<rbrace>
\<in> analz ?A"
by (rule analz.Snd)
hence E: "Auth_Data g b \<in> analz ?A"
by (rule analz.Fst)
have ?R
proof (rule disjE [OF \<open>?Q\<close>])
assume "Pri_AgrK b \<in> analz ?A"
moreover from this have "Pri_AgrK g \<in> analz ?A"
by (rule analz.Auth_Fst [OF E])
ultimately show ?R ..
next
assume "Pri_AgrK g \<in> analz ?A"
moreover from this have "Pri_AgrK b \<in> analz ?A"
by (rule analz.Auth_Snd [OF E])
ultimately show ?R
by simp
qed
}
moreover {
assume "?M' \<in> synth (analz ?A) \<and>
Key (sesK (c * f)) \<in> analz ?A"
hence "?M' \<in> synth (analz ?A)" ..
hence "\<lbrace>Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace> \<in> synth (analz ?A)"
by (rule synth_analz_snd)
hence "Auth_Data g b \<in> synth (analz ?A)"
by (rule synth_analz_fst)
hence "Auth_Data g b \<in> analz ?A \<or>
Pri_AgrK g \<in> analz ?A \<and> Pri_AgrK b \<in> analz ?A"
by (rule synth_auth_data)
moreover {
assume E: "Auth_Data g b \<in> analz ?A"
have ?R
proof (rule disjE [OF \<open>?Q\<close>])
assume "Pri_AgrK b \<in> analz ?A"
moreover from this have "Pri_AgrK g \<in> analz ?A"
by (rule analz.Auth_Fst [OF E])
ultimately show ?R ..
next
assume "Pri_AgrK g \<in> analz ?A"
moreover from this have "Pri_AgrK b \<in> analz ?A"
by (rule analz.Auth_Snd [OF E])
ultimately show ?R
by simp
qed
}
moreover {
assume "Pri_AgrK g \<in> analz ?A \<and> Pri_AgrK b \<in> analz ?A"
hence ?R
by simp
}
ultimately have ?R ..
}
ultimately show ?R ..
qed
show "Passwd m \<notin> analz (insert (Crypt (sesK (c * f)) ?M') ?A)"
proof (cases "Key (invK (sesK (c * f))) \<in> analz ?A",
simp_all add: analz_crypt_in analz_crypt_out analz_mpair analz_simp_insert_2 A)
assume E: "Key (invK (sesK (c * f))) \<in> analz ?A"
have "Key (pubSK CA) \<in> analz ?A"
using B by (rule pr_valid_key_analz)
hence F: "analz (insert (Crypt (priSK CA) ?H) ?A) =
{Crypt (priSK CA) ?H, ?H} \<union> analz ?A"
by (simp add: analz_crypt_in analz_simp_insert_2)
show "Passwd m \<notin> analz (insert (Auth_Data g b) (insert ?M ?A))"
proof (cases "Pri_AgrK g \<in> analz ?A \<or> Pri_AgrK b \<in> analz ?A", simp_all)
assume G: "Pri_AgrK g \<in> analz ?A \<or> Pri_AgrK b \<in> analz ?A"
hence H: "Pri_AgrK g \<in> analz (insert (Crypt (priSK CA) ?H) ?A) \<or>
Pri_AgrK b \<in> analz (insert (Crypt (priSK CA) ?H) ?A)"
using F by simp
have I: "Pri_AgrK b \<in> analz ?A \<and> Pri_AgrK g \<in> analz ?A"
using D and E and G by blast
hence "Pri_AgrK g \<in> analz (insert (Crypt (priSK CA) ?H) ?A)"
using F by simp
hence J: "Pri_AgrK g \<in> analz (insert (Pri_AgrK b)
(insert (Crypt (priSK CA) ?H) ?A))"
by (rule rev_subsetD, rule_tac analz_mono, blast)
have K: "Pri_AgrK b \<in> analz (insert (Crypt (priSK CA) ?H) ?A)"
using F and I by simp
show ?thesis
proof ((subst insert_commute, simp add: analz_mpair analz_simp_insert_2)+,
subst analz_auth_data_in [OF H], simp del: Un_insert_right,
subst analz_simp_insert_1 [OF J], subst analz_simp_insert_1 [OF K])
qed (simp add: A F)
next
assume G: "Pri_AgrK g \<notin> analz ?A \<and> Pri_AgrK b \<notin> analz ?A"
hence H: "Pri_AgrK g \<notin> analz (insert (Crypt (priSK CA) ?H) ?A)"
using F by simp
have I: "Pri_AgrK b \<notin> analz (insert (Crypt (priSK CA) ?H) ?A)"
using F and G by simp
show ?thesis
proof ((subst insert_commute, simp add: analz_mpair analz_simp_insert_2)+,
subst analz_auth_data_out [OF H I])
qed (simp add: A F)
qed
qed
next
fix evsC5 S A U m' n run s a b c f g X
let ?M = "\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>"
assume
A: "Passwd m \<notin> analz (A \<union> spies evsC5)" and
B: "0 < m" and
C: "(evsC5, S, A, U) \<in> protocol" and
D: "NonceS (S (User m', n, run)) = Some s" and
E: "IntMapK (S (User m', n, run)) = Some a" and
F: "ExtMapK (S (User m', n, run)) = Some b" and
G: "IntAgrK (S (User m', n, run)) = Some c" and
H: "ExtAgrK (S (User m', n, run)) = Some f" and
I: "Says n run 4 X (User m') (Crypt (sesK (c * f)) ?M) \<in> set evsC5"
from A show "Passwd m \<notin> analz (insert (Crypt (sesK (c * f)) (Passwd m'))
(A \<union> spies evsC5))"
proof (cases "Key (invK (sesK (c * f))) \<in> analz (A \<union> spies evsC5)",
simp_all add: analz_crypt_in analz_crypt_out analz_simp_insert_2, rule_tac notI)
case True
moreover assume "m = m'"
hence "User m' \<noteq> Spy"
using B by simp
hence "Key (sesK (c * f)) \<notin> A"
by (rule pr_key_secrecy_aux [OF C _ D E F G H I])
hence "Key (invK (sesK (c * f))) \<notin> analz (A \<union> spies evsC5)"
using C by (simp add: pr_key_analz invK_sesK)
ultimately show False
by contradiction
qed
next
fix evsFC5 A n d e
assume
A: "Passwd m \<notin> analz (A \<union> spies evsFC5)" and
B: "Crypt (sesK (d * e)) (Passwd n) \<in> synth (analz (A \<union> spies evsFC5))"
from A show "Passwd m \<notin> analz (insert (Crypt (sesK (d * e)) (Passwd n))
(A \<union> spies evsFC5))"
proof (cases "Key (invK (sesK (d * e))) \<in> analz (A \<union> spies evsFC5)",
simp_all add: analz_crypt_in analz_crypt_out analz_simp_insert_2, rule_tac notI)
case True
have "Crypt (sesK (d * e)) (Passwd n) \<in> analz (A \<union> spies evsFC5) \<or>
Passwd n \<in> synth (analz (A \<union> spies evsFC5)) \<and>
Key (sesK (d * e)) \<in> analz (A \<union> spies evsFC5)"
(is "?P \<or> ?Q")
using B by (rule synth_crypt)
moreover {
assume ?P
hence "Passwd n \<in> analz (A \<union> spies evsFC5)"
using True by (rule analz.Decrypt)
}
moreover {
assume ?Q
hence "Passwd n \<in> synth (analz (A \<union> spies evsFC5))" ..
hence "Passwd n \<in> analz (A \<union> spies evsFC5)"
by (rule synth_simp_intro, simp)
}
ultimately have "Passwd n \<in> analz (A \<union> spies evsFC5)" ..
moreover assume "m = n"
hence "Passwd n \<notin> analz (A \<union> spies evsFC5)"
using A by simp
ultimately show False
by contradiction
qed
next
fix evsR5 A d e
assume "Passwd m \<notin> analz (A \<union> spies evsR5)"
thus "Passwd m \<notin> analz (insert (Crypt (sesK (d * e)) (Number 0)) (A \<union> spies evsR5))"
by (cases "Key (invK (sesK (d * e))) \<in> analz (A \<union> spies evsR5)",
simp_all add: analz_crypt_in analz_crypt_out analz_simp_insert_2)
next
fix evsFR5 A c f
assume "Passwd m \<notin> analz (A \<union> spies evsFR5)"
thus "Passwd m \<notin> analz (insert (Crypt (sesK (c * f)) (Number 0)) (A \<union> spies evsFR5))"
by (cases "Key (invK (sesK (c * f))) \<in> analz (A \<union> spies evsFR5)",
simp_all add: analz_crypt_in analz_crypt_out analz_simp_insert_2)
qed
subsection "Authenticity theorems"
text \<open>
This section contains a series of lemmas culminating in the authenticity theorems
\<open>pr_user_authenticity\<close> and \<open>pr_card_authenticity\<close>. Structured Isar proofs are used.
\null
\<close>
lemma pr_passwd_parts [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Crypt (sesK K) (Passwd m) \<in> parts (A \<union> spies evs) \<longrightarrow>
(\<exists>n run. Says n run 5 (User m) (Card n) (Crypt (sesK K) (Passwd m)) \<in> set evs) \<or>
(\<exists>run. Says m run 5 Spy (Card m) (Crypt (sesK K) (Passwd m)) \<in> set evs)"
(is "_ \<Longrightarrow> ?M \<in> _ \<longrightarrow> ?P evs \<or> ?Q evs")
proof (erule protocol.induct, subst parts_simp, (simp, blast)+)
qed (simp_all add: parts_simp_insert parts_auth_data parts_crypt parts_mpair, blast+)
lemma pr_unique_run_1 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
\<lbrace>Key (sesK (d * e)), Agent (User m), Number n', Number run'\<rbrace> \<in> U \<longrightarrow>
IntAgrK (S (Card n, n, run)) = Some d \<longrightarrow>
ExtAgrK (S (Card n, n, run)) = Some e \<longrightarrow>
n' = n \<and> run' = run"
proof (erule protocol.induct, simp_all, rule_tac [3] conjI, (rule_tac [1-4] impI)+,
(rule_tac [5] impI)+, simp_all, (erule_tac [2-3] conjE)+, (rule_tac [!] ccontr))
fix evsR3 S A U s' a b c
assume "c * (s' + a * b) = e"
hence A: "d * e = c * d * (s' + a * b)"
by simp
assume
"(evsR3, S, A, U) \<in> protocol" and
"\<lbrace>Key (sesK (d * e)), Agent (User m), Number n', Number run'\<rbrace> \<in> U"
hence "Key (sesK (d * e)) \<in> U"
by (rule pr_sesk_user_2)
hence "Key (sesK (c * d * (s' + a * b))) \<in> U"
by (simp only: A)
moreover assume "Key (sesK (c * d * (s' + a * b))) \<notin> U"
ultimately show False
by contradiction
next
fix evsR3 S A U s a b c d'
assume "Key (sesK (c * d' * (s + a * b))) \<notin> U"
moreover assume "sesK (d * e) = sesK (c * d' * (s + a * b))"
with sesK_inj have "d * e = c * d' * (s + a * b)"
by (rule injD)
ultimately have "Key (sesK (d * e)) \<notin> U"
by simp
moreover assume
"(evsR3, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some e"
hence "Key (sesK (d * e)) \<in> U"
by (rule pr_sesk_card)
ultimately show False
by contradiction
next
fix evsFR3 S A U s a b c d'
assume "Key (sesK (c * d' * (s + a * b))) \<notin> U"
moreover assume "sesK (d * e) = sesK (c * d' * (s + a * b))"
with sesK_inj have "d * e = c * d' * (s + a * b)"
by (rule injD)
ultimately have "Key (sesK (d * e)) \<notin> U"
by simp
moreover assume
"(evsFR3, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some e"
hence "Key (sesK (d * e)) \<in> U"
by (rule pr_sesk_card)
ultimately show False
by contradiction
qed
lemma pr_unique_run_2:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
IntAgrK (S (User m, n', run')) = Some c \<Longrightarrow>
ExtAgrK (S (User m, n', run')) = Some f \<Longrightarrow>
IntAgrK (S (Card n, n, run)) = Some d \<Longrightarrow>
ExtAgrK (S (Card n, n, run)) = Some e \<Longrightarrow>
d * e = c * f \<Longrightarrow>
n' = n \<and> run' = run"
proof (frule pr_sesk_user_1, assumption+, drule sym [of "d * e"], simp)
qed (rule pr_unique_run_1)
lemma pr_unique_run_3 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
IntAgrK (S (Card n', n', run')) = Some d' \<longrightarrow>
ExtAgrK (S (Card n', n', run')) = Some e' \<longrightarrow>
IntAgrK (S (Card n, n, run)) = Some d \<longrightarrow>
ExtAgrK (S (Card n, n, run)) = Some e \<longrightarrow>
d * e = d' * e' \<longrightarrow>
n' = n \<and> run' = run"
proof (erule protocol.induct, simp_all, rule_tac [!] conjI, (rule_tac [!] impI)+,
simp_all, rule_tac [!] ccontr)
fix evsR3 S A U n' run' s' a b c
assume "c * (s' + a * b) = e"
hence A: "d * e = c * d * (s' + a * b)"
by simp
assume
"(evsR3, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n', n', run')) = Some d'" and
"ExtAgrK (S (Card n', n', run')) = Some e'"
hence "Key (sesK (d' * e')) \<in> U"
by (rule pr_sesk_card)
moreover assume "d * e = d' * e'"
ultimately have "Key (sesK (c * d * (s' + a * b))) \<in> U"
using A by simp
moreover assume "Key (sesK (c * d * (s' + a * b))) \<notin> U"
ultimately show False
by contradiction
next
fix evsR3 S A U s' a b c
assume "c * (s' + a * b) = e'"
hence A: "d' * e' = c * d' * (s' + a * b)"
by simp
assume
"(evsR3, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some e"
hence "Key (sesK (d * e)) \<in> U"
by (rule pr_sesk_card)
moreover assume "d * e = d' * e'"
ultimately have "Key (sesK (c * d' * (s' + a * b))) \<in> U"
using A by simp
moreover assume "Key (sesK (c * d' * (s' + a * b))) \<notin> U"
ultimately show False
by contradiction
qed
lemma pr_unique_run_4 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Says n' run' 5 X (Card n') (Crypt (sesK (d * e)) (Passwd m)) \<in> set evs \<longrightarrow>
IntAgrK (S (Card n, n, run)) = Some d \<longrightarrow>
ExtAgrK (S (Card n, n, run)) = Some e \<longrightarrow>
n' = n \<and> run' = run"
+using [[simproc del: defined_all]]
proof (erule protocol.induct, simp_all, (rule_tac [!] impI)+, rule_tac [1-3] impI,
simp_all, (erule_tac [2-3] conjE)+)
fix evsR3 S A U n run s' a b c
assume "c * (s' + a * b) = e"
hence A: "d * e = c * d * (s' + a * b)"
by simp
assume
"(evsR3, S, A, U) \<in> protocol" and
"Says n' run' 5 X (Card n') (Crypt (sesK (d * e)) (Passwd m)) \<in> set evsR3"
hence "Key (sesK (d * e)) \<in> U"
by (rule pr_sesk_passwd)
hence "Key (sesK (c * d * (s' + a * b))) \<in> U"
by (simp only: A)
moreover assume "Key (sesK (c * d * (s' + a * b))) \<notin> U"
ultimately show "n' = n \<and> run' = run"
by contradiction
next
fix evsC5 S A U m n' run' c f
assume
"(evsC5, S, A, U) \<in> protocol" and
"IntAgrK (S (User m, n', run')) = Some c" and
"ExtAgrK (S (User m, n', run')) = Some f" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some e"
moreover assume "sesK (d * e) = sesK (c * f)"
with sesK_inj have "d * e = c * f"
by (rule injD)
ultimately show "n' = n \<and> run' = run"
by (rule pr_unique_run_2)
next
fix evsFC5 S A U n' run' d' e'
assume
"(evsFC5, S, A, U) \<in> protocol" and
"IntAgrK (S (Card n', n', run')) = Some d'" and
"ExtAgrK (S (Card n', n', run')) = Some e'" and
"IntAgrK (S (Card n, n, run)) = Some d" and
"ExtAgrK (S (Card n, n, run)) = Some e"
moreover assume "sesK (d * e) = sesK (d' * e')"
with sesK_inj have "d * e = d' * e'"
by (rule injD)
ultimately show "n' = n \<and> run' = run"
by (rule pr_unique_run_3)
qed
theorem pr_user_authenticity [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Says n run 5 X (Card n) (Crypt (sesK K) (Passwd m)) \<in> set evs \<longrightarrow>
Says n run 5 (User m) (Card n) (Crypt (sesK K) (Passwd m)) \<in> set evs"
proof (erule protocol.induct, simp_all, rule impI, simp)
fix evsFC5 S A U n run d e
assume
A: "Says n run 5 Spy (Card n) (Crypt (sesK (d * e)) (Passwd n)) \<in> set evsFC5 \<longrightarrow>
Says n run 5 (User n) (Card n) (Crypt (sesK (d * e)) (Passwd n)) \<in> set evsFC5"
(is "_ \<longrightarrow> ?T") and
B: "(evsFC5, S, A, U) \<in> protocol" and
C: "IntAgrK (S (Card n, n, run)) = Some d" and
D: "ExtAgrK (S (Card n, n, run)) = Some e" and
E: "Crypt (sesK (d * e)) (Passwd n) \<in> synth (analz (A \<union> spies evsFC5))"
show "n = 0 \<or> ?T"
proof (cases "n = 0", simp_all)
assume "0 < n"
hence "User n \<noteq> Spy"
by simp
with B have F: "Passwd n \<notin> analz (A \<union> spies evsFC5)"
by (rule pr_passwd_secrecy)
have "Crypt (sesK (d * e)) (Passwd n) \<in> analz (A \<union> spies evsFC5) \<or>
Passwd n \<in> synth (analz (A \<union> spies evsFC5)) \<and>
Key (sesK (d * e)) \<in> analz (A \<union> spies evsFC5)"
(is "?P \<or> ?Q")
using E by (rule synth_crypt)
moreover have "\<not> ?Q"
proof
assume ?Q
hence "Passwd n \<in> synth (analz (A \<union> spies evsFC5))" ..
hence "Passwd n \<in> analz (A \<union> spies evsFC5)"
by (rule synth_simp_intro, simp)
thus False
using F by contradiction
qed
ultimately have ?P
by simp
hence "Crypt (sesK (d * e)) (Passwd n) \<in> parts (A \<union> spies evsFC5)"
by (rule subsetD [OF analz_parts_subset])
with B have
"(\<exists>n' run'. Says n' run' 5 (User n) (Card n') (Crypt (sesK (d * e)) (Passwd n))
\<in> set evsFC5) \<or>
(\<exists>run'. Says n run' 5 Spy (Card n) (Crypt (sesK (d * e)) (Passwd n))
\<in> set evsFC5)"
(is "(\<exists>n' run'. ?P n' run') \<or> (\<exists>run'. ?Q run')")
by (rule pr_passwd_parts)
moreover {
assume "\<exists>n' run'. ?P n' run'"
then obtain n' and run' where "?P n' run'"
by blast
moreover from this have "n' = n \<and> run' = run"
by (rule pr_unique_run_4 [OF B _ C D])
ultimately have ?T
by simp
}
moreover {
assume "\<exists>run'. ?Q run'"
then obtain run' where "?Q run'" ..
moreover from this have "n = n \<and> run' = run"
by (rule pr_unique_run_4 [OF B _ C D])
ultimately have "?Q run"
by simp
with A have ?T ..
}
ultimately show ?T ..
qed
qed
lemma pr_confirm_parts [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Crypt (sesK K) (Number 0) \<in> parts (A \<union> spies evs) \<longrightarrow>
Key (sesK K) \<notin> A \<longrightarrow>
(\<exists>n run X.
Says n run 5 X (Card n) (Crypt (sesK K) (Passwd n)) \<in> set evs \<and>
Says n run 5 (Card n) X (Crypt (sesK K) (Number 0)) \<in> set evs)"
(is "_ \<Longrightarrow> _ \<longrightarrow> _ \<longrightarrow> ?P K evs")
+using [[simproc del: defined_all]]
proof (erule protocol.induct, simp, subst parts_simp, simp, blast+,
simp_all add: parts_simp_insert parts_auth_data parts_crypt parts_mpair,
rule_tac [3] conjI, (rule_tac [!] impI)+, simp_all, blast+)
fix evsFR5 S A U c f
assume
A: "Crypt (sesK (c * f)) (Number 0) \<in> parts (A \<union> spies evsFR5) \<longrightarrow>
?P (c * f) evsFR5" and
B: "(evsFR5, S, A, U) \<in> protocol" and
C: "Key (sesK (c * f)) \<notin> A" and
D: "Crypt (sesK (c * f)) (Number 0) \<in> synth (analz (A \<union> spies evsFR5))"
show "?P (c * f) evsFR5"
proof -
have "Crypt (sesK (c * f)) (Number 0) \<in> analz (A \<union> spies evsFR5) \<or>
Number 0 \<in> synth (analz (A \<union> spies evsFR5)) \<and>
Key (sesK (c * f)) \<in> analz (A \<union> spies evsFR5)"
using D by (rule synth_crypt)
moreover have "Key (sesK (c * f)) \<notin> analz (A \<union> spies evsFR5)"
using B and C by (simp add: pr_key_analz)
ultimately have "Crypt (sesK (c * f)) (Number 0) \<in> analz (A \<union> spies evsFR5)"
by simp
hence "Crypt (sesK (c * f)) (Number 0) \<in> parts (A \<union> spies evsFR5)"
by (rule subsetD [OF analz_parts_subset])
with A show ?thesis ..
qed
qed
lemma pr_confirm_says [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Says n run 5 X Spy (Crypt (sesK K) (Number 0)) \<in> set evs \<longrightarrow>
Says n run 5 Spy (Card n) (Crypt (sesK K) (Passwd n)) \<in> set evs"
by (erule protocol.induct, simp_all)
lemma pr_passwd_says [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Says n run 5 X (Card n) (Crypt (sesK K) (Passwd m)) \<in> set evs \<longrightarrow>
X = User m \<or> X = Spy"
by (erule protocol.induct, simp_all)
lemma pr_unique_run_5 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
\<lbrace>Key (sesK K), Agent (User m'), Number n', Number run'\<rbrace> \<in> U \<longrightarrow>
\<lbrace>Key (sesK K), Agent (User m), Number n, Number run\<rbrace> \<in> U \<longrightarrow>
m = m' \<and> n = n' \<and> run = run'"
+using [[simproc del: defined_all]]
proof (erule protocol.induct, simp_all, blast, (rule conjI, rule impI)+,
rule_tac [2] impI, (rule_tac [3] impI)+, rule_tac [4] conjI, (rule_tac [4-5] impI)+,
simp_all, blast, rule_tac [!] ccontr)
fix evsR3 S A U s a b c d
assume
"(evsR3, S, A, U) \<in> protocol" and
"\<lbrace>Key (sesK (c * d * (s + a * b))), Agent (User m), Number n, Number run\<rbrace> \<in> U"
hence "Key (sesK (c * d * (s + a * b))) \<in> U"
by (rule pr_sesk_user_2)
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show False
by contradiction
next
fix evsR3 S A U s a b c d
assume
"(evsR3, S, A, U) \<in> protocol" and
"\<lbrace>Key (sesK (c * d * (s + a * b))), Agent (User m'), Number n', Number run'\<rbrace> \<in> U"
hence "Key (sesK (c * d * (s + a * b))) \<in> U"
by (rule pr_sesk_user_2)
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show False
by contradiction
next
fix evsFR3 S A U s a b c d
assume
"(evsFR3, S, A, U) \<in> protocol" and
"\<lbrace>Key (sesK (c * d * (s + a * b))), Agent (User m), Number n, Number run\<rbrace> \<in> U"
hence "Key (sesK (c * d * (s + a * b))) \<in> U"
by (rule pr_sesk_user_2)
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show False
by contradiction
next
fix evsFR3 S A U s a b c d
assume
"(evsFR3, S, A, U) \<in> protocol" and
"\<lbrace>Key (sesK (c * d * (s + a * b))), Agent (User m'), Number n', Number run'\<rbrace> \<in> U"
hence "Key (sesK (c * d * (s + a * b))) \<in> U"
by (rule pr_sesk_user_2)
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show False
by contradiction
qed
lemma pr_unique_run_6:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
\<lbrace>Key (sesK (c * f)), Agent (User m'), Number n', Number run'\<rbrace> \<in> U \<Longrightarrow>
IntAgrK (S (User m, n, run)) = Some c \<Longrightarrow>
ExtAgrK (S (User m, n, run)) = Some f \<Longrightarrow>
m = m' \<and> n = n' \<and> run = run'"
proof (frule pr_sesk_user_1, assumption+)
qed (rule pr_unique_run_5)
lemma pr_unique_run_7 [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Says n' run' 5 (User m') (Card n') (Crypt (sesK K) (Passwd m')) \<in> set evs \<longrightarrow>
\<lbrace>Key (sesK K), Agent (User m), Number n, Number run\<rbrace> \<in> U \<longrightarrow>
Key (sesK K) \<notin> A \<longrightarrow>
m' = m \<and> n' = n \<and> run' = run"
proof (erule protocol.induct, simp_all, (rule impI)+, (rule_tac [2-3] impI)+,
(erule_tac [3] conjE)+, (drule_tac [3] sym [of m'])+, drule_tac [3] sym [of 0],
simp_all)
fix evsR3 S A U n run s a b c d
assume
"(evsR3, S, A, U) \<in> protocol" and
"Says n' run' 5 (User m') (Card n')
(Crypt (sesK (c * d * (s + a * b))) (Passwd m')) \<in> set evsR3"
hence "Key (sesK (c * d * (s + a * b))) \<in> U"
by (rule pr_sesk_passwd)
moreover assume "Key (sesK (c * d * (s + a * b))) \<notin> U"
ultimately show "m' = m \<and> n' = n \<and> run' = run"
by contradiction
next
fix evsC5 S A U m' n' run' c f
assume
"(evsC5, S, A, U) \<in> protocol" and
"\<lbrace>Key (sesK (c * f)), Agent (User m), Number n, Number run\<rbrace> \<in> U" and
"IntAgrK (S (User m', n', run')) = Some c" and
"ExtAgrK (S (User m', n', run')) = Some f"
thus "m' = m \<and> n' = n \<and> run' = run"
by (rule pr_unique_run_6)
next
fix evsFC5 S A U run' d e
assume
A: "Says 0 run' 5 Spy (Card 0) (Crypt (sesK (d * e)) (Passwd 0)) \<in> set evsFC5 \<longrightarrow>
m = 0 \<and> n = 0 \<and> run' = run" and
B: "(evsFC5, S, A, U) \<in> protocol" and
C: "IntAgrK (S (Card 0, 0, run')) = Some d" and
D: "ExtAgrK (S (Card 0, 0, run')) = Some e" and
E: "Crypt (sesK (d * e)) (Passwd 0) \<in> synth (analz (A \<union> spies evsFC5))" and
F: "Key (sesK (d * e)) \<notin> A"
have "Crypt (sesK (d * e)) (Passwd 0) \<in> analz (A \<union> spies evsFC5) \<or>
Passwd 0 \<in> synth (analz (A \<union> spies evsFC5)) \<and>
Key (sesK (d * e)) \<in> analz (A \<union> spies evsFC5)"
using E by (rule synth_crypt)
moreover have "Key (sesK (d * e)) \<notin> analz (A \<union> spies evsFC5)"
using B and F by (simp add: pr_key_analz)
ultimately have "Crypt (sesK (d * e)) (Passwd 0) \<in> analz (A \<union> spies evsFC5)"
by simp
hence "Crypt (sesK (d * e)) (Passwd 0) \<in> parts (A \<union> spies evsFC5)"
by (rule subsetD [OF analz_parts_subset])
with B have
"(\<exists>n run. Says n run 5 Spy (Card n) (Crypt (sesK (d * e)) (Passwd 0))
\<in> set evsFC5) \<or>
(\<exists>run. Says 0 run 5 Spy (Card 0) (Crypt (sesK (d * e)) (Passwd 0))
\<in> set evsFC5)"
(is "(\<exists>n run. ?P n run) \<or> (\<exists>run. ?Q run)")
by (rule pr_passwd_parts)
moreover {
assume "\<exists>n run. ?P n run"
then obtain n'' and run'' where "?P n'' run''"
by blast
moreover from this have "n'' = 0 \<and> run'' = run'"
by (rule pr_unique_run_4 [OF B _ C D])
ultimately have "?P 0 run'"
by simp
}
moreover {
assume "\<exists>run. ?Q run"
then obtain run'' where "?Q run''" ..
hence "\<exists>n. ?P n run''" ..
then obtain n'' where "?P n'' run''" ..
moreover from this have "n'' = 0 \<and> run'' = run'"
by (rule pr_unique_run_4 [OF B _ C D])
ultimately have "?P 0 run'"
by simp
}
ultimately have "?P 0 run'" ..
with A show "m = 0 \<and> n = 0 \<and> run' = run" ..
qed
lemma pr_unique_run_8:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Says n' run' 5 (User m') (Card n') (Crypt (sesK (c * f)) (Passwd m')) \<in> set evs \<Longrightarrow>
IntAgrK (S (User m, n, run)) = Some c \<Longrightarrow>
ExtAgrK (S (User m, n, run)) = Some f \<Longrightarrow>
Key (sesK (c * f)) \<notin> A \<Longrightarrow>
m' = m \<and> n' = n \<and> run' = run"
proof (frule pr_sesk_user_1, assumption+)
qed (rule pr_unique_run_7)
lemma pr_unique_passwd_parts [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Crypt (sesK K) (Passwd m') \<in> parts (A \<union> spies evs) \<longrightarrow>
Crypt (sesK K) (Passwd m) \<in> parts (A \<union> spies evs) \<longrightarrow>
m' = m"
+using [[simproc del: defined_all]]
proof (erule protocol.induct, simp, subst parts_simp, simp, blast+,
simp_all add: parts_simp_insert parts_auth_data parts_crypt parts_mpair,
rule_tac [!] conjI, (rule_tac [!] impI)+, erule_tac [!] conjE, simp_all)
fix evsC5 S A U m'' n run s a b c f g X
assume
A: "(evsC5, S, A, U) \<in> protocol" and
B: "NonceS (S (User m'', n, run)) = Some s" and
C: "IntMapK (S (User m'', n, run)) = Some a" and
D: "ExtMapK (S (User m'', n, run)) = Some b" and
E: "IntAgrK (S (User m'', n, run)) = Some c" and
F: "ExtAgrK (S (User m'', n, run)) = Some f" and
G: "Crypt (sesK (c * f)) (Passwd m) \<in> parts (A \<union> spies evsC5)" and
H: "m' = m''" and
I: "Says n run 4 X (User m'') (Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>) \<in> set evsC5"
show "m'' = m"
proof (cases "m'' = 0", rule classical)
case True
moreover assume "m'' \<noteq> m"
ultimately have J: "User m \<noteq> Spy"
using H by simp
have
"(\<exists>n run. Says n run 5 (User m) (Card n) (Crypt (sesK (c * f)) (Passwd m))
\<in> set evsC5) \<or>
(\<exists>run. Says m run 5 Spy (Card m) (Crypt (sesK (c * f)) (Passwd m))
\<in> set evsC5)"
(is "(\<exists>n run. ?P n run) \<or> (\<exists>run. ?Q run)")
using A and G by (rule pr_passwd_parts)
moreover {
assume "\<exists>n run. ?P n run"
then obtain n' and run' where K: "?P n' run'"
by blast
with A and J have "Key (sesK (c * f)) \<notin> analz (A \<union> spies evsC5)"
by (rule pr_key_secrecy)
hence "Key (sesK (c * f)) \<notin> A"
using A by (simp add: pr_key_analz)
hence "m = m'' \<and> n' = n \<and> run' = run"
by (rule pr_unique_run_8 [OF A K E F])
hence ?thesis
by simp
}
moreover {
assume "\<exists>run. ?Q run"
then obtain run' where "?Q run'" ..
with A have K: "?P m run'"
by (rule pr_user_authenticity)
with A and J have "Key (sesK (c * f)) \<notin> analz (A \<union> spies evsC5)"
by (rule pr_key_secrecy)
hence "Key (sesK (c * f)) \<notin> A"
using A by (simp add: pr_key_analz)
hence "m = m'' \<and> m = n \<and> run' = run"
by (rule pr_unique_run_8 [OF A K E F])
hence ?thesis
by simp
}
ultimately show ?thesis ..
next
case False
hence "User m'' \<noteq> Spy"
by simp
hence J: "Key (sesK (c * f)) \<notin> A"
by (rule pr_key_secrecy_aux [OF A _ B C D E F I])
have
"(\<exists>n run. Says n run 5 (User m) (Card n) (Crypt (sesK (c * f)) (Passwd m))
\<in> set evsC5) \<or>
(\<exists>run. Says m run 5 Spy (Card m) (Crypt (sesK (c * f)) (Passwd m))
\<in> set evsC5)"
(is "(\<exists>n run. ?P n run) \<or> (\<exists>run. ?Q run)")
using A and G by (rule pr_passwd_parts)
moreover {
assume "\<exists>n run. ?P n run"
then obtain n' and run' where "?P n' run'"
by blast
hence "m = m'' \<and> n' = n \<and> run' = run"
by (rule pr_unique_run_8 [OF A _ E F J])
hence ?thesis
by simp
}
moreover {
assume "\<exists>run. ?Q run"
then obtain run' where "?Q run'" ..
with A have "?P m run'"
by (rule pr_user_authenticity)
hence "m = m'' \<and> m = n \<and> run' = run"
by (rule pr_unique_run_8 [OF A _ E F J])
hence ?thesis
by simp
}
ultimately show ?thesis ..
qed
next
fix evsC5 S A U m'' n run s a b c f g X
assume
A: "(evsC5, S, A, U) \<in> protocol" and
B: "NonceS (S (User m'', n, run)) = Some s" and
C: "IntMapK (S (User m'', n, run)) = Some a" and
D: "ExtMapK (S (User m'', n, run)) = Some b" and
E: "IntAgrK (S (User m'', n, run)) = Some c" and
F: "ExtAgrK (S (User m'', n, run)) = Some f" and
G: "Crypt (sesK (c * f)) (Passwd m') \<in> parts (A \<union> spies evsC5)" and
H: "m = m''" and
I: "Says n run 4 X (User m'') (Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>) \<in> set evsC5"
show "m' = m''"
proof (cases "m'' = 0", rule classical)
case True
moreover assume "m' \<noteq> m''"
ultimately have J: "User m' \<noteq> Spy"
using H by simp
have
"(\<exists>n run. Says n run 5 (User m') (Card n) (Crypt (sesK (c * f)) (Passwd m'))
\<in> set evsC5) \<or>
(\<exists>run. Says m' run 5 Spy (Card m') (Crypt (sesK (c * f)) (Passwd m'))
\<in> set evsC5)"
(is "(\<exists>n run. ?P n run) \<or> (\<exists>run. ?Q run)")
using A and G by (rule pr_passwd_parts)
moreover {
assume "\<exists>n run. ?P n run"
then obtain n' and run' where K: "?P n' run'"
by blast
with A and J have "Key (sesK (c * f)) \<notin> analz (A \<union> spies evsC5)"
by (rule pr_key_secrecy)
hence "Key (sesK (c * f)) \<notin> A"
using A by (simp add: pr_key_analz)
hence "m' = m'' \<and> n' = n \<and> run' = run"
by (rule pr_unique_run_8 [OF A K E F])
hence ?thesis
by simp
}
moreover {
assume "\<exists>run. ?Q run"
then obtain run' where "?Q run'" ..
with A have K: "?P m' run'"
by (rule pr_user_authenticity)
with A and J have "Key (sesK (c * f)) \<notin> analz (A \<union> spies evsC5)"
by (rule pr_key_secrecy)
hence "Key (sesK (c * f)) \<notin> A"
using A by (simp add: pr_key_analz)
hence "m' = m'' \<and> m' = n \<and> run' = run"
by (rule pr_unique_run_8 [OF A K E F])
hence ?thesis
by simp
}
ultimately show ?thesis ..
next
case False
hence "User m'' \<noteq> Spy"
by simp
hence J: "Key (sesK (c * f)) \<notin> A"
by (rule pr_key_secrecy_aux [OF A _ B C D E F I])
have
"(\<exists>n run. Says n run 5 (User m') (Card n) (Crypt (sesK (c * f)) (Passwd m'))
\<in> set evsC5) \<or>
(\<exists>run. Says m' run 5 Spy (Card m') (Crypt (sesK (c * f)) (Passwd m'))
\<in> set evsC5)"
(is "(\<exists>n run. ?P n run) \<or> (\<exists>run. ?Q run)")
using A and G by (rule pr_passwd_parts)
moreover {
assume "\<exists>n run. ?P n run"
then obtain n' and run' where "?P n' run'"
by blast
hence "m' = m'' \<and> n' = n \<and> run' = run"
by (rule pr_unique_run_8 [OF A _ E F J])
hence ?thesis
by simp
}
moreover {
assume "\<exists>run. ?Q run"
then obtain run' where "?Q run'" ..
with A have "?P m' run'"
by (rule pr_user_authenticity)
hence "m' = m'' \<and> m' = n \<and> run' = run"
by (rule pr_unique_run_8 [OF A _ E F J])
hence ?thesis
by simp
}
ultimately show ?thesis ..
qed
next
fix evsFC5 S A U n d e
assume
A: "Crypt (sesK (d * e)) (Passwd n) \<in> parts (A \<union> spies evsFC5) \<longrightarrow>
n = m" and
B: "(evsFC5, S, A, U) \<in> protocol" and
C: "Crypt (sesK (d * e)) (Passwd n) \<in> synth (analz (A \<union> spies evsFC5))" and
D: "Crypt (sesK (d * e)) (Passwd m) \<in> parts (A \<union> spies evsFC5)"
show "n = m"
proof (rule classical)
assume E: "n \<noteq> m"
have F: "Crypt (sesK (d * e)) (Passwd n) \<in> analz (A \<union> spies evsFC5) \<or>
Passwd n \<in> synth (analz (A \<union> spies evsFC5)) \<and>
Key (sesK (d * e)) \<in> analz (A \<union> spies evsFC5)"
using C by (rule synth_crypt)
have "Crypt (sesK (d * e)) (Passwd n) \<in> analz (A \<union> spies evsFC5)"
proof (rule disjE [OF F], assumption, erule conjE, cases "n = 0")
case True
hence G: "User m \<noteq> Spy"
using E by simp
have
"(\<exists>n run. Says n run 5 (User m) (Card n) (Crypt (sesK (d * e)) (Passwd m))
\<in> set evsFC5) \<or>
(\<exists>run. Says m run 5 Spy (Card m) (Crypt (sesK (d * e)) (Passwd m))
\<in> set evsFC5)"
(is "(\<exists>n run. ?P n run) \<or> (\<exists>run. ?Q run)")
using B and D by (rule pr_passwd_parts)
moreover {
assume "\<exists>n run. ?P n run"
then obtain n' and run where "?P n' run"
by blast
with B and G have "Key (sesK (d * e)) \<notin> analz (A \<union> spies evsFC5)"
by (rule pr_key_secrecy)
}
moreover {
assume "\<exists>run. ?Q run"
then obtain run where "?Q run" ..
with B have "?P m run"
by (rule pr_user_authenticity)
with B and G have "Key (sesK (d * e)) \<notin> analz (A \<union> spies evsFC5)"
by (rule pr_key_secrecy)
}
ultimately have "Key (sesK (d * e)) \<notin> analz (A \<union> spies evsFC5)" ..
moreover assume "Key (sesK (d * e)) \<in> analz (A \<union> spies evsFC5)"
ultimately show ?thesis
by contradiction
next
case False
hence "User n \<noteq> Spy"
by simp
with B have "Passwd n \<notin> analz (A \<union> spies evsFC5)"
by (rule pr_passwd_secrecy)
moreover assume "Passwd n \<in> synth (analz (A \<union> spies evsFC5))"
hence "Passwd n \<in> analz (A \<union> spies evsFC5)"
by (rule synth_simp_intro, simp)
ultimately show ?thesis
by contradiction
qed
hence "Crypt (sesK (d * e)) (Passwd n) \<in> parts (A \<union> spies evsFC5)"
by (rule subsetD [OF analz_parts_subset])
with A show ?thesis ..
qed
next
fix evsFC5 S A U n d e
assume
A: "Crypt (sesK (d * e)) (Passwd n) \<in> parts (A \<union> spies evsFC5) \<longrightarrow>
m' = n" and
B: "(evsFC5, S, A, U) \<in> protocol" and
C: "Crypt (sesK (d * e)) (Passwd n) \<in> synth (analz (A \<union> spies evsFC5))" and
D: "Crypt (sesK (d * e)) (Passwd m') \<in> parts (A \<union> spies evsFC5)"
show "m' = n"
proof (rule classical)
assume E: "m' \<noteq> n"
have F: "Crypt (sesK (d * e)) (Passwd n) \<in> analz (A \<union> spies evsFC5) \<or>
Passwd n \<in> synth (analz (A \<union> spies evsFC5)) \<and>
Key (sesK (d * e)) \<in> analz (A \<union> spies evsFC5)"
using C by (rule synth_crypt)
have "Crypt (sesK (d * e)) (Passwd n) \<in> analz (A \<union> spies evsFC5)"
proof (rule disjE [OF F], assumption, erule conjE, cases "n = 0")
case True
hence G: "User m' \<noteq> Spy"
using E by simp
have
"(\<exists>n run. Says n run 5 (User m') (Card n) (Crypt (sesK (d * e)) (Passwd m'))
\<in> set evsFC5) \<or>
(\<exists>run. Says m' run 5 Spy (Card m') (Crypt (sesK (d * e)) (Passwd m'))
\<in> set evsFC5)"
(is "(\<exists>n run. ?P n run) \<or> (\<exists>run. ?Q run)")
using B and D by (rule pr_passwd_parts)
moreover {
assume "\<exists>n run. ?P n run"
then obtain n' and run where "?P n' run"
by blast
with B and G have "Key (sesK (d * e)) \<notin> analz (A \<union> spies evsFC5)"
by (rule pr_key_secrecy)
}
moreover {
assume "\<exists>run. ?Q run"
then obtain run where "?Q run" ..
with B have "?P m' run"
by (rule pr_user_authenticity)
with B and G have "Key (sesK (d * e)) \<notin> analz (A \<union> spies evsFC5)"
by (rule pr_key_secrecy)
}
ultimately have "Key (sesK (d * e)) \<notin> analz (A \<union> spies evsFC5)" ..
moreover assume "Key (sesK (d * e)) \<in> analz (A \<union> spies evsFC5)"
ultimately show ?thesis
by contradiction
next
case False
hence "User n \<noteq> Spy"
by simp
with B have "Passwd n \<notin> analz (A \<union> spies evsFC5)"
by (rule pr_passwd_secrecy)
moreover assume "Passwd n \<in> synth (analz (A \<union> spies evsFC5))"
hence "Passwd n \<in> analz (A \<union> spies evsFC5)"
by (rule synth_simp_intro, simp)
ultimately show ?thesis
by contradiction
qed
hence "Crypt (sesK (d * e)) (Passwd n) \<in> parts (A \<union> spies evsFC5)"
by (rule subsetD [OF analz_parts_subset])
with A show ?thesis ..
qed
qed
theorem pr_card_authenticity [rule_format]:
"(evs, S, A, U) \<in> protocol \<Longrightarrow>
Says n run 5 (User m) (Card n) (Crypt (sesK K) (Passwd m)) \<in> set evs \<longrightarrow>
Says n run 5 X (User m) (Crypt (sesK K) (Number 0)) \<in> set evs \<longrightarrow>
n = m \<and>
(Says m run 5 (Card m) (User m) (Crypt (sesK K) (Number 0)) \<in> set evs \<or>
Says m run 5 (Card m) Spy (Crypt (sesK K) (Number 0)) \<in> set evs)"
proof (erule protocol.induct, simp_all, (rule_tac [1-2] impI)+, (erule_tac [2] conjE)+,
(rule_tac [3] impI, rule_tac [3] conjI)+, rule_tac [4] disjI1, rule_tac [5] impI,
(rule_tac [6] impI)+, simp_all)
fix evsC5 S A U m n run s a b c f g X'
assume
A: "Says n run 5 (User m) (Card n) (Crypt (sesK (c * f)) (Passwd m))
\<in> set evsC5 \<longrightarrow>
n = m \<and>
(Says m run 5 (Card m) (User m) (Crypt (sesK (c * f)) (Number 0))
\<in> set evsC5 \<or>
Says m run 5 (Card m) Spy (Crypt (sesK (c * f)) (Number 0))
\<in> set evsC5)" and
B: "(evsC5, S, A, U) \<in> protocol" and
C: "NonceS (S (User m, n, run)) = Some s" and
D: "IntMapK (S (User m, n, run)) = Some a" and
E: "ExtMapK (S (User m, n, run)) = Some b" and
F: "IntAgrK (S (User m, n, run)) = Some c" and
G: "ExtAgrK (S (User m, n, run)) = Some f" and
H: "Says n run 4 X' (User m) (Crypt (sesK (c * f))
\<lbrace>pubAK (c * (s + a * b)), Auth_Data g b, pubAK g,
Crypt (priSK CA) (Hash (pubAK g))\<rbrace>) \<in> set evsC5" and
I: "Says n run 5 X (User m) (Crypt (sesK (c * f)) (Number 0)) \<in> set evsC5"
show "n = m \<and>
(Says m run 5 (Card m) (User m) (Crypt (sesK (c * f)) (Number 0)) \<in> set evsC5 \<or>
Says m run 5 (Card m) Spy (Crypt (sesK (c * f)) (Number 0)) \<in> set evsC5)"
proof (cases "m = 0")
case True
hence "Says n run 5 X Spy (Crypt (sesK (c * f)) (Number 0)) \<in> set evsC5"
using I by simp
with B have "Says n run 5 Spy (Card n) (Crypt (sesK (c * f)) (Passwd n))
\<in> set evsC5"
by (rule pr_confirm_says)
with B have J: "Says n run 5 (User n) (Card n) (Crypt (sesK (c * f)) (Passwd n))
\<in> set evsC5"
by (rule pr_user_authenticity)
show ?thesis
proof (cases n)
case 0
hence "Says n run 5 (User m) (Card n) (Crypt (sesK (c * f)) (Passwd m))
\<in> set evsC5"
using J and True by simp
with A show ?thesis ..
next
case Suc
hence "User n \<noteq> Spy"
by simp
with B have "Key (sesK (c * f)) \<notin> analz (A \<union> spies evsC5)"
using J by (rule pr_key_secrecy)
hence "Key (sesK (c * f)) \<notin> A"
using B by (simp add: pr_key_analz)
hence "n = m \<and> n = n \<and> run = run"
by (rule pr_unique_run_8 [OF B J F G])
hence "Says n run 5 (User m) (Card n) (Crypt (sesK (c * f)) (Passwd m))
\<in> set evsC5"
using J by simp
with A show ?thesis ..
qed
next
case False
have "Crypt (sesK (c * f)) (Number 0) \<in> spies evsC5"
using I by (rule set_spies)
hence "Crypt (sesK (c * f)) (Number 0) \<in> A \<union> spies evsC5" ..
hence "Crypt (sesK (c * f)) (Number 0) \<in> parts (A \<union> spies evsC5)"
by (rule parts.Inj)
moreover have "User m \<noteq> Spy"
using False by simp
hence J: "Key (sesK (c * f)) \<notin> A"
by (rule pr_key_secrecy_aux [OF B _ C D E F G H])
ultimately have "\<exists>n run X.
Says n run 5 X (Card n) (Crypt (sesK (c * f)) (Passwd n)) \<in> set evsC5 \<and>
Says n run 5 (Card n) X (Crypt (sesK (c * f)) (Number 0)) \<in> set evsC5"
by (rule pr_confirm_parts [OF B])
then obtain n' and run' and X where
"Says n' run' 5 X (Card n') (Crypt (sesK (c * f)) (Passwd n')) \<in> set evsC5"
by blast
with B have
"Says n' run' 5 (User n') (Card n') (Crypt (sesK (c * f)) (Passwd n')) \<in> set evsC5"
by (rule pr_user_authenticity)
moreover from this have "n' = m \<and> n' = n \<and> run' = run"
by (rule pr_unique_run_8 [OF B _ F G J])
ultimately have
"Says n run 5 (User m) (Card n) (Crypt (sesK (c * f)) (Passwd m)) \<in> set evsC5"
by auto
with A show ?thesis ..
qed
next
fix evsFC5 S A U run d e
assume
"Says 0 run 5 Spy (Card 0) (Crypt (sesK (d * e)) (Passwd 0)) \<in> set evsFC5 \<longrightarrow>
Says 0 run 5 (Card 0) Spy (Crypt (sesK (d * e)) (Number 0)) \<in> set evsFC5"
moreover assume
"(evsFC5, S, A, U) \<in> protocol" and
"Says 0 run 5 X Spy (Crypt (sesK (d * e)) (Number 0)) \<in> set evsFC5"
hence "Says 0 run 5 Spy (Card 0) (Crypt (sesK (d * e)) (Passwd 0)) \<in> set evsFC5"
by (rule pr_confirm_says)
ultimately show
"Says 0 run 5 (Card 0) Spy (Crypt (sesK (d * e)) (Number 0)) \<in> set evsFC5" ..
next
fix evsR5 S A U n run d e X
assume "(evsR5, S, A, U) \<in> protocol"
moreover assume "Says n run 5 X (Card n) (Crypt (sesK (d * e)) (Passwd n))
\<in> set evsR5"
hence "Crypt (sesK (d * e)) (Passwd n) \<in> spies evsR5"
by (rule set_spies)
hence "Crypt (sesK (d * e)) (Passwd n) \<in> A \<union> spies evsR5" ..
hence "Crypt (sesK (d * e)) (Passwd n) \<in> parts (A \<union> spies evsR5)"
by (rule parts.Inj)
moreover assume "Says n run 5 X (Card n) (Crypt (sesK (d * e)) (Passwd m))
\<in> set evsR5"
hence "Crypt (sesK (d * e)) (Passwd m) \<in> spies evsR5"
by (rule set_spies)
hence "Crypt (sesK (d * e)) (Passwd m) \<in> A \<union> spies evsR5" ..
hence "Crypt (sesK (d * e)) (Passwd m) \<in> parts (A \<union> spies evsR5)"
by (rule parts.Inj)
ultimately show "n = m"
by (rule pr_unique_passwd_parts)
next
fix evsR5 S A U n run d e X
assume "(evsR5, S, A, U) \<in> protocol"
moreover assume "Says n run 5 X (Card n) (Crypt (sesK (d * e)) (Passwd m))
\<in> set evsR5"
hence "Crypt (sesK (d * e)) (Passwd m) \<in> spies evsR5"
by (rule set_spies)
hence "Crypt (sesK (d * e)) (Passwd m) \<in> A \<union> spies evsR5" ..
hence "Crypt (sesK (d * e)) (Passwd m) \<in> parts (A \<union> spies evsR5)"
by (rule parts.Inj)
moreover assume "Says n run 5 X (Card n) (Crypt (sesK (d * e)) (Passwd n))
\<in> set evsR5"
hence "Crypt (sesK (d * e)) (Passwd n) \<in> spies evsR5"
by (rule set_spies)
hence "Crypt (sesK (d * e)) (Passwd n) \<in> A \<union> spies evsR5" ..
hence "Crypt (sesK (d * e)) (Passwd n) \<in> parts (A \<union> spies evsR5)"
by (rule parts.Inj)
ultimately show "m = n"
by (rule pr_unique_passwd_parts)
next
fix evsR5 n' run' d e X
assume "n = m \<and>
(Says m run 5 (Card m) (User m) (Crypt (sesK K) (Number 0)) \<in> set evsR5 \<or>
Says m run 5 (Card m) Spy (Crypt (sesK K) (Number 0)) \<in> set evsR5)"
thus
"m = n' \<and> run = run' \<and> m = n' \<and> User m = X \<and> sesK K = sesK (d * e) \<or>
Says m run 5 (Card m) (User m) (Crypt (sesK K) (Number 0)) \<in> set evsR5 \<or>
m = n' \<and> run = run' \<and> m = n' \<and> Spy = X \<and> sesK K = sesK (d * e) \<or>
Says m run 5 (Card m) Spy (Crypt (sesK K) (Number 0)) \<in> set evsR5"
by blast
next
fix evsFR5 S A U m n run c f
assume
A: "(evsFR5, S, A, U) \<in> protocol" and
B: "0 < m" and
C: "IntAgrK (S (User m, n, run)) = Some c" and
D: "ExtAgrK (S (User m, n, run)) = Some f" and
E: "Crypt (sesK (c * f)) (Number 0) \<in> synth (analz (A \<union> spies evsFR5))" and
F: "Says n run 5 (User m) (Card n) (Crypt (sesK (c * f)) (Passwd m)) \<in> set evsFR5"
have "User m \<noteq> Spy"
using B by simp
with A have G: "Key (sesK (c * f)) \<notin> analz (A \<union> spies evsFR5)"
using F by (rule pr_key_secrecy)
moreover have "Crypt (sesK (c * f)) (Number 0) \<in> analz (A \<union> spies evsFR5) \<or>
Number 0 \<in> synth (analz (A \<union> spies evsFR5)) \<and>
Key (sesK (c * f)) \<in> analz (A \<union> spies evsFR5)"
using E by (rule synth_crypt)
ultimately have "Crypt (sesK (c * f)) (Number 0) \<in> analz (A \<union> spies evsFR5)"
by simp
hence "Crypt (sesK (c * f)) (Number 0) \<in> parts (A \<union> spies evsFR5)"
by (rule subsetD [OF analz_parts_subset])
moreover have H: "Key (sesK (c * f)) \<notin> A"
using A and G by (simp add: pr_key_analz)
ultimately have "\<exists>n run X.
Says n run 5 X (Card n) (Crypt (sesK (c * f)) (Passwd n)) \<in> set evsFR5 \<and>
Says n run 5 (Card n) X (Crypt (sesK (c * f)) (Number 0)) \<in> set evsFR5"
by (rule pr_confirm_parts [OF A])
then obtain n' and run' and X where I:
"Says n' run' 5 X (Card n') (Crypt (sesK (c * f)) (Passwd n')) \<in> set evsFR5 \<and>
Says n' run' 5 (Card n') X (Crypt (sesK (c * f)) (Number 0)) \<in> set evsFR5"
by blast
hence
"Says n' run' 5 X (Card n') (Crypt (sesK (c * f)) (Passwd n')) \<in> set evsFR5" ..
with A have J:
"Says n' run' 5 (User n') (Card n') (Crypt (sesK (c * f)) (Passwd n'))
\<in> set evsFR5"
by (rule pr_user_authenticity)
hence "Crypt (sesK (c * f)) (Passwd n') \<in> spies evsFR5"
by (rule set_spies)
hence "Crypt (sesK (c * f)) (Passwd n') \<in> A \<union> spies evsFR5" ..
hence "Crypt (sesK (c * f)) (Passwd n') \<in> parts (A \<union> spies evsFR5)"
by (rule parts.Inj)
moreover have "Crypt (sesK (c * f)) (Passwd m) \<in> spies evsFR5"
using F by (rule set_spies)
hence "Crypt (sesK (c * f)) (Passwd m) \<in> A \<union> spies evsFR5" ..
hence "Crypt (sesK (c * f)) (Passwd m) \<in> parts (A \<union> spies evsFR5)"
by (rule parts.Inj)
ultimately have "n' = m"
by (rule pr_unique_passwd_parts [OF A])
moreover from this have
"Says m run' 5 (User m) (Card m) (Crypt (sesK (c * f)) (Passwd m))
\<in> set evsFR5"
using J by simp
hence "m = m \<and> m = n \<and> run' = run"
by (rule pr_unique_run_8 [OF A _ C D H])
hence K: "n = m \<and> run' = run"
by simp
ultimately have L:
"Says m run 5 X (Card m) (Crypt (sesK (c * f)) (Passwd m)) \<in> set evsFR5 \<and>
Says m run 5 (Card m) X (Crypt (sesK (c * f)) (Number 0)) \<in> set evsFR5"
using I by simp
moreover from this have
"Says m run 5 X (Card m) (Crypt (sesK (c * f)) (Passwd m)) \<in> set evsFR5" ..
with A have "X = User m \<or> X = Spy"
by (rule pr_passwd_says)
thus "n = m \<and>
(Says m run 5 (Card m) (User m) (Crypt (sesK (c * f)) (Number 0)) \<in> set evsFR5 \<or>
Says m run 5 (Card m) Spy (Crypt (sesK (c * f)) (Number 0)) \<in> set evsFR5)"
by (rule disjE, insert L, simp_all add: K)
qed
end
diff --git a/thys/Pi_Calculus/Early_Semantics.thy b/thys/Pi_Calculus/Early_Semantics.thy
--- a/thys/Pi_Calculus/Early_Semantics.thy
+++ b/thys/Pi_Calculus/Early_Semantics.thy
@@ -1,1301 +1,1300 @@
(*
Title: The pi-calculus
Author/Maintainer: Jesper Bengtson (jebe.dk), 2012
*)
theory Early_Semantics
imports Agent
begin
declare name_fresh[simp del]
nominal_datatype freeRes = InputR name name ("_<_>" [110, 110] 110)
| OutputR name name ("_[_]" [110, 110] 110)
| TauR ("\<tau>" 110)
nominal_datatype residual = BoundOutputR name "\<guillemotleft>name\<guillemotright> pi" ("_<\<nu>_> \<prec> _" [110, 110, 110] 110)
| FreeR freeRes pi
lemma alphaBoundOutput:
fixes a :: name
and x :: name
and P :: pi
and x' :: name
assumes A1: "x' \<sharp> P"
shows "a<\<nu>x> \<prec> P = a<\<nu>x'> \<prec> ([(x, x')] \<bullet> P)"
proof(cases "x=x'")
assume "x=x'"
thus ?thesis by simp
next
assume "x \<noteq> x'"
with A1 show ?thesis
by(simp add: residual.inject alpha name_fresh_left name_calc)
qed
declare name_fresh[simp]
abbreviation Transitions_Freejudge ("_ \<prec> _" [80, 80] 80) where "\<alpha> \<prec> P' \<equiv> (FreeR \<alpha> P')"
inductive "TransitionsEarly" :: "pi \<Rightarrow> residual \<Rightarrow> bool" ("_ \<longmapsto> _" [80, 80] 80)
where
Tau: "\<tau>.(P) \<longmapsto> \<tau> \<prec> P"
| Input: "\<lbrakk>x \<noteq> a; x \<noteq> u\<rbrakk> \<Longrightarrow> a<x>.P \<longmapsto> a<u> \<prec> (P[x::=u])"
| Output: "a{b}.P \<longmapsto> a[b] \<prec> P"
| Match: "\<lbrakk>P \<longmapsto> V\<rbrakk> \<Longrightarrow> [b\<frown>b]P \<longmapsto> V"
| Mismatch: "\<lbrakk>P \<longmapsto> V; a \<noteq> b\<rbrakk> \<Longrightarrow> [a\<noteq>b]P \<longmapsto> V"
| Open: "\<lbrakk>P \<longmapsto> a[b] \<prec> P'; a \<noteq> b\<rbrakk> \<Longrightarrow> <\<nu>b>P \<longmapsto> a<\<nu>b> \<prec> P'"
| Sum1: "\<lbrakk>P \<longmapsto> V\<rbrakk> \<Longrightarrow> (P \<oplus> Q) \<longmapsto> V"
| Sum2: "\<lbrakk>Q \<longmapsto> V\<rbrakk> \<Longrightarrow> (P \<oplus> Q) \<longmapsto> V"
| Par1B: "\<lbrakk>P \<longmapsto> a<\<nu>x> \<prec> P'; x \<sharp> P; x \<sharp> Q; x \<noteq> a\<rbrakk> \<Longrightarrow> P \<parallel> Q \<longmapsto> a<\<nu>x> \<prec> (P' \<parallel> Q)"
| Par1F: "\<lbrakk>P \<longmapsto> \<alpha> \<prec> P'\<rbrakk> \<Longrightarrow> P \<parallel> Q \<longmapsto> \<alpha> \<prec> (P' \<parallel> Q)"
| Par2B: "\<lbrakk>Q \<longmapsto> a<\<nu>x> \<prec> Q'; x \<sharp> P; x \<sharp> Q; x \<noteq> a\<rbrakk> \<Longrightarrow> P \<parallel> Q \<longmapsto> a<\<nu>x> \<prec> (P \<parallel> Q')"
| Par2F: "\<lbrakk>Q \<longmapsto> \<alpha> \<prec> Q'\<rbrakk> \<Longrightarrow> P \<parallel> Q \<longmapsto> \<alpha> \<prec> (P \<parallel> Q')"
| Comm1: "\<lbrakk>P \<longmapsto> a<b> \<prec> P'; Q \<longmapsto> a[b] \<prec> Q'\<rbrakk> \<Longrightarrow> P \<parallel> Q \<longmapsto> \<tau> \<prec> P' \<parallel> Q'"
| Comm2: "\<lbrakk>P \<longmapsto> a[b] \<prec> P'; Q \<longmapsto> a<b> \<prec> Q'\<rbrakk> \<Longrightarrow> P \<parallel> Q \<longmapsto> \<tau> \<prec> P' \<parallel> Q'"
| Close1: "\<lbrakk>P \<longmapsto> a<x> \<prec> P'; Q \<longmapsto> a<\<nu>x> \<prec> Q'; x \<sharp> P; x \<sharp> Q; x \<noteq> a\<rbrakk> \<Longrightarrow> P \<parallel> Q \<longmapsto> \<tau> \<prec> <\<nu>x>(P' \<parallel> Q')"
| Close2: "\<lbrakk>P \<longmapsto> a<\<nu>x> \<prec> P'; Q \<longmapsto> a<x> \<prec> Q'; x \<sharp> P; x \<sharp> Q; x \<noteq> a\<rbrakk> \<Longrightarrow> P \<parallel> Q \<longmapsto> \<tau> \<prec> <\<nu>x>(P' \<parallel> Q')"
| ResB: "\<lbrakk>P \<longmapsto> a<\<nu>x> \<prec> P'; y \<noteq> a; y \<noteq> x; x \<sharp> P; x \<noteq> a\<rbrakk> \<Longrightarrow> <\<nu>y>P \<longmapsto> a<\<nu>x> \<prec> (<\<nu>y>P')"
| ResF: "\<lbrakk>P \<longmapsto> \<alpha> \<prec> P'; y \<sharp> \<alpha>\<rbrakk> \<Longrightarrow> <\<nu>y>P \<longmapsto> \<alpha> \<prec> <\<nu>y>P'"
| Bang: "\<lbrakk>P \<parallel> !P \<longmapsto> V\<rbrakk> \<Longrightarrow> !P \<longmapsto> V"
equivariance TransitionsEarly
nominal_inductive TransitionsEarly
by(auto simp add: abs_fresh fresh_fact2)
lemmas [simp] = freeRes.inject
lemma freshOutputAction:
fixes P :: pi
and a :: name
and b :: name
and P' :: pi
and c :: name
assumes "P \<longmapsto> a[b] \<prec> P'"
and "c \<sharp> P"
shows "c \<noteq> a" and "c \<noteq> b" and "c \<sharp> P'"
proof -
from assms have "c \<noteq> a \<and> c \<noteq> b \<and> c \<sharp> P'"
by(nominal_induct x2=="a[b] \<prec> P'" arbitrary: P' rule: TransitionsEarly.strong_induct) (fastforce simp add: residual.inject abs_fresh freeRes.inject)+
thus "c \<noteq> a" and "c \<noteq> b" and "c \<sharp> P'"
by blast+
qed
lemma freshInputAction:
fixes P :: pi
and a :: name
and b :: name
and P' :: pi
and c :: name
assumes "P \<longmapsto> a<b> \<prec> P'"
and "c \<sharp> P"
shows "c \<noteq> a"
using assms
by(nominal_induct x2=="a<b> \<prec> P'" arbitrary: P' rule: TransitionsEarly.strong_induct) (auto simp add: residual.inject abs_fresh)
lemma freshBoundOutputAction:
fixes P :: pi
and a :: name
and x :: name
and P' :: pi
and c :: name
assumes "P \<longmapsto> a<\<nu>x> \<prec> P'"
and "c \<sharp> P"
shows "c \<noteq> a"
using assms
by(nominal_induct x2=="a<\<nu>x> \<prec> P'" avoiding: x arbitrary: P' rule: TransitionsEarly.strong_induct) (auto simp add: residual.inject abs_fresh fresh_left calc_atm dest: freshOutputAction)
lemmas freshAction = freshOutputAction freshInputAction freshBoundOutputAction
lemma freshInputTransition:
fixes P :: pi
and a :: name
and u :: name
and P' :: pi
and c :: name
assumes "P \<longmapsto> a<u> \<prec> P'"
and "c \<sharp> P"
and "c \<noteq> u"
shows "c \<sharp> P'"
using assms
by(nominal_induct x2=="a<u> \<prec> P'" arbitrary: P' rule: TransitionsEarly.strong_induct)
(fastforce simp add: residual.inject name_fresh_abs fresh_fact1 fresh_fact2)+
lemma freshBoundOutputTransition:
fixes P :: pi
and a :: name
and x :: name
and P' :: pi
and c :: name
assumes "P \<longmapsto> a<\<nu>x> \<prec> P'"
and "c \<sharp> P"
and "c \<noteq> x"
shows "c \<sharp> P'"
using assms
apply(nominal_induct x2=="a<\<nu>x> \<prec> P'" avoiding: x arbitrary: P' rule: TransitionsEarly.strong_induct)
apply(fastforce simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(fastforce simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(fastforce simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(force simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(force simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(force simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(force simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(force simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(force simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(force simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(force simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(force simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction | simp | auto simp add: abs_fresh residual.inject alpha' calc_atm)
apply(fastforce simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction)
apply(fastforce simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction)
apply(fastforce simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction)
apply(fastforce simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction)
apply(auto simp add: residual.inject name_fresh_abs alpha' fresh_left calc_atm dest: freshOutputAction)
-apply force
done
lemma freshTauTransition:
fixes P :: pi
and P' :: pi
and c :: name
assumes "P \<longmapsto> \<tau> \<prec> P'"
and "c \<sharp> P"
shows "c \<sharp> P'"
using assms
apply(nominal_induct x2=="\<tau> \<prec> P'" arbitrary: P' rule: TransitionsEarly.strong_induct)
by(fastforce simp add: residual.inject abs_fresh dest: freshOutputAction freshInputTransition freshBoundOutputTransition)+
lemma freshFreeTransition:
fixes P :: pi
and \<alpha> :: freeRes
and P' :: pi
and c :: name
assumes "P \<longmapsto>\<alpha> \<prec> P'"
and "c \<sharp> P"
and "c \<sharp> \<alpha>"
shows "c \<sharp> P'"
using assms
by(nominal_induct \<alpha> rule: freeRes.strong_inducts)
(auto dest: freshInputTransition freshOutputAction freshTauTransition)
lemmas freshTransition = freshInputTransition freshOutputAction freshFreeTransition
freshBoundOutputTransition freshTauTransition
lemma substTrans[simp]: "b \<sharp> P \<Longrightarrow> ((P::pi)[a::=b])[b::=c] = P[a::=c]"
apply(simp add: injPermSubst[THEN sym])
apply(simp add: renaming)
by(simp add: pt_swap[OF pt_name_inst, OF at_name_inst])
lemma Input:
fixes a :: name
and x :: name
and u :: name
and P :: pi
shows "a<x>.P \<longmapsto>a<u> \<prec> P[x::=u]"
proof -
obtain y::name where "y \<noteq> a" and "y \<noteq> u" and "y \<sharp> P"
by(generate_fresh "name") (auto simp add: fresh_prod)
from \<open>y \<noteq> a\<close> \<open>y \<noteq> u\<close> have "a<y>.([(x, y)] \<bullet> P) \<longmapsto>a<u> \<prec> ([(x, y)] \<bullet> P)[y::=u]"
by(rule Input)
with \<open>y \<sharp> P\<close> show ?thesis by(simp add: alphaInput renaming name_swap)
qed
lemma Par1B:
fixes P :: pi
and a :: name
and x :: name
and P' :: pi
and Q :: pi
assumes "P \<longmapsto>a<\<nu>x> \<prec> P'"
and "x \<sharp> Q"
shows "P \<parallel> Q \<longmapsto> a<\<nu>x> \<prec> (P' \<parallel> Q)"
proof -
obtain y::name where "y \<sharp> P" and "y \<sharp> Q" and "y \<noteq> a" and "y \<sharp> P'"
by(generate_fresh "name") (auto simp add: fresh_prod)
from \<open>P \<longmapsto>a<\<nu>x> \<prec> P'\<close> \<open>y \<sharp> P'\<close> have "P \<longmapsto>a<\<nu>y> \<prec> ([(x, y)] \<bullet> P')"
by(simp add: alphaBoundOutput)
hence "P \<parallel> Q \<longmapsto>a<\<nu>y> \<prec> (([(x, y)] \<bullet> P') \<parallel> Q)" using \<open>y \<sharp> P\<close> \<open>y \<sharp> Q\<close> \<open>y \<noteq> a\<close>
by(rule Par1B)
with \<open>x \<sharp> Q\<close> \<open>y \<sharp> Q\<close> \<open>y \<sharp> P'\<close> show ?thesis
by(subst alphaBoundOutput) (auto simp add: name_fresh_fresh)
qed
lemma Par2B:
fixes Q :: pi
and a :: name
and x :: name
and Q' :: pi
and P :: pi
assumes "Q \<longmapsto>a<\<nu>x> \<prec> Q'"
and "x \<sharp> P"
shows "P \<parallel> Q \<longmapsto> a<\<nu>x> \<prec> (P \<parallel> Q')"
proof -
obtain y::name where "y \<sharp> P" and "y \<sharp> Q" and "y \<noteq> a" and "y \<sharp> Q'"
by(generate_fresh "name") (auto simp add: fresh_prod)
from \<open>Q \<longmapsto>a<\<nu>x> \<prec> Q'\<close> \<open>y \<sharp> Q'\<close> have "Q \<longmapsto>a<\<nu>y> \<prec> ([(x, y)] \<bullet> Q')"
by(simp add: alphaBoundOutput)
hence "P \<parallel> Q \<longmapsto>a<\<nu>y> \<prec> (P \<parallel> ([(x, y)] \<bullet> Q'))" using \<open>y \<sharp> P\<close> \<open>y \<sharp> Q\<close> \<open>y \<noteq> a\<close>
by(rule Par2B)
with \<open>x \<sharp> P\<close> \<open>y \<sharp> P\<close> \<open>y \<sharp> Q'\<close> show ?thesis
by(subst alphaBoundOutput[of y]) (auto simp add: name_fresh_fresh)
qed
lemma inputInduct[consumes 1, case_names cInput cMatch cMismatch cSum1 cSum2 cPar1 cPar2 cRes cBang]:
fixes P :: pi
and a :: name
and u :: name
and P' :: pi
and F :: "'a::fs_name \<Rightarrow> pi \<Rightarrow> name \<Rightarrow> name \<Rightarrow> pi \<Rightarrow> bool"
and C :: "'a::fs_name"
assumes Trans: "P \<longmapsto>a<u> \<prec> P'"
and "\<And>a x P u C. \<lbrakk>x \<sharp> C; x \<noteq> u; x \<noteq> a\<rbrakk> \<Longrightarrow> F C (a<x>.P) a u (P[x::=u])"
and "\<And>P a u P' b C. \<lbrakk>P \<longmapsto>a<u> \<prec> P'; \<And>C. F C P a u P'\<rbrakk> \<Longrightarrow> F C ([b\<frown>b]P) a u P'"
and "\<And>P a u P' b c C. \<lbrakk>P \<longmapsto>a<u> \<prec> P'; \<And>C. F C P a u P'; b\<noteq>c\<rbrakk> \<Longrightarrow> F C ([b\<noteq>c]P) a u P'"
and "\<And>P a u P' Q C. \<lbrakk>P \<longmapsto>a<u> \<prec> P'; \<And>C. F C P a u P'\<rbrakk> \<Longrightarrow> F C (P \<oplus> Q) a u P'"
and "\<And>Q a u Q' P C. \<lbrakk>Q \<longmapsto>a<u> \<prec> Q'; \<And>C. F C Q a u Q'\<rbrakk> \<Longrightarrow> F C (P \<oplus> Q) a u Q'"
and "\<And>P a u P' Q C. \<lbrakk>P \<longmapsto>a<u> \<prec> P'; \<And>C. F C P a u P'\<rbrakk> \<Longrightarrow> F C (P \<parallel> Q) a u (P' \<parallel> Q)"
and "\<And>Q a u Q' P C. \<lbrakk>Q \<longmapsto>a<u> \<prec> Q'; \<And>C. F C Q a u Q'\<rbrakk> \<Longrightarrow> F C (P \<parallel> Q) a u (P \<parallel> Q')"
and "\<And>P a u P' x C. \<lbrakk>P \<longmapsto>a<u> \<prec> P'; x \<noteq> a; x \<noteq> u; x \<sharp> C; \<And>C. F C P a u P'\<rbrakk> \<Longrightarrow> F C (<\<nu>x>P) a u (<\<nu>x>P')"
and "\<And>P a u P' C. \<lbrakk>P \<parallel> !P \<longmapsto>a<u> \<prec> P'; \<And>C. F C (P \<parallel> !P) a u P'\<rbrakk> \<Longrightarrow> F C (!P) a u P'"
shows "F C P a u P'"
using assms
by(nominal_induct x2=="a<u> \<prec> P'" avoiding: C arbitrary: P' rule: TransitionsEarly.strong_induct)
(auto simp add: residual.inject)
lemma inputAlpha:
assumes "P \<longmapsto>a<u> \<prec> P'"
and "u \<sharp> P"
and "r \<sharp> P'"
shows "P \<longmapsto>a<r> \<prec> ([(u, r)] \<bullet> P')"
using assms
proof(nominal_induct avoiding: r rule: inputInduct)
case(cInput a x P u r)
from \<open>x \<noteq> u\<close> \<open>u \<sharp> a<x>.P\<close>have "u \<noteq> a" and "u \<sharp> P" by(simp add: abs_fresh)+
have "a<x>.P \<longmapsto>a<r> \<prec> P[x::=r]"
by(rule Input)
thus ?case using \<open>r \<sharp> P[x::=u]\<close> \<open>u \<sharp> P\<close>
by(simp add: injPermSubst substTrans)
next
case(cMatch P a u P' b r)
thus ?case by(force intro: Match)
next
case(cMismatch P a u P' b c r)
thus ?case by(force intro: Mismatch)
next
case(cSum1 P a u P' Q r)
thus ?case by(force intro: Sum1)
next
case(cSum2 Q a u Q' P r)
thus ?case by(force intro: Sum2)
next
case(cPar1 P a u P' Q r)
thus ?case by(force intro: Par1F simp add: eqvts name_fresh_fresh)
next
case(cPar2 Q a u Q' P r)
thus ?case by(force intro: Par2F simp add: eqvts name_fresh_fresh)
next
case(cRes P a u P' x r)
thus ?case by(force intro: ResF simp add: eqvts calc_atm abs_fresh)
next
case(cBang P a u P' R)
thus ?case by(force intro: Bang)
qed
lemma Close1:
fixes P :: pi
and a :: name
and x :: name
and P' :: pi
and Q :: pi
and Q' :: pi
assumes "P \<longmapsto>a<x> \<prec> P'"
and "Q \<longmapsto>a<\<nu>x> \<prec> Q'"
and "x \<sharp> P"
shows "P \<parallel> Q \<longmapsto>\<tau> \<prec> <\<nu>x>(P' \<parallel> Q')"
proof -
obtain y::name where "y \<sharp> P" and "y \<sharp> Q" and "y \<noteq> a" and "y \<sharp> Q'" and "y \<sharp> P'"
by(generate_fresh "name") (auto simp add: fresh_prod)
from \<open>P \<longmapsto>a<x> \<prec> P'\<close> \<open>x \<sharp> P\<close> \<open>y \<sharp> P'\<close> have "P \<longmapsto>a<y> \<prec> ([(x, y)] \<bullet> P')"
by(rule inputAlpha)
moreover from \<open>Q \<longmapsto>a<\<nu>x> \<prec> Q'\<close> \<open>y \<sharp> Q'\<close> have "Q \<longmapsto>a<\<nu>y> \<prec> ([(x, y)] \<bullet> Q')"
by(simp add: alphaBoundOutput)
ultimately have "P \<parallel> Q \<longmapsto>\<tau> \<prec> <\<nu>y>(([(x, y)] \<bullet> P') \<parallel> ([(x, y)] \<bullet> Q'))" using \<open>y \<sharp> P\<close> \<open>y \<sharp> Q\<close> \<open>y \<noteq> a\<close>
by(rule Close1)
with \<open>y \<sharp> P'\<close> \<open>y \<sharp> Q'\<close> show ?thesis by(subst alphaRes) (auto simp add: name_fresh_fresh)
qed
lemma Close2:
fixes P :: pi
and a :: name
and x :: name
and P' :: pi
and Q :: pi
and Q' :: pi
assumes "P \<longmapsto>a<\<nu>x> \<prec> P'"
and "Q \<longmapsto>a<x> \<prec> Q'"
and "x \<sharp> Q"
shows "P \<parallel> Q \<longmapsto>\<tau> \<prec> <\<nu>x>(P' \<parallel> Q')"
proof -
obtain y::name where "y \<sharp> P" and "y \<sharp> Q" and "y \<noteq> a" and "y \<sharp> Q'" and "y \<sharp> P'"
by(generate_fresh "name") (auto simp add: fresh_prod)
from \<open>P \<longmapsto>a<\<nu>x> \<prec> P'\<close> \<open>y \<sharp> P'\<close> have "P \<longmapsto>a<\<nu>y> \<prec> ([(x, y)] \<bullet> P')"
by(simp add: alphaBoundOutput)
moreover from \<open>Q \<longmapsto>a<x> \<prec> Q'\<close> \<open>x \<sharp> Q\<close> \<open>y \<sharp> Q'\<close> have "Q \<longmapsto>a<y> \<prec> ([(x, y)] \<bullet> Q')"
by(rule inputAlpha)
ultimately have "P \<parallel> Q \<longmapsto>\<tau> \<prec> <\<nu>y>(([(x, y)] \<bullet> P') \<parallel> ([(x, y)] \<bullet> Q'))" using \<open>y \<sharp> P\<close> \<open>y \<sharp> Q\<close> \<open>y \<noteq> a\<close>
by(rule Close2)
with \<open>y \<sharp> P'\<close> \<open>y \<sharp> Q'\<close> show ?thesis by(subst alphaRes) (auto simp add: name_fresh_fresh)
qed
lemma ResB:
fixes P :: pi
and a :: name
and x :: name
and P' :: pi
and y :: name
assumes "P \<longmapsto>a<\<nu>x> \<prec> P'"
and "y \<noteq> a"
and "y \<noteq> x"
shows "<\<nu>y>P \<longmapsto>a<\<nu>x> \<prec> (<\<nu>y>P')"
proof -
obtain z :: name where "z \<sharp> P" and "z \<sharp> P'" and "z \<noteq> a" and "z \<noteq> y"
by(generate_fresh "name") (auto simp add: fresh_prod)
from \<open>P \<longmapsto>a<\<nu>x> \<prec> P'\<close> \<open>z \<sharp> P'\<close> have "P \<longmapsto>a<\<nu>z> \<prec> ([(x, z)] \<bullet> P')"
by(simp add: alphaBoundOutput)
hence "<\<nu>y>P \<longmapsto>a<\<nu>z> \<prec> (<\<nu>y>([(x, z)] \<bullet> P'))" using \<open>y \<noteq> a\<close> \<open>z \<noteq> y\<close> \<open>z \<sharp> P\<close> \<open>z \<noteq> a\<close>
by(rule_tac ResB) auto
thus ?thesis using \<open>z \<noteq> y\<close> \<open>y \<noteq> x\<close> \<open>z \<sharp> P'\<close>
by(subst alphaBoundOutput[where x'=z]) (auto simp add: eqvts calc_atm abs_fresh)
qed
lemma outputInduct[consumes 1, case_names Output Match Mismatch Sum1 Sum2 Par1 Par2 Res Bang]:
fixes P :: pi
and a :: name
and b :: name
and P' :: pi
and F :: "'a::fs_name \<Rightarrow> pi \<Rightarrow> name \<Rightarrow> name \<Rightarrow> pi \<Rightarrow> bool"
and C :: "'a::fs_name"
assumes Trans: "P \<longmapsto>a[b] \<prec> P'"
and "\<And>a b P C. F C (a{b}.P) a b P"
and "\<And>P a b P' c C. \<lbrakk>P \<longmapsto>OutputR a b \<prec> P'; \<And>C. F C P a b P'\<rbrakk> \<Longrightarrow> F C ([c\<frown>c]P) a b P'"
and "\<And>P a b P' c d C. \<lbrakk>P \<longmapsto>OutputR a b \<prec> P'; \<And>C. F C P a b P'; c\<noteq>d\<rbrakk> \<Longrightarrow> F C ([c\<noteq>d]P) a b P'"
and "\<And>P a b P' Q C. \<lbrakk>P \<longmapsto>OutputR a b \<prec> P'; \<And>C. F C P a b P'\<rbrakk> \<Longrightarrow> F C (P \<oplus> Q) a b P'"
and "\<And>Q a b Q' P C. \<lbrakk>Q \<longmapsto>OutputR a b \<prec> Q'; \<And>C. F C Q a b Q'\<rbrakk> \<Longrightarrow> F C (P \<oplus> Q) a b Q'"
and "\<And>P a b P' Q C. \<lbrakk>P \<longmapsto>OutputR a b \<prec> P'; \<And>C. F C P a b P'\<rbrakk> \<Longrightarrow> F C (P \<parallel> Q) a b (P' \<parallel> Q)"
and "\<And>Q a b Q' P C. \<lbrakk>Q \<longmapsto>OutputR a b \<prec> Q'; \<And>C. F C Q a b Q'\<rbrakk> \<Longrightarrow> F C (P \<parallel> Q) a b (P \<parallel> Q')"
and "\<And>P a b P' x C. \<lbrakk>P \<longmapsto>OutputR a b \<prec> P'; x \<noteq> a; x \<noteq> b; x \<sharp> C; \<And>C. F C P a b P'\<rbrakk> \<Longrightarrow>
F C (<\<nu>x>P) a b (<\<nu>x>P')"
and "\<And>P a b P' C. \<lbrakk>P \<parallel> !P \<longmapsto>OutputR a b \<prec> P'; \<And>C. F C (P \<parallel> !P) a b P'\<rbrakk> \<Longrightarrow> F C (!P) a b P'"
shows "F C P a b P'"
using assms
by(nominal_induct x2=="a[b] \<prec> P'" avoiding: C arbitrary: P' rule: TransitionsEarly.strong_induct)
(auto simp add: residual.inject)
lemma boundOutputInduct[consumes 2, case_names Match Mismatch Open Sum1 Sum2 Par1 Par2 Res Bang]:
fixes P :: pi
and a :: name
and x :: name
and P' :: pi
and F :: "('a::fs_name) \<Rightarrow> pi \<Rightarrow> name \<Rightarrow> name \<Rightarrow> pi \<Rightarrow> bool"
and C :: "'a::fs_name"
assumes a: "P \<longmapsto>a<\<nu>x> \<prec> P'"
and xFreshP: "x \<sharp> P"
and cMatch: "\<And>P a x P' b C. \<lbrakk>P \<longmapsto>a<\<nu>x> \<prec> P'; \<And>C. F C P a x P'\<rbrakk> \<Longrightarrow> F C ([b\<frown>b]P) a x P'"
and cMismatch: "\<And>P a x P' b c C. \<lbrakk>P \<longmapsto>a<\<nu>x> \<prec> P'; \<And>C. F C P a x P'; b \<noteq> c\<rbrakk> \<Longrightarrow> F C ([b\<noteq>c]P) a x P'"
and cOpen: "\<And>P a x P' C. \<lbrakk>P \<longmapsto>(OutputR a x) \<prec> P'; a \<noteq> x\<rbrakk> \<Longrightarrow> F C (<\<nu>x>P) a x P'"
and cSum1: "\<And>P Q a x P' C. \<lbrakk>P \<longmapsto>a<\<nu>x> \<prec> P'; \<And>C. F C P a x P'\<rbrakk> \<Longrightarrow> F C (P \<oplus> Q) a x P'"
and cSum2: "\<And>P Q a x Q' C. \<lbrakk>Q \<longmapsto>a<\<nu>x> \<prec> Q'; \<And>C. F C Q a x Q'\<rbrakk> \<Longrightarrow> F C (P \<oplus> Q) a x Q'"
and cPar1B: "\<And>P P' Q a x C. \<lbrakk>P \<longmapsto>a<\<nu>x> \<prec> P'; x \<sharp> Q; \<And>C. F C P a x P'\<rbrakk> \<Longrightarrow>
F C (P \<parallel> Q) a x (P' \<parallel> Q)"
and cPar2B: "\<And>P Q Q' a x C. \<lbrakk>Q \<longmapsto>a<\<nu>x> \<prec> Q'; x \<sharp> P; \<And>C. F C Q a x Q'\<rbrakk> \<Longrightarrow>
F C (P \<parallel> Q) a x (P \<parallel> Q')"
and cResB: "\<And>P P' a x y C. \<lbrakk>P \<longmapsto>a<\<nu>x> \<prec> P'; y \<noteq> a; y \<noteq> x; y \<sharp> C;
\<And>C. F C P a x P'\<rbrakk> \<Longrightarrow> F C (<\<nu>y>P) a x (<\<nu>y>P')"
and cBang: "\<And>P a x P' C. \<lbrakk>P \<parallel> !P \<longmapsto>a<\<nu>x> \<prec> P'; \<And>C. F C (P \<parallel> !P) a x P'\<rbrakk> \<Longrightarrow>
F C (!P) a x P'"
shows "F C P a x P'"
using assms
proof -
have Goal: "\<And>P Rs a x P' C. \<lbrakk>P \<longmapsto> Rs; Rs = a<\<nu>x> \<prec> P'; x \<sharp> P\<rbrakk> \<Longrightarrow> F C P a x P'"
proof -
fix P Rs a x P' C
assume "P \<longmapsto> Rs" and "Rs = a<\<nu>x> \<prec> P'" and "x \<sharp> P"
thus "F C P a x P'"
proof(nominal_induct avoiding: C a x P' rule: TransitionsEarly.strong_induct)
case(Tau P)
thus ?case by(simp add: residual.inject)
next
case(Input P a x)
thus ?case by(simp add: residual.inject)
next
case(Output P a b)
thus ?case by(simp add: residual.inject)
next
case(Match P Rs b C a x P')
thus ?case
by(force intro: cMatch simp add: residual.inject)
next
case(Mismatch P Rs b c C a x P')
thus ?case
by(force intro: cMismatch simp add: residual.inject)
next
case(Sum1 P Q Rs C)
thus ?case by(force intro: cSum1)
next
case(Sum2 P Q Rs C)
thus ?case by(force intro: cSum2)
next
case(Open P a b P' C a' x P'')
have "b \<sharp> x" by fact hence bineqx: "b \<noteq> x" by simp
moreover have "a<\<nu>b> \<prec> P' = a'<\<nu>x> \<prec> P''" by fact
ultimately have aeqa': "a=a'" and P'eqP'': "P'' = [(b, x)] \<bullet> P'"
by(simp add: residual.inject name_abs_eq)+
have "x \<sharp> <\<nu>b>P" by fact
with bineqx have xFreshP: "x \<sharp> P" by(simp add: name_fresh_abs)
have aineqb: "a \<noteq> b" by fact
have PTrans: "P \<longmapsto>a[b] \<prec> P'" by fact
with xFreshP have xineqa: "x \<noteq> a" by(force dest: freshAction)
from PTrans have "([(b, x)] \<bullet> P) \<longmapsto>[(b, x)] \<bullet> (a[b] \<prec> P')" by(rule TransitionsEarly.eqvt)
with P'eqP'' xineqa aineqb have Trans: "([(b, x)] \<bullet> P) \<longmapsto>a[x] \<prec> P''"
by(auto simp add: name_calc)
hence "F C (<\<nu>x>([(b, x)] \<bullet> P)) a x P''" using xineqa by(blast intro: cOpen)
with xFreshP aeqa' show ?case by(simp add: alphaRes)
next
case(Par1B P a x P' Q C a' x' P'')
have "x \<sharp> x'" by fact hence xineqx': "x \<noteq> x'" by simp
moreover have Eq: "a<\<nu>x> \<prec> (P' \<parallel> Q) = a'<\<nu>x'> \<prec> P''" by fact
hence aeqa': "a = a'" by(simp add: residual.inject)
have xFreshQ: "x \<sharp> Q" by fact
have "x' \<sharp> P \<parallel> Q" by fact
hence x'FreshP: "x' \<sharp> P" and x'FreshQ: "x' \<sharp> Q" by simp+
have P''eq: "P'' = ([(x, x')] \<bullet> P') \<parallel> Q"
proof -
from Eq xineqx' have "(P' \<parallel> Q) = [(x, x')] \<bullet> P''"
by(simp add: residual.inject name_abs_eq)
hence "([(x, x')] \<bullet> (P' \<parallel> Q)) = P''" by simp
with x'FreshQ xFreshQ show ?thesis by(simp add: name_fresh_fresh)
qed
have "x \<sharp> P''" by fact
with P''eq have x'FreshP': "x' \<sharp> P'" by(simp add: name_fresh_left name_calc)
have "P \<longmapsto>a<\<nu>x> \<prec> P'" by fact
with x'FreshP' aeqa' have "P \<longmapsto>a'<\<nu>x'> \<prec> ([(x, x')] \<bullet> P')"
by(simp add: alphaBoundOutput)
moreover have "\<And>C. F C P a x' ([(x, x')] \<bullet> P')"
proof -
fix C
have "\<And>C a' x' P''. \<lbrakk>a<\<nu>x> \<prec> P' = a'<\<nu>x'> \<prec> P''; x' \<sharp> P\<rbrakk> \<Longrightarrow> F C P a' x' P''" by fact
moreover with aeqa' xineqx' x'FreshP' have "a<\<nu>x> \<prec> P' = a'<\<nu>x'> \<prec> ([(x, x')] \<bullet> P')"
by(simp add: residual.inject name_abs_eq name_fresh_left name_calc)
ultimately show "F C P a x' ([(x, x')] \<bullet> P')" using x'FreshP aeqa' by blast
qed
ultimately have "F C (P \<parallel> Q) a' x' (([(x, x')] \<bullet> P') \<parallel> Q)" using x'FreshQ aeqa'
by(blast intro: cPar1B)
with P''eq show ?case by simp
next
case(Par1F P P' Q \<alpha>)
thus ?case by(simp add: residual.inject)
next
case(Par2B Q a x Q' P C a' x' Q'')
have "x \<sharp> x'" by fact hence xineqx': "x \<noteq> x'" by simp
moreover have Eq: "a<\<nu>x> \<prec> (P \<parallel> Q') = a'<\<nu>x'> \<prec> Q''" by fact
hence aeqa': "a = a'" by(simp add: residual.inject)
have xFreshP: "x \<sharp> P" by fact
have "x' \<sharp> P \<parallel> Q" by fact
hence x'FreshP: "x' \<sharp> P" and x'FreshQ: "x' \<sharp> Q" by simp+
have Q''eq: "Q'' = P \<parallel> ([(x, x')] \<bullet> Q')"
proof -
from Eq xineqx' have "(P \<parallel> Q') = [(x, x')] \<bullet> Q''"
by(simp add: residual.inject name_abs_eq)
hence "([(x, x')] \<bullet> (P \<parallel> Q')) = Q''" by simp
with x'FreshP xFreshP show ?thesis by(simp add: name_fresh_fresh)
qed
have "x \<sharp> Q''" by fact
with Q''eq have x'FreshQ': "x' \<sharp> Q'" by(simp add: name_fresh_left name_calc)
have "Q \<longmapsto>a<\<nu>x> \<prec> Q'" by fact
with x'FreshQ' aeqa' have "Q \<longmapsto>a'<\<nu>x'> \<prec> ([(x, x')] \<bullet> Q')"
by(simp add: alphaBoundOutput)
moreover have "\<And>C. F C Q a x' ([(x, x')] \<bullet> Q')"
proof -
fix C
have "\<And>C a' x' Q''. \<lbrakk>a<\<nu>x> \<prec> Q' = a'<\<nu>x'> \<prec> Q''; x' \<sharp> Q\<rbrakk> \<Longrightarrow> F C Q a' x' Q''" by fact
moreover with aeqa' xineqx' x'FreshQ' have "a<\<nu>x> \<prec> Q' = a'<\<nu>x'> \<prec> ([(x, x')] \<bullet> Q')"
by(simp add: residual.inject name_abs_eq name_fresh_left name_calc)
ultimately show "F C Q a x' ([(x, x')] \<bullet> Q')" using x'FreshQ aeqa' by blast
qed
ultimately have "F C (P \<parallel> Q) a' x' (P \<parallel> ([(x, x')] \<bullet> Q'))" using x'FreshP aeqa'
by(blast intro: cPar2B)
with Q''eq show ?case by simp
next
case(Par2F P P' Q \<alpha>)
thus ?case by(simp add: residual.inject)
next
case(Comm1 P P' Q Q' a b x)
thus ?case by(simp add: residual.inject)
next
case(Comm2 P P' Q Q' a b x)
thus ?case by(simp add: residual.inject)
next
case(Close1 P P' Q Q' a x y)
thus ?case by(simp add: residual.inject)
next
case(Close2 P P' Q Q' a x y)
thus ?case by(simp add: residual.inject)
next
case(ResB P a x P' y C a' x' P'')
have "x \<sharp> x'" by fact hence xineqx': "x \<noteq> x'" by simp
moreover have Eq: "a<\<nu>x> \<prec> (<\<nu>y>P') = a'<\<nu>x'> \<prec> P''" by fact
hence aeqa': "a = a'" by(simp add: residual.inject)
have "y \<sharp> x'" by fact hence yineqx': "y \<noteq> x'" by simp
moreover have "x' \<sharp> <\<nu>y>P" by fact
ultimately have x'FreshP: "x' \<sharp> P" by(simp add: name_fresh_abs)
have yineqx: "y \<noteq> x" and yineqa: "y \<noteq> a" and yFreshC: "y \<sharp> C" by fact+
have P''eq: "P'' = <\<nu>y>([(x, x')] \<bullet> P')"
proof -
from Eq xineqx' have "<\<nu>y>P' = [(x, x')] \<bullet> P''"
by(simp add: residual.inject name_abs_eq)
hence "([(x, x')] \<bullet> (<\<nu>y>P')) = P''" by simp
with yineqx' yineqx show ?thesis by(simp add: name_fresh_fresh)
qed
have "x \<sharp> P''" by fact
with P''eq yineqx have x'FreshP': "x' \<sharp> P'" by(simp add: name_fresh_left name_calc name_fresh_abs)
have "P \<longmapsto>a<\<nu>x> \<prec> P'" by fact
with x'FreshP' aeqa' have "P \<longmapsto>a'<\<nu>x'> \<prec> ([(x, x')] \<bullet> P')"
by(simp add: alphaBoundOutput)
moreover have "\<And>C. F C P a x' ([(x, x')] \<bullet> P')"
proof -
fix C
have "\<And>C a' x' P''. \<lbrakk>a<\<nu>x> \<prec> P' = a'<\<nu>x'> \<prec> P''; x' \<sharp> P\<rbrakk> \<Longrightarrow> F C P a' x' P''" by fact
moreover with aeqa' xineqx' x'FreshP' have "a<\<nu>x> \<prec> P' = a'<\<nu>x'> \<prec> ([(x, x')] \<bullet> P')"
by(simp add: residual.inject name_abs_eq name_fresh_left name_calc)
ultimately show "F C P a x' ([(x, x')] \<bullet> P')" using x'FreshP aeqa' by blast
qed
ultimately have "F C (<\<nu>y>P) a' x' (<\<nu>y>([(x, x')] \<bullet> P'))" using yineqx' yineqa yFreshC aeqa'
by(force intro: cResB)
with P''eq show ?case by simp
next
case(ResF P P' \<alpha> y)
thus ?case by(simp add: residual.inject)
next
case(Bang P Rs)
thus ?case by(force intro: cBang)
qed
qed
with a xFreshP show ?thesis by simp
qed
lemma tauInduct[consumes 1, case_names Tau Match Mismatch Sum1 Sum2 Par1 Par2 Comm1 Comm2 Close1 Close2 Res Bang]:
fixes P :: pi
and P' :: pi
and F :: "'a::fs_name \<Rightarrow> pi \<Rightarrow> pi \<Rightarrow> bool"
and C :: "'a::fs_name"
assumes Trans: "P \<longmapsto>\<tau> \<prec> P'"
and "\<And>P C. F C (\<tau>.(P)) P"
and "\<And>P P' a C. \<lbrakk>P \<longmapsto>\<tau> \<prec> P'; \<And>C. F C P P'\<rbrakk> \<Longrightarrow> F C ([a\<frown>a]P) P'"
and "\<And>P P' a b C. \<lbrakk>P \<longmapsto>\<tau> \<prec> P'; \<And>C. F C P P'; a \<noteq> b\<rbrakk> \<Longrightarrow> F C ([a\<noteq>b]P) P'"
and "\<And>P P' Q C. \<lbrakk>P \<longmapsto>\<tau> \<prec> P'; \<And>C. F C P P'\<rbrakk> \<Longrightarrow> F C (P \<oplus> Q) P'"
and "\<And>Q Q' P C. \<lbrakk>Q \<longmapsto>\<tau> \<prec> Q'; \<And>C. F C Q Q'\<rbrakk> \<Longrightarrow> F C (P \<oplus> Q) Q'"
and "\<And>P P' Q C. \<lbrakk>P \<longmapsto>\<tau> \<prec> P'; \<And>C. F C P P'\<rbrakk> \<Longrightarrow> F C (P \<parallel> Q) (P' \<parallel> Q)"
and "\<And>Q Q' P C. \<lbrakk>Q \<longmapsto>\<tau> \<prec> Q'; \<And>C. F C Q Q'\<rbrakk> \<Longrightarrow> F C (P \<parallel> Q) (P \<parallel> Q')"
and "\<And>P a b P' Q Q' C. \<lbrakk>P \<longmapsto>a<b> \<prec> P'; Q \<longmapsto>OutputR a b \<prec> Q'\<rbrakk> \<Longrightarrow> F C (P \<parallel> Q) (P' \<parallel> Q')"
and "\<And>P a b P' Q Q' C. \<lbrakk>P \<longmapsto>OutputR a b \<prec> P'; Q \<longmapsto>a<b> \<prec> Q'\<rbrakk> \<Longrightarrow> F C (P \<parallel> Q) (P' \<parallel> Q')"
and "\<And>P a x P' Q Q' C. \<lbrakk>P \<longmapsto>a<x> \<prec> P'; Q \<longmapsto>a<\<nu>x> \<prec> Q'; x \<sharp> P; x \<sharp> Q; x \<noteq> a; x \<sharp> C\<rbrakk> \<Longrightarrow> F C (P \<parallel> Q) (<\<nu>x>(P' \<parallel> Q'))"
and "\<And>P a x P' Q Q' C. \<lbrakk>P \<longmapsto>a<\<nu>x> \<prec> P'; Q \<longmapsto>a<x> \<prec> Q'; x \<sharp> P; x \<sharp> Q; x \<noteq> a; x \<sharp> C\<rbrakk> \<Longrightarrow> F C (P \<parallel> Q) (<\<nu>x>(P' \<parallel> Q'))"
and "\<And>P P' x C. \<lbrakk>P \<longmapsto>\<tau> \<prec> P'; x \<sharp> C; \<And>C. F C P P'\<rbrakk> \<Longrightarrow>
F C (<\<nu>x>P) (<\<nu>x>P')"
and "\<And>P P' C. \<lbrakk>P \<parallel> !P \<longmapsto>\<tau> \<prec> P'; \<And>C. F C (P \<parallel> !P) P'\<rbrakk> \<Longrightarrow> F C (!P) P'"
shows "F C P P'"
using \<open>P \<longmapsto>\<tau> \<prec> P'\<close>
by(nominal_induct x2=="\<tau> \<prec> P'" avoiding: C arbitrary: P' rule: TransitionsEarly.strong_induct)
(auto simp add: residual.inject intro: assms)+
inductive bangPred :: "pi \<Rightarrow> pi \<Rightarrow> bool"
where
aux1: "bangPred P (!P)"
| aux2: "bangPred P (P \<parallel> !P)"
inductive_cases tauCases'[simplified pi.distinct residual.distinct]: "\<tau>.(P) \<longmapsto> Rs"
inductive_cases inputCases'[simplified pi.inject residual.inject]: "a<b>.P \<longmapsto> Rs"
inductive_cases outputCases'[simplified pi.inject residual.inject]: "a{b}.P \<longmapsto> Rs"
inductive_cases matchCases'[simplified pi.inject residual.inject]: "[a\<frown>b]P \<longmapsto> Rs"
inductive_cases mismatchCases'[simplified pi.inject residual.inject]: "[a\<noteq>b]P \<longmapsto> Rs"
inductive_cases sumCases'[simplified pi.inject residual.inject]: "P \<oplus> Q \<longmapsto> Rs"
inductive_cases parCasesB'[simplified pi.distinct residual.distinct]: "A \<parallel> B \<longmapsto> b<\<nu>y> \<prec> A'"
inductive_cases parCasesF'[simplified pi.distinct residual.distinct]: "P \<parallel> Q \<longmapsto> \<alpha> \<prec> P'"
inductive_cases resCasesB'[simplified pi.distinct residual.distinct]: "<\<nu>x'>A \<longmapsto> a<\<nu>y'> \<prec> A'"
inductive_cases resCasesF'[simplified pi.distinct residual.distinct]: "<\<nu>x>A \<longmapsto> \<alpha> \<prec> A'"
lemma tauCases:
fixes P :: pi
and \<alpha> :: freeRes
and P' :: pi
assumes "\<tau>.(P) \<longmapsto>\<alpha> \<prec> P'"
and "Prop (\<tau>) P"
shows "Prop \<alpha> P'"
using assms
by(cases rule: tauCases') (auto simp add: pi.inject residual.inject)
lemma inputCases[consumes 1, case_names cInput]:
fixes a :: name
and x :: name
and P :: pi
and P' :: pi
assumes Input: "a<x>.P \<longmapsto>\<alpha> \<prec> P'"
and A: "\<And>u. Prop (a<u>) (P[x::=u])"
shows "Prop \<alpha> P'"
proof -
{
fix x P
assume "a<x>.P \<longmapsto>\<alpha> \<prec> P'"
moreover assume "(x::name) \<sharp> \<alpha>" and "x \<sharp> P'" and "x \<noteq> a"
moreover assume "\<And>u. Prop (a<u>) (P[x::=u])"
moreover obtain z::name where "z \<noteq> x" and "z \<sharp> P" and "z \<sharp> \<alpha>" and "z \<sharp> P'" and "z \<noteq> a"
by(generate_fresh "name", auto simp add: fresh_prod)
moreover obtain z'::name where "z' \<noteq> x" and "z' \<noteq> z" and "z' \<sharp> P" and "z' \<sharp> \<alpha>" and "z' \<sharp> P'" and "z' \<noteq> a"
by(generate_fresh "name", auto simp add: fresh_prod)
ultimately have "Prop \<alpha> P'"
by(cases rule: TransitionsEarly.strong_cases[where x=x and b=z and xa=z and xb=z and xc=z and xd=z and xe=z
and y=z' and ya=z'])
(auto simp add: pi.inject residual.inject abs_fresh alpha)
}
note Goal = this
obtain y::name where "y \<sharp> P" and "y \<sharp> \<alpha>" and "y \<sharp> P'" and "y \<noteq> a"
by(generate_fresh "name") (auto simp add: fresh_prod)
from Input \<open>y \<sharp> P\<close> have "a<y>.([(x, y)] \<bullet> P) \<longmapsto>\<alpha> \<prec> P'" by(simp add: alphaInput)
moreover note \<open>y \<sharp> \<alpha>\<close> \<open>y \<sharp> P'\<close> \<open>y \<noteq> a\<close>
moreover from A \<open>y \<sharp> P\<close> have "\<And>u. Prop (a<u>) (([(x, y)] \<bullet> P)[y::=u])"
by(simp add: renaming name_swap)
ultimately show ?thesis by(rule Goal)
qed
lemma outputCases:
fixes P :: pi
and \<alpha> :: freeRes
and P' :: pi
assumes "a{b}.P \<longmapsto>\<alpha> \<prec> P'"
and "Prop (OutputR a b) P"
shows "Prop \<alpha> P'"
using assms
by(cases rule: outputCases') (auto simp add: pi.inject residual.inject)
lemma zeroTrans[dest]:
fixes Rs :: residual
assumes "\<zero> \<longmapsto> e Rs"
shows "False"
using assms
by - (ind_cases "\<zero> \<longmapsto> e Rs")
lemma mismatchTrans[dest]:
fixes a :: name
and P :: pi
and Rs :: residual
assumes "[a\<noteq>a]P \<longmapsto> Rs"
shows "False"
using assms
by(erule_tac mismatchCases') auto
lemma matchCases[consumes 1, case_names Match]:
fixes a :: name
and b :: name
and P :: pi
and Rs :: residual
and F :: "name \<Rightarrow> name \<Rightarrow> bool"
assumes Trans: "[a\<frown>b]P \<longmapsto> Rs"
and cMatch: "P \<longmapsto> Rs \<Longrightarrow> F a a"
shows "F a b"
using assms
by(erule_tac matchCases', auto)
lemma mismatchCases[consumes 1, case_names Mismatch]:
fixes a :: name
and b :: name
and P :: pi
and Rs :: residual
and F :: "name \<Rightarrow> name \<Rightarrow> bool"
assumes Trans: "[a\<noteq>b]P \<longmapsto> Rs"
and cMatch: "\<lbrakk>P \<longmapsto> Rs; a \<noteq> b\<rbrakk> \<Longrightarrow> F a b"
shows "F a b"
using assms
by(erule_tac mismatchCases') auto
lemma sumCases[consumes 1, case_names Sum1 Sum2]:
fixes P :: pi
and Q :: pi
and Rs :: residual
assumes Trans: "P \<oplus> Q \<longmapsto> Rs"
and cSum1: "P \<longmapsto> Rs \<Longrightarrow> F"
and cSum2: "Q \<longmapsto> Rs \<Longrightarrow> F"
shows F
using assms
by(erule_tac sumCases') auto
lemma parCasesB[consumes 1, case_names cPar1 cPar2]:
fixes P :: pi
and Q :: pi
and a :: name
and x :: name
and PQ' :: pi
assumes Trans: "P \<parallel> Q \<longmapsto> a<\<nu>x> \<prec> PQ'"
and icPar1B: "\<And>P'. \<lbrakk>P \<longmapsto> a<\<nu>x> \<prec> P'; x \<sharp> Q\<rbrakk> \<Longrightarrow> F (P' \<parallel> Q)"
and icPar2B: "\<And>Q'. \<lbrakk>Q \<longmapsto> a<\<nu>x> \<prec> Q'; x \<sharp> P\<rbrakk> \<Longrightarrow> F (P \<parallel> Q')"
shows "F PQ'"
proof -
from Trans show ?thesis
proof(induct rule: parCasesB', auto simp add: pi.inject residual.inject)
fix P' y
assume PTrans: "P \<longmapsto> a<\<nu>y> \<prec> P'"
assume yFreshQ: "y \<sharp> (Q::pi)"
assume absEq: "[x].PQ' = [y].(P' \<parallel> Q)"
have "\<exists>c::name. c \<sharp> (P', x, y, Q)" by(blast intro: name_exists_fresh)
then obtain c where cFreshP': "c \<sharp> P'" and cineqx: "x \<noteq> c" and cineqy: "c \<noteq> y" and cFreshQ: "c \<sharp> Q"
by(force simp add: fresh_prod name_fresh)
from cFreshP' PTrans have Trans: "P \<longmapsto> a<\<nu>c> \<prec> ([(y, c)] \<bullet> P')" by(simp add: alphaBoundOutput)
from cFreshP' cFreshQ have "c \<sharp> P' \<parallel> Q" by simp
hence "[y].(P' \<parallel> Q) = [c].([(y, c)] \<bullet> (P' \<parallel> Q))"
by(auto simp add: alpha fresh_left calc_atm)
with yFreshQ cFreshQ have "[y].(P' \<parallel> Q) = [c].(([(y, c)] \<bullet> P') \<parallel> Q)"
by(simp add: name_fresh_fresh)
with cineqx absEq have L1: "PQ' = [(x, c)] \<bullet> (([(y, c)] \<bullet> P') \<parallel> Q)" and L2: "x \<sharp> ([(y, c)] \<bullet> P') \<parallel> Q"
by(simp add: name_abs_eq)+
from L2 have xFreshQ: "x \<sharp> Q" and xFreshP': "x \<sharp> [(y, c)] \<bullet> P'" by simp+
with cFreshQ L1 have L3: "PQ' = ([(x, c)] \<bullet> [(y, c)] \<bullet> P') \<parallel> Q" by(simp add: name_fresh_fresh)
from Trans xFreshP' have "P \<longmapsto> a<\<nu>x> \<prec> ([(x, c)] \<bullet> [(y, c)] \<bullet> P')" by(simp add: alphaBoundOutput name_swap)
thus ?thesis using xFreshQ L3
by(blast intro: icPar1B)
next
fix Q' y
assume QTrans: "Q \<longmapsto> a<\<nu>y> \<prec> Q'"
assume yFreshP: "y \<sharp> (P::pi)"
assume absEq: "[x].PQ' = [y].(P \<parallel> Q')"
have "\<exists>c::name. c \<sharp> (Q', x, y, P)" by(blast intro: name_exists_fresh)
then obtain c where cFreshQ': "c \<sharp> Q'" and cineqx: "x \<noteq> c" and cineqy: "c \<noteq> y" and cFreshP: "c \<sharp> P"
by(force simp add: fresh_prod name_fresh)
from cFreshQ' QTrans have Trans: "Q \<longmapsto> a<\<nu>c> \<prec> ([(y, c)] \<bullet> Q')" by(simp add: alphaBoundOutput)
from cFreshQ' cFreshP have "c \<sharp> P \<parallel> Q'" by simp
hence "[y].(P \<parallel> Q') = [c].([(y, c)] \<bullet> (P \<parallel> Q'))"
by(auto simp add: alpha fresh_left calc_atm)
with yFreshP cFreshP have "[y].(P \<parallel> Q') = [c].(P \<parallel> ([(y, c)] \<bullet> Q'))"
by(simp add: name_fresh_fresh)
with cineqx absEq have L1: "PQ' = [(x, c)] \<bullet> (P \<parallel> ([(y, c)] \<bullet> Q'))" and L2: "x \<sharp> P \<parallel> ([(y, c)] \<bullet> Q')"
by(simp add: name_abs_eq)+
from L2 have xFreshP: "x \<sharp> P" and xFreshQ': "x \<sharp> [(y, c)] \<bullet> Q'" by simp+
with cFreshP L1 have L3: "PQ' = P \<parallel> ([(x, c)] \<bullet> [(y, c)] \<bullet> Q')" by(simp add: name_fresh_fresh)
from Trans xFreshQ' have "Q \<longmapsto> a<\<nu>x> \<prec> ([(x, c)] \<bullet> [(y, c)] \<bullet> Q')" by(simp add: alphaBoundOutput name_swap)
thus ?thesis using xFreshP L3
by(blast intro: icPar2B)
qed
qed
lemma parCasesOutput[consumes 1, case_names Par1 Par2]:
fixes P :: pi
and Q :: pi
and a :: name
and b :: name
and P' :: pi
assumes "P \<parallel> Q \<longmapsto>a[b] \<prec> PQ'"
and "\<And>P'. \<lbrakk>P \<longmapsto>a[b] \<prec> P'\<rbrakk> \<Longrightarrow> F (P' \<parallel> Q)"
and "\<And>Q'. \<lbrakk>Q \<longmapsto>a[b] \<prec> Q'\<rbrakk> \<Longrightarrow> F (P \<parallel> Q')"
shows "F PQ'"
using assms
by(erule_tac parCasesF', auto simp add: pi.inject residual.inject)
lemma parCasesInput[consumes 1, case_names Par1 Par2]:
fixes P :: pi
and Q :: pi
and a :: name
and b :: name
and P' :: pi
assumes Trans: "P \<parallel> Q \<longmapsto>a<b> \<prec> PQ'"
and icPar1F: "\<And>P'. \<lbrakk>P \<longmapsto>a<b> \<prec> P'\<rbrakk> \<Longrightarrow> F (P' \<parallel> Q)"
and icPar2F: "\<And>Q'. \<lbrakk>Q \<longmapsto>a<b> \<prec> Q'\<rbrakk> \<Longrightarrow> F (P \<parallel> Q')"
shows "F PQ'"
using assms
by(erule_tac parCasesF') (auto simp add: pi.inject residual.inject)
lemma parCasesF[consumes 1, case_names cPar1 cPar2 cComm1 cComm2 cClose1 cClose2]:
fixes P :: pi
and Q :: pi
and \<alpha> :: freeRes
and P' :: pi
and C :: "'a::fs_name"
assumes Trans: "P \<parallel> Q \<longmapsto> \<alpha> \<prec> PQ'"
and icPar1F: "\<And>P'. \<lbrakk>P \<longmapsto> \<alpha> \<prec> P'\<rbrakk> \<Longrightarrow> F \<alpha> (P' \<parallel> Q)"
and icPar2F: "\<And>Q'. \<lbrakk>Q \<longmapsto> \<alpha> \<prec> Q'\<rbrakk> \<Longrightarrow> F \<alpha> (P \<parallel> Q')"
and icComm1: "\<And>P' Q' a b. \<lbrakk>P \<longmapsto> a<b> \<prec> P'; Q \<longmapsto> a[b] \<prec> Q'\<rbrakk> \<Longrightarrow> F (\<tau>) (P' \<parallel> Q')"
and icComm2: "\<And>P' Q' a b. \<lbrakk>P \<longmapsto> a[b] \<prec> P'; Q \<longmapsto> a<b> \<prec> Q'\<rbrakk> \<Longrightarrow> F (\<tau>) (P' \<parallel> Q')"
and icClose1: "\<And>P' Q' a x. \<lbrakk>P \<longmapsto> a<x> \<prec> P'; Q \<longmapsto> a<\<nu>x> \<prec> Q'; x \<sharp> P; x \<sharp> C\<rbrakk> \<Longrightarrow> F (\<tau>) (<\<nu>x>(P' \<parallel> Q'))"
and icClose2: "\<And>P' Q' a x. \<lbrakk>P \<longmapsto> a<\<nu>x> \<prec> P'; Q \<longmapsto> a<x> \<prec> Q'; x \<sharp> Q; x \<sharp> C\<rbrakk> \<Longrightarrow> F (\<tau>) (<\<nu>x>(P' \<parallel> Q'))"
shows "F \<alpha> PQ'"
proof -
from Trans show ?thesis
proof(rule parCasesF', auto)
fix Pa Pa' Qa \<alpha>'
assume Trans': "Pa \<longmapsto> \<alpha>' \<prec> Pa'"
assume Eq: "P \<parallel> Q = Pa \<parallel> Qa"
assume Eq': "\<alpha> \<prec> PQ' = \<alpha>' \<prec> Pa' \<parallel> Qa"
from Eq have "P = Pa" and "Q = Qa"
by(simp add: pi.inject)+
moreover with Eq' have "\<alpha> = \<alpha>'" and "PQ' = Pa' \<parallel> Q"
by(simp add: residual.inject)+
ultimately show ?thesis using icPar1F Trans'
by simp
next
fix Pa Qa Qa' \<alpha>'
assume Trans': "Qa \<longmapsto> \<alpha>' \<prec> Qa'"
assume Eq: "P \<parallel> Q = Pa \<parallel> Qa"
assume Eq': "\<alpha> \<prec> PQ' = \<alpha>' \<prec> Pa \<parallel> Qa'"
from Eq have "P = Pa" and "Q = Qa"
by(simp add: pi.inject)+
moreover with Eq' have "\<alpha> = \<alpha>'" and "PQ' = P \<parallel> Qa'"
by(simp add: residual.inject)+
ultimately show ?thesis using icPar2F Trans'
by simp
next
fix Pa Pa' Qa Qa' a b
assume TransP: "Pa \<longmapsto> a<b> \<prec> Pa'"
assume TransQ: "Qa \<longmapsto> a[b] \<prec> Qa'"
assume Eq: "P \<parallel> Q = Pa \<parallel> Qa"
assume Eq': "\<alpha> \<prec> PQ' = \<tau> \<prec> Pa' \<parallel> Qa'"
from TransP TransQ Eq Eq' icComm1 show ?thesis
by(simp add: pi.inject residual.inject)
next
fix Pa Pa' Qa Qa' a b x
assume TransP: "Pa \<longmapsto> (a::name)[b] \<prec> Pa'"
assume TransQ: "Qa \<longmapsto> a<b> \<prec> Qa'"
assume Eq: "P \<parallel> Q = Pa \<parallel> Qa"
assume Eq': "\<alpha> \<prec> PQ' = \<tau> \<prec> Pa' \<parallel> Qa'"
from TransP TransQ Eq Eq' icComm2 show ?thesis
by(simp add: pi.inject residual.inject)
next
fix Pa Pa' Qa Qa' a x
assume TransP: "Pa \<longmapsto> a<x> \<prec> Pa'"
assume TransQ: "Qa \<longmapsto> a<\<nu>x> \<prec> Qa'"
assume xFreshPa: "x \<sharp> Pa"
assume Eq: "P \<parallel> Q = Pa \<parallel> Qa"
assume Eq': "\<alpha> \<prec> PQ' = \<tau> \<prec> <\<nu>x>(Pa' \<parallel> Qa')"
have "\<exists>(c::name). c \<sharp> (Pa, Pa', x, Qa', a, C)"
by(blast intro: name_exists_fresh)
then obtain c::name where cFreshPa: "c \<sharp> Pa" and cFreshPa': "c \<sharp> Pa'" and cineqy: "c \<noteq> x" and cFreshQa': "c \<sharp> Qa'" and cFreshC: "c \<sharp> C" and cineqa: "c \<noteq> a"
by(force simp add: fresh_prod name_fresh)
from cFreshQa' have L1: "a<\<nu>x> \<prec> Qa' = a<\<nu>c> \<prec> ([(x, c)] \<bullet> Qa')"
by(simp add: alphaBoundOutput)
with cFreshQa' cFreshPa' have "c \<sharp> (Pa' \<parallel> Qa')"
by simp
then have L4: "<\<nu>x>(Pa' \<parallel> Qa') = <\<nu>c>(([(x, c)] \<bullet> Pa') \<parallel> ([(x, c)] \<bullet> Qa'))"
by(simp add: alphaRes)
have TransP: "Pa \<longmapsto> a<c> \<prec> [(x, c)] \<bullet> Pa'"
proof -
from xFreshPa TransP have xineqa: "x\<noteq>a" by(force dest: freshAction)
from TransP have "([(x, c)] \<bullet> Pa) \<longmapsto> [(x, c)] \<bullet> (a<x> \<prec> Pa')"
by(rule TransitionsEarly.eqvt)
with xineqa xFreshPa cFreshPa cineqa show ?thesis
by(simp add: name_fresh_fresh name_calc)
qed
with TransQ L1 L4 icClose1 Eq Eq' cFreshPa cFreshC show ?thesis
by(simp add: residual.inject, simp add: pi.inject)
next
fix Pa Pa' Qa Qa' a x
assume TransP: "Pa \<longmapsto> a<\<nu>x> \<prec> Pa'"
assume TransQ: "Qa \<longmapsto> a<x> \<prec> Qa'"
assume xFreshQa: "x \<sharp> Qa"
assume Eq: "P \<parallel> Q = Pa \<parallel> Qa"
assume Eq': "\<alpha> \<prec> PQ' = \<tau> \<prec> <\<nu>x>(Pa' \<parallel> Qa')"
have "\<exists>(c::name). c \<sharp> (Qa, Pa', x, Qa', a, C)"
by(blast intro: name_exists_fresh)
then obtain c::name where cFreshQa: "c \<sharp> Qa" and cFreshPa': "c \<sharp> Pa'" and cineqy: "c \<noteq> x" and cFreshQa': "c \<sharp> Qa'" and cFreshC: "c \<sharp> C" and cineqa: "c \<noteq> a"
by(force simp add: fresh_prod name_fresh)
from cFreshPa' have L1: "a<\<nu>x> \<prec> Pa' = a<\<nu>c> \<prec> ([(x, c)] \<bullet> Pa')"
by(simp add: alphaBoundOutput)
with cFreshQa' cFreshPa' have "c \<sharp> (Pa' \<parallel> Qa')"
by simp
then have L4: "<\<nu>x>(Pa' \<parallel> Qa') = <\<nu>c>(([(x, c)] \<bullet> Pa') \<parallel> ([(x, c)] \<bullet> Qa'))"
by(simp add: alphaRes)
have TransQ: "Qa \<longmapsto> a<c> \<prec> [(x, c)] \<bullet> Qa'"
proof -
from xFreshQa TransQ have xineqa: "x\<noteq>a" by(force dest: freshAction)
from TransQ have "([(x, c)] \<bullet> Qa) \<longmapsto> [(x, c)] \<bullet> (a<x> \<prec> Qa')"
by(rule TransitionsEarly.eqvt)
with xineqa xFreshQa cFreshQa cineqa show ?thesis
by(simp add: name_fresh_fresh name_calc)
qed
with TransP L1 L4 icClose2 Eq Eq' cFreshQa cFreshC show ?thesis
by(simp add: residual.inject, simp add: pi.inject)
qed
qed
lemma resCasesF[consumes 2, case_names Res]:
fixes x :: name
and P :: pi
and \<alpha> :: freeRes
and P' :: pi
assumes Trans: "<\<nu>x>P \<longmapsto> \<alpha> \<prec> RP'"
and xFreshAlpha: "x \<sharp> \<alpha>"
and rcResF: " \<And>P'. P \<longmapsto> \<alpha> \<prec> P' \<Longrightarrow> F (<\<nu>x>P')"
shows "F RP'"
proof -
from Trans show ?thesis
proof(induct rule: resCasesF', auto)
fix Pa Pa' \<beta> y
assume PTrans: "Pa \<longmapsto> \<beta> \<prec> Pa'"
assume yFreshBeta: "(y::name) \<sharp> \<beta>"
assume TermEq: "<\<nu>x>P = <\<nu>y>Pa"
assume ResEq: "\<alpha> \<prec> RP' = \<beta> \<prec> <\<nu>y>Pa'"
hence alphaeqbeta: "\<alpha> = \<beta>" and L2: "RP' = <\<nu>y>Pa'" by(simp add: residual.inject)+
have "\<exists>(c::name). c \<sharp> (Pa, \<alpha>, Pa', x, y)" by(blast intro: name_exists_fresh)
then obtain c::name where cFreshPa: "c \<sharp> Pa" and cFreshAlpha: "c \<sharp> \<alpha>" and cFreshPa': "c \<sharp> Pa'" and cineqx: "x \<noteq> c" and cineqy: "c \<noteq> y"
by(force simp add: fresh_prod name_fresh)
from cFreshPa have "<\<nu>y>Pa = <\<nu>c>([(y, c)] \<bullet> Pa)" by(rule alphaRes)
with TermEq cineqx have Peq: "P = [(x, c)] \<bullet> [(y, c)] \<bullet> Pa" and xeq: "x \<sharp> [(y, c)] \<bullet> Pa"
by(simp add: pi.inject name_abs_eq)+
from PTrans have "([(y, c)] \<bullet> Pa) \<longmapsto> [(y, c)] \<bullet> (\<beta> \<prec> Pa')" by(rule TransitionsEarly.eqvt)
with yFreshBeta cFreshAlpha alphaeqbeta have PTrans': "([(y, c)] \<bullet> Pa) \<longmapsto> \<alpha> \<prec> ([(y, c)] \<bullet> Pa')"
by(simp add: name_fresh_fresh)
from PTrans' have "([(x, c)] \<bullet> [(y, c)] \<bullet> Pa) \<longmapsto> [(x, c)] \<bullet> (\<alpha> \<prec> [(y, c)] \<bullet> Pa')"
by(rule TransitionsEarly.eqvt)
with xFreshAlpha cFreshAlpha Peq have PTrans'': "P \<longmapsto> \<alpha> \<prec> [(x, c)] \<bullet> [(y, c)] \<bullet> Pa'"
by(simp add: name_fresh_fresh)
from PTrans' xeq xFreshAlpha have xeq': "x \<sharp> [(y, c)] \<bullet> Pa'"
by(nominal_induct \<alpha> rule: freeRes.strong_induct)
(auto simp add: fresh_left calc_atm eqvts dest: freshTransition)
from cFreshPa' have "<\<nu>y>Pa' = <\<nu>c>([(y, c)] \<bullet> Pa')" by(rule alphaRes)
moreover from xeq' have "<\<nu>c>([(y, c)] \<bullet> Pa') = <\<nu>x>([(c, x)] \<bullet> [(y, c)] \<bullet> Pa')"
by(rule alphaRes)
ultimately have "RP' = <\<nu>x>([(x, c)] \<bullet> [(y, c)] \<bullet> Pa')" using ResEq
by(simp add: residual.inject name_swap)
with PTrans'' xFreshAlpha show ?thesis
by(blast intro: rcResF)
qed
qed
lemma resCasesB[consumes 2, case_names Open Res]:
fixes x :: name
and P :: pi
and a :: name
and y :: name
and RP' :: pi
assumes Trans: "<\<nu>y>P \<longmapsto> a<\<nu>x> \<prec> RP'"
and xineqy: "x \<noteq> y"
and rcOpen: "\<And>P'. \<lbrakk>P \<longmapsto>(OutputR a y) \<prec> P'; a \<noteq> y\<rbrakk> \<Longrightarrow> F ([(x, y)] \<bullet> P')"
and rcResB: "\<And>P'. \<lbrakk>P \<longmapsto>a<\<nu>x> \<prec> P'; y \<noteq> a\<rbrakk> \<Longrightarrow> F (<\<nu>y>P')"
shows "F RP'"
proof -
from Trans show ?thesis
proof(induct rule: resCasesB', auto)
fix Pa Pa' aa b
assume PTrans: "Pa \<longmapsto> (aa::name)[b] \<prec> Pa'"
assume aaineqb: "aa\<noteq>b"
assume TermEq: "<\<nu>y>P = <\<nu>b>Pa"
assume ResEq: "a<\<nu>x> \<prec> RP' = aa<\<nu>b> \<prec> Pa'"
have "\<exists>(c::name). c \<sharp> (x, a, aa, y, Pa, Pa', b)" by(blast intro: name_exists_fresh)
then obtain c where cineqx: "c\<noteq>x" and cFresha: "c \<sharp> a" and cineqy: "c \<noteq> y" and cineqaa: "c \<noteq> aa" and cFreshPa: "c \<sharp> Pa" and cFreshPa': "c \<sharp> Pa'" and cineqb: "c \<noteq> b"
by(force simp add: fresh_prod name_fresh)
from cFreshPa have "<\<nu>b>Pa = <\<nu>c>([(b, c)] \<bullet> Pa)" by(rule alphaRes)
with cineqy TermEq have PEq: "P = [(y, c)] \<bullet> [(b, c)] \<bullet> Pa" and yFreshPa: "y \<sharp> [(b, c)] \<bullet> Pa"
by(simp add: pi.inject name_abs_eq)+
from PTrans have "([(b, c)] \<bullet> Pa) \<longmapsto> ([(b, c)] \<bullet> (aa[b] \<prec> Pa'))" by(rule TransitionsEarly.eqvt)
with aaineqb cineqaa have L1: "([(b, c)] \<bullet> Pa) \<longmapsto> aa[c] \<prec> [(b, c)] \<bullet> Pa'" by(simp add: name_calc)
with yFreshPa have yineqaa: "y \<noteq> aa" by(force dest: freshAction)
from L1 yFreshPa cineqy have yFreshPa': "y \<sharp> [(b, c)] \<bullet> Pa'" by(force intro: freshTransition)
from L1 have "([(y, c)] \<bullet> [(b, c)] \<bullet> Pa) \<longmapsto> [(y, c)] \<bullet> (aa[c] \<prec> [(b, c)] \<bullet> Pa')"
by(rule TransitionsEarly.eqvt)
with cineqaa yineqaa cineqy PEq have PTrans: "P \<longmapsto> aa[y] \<prec> [(y, c)] \<bullet> [(b, c)] \<bullet> Pa'"
by(simp add: name_calc)
moreover from cFreshPa' have "aa<\<nu>b> \<prec> Pa' = aa<\<nu>c> \<prec> ([(b, c)] \<bullet> Pa')" by(rule alphaBoundOutput)
with ResEq cineqx have ResEq': "RP' = [(x, c)] \<bullet> [(b, c)] \<bullet> Pa'" and "x \<sharp> [(b, c)] \<bullet> Pa'"
by(simp add: residual.inject name_abs_eq)+
with xineqy cineqy cineqx yFreshPa' have "RP' = [(x, y)] \<bullet> [(y, c)] \<bullet> [(b, c)] \<bullet> Pa'"
by(subst pt_perm_compose[OF pt_name_inst, OF at_name_inst], simp add: name_calc name_fresh_fresh)
moreover from ResEq have "a=aa" by(simp add: residual.inject)
ultimately show ?thesis using yineqaa rcOpen
by blast
next
fix Pa Pa' aa xa ya
assume PTrans: "Pa \<longmapsto> aa<\<nu>xa> \<prec> Pa'"
assume yaFreshaa: "(ya::name) \<noteq> aa"
assume yaineqxa: "ya \<noteq> xa"
assume EqTrans: "<\<nu>y>P = <\<nu>ya>Pa"
assume EqRes: "a<\<nu>x> \<prec> RP' = aa<\<nu>xa> \<prec> (<\<nu>ya>Pa')"
hence aeqaa: "a = aa" by(simp add: residual.inject)
with yaFreshaa have yaFresha: "ya \<sharp> a" by simp
have "\<exists>(c::name). c \<sharp> (Pa', y, xa, ya, x, Pa, aa)" by(blast intro: name_exists_fresh)
then obtain c where cFreshPa': "c \<sharp> Pa'" and cineqy: "c \<noteq> y" and cineqxa: "c \<noteq> xa" and cineqya: "c \<noteq> ya" and cineqx: "c \<noteq> x" and cFreshP: "c \<sharp> Pa" and cFreshaa: "c \<sharp> aa"
by(force simp add: fresh_prod name_fresh)
have "\<exists>(d::name). d \<sharp> (Pa, a, x, Pa', c, xa, ya, y)" by(blast intro: name_exists_fresh)
then obtain d where dFreshPa: "d \<sharp> Pa" and dFresha: "d \<sharp> a" and dineqx: "d \<noteq> x" and dFreshPa': "d \<sharp> Pa'" and dineqc: "d\<noteq>c" and dineqxa: "d \<noteq> xa" and dineqya: "d \<noteq> ya" and dineqy: "d \<noteq> y"
by(force simp add: fresh_prod name_fresh)
from dFreshPa have "<\<nu>ya>Pa = <\<nu>d>([(ya, d)] \<bullet> Pa)" by(rule alphaRes)
with EqTrans dineqy have PEq: "P = [(y, d)] \<bullet> [(ya, d)] \<bullet> Pa"
and yFreshPa: "y \<sharp> [(ya, d)] \<bullet> Pa"
by(simp add: pi.inject name_abs_eq)+
from dFreshPa' have L1: "<\<nu>ya>Pa' = <\<nu>d>([(ya, d)] \<bullet> Pa')" by(rule alphaRes)
from cFreshPa' dineqc cineqya have "c \<sharp> <\<nu>d>([(ya, d)] \<bullet> Pa')"
by(simp add: name_fresh_abs name_calc name_fresh_left)
hence "aa<\<nu>xa> \<prec> (<\<nu>d>([(ya, d)] \<bullet> Pa')) = aa<\<nu>c> \<prec> ([(xa, c)] \<bullet> <\<nu>d>([(ya, d)] \<bullet> Pa'))" (is "?LHS = _")
by(rule alphaBoundOutput)
with dineqxa dineqc have "?LHS = aa<\<nu>c> \<prec> (<\<nu>d>([(xa, c)] \<bullet> [(ya, d)] \<bullet> Pa'))"
by(simp add: name_calc)
with L1 EqRes cineqx dineqc dineqx have
RP'Eq: "RP' = <\<nu>d>([(x, c)] \<bullet> [(xa, c)] \<bullet> [(ya, d)] \<bullet> Pa')"
and xFreshPa': "x \<sharp> [(xa, c)] \<bullet> [(ya, d)] \<bullet> Pa'"
by(simp add: residual.inject name_abs_eq name_fresh_abs name_calc)+
from PTrans aeqaa have "([(ya, d)] \<bullet> Pa) \<longmapsto> [(ya, d)] \<bullet> (a<\<nu>xa> \<prec> Pa')"
by(blast intro: TransitionsEarly.eqvt)
with yaineqxa yaFresha dineqxa dFresha have L1:
"([(ya, d)] \<bullet> Pa) \<longmapsto> a<\<nu>xa> \<prec> ([(ya, d)] \<bullet> Pa')" by(simp add: name_calc name_fresh_fresh)
with yFreshPa have yineqa: "y \<noteq> a" by(force dest: freshAction)
from dineqc cineqya cFreshPa' have "c \<sharp> [(ya, d)] \<bullet> Pa'"
by(simp add: name_fresh_left name_calc)
hence "a<\<nu>xa> \<prec> ([(ya, d)] \<bullet> Pa') = a<\<nu>c> \<prec> ([(xa, c)] \<bullet> [(ya, d)] \<bullet> Pa')" (is "?LHS = _")
by(rule alphaBoundOutput)
with xFreshPa' have L2: "?LHS = a<\<nu>x> \<prec> ([(c, x)] \<bullet> [(xa, c)] \<bullet> [(ya, d)] \<bullet> Pa')"
by(simp add: alphaBoundOutput)
with L1 PEq have "P \<longmapsto> [(y, d)] \<bullet> (a<\<nu>x> \<prec> ([(c, x)] \<bullet> [(xa, c)] \<bullet> [(ya, d)] \<bullet> Pa'))"
by(force intro: TransitionsEarly.eqvt simp del: residual.perm)
with yineqa dFresha xineqy dineqx have Trans: "P \<longmapsto> a<\<nu>x> \<prec> ([(y, d)] \<bullet> [(c, x)] \<bullet> [(xa, c)] \<bullet> [(ya, d)] \<bullet> Pa')"
by(simp add: name_calc name_fresh_fresh)
from L1 L2 yFreshPa xineqy have "y \<sharp> [(c, x)] \<bullet> [(xa, c)] \<bullet> [(ya, d)] \<bullet> Pa'"
by(force intro: freshTransition)
with RP'Eq have "RP' = <\<nu>y>([(y, d)] \<bullet> [(c, x)] \<bullet> [(xa, c)] \<bullet> [(ya, d)] \<bullet> Pa')"
by(simp add: alphaRes name_swap)
with Trans yineqa show ?thesis
by(blast intro: rcResB)
qed
qed
lemma bangInduct[consumes 1, case_names Par1B Par1F Par2B Par2F Comm1 Comm2 Close1 Close2 Bang]:
fixes F :: "'a::fs_name \<Rightarrow> pi \<Rightarrow> residual \<Rightarrow> bool"
and P :: pi
and Rs :: residual
and C :: "'a::fs_name"
assumes Trans: "!P \<longmapsto> Rs"
and cPar1B: "\<And>a x P' C. \<lbrakk>P \<longmapsto> a<\<nu>x> \<prec> P'; x \<sharp> P; x \<sharp> C\<rbrakk> \<Longrightarrow> F C (P \<parallel> !P) (a<\<nu>x> \<prec> (P' \<parallel> !P))"
and cPar1F: "\<And>(\<alpha>::freeRes) (P'::pi) C. \<lbrakk>P \<longmapsto> \<alpha> \<prec> P'\<rbrakk> \<Longrightarrow> F C (P \<parallel> !P) (\<alpha> \<prec> P' \<parallel> !P)"
and cPar2B: "\<And>a x P' C. \<lbrakk>!P \<longmapsto> a<\<nu>x> \<prec> P'; x \<sharp> P; x \<sharp> C; \<And>C. F C (!P) (a<\<nu>x> \<prec> P')\<rbrakk> \<Longrightarrow> F C (P \<parallel> !P) (a<\<nu>x> \<prec> (P \<parallel> P'))"
and cPar2F: "\<And>\<alpha> P' C. \<lbrakk>!P \<longmapsto> \<alpha> \<prec> P'; \<And>C. F C (!P) (\<alpha> \<prec> P')\<rbrakk> \<Longrightarrow> F C (P \<parallel> !P) (\<alpha> \<prec> P \<parallel> P')"
and cComm1: "\<And>a P' b P'' C. \<lbrakk>P \<longmapsto> a<b> \<prec> P'; !P \<longmapsto> (OutputR a b) \<prec> P''; \<And>C. F C (!P) ((OutputR a b) \<prec> P'')\<rbrakk> \<Longrightarrow>
F C (P \<parallel> !P) (\<tau> \<prec> P' \<parallel> P'')"
and cComm2: "\<And>a b P' P'' C. \<lbrakk>P \<longmapsto> (OutputR a b) \<prec> P'; !P \<longmapsto> a<b> \<prec> P''; \<And>C. F C (!P) (a<b> \<prec> P'')\<rbrakk> \<Longrightarrow>
F C (P \<parallel> !P) (\<tau> \<prec> P' \<parallel> P'')"
and cClose1: "\<And>a x P' P'' C. \<lbrakk>P \<longmapsto> a<x> \<prec> P'; !P \<longmapsto> a<\<nu>x> \<prec> P''; x \<sharp> P; x \<sharp> C; \<And>C. F C (!P) (a<\<nu>x> \<prec> P'')\<rbrakk> \<Longrightarrow>
F C (P \<parallel> !P) (\<tau> \<prec> <\<nu>x>(P' \<parallel> P''))"
and cClose2: "\<And>a x P' P'' C. \<lbrakk>P \<longmapsto> a<\<nu>x> \<prec> P'; !P \<longmapsto> a<x> \<prec> P''; x \<sharp> P; x \<sharp> C; \<And>C. F C (!P) (a<x> \<prec> P'')\<rbrakk> \<Longrightarrow>
F C (P \<parallel> !P) (\<tau> \<prec> <\<nu>x>(P' \<parallel> P''))"
and cBang: "\<And>Rs C. \<lbrakk>P \<parallel> !P \<longmapsto> Rs; \<And>C. F C (P \<parallel> !P) Rs\<rbrakk> \<Longrightarrow> F C (!P) Rs"
shows "F C (!P) Rs"
proof -
have "\<And>X Y C. \<lbrakk>X \<longmapsto> Y; bangPred P X\<rbrakk> \<Longrightarrow> F C X Y"
proof -
fix X Y C
assume "X \<longmapsto> Y" and "bangPred P X"
thus "F C X Y"
proof(nominal_induct avoiding: C rule: TransitionsEarly.strong_induct)
case(Tau Pa)
thus ?case
apply -
by(ind_cases "bangPred P (\<tau>.(Pa))")
next
case(Input x a u Pa C)
thus ?case
by - (ind_cases "bangPred P (a<x>.Pa)")
next
case(Output a b Pa C)
thus ?case
by - (ind_cases "bangPred P (a{b}.Pa)")
next
case(Match Pa Rs b C)
thus ?case
by - (ind_cases "bangPred P ([b\<frown>b]Pa)")
next
case(Mismatch Pa Rs a b C)
thus ?case
by - (ind_cases "bangPred P ([a \<noteq> b]Pa)")
next
case(Open Pa a b Pa')
thus ?case
by - (ind_cases "bangPred P (<\<nu>b>Pa)")
next
case(Sum1 Pa Rs Q)
thus ?case
by - (ind_cases "bangPred P (Pa \<oplus> Q)")
next
case(Sum2 Q Rs Pa)
thus ?case
by - (ind_cases "bangPred P (Pa \<oplus> Q)")
next
case(Par1B Pa a x P' Q C)
thus ?case
by - (ind_cases "bangPred P (Pa \<parallel> Q)", auto simp add: pi.inject cPar1B)
next
case(Par1F Pa \<alpha> P' Q C)
thus ?case
by - (ind_cases "bangPred P (Pa \<parallel> Q)", auto simp add: pi.inject cPar1F)
next
case(Par2B Q a x Q' Pa)
thus ?case
by - (ind_cases "bangPred P (Pa \<parallel> Q)", auto simp add: pi.inject aux1 cPar2B)
next
case(Par2F Q \<alpha> Q' Pa)
thus ?case
by - (ind_cases "bangPred P (Pa \<parallel> Q)", auto simp add: pi.inject intro: cPar2F aux1)
next
case(Comm1 Pa a b Pa' Q Q' C)
thus ?case
by - (ind_cases "bangPred P (Pa \<parallel> Q)", auto simp add: pi.inject intro: cComm1 aux1)
next
case(Comm2 Pa a b Pa' Q P'' C)
thus ?case
by - (ind_cases "bangPred P (Pa \<parallel> Q)", auto simp add: pi.inject intro: cComm2 aux1)
next
case(Close1 Pa a x Pa' Q Q'' C)
thus ?case
by - (ind_cases "bangPred P (Pa \<parallel> Q)", auto simp add: pi.inject aux1 cClose1)
next
case(Close2 Pa a x Pa' Q Q' C)
thus ?case
by - (ind_cases "bangPred P (Pa \<parallel> Q)", auto simp add: pi.inject aux1 cClose2)
next
case(ResB Pa a x Pa' y)
thus ?case
by - (ind_cases "bangPred P (<\<nu>y>Pa)")
next
case(ResF Pa \<alpha> Pa' y)
thus ?case
by - (ind_cases "bangPred P (<\<nu>y>Pa)")
next
case(Bang Pa Rs)
thus ?case
by - (ind_cases "bangPred P (!Pa)", auto simp add: pi.inject intro: aux2 cBang)
qed
qed
with Trans show ?thesis by(force intro: bangPred.aux1)
qed
end
diff --git a/thys/Refine_Monadic/Refine_Foreach.thy b/thys/Refine_Monadic/Refine_Foreach.thy
--- a/thys/Refine_Monadic/Refine_Foreach.thy
+++ b/thys/Refine_Monadic/Refine_Foreach.thy
@@ -1,1895 +1,1896 @@
section \<open>Foreach Loops\<close>
theory Refine_Foreach
imports
Refine_While
Refine_Pfun
Refine_Transfer
Refine_Heuristics
(*"../Collections/Lib/SetIterator"
"../Collections/Lib/Proper_Iterator"*)
begin
text \<open>
A common pattern for loop usage is iteration over the elements of a set.
This theory provides the \<open>FOREACH\<close>-combinator, that iterates over
each element of a set.
\<close>
subsection \<open>Auxilliary Lemmas\<close>
text \<open>The following lemma is commonly used when reasoning about iterator
invariants.
It helps converting the set of elements that remain to be iterated over to
the set of elements already iterated over.\<close>
lemma it_step_insert_iff:
"it \<subseteq> S \<Longrightarrow> x\<in>it \<Longrightarrow> S-(it-{x}) = insert x (S-it)" by auto
subsection \<open>Definition\<close>
text \<open>
Foreach-loops come in different versions, depending on whether they have an
annotated invariant (I), a termination condition (C), and an order (O).
Note that asserting that the set is finite is not necessary to guarantee
termination. However, we currently provide only iteration over finite sets,
as this also matches the ICF concept of iterators.
\<close>
definition "FOREACH_body f \<equiv> \<lambda>(xs, \<sigma>). do {
let x = hd xs; \<sigma>'\<leftarrow>f x \<sigma>; RETURN (tl xs,\<sigma>')
}"
definition FOREACH_cond where "FOREACH_cond c \<equiv> (\<lambda>(xs,\<sigma>). xs\<noteq>[] \<and> c \<sigma>)"
text \<open>Foreach with continuation condition, order and annotated invariant:\<close>
definition FOREACHoci ("FOREACH\<^sub>O\<^sub>C\<^bsup>_,_\<^esup>") where "FOREACHoci R \<Phi> S c f \<sigma>0 \<equiv> do {
ASSERT (finite S);
xs \<leftarrow> SPEC (\<lambda>xs. distinct xs \<and> S = set xs \<and> sorted_wrt R xs);
(_,\<sigma>) \<leftarrow> WHILEIT
(\<lambda>(it,\<sigma>). \<exists>xs'. xs = xs' @ it \<and> \<Phi> (set it) \<sigma>) (FOREACH_cond c) (FOREACH_body f) (xs,\<sigma>0);
RETURN \<sigma> }"
text \<open>Foreach with continuation condition and annotated invariant:\<close>
definition FOREACHci ("FOREACH\<^sub>C\<^bsup>_\<^esup>") where "FOREACHci \<equiv> FOREACHoci (\<lambda>_ _. True)"
text \<open>Foreach with continuation condition:\<close>
definition FOREACHc ("FOREACH\<^sub>C") where "FOREACHc \<equiv> FOREACHci (\<lambda>_ _. True)"
text \<open>Foreach with annotated invariant:\<close>
definition FOREACHi ("FOREACH\<^bsup>_\<^esup>") where
"FOREACHi \<Phi> S \<equiv> FOREACHci \<Phi> S (\<lambda>_. True)"
text \<open>Foreach with annotated invariant and order:\<close>
definition FOREACHoi ("FOREACH\<^sub>O\<^bsup>_,_\<^esup>") where
"FOREACHoi R \<Phi> S \<equiv> FOREACHoci R \<Phi> S (\<lambda>_. True)"
text \<open>Basic foreach\<close>
definition "FOREACH S \<equiv> FOREACHc S (\<lambda>_. True)"
lemmas FOREACH_to_oci_unfold
= FOREACHci_def FOREACHc_def FOREACHi_def FOREACHoi_def FOREACH_def
subsection \<open>Proof Rules\<close>
lemma FOREACHoci_rule[refine_vcg]:
assumes FIN: "finite S"
assumes I0: "I S \<sigma>0"
assumes IP:
"\<And>x it \<sigma>. \<lbrakk> c \<sigma>; x\<in>it; it\<subseteq>S; I it \<sigma>; \<forall>y\<in>it - {x}. R x y;
\<forall>y\<in>S - it. R y x \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (it-{x}))"
assumes II1: "\<And>\<sigma>. \<lbrakk>I {} \<sigma>\<rbrakk> \<Longrightarrow> P \<sigma>"
assumes II2: "\<And>it \<sigma>. \<lbrakk> it\<noteq>{}; it\<subseteq>S; I it \<sigma>; \<not>c \<sigma>;
\<forall>x\<in>it. \<forall>y\<in>S - it. R y x \<rbrakk> \<Longrightarrow> P \<sigma>"
shows "FOREACHoci R I S c f \<sigma>0 \<le> SPEC P"
unfolding FOREACHoci_def
apply (intro refine_vcg)
apply (rule FIN)
apply (subgoal_tac "wf (measure (\<lambda>(xs, _). length xs))")
apply assumption
apply simp
apply (insert I0, simp add: I0) []
unfolding FOREACH_body_def FOREACH_cond_def
apply (rule refine_vcg)+
apply ((simp, elim conjE exE)+) []
apply (rename_tac xs'' s xs' \<sigma> xs)
defer
apply (simp, elim conjE exE)+
apply (rename_tac x s xs' \<sigma> xs)
defer
proof -
fix xs' \<sigma> xs
assume I_xs': "I (set xs') \<sigma>"
and sorted_xs_xs': "sorted_wrt R (xs @ xs')"
and dist: "distinct xs" "distinct xs'" "set xs \<inter> set xs' = {}"
and S_eq: "S = set xs \<union> set xs'"
from S_eq have "set xs' \<subseteq> S" by simp
from dist S_eq have S_diff: "S - set xs' = set xs" by blast
{ assume "xs' \<noteq> []" "c \<sigma>"
from \<open>xs' \<noteq> []\<close> obtain x xs'' where xs'_eq: "xs' = x # xs''" by (cases xs', auto)
have x_in_xs': "x \<in> set xs'" and x_nin_xs'': "x \<notin> set xs''"
using \<open>distinct xs'\<close> unfolding xs'_eq by simp_all
from IP[of \<sigma> x "set xs'", OF \<open>c \<sigma>\<close> x_in_xs' \<open>set xs' \<subseteq> S\<close> \<open>I (set xs') \<sigma>\<close>] x_nin_xs''
sorted_xs_xs' S_diff
show "f (hd xs') \<sigma> \<le> SPEC
(\<lambda>x. (\<exists>xs'a. xs @ xs' = xs'a @ tl xs') \<and>
I (set (tl xs')) x)"
apply (simp add: xs'_eq)
apply (simp add: sorted_wrt_append)
done
}
{ assume "xs' = [] \<or> \<not>(c \<sigma>)"
show "P \<sigma>"
proof (cases "xs' = []")
case True thus "P \<sigma>" using \<open>I (set xs') \<sigma>\<close> by (simp add: II1)
next
case False note xs'_neq_nil = this
with \<open>xs' = [] \<or> \<not> c \<sigma>\<close> have "\<not> c \<sigma>" by simp
from II2 [of "set xs'" \<sigma>] S_diff sorted_xs_xs'
show "P \<sigma>"
apply (simp add: xs'_neq_nil S_eq \<open>\<not> c \<sigma>\<close> I_xs')
apply (simp add: sorted_wrt_append)
done
qed
}
qed
lemma FOREACHoi_rule[refine_vcg]:
assumes FIN: "finite S"
assumes I0: "I S \<sigma>0"
assumes IP:
"\<And>x it \<sigma>. \<lbrakk> x\<in>it; it\<subseteq>S; I it \<sigma>; \<forall>y\<in>it - {x}. R x y;
\<forall>y\<in>S - it. R y x \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (it-{x}))"
assumes II1: "\<And>\<sigma>. \<lbrakk>I {} \<sigma>\<rbrakk> \<Longrightarrow> P \<sigma>"
shows "FOREACHoi R I S f \<sigma>0 \<le> SPEC P"
unfolding FOREACHoi_def
by (rule FOREACHoci_rule) (simp_all add: assms)
lemma FOREACHci_rule[refine_vcg]:
assumes FIN: "finite S"
assumes I0: "I S \<sigma>0"
assumes IP:
"\<And>x it \<sigma>. \<lbrakk> x\<in>it; it\<subseteq>S; I it \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (it-{x}))"
assumes II1: "\<And>\<sigma>. \<lbrakk>I {} \<sigma>\<rbrakk> \<Longrightarrow> P \<sigma>"
assumes II2: "\<And>it \<sigma>. \<lbrakk> it\<noteq>{}; it\<subseteq>S; I it \<sigma>; \<not>c \<sigma> \<rbrakk> \<Longrightarrow> P \<sigma>"
shows "FOREACHci I S c f \<sigma>0 \<le> SPEC P"
unfolding FOREACHci_def
by (rule FOREACHoci_rule) (simp_all add: assms)
subsubsection \<open>Refinement:\<close>
text \<open>
Refinement rule using a coupling invariant over sets of remaining
items and the state.
\<close>
lemma FOREACHoci_refine_genR:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa" \<comment> \<open>Abstraction mapping of elements\<close>
fixes S :: "'S set" \<comment> \<open>Concrete set\<close>
fixes S' :: "'Sa set" \<comment> \<open>Abstract set\<close>
fixes \<sigma>0 :: "'\<sigma>"
fixes \<sigma>0' :: "'\<sigma>a"
fixes R :: "(('S set \<times> '\<sigma>) \<times> ('Sa set \<times> '\<sigma>a)) set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS[simp]: "S' = \<alpha>`S"
assumes RR_OK: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; RR x y\<rbrakk> \<Longrightarrow> RR' (\<alpha> x) (\<alpha> y)"
assumes REF0: "((S,\<sigma>0),(\<alpha>`S,\<sigma>0')) \<in> R"
assumes REFC: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; \<Phi> it \<sigma>;
\<forall>x\<in>S-it. \<forall>y\<in>it. RR x y; \<forall>x\<in>S'-it'. \<forall>y\<in>it'. RR' x y;
it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R
\<rbrakk> \<Longrightarrow> c \<sigma> \<longleftrightarrow> c' \<sigma>'"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>';
\<forall>x\<in>S-it. \<forall>y\<in>it. RR x y; \<forall>x\<in>S'-it'. \<forall>y\<in>it'. RR' x y;
it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi> it \<sigma>; \<Phi>' it' \<sigma>';
\<forall>x\<in>S-it. \<forall>y\<in>it. RR x y; \<forall>x\<in>S'-it'. \<forall>y\<in>it'. RR' x y;
x'=\<alpha> x; it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R;
x\<in>it; \<forall>y\<in>it-{x}. RR x y;
x'\<in>it'; \<forall>y'\<in>it'-{x'}. RR' x' y';
c \<sigma>; c' \<sigma>'
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). ((it-{x},\<sigma>),(it'-{x'},\<sigma>'))\<in>R}) (f' x' \<sigma>')"
assumes REF_R_DONE: "\<And>\<sigma> \<sigma>'. \<lbrakk> \<Phi> {} \<sigma>; \<Phi>' {} \<sigma>'; (({},\<sigma>),({},\<sigma>'))\<in>R \<rbrakk>
\<Longrightarrow> (\<sigma>,\<sigma>')\<in>R'"
assumes REF_R_BRK: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi> it \<sigma>; \<Phi>' it' \<sigma>';
\<forall>x\<in>S-it. \<forall>y\<in>it. RR x y; \<forall>x\<in>S'-it'. \<forall>y\<in>it'. RR' x y;
it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R;
it\<noteq>{}; it'\<noteq>{};
\<not>c \<sigma>; \<not>c' \<sigma>'
\<rbrakk> \<Longrightarrow> (\<sigma>,\<sigma>')\<in>R'"
shows "FOREACHoci RR \<Phi> S c f \<sigma>0 \<le> \<Down>R' (FOREACHoci RR' \<Phi>' S' c' f' \<sigma>0')"
(* TODO: Clean up this mess !!! *)
+ supply [[simproc del: defined_all]]
unfolding FOREACHoci_def
apply (refine_rcg WHILEIT_refine_genR[where
R'="{((xs,\<sigma>),(xs',\<sigma>')) .
xs'=map \<alpha> xs \<and>
set xs \<subseteq> S \<and> set xs' \<subseteq> S' \<and>
(\<forall>x\<in>S - set xs. \<forall>y\<in>set xs. RR x y) \<and>
(\<forall>x\<in>S' - set xs'. \<forall>y\<in>set xs'. RR' x y) \<and>
((set xs,\<sigma>),(set xs',\<sigma>')) \<in> R }"
])
using REFS INJ apply (auto dest: finite_imageD) []
apply (rule intro_prgR[where R="{(xs,xs') . xs'=map \<alpha> xs }"])
apply (rule SPEC_refine)
using INJ RR_OK
apply (auto
simp add: distinct_map sorted_wrt_map
intro: sorted_wrt_mono_rel[of _ RR]) []
using REF0 apply auto []
apply simp apply (rule conjI)
using INJ apply clarsimp
apply (erule map_eq_appendE)
apply clarsimp
apply (rule_tac x=l in exI)
apply simp
apply (subst inj_on_map_eq_map[where f=\<alpha>,symmetric])
apply (rule subset_inj_on, assumption, blast)
apply assumption
apply (simp split: prod.split_asm, elim conjE)
apply (rule REFPHI, auto) []
apply (simp add: FOREACH_cond_def split: prod.split prod.split_asm,
intro allI impI conj_cong) []
apply auto []
apply (rule REFC, auto) []
unfolding FOREACH_body_def
apply refine_rcg
apply (rule REFSTEP) []
prefer 3 apply auto []
prefer 3 apply auto []
apply simp_all[13]
apply auto []
apply (rename_tac a b d e f g h i)
apply (case_tac h, auto simp: FOREACH_cond_def) []
apply auto []
apply (auto simp: FOREACH_cond_def) []
apply (clarsimp simp: FOREACH_cond_def)
apply (rule ccontr)
apply (rename_tac a b d e f)
apply (case_tac b)
apply (auto simp: sorted_wrt_append) [2]
apply (auto simp: FOREACH_cond_def) []
apply (rename_tac a b d e)
apply (case_tac b)
apply (auto) [2]
apply (clarsimp simp: FOREACH_cond_def)
apply (rule ccontr)
apply (rename_tac a b d e f)
apply (case_tac b)
apply (auto simp: sorted_wrt_append) [2]
apply (clarsimp simp: FOREACH_cond_def)
apply (clarsimp simp: FOREACH_cond_def)
apply (clarsimp simp: map_tl)
apply (intro conjI)
apply (rename_tac a b d e f g)
apply (case_tac b, auto) []
apply (rename_tac a b d e f g)
apply (case_tac b, auto) []
apply (rename_tac a b d e f g)
apply (case_tac b, auto simp: sorted_wrt_append) []
apply (rename_tac a b d e f g)
apply (case_tac b, auto simp: sorted_wrt_append) []
apply (rename_tac a b d e f g)
apply (case_tac b, auto) []
apply (rule introR[where R="{((xs,\<sigma>),(xs',\<sigma>')).
xs'=map \<alpha> xs \<and> \<Phi> (set xs) \<sigma> \<and> \<Phi>' (set xs') \<sigma>' \<and>
set xs \<subseteq> S \<and> set xs' \<subseteq> S' \<and>
(\<forall>x\<in>S - set xs. \<forall>y\<in>set xs. RR x y) \<and>
(\<forall>x\<in>S' - set xs'. \<forall>y\<in>set xs'. RR' x y) \<and>
((set xs,\<sigma>),(set xs',\<sigma>')) \<in> R \<and>
\<not> FOREACH_cond c (xs,\<sigma>) \<and> \<not> FOREACH_cond c' (xs',\<sigma>')
}
"])
apply auto []
apply (simp add: FOREACH_cond_def, elim conjE)
apply (elim disjE1, simp_all) []
using REF_R_DONE apply auto []
using REF_R_BRK apply auto []
done
lemma FOREACHoci_refine:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes RR_OK: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; RR x y\<rbrakk> \<Longrightarrow> RR' (\<alpha> x) (\<alpha> y)"
assumes REFPHI0: "\<Phi>'' S \<sigma>0 (\<alpha>`S) \<sigma>0'"
assumes REFC: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; \<Phi>'' it \<sigma> it' \<sigma>'; \<Phi> it \<sigma>; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> c \<sigma> \<longleftrightarrow> c' \<sigma>'"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; \<Phi>'' it \<sigma> it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>\<forall>y\<in>it-{x}. RR x y;
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi> it \<sigma>; \<Phi>' it' \<sigma>'; \<Phi>'' it \<sigma> it' \<sigma>'; c \<sigma>; c' \<sigma>';
(\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). (\<sigma>, \<sigma>') \<in> R \<and> \<Phi>'' (it - {x}) \<sigma> (it' - {x'}) \<sigma>'}) (f' x' \<sigma>')"
shows "FOREACHoci RR \<Phi> S c f \<sigma>0 \<le> \<Down>R (FOREACHoci RR' \<Phi>' S' c' f' \<sigma>0')"
apply (rule FOREACHoci_refine_genR[
where R = "{((it,\<sigma>),(it',\<sigma>')). (\<sigma>,\<sigma>')\<in>R \<and> \<Phi>'' it \<sigma> it' \<sigma>'}"
])
apply fact
apply fact
apply fact
using REF0 REFPHI0 apply blast
using REFC apply auto []
using REFPHI apply auto []
using REFSTEP apply auto []
apply auto []
apply auto []
done
lemma FOREACHoci_refine_rcg[refine]:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes RR_OK: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; RR x y\<rbrakk> \<Longrightarrow> RR' (\<alpha> x) (\<alpha> y)"
assumes REFC: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; \<Phi> it \<sigma>; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> c \<sigma> \<longleftrightarrow> c' \<sigma>'"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk> \<forall>y\<in>it-{x}. RR x y;
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi> it \<sigma>; \<Phi>' it' \<sigma>'; c \<sigma>; c' \<sigma>';
(\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma> \<le> \<Down>R (f' x' \<sigma>')"
shows "FOREACHoci RR \<Phi> S c f \<sigma>0 \<le> \<Down>R (FOREACHoci RR' \<Phi>' S' c' f' \<sigma>0')"
apply (rule FOREACHoci_refine[where \<Phi>''="\<lambda>_ _ _ _. True"])
apply (rule assms)+
using assms by simp_all
lemma FOREACHoci_weaken:
assumes IREF: "\<And>it \<sigma>. it\<subseteq>S \<Longrightarrow> I it \<sigma> \<Longrightarrow> I' it \<sigma>"
shows "FOREACHoci RR I' S c f \<sigma>0 \<le> FOREACHoci RR I S c f \<sigma>0"
apply (rule FOREACHoci_refine_rcg[where \<alpha>=id and R=Id, simplified])
apply (auto intro: IREF)
done
lemma FOREACHoci_weaken_order:
assumes RRREF: "\<And>x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> RR x y \<Longrightarrow> RR' x y"
shows "FOREACHoci RR I S c f \<sigma>0 \<le> FOREACHoci RR' I S c f \<sigma>0"
apply (rule FOREACHoci_refine_rcg[where \<alpha>=id and R=Id, simplified])
apply (auto intro: RRREF)
done
subsubsection \<open>Rules for Derived Constructs\<close>
lemma FOREACHoi_refine_genR:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa" \<comment> \<open>Abstraction mapping of elements\<close>
fixes S :: "'S set" \<comment> \<open>Concrete set\<close>
fixes S' :: "'Sa set" \<comment> \<open>Abstract set\<close>
fixes \<sigma>0 :: "'\<sigma>"
fixes \<sigma>0' :: "'\<sigma>a"
fixes R :: "(('S set \<times> '\<sigma>) \<times> ('Sa set \<times> '\<sigma>a)) set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS[simp]: "S' = \<alpha>`S"
assumes RR_OK: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; RR x y\<rbrakk> \<Longrightarrow> RR' (\<alpha> x) (\<alpha> y)"
assumes REF0: "((S,\<sigma>0),(\<alpha>`S,\<sigma>0')) \<in> R"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>';
\<forall>x\<in>S-it. \<forall>y\<in>it. RR x y; \<forall>x\<in>S'-it'. \<forall>y\<in>it'. RR' x y;
it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi> it \<sigma>; \<Phi>' it' \<sigma>';
\<forall>x\<in>S-it. \<forall>y\<in>it. RR x y; \<forall>x\<in>S'-it'. \<forall>y\<in>it'. RR' x y;
x'=\<alpha> x; it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R;
x\<in>it; \<forall>y\<in>it-{x}. RR x y;
x'\<in>it'; \<forall>y'\<in>it'-{x'}. RR' x' y'
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). ((it-{x},\<sigma>),(it'-{x'},\<sigma>'))\<in>R}) (f' x' \<sigma>')"
assumes REF_R_DONE: "\<And>\<sigma> \<sigma>'. \<lbrakk> \<Phi> {} \<sigma>; \<Phi>' {} \<sigma>'; (({},\<sigma>),({},\<sigma>'))\<in>R \<rbrakk>
\<Longrightarrow> (\<sigma>,\<sigma>')\<in>R'"
shows "FOREACHoi RR \<Phi> S f \<sigma>0 \<le> \<Down>R' (FOREACHoi RR' \<Phi>' S' f' \<sigma>0')"
unfolding FOREACHoi_def
apply (rule FOREACHoci_refine_genR)
apply (fact | simp)+
using REFSTEP apply auto []
apply (fact | simp)+
done
lemma FOREACHoi_refine:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes RR_OK: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; RR x y\<rbrakk> \<Longrightarrow> RR' (\<alpha> x) (\<alpha> y)"
assumes REFPHI0: "\<Phi>'' S \<sigma>0 (\<alpha>`S) \<sigma>0'"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; \<Phi>'' it \<sigma> it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>\<forall>y\<in>it-{x}. RR x y;
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi> it \<sigma>; \<Phi>' it' \<sigma>'; \<Phi>'' it \<sigma> it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). (\<sigma>, \<sigma>') \<in> R \<and> \<Phi>'' (it - {x}) \<sigma> (it' - {x'}) \<sigma>'}) (f' x' \<sigma>')"
shows "FOREACHoi RR \<Phi> S f \<sigma>0 \<le> \<Down>R (FOREACHoi RR' \<Phi>' S' f' \<sigma>0')"
unfolding FOREACHoi_def
apply (rule FOREACHoci_refine [of \<alpha> _ _ _ _ _ _ _ \<Phi>''])
apply (simp_all add: assms)
done
lemma FOREACHoi_refine_rcg[refine]:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes RR_OK: "\<And>x y. \<lbrakk>x \<in> S; y \<in> S; RR x y\<rbrakk> \<Longrightarrow> RR' (\<alpha> x) (\<alpha> y)"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk> \<forall>y\<in>it-{x}. RR x y;
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi> it \<sigma>; \<Phi>' it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma> \<le> \<Down>R (f' x' \<sigma>')"
shows "FOREACHoi RR \<Phi> S f \<sigma>0 \<le> \<Down>R (FOREACHoi RR' \<Phi>' S' f' \<sigma>0')"
apply (rule FOREACHoi_refine[where \<Phi>''="\<lambda>_ _ _ _. True"])
apply (rule assms)+
using assms by simp_all
lemma FOREACHci_refine_genR:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa" \<comment> \<open>Abstraction mapping of elements\<close>
fixes S :: "'S set" \<comment> \<open>Concrete set\<close>
fixes S' :: "'Sa set" \<comment> \<open>Abstract set\<close>
fixes \<sigma>0 :: "'\<sigma>"
fixes \<sigma>0' :: "'\<sigma>a"
fixes R :: "(('S set \<times> '\<sigma>) \<times> ('Sa set \<times> '\<sigma>a)) set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS[simp]: "S' = \<alpha>`S"
assumes REF0: "((S,\<sigma>0),(\<alpha>`S,\<sigma>0')) \<in> R"
assumes REFC: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; \<Phi> it \<sigma>;
it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R
\<rbrakk> \<Longrightarrow> c \<sigma> \<longleftrightarrow> c' \<sigma>'"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>';
it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi> it \<sigma>; \<Phi>' it' \<sigma>';
x'=\<alpha> x; it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R;
x\<in>it; x'\<in>it';
c \<sigma>; c' \<sigma>'
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). ((it-{x},\<sigma>),(it'-{x'},\<sigma>'))\<in>R}) (f' x' \<sigma>')"
assumes REF_R_DONE: "\<And>\<sigma> \<sigma>'. \<lbrakk> \<Phi> {} \<sigma>; \<Phi>' {} \<sigma>'; (({},\<sigma>),({},\<sigma>'))\<in>R \<rbrakk>
\<Longrightarrow> (\<sigma>,\<sigma>')\<in>R'"
assumes REF_R_BRK: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi> it \<sigma>; \<Phi>' it' \<sigma>';
it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R;
it\<noteq>{}; it'\<noteq>{};
\<not>c \<sigma>; \<not>c' \<sigma>'
\<rbrakk> \<Longrightarrow> (\<sigma>,\<sigma>')\<in>R'"
shows "FOREACHci \<Phi> S c f \<sigma>0 \<le> \<Down>R' (FOREACHci \<Phi>' S' c' f' \<sigma>0')"
unfolding FOREACHci_def
apply (rule FOREACHoci_refine_genR)
apply (fact|simp)+
using REFC apply auto []
using REFPHI apply auto []
using REFSTEP apply auto []
apply (fact|simp)+
using REF_R_BRK apply auto []
done
lemma FOREACHci_refine:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes REFPHI0: "\<Phi>'' S \<sigma>0 (\<alpha>`S) \<sigma>0'"
assumes REFC: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; \<Phi>'' it \<sigma> it' \<sigma>'; \<Phi> it \<sigma>; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> c \<sigma> \<longleftrightarrow> c' \<sigma>'"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; \<Phi>'' it \<sigma> it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi> it \<sigma>; \<Phi>' it' \<sigma>'; \<Phi>'' it \<sigma> it' \<sigma>'; c \<sigma>; c' \<sigma>';
(\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). (\<sigma>, \<sigma>') \<in> R \<and> \<Phi>'' (it - {x}) \<sigma> (it' - {x'}) \<sigma>'}) (f' x' \<sigma>')"
shows "FOREACHci \<Phi> S c f \<sigma>0 \<le> \<Down>R (FOREACHci \<Phi>' S' c' f' \<sigma>0')"
unfolding FOREACHci_def
apply (rule FOREACHoci_refine [of \<alpha> _ _ _ _ _ _ _ \<Phi>''])
apply (simp_all add: assms)
done
lemma FOREACHci_refine_rcg[refine]:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes REFC: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; \<Phi> it \<sigma>; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> c \<sigma> \<longleftrightarrow> c' \<sigma>'"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi> it \<sigma>; \<Phi>' it' \<sigma>'; c \<sigma>; c' \<sigma>';
(\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma> \<le> \<Down>R (f' x' \<sigma>')"
shows "FOREACHci \<Phi> S c f \<sigma>0 \<le> \<Down>R (FOREACHci \<Phi>' S' c' f' \<sigma>0')"
apply (rule FOREACHci_refine[where \<Phi>''="\<lambda>_ _ _ _. True"])
apply (rule assms)+
using assms by auto
lemma FOREACHci_weaken:
assumes IREF: "\<And>it \<sigma>. it\<subseteq>S \<Longrightarrow> I it \<sigma> \<Longrightarrow> I' it \<sigma>"
shows "FOREACHci I' S c f \<sigma>0 \<le> FOREACHci I S c f \<sigma>0"
apply (rule FOREACHci_refine_rcg[where \<alpha>=id and R=Id, simplified])
apply (auto intro: IREF)
done
lemma FOREACHi_rule[refine_vcg]:
assumes FIN: "finite S"
assumes I0: "I S \<sigma>0"
assumes IP:
"\<And>x it \<sigma>. \<lbrakk> x\<in>it; it\<subseteq>S; I it \<sigma> \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (it-{x}))"
assumes II: "\<And>\<sigma>. \<lbrakk>I {} \<sigma>\<rbrakk> \<Longrightarrow> P \<sigma>"
shows "FOREACHi I S f \<sigma>0 \<le> SPEC P"
unfolding FOREACHi_def
apply (rule FOREACHci_rule[of S I])
using assms by auto
lemma FOREACHc_rule:
assumes FIN: "finite S"
assumes I0: "I S \<sigma>0"
assumes IP:
"\<And>x it \<sigma>. \<lbrakk> x\<in>it; it\<subseteq>S; I it \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (it-{x}))"
assumes II1: "\<And>\<sigma>. \<lbrakk>I {} \<sigma>\<rbrakk> \<Longrightarrow> P \<sigma>"
assumes II2: "\<And>it \<sigma>. \<lbrakk> it\<noteq>{}; it\<subseteq>S; I it \<sigma>; \<not>c \<sigma> \<rbrakk> \<Longrightarrow> P \<sigma>"
shows "FOREACHc S c f \<sigma>0 \<le> SPEC P"
unfolding FOREACHc_def
apply (rule order_trans[OF FOREACHci_weaken], rule TrueI)
apply (rule FOREACHci_rule[where I=I])
using assms by auto
lemma FOREACH_rule:
assumes FIN: "finite S"
assumes I0: "I S \<sigma>0"
assumes IP:
"\<And>x it \<sigma>. \<lbrakk> x\<in>it; it\<subseteq>S; I it \<sigma> \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (it-{x}))"
assumes II: "\<And>\<sigma>. \<lbrakk>I {} \<sigma>\<rbrakk> \<Longrightarrow> P \<sigma>"
shows "FOREACH S f \<sigma>0 \<le> SPEC P"
unfolding FOREACH_def FOREACHc_def
apply (rule order_trans[OF FOREACHci_weaken], rule TrueI)
apply (rule FOREACHci_rule[where I=I])
using assms by auto
lemma FOREACHc_refine_genR:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa" \<comment> \<open>Abstraction mapping of elements\<close>
fixes S :: "'S set" \<comment> \<open>Concrete set\<close>
fixes S' :: "'Sa set" \<comment> \<open>Abstract set\<close>
fixes \<sigma>0 :: "'\<sigma>"
fixes \<sigma>0' :: "'\<sigma>a"
fixes R :: "(('S set \<times> '\<sigma>) \<times> ('Sa set \<times> '\<sigma>a)) set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS[simp]: "S' = \<alpha>`S"
assumes REF0: "((S,\<sigma>0),(\<alpha>`S,\<sigma>0')) \<in> R"
assumes REFC: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S';
it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R
\<rbrakk> \<Longrightarrow> c \<sigma> \<longleftrightarrow> c' \<sigma>'"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S';
x'=\<alpha> x; it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R;
x\<in>it; x'\<in>it';
c \<sigma>; c' \<sigma>'
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). ((it-{x},\<sigma>),(it'-{x'},\<sigma>'))\<in>R}) (f' x' \<sigma>')"
assumes REF_R_DONE: "\<And>\<sigma> \<sigma>'. \<lbrakk> (({},\<sigma>),({},\<sigma>'))\<in>R \<rbrakk>
\<Longrightarrow> (\<sigma>,\<sigma>')\<in>R'"
assumes REF_R_BRK: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S';
it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R;
it\<noteq>{}; it'\<noteq>{};
\<not>c \<sigma>; \<not>c' \<sigma>'
\<rbrakk> \<Longrightarrow> (\<sigma>,\<sigma>')\<in>R'"
shows "FOREACHc S c f \<sigma>0 \<le> \<Down>R' (FOREACHc S' c' f' \<sigma>0')"
unfolding FOREACHc_def
apply (rule FOREACHci_refine_genR)
apply simp_all
apply (fact|simp)+
using REFC apply auto []
using REFSTEP apply auto []
using REF_R_DONE apply auto []
using REF_R_BRK apply auto []
done
lemma FOREACHc_refine:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes REFPHI0: "\<Phi>'' S \<sigma>0 (\<alpha>`S) \<sigma>0'"
assumes REFC: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>'' it \<sigma> it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> c \<sigma> \<longleftrightarrow> c' \<sigma>'"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi>'' it \<sigma> it' \<sigma>'; c \<sigma>; c' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). (\<sigma>, \<sigma>') \<in> R \<and> \<Phi>'' (it - {x}) \<sigma> (it' - {x'}) \<sigma>'}) (f' x' \<sigma>')"
shows "FOREACHc S c f \<sigma>0 \<le> \<Down>R (FOREACHc S' c' f' \<sigma>0')"
unfolding FOREACHc_def
apply (rule FOREACHci_refine[where \<Phi>''=\<Phi>'', OF INJ REFS REF0 REFPHI0])
apply (erule (4) REFC)
apply (rule TrueI)
apply (erule (9) REFSTEP)
done
lemma FOREACHc_refine_rcg[refine]:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes REFC: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> c \<sigma> \<longleftrightarrow> c' \<sigma>'"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; c \<sigma>; c' \<sigma>';
(\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma> \<le> \<Down>R (f' x' \<sigma>')"
shows "FOREACHc S c f \<sigma>0 \<le> \<Down>R (FOREACHc S' c' f' \<sigma>0')"
unfolding FOREACHc_def
apply (rule FOREACHci_refine_rcg)
apply (rule assms)+
using assms by auto
lemma FOREACHi_refine_genR:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa" \<comment> \<open>Abstraction mapping of elements\<close>
fixes S :: "'S set" \<comment> \<open>Concrete set\<close>
fixes S' :: "'Sa set" \<comment> \<open>Abstract set\<close>
fixes \<sigma>0 :: "'\<sigma>"
fixes \<sigma>0' :: "'\<sigma>a"
fixes R :: "(('S set \<times> '\<sigma>) \<times> ('Sa set \<times> '\<sigma>a)) set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS[simp]: "S' = \<alpha>`S"
assumes REF0: "((S,\<sigma>0),(\<alpha>`S,\<sigma>0')) \<in> R"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>';
it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S'; \<Phi> it \<sigma>; \<Phi>' it' \<sigma>';
x'=\<alpha> x; it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R;
x\<in>it; x'\<in>it'
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). ((it-{x},\<sigma>),(it'-{x'},\<sigma>'))\<in>R}) (f' x' \<sigma>')"
assumes REF_R_DONE: "\<And>\<sigma> \<sigma>'. \<lbrakk> \<Phi> {} \<sigma>; \<Phi>' {} \<sigma>'; (({},\<sigma>),({},\<sigma>'))\<in>R \<rbrakk>
\<Longrightarrow> (\<sigma>,\<sigma>')\<in>R'"
shows "FOREACHi \<Phi> S f \<sigma>0 \<le> \<Down>R' (FOREACHi \<Phi>' S' f' \<sigma>0')"
unfolding FOREACHi_def
apply (rule FOREACHci_refine_genR)
apply (fact|simp)+
using REFSTEP apply auto []
apply (fact|simp)+
done
lemma FOREACHi_refine:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes REFPHI0: "\<Phi>'' S \<sigma>0 (\<alpha>`S) \<sigma>0'"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; \<Phi>'' it \<sigma> it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi> it \<sigma>; \<Phi>' it' \<sigma>'; \<Phi>'' it \<sigma> it' \<sigma>';
(\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). (\<sigma>, \<sigma>') \<in> R \<and> \<Phi>'' (it - {x}) \<sigma> (it' - {x'}) \<sigma>'}) (f' x' \<sigma>')"
shows "FOREACHi \<Phi> S f \<sigma>0 \<le> \<Down>R (FOREACHi \<Phi>' S' f' \<sigma>0')"
unfolding FOREACHi_def
apply (rule FOREACHci_refine[where \<Phi>''=\<Phi>'', OF INJ REFS REF0 REFPHI0])
apply (rule refl)
apply (erule (5) REFPHI)
apply (erule (9) REFSTEP)
done
lemma FOREACHi_refine_rcg[refine]:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes REFPHI: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> \<Phi> it \<sigma>"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi> it \<sigma>; \<Phi>' it' \<sigma>';
(\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma> \<le> \<Down>R (f' x' \<sigma>')"
shows "FOREACHi \<Phi> S f \<sigma>0 \<le> \<Down>R (FOREACHi \<Phi>' S' f' \<sigma>0')"
unfolding FOREACHi_def
apply (rule FOREACHci_refine_rcg)
apply (rule assms)+
using assms apply auto
done
lemma FOREACH_refine_genR:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa" \<comment> \<open>Abstraction mapping of elements\<close>
fixes S :: "'S set" \<comment> \<open>Concrete set\<close>
fixes S' :: "'Sa set" \<comment> \<open>Abstract set\<close>
fixes \<sigma>0 :: "'\<sigma>"
fixes \<sigma>0' :: "'\<sigma>a"
fixes R :: "(('S set \<times> '\<sigma>) \<times> ('Sa set \<times> '\<sigma>a)) set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS[simp]: "S' = \<alpha>`S"
assumes REF0: "((S,\<sigma>0),(\<alpha>`S,\<sigma>0')) \<in> R"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
it\<subseteq>S; it'\<subseteq>S';
x'=\<alpha> x; it'=\<alpha>`it; ((it,\<sigma>),(it',\<sigma>'))\<in>R;
x\<in>it; x'\<in>it'
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). ((it-{x},\<sigma>),(it'-{x'},\<sigma>'))\<in>R}) (f' x' \<sigma>')"
assumes REF_R_DONE: "\<And>\<sigma> \<sigma>'. \<lbrakk> (({},\<sigma>),({},\<sigma>'))\<in>R \<rbrakk>
\<Longrightarrow> (\<sigma>,\<sigma>')\<in>R'"
shows "FOREACH S f \<sigma>0 \<le> \<Down>R' (FOREACH S' f' \<sigma>0')"
unfolding FOREACH_def
apply (rule FOREACHc_refine_genR)
apply (fact|simp)+
using REFSTEP apply auto []
apply (fact|simp)+
done
lemma FOREACH_refine:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes REFPHI0: "\<Phi>'' S \<sigma>0 (\<alpha>`S) \<sigma>0'"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi>'' it \<sigma> it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma>
\<le> \<Down>({(\<sigma>, \<sigma>'). (\<sigma>, \<sigma>') \<in> R \<and> \<Phi>'' (it - {x}) \<sigma> (it' - {x'}) \<sigma>'}) (f' x' \<sigma>')"
shows "FOREACH S f \<sigma>0 \<le> \<Down>R (FOREACH S' f' \<sigma>0')"
unfolding FOREACH_def
apply (rule FOREACHc_refine[where \<Phi>''=\<Phi>'', OF INJ REFS REF0 REFPHI0])
apply (rule refl)
apply (erule (7) REFSTEP)
done
lemma FOREACH_refine_rcg[refine]:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
(\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma> \<le> \<Down>R (f' x' \<sigma>')"
shows "FOREACH S f \<sigma>0 \<le> \<Down>R (FOREACH S' f' \<sigma>0')"
unfolding FOREACH_def
apply (rule FOREACHc_refine_rcg)
apply (rule assms)+
using assms by auto
lemma FOREACHci_refine_rcg'[refine]:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes REFC: "\<And>it \<sigma> it' \<sigma>'. \<lbrakk>
it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S'; \<Phi>' it' \<sigma>'; (\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> c \<sigma> \<longleftrightarrow> c' \<sigma>'"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi>' it' \<sigma>'; c \<sigma>; c' \<sigma>';
(\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma> \<le> \<Down>R (f' x' \<sigma>')"
shows "FOREACHc S c f \<sigma>0 \<le> \<Down>R (FOREACHci \<Phi>' S' c' f' \<sigma>0')"
unfolding FOREACHc_def
apply (rule FOREACHci_refine_rcg)
apply (rule assms)
apply (rule assms)
apply (rule assms)
apply (erule (4) REFC)
apply (rule TrueI)
apply (rule REFSTEP, assumption+)
done
lemma FOREACHi_refine_rcg'[refine]:
fixes \<alpha> :: "'S \<Rightarrow> 'Sa"
fixes S :: "'S set"
fixes S' :: "'Sa set"
assumes INJ: "inj_on \<alpha> S"
assumes REFS: "S' = \<alpha>`S"
assumes REF0: "(\<sigma>0,\<sigma>0')\<in>R"
assumes REFSTEP: "\<And>x it \<sigma> x' it' \<sigma>'. \<lbrakk>
x'=\<alpha> x; x\<in>it; x'\<in>it'; it'=\<alpha>`it; it\<subseteq>S; it'\<subseteq>S';
\<Phi>' it' \<sigma>';
(\<sigma>,\<sigma>')\<in>R
\<rbrakk> \<Longrightarrow> f x \<sigma> \<le> \<Down>R (f' x' \<sigma>')"
shows "FOREACH S f \<sigma>0 \<le> \<Down>R (FOREACHi \<Phi>' S' f' \<sigma>0')"
unfolding FOREACH_def FOREACHi_def
apply (rule FOREACHci_refine_rcg')
apply (rule assms)+
apply simp
apply (rule REFSTEP, assumption+)
done
subsubsection \<open>Alternative set of FOREACHc-rules\<close>
text \<open>Here, we provide an alternative set of FOREACH rules with
interruption. In some cases, they are easier to use, as they avoid
redundancy between the final cases for interruption and non-interruption\<close>
lemma FOREACHoci_rule':
assumes FIN: "finite S"
assumes I0: "I S \<sigma>0"
assumes IP:
"\<And>x it \<sigma>. \<lbrakk> c \<sigma>; x\<in>it; it\<subseteq>S; I it \<sigma>; \<forall>y\<in>it - {x}. R x y;
\<forall>y\<in>S - it. R y x \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (it-{x}))"
assumes II1: "\<And>\<sigma>. \<lbrakk>I {} \<sigma>; c \<sigma>\<rbrakk> \<Longrightarrow> P \<sigma>"
assumes II2: "\<And>it \<sigma>. \<lbrakk> it\<subseteq>S; I it \<sigma>; \<not>c \<sigma>;
\<forall>x\<in>it. \<forall>y\<in>S - it. R y x \<rbrakk> \<Longrightarrow> P \<sigma>"
shows "FOREACHoci R I S c f \<sigma>0 \<le> SPEC P"
apply (rule FOREACHoci_rule[OF FIN, where I=I, OF I0])
apply (rule IP, assumption+)
apply (case_tac "c \<sigma>")
apply (blast intro: II1)
apply (blast intro: II2)
apply (blast intro: II2)
done
lemma FOREACHci_rule'[refine_vcg]:
assumes FIN: "finite S"
assumes I0: "I S \<sigma>0"
assumes IP:
"\<And>x it \<sigma>. \<lbrakk> x\<in>it; it\<subseteq>S; I it \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (it-{x}))"
assumes II1: "\<And>\<sigma>. \<lbrakk>I {} \<sigma>; c \<sigma>\<rbrakk> \<Longrightarrow> P \<sigma>"
assumes II2: "\<And>it \<sigma>. \<lbrakk> it\<subseteq>S; I it \<sigma>; \<not>c \<sigma> \<rbrakk> \<Longrightarrow> P \<sigma>"
shows "FOREACHci I S c f \<sigma>0 \<le> SPEC P"
unfolding FOREACHci_def
by (rule FOREACHoci_rule') (simp_all add: assms)
lemma FOREACHc_rule':
assumes FIN: "finite S"
assumes I0: "I S \<sigma>0"
assumes IP:
"\<And>x it \<sigma>. \<lbrakk> x\<in>it; it\<subseteq>S; I it \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (it-{x}))"
assumes II1: "\<And>\<sigma>. \<lbrakk>I {} \<sigma>; c \<sigma>\<rbrakk> \<Longrightarrow> P \<sigma>"
assumes II2: "\<And>it \<sigma>. \<lbrakk> it\<subseteq>S; I it \<sigma>; \<not>c \<sigma> \<rbrakk> \<Longrightarrow> P \<sigma>"
shows "FOREACHc S c f \<sigma>0 \<le> SPEC P"
unfolding FOREACHc_def
apply (rule order_trans[OF FOREACHci_weaken], rule TrueI)
apply (rule FOREACHci_rule'[where I=I])
using assms by auto
subsection \<open>FOREACH with empty sets\<close>
lemma FOREACHoci_emp [simp] :
"FOREACHoci R \<Phi> {} c f \<sigma> = do {ASSERT (\<Phi> {} \<sigma>); RETURN \<sigma>}"
by (simp add: FOREACHoci_def FOREACH_cond_def bind_RES)
(simp add: WHILEIT_unfold)
lemma FOREACHoi_emp [simp] :
"FOREACHoi R \<Phi> {} f \<sigma> = do {ASSERT (\<Phi> {} \<sigma>); RETURN \<sigma>}"
by (simp add: FOREACHoi_def)
lemma FOREACHci_emp [simp] :
"FOREACHci \<Phi> {} c f \<sigma> = do {ASSERT (\<Phi> {} \<sigma>); RETURN \<sigma>}"
by (simp add: FOREACHci_def)
lemma FOREACHc_emp [simp] :
"FOREACHc {} c f \<sigma> = RETURN \<sigma>"
by (simp add: FOREACHc_def)
lemma FOREACH_emp [simp] :
"FOREACH {} f \<sigma> = RETURN \<sigma>"
by (simp add: FOREACH_def)
lemma FOREACHi_emp [simp] :
"FOREACHi \<Phi> {} f \<sigma> = do {ASSERT (\<Phi> {} \<sigma>); RETURN \<sigma>}"
by (simp add: FOREACHi_def)
subsection "Monotonicity"
(* TODO: Move to RefineG_Domain *)
definition "lift_refl P c f g == \<forall>x. P c (f x) (g x)"
definition "lift_mono P c f g == \<forall>x y. c x y \<longrightarrow> P c (f x) (g y)"
definition "lift_mono1 P c f g == \<forall>x y. (\<forall>a. c (x a) (y a)) \<longrightarrow> P c (f x) (g y)"
definition "lift_mono2 P c f g == \<forall>x y. (\<forall>a b. c (x a b) (y a b)) \<longrightarrow> P c (f x) (g y)"
definition "trimono_spec L f == ((L id (\<le>) f f) \<and> (L id flat_ge f f))"
lemmas trimono_atomize = atomize_imp atomize_conj atomize_all
lemmas trimono_deatomize = trimono_atomize[symmetric]
lemmas trimono_spec_defs = trimono_spec_def lift_refl_def[abs_def] comp_def id_def
lift_mono_def[abs_def] lift_mono1_def[abs_def] lift_mono2_def[abs_def]
trimono_deatomize
locale trimono_spec begin
abbreviation "R \<equiv> lift_refl"
abbreviation "M \<equiv> lift_mono"
abbreviation "M1 \<equiv> lift_mono1"
abbreviation "M2 \<equiv> lift_mono2"
end
context begin interpretation trimono_spec .
lemma FOREACHoci_mono[unfolded trimono_spec_defs,refine_mono]:
"trimono_spec (R o R o R o R o M2 o R) FOREACHoci"
"trimono_spec (R o R o R o M2 o R) FOREACHoi"
"trimono_spec (R o R o R o M2 o R) FOREACHci"
"trimono_spec (R o R o M2 o R) FOREACHc"
"trimono_spec (R o R o M2 o R) FOREACHi"
"trimono_spec (R o M2 o R) FOREACH"
apply (unfold trimono_spec_defs)
apply -
unfolding FOREACHoci_def FOREACH_to_oci_unfold FOREACH_body_def
apply (refine_mono)+
done
end
subsection \<open>Nres-Fold with Interruption (nfoldli)\<close>
text \<open>
A foreach-loop can be conveniently expressed as an operation that converts
the set to a list, followed by folding over the list.
This representation is handy for automatic refinement, as the complex
foreach-operation is expressed by two relatively simple operations.
\<close>
text \<open>We first define a fold-function in the nres-monad\<close>
partial_function (nrec) nfoldli where
"nfoldli l c f s = (case l of
[] \<Rightarrow> RETURN s
| x#ls \<Rightarrow> if c s then do { s\<leftarrow>f x s; nfoldli ls c f s} else RETURN s
)"
lemma nfoldli_simps[simp]:
"nfoldli [] c f s = RETURN s"
"nfoldli (x#ls) c f s =
(if c s then do { s\<leftarrow>f x s; nfoldli ls c f s} else RETURN s)"
apply (subst nfoldli.simps, simp)+
done
lemma param_nfoldli[param]:
shows "(nfoldli,nfoldli) \<in>
\<langle>Ra\<rangle>list_rel \<rightarrow> (Rb\<rightarrow>Id) \<rightarrow> (Ra\<rightarrow>Rb\<rightarrow>\<langle>Rb\<rangle>nres_rel) \<rightarrow> Rb \<rightarrow> \<langle>Rb\<rangle>nres_rel"
apply (intro fun_relI)
proof goal_cases
case (1 l l' c c' f f' s s')
thus ?case
apply (induct arbitrary: s s')
apply (simp only: nfoldli_simps True_implies_equals)
apply parametricity
apply (simp only: nfoldli_simps True_implies_equals)
apply (parametricity)
done
qed
lemma nfoldli_no_ctd[simp]: "\<not>ctd s \<Longrightarrow> nfoldli l ctd f s = RETURN s"
by (cases l) auto
lemma nfoldli_append[simp]: "nfoldli (l1@l2) ctd f s = nfoldli l1 ctd f s \<bind> nfoldli l2 ctd f"
by (induction l1 arbitrary: s) (auto simp: pw_eq_iff refine_pw_simps)
lemma nfoldli_map: "nfoldli (map f l) ctd g s = nfoldli l ctd (g o f) s"
apply (induction l arbitrary: s)
by (auto simp: pw_eq_iff refine_pw_simps)
lemma nfoldli_nfoldli_prod_conv:
"nfoldli l2 ctd (\<lambda>i. nfoldli l1 ctd (f i)) s = nfoldli (List.product l2 l1) ctd (\<lambda>(i,j). f i j) s"
proof -
have [simp]: "nfoldli (map (Pair a) l) ctd (\<lambda>(x, y). f x y) s = nfoldli l ctd (f a) s"
for a l s
apply (induction l arbitrary: s)
by (auto simp: pw_eq_iff refine_pw_simps)
show ?thesis
by (induction l2 arbitrary: l1 s) auto
qed
text \<open>The fold-function over the nres-monad is transfered to a plain
foldli function\<close>
lemma nfoldli_transfer_plain[refine_transfer]:
assumes "\<And>x s. RETURN (f x s) \<le> f' x s"
shows "RETURN (foldli l c f s) \<le> (nfoldli l c f' s)"
using assms
apply (induct l arbitrary: s)
apply (auto)
by (metis (lifting) plain_bind)
lemma nfoldli_transfer_dres[refine_transfer]:
fixes l :: "'a list" and c:: "'b \<Rightarrow> bool"
assumes FR: "\<And>x s. nres_of (f x s) \<le> f' x s"
shows "nres_of
(foldli l (case_dres False False c) (\<lambda>x s. s\<bind>f x) (dRETURN s))
\<le> (nfoldli l c f' s)"
proof (induct l arbitrary: s)
case Nil thus ?case by auto
next
case (Cons a l)
thus ?case
apply (auto)
apply (cases "f a s")
apply (cases l, simp_all) []
apply simp
apply (rule order_trans[rotated])
apply (rule bind_mono)
apply (rule FR)
apply assumption
apply simp
apply simp
using FR[of a s]
apply simp
done
qed
lemma nfoldli_mono[refine_mono]:
"\<lbrakk> \<And>x s. f x s \<le> f' x s \<rbrakk> \<Longrightarrow> nfoldli l c f \<sigma> \<le> nfoldli l c f' \<sigma>"
"\<lbrakk> \<And>x s. flat_ge (f x s) (f' x s) \<rbrakk> \<Longrightarrow> flat_ge (nfoldli l c f \<sigma>) (nfoldli l c f' \<sigma>)"
apply (induct l arbitrary: \<sigma>)
apply auto
apply refine_mono
apply (induct l arbitrary: \<sigma>)
apply auto
apply refine_mono
done
text \<open>We relate our fold-function to the while-loop that we used in
the original definition of the foreach-loop\<close>
lemma nfoldli_while: "nfoldli l c f \<sigma>
\<le>
(WHILE\<^sub>T\<^bsup>I\<^esup>
(FOREACH_cond c) (FOREACH_body f) (l, \<sigma>) \<bind>
(\<lambda>(_, \<sigma>). RETURN \<sigma>))"
proof (induct l arbitrary: \<sigma>)
case Nil thus ?case by (subst WHILEIT_unfold) (auto simp: FOREACH_cond_def)
next
case (Cons x ls)
show ?case
proof (cases "c \<sigma>")
case False thus ?thesis
apply (subst WHILEIT_unfold)
unfolding FOREACH_cond_def
by simp
next
case [simp]: True
from Cons show ?thesis
apply (subst WHILEIT_unfold)
unfolding FOREACH_cond_def FOREACH_body_def
apply clarsimp
apply (rule Refine_Basic.bind_mono)
apply simp_all
done
qed
qed
lemma while_nfoldli:
"do {
(_,\<sigma>) \<leftarrow> WHILE\<^sub>T (FOREACH_cond c) (FOREACH_body f) (l,\<sigma>);
RETURN \<sigma>
} \<le> nfoldli l c f \<sigma>"
apply (induct l arbitrary: \<sigma>)
apply (subst WHILET_unfold)
apply (simp add: FOREACH_cond_def)
apply (subst WHILET_unfold)
apply (auto
simp: FOREACH_cond_def FOREACH_body_def
intro: bind_mono)
done
lemma while_eq_nfoldli: "do {
(_,\<sigma>) \<leftarrow> WHILE\<^sub>T (FOREACH_cond c) (FOREACH_body f) (l,\<sigma>);
RETURN \<sigma>
} = nfoldli l c f \<sigma>"
apply (rule antisym)
apply (rule while_nfoldli)
apply (rule order_trans[OF nfoldli_while[where I="\<lambda>_. True"]])
apply (simp add: WHILET_def)
done
lemma nfoldli_rule:
assumes I0: "I [] l0 \<sigma>0"
assumes IS: "\<And>x l1 l2 \<sigma>. \<lbrakk> l0=l1@x#l2; I l1 (x#l2) \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (l1@[x]) l2)"
assumes FNC: "\<And>l1 l2 \<sigma>. \<lbrakk> l0=l1@l2; I l1 l2 \<sigma>; \<not>c \<sigma> \<rbrakk> \<Longrightarrow> P \<sigma>"
assumes FC: "\<And>\<sigma>. \<lbrakk> I l0 [] \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> P \<sigma>"
shows "nfoldli l0 c f \<sigma>0 \<le> SPEC P"
apply (rule order_trans[OF nfoldli_while[
where I="\<lambda>(l2,\<sigma>). \<exists>l1. l0=l1@l2 \<and> I l1 l2 \<sigma>"]])
unfolding FOREACH_cond_def FOREACH_body_def
apply (refine_rcg WHILEIT_rule[where R="measure (length o fst)"] refine_vcg)
apply simp
using I0 apply simp
apply (case_tac a, simp)
apply simp
apply (elim exE conjE)
apply (rule order_trans[OF IS], assumption+)
apply auto []
apply simp
apply (elim exE disjE2)
using FC apply auto []
using FNC apply auto []
done
lemma nfoldli_leof_rule:
assumes I0: "I [] l0 \<sigma>0"
assumes IS: "\<And>x l1 l2 \<sigma>. \<lbrakk> l0=l1@x#l2; I l1 (x#l2) \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> f x \<sigma> \<le>\<^sub>n SPEC (I (l1@[x]) l2)"
assumes FNC: "\<And>l1 l2 \<sigma>. \<lbrakk> l0=l1@l2; I l1 l2 \<sigma>; \<not>c \<sigma> \<rbrakk> \<Longrightarrow> P \<sigma>"
assumes FC: "\<And>\<sigma>. \<lbrakk> I l0 [] \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> P \<sigma>"
shows "nfoldli l0 c f \<sigma>0 \<le>\<^sub>n SPEC P"
proof -
{
fix l1 l2 \<sigma>
assume "l0=l1@l2" "I l1 l2 \<sigma>"
hence "nfoldli l2 c f \<sigma> \<le>\<^sub>n SPEC P"
proof (induction l2 arbitrary: l1 \<sigma>)
case Nil thus ?case
apply simp
apply (cases "c \<sigma>")
apply (rule FC; auto; fail)
apply (rule FNC[of l1 "[]"]; auto; fail)
done
next
case (Cons x l2)
note [refine_vcg] = Cons.IH[of "l1@[x]",THEN leof_trans] IS[of l1 x l2 \<sigma>,THEN leof_trans]
show ?case
apply (simp split del: if_split)
apply refine_vcg
using Cons.prems FNC by auto
qed
} from this[of "[]" l0 \<sigma>0] I0 show ?thesis by auto
qed
lemma nfoldli_refine[refine]:
assumes "(li, l) \<in> \<langle>S\<rangle>list_rel"
and "(si, s) \<in> R"
and CR: "(ci, c) \<in> R \<rightarrow> bool_rel"
and [refine]: "\<And>xi x si s. \<lbrakk> (xi,x)\<in>S; (si,s)\<in>R; c s \<rbrakk> \<Longrightarrow> fi xi si \<le> \<Down>R (f x s)"
shows "nfoldli li ci fi si \<le> \<Down> R (nfoldli l c f s)"
using assms(1,2)
proof (induction arbitrary: si s rule: list_rel_induct)
case Nil thus ?case by simp
next
case (Cons xi x li l)
note [refine] = Cons
show ?case
apply (simp split del: if_split)
apply refine_rcg
using CR Cons.prems by (auto dest: fun_relD)
qed
(* Refine, establishing additional invariant *)
lemma nfoldli_invar_refine:
assumes "(li,l)\<in>\<langle>S\<rangle>list_rel"
assumes "(si,s)\<in>R"
assumes "I [] li si"
assumes COND: "\<And>l1i l2i l1 l2 si s. \<lbrakk>
li=l1i@l2i; l=l1@l2; (l1i,l1)\<in>\<langle>S\<rangle>list_rel; (l2i,l2)\<in>\<langle>S\<rangle>list_rel;
I l1i l2i si; (si,s)\<in>R\<rbrakk> \<Longrightarrow> (ci si, c s)\<in>bool_rel"
assumes INV: "\<And>l1i xi l2i si. \<lbrakk>li=l1i@xi#l2i; I l1i (xi#l2i) si\<rbrakk> \<Longrightarrow> fi xi si \<le>\<^sub>n SPEC (I (l1i@[xi]) l2i)"
assumes STEP: "\<And>l1i xi l2i l1 x l2 si s. \<lbrakk>
li=l1i@xi#l2i; l=l1@x#l2; (l1i,l1)\<in>\<langle>S\<rangle>list_rel; (xi,x)\<in>S; (l2i,l2)\<in>\<langle>S\<rangle>list_rel;
I l1i (xi#l2i) si; (si,s)\<in>R\<rbrakk> \<Longrightarrow> fi xi si \<le> \<Down>R (f x s)"
shows "nfoldli li ci fi si \<le> \<Down>R (nfoldli l c f s)"
proof -
{
have [refine_dref_RELATES]: "RELATES R" "RELATES S" by (auto simp: RELATES_def)
note [refine del] = nfoldli_refine
fix l1i l2i l1 l2 si s
assume "(l2i,l2) \<in> \<langle>S\<rangle>list_rel" "(l1i,l1) \<in> \<langle>S\<rangle>list_rel"
and "li=l1i@l2i" "l=l1@l2"
and "(si,s)\<in>R" "I l1i l2i si"
hence "nfoldli l2i ci fi si \<le> \<Down>R (nfoldli l2 c f s)"
proof (induction arbitrary: si s l1i l1 rule: list_rel_induct)
case Nil thus ?case by auto
next
case (Cons xi x l2i l2)
show ?case
apply (simp split del: if_split)
apply (refine_rcg bind_refine')
apply (refine_dref_type)
subgoal using COND[of l1i "xi#l2i" l1 "x#l2" si s] Cons.prems Cons.hyps by auto
subgoal apply (rule STEP) using Cons.prems Cons.hyps by auto
subgoal for si' s'
apply (rule Cons.IH[of "l1i@[xi]" "l1@[x]"])
using Cons.prems Cons.hyps
apply (auto simp: list_rel_append1) apply force
using INV[of l1i xi l2i si]
by (auto simp: pw_leof_iff)
subgoal using Cons.prems by simp
done
qed
}
from this[of li l "[]" "[]" si s] assms(1,2,3) show ?thesis by auto
qed
lemma foldli_mono_dres_aux1:
fixes \<sigma> :: "'a :: {order_bot, order_top}"
assumes COND: "\<And>\<sigma> \<sigma>'. \<sigma>\<le>\<sigma>' \<Longrightarrow> c \<sigma> \<noteq> c \<sigma>' \<Longrightarrow> \<sigma>=bot \<or> \<sigma>'=top "
assumes STRICT: "\<And>x. f x bot = bot" "\<And>x. f' x top = top"
assumes B: "\<sigma>\<le>\<sigma>'"
assumes A: "\<And>a x x'. x\<le>x' \<Longrightarrow> f a x \<le> f' a x'"
shows "foldli l c f \<sigma> \<le> foldli l c f' \<sigma>'"
proof -
{ fix l
have "foldli l c f bot = bot" by (induct l) (auto simp: STRICT)
} note [simp] = this
{ fix l
have "foldli l c f' top = top" by (induct l) (auto simp: STRICT)
} note [simp] = this
show ?thesis
using B
apply (induct l arbitrary: \<sigma> \<sigma>')
apply (auto simp: A STRICT dest!: COND)
done
qed
lemma foldli_mono_dres_aux2:
assumes STRICT: "\<And>x. f x bot = bot" "\<And>x. f' x top = top"
assumes A: "\<And>a x x'. x\<le>x' \<Longrightarrow> f a x \<le> f' a x'"
shows "foldli l (case_dres False False c) f \<sigma>
\<le> foldli l (case_dres False False c) f' \<sigma>"
apply (rule foldli_mono_dres_aux1)
apply (simp_all split: dres.split_asm add: STRICT A)
done
lemma foldli_mono_dres[refine_mono]:
assumes A: "\<And>a x. f a x \<le> f' a x"
shows "foldli l (case_dres False False c) (\<lambda>x s. dbind s (f x)) \<sigma>
\<le> foldli l (case_dres False False c) (\<lambda>x s. dbind s (f' x)) \<sigma>"
apply (rule foldli_mono_dres_aux2)
apply (simp_all)
apply (rule dbind_mono)
apply (simp_all add: A)
done
partial_function (drec) dfoldli where
"dfoldli l c f s = (case l of
[] \<Rightarrow> dRETURN s
| x#ls \<Rightarrow> if c s then do { s\<leftarrow>f x s; dfoldli ls c f s} else dRETURN s
)"
lemma dfoldli_simps[simp]:
"dfoldli [] c f s = dRETURN s"
"dfoldli (x#ls) c f s =
(if c s then do { s\<leftarrow>f x s; dfoldli ls c f s} else dRETURN s)"
apply (subst dfoldli.simps, simp)+
done
lemma dfoldli_mono[refine_mono]:
"\<lbrakk> \<And>x s. f x s \<le> f' x s \<rbrakk> \<Longrightarrow> dfoldli l c f \<sigma> \<le> dfoldli l c f' \<sigma>"
"\<lbrakk> \<And>x s. flat_ge (f x s) (f' x s) \<rbrakk> \<Longrightarrow> flat_ge (dfoldli l c f \<sigma>) (dfoldli l c f' \<sigma>)"
apply (induct l arbitrary: \<sigma>)
apply auto
apply refine_mono
apply (induct l arbitrary: \<sigma>)
apply auto
apply refine_mono
done
lemma foldli_dres_pres_FAIL[simp]:
"foldli l (case_dres False False c) (\<lambda>x s. dbind s (f x)) dFAIL = dFAIL"
by (cases l) auto
lemma foldli_dres_pres_SUCCEED[simp]:
"foldli l (case_dres False False c) (\<lambda>x s. dbind s (f x)) dSUCCEED = dSUCCEED"
by (cases l) auto
lemma dfoldli_by_foldli: "dfoldli l c f \<sigma>
= foldli l (case_dres False False c) (\<lambda>x s. dbind s (f x)) (dRETURN \<sigma>)"
apply (induction l arbitrary: \<sigma>)
apply simp
apply (clarsimp intro!: ext)
apply (rename_tac a l x)
apply (case_tac "f a x")
apply auto
done
lemma foldli_mono_dres_flat[refine_mono]:
assumes A: "\<And>a x. flat_ge (f a x) (f' a x)"
shows "flat_ge (foldli l (case_dres False False c) (\<lambda>x s. dbind s (f x)) \<sigma>)
(foldli l (case_dres False False c) (\<lambda>x s. dbind s (f' x)) \<sigma>)"
apply (cases \<sigma>)
apply (simp_all add: dfoldli_by_foldli[symmetric])
using A apply refine_mono
done
lemma dres_foldli_ne_bot[refine_transfer]:
assumes 1: "\<sigma> \<noteq> dSUCCEED"
assumes 2: "\<And>x \<sigma>. f x \<sigma> \<noteq> dSUCCEED"
shows "foldli l c (\<lambda>x s. s \<bind> f x) \<sigma> \<noteq> dSUCCEED"
using 1 apply (induct l arbitrary: \<sigma>)
apply simp
apply (simp split: dres.split, intro allI impI)
apply rprems
using 2
apply (simp add: dres_ne_bot_basic)
done
subsection \<open>LIST FOREACH combinator\<close>
text \<open>
Foreach-loops are mapped to the combinator \<open>LIST_FOREACH\<close>, that
takes as first argument an explicit \<open>to_list\<close> operation.
This mapping is done during operation identification.
It is then the responsibility of the various implementations to further map
the \<open>to_list\<close> operations to custom \<open>to_list\<close> operations, like
\<open>set_to_list\<close>, \<open>map_to_list\<close>, \<open>nodes_to_list\<close>, etc.
\<close>
text \<open>We define a relation between distinct lists and sets.\<close>
definition [to_relAPP]: "list_set_rel R \<equiv> \<langle>R\<rangle>list_rel O br set distinct"
lemma autoref_nfoldli[autoref_rules]:
shows "(nfoldli, nfoldli)
\<in> \<langle>Ra\<rangle>list_rel \<rightarrow> (Rb \<rightarrow> bool_rel) \<rightarrow> (Ra \<rightarrow> Rb \<rightarrow> \<langle>Rb\<rangle>nres_rel) \<rightarrow> Rb \<rightarrow> \<langle>Rb\<rangle>nres_rel"
by (rule param_nfoldli)
text \<open>This constant is a placeholder to be converted to
custom operations by pattern rules\<close>
definition "it_to_sorted_list R s
\<equiv> SPEC (\<lambda>l. distinct l \<and> s = set l \<and> sorted_wrt R l)"
definition "LIST_FOREACH \<Phi> tsl c f \<sigma>0 \<equiv> do {
xs \<leftarrow> tsl;
(_,\<sigma>) \<leftarrow> WHILE\<^sub>T\<^bsup>\<lambda>(it, \<sigma>). \<exists>xs'. xs = xs' @ it \<and> \<Phi> (set it) \<sigma>\<^esup>
(FOREACH_cond c) (FOREACH_body f) (xs, \<sigma>0);
RETURN \<sigma>}"
lemma FOREACHoci_by_LIST_FOREACH:
"FOREACHoci R \<Phi> S c f \<sigma>0 = do {
ASSERT (finite S);
LIST_FOREACH \<Phi> (it_to_sorted_list R S) c f \<sigma>0
}"
unfolding OP_def FOREACHoci_def LIST_FOREACH_def it_to_sorted_list_def
by simp
text \<open>Patterns that convert FOREACH-constructs
to \<open>LIST_FOREACH\<close>
\<close>
context begin interpretation autoref_syn .
lemma FOREACH_patterns[autoref_op_pat_def]:
"FOREACH\<^bsup>I\<^esup> s f \<equiv> FOREACH\<^sub>O\<^sub>C\<^bsup>\<lambda>_ _. True,I\<^esup> s (\<lambda>_. True) f"
"FOREACHci I s c f \<equiv> FOREACHoci (\<lambda>_ _. True) I s c f"
"FOREACH\<^sub>O\<^sub>C\<^bsup>R,\<Phi>\<^esup> s c f \<equiv> \<lambda>\<sigma>. do {
ASSERT (finite s);
Autoref_Tagging.OP (LIST_FOREACH \<Phi>) (it_to_sorted_list R s) c f \<sigma>
}"
"FOREACH s f \<equiv> FOREACHoci (\<lambda>_ _. True) (\<lambda>_ _. True) s (\<lambda>_. True) f"
"FOREACHoi R I s f \<equiv> FOREACHoci R I s (\<lambda>_. True) f"
"FOREACHc s c f \<equiv> FOREACHoci (\<lambda>_ _. True) (\<lambda>_ _. True) s c f"
unfolding
FOREACHoci_by_LIST_FOREACH[abs_def]
FOREACHc_def[abs_def]
FOREACH_def[abs_def]
FOREACHci_def[abs_def]
FOREACHi_def[abs_def]
FOREACHoi_def[abs_def]
by simp_all
(*lemma FOREACH_patterns[autoref_op_pat]:
"FOREACHoci R \<Phi> s c f \<sigma> \<equiv> do {
ASSERT (finite s);
OP (LIST_FOREACH \<Phi>) (it_to_sorted_list R s) c f \<sigma>
}"
"FOREACHc s c f \<sigma> \<equiv> FOREACHoci (\<lambda>_ _. True) (\<lambda>_ _. True) s c f \<sigma>"
"FOREACH s f \<sigma> \<equiv> FOREACHoci (\<lambda>_ _. True) (\<lambda>_ _. True) s (\<lambda>_. True) f \<sigma>"
"FOREACHci I s c f \<sigma> \<equiv> FOREACHoci (\<lambda>_ _. True) I s c f \<sigma>"
"FOREACHi I s f \<sigma> \<equiv> FOREACHoci (\<lambda>_ _. True) I s (\<lambda>_. True) f \<sigma>"
"FOREACHoi R I s f \<sigma> \<equiv> FOREACHoci R I s (\<lambda>_. True) f \<sigma>"
unfolding
FOREACHoci_by_LIST_FOREACH[abs_def]
FOREACHc_def[abs_def]
FOREACH_def[abs_def]
FOREACHci_def[abs_def]
FOREACHi_def[abs_def]
FOREACHoi_def[abs_def]
by simp_all*)
end
definition "LIST_FOREACH' tsl c f \<sigma> \<equiv> do {xs \<leftarrow> tsl; nfoldli xs c f \<sigma>}"
lemma LIST_FOREACH'_param[param]:
shows "(LIST_FOREACH',LIST_FOREACH')
\<in> (\<langle>\<langle>Rv\<rangle>list_rel\<rangle>nres_rel \<rightarrow> (R\<sigma>\<rightarrow>bool_rel)
\<rightarrow> (Rv \<rightarrow> R\<sigma> \<rightarrow> \<langle>R\<sigma>\<rangle>nres_rel) \<rightarrow> R\<sigma> \<rightarrow> \<langle>R\<sigma>\<rangle>nres_rel)"
unfolding LIST_FOREACH'_def[abs_def]
by parametricity
lemma LIST_FOREACH_autoref[autoref_rules]:
shows "(LIST_FOREACH', LIST_FOREACH \<Phi>) \<in>
(\<langle>\<langle>Rv\<rangle>list_rel\<rangle>nres_rel \<rightarrow> (R\<sigma>\<rightarrow>bool_rel)
\<rightarrow> (Rv \<rightarrow> R\<sigma> \<rightarrow> \<langle>R\<sigma>\<rangle>nres_rel) \<rightarrow> R\<sigma> \<rightarrow> \<langle>R\<sigma>\<rangle>nres_rel)"
proof (intro fun_relI nres_relI)
fix tsl tsl' c c' f f' \<sigma> \<sigma>'
assume [param]:
"(tsl,tsl')\<in>\<langle>\<langle>Rv\<rangle>list_rel\<rangle>nres_rel"
"(c,c')\<in>R\<sigma>\<rightarrow>bool_rel"
"(f,f')\<in>Rv\<rightarrow>R\<sigma>\<rightarrow>\<langle>R\<sigma>\<rangle>nres_rel"
"(\<sigma>,\<sigma>')\<in>R\<sigma>"
have "LIST_FOREACH' tsl c f \<sigma> \<le> \<Down>R\<sigma> (LIST_FOREACH' tsl' c' f' \<sigma>')"
apply (rule nres_relD)
by parametricity
also have "LIST_FOREACH' tsl' c' f' \<sigma>'
\<le> LIST_FOREACH \<Phi> tsl' c' f' \<sigma>'"
apply (rule refine_IdD)
unfolding LIST_FOREACH_def LIST_FOREACH'_def
apply refine_rcg
apply simp
apply (rule nfoldli_while)
done
finally show
"LIST_FOREACH' tsl c f \<sigma> \<le> \<Down> R\<sigma> (LIST_FOREACH \<Phi> tsl' c' f' \<sigma>')"
.
qed
context begin interpretation trimono_spec .
lemma LIST_FOREACH'_mono[unfolded trimono_spec_defs,refine_mono]:
"trimono_spec (R o R o M2 o R) LIST_FOREACH'"
apply (unfold trimono_spec_defs)
apply -
unfolding LIST_FOREACH'_def
by refine_mono+
end
lemma LIST_FOREACH'_transfer_plain[refine_transfer]:
assumes "RETURN tsl \<le> tsl'"
assumes "\<And>x \<sigma>. RETURN (f x \<sigma>) \<le> f' x \<sigma>"
shows "RETURN (foldli tsl c f \<sigma>) \<le> LIST_FOREACH' tsl' c f' \<sigma>"
apply (rule order_trans[rotated])
unfolding LIST_FOREACH'_def
using assms
apply refine_transfer
by simp
thm refine_transfer
lemma LIST_FOREACH'_transfer_nres[refine_transfer]:
assumes "nres_of tsl \<le> tsl'"
assumes "\<And>x \<sigma>. nres_of (f x \<sigma>) \<le> f' x \<sigma>"
shows "nres_of (
do {
xs\<leftarrow>tsl;
foldli xs (case_dres False False c) (\<lambda>x s. s\<bind>f x) (dRETURN \<sigma>)
}) \<le> LIST_FOREACH' tsl' c f' \<sigma>"
unfolding LIST_FOREACH'_def
using assms
by refine_transfer
text \<open>Simplification rules to summarize iterators\<close>
lemma [refine_transfer_post_simp]:
"do {
xs \<leftarrow> dRETURN tsl;
foldli xs c f \<sigma>
} = foldli tsl c f \<sigma>"
by simp
lemma [refine_transfer_post_simp]:
"(let xs = tsl in foldli xs c f \<sigma>) = foldli tsl c f \<sigma>"
by simp
lemma LFO_pre_refine: (* TODO: Generalize! *)
assumes "(li,l)\<in>\<langle>A\<rangle>list_set_rel"
assumes "(ci,c)\<in>R \<rightarrow> bool_rel"
assumes "(fi,f)\<in>A\<rightarrow>R\<rightarrow>\<langle>R\<rangle>nres_rel"
assumes "(s0i,s0)\<in>R"
shows "LIST_FOREACH' (RETURN li) ci fi s0i \<le> \<Down>R (FOREACHci I l c f s0)"
proof -
from assms(1) have [simp]: "finite l" by (auto simp: list_set_rel_def br_def)
show ?thesis
unfolding FOREACHc_def FOREACHci_def FOREACHoci_by_LIST_FOREACH
apply simp
apply (rule LIST_FOREACH_autoref[param_fo, THEN nres_relD])
using assms
apply auto
apply (auto simp: it_to_sorted_list_def nres_rel_def pw_le_iff refine_pw_simps
list_set_rel_def br_def)
done
qed
lemma LFOci_refine: (* TODO: Generalize! *)
assumes "(li,l)\<in>\<langle>A\<rangle>list_set_rel"
assumes "\<And>s si. (si,s)\<in>R \<Longrightarrow> ci si \<longleftrightarrow> c s"
assumes "\<And>x xi s si. \<lbrakk>(xi,x)\<in>A; (si,s)\<in>R\<rbrakk> \<Longrightarrow> fi xi si \<le> \<Down>R (f x s)"
assumes "(s0i,s0)\<in>R"
shows "nfoldli li ci fi s0i \<le> \<Down>R (FOREACHci I l c f s0)"
proof -
from assms LFO_pre_refine[of li l A ci c R fi f s0i s0] show ?thesis
unfolding fun_rel_def nres_rel_def LIST_FOREACH'_def
apply (simp add: pw_le_iff refine_pw_simps)
apply blast+
done
qed
lemma LFOc_refine: (* TODO: Generalize! *)
assumes "(li,l)\<in>\<langle>A\<rangle>list_set_rel"
assumes "\<And>s si. (si,s)\<in>R \<Longrightarrow> ci si \<longleftrightarrow> c s"
assumes "\<And>x xi s si. \<lbrakk>(xi,x)\<in>A; (si,s)\<in>R\<rbrakk> \<Longrightarrow> fi xi si \<le> \<Down>R (f x s)"
assumes "(s0i,s0)\<in>R"
shows "nfoldli li ci fi s0i \<le> \<Down>R (FOREACHc l c f s0)"
unfolding FOREACHc_def
by (rule LFOci_refine[OF assms])
lemma LFO_refine: (* TODO: Generalize! *)
assumes "(li,l)\<in>\<langle>A\<rangle>list_set_rel"
assumes "\<And>x xi s si. \<lbrakk>(xi,x)\<in>A; (si,s)\<in>R\<rbrakk> \<Longrightarrow> fi xi si \<le> \<Down>R (f x s)"
assumes "(s0i,s0)\<in>R"
shows "nfoldli li (\<lambda>_. True) fi s0i \<le> \<Down>R (FOREACH l f s0)"
unfolding FOREACH_def
apply (rule LFOc_refine)
apply (rule assms | simp)+
done
lemma LFOi_refine: (* TODO: Generalize! *)
assumes "(li,l)\<in>\<langle>A\<rangle>list_set_rel"
assumes "\<And>x xi s si. \<lbrakk>(xi,x)\<in>A; (si,s)\<in>R\<rbrakk> \<Longrightarrow> fi xi si \<le> \<Down>R (f x s)"
assumes "(s0i,s0)\<in>R"
shows "nfoldli li (\<lambda>_. True) fi s0i \<le> \<Down>R (FOREACHi I l f s0)"
unfolding FOREACHi_def
apply (rule LFOci_refine)
apply (rule assms | simp)+
done
lemma LIST_FOREACH'_refine: "LIST_FOREACH' tsl' c' f' \<sigma>' \<le> LIST_FOREACH \<Phi> tsl' c' f' \<sigma>'"
apply (rule refine_IdD)
unfolding LIST_FOREACH_def LIST_FOREACH'_def
apply refine_rcg
apply simp
apply (rule nfoldli_while)
done
lemma LIST_FOREACH'_eq: "LIST_FOREACH (\<lambda>_ _. True) tsl' c' f' \<sigma>' = (LIST_FOREACH' tsl' c' f' \<sigma>')"
apply (rule antisym)
subgoal
apply (rule refine_IdD)
unfolding LIST_FOREACH_def LIST_FOREACH'_def while_eq_nfoldli[symmetric]
apply (refine_rcg WHILEIT_refine_new_invar)
unfolding WHILET_def
apply (rule WHILEIT_refine_new_invar)
apply refine_dref_type
apply clarsimp_all
unfolding FOREACH_body_def FOREACH_cond_def
apply refine_vcg
apply (auto simp: refine_pw_simps pw_le_iff neq_Nil_conv)
done
subgoal by (rule LIST_FOREACH'_refine)
done
subsection \<open>FOREACH with duplicates\<close>
definition "FOREACHcd S c f \<sigma> \<equiv> do {
ASSERT (finite S);
l \<leftarrow> SPEC (\<lambda>l. set l = S);
nfoldli l c f \<sigma>
}"
lemma FOREACHcd_rule:
assumes "finite S\<^sub>0"
assumes I0: "I {} S\<^sub>0 \<sigma>\<^sub>0"
assumes STEP: "\<And>S1 S2 x \<sigma>. \<lbrakk> S\<^sub>0 = insert x (S1\<union>S2); I S1 (insert x S2) \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (insert x S1) S2)"
assumes INTR: "\<And>S1 S2 \<sigma>. \<lbrakk> S\<^sub>0 = S1\<union>S2; I S1 S2 \<sigma>; \<not>c \<sigma> \<rbrakk> \<Longrightarrow> \<Phi> \<sigma>"
assumes COMPL: "\<And>\<sigma>. \<lbrakk> I S\<^sub>0 {} \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> \<Phi> \<sigma>"
shows "FOREACHcd S\<^sub>0 c f \<sigma>\<^sub>0 \<le> SPEC \<Phi>"
unfolding FOREACHcd_def
apply refine_vcg
apply fact
apply (rule nfoldli_rule[where I = "\<lambda>l1 l2 \<sigma>. I (set l1) (set l2) \<sigma>"])
subgoal using I0 by auto
subgoal using STEP by auto
subgoal using INTR by auto
subgoal using COMPL by auto
done
definition FOREACHcdi
:: "('a set \<Rightarrow> 'a set \<Rightarrow> 'b \<Rightarrow> bool)
\<Rightarrow> 'a set \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'b nres) \<Rightarrow> 'b \<Rightarrow> 'b nres"
where
"FOREACHcdi I \<equiv> FOREACHcd"
lemma FOREACHcdi_rule[refine_vcg]:
assumes "finite S\<^sub>0"
assumes I0: "I {} S\<^sub>0 \<sigma>\<^sub>0"
assumes STEP: "\<And>S1 S2 x \<sigma>. \<lbrakk> S\<^sub>0 = insert x (S1\<union>S2); I S1 (insert x S2) \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> f x \<sigma> \<le> SPEC (I (insert x S1) S2)"
assumes INTR: "\<And>S1 S2 \<sigma>. \<lbrakk> S\<^sub>0 = S1\<union>S2; I S1 S2 \<sigma>; \<not>c \<sigma> \<rbrakk> \<Longrightarrow> \<Phi> \<sigma>"
assumes COMPL: "\<And>\<sigma>. \<lbrakk> I S\<^sub>0 {} \<sigma>; c \<sigma> \<rbrakk> \<Longrightarrow> \<Phi> \<sigma>"
shows "FOREACHcdi I S\<^sub>0 c f \<sigma>\<^sub>0 \<le> SPEC \<Phi>"
unfolding FOREACHcdi_def
using assms
by (rule FOREACHcd_rule)
lemma FOREACHcd_refine[refine]:
assumes [simp]: "finite s'"
assumes S: "(s',s)\<in>\<langle>S\<rangle>set_rel"
assumes SV: "single_valued S"
assumes R0: "(\<sigma>',\<sigma>)\<in>R"
assumes C: "\<And>\<sigma>' \<sigma>. (\<sigma>',\<sigma>)\<in>R \<Longrightarrow> (c' \<sigma>', c \<sigma>)\<in>bool_rel"
assumes F: "\<And>x' x \<sigma>' \<sigma>. \<lbrakk>(x', x) \<in> S; (\<sigma>', \<sigma>) \<in> R\<rbrakk>
\<Longrightarrow> f' x' \<sigma>' \<le> \<Down> R (f x \<sigma>)"
shows "FOREACHcd s' c' f' \<sigma>' \<le> \<Down>R (FOREACHcdi I s c f \<sigma>)"
proof -
have [refine_dref_RELATES]: "RELATES S" by (simp add: RELATES_def)
from SV obtain \<alpha> I where [simp]: "S=br \<alpha> I" by (rule single_valued_as_brE)
with S have [simp]: "s=\<alpha>`s'" and [simp]: "\<forall>x\<in>s'. I x"
by (auto simp: br_set_rel_alt)
show ?thesis
unfolding FOREACHcd_def FOREACHcdi_def
apply refine_rcg
apply refine_dref_type
subgoal by simp
subgoal
apply (auto simp: pw_le_iff refine_pw_simps)
using S
apply (rule_tac x="map \<alpha> x" in exI)
apply (auto simp: map_in_list_rel_conv)
done
subgoal using R0 by auto
subgoal using C by auto
subgoal using F by auto
done
qed
lemma FOREACHc_refines_FOREACHcd_aux:
shows "FOREACHc s c f \<sigma> \<le> FOREACHcd s c f \<sigma>"
unfolding FOREACHc_def FOREACHci_def FOREACHoci_by_LIST_FOREACH LIST_FOREACH'_eq
LIST_FOREACH'_def FOREACHcd_def it_to_sorted_list_def
apply (rule refine_IdD)
apply (refine_rcg)
apply refine_dref_type
apply auto
done
lemmas FOREACHc_refines_FOREACHcd[refine]
= order_trans[OF FOREACHc_refines_FOREACHcd_aux FOREACHcd_refine]
subsection \<open>Miscellanneous Utility Lemmas\<close>
(* TODO: Can we make this somewhat more general ? *)
lemma map_foreach:
assumes "finite S"
shows "FOREACH S (\<lambda>x \<sigma>. RETURN (insert (f x) \<sigma>)) R0 \<le> SPEC ((=) (R0 \<union> f`S))"
apply (rule FOREACH_rule[where I="\<lambda>it \<sigma>. \<sigma>=R0 \<union> f`(S-it)"])
apply (auto intro: assms)
done
lemma map_sigma_foreach:
fixes f :: "'a \<times> 'b \<Rightarrow> 'c"
assumes "finite A"
assumes "\<And>x. x\<in>A \<Longrightarrow> finite (B x)"
shows "FOREACH A (\<lambda>a \<sigma>.
FOREACH (B a) (\<lambda>b \<sigma>. RETURN (insert (f (a,b)) \<sigma>)) \<sigma>
) R0 \<le> SPEC ((=) (R0 \<union> f`Sigma A B))"
apply (rule FOREACH_rule[where I="\<lambda>it \<sigma>. \<sigma>=R0 \<union> f`(Sigma (A-it) B)"])
apply (auto intro: assms) [2]
apply (rule_tac I="\<lambda>it' \<sigma>. \<sigma>=R0 \<union> f`(Sigma (A - it) B)
\<union> f`({x} \<times> (B x - it'))"
in FOREACH_rule)
apply (auto intro: assms) [2]
apply (rule refine_vcg)
apply auto []
apply auto []
apply auto []
done
lemma map_sigma_sigma_foreach:
fixes f :: "'a \<times> ('b \<times> 'c) \<Rightarrow> 'd"
assumes "finite A"
assumes "\<And>a. a\<in>A \<Longrightarrow> finite (B a)"
assumes "\<And>a b. \<lbrakk>a\<in>A; b\<in>B a\<rbrakk> \<Longrightarrow> finite (C a b)"
shows "FOREACH A (\<lambda>a \<sigma>.
FOREACH (B a) (\<lambda>b \<sigma>.
FOREACH (C a b) (\<lambda>c \<sigma>.
RETURN (insert (f (a,(b,c))) \<sigma>)) \<sigma>) \<sigma>
) R0 \<le> SPEC ((=) (R0 \<union> f`Sigma A (\<lambda>a. Sigma (B a) (C a))))"
apply (rule FOREACH_rule[where
I="\<lambda>it \<sigma>. \<sigma>=R0 \<union> f`(Sigma (A-it) (\<lambda>a. Sigma (B a) (C a)))"])
apply (auto intro: assms) [2]
apply (rule_tac
I="\<lambda>it' \<sigma>. \<sigma>=R0 \<union> f`(Sigma (A - it) (\<lambda>a. Sigma (B a) (C a)))
\<union> f`({x} \<times> ( Sigma (B x - it') (C x)))"
in FOREACH_rule)
apply (auto intro: assms) [2]
apply (rule_tac
I="\<lambda>it'' \<sigma>. \<sigma>=R0 \<union> f`(Sigma (A - it) (\<lambda>a. Sigma (B a) (C a)))
\<union> f`({x} \<times> ( Sigma (B x - ita) (C x)))
\<union> f`({x} \<times> ({xa} \<times> (C x xa - it'')))
"
in FOREACH_rule)
apply (auto intro: assms) [2]
apply auto
done
lemma bij_set_rel_for_inj:
fixes R
defines "\<alpha> \<equiv> fun_of_rel R"
assumes "bijective R" "(s,s')\<in>\<langle>R\<rangle>set_rel"
shows "inj_on \<alpha> s" "s' = \<alpha>`s"
\<comment> \<open>To be used when generating refinement conditions for foreach-loops\<close>
using assms
unfolding bijective_def set_rel_def \<alpha>_def fun_of_rel_def[abs_def]
apply (auto intro!: inj_onI ImageI simp: image_def)
apply (metis (mono_tags) Domain.simps contra_subsetD tfl_some)
apply (metis (mono_tags) someI)
apply (metis DomainE contra_subsetD tfl_some)
done
lemma nfoldli_by_idx_gen:
shows "nfoldli (drop k l) c f s = nfoldli [k..<length l] c (\<lambda>i s. do {
ASSERT (i<length l);
let x = l!i;
f x s
}) s"
proof (cases "k\<le>length l")
case False thus ?thesis by auto
next
case True thus ?thesis
proof (induction arbitrary: s rule: inc_induct)
case base thus ?case
by auto
next
case (step k)
from step.hyps have 1: "drop k l = l!k # drop (Suc k) l"
by (auto simp: Cons_nth_drop_Suc)
from step.hyps have 2: "[k..<length l] = k#[Suc k..<length l]"
by (auto simp: upt_conv_Cons)
show ?case
unfolding 1 2
by (auto simp: step.IH[abs_def] step.hyps)
qed
qed
lemma nfoldli_by_idx:
"nfoldli l c f s = nfoldli [0..<length l] c (\<lambda>i s. do {
ASSERT (i<length l);
let x = l!i;
f x s
}) s"
using nfoldli_by_idx_gen[of 0] by auto
lemma nfoldli_map_inv:
assumes "inj g"
shows "nfoldli l c f = nfoldli (map g l) c (\<lambda>x s. f (the_inv g x) s)"
using assms
apply (induction l)
subgoal by auto
subgoal by (auto simp: the_inv_f_f)
done
lemma nfoldli_shift:
fixes ofs :: nat
shows "nfoldli l c f = nfoldli (map (\<lambda>i. i+ofs) l) c (\<lambda>x s. do {ASSERT (x\<ge>ofs); f (x - ofs) s})"
by (induction l) auto
lemma nfoldli_foreach_shift:
shows "nfoldli [a..<b] c f = nfoldli [a+ofs..<b+ofs] c (\<lambda>x s. do{ASSERT(x\<ge>ofs); f (x - ofs) s})"
using nfoldli_shift[of "[a..<b]" c f ofs]
by (auto simp: map_add_upt')
(* TODO: Move. Probably we have this already with FOREACH, or iterator! *)
lemma member_by_nfoldli: "nfoldli l (\<lambda>f. \<not>f) (\<lambda>y _. RETURN (y=x)) False \<le> SPEC (\<lambda>r. r \<longleftrightarrow> x\<in>set l)"
proof -
have "nfoldli l (\<lambda>f. \<not>f) (\<lambda>y _. RETURN (y=x)) s \<le> SPEC (\<lambda>r. r \<longleftrightarrow> s \<or> x\<in>set l)" for s
apply (induction l arbitrary: s)
subgoal by auto
subgoal for a l s
apply clarsimp
apply (rule order_trans)
apply rprems
by auto
done
from this[of False] show ?thesis by auto
qed
definition sum_impl :: "('a \<Rightarrow> 'b::comm_monoid_add nres) \<Rightarrow> 'a set \<Rightarrow> 'b nres" where
"sum_impl g S \<equiv> FOREACH S (\<lambda>x a. do { b \<leftarrow> g x; RETURN (a+b)}) 0"
lemma sum_impl_correct:
assumes [simp]: "finite S"
assumes [refine_vcg]: "\<And>x. x\<in>S \<Longrightarrow> gi x \<le> SPEC (\<lambda>r. r=g x)"
shows "sum_impl gi S \<le> SPEC (\<lambda>r. r=sum g S)"
unfolding sum_impl_def
apply (refine_vcg FOREACH_rule[where I="\<lambda>it a. a = sum g (S - it)"])
apply (auto simp: it_step_insert_iff algebra_simps)
done
end
diff --git a/thys/SIFPL/VS_OBJ.thy b/thys/SIFPL/VS_OBJ.thy
--- a/thys/SIFPL/VS_OBJ.thy
+++ b/thys/SIFPL/VS_OBJ.thy
@@ -1,1986 +1,1987 @@
(*File: VS_OBJ.thy*)
(*Authors: Lennart Beringer and Martin Hofmann, LMU Munich 2008*)
theory VS_OBJ imports VDM_OBJ PBIJ begin
subsection\<open>Non-interference\<close>
text\<open>\label{sec:ObjNI}\<close>
subsubsection\<open>Indistinguishability relations\<close>
text\<open>We have the usual two security types.\<close>
datatype TP = low | high
text\<open>Global contexts assigns security types to program
variables and field names.\<close>
consts CONTEXT :: "Var \<Rightarrow> TP"
consts GAMMA :: "Field \<Rightarrow> TP"
text\<open>Indistinguishability of values depends on a partial bijection
$\beta$.\<close>
inductive_set twiddleVal::"(PBij \<times> Val \<times> Val) set"
where
twiddleVal_Null: "(\<beta>, RVal Nullref, RVal Nullref) : twiddleVal"
| twiddleVal_Loc: "(l1,l2) : \<beta> \<Longrightarrow>
(\<beta>, RVal (Loc l1), RVal (Loc l2)) : twiddleVal"
| twiddleVal_IVal: "i1 = i2 \<Longrightarrow> (\<beta>, IVal i1, IVal i2) : twiddleVal"
text\<open>For stores, indistinguishability is as follows.\<close>
definition twiddleStore::"PBij \<Rightarrow> Store \<Rightarrow> Store \<Rightarrow> bool"
where "twiddleStore \<beta> s1 s2 =
(\<forall> x. CONTEXT x = low \<longrightarrow> (\<beta>, s1 x, s2 x) : twiddleVal)"
abbreviation twiddleStore_syntax ("_ \<approx>\<^sub>_ _" [100,100] 50)
where "s \<approx>\<^sub>\<beta> t == twiddleStore \<beta> s t"
text\<open>On objects, we require the values in low fields to be related,
and the sets of defined low fields to be equal.\<close>
definition LowDom::"((Field \<times> Val) list) \<Rightarrow> Field set"
where "LowDom F = {f . \<exists> v . lookup F f = Some v \<and> GAMMA f = low}"
definition twiddleObj::"PBij \<Rightarrow> Object \<Rightarrow> Object \<Rightarrow> bool"
where "twiddleObj \<beta> o1 o2 = ((fst o1 = fst o2) \<and>
LowDom (snd o1) = LowDom (snd o2) \<and>
(\<forall> f v w . lookup (snd o1) f = Some v \<longrightarrow>
lookup (snd o2) f = Some w \<longrightarrow>
GAMMA f = low \<longrightarrow>
(\<beta>, v, w) : twiddleVal))"
text\<open>On heaps, we require locations related by $\beta$ to contain
indistinguishable objects. Domain and codomain of the bijection
should be subsets of the domains of the heaps, of course.\<close>
definition twiddleHeap::"PBij \<Rightarrow> Heap \<Rightarrow> Heap \<Rightarrow> bool"
where "twiddleHeap \<beta> h1 h2 = (\<beta>:Pbij \<and>
Pbij_Dom \<beta> \<subseteq> Dom h1 \<and>
Pbij_Rng \<beta> \<subseteq> Dom h2 \<and>
(\<forall> l ll v w . (l,ll):\<beta> \<longrightarrow>
lookup h1 l = Some v \<longrightarrow>
lookup h2 ll = Some w \<longrightarrow>
twiddleObj \<beta> v w))"
text\<open>We also define a predicate which expresses when a state does not
contain dangling pointers.\<close>
(*The notion used in our paper:
definition noDPs::"State \<Rightarrow> bool"
where "noDPs S = (case S of (s,h) \<Rightarrow> ((\<forall> x l . s x = RVal(Loc l) \<longrightarrow> l:Dom h) \<and>
(\<forall> ll c F f l . lookup h ll = Some(c,F) \<longrightarrow>
lookup F f = Some(RVal(Loc l))
\<longrightarrow> l:Dom h)))"
*)
definition noLowDPs::"State \<Rightarrow> bool"
where "noLowDPs S = (case S of (s,h) \<Rightarrow>
((\<forall> x l . CONTEXT x = low \<longrightarrow> s x = RVal(Loc l) \<longrightarrow> l:Dom h) \<and>
(\<forall> ll c F f l . lookup h ll = Some(c,F) \<longrightarrow> GAMMA f = low \<longrightarrow>
lookup F f = Some(RVal(Loc l)) \<longrightarrow> l:Dom h)))"
text\<open>The motivation for introducing this notion stems from the
intended interpretation of the proof rule for skip, where the initial
and final states should be low equivalent. However, in the presence of
dangling pointers, indistinguishability does not hold as such a
dangling pointer is not in the expected bijection \<open>mkId\<close>. In
contrast, for the notion of indistinguishability we use (see the
following definition), reflexivity indeed holds, as proven in lemma
\<open>twiddle_mkId\<close> below. As a small improvement in comparison to
our paper, we now allow dangling pointers in high variables and high
fields since these are harmless.\<close>
definition twiddle::"PBij \<Rightarrow> State \<Rightarrow> State \<Rightarrow> bool"
where "twiddle \<beta> s t = (noLowDPs s \<and> noLowDPs t \<and>
(fst s) \<approx>\<^sub>\<beta> (fst t) \<and> twiddleHeap \<beta> (snd s) (snd t))"
abbreviation twiddle_syntax ("_ \<equiv>\<^sub>_ _" [100,100] 50)
where "s \<equiv>\<^sub>\<beta> t == twiddle \<beta> s t"
text\<open>The following properties are easily proven by unfolding the
definitions.\<close>
lemma twiddleHeap_isPbij:"twiddleHeap \<beta> h hh \<Longrightarrow> \<beta>:Pbij"
(*<*)
by (simp add:twiddleHeap_def)
(*>*)
lemma isPBij:"s \<equiv>\<^sub>\<beta> t \<Longrightarrow> \<beta>:Pbij"
(*<*)
apply (simp add: twiddle_def, clarsimp)
by (erule twiddleHeap_isPbij)
(*>*)
lemma twiddleVal_inverse:
"(\<beta>, w, v) \<in> twiddleVal \<Longrightarrow> (Pbij_inverse \<beta>, v, w) \<in> twiddleVal"
(*<*)
apply (erule twiddleVal.cases)
apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc)
apply (erule Pbij_inverseI)
apply clarsimp apply (rule twiddleVal_IVal) apply simp
done
(*>*)
lemma twiddleStore_inverse: "s \<approx>\<^sub>\<beta> t \<Longrightarrow> t \<approx>\<^sub>(Pbij_inverse \<beta>) s"
(*<*)
apply (simp add: twiddleStore_def, clarsimp)
apply (rule twiddleVal_inverse) apply fast
done
(*>*)
lemma twiddleHeap_inverse:
"twiddleHeap \<beta> s t \<Longrightarrow> twiddleHeap (Pbij_inverse \<beta>) t s"
(*<*)
apply (simp add: twiddleHeap_def, clarsimp)
apply (rule, erule Pbij_inverse_Pbij)
apply (rule, simp add: Pbij_inverse_Rng)
apply (rule, simp add: Pbij_inverse_Dom)
apply clarsimp
apply (erule_tac x=ll in allE, erule_tac x=l in allE, erule impE, erule Pbij_inverseD)
apply clarsimp
apply (simp add: twiddleObj_def, clarsimp)
apply (erule_tac x=f in allE, clarsimp)
apply (erule twiddleVal_inverse)
done
(*>*)
lemma Pbij_inverse_twiddle: "\<lbrakk>s \<equiv>\<^sub>\<beta> t\<rbrakk> \<Longrightarrow> t \<equiv>\<^sub>(Pbij_inverse \<beta>) s"
(*<*)
apply (simp add: twiddle_def, clarsimp)
apply (rule, erule twiddleStore_inverse)
apply (erule twiddleHeap_inverse)
done
(*>*)
lemma twiddleVal_betaExtend[rule_format]:
"(\<beta>,v,w):twiddleVal \<Longrightarrow> \<forall> \<gamma>. Pbij_extends \<gamma> \<beta> \<longrightarrow> (\<gamma>,v,w):twiddleVal"
(*<*)
apply (erule twiddleVal.induct)
apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (simp add: Pbij_extends_def) apply fast
apply clarsimp apply (rule twiddleVal_IVal) apply simp
done
(*>*)
lemma twiddleObj_betaExtend[rule_format]:
"\<lbrakk>twiddleObj \<beta> o1 o2; Pbij_extends \<gamma> \<beta>\<rbrakk> \<Longrightarrow> twiddleObj \<gamma> o1 o2"
(*<*)
apply (simp add: twiddleObj_def, clarsimp)
apply (erule_tac x=f in allE, erule_tac x=v in allE, clarsimp)
apply (erule twiddleVal_betaExtend) apply assumption
done
(*>*)
lemma twiddleVal_compose:
"\<lbrakk>(\<beta>, v, u) \<in> twiddleVal; (\<gamma>, u, w) \<in> twiddleVal\<rbrakk>
\<Longrightarrow> (Pbij_compose \<beta> \<gamma>, v, w) \<in> twiddleVal"
(*<*)
apply (erule twiddleVal.cases)
apply clarsimp
apply (erule twiddleVal.cases)
apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp
apply clarsimp
apply clarsimp
apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp apply (rule twiddleVal_Loc) apply (erule Pbij_composeI, assumption)
apply clarsimp
apply clarsimp
apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp
apply clarsimp apply (rule twiddleVal_IVal) apply simp
done
(*>*)
lemma twiddleHeap_compose:
"\<lbrakk> twiddleHeap \<beta> h1 h2; twiddleHeap \<gamma> h2 h3; \<beta> \<in> Pbij; \<gamma> \<in> Pbij\<rbrakk>
\<Longrightarrow> twiddleHeap (Pbij_compose \<beta> \<gamma>) h1 h3"
(*<*)
apply (simp add: twiddleHeap_def)
apply rule apply (erule Pbij_compose_Pbij) apply assumption
apply rule apply (subgoal_tac "Pbij_Dom (Pbij_compose \<beta> \<gamma>) \<subseteq> Pbij_Dom \<beta>", fast) apply (rule Pbij_compose_Dom)
apply rule apply (subgoal_tac "Pbij_Rng (Pbij_compose \<beta> \<gamma>) \<subseteq> Pbij_Rng \<gamma>", fast) apply (rule Pbij_compose_Rng)
apply (erule conjE)+
apply (rule, rule, rule)
apply (subgoal_tac "\<exists> l1 . (l,l1) : \<beta> \<and> (l1,ll):\<gamma>", erule exE, erule conjE)
prefer 2 apply (simp add: Pbij_compose_def)
apply (subgoal_tac "\<exists> x y . lookup h2 l1 = Some(x,y)", (erule exE)+)
prefer 2 apply (subgoal_tac "l1 : Dom h2", simp add: Dom_def)
apply (simp add:Pbij_compose_def Pbij_Dom_def) apply clarsimp apply fast
apply (rule, rule, rule)
apply (rule, rule, rule)
apply (erule_tac x=l in allE, erule_tac x=l1 in allE, erule impE, assumption)
apply (erule_tac x=l1 in allE, erule_tac x=ll in allE, erule impE, assumption)
apply (erule_tac x=a in allE, erule_tac x=b in allE, erule impE, assumption)
apply (erule_tac x=x in allE, erule_tac x=y in allE, erule impE, assumption)
apply (erule_tac x=aa in allE, erule_tac x=ba in allE, erule impE, assumption)
apply (erule_tac x=x in allE, erule_tac x=y in allE, erule impE, assumption)
apply (simp add: twiddleObj_def)
apply clarsimp
apply (subgoal_tac "\<exists> u . lookup y f = Some u", erule exE)
prefer 2 apply (simp add: LowDom_def) apply (rotate_tac -6) apply (erule thin_rl, erule thin_rl) apply (erule thin_rl) apply fast
apply (erule_tac x=f in allE, erule_tac x=u in allE, clarsimp)
apply (erule_tac x=f in allE, erule_tac x=v in allE, clarsimp)
apply (erule twiddleVal_compose) apply assumption
done
(*>*)
lemma twiddleStore_compose:
"\<lbrakk>s \<approx>\<^sub>\<beta> r; r\<approx>\<^sub>\<gamma> t\<rbrakk> \<Longrightarrow> s \<approx>\<^sub>(Pbij_compose \<beta> \<gamma>) t"
(*<*)
apply (simp add:twiddleStore_def)
apply clarsimp apply (erule_tac x=x in allE, clarsimp)+
apply (erule twiddleVal_compose) apply assumption
done
(*>*)
lemma twiddle_compose:
"\<lbrakk>s \<equiv>\<^sub>\<beta> r; r \<equiv>\<^sub>\<gamma> t\<rbrakk> \<Longrightarrow> s \<equiv>\<^sub>(Pbij_compose \<beta> \<gamma>) t"
(*<*)
apply (simp add: twiddle_def)
apply (erule conjE)+
apply rule apply (erule twiddleStore_compose) apply assumption
apply (rule twiddleHeap_compose) apply assumption+
apply (simp add: twiddleHeap_def)
apply (simp add: twiddleHeap_def)
done
(*>*)
lemma twiddle_mkId: "noLowDPs (s,h) \<Longrightarrow> (s,h) \<equiv>\<^sub>(mkId h) (s,h)"
(*<*)
apply (simp add: twiddle_def)
apply rule
apply (simp add: twiddleStore_def)
apply (rule, rule)
apply (case_tac "s x")
apply (rename_tac Var Ref)
apply clarsimp apply (case_tac Ref)
apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (rule mkId4) apply (simp add: noLowDPs_def)
apply clarsimp apply (simp add: twiddleVal_IVal)
(*twiddleHeap*)
apply (simp add: twiddleHeap_def)
apply (rule, rule mkId1)
apply (rule, simp add: mkId2)
apply (rule, simp add: mkId2b Dom_def)
apply clarsimp
apply (simp add: twiddleObj_def) apply (drule mkId4b) apply clarsimp
apply (case_tac v, clarsimp)
apply (rename_tac Ref)
apply (case_tac Ref, clarsimp)
apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (rule mkId4) apply (simp add: noLowDPs_def)
apply clarsimp apply (rule twiddleVal_IVal) apply simp
done(*>*)
text\<open>We call expressions (semantically) low if the following
predicate is satisfied. In particular, this means that if $e$
evaluates in $s$ (respecitvely, $t$) to some location $l$, then $l
\in Pbij\_dom(\beta)$ ($l \in Pbij\_cod(\beta)$) holds.\<close>
definition Expr_low::"Expr \<Rightarrow> bool"
where "Expr_low e = (\<forall> s t \<beta>. s \<approx>\<^sub>\<beta> t \<longrightarrow> (\<beta>, evalE e s, evalE e t):twiddleVal)"
text\<open>A similar notion is defined for boolean expressions, but the
fact that these evaluate to (meta-logical) boolean values allows us to
replace indistinguishability by equality.\<close>
definition BExpr_low::"BExpr \<Rightarrow> bool"
where "BExpr_low b = (\<forall> s t \<beta> . s \<approx>\<^sub>\<beta> t \<longrightarrow> evalB b s = evalB b t)"
subsubsection\<open>Definition and characterisation of security\<close>
text\<open>Now, the notion of security, as defined in the paper. Banerjee
and Naumann's paper~\cite{DBLP:journals/jfp/BanerjeeN05} and the
Mobius Deliverable 2.3~\cite{MobiusDeliverable2.3} contain similar
notions.\<close>
definition secure::"OBJ \<Rightarrow> bool"
where "secure c = (\<forall> s ss t tt \<beta> .
s \<equiv>\<^sub>\<beta> ss \<longrightarrow> (s,c \<Down> t) \<longrightarrow> (ss,c \<Down> tt) \<longrightarrow>
(\<exists> \<gamma> . t \<equiv>\<^sub>\<gamma> tt \<and> Pbij_extends \<gamma> \<beta>))"
(*<*)
lemma Skip1: "secure Skip"
apply (simp only: secure_def Sem_def twiddle_def)
apply (rule, rule, rule, rule, rule, rule)
apply (rule, rule, erule exE, erule exE)
apply (erule Sem_eval_cases)
apply (erule Sem_eval_cases)
apply (rule_tac x=\<beta> in exI)
apply simp
apply (rule Pbij_extends_reflexive)
done
lemma AssignAux:
"\<lbrakk> Expr_low e; s \<equiv>\<^sub>\<beta> t\<rbrakk>
\<Longrightarrow> (update (fst s) x (evalE e (fst s)), snd s) \<equiv>\<^sub>\<beta>
(update (fst t) x (evalE e (fst t)), snd t)"
apply (simp only: twiddle_def Expr_low_def)
apply rule apply (simp add: noLowDPs_def)
apply (case_tac s, clarsimp, hypsubst_thin) apply (rename_tac s h y t k l)
apply (case_tac "x=y", clarsimp) apply (simp add: update_def)
apply (erule_tac x=s in allE)
apply (erule_tac x=t in allE, clarsimp)
apply (erule_tac x=\<beta> in allE, erule impE, assumption)
apply (erule twiddleVal.cases, clarsimp) prefer 2 apply clarsimp apply clarsimp
apply (simp add: twiddleHeap_def) apply (simp add: Pbij_Dom_def) apply fast
apply (simp add: update_def)
apply rule apply (simp add: noLowDPs_def)
apply (case_tac s, clarsimp, hypsubst_thin) apply (rename_tac s h t k y l)
apply (case_tac "x=y", clarsimp) apply (simp add: update_def)
apply (erule_tac x=s in allE)
apply (erule_tac x=t in allE)
apply (erule_tac x=\<beta> in allE, clarsimp)
apply (erule twiddleVal.cases, clarsimp) prefer 2 apply clarsimp apply clarsimp
apply (simp add: twiddleHeap_def) apply (simp add: Pbij_Rng_def) apply fast
apply (simp add: update_def)
apply rule
prefer 2 apply simp
apply (unfold twiddleStore_def)
apply (rule, rule)
apply (case_tac "x=xa", clarsimp)
apply (erule_tac x="fst s" in allE)
apply (erule_tac x="fst t" in allE)
apply (erule_tac x=\<beta> in allE, clarsimp)
apply (simp add: update_def)
apply (simp add: update_def)
done
lemma Assign1:
"Expr_low e \<Longrightarrow> secure (Assign x e)"
apply (simp only: secure_def Sem_def)
apply (rule, rule, rule, rule, rule, rule)
apply (rule, rule, erule exE, erule exE)
apply (erule Sem_eval_cases)
apply (erule Sem_eval_cases)
apply (rule_tac x=\<beta> in exI)
apply simp
apply rule
apply (rule AssignAux) apply fastforce apply assumption
apply (rule Pbij_extends_reflexive)
done
lemma Comp1: "\<lbrakk>secure c1; secure c2\<rbrakk> \<Longrightarrow> secure (Comp c1 c2)"
apply (simp only: secure_def Sem_def)
apply (rule, rule, rule, rule, rule, rule)
apply (rule, rule)
apply (erule exE, erule exE)
apply (erule Sem_eval_cases, erule Sem_eval_cases)
apply (erule_tac x=s in allE, rotate_tac -1)
apply (erule_tac x=ss in allE, rotate_tac -1)
apply (erule_tac x=r in allE, rotate_tac -1)
apply (erule_tac x=ra in allE, rotate_tac -1)
apply (erule_tac x=\<beta> in allE, rotate_tac -1)
apply (erule impE, assumption)
apply (erule impE, rule, assumption)
apply (erule impE, rule, assumption)
apply (erule exE, erule conjE)
apply (erule_tac x=r in allE, rotate_tac -1)
apply (erule_tac x=ra in allE, rotate_tac -1)
apply (erule_tac x=t in allE, rotate_tac -1)
apply (erule_tac x=tt in allE, rotate_tac -1)
apply (erule_tac x=\<gamma> in allE, rotate_tac -1)
apply (erule impE, assumption)
apply (erule impE, rule, assumption)
apply (erule impE, rule, assumption)
apply (erule exE, erule conjE)
apply (rule_tac x=\<gamma>' in exI, simp)
apply (rule Pbij_extends_transitive)
apply (assumption, assumption)
done
lemma Iff1:
"\<lbrakk>BExpr_low b; secure c1; secure c2\<rbrakk> \<Longrightarrow> secure (Iff b c1 c2)"
apply (simp only: secure_def Sem_def BExpr_low_def)
apply (rule, rule, rule, rule, rule, rule)
apply (rule, rule)
apply (erule exE, erule exE)
apply (erule Sem_eval_cases, erule Sem_eval_cases)
apply (rotate_tac 2, erule thin_rl)
apply (erule_tac x=s in allE, rotate_tac -1)
apply (erule_tac x=ss in allE, rotate_tac -1)
apply (erule_tac x=t in allE, rotate_tac -1)
apply (erule_tac x=tt in allE, rotate_tac -1)
apply (erule_tac x=\<beta> in allE, clarsimp)
apply (erule impE, rule, assumption)
apply (erule impE, rule, assumption)
apply assumption
apply (erule_tac x="fst s" in allE, erule thin_rl, erule thin_rl)
apply (erule_tac x="fst ss" in allE)
apply (erule_tac x=\<beta> in allE, erule impE, simp add: twiddle_def)
apply simp
apply (erule Sem_eval_cases)
apply (erule_tac x="fst s" in allE, erule thin_rl, erule thin_rl)
apply (erule_tac x="fst ss" in allE)
apply (erule_tac x=\<beta> in allE, erule impE, simp add: twiddle_def)
apply simp
apply (erule thin_rl, erule thin_rl)
apply (erule_tac x=s in allE, rotate_tac -1)
apply (erule_tac x=ss in allE, rotate_tac -1)
apply (erule_tac x=t in allE, rotate_tac -1)
apply (erule_tac x=tt in allE, rotate_tac -1)
apply (erule_tac x=\<beta> in allE, rotate_tac -1)
apply (erule impE, assumption)
apply (erule impE, rule, assumption)
apply (erule impE, rule, assumption)
apply assumption
done
(*>*)
text\<open>The type of invariants $\Phi$ includes a component that holds a
partial bijection.\<close>
type_synonym TT = "(State \<times> State \<times> PBij) \<Rightarrow> bool"
text\<open>The operator constructing an assertion from an invariant.\<close>
definition Sec :: "TT \<Rightarrow> Assn"
where "Sec \<Phi> s t =
((\<forall> r \<beta>. s \<equiv>\<^sub>\<beta> r \<longrightarrow> \<Phi>(t,r,\<beta>)) \<and>
(\<forall> r \<beta> . \<Phi> (r,s,\<beta>) \<longrightarrow> (\<exists> \<gamma> . r \<equiv>\<^sub>\<gamma> t \<and> Pbij_extends \<gamma> \<beta>)))"
text\<open>The lemmas proving that the operator ensures security, and that
secure programs satisfy an assertion formed by the operator, are
proven in a similar way as in Section \ref{sec:BaseLineNI}.\<close>
lemma Prop1A:"\<Turnstile> c : Sec \<Phi> \<Longrightarrow> secure c"
(*<*)
apply (simp only: VDM_valid_def secure_def Sec_def)
apply (rule, rule, rule, rule)
apply (rule, rule, rule, rule)
apply (subgoal_tac "(\<forall>r \<beta>. s \<equiv>\<^sub>\<beta> r \<longrightarrow> \<Phi>(t, r, \<beta>)) \<and>
(\<forall>r \<beta>. \<Phi>(r, s, \<beta>) \<longrightarrow> (\<exists>\<gamma>. r \<equiv>\<^sub>\<gamma> t \<and> Pbij_extends \<gamma> \<beta>))")
prefer 2 apply fast
apply (erule_tac x=ss in allE, erule_tac x=tt in allE, erule impE, assumption)
apply clarsimp
done
(*>*)
lemma Prop1B:
"secure c \<Longrightarrow> \<Turnstile> c : Sec (\<lambda> (r, t, \<beta>). \<exists> s. s , c \<Down> r \<and> s \<equiv>\<^sub>\<beta> t)"
(*<*)
apply (simp only: VDM_valid_def Sec_def)
apply (rule, rule) apply (rule, rule)
apply (rule, rule, rule)
apply simp
apply (case_tac s, clarsimp)
apply (rule_tac x=ac in exI, rule_tac x=bc in exI, simp)
apply (rule, rule)
apply (rule)
apply simp
apply ((erule exE)+, (erule conjE)+)
apply (unfold secure_def)
apply (erule_tac x="(a,b)" in allE, erule_tac x=s in allE)
apply (erule_tac x=r in allE, erule_tac x=t in allE)
apply (erule_tac x=\<beta> in allE)
apply (erule impE, assumption)+
apply (erule exE, erule conjE)
apply (rule_tac x=\<gamma> in exI, simp)
(* apply (rule Pbij_extends_RCompl)*)
done
(*>*)
lemma Prop1BB:"secure c \<Longrightarrow> \<exists> \<Phi> . \<Turnstile> c : Sec \<Phi>"
(*<*)
by (drule Prop1B, fast)
(*>*)
lemma Prop1:
"secure c = (\<Turnstile> c : Sec (\<lambda> (r, t,\<beta>) . \<exists> s . (s , c \<Down> r) \<and> s \<equiv>\<^sub>\<beta> t))"
(*<*)
apply rule
apply (erule Prop1B)
apply (erule Prop1A)
done
(*>*)
subsection\<open>Derivation of proof rules\<close>
text\<open>\label{sec:ObjDerivedRules}\<close>
subsubsection\<open>Low proof rules\<close>
(*<*)
lemma SkipSem: "\<Turnstile> Skip : Sec (\<lambda> (s, t, \<beta>) . s \<equiv>\<^sub>\<beta> t)"
apply (simp only: VDM_valid_def Sec_def Sem_def)
apply (rule, rule, rule)
apply (erule exE)
apply (erule Sem_eval_cases)
apply rule
apply simp
apply (rule, rule, rule)
apply (rule_tac x=\<beta> in exI, simp)
apply (rule Pbij_extends_reflexive)
done
(*>*)
lemma SKIP: "G \<rhd> Skip : Sec (\<lambda> (s, t, \<beta>) . s \<equiv>\<^sub>\<beta> t)"
(*<*)
apply (rule VDMConseq, rule VDMSkip)
apply (simp only: Sec_def)
apply (rule, rule, rule)
apply rule
apply simp
apply (rule, rule, rule)
apply (rule_tac x=\<beta> in exI,simp)
apply (rule Pbij_extends_reflexive)
done
(*>*)
lemma ASSIGN:
"Expr_low e
\<Longrightarrow> G \<rhd> Assign x e : Sec (\<lambda> (s, t, \<beta>) .
\<exists> r . s = (update (fst r) x (evalE e (fst r)), snd r)
\<and> r \<equiv>\<^sub>\<beta> t)"
(*<*)
apply (rule VDMConseq, rule VDMAssign)
apply (simp only: Sec_def Expr_low_def)
apply (rule, rule, rule)
apply rule
apply (rule, rule, rule)
apply simp
apply (erule_tac x="fst s" in allE, erule_tac x="fst r" in allE, erule_tac x=\<beta> in allE, erule impE)
apply (simp add: twiddle_def)
apply (simp add: twiddle_def)
apply (simp add: twiddleStore_def)
apply (erule conjE)
apply (rule_tac x="fst s" in exI, simp)
apply (rule, rule, rule)
apply simp
apply (erule exE)+
apply (erule conjE)+
apply (rule_tac x=\<beta> in exI)
apply rule prefer 2 apply (rule Pbij_extends_reflexive)
apply clarsimp
apply (simp add: twiddle_def)
apply (simp add: twiddleStore_def)
apply clarsimp
apply (case_tac "xa=x")
apply (simp add: update_def noLowDPs_def)
apply (rule, clarsimp) apply (erule_tac x=ac in allE, erule_tac x=a in allE, erule_tac x=\<beta> in allE, erule impE, assumption)
apply (erule twiddleVal.cases, clarsimp)
apply (simp add: twiddleHeap_def Pbij_Dom_def, clarsimp) apply fast
apply clarsimp
apply clarsimp apply (erule_tac x=ac in allE, erule_tac x=a in allE, erule_tac x=\<beta> in allE, erule impE, assumption)
apply (erule twiddleVal.cases, clarsimp)
apply (simp add: twiddleHeap_def Pbij_Rng_def, clarsimp) apply fast
apply clarsimp
apply (simp add: update_def noLowDPs_def)
apply (rule, clarsimp) apply (erule_tac x=ac in allE, erule_tac x=a in allE, erule_tac x=\<beta> in allE, erule impE, assumption)
apply (erule twiddleVal.cases, clarsimp)
apply (simp add: twiddleHeap_def Pbij_Dom_def, clarsimp) apply fast
apply clarsimp
apply clarsimp apply (erule_tac x=ac in allE, erule_tac x=a in allE, erule_tac x=\<beta> in allE, erule impE, assumption)
apply (erule twiddleVal.cases, clarsimp)
apply (simp add: twiddleHeap_def Pbij_Rng_def, clarsimp) apply fast
apply clarsimp
done
(*>*)
(*<*)
lemma CompSem:
"\<lbrakk> \<Turnstile> c1 : Sec \<Phi>; \<Turnstile> c2 : Sec \<Psi>\<rbrakk> \<Longrightarrow>
\<Turnstile> (Comp c1 c2) : Sec (\<lambda> (s, t, \<beta>) . \<exists> r . \<Phi>(r, t, \<beta>) \<and>
(\<forall> w \<gamma>. r \<equiv>\<^sub>\<gamma> w \<longrightarrow> \<Psi>(s, w, \<gamma>)))"
apply (simp only: VDM_valid_def Sec_def Sem_def)
apply (rule, rule, rule)
apply (erule exE)
apply (erule Sem_eval_cases)
apply (erule_tac x=s in allE, rotate_tac -1)
apply (erule_tac x=r in allE, rotate_tac -1)
apply (erule impE, rule, assumption)
apply (erule_tac x=r in allE, rotate_tac -1)
apply (erule_tac x=t in allE, rotate_tac -1)
apply (erule impE, rule, assumption)
apply (erule thin_rl, erule thin_rl)
apply (erule conjE)+
apply rule
apply (rule, rule, rule)
apply (erule_tac x=ra in allE, rotate_tac -1)
apply (erule_tac x=\<beta> in allE, rotate_tac -1)
apply (erule impE, assumption)
apply simp
apply (case_tac r, clarsimp)
apply (rule_tac x=ad in exI, rule_tac x=bd in exI, simp)
apply (rule, rule, rule)
apply simp
apply (erule exE)+ apply (erule conjE)+
apply (rotate_tac 2, erule_tac x=a in allE)
apply (rotate_tac -1, erule_tac x=b in allE)
apply (rotate_tac -1, erule_tac x=\<beta> in allE)
apply (rotate_tac -1, erule impE, assumption)
apply (erule exE, erule conjE)
apply (case_tac r, clarsimp)
apply (rotate_tac 3, erule_tac x=aaa in allE)
apply (rotate_tac -1, erule_tac x=baa in allE)
apply (rotate_tac -1, erule_tac x=\<gamma> in allE)
apply (erule impE, assumption)
apply (rotate_tac 5, erule_tac x=ac in allE)
apply (rotate_tac -1, erule_tac x=bc in allE)
apply (rotate_tac -1, erule_tac x=\<gamma> in allE)
apply (erule impE, assumption)
apply (erule exE)
apply (erule conjE)
apply (rule_tac x=\<gamma>' in exI, rule, assumption)
apply (erule Pbij_extends_transitive) apply assumption
done
(*>*)
lemma COMP:
"\<lbrakk> G \<rhd> c1 : Sec \<Phi>; G \<rhd> c2 : Sec \<Psi>\<rbrakk>
\<Longrightarrow> G \<rhd> (Comp c1 c2) : Sec (\<lambda> (s, t, \<beta>) .
\<exists> r . \<Phi>(r, t, \<beta>) \<and> (\<forall> w \<gamma>. r \<equiv>\<^sub>\<gamma> w \<longrightarrow> \<Psi>(s, w, \<gamma>)))"
(*<*)
apply (rule VDMConseq, rule VDMComp)
apply (assumption, assumption)
apply (simp only: Sec_def)
apply (rule, rule, rule)
apply (erule exE)
apply (erule conjE)+
apply (erule thin_rl, erule thin_rl)
apply rule
apply (rule, rule, rule)
apply simp
apply (case_tac ra, clarsimp)
apply (erule_tac x=ad in allE, rotate_tac -1)
apply (erule_tac x=bd in allE, rotate_tac -1)
apply (erule_tac x=\<beta> in allE, rotate_tac -1)
apply (erule impE, assumption)
apply (rule_tac x=ab in exI, rule_tac x=bb in exI, simp)
apply (rule, rule, rule)
apply simp
apply (erule exE)+ apply (erule conjE)+
apply (rotate_tac 1, erule_tac x=a in allE)
apply (rotate_tac -1, erule_tac x=b in allE)
apply (rotate_tac -1, erule_tac x=\<beta> in allE)
apply (rotate_tac -1, erule impE, assumption)
apply (erule exE, erule conjE)
apply (case_tac r, clarsimp)
apply (rotate_tac 3, erule_tac x=aaa in allE)
apply (rotate_tac -1, erule_tac x=baa in allE)
apply (rotate_tac -1, erule_tac x=\<gamma> in allE)
apply (erule impE, assumption)
apply (rotate_tac 4, erule_tac x=ac in allE)
apply (rotate_tac -1, erule_tac x=bc in allE)
apply (rotate_tac -1, erule_tac x=\<gamma> in allE)
apply (erule impE, assumption)
apply (erule exE)
apply (erule conjE)
apply (rule_tac x=\<gamma>' in exI, rule, assumption)
apply (erule Pbij_extends_transitive) apply assumption
done(*>*)
(*<*)
lemma IffSem:
"\<lbrakk> BExpr_low b; \<Turnstile> c1 : Sec \<Phi>; \<Turnstile> c2 : Sec \<Psi>\<rbrakk> \<Longrightarrow>
\<Turnstile> (Iff b c1 c2) : (Sec (\<lambda> (s, t, \<beta>) .
(evalB b (fst t) \<longrightarrow> \<Phi>(s,t, \<beta>)) \<and>
((\<not> evalB b (fst t)) \<longrightarrow> \<Psi>(s,t,\<beta>))))"
apply (simp only: VDM_valid_def Sec_def Sem_def BExpr_low_def)
apply (rule, rule, rule)
apply (erule exE)
apply (erule Sem_eval_cases)
apply rule
apply (rule, rule, rule)
apply simp
apply (erule_tac x="fst s" in allE, rotate_tac -1)
apply (erule_tac x="fst r" in allE, rotate_tac -1)
apply (erule impE, rule_tac x=\<beta> in exI, simp add: twiddle_def)
apply (case_tac s, clarsimp)
apply (erule_tac x=ac in allE, rotate_tac -1)
apply (erule_tac x=bc in allE, rotate_tac -1)
apply (erule_tac x=aa in allE, rotate_tac -1)
apply (erule_tac x=ba in allE, rotate_tac -1)
apply (erule impE, rule, assumption)
(* apply (simp add: twiddle_def)*)
apply clarsimp
apply (rule, rule, rule)
apply simp
apply (case_tac s, clarsimp)
apply (rotate_tac 1)
apply (erule_tac x=ac in allE, rotate_tac -1)
apply (erule_tac x=bc in allE, rotate_tac -1)
apply (erule_tac x=aa in allE, rotate_tac -1)
apply (erule_tac x=ba in allE, rotate_tac -1)
apply (erule impE, rule, assumption)
apply clarsimp
apply rule
apply (rule, rule, rule, rule)
apply simp
apply (case_tac s, clarsimp)
apply (erule_tac x=ac in allE, rotate_tac -1)
apply (erule_tac x=ab in allE, rotate_tac -1)
apply (erule impE) apply(rule_tac x=\<beta> in exI, simp add: twiddle_def)
apply clarsimp apply (erule thin_rl, erule_tac x=ac in allE)
apply (erule_tac x=bc in allE, erule_tac x=aa in allE, erule_tac x=ba in allE, erule impE)
apply (rule, assumption)
apply clarsimp
apply clarsimp apply (erule thin_rl, erule thin_rl, erule_tac x=a in allE)
apply (erule_tac x=ba in allE, erule_tac x=aa in allE, erule_tac x=baa in allE, erule impE)
apply (rule, assumption)
apply clarsimp
done
(*>*)
lemma IFF:
"\<lbrakk> BExpr_low b; G \<rhd> c1 : Sec \<Phi>; G \<rhd> c2 : Sec \<Psi>\<rbrakk>
\<Longrightarrow> G \<rhd> (Iff b c1 c2) : Sec (\<lambda> (s,t,\<beta>) .
(evalB b (fst t) \<longrightarrow> \<Phi>(s,t,\<beta>)) \<and>
((\<not> evalB b (fst t)) \<longrightarrow> \<Psi>(s,t,\<beta>)))"
(*<*)
apply (rule VDMConseq, rule VDMIff)
apply (assumption, assumption)
apply (simp only: Sec_def BExpr_low_def)
apply (rule, rule, rule)
apply (rule, rule, rule, rule)
apply (erule_tac x="fst s" in allE, rotate_tac -1)
apply (erule_tac x="fst r" in allE, rotate_tac -1)
apply (erule_tac x=\<beta> in allE, rotate_tac -1)
apply (simp add: twiddle_def)
apply clarsimp
apply (rule, rule, rule)
apply (case_tac "evalB b (fst s)")
apply clarsimp
apply clarsimp
done
(*>*)
(*<*)
lemma noLowDPs_NEW:
"noLowDPs (s,h) \<Longrightarrow> noLowDPs (update s x (RVal (Loc l)), (l, C, []) # h)"
apply (simp add: noLowDPs_def update_def, clarsimp)
apply (rule, clarsimp)
apply (rule, clarsimp) apply (simp add: Dom_def)
apply clarsimp apply (simp add: Dom_def)
apply clarsimp apply (erule_tac x=ll in allE, clarsimp) apply (simp add: Dom_def)
done
(*>*)
lemma NEW:
"CONTEXT x = low
\<Longrightarrow> G \<rhd> (New x C) : Sec (\<lambda> (s,t,\<beta>) .
\<exists> l r . l \<notin> Dom (snd r) \<and> r \<equiv>\<^sub>\<beta> t \<and>
s = (update (fst r) x (RVal (Loc l)),
(l,(C,[])) # (snd r)))"
(*<*)
apply (rule VDMConseq, rule VDMNew)
apply (simp only: Sec_def)
apply (rule, rule, rule)
apply (erule exE, (erule conjE)+)
apply rule
(*First part of Sec*)
apply (rule, rule, rule)
apply simp
apply (rule, rule, assumption)
apply (rule_tac x="fst s" in exI, simp)
(*Second part of Sec*)
apply (rule, rule, rule)
apply simp
apply (erule exE)+
apply (erule conjE)+
apply (rule_tac x="insert (la,l) \<beta>" in exI)
apply rule
(*Show twiddle*)
apply (frule isPBij)
apply (simp add: twiddle_def, clarsimp)
apply rule apply (erule noLowDPs_NEW)
apply rule apply (erule noLowDPs_NEW)
(*twiddleStore*)
apply (simp add: twiddleStore_def)
apply (rule, rule)
apply (case_tac "x=xa")
apply (simp add: update_def) apply clarsimp
apply (rule twiddleVal_Loc) apply simp
apply (simp add: update_def) apply clarsimp
apply (erule_tac x=xa in allE, clarsimp)
apply (rule twiddleVal_betaExtend, assumption) apply (simp add: Pbij_extends_def) apply fast
(*twiddleHeap*)
apply (simp add: twiddleHeap_def) apply clarsimp
apply rule apply (erule Pbij_insert) apply fast apply fast
apply (rule, simp add: Pbij_Dom_def Dom_def) apply fast
apply (rule, simp add: Pbij_Rng_def Dom_def) apply fast
apply (rule, rule)
apply clarsimp
apply (rule, clarsimp)
apply (rule, clarsimp) apply (simp add: twiddleObj_def)
apply clarsimp apply (simp add: twiddleObj_def)
apply clarsimp apply (simp add: Pbij_Dom_def) apply fast
apply clarsimp
apply (rule, clarsimp) apply (simp add: Pbij_Rng_def) apply fast
apply clarsimp
apply (erule_tac x=lb in allE, erule_tac x=ll in allE, clarsimp)
apply (rule twiddleObj_betaExtend) apply assumption+ apply (simp add: Pbij_extends_def) apply fast
(*Pbij_extends*)
apply (simp add: Pbij_extends_def) apply fast
done
(*>*)
lemma GET:
"\<lbrakk> CONTEXT y = low; GAMMA f = low\<rbrakk>
\<Longrightarrow> G \<rhd> Get x y f : Sec (\<lambda> (s,t,\<beta>) .
\<exists> r l C Flds v. (fst r) y = RVal(Loc l) \<and>
lookup (snd r) l = Some(C,Flds) \<and>
lookup Flds f = Some v \<and> r \<equiv>\<^sub>\<beta> t \<and>
s = (update (fst r) x v, snd r))"
(*<*)
apply (rule VDMConseq, rule VDMGet)
apply (simp only: Sec_def)
apply (rule, rule, rule)
apply (erule exE)+
apply (erule conjE)+
apply rule
apply (rule, rule, rule)
apply simp
apply (rule, rule, rule)
apply (rule, assumption)
apply (rule, rule, rule, assumption)
apply (rule, rule, assumption)
apply simp
apply (rule, rule, rule)
apply simp
apply (erule exE)+
apply (erule conjE)+
apply (erule exE)+
apply (erule conjE)+
apply (erule exE)+
apply (erule conjE)+
apply (rule_tac x=\<beta> in exI, rule)
prefer 2 apply (rule Pbij_extends_reflexive)
apply (simp add: twiddle_def, (erule conjE)+)
apply (rule, simp add: noLowDPs_def update_def, clarsimp)
apply (rule, simp add: noLowDPs_def update_def)
apply (simp add: twiddleStore_def) apply clarsimp
apply (case_tac "x=xa", clarsimp) prefer 2 apply (simp add: update_def)
apply (simp add:update_def) apply (simp add: twiddleHeap_def) apply clarsimp
apply (erule_tac x=y in allE, clarsimp)
apply (erule twiddleVal.cases)
apply clarsimp
prefer 2 apply clarsimp
apply clarsimp
apply (erule_tac x=l1 in allE)
apply (erule_tac x=l2 in allE, clarsimp)
apply (simp add: twiddleObj_def)
done
(*>*)
lemma PUT:
"\<lbrakk> CONTEXT x = low; GAMMA f = low; Expr_low e\<rbrakk>
\<Longrightarrow> G \<rhd> Put x f e: Sec (\<lambda> (s,t,\<beta>) .
\<exists> r l C Flds. (fst r) x = RVal(Loc l) \<and> r \<equiv>\<^sub>\<beta> t \<and>
lookup (snd r) l = Some(C,Flds) \<and>
s = (fst r,
(l,(C,(f,evalE e (fst r)) # Flds)) # (snd r)))"
(*<*)
apply (rule VDMConseq, rule VDMPut)
apply (simp only: Sec_def Expr_low_def)
apply (rule, rule, rule)
apply (erule exE)+
apply (erule conjE)+
apply rule
apply (rule, rule, rule)
apply simp
apply (rule, rule, rule)
apply simp
apply (erule exE)+
apply (erule conjE)+
apply (erule exE)+
apply (erule conjE)+
apply (rule_tac x=\<beta> in exI, rule)
prefer 2 apply (rule Pbij_extends_reflexive)
apply (simp add: twiddle_def)
apply (simp add: twiddleStore_def)
apply (simp add: twiddleHeap_def)
apply clarsimp
apply (rule, rotate_tac -1, erule thin_rl, simp add: noLowDPs_def)
apply (rule, clarsimp) apply (subgoal_tac "lb : Dom bc", simp add: Dom_def)
apply (erule_tac x=xa in allE, erule impE, assumption,
erule_tac x=lb in allE, erule mp, assumption)
apply clarsimp apply (rule, clarsimp)
apply (rule, clarsimp)
apply (erule_tac x=ac in allE, erule_tac x=a in allE, erule_tac x=\<beta> in allE, erule impE, assumption)
apply (erule twiddleVal.cases, clarsimp)
apply clarsimp apply (subgoal_tac "la : Dom bc", simp add: Dom_def)
apply (simp add: Pbij_Dom_def) apply fast
apply clarsimp
apply clarsimp apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=fa in allE, clarsimp) apply (simp add: Dom_def)
apply clarsimp apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=fa in allE, clarsimp) apply (simp add: Dom_def)
apply (rule, rotate_tac -1, erule thin_rl, simp add: noLowDPs_def)
apply (rule, clarsimp) apply (subgoal_tac "lb : Dom b", simp add: Dom_def)
apply (erule_tac x=xa in allE, erule impE, assumption,
erule_tac x=lb in allE, erule mp, assumption)
apply clarsimp apply (rule, clarsimp)
apply (rule, clarsimp)
apply (erule_tac x=ac in allE, erule_tac x=a in allE, erule_tac x=\<beta> in allE, erule impE, assumption)
apply (erule twiddleVal.cases, clarsimp)
apply clarsimp apply (subgoal_tac "l : Dom b", simp add: Dom_def)
apply (simp add: Pbij_Rng_def) apply fast
apply clarsimp
apply clarsimp apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=fa in allE, clarsimp) apply (simp add: Dom_def)
apply clarsimp apply (erule_tac x=ll in allE, clarsimp)
apply (erule_tac x=fa in allE, clarsimp) apply (simp add: Dom_def)
apply (rule, rotate_tac -1, erule thin_rl, simp add: Dom_def) apply fast
apply (rule, rotate_tac -1, erule thin_rl, simp add: Dom_def) apply fast
apply clarsimp
apply (rule, rule, rule)
apply rule apply clarsimp
apply (erule_tac x=lb in allE, rotate_tac -1)
apply (erule_tac x=ll in allE, rotate_tac -1, clarsimp)
apply (erule_tac x=ac in allE, erule_tac x=a in allE, erule_tac x=\<beta> in allE, erule impE, assumption)
apply (simp add: twiddleObj_def, clarsimp)
apply (simp add: LowDom_def) apply (rotate_tac -1, erule thin_rl)
apply rule apply (rule, clarsimp) apply fast
apply (rule, clarsimp) apply fast
apply clarsimp apply (erule_tac x=x in allE, clarsimp)
apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp apply (simp add: Pbij_def) apply fast
apply clarsimp
apply clarsimp apply (erule_tac x=x in allE, clarsimp)
apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp apply (simp add: Pbij_def)
apply clarsimp
done
(*>*)
text\<open>Again, we define a fixed point operator over invariants.\<close>
definition FIX::"(TT \<Rightarrow> TT) \<Rightarrow> TT"
where "FIX \<phi> = (\<lambda> (s,t,\<beta>).
\<forall> \<Phi> . (\<forall> ss tt \<gamma>. \<phi> \<Phi> (ss, tt,\<gamma>) \<longrightarrow> \<Phi> (ss, tt,\<gamma>)) \<longrightarrow> \<Phi> (s, t,\<beta>))"
definition Monotone::"(TT \<Rightarrow> TT) \<Rightarrow> bool"
where "Monotone \<phi> =
(\<forall> \<Phi> \<Psi> . (\<forall> s t \<beta>. \<Phi>(s,t,\<beta>) \<longrightarrow> \<Psi>(s,t,\<beta>)) \<longrightarrow>
(\<forall> s t \<beta>. \<phi> \<Phi> (s,t,\<beta>) \<longrightarrow> \<phi> \<Psi> (s,t,\<beta>)))"
(*<*)
lemma Fix2: "\<lbrakk>Monotone \<phi>; \<phi> (FIX \<phi>) (s, t,\<beta>)\<rbrakk> \<Longrightarrow> FIX \<phi> (s,t,\<beta>)"
apply (unfold FIX_def)
apply (rule, rule)
apply (rule, rule)
apply (subgoal_tac "\<phi> \<Phi> (s,t,\<beta>)") apply fast
apply (subgoal_tac "\<forall> r u \<gamma>. FIX \<phi> (r,u,\<gamma>) \<longrightarrow> \<Phi>(r,u,\<gamma>)")
prefer 2 apply (erule thin_rl) apply (simp add: FIX_def) apply clarsimp
apply (erule_tac x=\<Phi> in allE, simp)
apply (unfold Monotone_def)
apply (erule_tac x="FIX \<phi>" in allE, erule_tac x=\<Phi> in allE)
apply (erule impE) apply assumption
apply (unfold FIX_def) apply fast
done
lemma Fix1: "\<lbrakk>Monotone \<phi>; FIX \<phi> (s,t,\<beta>)\<rbrakk> \<Longrightarrow> \<phi> (FIX \<phi>) (s,t,\<beta>)"
apply (simp add: FIX_def)
apply (erule_tac x="\<phi>(FIX \<phi>)" in allE)
apply (erule impE)
prefer 2 apply (simp add: FIX_def)
apply (subgoal_tac "\<forall> r u \<gamma>. \<phi> (FIX \<phi>) (r,u,\<gamma>) \<longrightarrow> FIX \<phi> (r,u,\<gamma>)")
prefer 2 apply clarsimp apply (erule Fix2) apply assumption
apply (unfold Monotone_def)
apply (erule_tac x="\<phi> (FIX \<phi>)" in allE, erule_tac x="FIX \<phi>" in allE, erule impE) apply assumption
apply simp
done
(*>*)
text\<open>For monotone transformers, the construction indeed
yields a fixed point.\<close>
lemma Fix_lemma:"Monotone \<phi> \<Longrightarrow> \<phi> (FIX \<phi>) = FIX \<phi>"
(*<*)
apply (rule ext, rule iffI)
apply clarsimp apply (erule Fix2) apply assumption
apply clarsimp apply (erule Fix1) apply assumption
done
(*>*)
text\<open>The operator used in the while rule is defined by\<close>
definition PhiWhileOp::"BExpr \<Rightarrow> TT \<Rightarrow> TT \<Rightarrow> TT"
where "PhiWhileOp b \<Phi> = (\<lambda> \<Psi> . (\<lambda> (s, t, \<beta>).
(evalB b (fst t) \<longrightarrow>
(\<exists> r. \<Phi> (r, t, \<beta>) \<and> (\<forall> w \<gamma>. r \<equiv>\<^sub>\<gamma> w \<longrightarrow> \<Psi>(s, w, \<gamma>)))) \<and>
(\<not> evalB b (fst t) \<longrightarrow> s\<equiv>\<^sub>\<beta> t)))"
text\<open>and is monotone:\<close>
lemma PhiWhileOp_Monotone: "Monotone (PhiWhileOp b \<Phi>)"
(*<*)
apply (simp add: PhiWhileOp_def Monotone_def) apply clarsimp
apply (rule_tac x=ab in exI, rule_tac x=bb in exI, simp)
done
(*>*)
text\<open>Hence, we can define its fixed point:\<close>
definition PhiWhile::"BExpr \<Rightarrow> TT \<Rightarrow> TT"
where "PhiWhile b \<Phi> = FIX (PhiWhileOp b \<Phi>)"
text\<open>The while rule may now be given as follows:\<close>
lemma WHILE:
"\<lbrakk> BExpr_low b; G \<rhd> c : (Sec \<Phi>)\<rbrakk>
\<Longrightarrow> G \<rhd> (While b c) : (Sec (PhiWhile b \<Phi>))"
(*<*)
apply (rule VDMConseq)
apply (rule VDMWhile)
prefer 4 apply (subgoal_tac "\<forall>s t. (Sec (PhiWhileOp b \<Phi> (PhiWhile b \<Phi>))) s t \<and>
\<not> evalB b (fst t) \<longrightarrow> Sec (PhiWhile b \<Phi>) s t", assumption)
prefer 2 apply assumption
apply clarsimp apply (subgoal_tac "PhiWhileOp b \<Phi> (PhiWhile b \<Phi>)= PhiWhile b \<Phi>", clarsimp)
apply (simp add: PhiWhile_def) apply (rule Fix_lemma) apply (rule PhiWhileOp_Monotone)
apply clarsimp apply (simp add: Sec_def)
apply (rule, clarsimp) apply (simp add: PhiWhileOp_def BExpr_low_def)
apply clarsimp apply (simp add: twiddle_def) apply (erule_tac x=a in allE, erule_tac x=aa in allE, clarsimp)
apply clarsimp apply (simp add: PhiWhileOp_def)
apply (rule_tac x=\<beta> in exI, simp) apply (rule Pbij_extends_reflexive)
apply clarsimp apply (simp add: Sec_def)
apply rule
prefer 2 apply clarsimp
apply (subgoal_tac "\<exists> r1 r2 \<gamma> . \<Phi> ((r1,r2), (a,ba),\<gamma>) \<and> Pbij_extends \<gamma> \<beta> \<and>
(\<forall> w1 w2 \<gamma> . (r1,r2) \<equiv>\<^sub>\<gamma> (w1,w2) \<longrightarrow>
(PhiWhile b \<Phi>) ((ac,bc), (w1,w2),\<gamma>))")
prefer 2 apply (simp add: PhiWhileOp_def)
apply (erule exE)+ apply (erule conjE)+ apply (rule_tac x=ad in exI, rule_tac x=bd in exI, rule)
apply (rule, assumption)
apply (rule, rule Pbij_extends_reflexive)
apply assumption
apply (erule exE)+ apply (erule conjE)+
apply (rotate_tac 4, erule_tac x=r1 in allE,
rotate_tac -1, erule_tac x=r2 in allE,
rotate_tac -1, erule_tac x=\<gamma> in allE, erule impE, assumption)
apply (erule exE, erule conjE)
apply (rotate_tac 4, erule_tac x=aa in allE,
rotate_tac -1, erule_tac x=baa in allE,
rotate_tac -1, erule_tac x=\<gamma>' in allE, erule impE, assumption)
apply (rotate_tac 8, erule_tac x=ac in allE,
rotate_tac -1, erule_tac x=bc in allE,
rotate_tac -1, erule_tac x=\<gamma>' in allE, erule impE)
apply (subgoal_tac "PhiWhileOp b \<Phi> (PhiWhile b \<Phi>) = (PhiWhile b \<Phi>)", clarsimp)
apply (simp add: PhiWhile_def)
apply (rule Fix_lemma) apply (rule PhiWhileOp_Monotone)
apply (erule exE, erule conjE)
apply (rule, rule, assumption) apply (erule Pbij_extends_transitive)
apply (erule Pbij_extends_transitive) apply assumption
apply clarsimp
apply (simp only: BExpr_low_def)
apply (erule_tac x=a in allE, rotate_tac -1)
apply (erule_tac x=ac in allE, rotate_tac -1)
apply (erule_tac x=\<beta> in allE, rotate_tac -1)
apply (erule impE, simp add: twiddle_def)
apply (simp (no_asm_simp) add: PhiWhileOp_def)
apply clarsimp
apply (erule thin_rl)
apply (erule_tac x=ac in allE, rotate_tac -1)
apply (erule_tac x=bc in allE, rotate_tac -1)
apply (erule_tac x=\<beta> in allE, rotate_tac -1)
apply (erule impE, assumption)
apply (rule_tac x=aa in exI, rule_tac x=baa in exI, rule, assumption)
apply clarsimp
apply (rotate_tac 2)
apply (erule_tac x=ad in allE, rotate_tac -1)
apply (erule_tac x=bd in allE, rotate_tac -1)
apply (erule_tac x=\<gamma> in allE, rotate_tac -1)
apply (erule impE, assumption)
apply (subgoal_tac "PhiWhileOp b \<Phi> (PhiWhile b \<Phi>) = PhiWhile b \<Phi>", clarsimp)
apply (simp add: PhiWhile_def)
apply (rule Fix_lemma) apply (rule PhiWhileOp_Monotone)
done
(*>*)
text\<open>Operator $\mathit{PhiWhile}$ is itself monotone in $\Phi$:\<close>
lemma PhiWhileMonotone: "Monotone (\<lambda> \<Phi> . PhiWhile b \<Phi>)"
(*<*)
apply (simp add: Monotone_def) apply clarsimp
apply (simp add: PhiWhile_def)
apply (simp add: FIX_def) apply clarsimp
apply (erule_tac x=\<Phi>' in allE, erule mp)
apply (clarsimp) apply (erule_tac x=a in allE, erule_tac x=ba in allE,
erule_tac x=aa in allE, erule_tac x=baa in allE,
erule_tac x=\<gamma> in allE, erule mp)
apply (simp add: PhiWhileOp_def) apply clarsimp
apply (rule_tac x=ab in exI, rule_tac x=bb in exI, simp)
done
(*>*)
text\<open>We now give an alternative formulation using an inductive
relation equivalent to FIX. First, the definition of the variant.\<close>
inductive_set var::"(BExpr \<times> TT \<times> PBij \<times> State \<times> State) set"
where
varFalse: "\<lbrakk>\<not> evalB b (fst t); s \<equiv>\<^sub>\<beta> t\<rbrakk> \<Longrightarrow> (b,\<Phi>,\<beta>,s,t):var"
| varTrue:
"\<lbrakk> evalB b (fst t); \<Phi>(r,t,\<beta>); \<forall> w \<gamma>. r \<equiv>\<^sub>\<gamma> w \<longrightarrow> (b,\<Phi>,\<gamma>,s,w): var\<rbrakk>
\<Longrightarrow> (b,\<Phi>,\<beta>,s,t):var"
text\<open>The equivalence of the invariant with the fixed point
construction.\<close>
(*<*)
lemma varFIX: "(b,\<Phi>,\<beta>,s,t):var \<Longrightarrow> PhiWhile b \<Phi> (s,t,\<beta>)"
apply (erule var.induct)
apply (simp add: PhiWhile_def)
apply (subgoal_tac "PhiWhileOp b \<Phi> (FIX (PhiWhileOp b \<Phi>)) (s,t,\<beta>)")
apply (subgoal_tac "PhiWhileOp b \<Phi> (FIX (PhiWhileOp b \<Phi>)) = FIX (PhiWhileOp b \<Phi>)", clarsimp)
apply (rule Fix_lemma) apply (rule PhiWhileOp_Monotone)
apply (simp add: PhiWhileOp_def)
apply (simp (no_asm_simp) add: PhiWhile_def)
apply (subgoal_tac "PhiWhileOp b \<Phi> (FIX (PhiWhileOp b \<Phi>)) (s,t,\<beta>)")
apply (subgoal_tac "PhiWhileOp b \<Phi> (FIX (PhiWhileOp b \<Phi>)) = FIX (PhiWhileOp b \<Phi>)", clarsimp)
apply (rule Fix_lemma) apply (rule PhiWhileOp_Monotone)
apply (simp add: PhiWhileOp_def)
apply (case_tac r, clarsimp)
apply (rule_tac x=ac in exI, rule_tac x=baa in exI, rule) apply assumption
apply clarsimp
apply (erule_tac x=aa in allE)
apply (erule_tac x=bb in allE)
apply (erule_tac x=\<gamma> in allE, clarsimp)
apply (simp add: PhiWhile_def)
apply (simp add: PhiWhileOp_def)
done
lemma FIXvar: "PhiWhile b \<Phi> (s,t,\<beta>) \<Longrightarrow> (b,\<Phi>,\<beta>,s,t):var"
apply (simp add: PhiWhile_def)
apply (subgoal_tac "PhiWhileOp b \<Phi> (FIX (PhiWhileOp b \<Phi>)) (s, t,\<beta>)")
prefer 2
apply (subgoal_tac "PhiWhileOp b \<Phi> (FIX (PhiWhileOp b \<Phi>)) = FIX (PhiWhileOp b \<Phi>)", clarsimp)
apply (rule Fix_lemma) apply (rule PhiWhileOp_Monotone)
apply (erule thin_rl, simp add: PhiWhileOp_def) apply clarsimp
apply (case_tac "evalB b (fst t)")
prefer 2 apply clarsimp apply (rule varFalse) apply assumption+
apply clarsimp apply (rule varTrue) apply assumption apply assumption
apply (rule, rule, rule)
apply (case_tac w, clarsimp)
apply (erule_tac x=aaa in allE, erule_tac x=baa in allE, erule_tac x=\<gamma> in allE, clarsimp)
apply (unfold FIX_def) apply clarify
apply (erule_tac x="\<lambda> (x,y,\<beta>) . (b,\<Phi>,\<beta>,x,y):var" in allE, erule impE) prefer 2 apply simp
apply clarsimp
apply (case_tac "evalB b ab")
prefer 2 apply clarsimp apply (rule varFalse) apply simp apply assumption
apply clarsimp apply (rule varTrue) apply simp apply assumption apply simp
done
(*>*)
lemma varFIXvar: "(PhiWhile b \<Phi> (s,t,\<beta>)) = ((b,\<Phi>,\<beta>,s,t):var)"
(*<*)
apply rule
apply (erule FIXvar)
apply (erule varFIX)
done
(*>*)
(*<*)
lemma FIXvarFIX':
"(PhiWhile b \<Phi>) = (\<lambda> (s,t,\<beta>) . (b,\<Phi>,\<beta>,s,t):var)"
apply (rule ext, rule iffI)
apply (case_tac x, clarsimp) apply (erule FIXvar)
apply (case_tac x, clarsimp) apply (simp add: varFIXvar)
done
lemma FIXvarFIX: "(PhiWhile b) = (\<lambda> \<Phi> . (\<lambda> (s,t,\<beta>) . (b,\<Phi>,\<beta>,s,t):var))"
apply rule apply (rule FIXvarFIX')
done
(*>*)
text\<open>Using this equivalence we can derive the while rule given in our
paper from \<open>WHILE\<close>.\<close>
lemma WHILE_IND:
"\<lbrakk> BExpr_low b; G \<rhd> c : (Sec \<Phi>)\<rbrakk>
\<Longrightarrow> G \<rhd> While b c: (Sec (\<lambda>(s,t,\<gamma>). (b,\<Phi>,\<gamma>,s,t):var))"
(*<*)
apply (rule VDMConseq)
apply (erule WHILE) apply assumption
apply clarsimp
apply (simp add: FIXvarFIX')
done
(*>*)
text\<open>We can also show the following property.\<close>
(*<*)
lemma varMonotoneAux[rule_format]:
"(b, \<Phi>, \<beta>, s, t) \<in> var \<Longrightarrow>
(\<forall>s t \<gamma>. \<Phi> (s, t, \<gamma>) \<longrightarrow> \<Psi> (s, t, \<gamma>)) \<longrightarrow>
(b, \<Psi>, \<beta>, s, t) \<in> var"
apply (erule var.induct)
apply clarsimp apply (rule varFalse) apply simp apply assumption
apply clarsimp apply (rule varTrue) apply simp apply fast apply clarsimp
done
(*>*)
lemma var_Monotone:
"Monotone (\<lambda> \<Phi> . (\<lambda> (s,t,\<beta>) .(b,\<Phi>,\<beta>,s,t):var))"
(*<*)
apply (simp add: Monotone_def)
apply clarsimp
apply (rule varMonotoneAux) apply assumption apply clarsimp
done
(*>*)
(*<*)
lemma varMonotone_byFIX: "Monotone (\<lambda> \<Phi> . (\<lambda> (s,t,\<beta>) .(b,\<Phi>,\<beta>,s,t):var))"
apply (subgoal_tac "Monotone (\<lambda> \<Phi> . PhiWhile b \<Phi>)")
apply (simp add: FIXvarFIX)
apply (rule PhiWhileMonotone)
done
(*>*)
text\<open>The call rule is formulated for an arbitrary fixed point of
a monotone transformer.\<close>
lemma CALL:
"\<lbrakk>({Sec (FIX \<phi>)} \<union> X) \<rhd> body : Sec (\<phi> (FIX \<phi>)); Monotone \<phi>\<rbrakk>
\<Longrightarrow> X \<rhd> Call : Sec (FIX \<phi>)"
(*<*)
apply (rule VDMCall)
apply (subgoal_tac "\<phi> (FIX \<phi>) = FIX \<phi>", clarsimp)
apply (erule Fix_lemma)
done
(*>*)
subsubsection\<open>High proof rules\<close>
definition HighSec::Assn
where "HighSec = (\<lambda> s t . noLowDPs s \<longrightarrow> s \<equiv>\<^sub>(mkId (snd s)) t)"
lemma CAST[rule_format]:
"G \<rhd> c : HighSec \<Longrightarrow> G \<rhd> c : Sec (\<lambda> (s, t, \<beta>) . s \<equiv>\<^sub>\<beta> t)"
(*<*)
apply (erule VDMConseq) apply (simp add:Sec_def) apply clarsimp
apply (rule, clarsimp)
apply (simp add: HighSec_def)
apply (simp add: twiddle_def, clarsimp)
apply (rule, simp add: twiddleStore_def, clarsimp)
apply (erule_tac x=x in allE, clarsimp)
apply (erule_tac x=x in allE, clarsimp)
apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp apply (drule mkId4b, clarsimp) apply (erule twiddleVal_Loc)
apply clarsimp
apply clarsimp
apply (simp add: twiddleHeap_def, clarsimp)
apply (rule, simp add: mkId2 mkId2b)
apply (simp add: mkId2 mkId2b, clarsimp)
apply (subgoal_tac "l:Dom b")
prefer 2 apply (simp add: Pbij_Dom_def) apply fast
apply (erule_tac x=l in allE, erule_tac x=ll in allE, erule impE, assumption)
apply (erule_tac x=l in allE, erule_tac x=l in allE, erule impE, erule mkId4)
apply (subgoal_tac "\<exists> x y . lookup b l = Some (x,y)", clarsimp)
prefer 2 apply (simp add: Dom_def)
apply (simp add: twiddleObj_def, clarsimp)
apply (subgoal_tac "\<exists> u . lookup y f = Some u", clarsimp)
prefer 2 apply (rotate_tac -6, erule thin_rl, erule thin_rl, erule thin_rl)
apply (simp add: LowDom_def) apply fastforce
apply (erule_tac x=f in allE, erule_tac x=f in allE, clarsimp)
apply (erule twiddleVal.cases)
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp
apply clarsimp
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp apply (drule mkId4b, clarsimp)
apply (erule twiddleVal_Loc)
apply clarsimp
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp
apply clarsimp apply (rule twiddleVal_IVal) apply simp
apply (simp add: HighSec_def)
apply (simp add: twiddle_def, clarsimp)
apply (rule_tac x=\<beta> in exI)
apply (rule, simp add: twiddleStore_def, clarsimp)
apply (erule_tac x=x in allE, clarsimp)
apply (erule_tac x=x in allE, clarsimp)
apply (erule twiddleVal.cases)
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp
apply clarsimp
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp apply (drule mkId4b, clarsimp)
apply (erule twiddleVal_Loc)
apply clarsimp
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp
apply clarsimp apply (rule twiddleVal_IVal) apply simp
apply (rule, simp add: twiddleHeap_def, clarsimp)
apply (rule, simp add: mkId2 mkId2b)
apply (simp add: mkId2 mkId2b, clarsimp)
apply (subgoal_tac "ll:Dom b")
prefer 2 apply (simp add: Pbij_Rng_def) apply fast
apply (erule_tac x=l in allE, erule_tac x=ll in allE, erule impE, assumption)
apply (erule_tac x=ll in allE, erule_tac x=ll in allE, erule impE, erule mkId4)
apply (subgoal_tac "\<exists> x y . lookup b ll = Some (x,y)", clarsimp)
prefer 2 apply (simp add: Dom_def)
apply (simp add: twiddleObj_def, clarsimp)
apply (subgoal_tac "\<exists> u . lookup y f = Some u", clarsimp)
prefer 2 apply (rotate_tac -8, erule thin_rl, erule thin_rl, erule thin_rl)
apply (simp add: LowDom_def) apply fastforce
apply (erule_tac x=f in allE, erule_tac x=f in allE, clarsimp)
apply (erule twiddleVal.cases)
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp
apply clarsimp
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp apply (drule mkId4b, clarsimp)
apply (erule twiddleVal_Loc)
apply clarsimp
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp
apply clarsimp apply (rule twiddleVal_IVal) apply simp
apply (rule Pbij_extends_reflexive)
done
(*>*)
lemma SkipHigh: "G \<rhd> Skip: HighSec"
(*<*)
apply (rule VDMConseq, rule VDMSkip)
apply (simp add: HighSec_def) apply clarsimp
(*apply (rule_tac x="mkId b" in exI)*)
apply (erule twiddle_mkId)
done
(*>*)
text\<open>We define a predicate expressing when locations obtained by
evaluating an expression are non-dangling.\<close>
definition Expr_good::"Expr \<Rightarrow> State \<Rightarrow> bool"
where "Expr_good e s =
(\<forall> l . evalE e (fst s) = RVal(Loc l) \<longrightarrow> l : Dom (snd s))"
lemma AssignHigh:
"\<lbrakk> CONTEXT x = high; \<forall> s . noLowDPs s \<longrightarrow> Expr_good e s\<rbrakk>
\<Longrightarrow> G \<rhd> Assign x e: HighSec"
(*<*)
apply (rule VDMConseq, rule VDMAssign, unfold HighSec_def)
apply (rule, rule, rule, rule)
apply (simp add: twiddle_def)
apply rule
(*noLowDPs t*)
apply (simp add: noLowDPs_def update_def) apply clarsimp
apply rule
(*twiddleStore*)
apply (simp add: twiddleStore_def update_def)
apply (rule, rule)
apply clarsimp
apply clarsimp
apply (case_tac "a xa")
apply (rename_tac Ref)
apply clarsimp apply (case_tac Ref)
apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (rule mkId4) apply (simp add: noLowDPs_def)
apply clarsimp apply (simp add: twiddleVal_IVal)
(*twiddleHeap*)
apply (simp add: twiddleHeap_def update_def) apply clarsimp
apply (rule, rule mkId1)
apply (rule, simp add: mkId2)
apply (rule, simp add: mkId2b)
apply clarsimp apply (drule mkId4b) apply clarsimp
(*l=ll*)
apply (simp add: twiddleObj_def)
apply clarsimp
apply (case_tac v, clarsimp)
apply (rename_tac Ref)
apply (case_tac Ref, clarsimp)
apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (rule mkId4) apply (simp add: noLowDPs_def)
apply clarsimp apply (rule twiddleVal_IVal) apply simp
done
(*>*)
lemma NewHigh:
"CONTEXT x = high \<Longrightarrow> G \<rhd> New x C : HighSec"
(*<*)
apply (rule VDMConseq, rule VDMNew, unfold HighSec_def)
apply (rule, rule, rule)
apply (erule exE, (erule conjE)+)
apply (rule, simp add: twiddle_def, rule)
(*noLowDPs t*)
apply (simp add: noLowDPs_def, clarsimp) apply (rename_tac l s h) apply rule
(*store content*)
apply clarsimp apply (simp add: update_def)
apply (case_tac "x=xa", clarsimp) apply (simp add: Dom_def)
(*apply clarsimp apply (erule_tac x=xa in allE, erule_tac x=la in allE, clarsimp)
apply (simp add: Dom_def)*)
(*Field content*)
apply clarsimp apply (simp add: Dom_def)
(*apply (rule_tac x="mkId (snd s)" in exI)*)
apply rule
(*twiddleStore*)
apply (simp add: twiddleStore_def)
apply (rule, rule) apply (simp add: update_def) apply clarsimp
apply (rename_tac s h l xa)
apply (case_tac "x=xa", clarsimp) apply clarsimp
apply (case_tac "s xa")
apply (rename_tac Ref)
apply clarsimp apply (case_tac Ref)
apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (rule mkId4) apply (simp add: noLowDPs_def)
apply clarsimp apply (simp add: twiddleVal_IVal)
(*twiddleHeap*)
apply (simp add: twiddleHeap_def) apply clarsimp
apply (rename_tac s h l)
apply (rule, rule mkId1)
apply (rule, simp add: mkId2)
apply (rule, simp add: mkId2b Dom_def) apply fastforce
apply clarsimp
apply rule
apply clarsimp apply (drule mkId4b) apply clarsimp
apply clarsimp apply (drule mkId4b) apply clarsimp
apply (rename_tac C F)
apply (simp add: twiddleObj_def) apply clarsimp
apply (case_tac v, clarsimp)
apply (rename_tac Ref)
apply (case_tac Ref, clarsimp)
apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (rule mkId4) apply (simp add: noLowDPs_def)
apply clarsimp apply (rule twiddleVal_IVal) apply simp
done
(*>*)
lemma GetHigh:
"\<lbrakk> CONTEXT x = high \<rbrakk> \<Longrightarrow> G \<rhd> Get x y f : HighSec"
(*<*)
apply (rule VDMConseq, rule VDMGet, unfold HighSec_def)
apply (rule, rule, rule)
apply (erule exE)+
apply (erule conjE)+
apply (rule, simp add: twiddle_def, rule)
(*noLowDPs t*)
apply (simp add: noLowDPs_def, clarsimp)
(*store content*)
apply (rename_tac s h xa la)
apply (simp add: update_def)
apply (case_tac "x=xa", clarsimp) apply (simp add: Dom_def)
(*Field content already discharged*)
(*apply (rule_tac x="mkId (snd s)" in exI)*)
apply rule
(*twiddleStore*)
apply (simp add: twiddleStore_def)
apply (rule, rule) apply (simp add: update_def) apply clarsimp
apply (rename_tac s h l C Flds v xa)
apply (case_tac "x=xa", clarsimp) apply clarsimp
apply (case_tac "s xa")
apply (rename_tac Ref)
apply clarsimp apply (case_tac Ref)
apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (rule mkId4) apply (simp add: noLowDPs_def)
apply clarsimp apply (simp add: twiddleVal_IVal)
(*twiddleHeap*)
apply (simp add: twiddleHeap_def) apply clarsimp
apply (rename_tac s h l C Flds v)
apply (rule, rule mkId1)
apply (rule, simp add: mkId2)
apply (rule, simp add: mkId2b)
apply clarsimp apply (drule mkId4b) apply clarsimp
apply (rename_tac D F)
apply (simp add: twiddleObj_def) apply clarsimp
apply (case_tac va, clarsimp)
apply (rename_tac Ref)
apply (case_tac Ref, clarsimp)
apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (rule mkId4) apply (simp add: noLowDPs_def)
apply clarsimp apply (rule twiddleVal_IVal) apply simp
done
(*>*)
lemma PutHigh:
"\<lbrakk> GAMMA f = high; \<forall> s . noLowDPs s \<longrightarrow> Expr_good e s\<rbrakk>
\<Longrightarrow> G \<rhd> Put x f e: HighSec"
(*<*)
apply (rule VDMConseq, rule VDMPut, unfold HighSec_def)
apply (rule, rule, rule)
apply (erule exE)+
apply (erule conjE)+
apply (rule, simp add: twiddle_def, rule)
(*noLowDPs t*)
apply (simp add: noLowDPs_def, clarsimp)
apply (rename_tac s h)
apply rule
(*store content*)
apply clarsimp apply (erule_tac x=xa in allE, erule_tac x=la in allE, clarsimp)
apply (simp add: Dom_def)
(*Field content*)
apply clarsimp apply rule
(*l=ll*)
apply clarsimp apply rule
(*f=fa*)
apply clarsimp
(*f\<noteq>fa*)
apply clarsimp apply (simp add: Dom_def)
(*l\<noteq>ll*)
apply clarsimp apply (simp add: Dom_def)
apply rule
(*twiddleStore*)
apply (simp add: twiddleStore_def)
apply (rule, rule)
apply (case_tac "fst s xa")
apply (rename_tac Ref)
apply clarsimp apply (case_tac Ref)
apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (rule mkId4) apply (simp add: noLowDPs_def)
apply clarsimp apply (simp add: twiddleVal_IVal)
(*twiddleHeap*)
apply (simp add: twiddleHeap_def) apply clarsimp
apply (rule, simp add: mkId1)
apply (rule, simp add: mkId2)
apply (rule, simp add: mkId2b Dom_def) apply fast
apply clarsimp apply rule
(*l=ll*)
apply clarsimp apply (drule mkId4b) apply clarsimp
apply (simp add: twiddleObj_def)
apply (rule, simp add: LowDom_def) apply (rotate_tac -1, erule thin_rl) apply fastforce
apply clarsimp apply rule
(*f=fa*)
apply clarsimp
(*f\<noteq>fa*)
apply clarsimp
apply (case_tac v, clarsimp)
apply (rename_tac Ref)
apply (case_tac Ref, clarsimp)
apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (rule mkId4) apply (simp add: noLowDPs_def)
apply clarsimp apply (rule twiddleVal_IVal) apply simp
(*l\<noteq>ll*)
apply clarsimp apply (drule mkId4b) apply clarsimp
apply (simp add: twiddleObj_def)
apply clarsimp
apply (case_tac v, clarsimp)
apply (rename_tac Ref)
apply (case_tac Ref, clarsimp)
apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_Loc) apply (rule mkId4) apply (simp add: noLowDPs_def)
apply clarsimp apply (rule twiddleVal_IVal) apply simp
done
(*>*)
(*<*)
lemma PutHigh2:
"\<lbrakk> GAMMA f = high; \<forall> s . Expr_good e s\<rbrakk> \<Longrightarrow> G \<rhd> Put x f e: HighSec"
apply (rule VDMConseq, erule PutHigh) apply simp apply simp
done
(*>*)
(*<*)
lemma twiddle_mkIDs_compose:
"\<lbrakk>(a,b) \<equiv>\<^sub>(mkId b) (ab, bb); (ab,bb) \<equiv>\<^sub>(mkId bb) (aa, ba)\<rbrakk>
\<Longrightarrow> (a,b) \<equiv>\<^sub>(mkId b) (aa, ba)"
+supply [[simproc del: defined_all]]
apply (simp add: twiddle_def)
apply (rule, simp add: twiddleStore_def, clarsimp)
apply (erule_tac x=x in allE, clarsimp)
apply (erule_tac x=x in allE, clarsimp)
apply (erule twiddleVal.cases)
apply clarsimp
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp apply clarsimp prefer 2 apply clarsimp
apply (drule mkId4b, clarsimp) apply (drule mkId4b, clarsimp)
apply (rule twiddleVal_Loc) apply (erule mkId4)
apply clarsimp
apply (simp add: twiddleHeap_def, clarsimp)
apply (simp add: mkId2 mkId2b)
apply rule apply fast
apply clarsimp apply (drule mkId4b, clarsimp)
apply (erule_tac x=ll in allE, erule_tac x=ll in allE, erule impE, erule mkId4) apply clarsimp
apply (subgoal_tac "ll:Dom bb") prefer 2 apply fast
apply (erule_tac x=ll in allE, erule_tac x=ll in allE, erule impE, erule mkId4) apply clarsimp
apply (subgoal_tac "\<exists> x y . lookup bb ll = Some(x,y)", clarsimp)
prefer 2 apply (simp add: Dom_def)
apply (simp add: twiddleObj_def, clarsimp)
apply (subgoal_tac "\<exists> u . lookup bc f = Some u", clarsimp)
prefer 2 apply (simp add: LowDom_def)
apply (subgoal_tac "\<exists> u . lookup y f = Some u", clarsimp)
prefer 2 apply (rotate_tac -7, erule thin_rl, erule thin_rl) apply (simp add: LowDom_def) apply fast
apply (erule_tac x=f in allE, clarsimp)
apply (erule_tac x=f in allE, clarsimp)
apply (erule twiddleVal.cases)
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp apply (rule twiddleVal_Null) apply clarsimp apply clarsimp
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp apply clarsimp prefer 2 apply clarsimp
apply (drule mkId4b, clarsimp) apply (drule mkId4b, clarsimp)
apply (rule twiddleVal_Loc) apply (erule mkId4)
apply clarsimp apply (erule twiddleVal.cases)
apply clarsimp apply clarsimp apply clarsimp apply (rule twiddleVal_IVal) apply clarsimp
done
(*>*)
(*<*)
lemma twiddle_mkIDs_compose':
"\<lbrakk> s \<equiv>\<^sub>(mkId (snd s)) r; r \<equiv>\<^sub>(mkId (snd r)) t\<rbrakk> \<Longrightarrow> s \<equiv>\<^sub>(mkId (snd s)) t"
apply (case_tac s, clarsimp)
apply (case_tac t, clarsimp)
apply (case_tac r, clarsimp)
apply (erule twiddle_mkIDs_compose, assumption)
done
(*>*)
lemma CompHigh:
"\<lbrakk> G \<rhd> c: HighSec; G \<rhd> d:HighSec\<rbrakk> \<Longrightarrow> G \<rhd> Comp c d: HighSec"
(*<*)
apply (rule VDMConseq, erule VDMComp, assumption)
apply (unfold HighSec_def)
apply (rule, rule, rule)
apply (erule thin_rl, erule thin_rl, erule exE)
apply (erule conjE) apply clarsimp
apply (erule impE) prefer 2 apply (erule twiddle_mkIDs_compose, assumption)
apply (simp add: noLowDPs_def twiddle_def)
done
(*>*)
lemma IfHigh:
"\<lbrakk> G \<rhd> c: HighSec; G \<rhd> d:HighSec\<rbrakk> \<Longrightarrow> G \<rhd> Iff b c d: HighSec"
(*<*)
apply (rule VDMConseq, rule VDMIff)
apply (assumption, assumption)
apply (erule thin_rl, erule thin_rl)
apply clarsimp
done(*>*)
lemma WhileHigh: "\<lbrakk> G \<rhd> c: HighSec\<rbrakk> \<Longrightarrow> G \<rhd> While b c: HighSec"
(*<*)
apply (rule VDMConseq, erule VDMWhile [of G c HighSec b HighSec])
apply (simp add: HighSec_def) apply clarsimp
(*apply (rule_tac x="mkId ba" in exI)*)
apply (erule twiddle_mkId)
apply clarsimp
apply (simp add: HighSec_def) apply clarsimp
(*apply (rule_tac x="Pbij_compose \<beta> \<beta>'" in exI)*)
apply (erule impE) prefer 2 apply (erule twiddle_mkIDs_compose, assumption)
(*apply (erule twiddle_compose, assumption)*)
apply (simp add: noLowDPs_def twiddle_def)
apply clarsimp
done(*>*)
lemma CallHigh:
"({HighSec} \<union> G) \<rhd> body : HighSec \<Longrightarrow> G \<rhd> Call : HighSec"
(*<*)
by (erule VDMCall)
(*>*)
text\<open>We combine all rules to an inductive derivation system.\<close>
inductive_set Deriv::"(Assn set \<times> OBJ \<times> Assn) set"
where
D_CAST:
"(G, c, HighSec):Deriv \<Longrightarrow>
(G, c, Sec (\<lambda> (s, t, \<beta>). s \<equiv>\<^sub>\<beta> t)):Deriv"
| D_SKIP: "(G, Skip, Sec (\<lambda> (s, t, \<beta>) . s \<equiv>\<^sub>\<beta> t)) : Deriv"
| D_ASSIGN:
"Expr_low e \<Longrightarrow>
(G, Assign x e, Sec (\<lambda> (s, t, \<beta>) .
\<exists> r . s = (update (fst r) x (evalE e (fst r)), snd r)
\<and> r \<equiv>\<^sub>\<beta> t)):Deriv"
| D_COMP:
"\<lbrakk> (G, c1, Sec \<Phi>):Deriv; (G, c2, Sec \<Psi>):Deriv\<rbrakk> \<Longrightarrow>
(G, Comp c1 c2, Sec (\<lambda> (s, t, \<beta>) .
\<exists> r . \<Phi>(r, t, \<beta>) \<and>
(\<forall> w \<gamma>. r \<equiv>\<^sub>\<gamma> w \<longrightarrow> \<Psi>(s, w, \<gamma>)))):Deriv"
| D_IFF:
"\<lbrakk> BExpr_low b; (G, c1, Sec \<Phi>):Deriv; (G, c2, Sec \<Psi>):Deriv\<rbrakk> \<Longrightarrow>
(G, Iff b c1 c2, Sec (\<lambda> (s,t,\<beta>) .
(evalB b (fst t) \<longrightarrow> \<Phi>(s,t,\<beta>)) \<and>
((\<not> evalB b (fst t)) \<longrightarrow> \<Psi>(s,t,\<beta>)))):Deriv"
| D_NEW:
"CONTEXT x = low \<Longrightarrow>
(G, New x C, Sec (\<lambda> (s,t,\<beta>) .
\<exists> l r . l \<notin> Dom (snd r) \<and> r \<equiv>\<^sub>\<beta> t \<and>
s = (update (fst r) x (RVal (Loc l)),
(l,(C,[])) # (snd r)))):Deriv"
| D_GET:
"\<lbrakk> CONTEXT y = low; GAMMA f = low\<rbrakk> \<Longrightarrow>
(G, Get x y f, Sec (\<lambda> (s,t,\<beta>) .
\<exists> r l C Flds v. (fst r) y = RVal(Loc l) \<and>
lookup (snd r) l = Some(C,Flds) \<and>
lookup Flds f = Some v \<and> r \<equiv>\<^sub>\<beta> t \<and>
s = (update (fst r) x v, snd r))):Deriv"
| D_PUT:
"\<lbrakk> CONTEXT x = low; GAMMA f = low; Expr_low e\<rbrakk> \<Longrightarrow>
(G, Put x f e, Sec (\<lambda> (s,t,\<beta>) .
\<exists> r l C F h. (fst r) x = RVal(Loc l) \<and> r \<equiv>\<^sub>\<beta> t \<and>
lookup (snd r) l = Some(C,F) \<and>
h = (l,(C,(f,evalE e (fst r)) # F)) # (snd r) \<and>
s = (fst r, h))):Deriv"
| D_WHILE:
"\<lbrakk> BExpr_low b; (G, c, Sec \<Phi>):Deriv\<rbrakk>
\<Longrightarrow> (G, While b c, Sec (PhiWhile b \<Phi>)):Deriv"
| D_CALL:
"\<lbrakk>({Sec (FIX \<phi>)} \<union> G, body, Sec (\<phi> (FIX \<phi>))):Deriv; Monotone \<phi>\<rbrakk>
\<Longrightarrow> (G, Call, Sec (FIX \<phi>)):Deriv"
| D_SKIP_H: "(G, Skip, HighSec):Deriv"
| D_ASSIGN_H:
"\<lbrakk> CONTEXT x = high; \<forall> s . noLowDPs s \<longrightarrow> Expr_good e s\<rbrakk>
\<Longrightarrow> (G, Assign x e, HighSec):Deriv"
| D_NEW_H: "CONTEXT x = high \<Longrightarrow> (G, New x C, HighSec):Deriv"
| D_GET_H: "CONTEXT x = high \<Longrightarrow> (G, Get x y f, HighSec):Deriv"
| D_PUT_H:
"\<lbrakk> GAMMA f = high; \<forall> s . noLowDPs s \<longrightarrow> Expr_good e s\<rbrakk>
\<Longrightarrow> (G, Put x f e, HighSec):Deriv"
| D_COMP_H:
"\<lbrakk> (G, c, HighSec):Deriv; (G, d, HighSec):Deriv\<rbrakk>
\<Longrightarrow> (G, Comp c d, HighSec):Deriv"
| D_IFF_H:
"\<lbrakk> (G, c, HighSec):Deriv; (G, d, HighSec):Deriv\<rbrakk>
\<Longrightarrow> (G, Iff b c d, HighSec):Deriv"
| D_WHILE_H:
"\<lbrakk> (G, c, HighSec):Deriv\<rbrakk> \<Longrightarrow> (G, While b c, HighSec):Deriv"
| D_CALL_H:
"({HighSec} \<union> G, body, HighSec):Deriv \<Longrightarrow> (G, Call, HighSec):Deriv"
text\<open>By construction, all derivations represent legal derivations in
the program logic. Here's an explicit lemma to this effect.\<close>
lemma Deriv_derivable: "(G,c,A):Deriv \<Longrightarrow> G \<rhd> c: A"
(*<*)
apply (erule Deriv.induct)
apply (erule CAST)
apply (rule SKIP)
apply (erule ASSIGN)
apply (erule COMP) apply assumption
apply (erule IFF) apply assumption apply assumption
apply (erule NEW)
apply (erule GET) apply assumption
apply (drule_tac G=G in PUT) apply assumption apply assumption apply simp
apply (erule WHILE) apply assumption
apply (erule CALL) apply assumption
apply (rule SkipHigh)
apply (erule AssignHigh) apply assumption
apply (erule NewHigh)
apply (erule GetHigh)
apply (erule PutHigh) apply assumption
apply (erule CompHigh) apply assumption
apply (erule IfHigh) apply assumption
apply (erule WhileHigh)
apply (erule CallHigh)
done
(*>*)
subsection\<open>Type system\<close>
text\<open>\label{sec:ObjTypeSystem}\<close>
text\<open>We now give a type system in the style of Volpano et al.~and
then prove its embedding into the system of derived rules. First, type
systems for expressions and boolean expressions. These are similar to
the ones in Section \ref{sec:BaseLineNI} but require some side
conditions regarding the (semantically modelled) operators.\<close>
definition opEGood::"(Val \<Rightarrow> Val \<Rightarrow> Val) \<Rightarrow> bool"
where "opEGood f = (\<forall> \<beta> v v' w w' . (\<beta>, v, v') \<in> twiddleVal\<longrightarrow>
(\<beta>, w, w') \<in> twiddleVal \<longrightarrow> (\<beta>, f v w, f v' w') \<in> twiddleVal)"
definition compBGood::"(Val \<Rightarrow> Val \<Rightarrow> bool) \<Rightarrow> bool"
where "compBGood f = (\<forall> \<beta> v v' w w' . (\<beta>, v, v') \<in> twiddleVal\<longrightarrow>
(\<beta>, w, w') \<in> twiddleVal \<longrightarrow> f v w = f v' w')"
inductive_set VS_expr:: "(Expr \<times> TP) set"
where
VS_exprVar: "CONTEXT x = t \<Longrightarrow> (varE x, t) : VS_expr"
|
VS_exprVal:
"\<lbrakk>v=RVal Nullref \<or> (\<exists> i . v=IVal i)\<rbrakk> \<Longrightarrow> (valE v, low) : VS_expr"
|
VS_exprOp:
"\<lbrakk>(e1,t) : VS_expr; (e2,t):VS_expr; opEGood f\<rbrakk>
\<Longrightarrow> (opE f e1 e2,t) : VS_expr"
|
VS_exprHigh: "(e, high) : VS_expr"
inductive_set VS_Bexpr:: "(BExpr \<times> TP) set"
where
VS_BexprOp:
"\<lbrakk>(e1,t):VS_expr; (e2,t):VS_expr; compBGood f\<rbrakk>
\<Longrightarrow> (compB f e1 e2,t):VS_Bexpr"
|
VS_BexprHigh: "(e,high) : VS_Bexpr"
text\<open>Next, the core of the type system, the rules for commands. The
second side conditions of rules \<open>VS_comAssH\<close> and \<open>VS_comPutH\<close> could be strengthened to $\forall\; s .\;
\mathit{Epxr\_good}\; e\; s$.\<close>
inductive_set VS_com:: "(TP \<times> OBJ) set"
where
VS_comGetL:
"\<lbrakk> CONTEXT y = low; GAMMA f = low\<rbrakk>
\<Longrightarrow> (low, Get x y f):VS_com"
| VS_comGetH: "CONTEXT x = high \<Longrightarrow> (high, Get x y f):VS_com"
| VS_comPutL:
"\<lbrakk> CONTEXT x = low; GAMMA f = low; (e, low) : VS_expr\<rbrakk>
\<Longrightarrow> (low,Put x f e):VS_com"
| VS_comPutH:
"\<lbrakk> GAMMA f = high; \<forall> s . noLowDPs s \<longrightarrow> Expr_good e s\<rbrakk>
\<Longrightarrow> (high,Put x f e):VS_com"
| VS_comNewL:
"CONTEXT x = low \<Longrightarrow> (low, New x c) : VS_com"
| VS_comNewH:
"CONTEXT x = high \<Longrightarrow> (high, New x C):VS_com"
| VS_comAssL:
"\<lbrakk>CONTEXT x = low; (e,low):VS_expr\<rbrakk>
\<Longrightarrow> (low,Assign x e) : VS_com"
| VS_comAssH:
"\<lbrakk>CONTEXT x = high; \<forall> s . noLowDPs s \<longrightarrow> Expr_good e s\<rbrakk>
\<Longrightarrow> (high,Assign x e) : VS_com"
| VS_comSkip: "(pc,Skip) : VS_com"
| VS_comComp:
"\<lbrakk> (pc,c):VS_com; (pc,d):VS_com\<rbrakk> \<Longrightarrow> (pc,Comp c d) : VS_com"
| VS_comIf:
"\<lbrakk> (b,pc):VS_Bexpr; (pc,c):VS_com; (pc,d):VS_com\<rbrakk>
\<Longrightarrow> (pc,Iff b c d):VS_com"
| VS_comWhile:
"\<lbrakk>(b,pc):VS_Bexpr; (pc,c):VS_com\<rbrakk> \<Longrightarrow> (pc,While b c):VS_com"
| VS_comSub: "(high,c) : VS_com \<Longrightarrow> (low,c):VS_com"
text\<open>In order to prove the type system sound, we first define the
interpretation of expression typings\ldots\<close>
primrec SemExpr::"Expr \<Rightarrow> TP \<Rightarrow> bool"
where
"SemExpr e low = Expr_low e" |
"SemExpr e high = True"
text\<open>\ldots and show the soundness of the typing rules.\<close>
lemma ExprSound: "(e,tp):VS_expr \<Longrightarrow> SemExpr e tp"
(*<*)
apply (erule VS_expr.induct)
apply clarsimp
apply (case_tac "CONTEXT x", clarsimp) apply (simp add: Expr_low_def twiddleStore_def)
apply clarsimp
apply (simp add: Expr_low_def)
apply (erule disjE) apply clarsimp apply (rule twiddleVal_Null)
apply clarsimp apply (rule twiddleVal_IVal) apply simp
apply (case_tac t, clarsimp)
apply (simp add: Expr_low_def, clarsimp)
apply (erule_tac x=s in allE, erule_tac x=t in allE, erule_tac x=\<beta> in allE, erule impE, assumption)
apply (erule_tac x=s in allE, erule_tac x=t in allE, erule_tac x=\<beta> in allE, erule impE, assumption)
apply (simp add: opEGood_def)
apply simp
apply clarsimp
done
(*>*)
text\<open>Likewise for the boolean expressions.\<close>
primrec SemBExpr::"BExpr \<Rightarrow> TP \<Rightarrow> bool"
where
"SemBExpr b low = BExpr_low b" |
"SemBExpr b high = True"
lemma BExprSound: "(e,tp):VS_Bexpr \<Longrightarrow> SemBExpr e tp"
(*<*)
apply (erule VS_Bexpr.induct, simp_all)
apply (drule ExprSound)
apply (drule ExprSound)
apply (case_tac t, simp_all add: BExpr_low_def Expr_low_def) apply clarsimp
apply (erule_tac x=s in allE, erule_tac x=ta in allE, erule_tac x=\<beta> in allE, erule impE, assumption)
apply (erule_tac x=s in allE, erule_tac x=ta in allE, erule_tac x=\<beta> in allE, erule impE, assumption)
apply (simp add: compBGood_def)
done
(*>*)
text\<open>Using these auxiliary lemmas we can prove the embedding of the
type system for commands into the system of derived proof rules, by
induction on the typing rules.\<close>
theorem VS_com_Deriv[rule_format]:
"(t,c):VS_com \<Longrightarrow> (t=high \<longrightarrow> (G, c, HighSec):Deriv) \<and>
(t=low \<longrightarrow> (\<exists> \<Phi> . (G, c, Sec \<Phi>):Deriv))"
(*<*)
apply (erule VS_com.induct)
apply clarsimp
apply (rule, erule D_GET) apply assumption
apply clarsimp apply (erule D_GET_H)
apply clarsimp
apply (rule, erule D_PUT) apply assumption apply (drule ExprSound) apply simp
apply clarsimp apply (erule D_PUT_H) apply simp
apply clarsimp apply (rule, erule D_NEW)
apply clarsimp apply (erule D_NEW_H)
apply clarsimp apply (rule, rule D_ASSIGN) apply (drule ExprSound) apply simp
apply clarsimp apply (erule D_ASSIGN_H) apply simp
apply rule
apply clarsimp apply (rule D_SKIP_H)
apply clarsimp apply (rule, rule D_SKIP)
apply (erule conjE)+ apply rule
apply rule apply (erule impE, assumption) apply (erule impE, assumption)
apply (erule D_COMP_H) apply assumption
apply rule apply (erule impE, assumption, erule exE) apply (erule impE, assumption, erule exE)
apply (rule, erule D_COMP) apply assumption
apply (erule conjE)+ apply rule
apply rule apply (erule impE, assumption) apply (erule impE, assumption)
apply (erule D_IFF_H) apply assumption
apply rule apply (erule impE, assumption, erule exE) apply (erule impE, assumption, erule exE)
apply (rule, rule D_IFF) prefer 2 apply assumption prefer 2 apply assumption
apply (drule BExprSound) apply simp
apply (erule conjE)+ apply rule
apply (rule, erule impE, assumption)
apply (erule D_WHILE_H)
apply (rule, erule impE, assumption, erule exE)
apply (rule, rule D_WHILE) prefer 2 apply assumption
apply (drule BExprSound) apply simp
apply simp apply (rule, erule D_CAST)
done
(*>*)
text\<open>Combining this result with the derivability of the derived proof
system and the soundness theorem of the program logic yields non-interference
of programs that are low typeable.\<close>
theorem VS_SOUND: "(low,c):VS_com \<Longrightarrow> secure c"
(*<*)
apply (drule_tac G="{}" in VS_com_Deriv, simp)
apply (erule exE)
apply (drule Deriv_derivable)
apply (drule VDM_Sound_emptyCtxt)
apply (erule Prop1A)
done
(*>*)
text\<open>End of theory \<open>VS_OBJ\<close>\<close>
end
diff --git a/thys/Separation_Algebra/ex/Simple_Separation_Example.thy b/thys/Separation_Algebra/ex/Simple_Separation_Example.thy
--- a/thys/Separation_Algebra/ex/Simple_Separation_Example.thy
+++ b/thys/Separation_Algebra/ex/Simple_Separation_Example.thy
@@ -1,108 +1,109 @@
(* Title: Adaptation of example from HOL/Hoare/Separation
Author: Gerwin Klein, 2012
Maintainers: Gerwin Klein <kleing at cse.unsw.edu.au>
Rafal Kolanski <rafal.kolanski at nicta.com.au>
*)
section "Example from HOL/Hoare/Separation"
theory Simple_Separation_Example
imports "HOL-Hoare.Hoare_Logic_Abort" "../Sep_Heap_Instance"
"../Sep_Tactics"
begin
declare [[syntax_ambiguity_warning = false]]
type_synonym heap = "(nat \<Rightarrow> nat option)"
(* This syntax isn't ideal, but this is what is used in the original *)
definition maps_to:: "nat \<Rightarrow> nat \<Rightarrow> heap \<Rightarrow> bool" ("_ \<mapsto> _" [56,51] 56)
where "x \<mapsto> y \<equiv> \<lambda>h. h = [x \<mapsto> y]"
(* If you don't mind syntax ambiguity: *)
notation pred_ex (binder "\<exists>" 10)
(* could be generated automatically *)
definition maps_to_ex :: "nat \<Rightarrow> heap \<Rightarrow> bool" ("_ \<mapsto> -" [56] 56)
where "x \<mapsto> - \<equiv> \<exists>y. x \<mapsto> y"
(* could be generated automatically *)
lemma maps_to_maps_to_ex [elim!]:
"(p \<mapsto> v) s \<Longrightarrow> (p \<mapsto> -) s"
by (auto simp: maps_to_ex_def)
(* The basic properties of maps_to: *)
lemma maps_to_write:
"(p \<mapsto> - ** P) H \<Longrightarrow> (p \<mapsto> v ** P) (H (p \<mapsto> v))"
apply (clarsimp simp: sep_conj_def maps_to_def maps_to_ex_def split: option.splits)
apply (rule_tac x=y in exI)
apply (auto simp: sep_disj_fun_def map_convs map_add_def split: option.splits)
done
lemma points_to:
"(p \<mapsto> v ** P) H \<Longrightarrow> the (H p) = v"
by (auto elim!: sep_conjE
simp: sep_disj_fun_def maps_to_def map_convs map_add_def
split: option.splits)
(* This differs from the original and uses separation logic for the definition. *)
primrec
list :: "nat \<Rightarrow> nat list \<Rightarrow> heap \<Rightarrow> bool"
where
"list i [] = (\<langle>i=0\<rangle> and \<box>)"
| "list i (x#xs) = (\<langle>i=x \<and> i\<noteq>0\<rangle> and (EXS j. i \<mapsto> j ** list j xs))"
lemma list_empty [simp]:
shows "list 0 xs = (\<lambda>s. xs = [] \<and> \<box> s)"
by (cases xs) auto
(* The examples from Hoare/Separation.thy *)
lemma "VARS x y z w h
{(x \<mapsto> y ** z \<mapsto> w) h}
SKIP
{x \<noteq> z}"
apply vcg
apply(auto elim!: sep_conjE simp: maps_to_def sep_disj_fun_def domain_conv)
done
lemma "VARS H x y z w
{(P ** Q) H}
SKIP
{(Q ** P) H}"
apply vcg
apply(simp add: sep_conj_commute)
done
lemma "VARS H
{p\<noteq>0 \<and> (p \<mapsto> - ** list q qs) H}
H := H(p \<mapsto> q)
{list p (p#qs) H}"
apply vcg
apply (auto intro: maps_to_write)
done
lemma "VARS H p q r
{(list p Ps ** list q Qs) H}
WHILE p \<noteq> 0
INV {\<exists>ps qs. (list p ps ** list q qs) H \<and> rev ps @ qs = rev Ps @ Qs}
DO r := p; p := the(H p); H := H(r \<mapsto> q); q := r OD
{list q (rev Ps @ Qs) H}"
+ supply [[simproc del: defined_all]]
apply vcg
apply fastforce
apply clarsimp
apply (case_tac ps, simp)
apply (rename_tac p ps')
apply (clarsimp simp: sep_conj_exists sep_conj_ac)
apply (sep_subst points_to)
apply (rule_tac x = "ps'" in exI)
apply (rule_tac x = "p # qs" in exI)
apply (simp add: sep_conj_exists sep_conj_ac)
apply (rule exI)
apply (sep_rule maps_to_write)
apply ((sep_cancel add: maps_to_maps_to_ex)+)[1]
apply clarsimp
done
end
diff --git a/thys/Separation_Logic_Imperative_HOL/Examples/Circ_List.thy b/thys/Separation_Logic_Imperative_HOL/Examples/Circ_List.thy
--- a/thys/Separation_Logic_Imperative_HOL/Examples/Circ_List.thy
+++ b/thys/Separation_Logic_Imperative_HOL/Examples/Circ_List.thy
@@ -1,250 +1,250 @@
section \<open>Circular Singly Linked Lists\<close>
theory Circ_List
imports List_Seg Imp_List_Spec
begin
text \<open>
Example of circular lists, with efficient append, prepend, pop, and rotate
operations.
\<close>
subsection \<open>Datatype Definition\<close>
type_synonym 'a cs_list = "'a node ref option"
text \<open>A circular list is described by a list segment, with special
cases for the empty list:\<close>
fun cs_list :: "'a::heap list \<Rightarrow> 'a node ref option \<Rightarrow> assn" where
"cs_list [] None = emp"
| "cs_list (x#l) (Some p) = lseg (x#l) (Some p) (Some p)"
| "cs_list _ _ = false"
lemma [simp]: "cs_list l None = \<up>(l=[])"
by (cases l) auto
lemma [simp]:
"cs_list l (Some p)
= (\<exists>\<^sub>Ax ls. \<up>(l=x#ls) * lseg (x#ls) (Some p) (Some p))"
apply (rule ent_iffI)
apply (cases l)
apply simp
apply sep_auto
apply (cases l)
apply simp
apply sep_auto
done
subsection \<open>Precision\<close>
lemma cs_prec:
"precise cs_list"
apply rule
apply (case_tac p)
apply clarsimp
apply clarsimp
apply (subgoal_tac "x=xa \<and> n=na", simp)
apply (erule prec_frame_expl[OF lseg_prec1])
apply frame_inference
apply frame_inference
apply (drule prec_frame[OF sngr_prec])
apply frame_inference
apply frame_inference
apply simp
done
lemma cs_imp_list_impl: "imp_list cs_list"
apply unfold_locales
apply (rule cs_prec)
done
interpretation cs: imp_list cs_list by (rule cs_imp_list_impl)
subsection \<open>Operations\<close>
subsubsection \<open>Allocate Empty List\<close>
definition cs_empty :: "'a::heap cs_list Heap" where
"cs_empty \<equiv> return None"
lemma cs_empty_rule: "<emp> cs_empty <cs_list []>"
unfolding cs_empty_def
by sep_auto
lemma cs_empty_impl: "imp_list_empty cs_list cs_empty"
by unfold_locales (sep_auto heap: cs_empty_rule)
interpretation cs: imp_list_empty cs_list cs_empty by (rule cs_empty_impl)
subsubsection \<open>Prepend Element\<close>
fun cs_prepend :: "'a \<Rightarrow> 'a::heap cs_list \<Rightarrow> 'a cs_list Heap" where
"cs_prepend x None = do {
p \<leftarrow> ref (Node x None);
p:=Node x (Some p);
return (Some p)
}"
| "cs_prepend x (Some p) = do {
n \<leftarrow> !p;
q \<leftarrow> ref (Node (val n) (next n));
p := Node x (Some q);
return (Some p)
}"
declare cs_prepend.simps [simp del]
lemma cs_prepend_rule:
"<cs_list l p> cs_prepend x p <cs_list (x#l)>"
apply (cases p)
apply simp_all
apply (sep_auto simp: cs_prepend.simps)
apply (sep_auto simp: cs_prepend.simps)
done
lemma cs_prepend_impl: "imp_list_prepend cs_list cs_prepend"
by unfold_locales (sep_auto heap: cs_prepend_rule)
interpretation cs: imp_list_prepend cs_list cs_prepend
by (rule cs_prepend_impl)
subsubsection \<open>Append Element\<close>
fun cs_append :: "'a \<Rightarrow> 'a::heap cs_list \<Rightarrow> 'a cs_list Heap" where
"cs_append x None = do {
p \<leftarrow> ref (Node x None);
p:=Node x (Some p);
return (Some p) }"
| "cs_append x (Some p) = do {
n \<leftarrow> !p;
q \<leftarrow> ref (Node (val n) (next n));
p := Node x (Some q);
return (Some q)
}"
declare cs_append.simps [simp del]
lemma cs_append_rule:
"<cs_list l p> cs_append x p <cs_list (l@[x])>"
apply (cases p)
apply simp_all
apply (sep_auto simp: cs_append.simps)
apply (sep_auto simp: cs_append.simps)
apply (rule ent_frame_fwd)
- apply (rule_tac s=pp in lseg_append) (* frame_inference does no backtracking
+ apply (rule_tac s=a in lseg_append) (* frame_inference does no backtracking
on instantiating schematics, hence we have to give it some help here. *)
apply frame_inference
apply (sep_auto)
done
lemma cs_append_impl: "imp_list_append cs_list cs_append"
by unfold_locales (sep_auto heap: cs_append_rule)
interpretation cs: imp_list_append cs_list cs_append
by (rule cs_append_impl)
subsubsection \<open>Pop First Element\<close>
fun cs_pop :: "'a::heap cs_list \<Rightarrow> ('a\<times>'a cs_list) Heap" where
"cs_pop None = raise STR ''Pop from empty list''"
| "cs_pop (Some p) = do {
n1 \<leftarrow> !p;
if next n1 = Some p then
return (val n1,None) \<comment> \<open>Singleton list becomes empty list\<close>
else do {
let p2 = the (next n1);
n2 \<leftarrow> !p2;
p := Node (val n2) (next n2);
return (val n1,Some p)
}
}"
declare cs_pop.simps[simp del]
lemma cs_pop_rule:
"<cs_list (x#l) p> cs_pop p <\<lambda>(y,p'). cs_list l p' * true * \<up>(y=x)>"
apply (cases p)
apply (sep_auto simp: cs_pop.simps)
apply (cases l)
apply (sep_auto simp: cs_pop.simps dflt_simps: option.sel)
apply (sep_auto
simp: cs_pop.simps
dflt_simps: option.sel
eintros del: exI)
(* Some unfortunate quantifier fiddling :( *)
apply (rule_tac x=aa in exI)
apply (rule_tac x=list in exI)
- apply (rule_tac x=pp in exI)
+ apply (rule_tac x=a in exI)
apply clarsimp
apply (rule exI)
apply sep_auto
done
lemma cs_pop_impl: "imp_list_pop cs_list cs_pop"
apply unfold_locales
apply (sep_auto heap: cs_pop_rule elim!: neq_NilE)
done
interpretation cs: imp_list_pop cs_list cs_pop by (rule cs_pop_impl)
subsubsection \<open>Rotate\<close>
fun cs_rotate :: "'a::heap cs_list \<Rightarrow> 'a cs_list Heap" where
"cs_rotate None = return None"
| "cs_rotate (Some p) = do {
n \<leftarrow> !p;
return (next n)
}"
declare cs_rotate.simps [simp del]
lemma cs_rotate_rule:
"<cs_list l p> cs_rotate p <cs_list (rotate1 l)>"
apply (cases p)
apply (sep_auto simp: cs_rotate.simps)
apply (cases l)
apply simp
apply (case_tac list)
apply simp
apply (sep_auto simp: cs_rotate.simps)
apply (sep_auto simp: cs_rotate.simps)
apply (rule ent_frame_fwd)
- apply (rule_tac s="pp" in lseg_append)
+ apply (rule_tac s="a" in lseg_append)
apply frame_inference
apply sep_auto
done
lemma cs_rotate_impl: "imp_list_rotate cs_list cs_rotate"
apply unfold_locales
apply (sep_auto heap: cs_rotate_rule)
done
interpretation cs: imp_list_rotate cs_list cs_rotate by (rule cs_rotate_impl)
subsection \<open>Test\<close>
definition "test \<equiv> do {
l \<leftarrow> cs_empty;
l \<leftarrow> cs_append ''a'' l;
l \<leftarrow> cs_append ''b'' l;
l \<leftarrow> cs_append ''c'' l;
l \<leftarrow> cs_prepend ''0'' l;
l \<leftarrow> cs_rotate l;
(v1,l)\<leftarrow>cs_pop l;
(v2,l)\<leftarrow>cs_pop l;
(v3,l)\<leftarrow>cs_pop l;
(v4,l)\<leftarrow>cs_pop l;
return [v1,v2,v3,v4]
}"
definition "test_result \<equiv> [''a'', ''b'', ''c'', ''0'']"
lemma "<emp> test <\<lambda>r. \<up>(r=test_result) * true>"
unfolding test_def test_result_def
apply (sep_auto)
done
export_code test checking SML_imp
ML_val \<open>
val res = @{code test} ();
if res = @{code test_result} then () else raise Match;
\<close>
hide_const (open) test test_result
end
diff --git a/thys/Separation_Logic_Imperative_HOL/Examples/Idioms.thy b/thys/Separation_Logic_Imperative_HOL/Examples/Idioms.thy
--- a/thys/Separation_Logic_Imperative_HOL/Examples/Idioms.thy
+++ b/thys/Separation_Logic_Imperative_HOL/Examples/Idioms.thy
@@ -1,257 +1,257 @@
section \<open>Common Proof Methods and Idioms\<close>
theory Idioms
imports "../Sep_Main" Open_List Circ_List Hash_Set_Impl
begin
text_raw\<open>\label{thy:ex:idioms}\<close>
text \<open>
This theory gives a short documentation of common proof techniques and
idioms for the separation logic framework. For this purpose, it presents
some proof snippets (inspired by the other example theories), and heavily
comments on them.
\<close>
subsection \<open>The Method \<open>sep_auto\<close>\<close>
text \<open>The most versatile method of our framework is \<open>sep_auto\<close>,
which integrates the verification condition generator, the entailment
solver and some pre- and postprocessing tactics based on the simplifier
and classical reasoner. It can be applied to a Hoare-triple or entailment
subgoal, and will try to solve it, and any emerging new goals. It stops
when the goal is either solved or it gets stuck somewhere.\<close>
text \<open>As a simple example for \<open>sep_auto\<close> consider the following
program that does some operations on two circular lists:\<close>
definition "test \<equiv> do {
l1 \<leftarrow> cs_empty;
l2 \<leftarrow> cs_empty;
l1 \<leftarrow> cs_append ''a'' l1;
l2 \<leftarrow> cs_append ''c'' l2;
l1 \<leftarrow> cs_append ''b'' l1;
l2 \<leftarrow> cs_append ''e'' l2;
l2 \<leftarrow> cs_prepend ''d'' l2;
l2 \<leftarrow> cs_rotate l2;
return (l1,l2)
}"
text \<open>The \<open>sep_auto\<close> method does all the
necessary frame-inference automatically, and thus manages to prove
the following lemma in one step:\<close>
lemma "<emp>
test
<\<lambda>(l1,l2). cs_list [''a'',''b''] l1
* cs_list [''c'',''e'',''d''] l2>\<^sub>t"
unfolding test_def
apply (sep_auto)
done
text \<open>\<open>sep_auto\<close> accepts all the section-options of the classical
reasoner and simplifier, e.g., \<open>simp add/del:\<close>, \<open>intro:\<close>.
Moreover, it has some more section options, the most useful being
\<open>heap add/del:\<close> to add or remove Hoare-rules that are applied
with frame-inference. A complete documentation of the accepted options can
be found in Section~\ref{sec:auto:overview}.
\<close>
text \<open>As a typical example, consider the following proof:\<close>
lemma complete_ht_rehash:
"<is_hashtable l ht> ht_rehash ht
<\<lambda>r. is_hashtable l ht * is_hashtable (ls_rehash l) r>"
proof -
have LEN: " l \<noteq> [] \<Longrightarrow> Suc 0 < 2 * length l" by (cases l) auto
show ?thesis
apply (rule cons_pre_rule[OF ht_imp_len])
unfolding ht_rehash_def
apply (sep_auto
heap: complete_ht_new_sz complete_ht_copy
simp: ls_rehash_def LEN
) \<comment> \<open>Here we add a heap-rule, and some simp-rules\<close>
done
qed
subsection \<open>Applying Single Rules\<close>
text \<open>\paragraph{Hoare Triples} In this example, we show how to do
a proof step-by-step.\<close>
lemma
"<os_list xs n> os_prepend x n <os_list (x # xs)>"
unfolding os_prepend_def
txt \<open>The rules to deconstruct compound statements are contained in the
\<open>sep_decon_rules\<close> collection\<close>
thm sep_decon_rules
apply (rule sep_decon_rules)
txt \<open>The rules for statement that deend on the heap are
contained in the \<open>sep_heap_rules\<close> collection. The
\<open>fi_rule\<close>-lemma prepares frame inference for them\<close>
apply (rule sep_heap_rules[THEN fi_rule])
apply frame_inference \<comment> \<open>This method does the frame-inference\<close>
txt \<open>The consequence rule comes in three versions,
\<open>const_rule\<close>, \<open>cons_pre_rule\<close>,
and \<open>cons_post_rule\<close>\<close>
apply (rule cons_post_rule)
apply (rule sep_decon_rules)
txt \<open>A simplification unfolds \<open>os_list\<close> and extract the
pure part of the assumption\<close>
apply (clarsimp)
txt \<open>We can use \<open>ent_ex_postI\<close> to manually introduce
existentials in entailsments\<close>
apply (rule_tac x=xa in ent_ex_postI)
apply (rule_tac x=n in ent_ex_postI)
txt \<open>The simplifier has a setup for assertions, so it will do the rest\<close>
apply simp
done
text \<open>Note that the proof above can be done with \<open>sep_auto\<close>,
the "Swiss army knife" of our framework\<close>
lemma
"<os_list xs n> os_prepend x n <os_list (x # xs)>"
unfolding os_prepend_def by sep_auto
text \<open>\paragraph{Entailment} This example presents an actual proof
from the circular list theory, where we have to manually apply a
rule and give some hints to frame inference\<close>
lemma cs_append_rule:
"<cs_list l p> cs_append x p <cs_list (l@[x])>"
apply (cases p)
apply (sep_auto simp: cs_append.simps)
apply (sep_auto simp: cs_append.simps heap: lseg_append)
txt \<open>At this point, we are left with an entailment subgoal that sep-auto
cannot solve. A closer look reveals that we could use the rule
\<open>lseg_append\<close>.
With the \<open>ent_frame_fwd\<close>-rule, we can manually apply a rule to
solve an entailment, involving frame inference. In this case, we have
the additional problem that frame-inference guesses
a wrong instantiation, and is not able to infer the frame.
So we have to pre-instantiate the rule, as done below.\<close>
- apply (rule_tac s1=pp in ent_frame_fwd[OF lseg_append])
+ apply (rule_tac s1=a in ent_frame_fwd[OF lseg_append])
apply frame_inference \<comment> \<open>Now frame-inference is able to infer the frame\<close>
txt \<open>Now we are left with a trivial entailment, modulo commutativity of
star. This can be handled by the entailment solver:\<close>
apply solve_entails
done
subsection \<open>Functions with Explicit Recursion\<close>
text \<open>If the termination argument of a function depends on one of
its parameters, we can use the function package. For example,
the following function inserts elements from a list into a hash-set:\<close>
fun ins_from_list
:: "('x::{heap,hashable}) list \<Rightarrow> 'x hashset \<Rightarrow> 'x hashset Heap"
where
"ins_from_list [] hs = return hs" |
"ins_from_list (x # l) hs = do { hs \<leftarrow> hs_ins x hs; ins_from_list l hs }"
text \<open>Proofs over such functions are usually done by structural
induction on the explicit parameter, in this case, on the list\<close>
lemma ins_from_list_correct:
"<is_hashset s hs> ins_from_list l hs <is_hashset (s\<union>set l)>\<^sub>t"
proof (induction l arbitrary: hs s)
case (Cons x l)
txt \<open>In the induction step, the induction hypothesis has to be
declared as a heap-rule, as \<open>sep_auto\<close> currently does not
look for potential heap-rules among the premises of the subgoal\<close>
show ?case by (sep_auto heap: Cons.IH)
qed sep_auto
subsection \<open>
Functions with Recursion Involving the Heap
\<close>
text \<open>If the termination argument of a function depends on data stored on
the heap, \<open>partial_function\<close> is a useful tool.
Note that, despite the name, proving a Hoare-Triple \<open><\<dots>> \<dots> <\<dots>>\<close>
for something defined with \<open>partial_function\<close> implies total
correctness.
\<close>
text \<open>In the following example, we compute the sum of a list, using an
iterator. Note that the partial-function package does not provide a
code generator setup by default, so we have to add a \<open>[code]\<close>
attribute manually\<close>
partial_function (heap) os_sum' :: "int os_list_it \<Rightarrow> int \<Rightarrow> int Heap"
where [code]:
"os_sum' it s = do {
b \<leftarrow> os_it_has_next it;
if b then do {
(x,it') \<leftarrow> os_it_next it;
os_sum' it' (s+x)
} else return s
}"
text \<open>The proof that the function is correct can be done by induction
over the representation of the list that we still have to iterate over.
Note that for iterators over sets, we need induction on finite sets,
cf. also \<open>To_List_Ga.thy\<close>\<close>
lemma os_sum'_rule:
"<os_is_it l p l' it>
os_sum' it s
<\<lambda>r. os_list l p * \<up>(r = s + sum_list l')>\<^sub>t"
proof (induct l' arbitrary: it s)
case Nil thus ?case
txt \<open>To unfold the definition of a partial function, we have to use
\<open>subst\<close>. Note that \<open>simp\<close> would loop, unfolding the
function arbitrarily deep\<close>
apply (subst os_sum'.simps)
txt \<open>\<open>sep_auto\<close> accepts all the section parameters that
\<open>auto\<close> does, eg. \<open>intro:\<close>\<close>
apply (sep_auto intro: os.quit_iteration)
done
next
case (Cons x l')
show ?case
apply (subst os_sum'.simps)
txt \<open>Additionally, \<open>sep_auto\<close> accepts some more section
parameters. The most common one, \<open>heap:\<close>, declares rules
to be used with frame inference. See Section~\ref{sec:auto:overview}
for a complete overview.\<close>
apply (sep_auto heap: Cons.hyps)
done
qed
subsection \<open>Precision Proofs\<close>
text \<open>
Precision lemmas show that an assertion uniquely determines some of its
parameters. Our example shows that two list segments from the same start
pointer and with the same list, also have to end at the same end pointer.
\<close>
lemma lseg_prec3:
"\<forall>q q'. h \<Turnstile> (lseg l p q * F1) \<and>\<^sub>A (lseg l p q' * F2) \<longrightarrow> q=q'"
apply (intro allI)
proof (induct l arbitrary: p F1 F2)
case Nil thus ?case
apply simp \<comment> \<open>A precision solver for references and arrays is included
in the standard simplifier setup. Building a general precision solver
remains future work.\<close>
by metis \<comment> \<open>Unfortunately, the simplifier cannot cope with arbitrarily
directed equations, so we have to use some more powerful tool\<close>
next
case (Cons x l)
show ?case
apply clarsimp
apply (subgoal_tac "na=n")
txt \<open>The \<open>prec_frame\<close> and \<open>prec_frame'\<close> rules are
useful to do precision proofs\<close>
apply (erule prec_frame'[OF Cons.hyps])
apply frame_inference
apply frame_inference
apply (drule prec_frame[OF sngr_prec])
apply frame_inference
apply frame_inference
apply simp
done
qed
end
diff --git a/thys/Shivers-CFA/CPSUtils.thy b/thys/Shivers-CFA/CPSUtils.thy
--- a/thys/Shivers-CFA/CPSUtils.thy
+++ b/thys/Shivers-CFA/CPSUtils.thy
@@ -1,291 +1,290 @@
section \<open>Syntax tree helpers\<close>
theory CPSUtils
imports CPSScheme
begin
text \<open>
This theory defines the sets \<open>lambdas p\<close>, \<open>calls p\<close>, \<open>calls p\<close>, \<open>vars p\<close>, \<open>labels p\<close> and \<open>prims p\<close> as the subexpressions of the program \<open>p\<close>. Finiteness is shown for each of these sets, and some rules about how these sets relate. All these rules are proven more or less the same ways, which is very inelegant due to the nesting of the type and the shape of the derived induction rule.
It would be much nicer to start with these rules and define the set inductively. Unfortunately, that approach would make it very hard to show the finiteness of the sets in question.
\<close>
fun lambdas :: "lambda \<Rightarrow> lambda set"
and lambdasC :: "call \<Rightarrow> lambda set"
and lambdasV :: "val \<Rightarrow> lambda set"
where "lambdas (Lambda l vs c) = ({Lambda l vs c} \<union> lambdasC c)"
| "lambdasC (App l d ds) = lambdasV d \<union> \<Union> (lambdasV ` set ds)"
| "lambdasC (Let l binds c') = (\<Union>(_, y)\<in>set binds. lambdas y) \<union> lambdasC c'"
| "lambdasV (L l) = lambdas l"
| "lambdasV _ = {}"
fun calls :: "lambda \<Rightarrow> call set"
and callsC :: "call \<Rightarrow> call set"
and callsV :: "val \<Rightarrow> call set"
where "calls (Lambda l vs c) = callsC c"
| "callsC (App l d ds) = {App l d ds} \<union> callsV d \<union> (\<Union>(callsV ` (set ds)))"
| "callsC (Let l binds c') = {call.Let l binds c'} \<union> ((\<Union>(_, y)\<in>set binds. calls y) \<union> callsC c')"
| "callsV (L l) = calls l"
| "callsV _ = {}"
lemma finite_lambdas[simp]: "finite (lambdas l)" and "finite (lambdasC c)" "finite (lambdasV v)"
by (induct rule: lambdas_lambdasC_lambdasV.induct, auto)
lemma finite_calls[simp]: "finite (calls l)" and "finite (callsC c)" "finite (callsV v)"
by (induct rule: calls_callsC_callsV.induct, auto)
fun vars :: "lambda \<Rightarrow> var set"
and varsC :: "call \<Rightarrow> var set"
and varsV :: "val \<Rightarrow> var set"
where "vars (Lambda _ vs c) = set vs \<union> varsC c"
| "varsC (App _ a as) = varsV a \<union> \<Union>(varsV ` (set as))"
| "varsC (Let _ binds c') = (\<Union>(v, l)\<in>set binds. {v} \<union> vars l) \<union> varsC c'"
| "varsV (L l) = vars l"
| "varsV (R _ v) = {v}"
| "varsV _ = {}"
lemma finite_vars[simp]: "finite (vars l)" and "finite (varsC c)" "finite (varsV v)"
by (induct rule: vars_varsC_varsV.induct, auto)
fun label :: "lambda + call \<Rightarrow> label"
where "label (Inl (Lambda l _ _)) = l"
| "label (Inr (App l _ _)) = l"
| "label (Inr (Let l _ _)) = l"
fun labels :: "lambda \<Rightarrow> label set"
and labelsC :: "call \<Rightarrow> label set"
and labelsV :: "val \<Rightarrow> label set"
where "labels (Lambda l vs c) = {l} \<union> labelsC c"
| "labelsC (App l a as) = {l} \<union> labelsV a \<union> \<Union>(labelsV ` (set as))"
| "labelsC (Let l binds c') = {l} \<union> (\<Union>(v, y)\<in>set binds. labels y) \<union> labelsC c'"
| "labelsV (L l) = labels l"
| "labelsV (R l _) = {l}"
| "labelsV _ = {}"
lemma finite_labels[simp]: "finite (labels l)" and "finite (labelsC c)" "finite (labelsV v)"
by (induct rule: labels_labelsC_labelsV.induct, auto)
fun prims :: "lambda \<Rightarrow> prim set"
and primsC :: "call \<Rightarrow> prim set"
and primsV :: "val \<Rightarrow> prim set"
where "prims (Lambda _ vs c) = primsC c"
| "primsC (App _ a as) = primsV a \<union> \<Union>(primsV ` (set as))"
| "primsC (Let _ binds c') = (\<Union>(_, y)\<in>set binds. prims y) \<union> primsC c'"
| "primsV (L l) = prims l"
| "primsV (R l v) = {}"
| "primsV (P prim) = {prim}"
| "primsV (C l v) = {}"
lemma finite_prims[simp]: "finite (prims l)" and "finite (primsC c)" "finite (primsV v)"
by (induct rule: labels_labelsC_labelsV.induct, auto)
fun vals :: "lambda \<Rightarrow> val set"
and valsC :: "call \<Rightarrow> val set"
and valsV :: "val \<Rightarrow> val set"
where "vals (Lambda _ vs c) = valsC c"
| "valsC (App _ a as) = valsV a \<union> \<Union>(valsV ` (set as))"
| "valsC (Let _ binds c') = (\<Union>(_, y)\<in>set binds. vals y) \<union> valsC c'"
| "valsV (L l) = {L l} \<union> vals l"
| "valsV (R l v) = {R l v}"
| "valsV (P prim) = {P prim}"
| "valsV (C l v) = {C l v}"
lemma
fixes list2 :: "(var \<times> lambda) list" and t :: "var\<times>lambda"
shows lambdas1: "Lambda l vs c \<in> lambdas x \<Longrightarrow> c \<in> calls x"
and "Lambda l vs c \<in> lambdasC y \<Longrightarrow> c \<in> callsC y"
and "Lambda l vs c \<in> lambdasV z \<Longrightarrow> c \<in> callsV z"
and "\<forall>z\<in> set list. Lambda l vs c \<in> lambdasV z \<longrightarrow> c \<in> callsV z"
and "\<forall>x\<in> set list2. Lambda l vs c \<in> lambdas (snd x) \<longrightarrow> c \<in> calls (snd x)"
and "Lambda l vs c \<in> lambdas (snd t) \<Longrightarrow> c \<in> calls (snd t)"
apply (induct rule:mutual_lambda_call_var_inducts)
apply auto
apply (case_tac c, auto)[1]
apply (rule_tac x="((a, b), ba)" in bexI, auto)
done
lemma
shows lambdas2: "Lambda l vs c \<in> lambdas x \<Longrightarrow> l \<in> labels x"
and "Lambda l vs c \<in> lambdasC y \<Longrightarrow> l \<in> labelsC y"
and "Lambda l vs c \<in> lambdasV z \<Longrightarrow> l \<in> labelsV z"
and "\<forall>z\<in> set list. Lambda l vs c \<in> lambdasV z \<longrightarrow> l \<in> labelsV z"
and "\<forall>x\<in> set (list2 :: (var \<times> lambda) list) . Lambda l vs c \<in> lambdas (snd x) \<longrightarrow> l \<in> labels (snd x)"
and "Lambda l vs c \<in> lambdas (snd (t:: var\<times>lambda)) \<Longrightarrow> l \<in> labels (snd t)"
apply (induct rule:mutual_lambda_call_var_inducts)
apply auto
apply (rule_tac x="((a, b), ba)" in bexI, auto)
done
lemma
shows lambdas3: "Lambda l vs c \<in> lambdas x \<Longrightarrow> set vs \<subseteq> vars x"
and "Lambda l vs c \<in> lambdasC y \<Longrightarrow> set vs \<subseteq> varsC y"
and "Lambda l vs c \<in> lambdasV z \<Longrightarrow> set vs \<subseteq> varsV z"
and "\<forall>z\<in> set list. Lambda l vs c \<in> lambdasV z \<longrightarrow> set vs \<subseteq> varsV z"
and "\<forall>x\<in> set (list2 :: (var \<times> lambda) list) . Lambda l vs c \<in> lambdas (snd x) \<longrightarrow> set vs \<subseteq> vars (snd x)"
and "Lambda l vs c \<in> lambdas (snd (t:: var\<times>lambda)) \<Longrightarrow> set vs \<subseteq> vars (snd t)"
apply (induct x and y and z and list and list2 and t rule:mutual_lambda_call_var_inducts)
apply auto
apply (erule_tac x="((aa, ba), bb)" in ballE)
apply (rule_tac x="((aa, ba), bb)" in bexI, auto)
done
lemma
shows app1: "App l d ds \<in> calls x \<Longrightarrow> d \<in> vals x"
and "App l d ds \<in> callsC y \<Longrightarrow> d \<in> valsC y"
and "App l d ds \<in> callsV z \<Longrightarrow> d \<in> valsV z"
and "\<forall>z\<in> set list. App l d ds \<in> callsV z \<longrightarrow> d \<in> valsV z"
and "\<forall>x\<in> set (list2 :: (var \<times> lambda) list) . App l d ds \<in> calls (snd x) \<longrightarrow> d \<in> vals (snd x)"
and "App l d ds \<in> calls (snd (t:: var\<times>lambda)) \<Longrightarrow> d \<in> vals (snd t)"
apply (induct x and y and z and list and list2 and t rule:mutual_lambda_call_var_inducts)
apply auto
apply (case_tac d, auto)
apply (erule_tac x="((a, b), ba)" in ballE)
apply (rule_tac x="((a, b), ba)" in bexI, auto)
done
lemma
shows app2: "App l d ds \<in> calls x \<Longrightarrow> set ds \<subseteq> vals x"
and "App l d ds \<in> callsC y \<Longrightarrow> set ds \<subseteq> valsC y"
and "App l d ds \<in> callsV z \<Longrightarrow> set ds \<subseteq> valsV z"
and "\<forall>z\<in> set list. App l d ds \<in> callsV z \<longrightarrow> set ds \<subseteq> valsV z"
and "\<forall>x\<in> set (list2 :: (var \<times> lambda) list) . App l d ds \<in> calls (snd x) \<longrightarrow> set ds \<subseteq> vals (snd x)"
and "App l d ds \<in> calls (snd (t:: var\<times>lambda)) \<Longrightarrow> set ds \<subseteq> vals (snd t)"
apply (induct x and y and z and list and list2 and t rule:mutual_lambda_call_var_inducts)
apply auto
apply (case_tac x, auto)
apply (erule_tac x="((a, b), ba)" in ballE)
apply (rule_tac x="((a, b), ba)" in bexI, auto)
done
lemma
shows let1: "Let l binds c' \<in> calls x \<Longrightarrow> l \<in> labels x"
and "Let l binds c' \<in> callsC y \<Longrightarrow> l \<in> labelsC y"
and "Let l binds c' \<in> callsV z \<Longrightarrow> l \<in> labelsV z"
and "\<forall>z\<in> set list. Let l binds c' \<in> callsV z \<longrightarrow> l \<in> labelsV z"
and "\<forall>x\<in> set (list2 :: (var \<times> lambda) list) . Let l binds c' \<in> calls (snd x) \<longrightarrow> l \<in> labels (snd x)"
and "Let l binds c' \<in> calls (snd (t:: var\<times>lambda)) \<Longrightarrow> l \<in> labels (snd t)"
apply (induct x and y and z and list and list2 and t rule:mutual_lambda_call_var_inducts)
apply auto
apply (erule_tac x="((a, b), ba)" in ballE)
apply (rule_tac x="((a, b), ba)" in bexI, auto)
done
lemma
shows let2: "Let l binds c' \<in> calls x \<Longrightarrow> c' \<in> calls x"
and "Let l binds c' \<in> callsC y \<Longrightarrow> c' \<in> callsC y"
and "Let l binds c' \<in> callsV z \<Longrightarrow> c' \<in> callsV z"
and "\<forall>z\<in> set list. Let l binds c' \<in> callsV z \<longrightarrow> c' \<in> callsV z"
and "\<forall>x\<in> set (list2 :: (var \<times> lambda) list) . Let l binds c' \<in> calls (snd x) \<longrightarrow> c' \<in> calls (snd x)"
and "Let l binds c' \<in> calls (snd (t:: var\<times>lambda)) \<Longrightarrow> c' \<in> calls (snd t)"
apply (induct x and y and z and list and list2 and t rule:mutual_lambda_call_var_inducts)
apply auto
apply (case_tac c', auto)
apply (erule_tac x="((a, b), ba)" in ballE)
apply (rule_tac x="((a, b), ba)" in bexI, auto)
done
lemma
shows let3: "Let l binds c' \<in> calls x \<Longrightarrow> fst ` set binds \<subseteq> vars x"
and "Let l binds c' \<in> callsC y \<Longrightarrow> fst ` set binds \<subseteq> varsC y"
and "Let l binds c' \<in> callsV z \<Longrightarrow> fst ` set binds \<subseteq> varsV z"
and "\<forall>z\<in> set list. Let l binds c' \<in> callsV z \<longrightarrow> fst ` set binds \<subseteq> varsV z"
and "\<forall>x\<in> set (list2 :: (var \<times> lambda) list) . Let l binds c' \<in> calls (snd x) \<longrightarrow> fst ` set binds \<subseteq> vars (snd x)"
and "Let l binds c' \<in> calls (snd (t:: var\<times>lambda)) \<Longrightarrow> fst ` set binds \<subseteq> vars (snd t)"
-apply (induct x and y and z and list and list2 and t rule:mutual_lambda_call_var_inducts)
-apply auto
-apply (erule_tac x="((ab, bc), bd)" in ballE)
-apply (rule_tac x="((ab, bc), bd)" in bexI, auto)
-done
+ apply (induct x and y and z and list and list2 and t rule:mutual_lambda_call_var_inducts)
+ apply auto
+ apply fastforce
+ done
lemma
shows let4: "Let l binds c' \<in> calls x \<Longrightarrow> snd ` set binds \<subseteq> lambdas x"
and "Let l binds c' \<in> callsC y \<Longrightarrow> snd ` set binds \<subseteq> lambdasC y"
and "Let l binds c' \<in> callsV z \<Longrightarrow> snd ` set binds \<subseteq> lambdasV z"
and "\<forall>z\<in> set list. Let l binds c' \<in> callsV z \<longrightarrow> snd ` set binds \<subseteq> lambdasV z"
and "\<forall>x\<in> set (list2 :: (var \<times> lambda) list) . Let l binds c' \<in> calls (snd x) \<longrightarrow> snd ` set binds \<subseteq> lambdas (snd x)"
and "Let l binds c' \<in> calls (snd (t:: var\<times>lambda)) \<Longrightarrow> snd ` set binds \<subseteq> lambdas (snd t)"
apply (induct x and y and z and list and list2 and t rule:mutual_lambda_call_var_inducts)
apply auto
apply (rule_tac x="((a, b), ba)" in bexI, auto)
apply (case_tac ba, auto)
apply (erule_tac x="((aa, bb), bc)" in ballE)
apply (rule_tac x="((aa, bb), bc)" in bexI, auto)
done
lemma
shows vals1: "P prim \<in> vals p \<Longrightarrow> prim \<in> prims p"
and "P prim \<in> valsC y \<Longrightarrow> prim \<in> primsC y"
and "P prim \<in> valsV z \<Longrightarrow> prim \<in> primsV z"
and "\<forall>z\<in> set list. P prim \<in> valsV z \<longrightarrow> prim \<in> primsV z"
and "\<forall>x\<in> set (list2 :: (var \<times> lambda) list) . P prim \<in> vals (snd x) \<longrightarrow> prim \<in> prims (snd x)"
and "P prim \<in> vals (snd (t:: var\<times>lambda)) \<Longrightarrow> prim \<in> prims (snd t)"
apply (induct rule:mutual_lambda_call_var_inducts)
apply auto
apply (erule_tac x="((a, b), ba)" in ballE)
apply (rule_tac x="((a, b), ba)" in bexI, auto)
done
lemma
shows vals2: "R l var \<in> vals p \<Longrightarrow> var \<in> vars p"
and "R l var \<in> valsC y \<Longrightarrow> var \<in> varsC y"
and "R l var \<in> valsV z \<Longrightarrow> var \<in> varsV z"
and "\<forall>z\<in> set list. R l var \<in> valsV z \<longrightarrow> var \<in> varsV z"
and "\<forall>x\<in> set (list2 :: (var \<times> lambda) list) . R l var \<in> vals (snd x) \<longrightarrow> var \<in> vars (snd x)"
and "R l var \<in> vals (snd (t:: var\<times>lambda)) \<Longrightarrow> var \<in> vars (snd t)"
apply (induct rule:mutual_lambda_call_var_inducts)
apply auto
apply (erule_tac x="((a, b), ba)" in ballE)
apply (rule_tac x="((a, b), ba)" in bexI, auto)
done
lemma
shows vals3: "L l \<in> vals p \<Longrightarrow> l \<in> lambdas p"
and "L l \<in> valsC y \<Longrightarrow> l \<in> lambdasC y"
and "L l \<in> valsV z \<Longrightarrow> l \<in> lambdasV z"
and "\<forall>z\<in> set list. L l \<in> valsV z \<longrightarrow> l \<in> lambdasV z"
and "\<forall>x\<in> set (list2 :: (var \<times> lambda) list) . L l \<in> vals (snd x) \<longrightarrow> l \<in> lambdas (snd x)"
and "L l \<in> vals (snd (t:: var\<times>lambda)) \<Longrightarrow> l \<in> lambdas (snd t)"
apply (induct rule:mutual_lambda_call_var_inducts)
apply auto
apply (erule_tac x="((a, b), ba)" in ballE)
apply (rule_tac x="((a, b), ba)" in bexI, auto)
apply (case_tac l, auto)
done
definition nList :: "'a set => nat => 'a list set"
where "nList A n \<equiv> {l. set l \<le> A \<and> length l = n}"
lemma finite_nList[intro]:
assumes finA: "finite A"
shows "finite (nList A n)"
proof(induct n)
case 0 thus ?case by (simp add:nList_def) next
case (Suc n) hence finn: "finite (nList (A) n)" by simp
have "nList A (Suc n) = (case_prod (#)) ` (A \<times> nList A n)" (is "?lhs = ?rhs")
proof(rule subset_antisym[OF subsetI subsetI])
fix l assume "l \<in> ?lhs" thus "l \<in> ?rhs"
by (cases l, auto simp add:nList_def)
next
fix l assume "l \<in> ?rhs" thus "l \<in> ?lhs"
by (auto simp add:nList_def)
qed
thus "finite ?lhs" using finA and finn
by auto
qed
definition NList :: "'a set => nat set => 'a list set"
where "NList A N \<equiv> \<Union> n \<in> N. nList A n"
lemma finite_Nlist[intro]:
"\<lbrakk> finite A; finite N \<rbrakk> \<Longrightarrow> finite (NList A N)"
unfolding NList_def by auto
definition call_list_lengths
where "call_list_lengths p = {0,1,2,3} \<union> (\<lambda>c. case c of (App _ _ ds) \<Rightarrow> length ds | _ \<Rightarrow> 0) ` calls p"
lemma finite_call_list_lengths[simp]: "finite (call_list_lengths p)"
unfolding call_list_lengths_def by auto
end
diff --git a/thys/Stable_Matching/Contracts.thy b/thys/Stable_Matching/Contracts.thy
--- a/thys/Stable_Matching/Contracts.thy
+++ b/thys/Stable_Matching/Contracts.thy
@@ -1,2680 +1,2682 @@
(*<*)
theory Contracts
imports
Choice_Functions
"HOL-Library.Dual_Ordered_Lattice"
"HOL-Library.Bourbaki_Witt_Fixpoint"
"HOL-Library.While_Combinator"
"HOL-Library.Product_Order"
begin
(*>*)
section\<open> \citet{HatfieldMilgrom:2005}: Matching with contracts \label{sec:contracts} \<close>
text\<open>
We take the original paper on matching with contracts by
\citet{HatfieldMilgrom:2005} as our roadmap, which follows a similar
path to \citet[\S2.5]{RothSotomayor:1990}. We defer further motivation
to these texts. Our first move is to capture the scenarios described
in their {\S}I(A) (p916) in a locale.
\<close>
locale Contracts =
fixes Xd :: "'x::finite \<Rightarrow> 'd::finite"
fixes Xh :: "'x \<Rightarrow> 'h::finite"
fixes Pd :: "'d \<Rightarrow> 'x rel"
fixes Ch :: "'h \<Rightarrow> 'x cfun"
assumes Pd_linear: "\<forall>d. Linear_order (Pd d)"
assumes Pd_range: "\<forall>d. Field (Pd d) \<subseteq> {x. Xd x = d}"
assumes Ch_range: "\<forall>h. \<forall>X. Ch h X \<subseteq> {x\<in>X. Xh x = h}"
assumes Ch_singular: "\<forall>h. \<forall>X. inj_on Xd (Ch h X)"
begin
text \<open>
The set of contracts is modelled by the type @{typ "'x"}, a free type
variable that will later be interpreted by some non-empty
set. Similarly @{typ "'d"} and @{typ "'h"} track the names of doctors
and hospitals respectively. All of these are finite by virtue of
belonging to the \<open>finite\<close> type class.
We fix four constants:
\begin{itemize}
\item \<open>Xd\<close> (\<open>Xh\<close>) projects the name of the
relevant doctor (hospital) from a contract;
\item \<open>Pd\<close> maps doctors to their linear preferences over
some subset of contracts that name them (assumptions @{thm [source]
Pd_linear} and @{thm [source] Pd_range}); and
\item \<open>Ch\<close> maps hospitals to their choice functions
(\S\ref{sec:cf}), which are similarly constrained to contracts that
name them (assumption @{thm [source] Ch_range}). Moreover their
choices must name each doctor at most once, i.e., \<open>Xd\<close>
must be injective on these (assumption @{thm [source]
"Ch_singular"}).
\end{itemize}
The reader familiar with the literature will note that we do not have
a null contract (also said to represent the @{emph \<open>outside option\<close>} of
unemployment), and instead use partiality of the doctors'
preferences. This provides two benefits: firstly, \<open>Xh\<close> is
a total function here, and secondly we achieve some economy of
description when instantiating this locale as \<open>Pd\<close> only
has to rank the relevant contracts.
We note in passing that neither the doctors' nor hospitals' choice
functions are required to be decisive, unlike the standard literature
on choice functions (\S\ref{sec:cf}).
In addition to the above, the following constitute the definitions
that must be trusted for the results we prove to be meaningful.
\<close>
definition Cd :: "'d \<Rightarrow> 'x cfun" where
"Cd d \<equiv> set_option \<circ> MaxR.MaxR_opt (Pd d)"
definition CD_on :: "'d set \<Rightarrow> 'x cfun" where
"CD_on ds X = (\<Union>d\<in>ds. Cd d X)"
abbreviation CD :: "'x set \<Rightarrow> 'x set" where
"CD \<equiv> CD_on UNIV"
definition CH :: "'x cfun" where
"CH X = (\<Union>h. Ch h X)"
text\<open>
The function @{const "Cd"} constructs a choice function from the
doctor's linear preferences (see \S\ref{sec:cf-linear}). Both @{const
"CD"} and @{const "CH"} simply aggregate opinions in the obvious
way. The functions @{const "CD_on"} is parameterized with a set of
doctors to support the proofs of \S\ref{sec:contracts-vacancy-chain}.
We also define \<open>RD\<close> (\<open>Rh\<close>,
\<open>RH\<close>) to be the subsets of a given set of contracts that
are rejected by the doctors (hospitals). (The abbreviation @{const
"Rf"} is defined in \S\ref{sec:cf-rf}.)
\<close>
abbreviation (input) RD_on :: "'d set \<Rightarrow> 'x cfun" where
"RD_on ds \<equiv> Rf (CD_on ds)"
abbreviation (input) RD :: "'x cfun" where
"RD \<equiv> RD_on UNIV"
abbreviation (input) Rh :: "'h \<Rightarrow> 'x cfun" where
"Rh h \<equiv> Rf (Ch h)"
abbreviation (input) RH :: "'x cfun" where
"RH \<equiv> Rf CH"
text \<open>
A @{emph \<open>mechanism\<close>} maps doctor and hospital preferences into a match
(here a set of contracts).
\<close>
type_synonym (in -) ('d, 'h, 'x) mechanism = "('d \<Rightarrow> 'x rel) \<Rightarrow> ('h \<Rightarrow> 'x cfun) \<Rightarrow> 'd set \<Rightarrow> 'x set"
(*<*)
(* Pd *)
lemmas Pd_linear' = Pd_linear[rule_format]
lemmas Pd_range' = subsetD[OF Pd_range[rule_format], simplified, of x d] for x d
lemma Pd_refl:
assumes "x \<in> Field (Pd d)"
shows "(x, x) \<in> Pd d"
using assms Pd_linear' by (meson subset_refl underS_incl_iff)
lemma Pd_Xd:
assumes "(x, y) \<in> Pd d"
shows "Xd x = d \<and> Xd y = d"
using assms Pd_range contra_subsetD unfolding Field_def by blast
lemma Above_Pd_Xd:
assumes "x \<in> Above (Pd d) X"
shows "Xd x = d"
using assms by (blast dest: Above_Field Pd_range')
lemma AboveS_Pd_Xd:
assumes "x \<in> AboveS (Pd d) X"
shows "Xd x = d"
using assms by (blast dest: AboveS_Field Pd_range')
(* Cd *)
interpretation Cd: linear_cf "Pd d" "Cd d" for d
using Cd_def Pd_linear by unfold_locales simp_all
lemmas Cd_domain = Cd.domain
lemmas Cd_f_range = Cd.f_range
lemmas Cd_range = Cd.range
lemmas Cd_range' = Cd.range'
lemmas Rf_Cd_mono = Cd.Rf_mono_on[of UNIV, unfolded mono_on_mono]
lemmas Cd_Chernoff = Cd.Chernoff
lemmas Cd_path_independent = Cd.path_independent
lemmas Cd_iia = Cd.iia
lemmas Cd_irc = Cd.irc
lemmas Cd_lad = Cd.lad
lemmas Cd_mono = Cd.mono
lemmas Cd_greatest = Cd.greatest
lemmas Cd_preferred = Cd.preferred
lemmas Cd_singleton = Cd.singleton
lemmas Cd_union = Cd.union
lemmas Cd_idem = iia_f_idem[OF Cd.f_range[of UNIV d, folded Cd_def] Cd_iia[of UNIV], simplified] for d
lemma Cd_Xd:
shows "x \<in> Cd d X \<Longrightarrow> Xd x = d"
using Pd_range Cd_range by fastforce
lemma Cd_inj_on_Xd:
shows "inj_on Xd (Cd d X)"
by (rule inj_onI) (clarsimp simp: Cd_Xd Cd_singleton)
lemma Cd_range_disjoint:
assumes "d \<noteq> d'"
shows "Cd d A \<inter> Cd d' A = {}"
using assms Cd_range Pd_range by blast
lemma Cd_single:
assumes "x \<in> X"
assumes "inj_on Xd X"
assumes "x \<in> Field (Pd d)"
shows "x \<in> Cd d X"
using assms Pd_linear unfolding Cd_greatest greatest_def
by clarsimp (metis Pd_Xd inj_on_eq_iff subset_refl underS_incl_iff)
lemma Cd_Above:
shows "Cd d X = Above (Pd d) (X \<inter> Field (Pd d)) \<inter> X"
unfolding Cd_greatest greatest_Above Above_def by blast
(* Code generator setup. Repeats a lot of stuff. *)
definition maxR :: "'d \<Rightarrow> 'x \<Rightarrow> 'x \<Rightarrow> 'x" where
"maxR d x y = (if (x, y) \<in> Pd d then y else x)"
definition MaxR_f :: "'d \<Rightarrow> 'x \<Rightarrow> 'x option \<Rightarrow> 'x option" where
"MaxR_f d = (\<lambda>x acc. if x \<in> Field (Pd d) then Some (case acc of None \<Rightarrow> x | Some y \<Rightarrow> maxR d x y) else acc)"
lemma MaxR_maxR:
shows "MaxR.maxR (Pd d) = maxR d"
by (simp add: fun_eq_iff maxR_def Cd.maxR_code)
lemma MaxR_MaxR_f:
shows "MaxR.MaxR_f (Pd d) = MaxR_f d"
by (simp add: fun_eq_iff Cd.MaxR_f_code MaxR_f_def MaxR_maxR cong: option.case_cong)
lemmas Cd_code[code] = Cd.code[unfolded MaxR_MaxR_f]
lemma Cd_simps[simp, nitpick_simp]:
shows "Cd d {} = {}"
"Cd d (insert x A) = (if x \<in> Field (Pd d) then if Cd d A = {} then {x} else {maxR d x y |y. y \<in> Cd d A} else Cd d A)"
unfolding Cd.simps MaxR_maxR by simp_all
(* CD *)
lemma CD_on_def2:
shows "CD_on ds A = (\<Union>d\<in>ds. Cd d (A \<inter> Field (Pd d)))"
using Cd_domain unfolding CD_on_def by blast
lemma CD_on_Xd:
assumes "x \<in> CD_on ds A"
shows "Xd x \<in> ds"
using assms Cd_Xd unfolding CD_on_def by blast
lemma mem_CD_on_Cd:
shows "x \<in> CD_on ds X \<longleftrightarrow> (x \<in> Cd (Xd x) X \<and> Xd x \<in> ds)"
unfolding CD_on_def using Cd_range Cd_Xd by blast
lemma CD_on_domain:
assumes "d \<in> ds"
shows "CD_on ds A \<inter> Field (Pd d) = Cd d (A \<inter> Field (Pd d))"
unfolding CD_on_def2 using assms Cd_range by (force dest: Pd_range')
lemma CD_on_range:
shows "CD_on ds A \<subseteq> A \<inter> (\<Union>d\<in>ds. Field (Pd d))"
using Cd_range unfolding CD_on_def by blast
lemmas CD_on_range' = subsetD[OF CD_on_range]
lemma CD_on_f_range_on:
shows "f_range_on A (CD_on ds)"
by (rule f_range_onI) (meson CD_on_range Int_subset_iff)
lemma RD_on_mono:
shows "mono (RD_on ds)"
unfolding CD_on_def by (rule monoI) (auto dest: monoD[OF Rf_Cd_mono])
lemma CD_on_Chernoff:
shows "Chernoff (CD_on ds)"
using mono_on_mono RD_on_mono[of ds] Rf_mono_on_iia_on[of UNIV] Chernoff_on_iia_on by (simp add: fun_eq_iff) blast
lemma CD_on_irc:
shows "irc (CD_on ds)"
by (rule ircI) (fastforce simp: CD_on_def ircD[OF Cd_irc] simp del: Cd_simps cong: SUP_cong)
lemmas CD_on_consistency = irc_on_consistency_on[OF CD_on_irc, simplified]
lemma CD_on_path_independent:
shows "path_independent (\<lambda>X. CD_on ds X)"
using CD_on_f_range_on CD_on_Chernoff CD_on_consistency by (blast intro: path_independent_onI2)
lemma CD_on_simps:
shows "CD_on ds {} = {}"
using CD_on_range by blast
lemmas CD_on_iia = RD_on_mono[unfolded Rf_mono_iia]
lemmas CD_on_idem = iia_f_idem[OF CD_on_f_range_on CD_on_iia, simplified]
lemma CD_on_inj_on_Xd:
shows "inj_on Xd (CD_on ds X)"
unfolding CD_on_def by (rule inj_onI) (clarsimp simp: Cd_Xd Cd_singleton)
lemma CD_on_card:
shows "card (CD_on ds X) = (\<Sum>d\<in>ds. card (Cd d X))"
unfolding CD_on_def by (simp add: card_UN_disjoint Cd_range_disjoint)
lemma CD_on_closed:
assumes "inj_on Xd X"
assumes "X \<subseteq> (\<Union>d\<in>ds. Field (Pd d))"
shows "CD_on ds X = X"
using assms Cd_domain Cd_single[OF _ assms(1)] unfolding CD_on_def2 by (force dest: Cd_range')
(* Ch *)
lemmas Ch_singular' = Ch_singular[rule_format]
lemmas Ch_range' = subsetD[OF Ch_range[rule_format], simplified, of x h X] for x h X
lemma Ch_simps:
shows "Ch h {} = {}"
using Ch_range by blast
lemma Ch_range_disjoint:
assumes "h \<noteq> h'"
shows "Ch h A \<inter> Ch h' A = {}"
using assms Ch_range by blast
lemma Ch_f_range:
shows "f_range (Ch h)"
using Ch_range unfolding f_range_on_def by blast
(* CH *)
lemma CH_card:
shows "card (CH X) = (\<Sum>h\<in>UNIV. card (Ch h X))"
unfolding CH_def by (simp add: card_UN_disjoint Ch_range_disjoint)
lemma CH_simps:
shows "CH {} = {}"
unfolding CH_def by (simp add: Ch_simps)
lemma CH_range:
shows "CH A \<subseteq> A"
unfolding CH_def using Ch_range by blast
lemmas CH_range' = subsetD[OF CH_range]
lemmas CH_f_range_on = f_range_onI[OF CH_range]
lemma mem_CH_Ch:
shows "x \<in> CH X \<longleftrightarrow> x \<in> Ch (Xh x) X"
unfolding CH_def using Ch_range by blast
lemma mem_Ch_CH:
assumes "x \<in> Ch h X"
shows "x \<in> CH X"
unfolding CH_def using assms Ch_range by blast
(*>*)
text\<open>
An @{emph \<open>allocation\<close>} is a set of contracts where each names a distinct
doctor. (Hospitals can contract multiple doctors.)
\<close>
abbreviation (input) allocation :: "'x set \<Rightarrow> bool" where
"allocation \<equiv> inj_on Xd"
text\<open>
We often wish to extract a doctor's or a hospital's contract from an
@{const "allocation"}.
\<close>
definition dX :: "'x set \<Rightarrow> 'd \<Rightarrow> 'x set" where
"dX X d = {x \<in> X. Xd x = d}"
definition hX :: "'x set \<Rightarrow> 'h \<Rightarrow> 'x set" where
"hX X h = {x \<in> X. Xh x = h}"
(*<*)
lemma dX_union:
shows "dX (X \<union> Y) d = dX X d \<union> dX Y d"
unfolding dX_def by auto
lemma dX_range:
shows "\<forall>d. dX X d \<subseteq> {x. Xd x = d}"
unfolding dX_def by clarsimp
lemma dX_range':
assumes "x \<in> dX X d"
shows "x \<in> X \<and> Xd x = d"
using assms unfolding dX_def by simp
lemma dX_empty_or_singleton:
assumes "allocation X"
shows "\<forall>d. dX X d = {} \<or> (\<exists>x. dX X d = {x})"
unfolding dX_def using \<open>allocation X\<close> by (fastforce dest: inj_onD)
lemma dX_linear:
assumes "allocation X"
shows "Linear_order (dX X d \<times> dX X d)"
using spec[OF dX_empty_or_singleton[OF \<open>allocation X\<close>], where x=d] by fastforce
lemma dX_singular:
assumes "allocation X"
assumes "x \<in> X"
assumes "d = Xd x"
shows "dX X d = {x}"
using assms unfolding dX_def by (fastforce dest: inj_onD)
lemma dX_Int_Field_Pd:
assumes "dX X d \<subseteq> Field (Pd d)"
shows "X \<inter> Field (Pd d) = dX X d"
using assms unfolding dX_def by (fastforce dest: Pd_range')
lemma Cd_Above_dX:
assumes "dX X d \<subseteq> Field (Pd d)"
shows "Cd d X = Above (Pd d) (dX X d) \<inter> X"
using assms unfolding Cd_greatest greatest_Above Above_def dX_def by (auto dest: Pd_range')
(*>*)
text\<open>
@{emph \<open>Stability\<close>} is the key property we look for in a match (here a
set of contracts), and consists of two parts.
Firstly, we ask that it be @{emph \<open>individually rational\<close>}, i.e., the
contracts in the match are actually acceptable to all
participants. Note that this implies the match is an @{const
"allocation"}.
\<close>
definition individually_rational_on :: "'d set \<Rightarrow> 'x set \<Rightarrow> bool" where
"individually_rational_on ds X \<longleftrightarrow> CD_on ds X = X \<and> CH X = X"
abbreviation individually_rational :: "'x set \<Rightarrow> bool" where
"individually_rational \<equiv> individually_rational_on UNIV"
text\<open>
The second condition requires that there be no coalition of a hospital
and one or more doctors who prefer another set of contracts involving
them; the hospital strictly, the doctors weakly. Contrast this
definition with the classical one for stable marriages given in
\S\ref{sec:sotomayor}.
\<close>
definition blocking_on :: "'d set \<Rightarrow> 'x set \<Rightarrow> 'h \<Rightarrow> 'x set \<Rightarrow> bool" where
"blocking_on ds X h X' \<longleftrightarrow> X' \<noteq> Ch h X \<and> X' = Ch h (X \<union> X') \<and> X' \<subseteq> CD_on ds (X \<union> X')"
definition stable_no_blocking_on :: "'d set \<Rightarrow> 'x set \<Rightarrow> bool" where
"stable_no_blocking_on ds X \<longleftrightarrow> (\<forall>h X'. \<not>blocking_on ds X h X')"
abbreviation stable_no_blocking :: "'x set \<Rightarrow> bool" where
"stable_no_blocking \<equiv> stable_no_blocking_on UNIV"
definition stable_on :: "'d set \<Rightarrow> 'x set \<Rightarrow> bool" where
"stable_on ds X \<longleftrightarrow> individually_rational_on ds X \<and> stable_no_blocking_on ds X"
abbreviation stable :: "'x set \<Rightarrow> bool" where
"stable \<equiv> stable_on UNIV"
(*<*)
lemma stable_onI:
assumes "individually_rational_on ds X"
assumes "stable_no_blocking_on ds X"
shows "stable_on ds X"
unfolding stable_on_def using assms by blast
lemma individually_rational_onI:
assumes "CD_on ds X = X"
assumes "CH X = X"
shows "individually_rational_on ds X"
unfolding individually_rational_on_def using assms by blast
lemma individually_rational_on_CD_on:
assumes "individually_rational_on ds X"
shows "CD_on ds X = X"
using assms unfolding individually_rational_on_def by blast
lemma individually_rational_on_Cd:
assumes "individually_rational_on ds X"
shows "Cd d X = dX X d"
using individually_rational_on_CD_on[OF assms]
by (auto simp: dX_def mem_CD_on_Cd dest: Cd_range' Cd_Xd)
lemma individually_rational_on_empty:
shows "individually_rational_on ds {}"
by (simp add: CD_on_simps CH_simps individually_rational_onI)
lemma blocking_onI:
assumes "X'' \<noteq> Ch h X"
assumes "X'' = Ch h (X \<union> X'')"
assumes "\<And>x. x \<in> X'' \<Longrightarrow> x \<in> CD_on ds (X \<union> X'')"
shows "blocking_on ds X h X''"
unfolding blocking_on_def using assms by blast
lemma blocking_on_imp_not_stable:
assumes "blocking_on ds X h X''"
shows "\<not>stable_on ds X"
unfolding stable_on_def stable_no_blocking_on_def using assms by blast
lemma blocking_on_allocation:
assumes "blocking_on ds X h X''"
shows "allocation X''"
using assms unfolding blocking_on_def by (metis Ch_singular')
lemma blocking_on_Field:
assumes "blocking_on ds X h X''"
shows "dX X'' d \<subseteq> Field (Pd d)"
using assms blocking_on_allocation[OF assms] unfolding blocking_on_def dX_def
by (force simp: Pd_range' dest: CD_on_range')
lemma blocking_on_CD_on:
assumes "blocking_on ds X h X''"
shows "X'' \<subseteq> CD_on ds (X \<union> X'')"
using assms unfolding blocking_on_def by blast
lemma blocking_on_CD_on':
assumes "blocking_on ds X h X''"
assumes "x \<in> X''"
shows "x \<in> CD_on ds (X \<union> X'')"
using assms unfolding blocking_on_def by blast
lemma blocking_on_Cd:
assumes "blocking_on ds X h X''"
shows "dX X'' d \<subseteq> Cd d (X \<union> X'')"
using assms unfolding blocking_on_def by (force dest: dX_range' simp: mem_CD_on_Cd)
lemma stable_no_blocking_onI:
assumes "\<And>h X''. \<lbrakk>X'' = Ch h (X \<union> X''); X'' \<noteq> Ch h X; X'' \<subseteq> CD_on ds (X \<union> X'')\<rbrakk> \<Longrightarrow> False"
shows "stable_no_blocking_on ds X"
unfolding stable_no_blocking_on_def blocking_on_def using assms by blast
lemma stable_no_blocking_onI2:
assumes "\<And>h X''. blocking_on ds X h X'' \<Longrightarrow> False"
shows "stable_no_blocking_on ds X"
unfolding stable_no_blocking_on_def using assms by blast
lemma "stable_no_blocking_on ds UNIV"
using stable_no_blocking_onI by fastforce
lemma
assumes "stable_on ds X"
shows stable_on_CD_on: "CD_on ds X = X"
and stable_on_Xd: "x \<in> X \<Longrightarrow> Xd x \<in> ds"
and stable_on_range': "x \<in> X \<Longrightarrow> x \<in> Field (Pd (Xd x))"
and stable_on_CH: "CH X = X"
and stable_on_no_blocking_on: "stable_no_blocking_on ds X"
using assms mem_CD_on_Cd Cd_range' Pd_range'
unfolding stable_on_def individually_rational_on_def by blast+
lemma stable_on_allocation:
assumes "stable_on ds X"
shows "allocation X"
using assms unfolding stable_on_def individually_rational_on_def by (metis CD_on_inj_on_Xd)
lemma stable_on_blocking_onD:
assumes "stable_on ds X"
shows "\<lbrakk>X'' = Ch h (X \<union> X''); X'' \<subseteq> CD_on ds (X \<union> X'')\<rbrakk> \<Longrightarrow> X'' = Ch h X"
using \<open>stable_on ds X\<close> unfolding stable_on_def individually_rational_on_def stable_no_blocking_on_def blocking_on_def by blast
lemma not_stable_on_cases[consumes 1, case_names not_individually_rational not_no_blocking]:
assumes "\<not> stable_on ds X"
assumes "\<not> individually_rational_on ds X \<Longrightarrow> P"
assumes "\<not> stable_no_blocking_on ds X \<Longrightarrow> P"
shows "P"
using assms unfolding stable_on_def by blast
(*>*)
text\<open>\<close>
end
subsection\<open> Theorem~1: Existence of stable pairs \<close>
text\<open>
We proceed to define a function whose fixed points capture all stable
matches. \citet[I(B), p917]{HatfieldMilgrom:2005} provide the
following intuition:
\begin{quote}
The first theorem states that a set of contracts is stable if any
alternative contract would be rejected by some doctor or some hospital
from its suitably defined opportunity set. In the formulas below,
think of the doctors' opportunity set as @{term "XD"} and the
hospitals' opportunity set as @{term "XH"}. If @{term "X'"} is the
corresponding stable set, then @{term "XD"} must include, in addition
to @{term "X'"}, all contracts that would not be rejected by the
hospitals, and @{term "XH"} must similarly include @{term "X'"} and
all contracts that would not be rejected by the doctors. If @{term
"X'"} is stable, then every alternative contract is rejected by
somebody, so @{term "X = XH \<union> XD"} [where @{term "X"} is the
set of all contracts]. This logic is summarized in the first theorem.
\end{quote}
See also \citet[p6,\S4]{Fleiner:2003} and \citet[\S2]{Fleiner:2002},
from whom we adopt the term @{emph \<open>stable pair\<close>}.
\<close>
context Contracts
begin
definition stable_pair_on :: "'d set \<Rightarrow> 'x set \<times> 'x set \<Rightarrow> bool" where
"stable_pair_on ds = (\<lambda>(XD, XH). XD = - RH XH \<and> XH = - RD_on ds XD)"
abbreviation stable_pair :: "'x set \<times> 'x set \<Rightarrow> bool" where
"stable_pair \<equiv> stable_pair_on UNIV"
abbreviation match :: "'x set \<times> 'x set \<Rightarrow> 'x set" where
"match X \<equiv> fst X \<inter> snd X"
text \<open>
\citet[Theorem~1]{HatfieldMilgrom:2005} state that every solution
@{term "(XD, XH)"} of @{const "stable_pair"} yields a stable match
@{term "XD \<inter> XH"}, and conversely, i.e., every stable match is
the intersection of some stable pair. \citet{AygunSonmez:2012-WP2}
show that neither is the case without further restrictions on the
hospitals' choice functions @{term "Ch"}; we exhibit their
counterexample below.
Even so we can establish some properties in the present setting:
\<close>
lemma stable_pair_on_CD_on:
assumes "stable_pair_on ds XD_XH"
shows "match XD_XH = CD_on ds (fst XD_XH)"
using %invisible assms CD_on_range unfolding stable_pair_on_def split_def fst_conv snd_conv
by blast
lemma stable_pair_on_CH:
assumes "stable_pair_on ds XD_XH"
shows "match XD_XH = CH (snd XD_XH)"
using %invisible assms CH_range unfolding stable_pair_on_def split_def fst_conv snd_conv
by blast
lemma stable_pair_on_CD_on_CH:
assumes "stable_pair_on ds XD_XH"
shows "CD_on ds (fst XD_XH) = CH (snd XD_XH)"
using %invisible assms stable_pair_on_CD_on stable_pair_on_CH by blast
lemma stable_pair_on_allocation:
assumes "stable_pair_on ds XD_XH"
shows "allocation (match XD_XH)"
unfolding %invisible stable_pair_on_CD_on[OF assms] by (rule CD_on_inj_on_Xd)
(*<*)
lemma stable_pair_onI:
assumes "fst XD_XH = - RH (snd XD_XH)"
assumes "snd XD_XH = - RD_on ds (fst XD_XH)"
shows "stable_pair_on ds XD_XH"
using assms unfolding stable_pair_on_def split_def by blast
lemma stable_pair_onE:
shows "\<lbrakk>stable_pair_on ds XD_XH; \<lbrakk>- RH (snd XD_XH) = fst XD_XH; - RD_on ds (fst XD_XH) = snd XD_XH\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P"
unfolding stable_pair_on_def split_def by blast
lemma stable_pair_on_Cd:
assumes "stable_pair_on ds XD_XH"
assumes "d \<in> ds"
shows "Cd d (fst XD_XH) = match XD_XH \<inter> Field (Pd d)"
using stable_pair_on_CD_on[OF \<open>stable_pair_on ds XD_XH\<close>] CD_on_domain Cd_domain \<open>d \<in> ds\<close> by simp
lemma stable_pair_on_Cd_match:
assumes "stable_pair_on ds XD_XH"
assumes "d \<in> ds"
shows "Cd d (match XD_XH) = Cd d (fst XD_XH)"
using assms by (metis Cd_domain Cd_idem stable_pair_on_Cd)
lemma stable_pair_on_Xd:
assumes "stable_pair_on ds XD_XH"
assumes "x \<in> match XD_XH"
shows "Xd x \<in> ds"
using assms CD_on_Xd unfolding stable_pair_on_def split_def by blast
lemma stable_pair_on_match_Cd:
assumes "stable_pair_on ds XD_XH"
assumes "x \<in> match XD_XH"
shows "x \<in> Cd (Xd x) (match XD_XH)"
using assms by (metis (full_types) CD_on_def Cd_Xd UN_iff stable_pair_on_CD_on stable_pair_on_Cd_match)
(*>*)
text\<open>
We run out of steam on the following two lemmas, which are the
remaining requirements for stability.
\<close>
lemma
assumes "stable_pair_on ds XD_XH"
shows "individually_rational_on ds (match XD_XH)"
oops
lemma
assumes "stable_pair_on ds XD_XH"
shows "stable_no_blocking (match XD_XH)"
oops
text\<open>
\citet{HatfieldMilgrom:2005} also claim that the converse holds:
\<close>
lemma
assumes "stable_on ds X"
obtains XD_XH where "stable_pair_on ds XD_XH" and "X = match XD_XH"
oops
text\<open>
Again, the following counterexample shows that the @{const
substitutes} condition on @{term "Ch"} is too weak to guarantee
this. We show it holds under stronger assumptions in
\S\ref{sec:contracts-t1-converse}.
\<close>
end
subsubsection\<open> Theorem~1 does not hold \citep{AygunSonmez:2012-WP2} \label{sec:contracts-t1-counterexample} \<close>
text\<open>
The following counterexample, due to \citet[\S3:
Example~2]{AygunSonmez:2012-WP2}, comprehensively demonstrates that
\citet[Theorem~1]{HatfieldMilgrom:2005} does not hold.
We create three types: \<open>D2\<close> consists of two elements,
representing the doctors, and \<open>H\<close> is the type of the single
hospital. There are four contracts in the type \<open>X4\<close>.
\<close>
datatype D2 = D1 | D2
datatype H1 = H
datatype X4 = Xd1 | Xd1' | Xd2 | Xd2'
(*<*)
lemma D2_UNIV:
shows "UNIV = set [D1, D2]"
using D2.exhaust by auto
instantiation D2 :: enum
begin
definition "enum_class.enum = [D1, D2]"
definition "enum_class.enum_all P = (P D1 \<and> P D2)"
definition "enum_class.enum_ex P = (P D1 \<or> P D2)"
instance
by standard (simp_all add: enum_D2_def enum_all_D2_def enum_ex_D2_def D2_UNIV)
end
lemma D2_ALL:
shows "(\<forall>d. P d) = (\<forall>d\<in>{D1, D2}. P d)"
using D2_UNIV by auto
lemma D2_UNION:
shows "(\<Union>d. P d) = (\<Union>d\<in>{D1, D2}. P d)"
using D2_UNIV by auto
instance H1 :: finite
by standard (metis (full_types) H1.exhaust ex_new_if_finite finite.intros(1) finite_insert insert_subset subset_insertI)
lemma X4_UNIV:
shows "UNIV = set [Xd1, Xd1', Xd2, Xd2']"
using X4.exhaust by auto
lemmas X4_pow = subset_subseqs[OF subset_trans[OF subset_UNIV Set.equalityD1[OF X4_UNIV]]]
instance X4 :: finite
by standard (simp add: X4_UNIV)
lemma X4_ALL:
shows "(\<forall>X''. P X'') \<longleftrightarrow> (\<forall>X''\<in>set ` set (subseqs [Xd1, Xd1', Xd2, Xd2']). P X'')"
using X4_pow by blast
(*>*)
primrec X4d :: "X4 \<Rightarrow> D2" where
"X4d Xd1 = D1"
| "X4d Xd1' = D1"
| "X4d Xd2 = D2"
| "X4d Xd2' = D2"
abbreviation X4h :: "X4 \<Rightarrow> H1" where
"X4h _ \<equiv> H"
primrec PX4d :: "D2 \<Rightarrow> X4 rel" where
"PX4d D1 = linord_of_list [Xd1', Xd1]"
| "PX4d D2 = linord_of_list [Xd2, Xd2']"
function CX4h :: "H1 \<Rightarrow> X4 cfun" where
"CX4h _ {Xd1} = {Xd1}"
| "CX4h _ {Xd1'} = {Xd1'}"
| "CX4h _ {Xd2} = {Xd2}"
| "CX4h _ {Xd2'} = {Xd2'}"
| "CX4h _ {Xd1, Xd1'} = {Xd1}"
| "CX4h _ {Xd1, Xd2} = {Xd1, Xd2}"
| "CX4h _ {Xd1, Xd2'} = {Xd2'}"
| "CX4h _ {Xd1', Xd2} = {Xd1'}"
| "CX4h _ {Xd1', Xd2'} = {Xd1', Xd2'}"
| "CX4h _ {Xd2, Xd2'} = {Xd2}"
| "CX4h _ {Xd1, Xd1', Xd2} = {}"
| "CX4h _ {Xd1, Xd1', Xd2'} = {}"
| "CX4h _ {Xd1, Xd2, Xd2'} = {}"
| "CX4h _ {Xd1', Xd2, Xd2'} = {}"
| "CX4h _ {Xd1, Xd1', Xd2, Xd2'} = {}"
| "CX4h _ {} = {}"
-apply %invisible (case_tac x)
-apply (cut_tac X=b in X4_pow)
-apply auto
+ apply %invisible (case_tac x)
+ apply (cut_tac X=b in X4_pow)
+ apply simp
+ apply force
+ apply auto
done
(*<*)
termination by %invisible lexicographic_order
lemma PX4d_linear:
shows "Linear_order (PX4d d)"
by (cases d) (simp_all add: linord_of_list_Linear_order)
lemma PX4d_range:
shows "Field (PX4d d) \<subseteq> {x. X4d x = d}"
by (cases d) simp_all
lemma CX4h_range:
shows "CX4h h X \<subseteq> {x \<in> X. H = h}"
by (cases "(h, X)" rule: CX4h.cases) (auto simp: spec[OF H1.nchotomy, of h])
lemma CX4h_singular:
shows "inj_on X4d (CX4h h X)"
by (cases "(h, X)" rule: CX4h.cases) auto
(*>*)
text\<open>\<close>
interpretation StableNoDecomp: Contracts X4d X4h PX4d CX4h
using %invisible PX4d_linear PX4d_range CX4h_range CX4h_singular by unfold_locales blast+
text\<open>
There are two stable matches in this model.
\<close>
(*<*)
lemma Xd1_Xd2_stable:
shows "StableNoDecomp.stable {Xd1, Xd2}"
proof(rule StableNoDecomp.stable_onI)
show "StableNoDecomp.individually_rational {Xd1, Xd2}"
by (simp add: StableNoDecomp.individually_rational_on_def StableNoDecomp.CD_on_def
StableNoDecomp.CH_def insert_commute D2_UNION cong add: SUP_cong_simp)
show "StableNoDecomp.stable_no_blocking {Xd1, Xd2}"
apply (rule StableNoDecomp.stable_no_blocking_onI)
apply (rule_tac x="(H, X'')" in CX4h.cases)
apply (simp_all add: insert_commute)
done
qed
lemma Xd1'_Xd2'_stable:
shows "StableNoDecomp.stable {Xd1', Xd2'}"
proof(rule StableNoDecomp.stable_onI)
show "StableNoDecomp.individually_rational {Xd1', Xd2'}"
by (simp add: StableNoDecomp.individually_rational_on_def StableNoDecomp.CD_on_def
StableNoDecomp.CH_def insert_commute D2_UNION cong add: SUP_cong_simp)
show "StableNoDecomp.stable_no_blocking {Xd1', Xd2'}"
apply (rule StableNoDecomp.stable_no_blocking_onI)
apply (rule_tac x="(H, X'')" in CX4h.cases)
apply (simp_all add: insert_commute)
done
qed
(*>*)
text\<open>\<close>
lemma stable:
shows "StableNoDecomp.stable X \<longleftrightarrow> X = {Xd1, Xd2} \<or> X = {Xd1', Xd2'}"
(*<*)
(is "?lhs = ?rhs")
proof(rule iffI)
assume ?lhs then show ?rhs
using X4_pow[where X=X]
unfolding StableNoDecomp.stable_on_def StableNoDecomp.individually_rational_on_def
StableNoDecomp.stable_no_blocking_on_def StableNoDecomp.blocking_on_def
StableNoDecomp.CD_on_def StableNoDecomp.CH_def
by simp (elim disjE, simp_all add: D2_UNION X4_ALL insert_commute StableNoDecomp.maxR_def cong add: SUP_cong_simp)
next
assume ?rhs then show ?lhs
using Xd1_Xd2_stable Xd1'_Xd2'_stable by blast
qed
(*>*)
text\<open>
However neither arises from a pair \<open>XD, XH\<close> that satisfy
@{const "StableNoDecomp.stable_pair"}:
\<close>
lemma StableNoDecomp_XD_XH:
shows "StableNoDecomp.stable_pair (XD, XH) \<longleftrightarrow> (XD = {} \<and> XH = {Xd1, Xd1', Xd2, Xd2'})"
(*<*)
(is "?lhs = ?rhs")
proof(rule iffI)
note image_cong_simp [cong del] note INF_cong_simp [cong] note SUP_cong_simp [cong]
assume ?lhs then show ?rhs (* Expand the Cartesian product and check. *)
using X4_pow [of XD] X4_pow [of XH]
apply simp
apply (erule StableNoDecomp.stable_pair_onE)
apply (elim disjE)
apply (simp_all add: StableNoDecomp.CD_on_def StableNoDecomp.CH_def)
unfolding X4_UNIV [simplified]
apply (auto simp: D2_ALL D2_UNION X4_ALL insert_commute StableNoDecomp.maxR_def linord_of_list_linord_of_listP)
done
next
assume ?rhs then show ?lhs
unfolding StableNoDecomp.stable_pair_on_def using X4.exhaust by (auto simp: StableNoDecomp.CH_def)
qed
(*>*)
text\<open>\<close>
proposition
assumes "StableNoDecomp.stable_pair (XD, XH)"
shows "\<not>StableNoDecomp.stable (XD \<inter> XH)"
using %invisible assms
apply (subst (asm) StableNoDecomp_XD_XH)
apply (simp add: StableNoDecomp.stable_on_def StableNoDecomp.stable_no_blocking_on_def StableNoDecomp.blocking_on_def StableNoDecomp.individually_rational_on_empty)
apply (auto simp: StableNoDecomp.mem_CD_on_Cd MaxR_def exI[where x=D1] exI[where x=H] exI[where x="{Xd1}"])
done
text\<open>
Moreover the converse of Theorem~1 does not hold either: the single
decomposition that satisfies @{const "StableNoDecomp.stable_pair"} (@{thm
[source] "StableNoDecomp_XD_XH"}) does not yield a stable match:
\<close>
proposition
assumes "StableNoDecomp.stable X"
shows "\<not>(\<exists>XD XH. StableNoDecomp.stable_pair (XD, XH) \<and> X = XD \<inter> XH)"
using %invisible assms StableNoDecomp_XD_XH stable by fastforce
text\<open>
So there is not hope for \citet[Theorem~1]{HatfieldMilgrom:2005} as it
stands. Note that the counterexample satisfies the @{const "substitutes"}
condition (see \S\ref{sec:cf-substitutes}):
\<close>
lemma
shows "substitutes (CX4h H)"
proof %invisible (rule substitutes_onI)
fix A a b assume "b \<notin> CX4h H (insert b A)"
then show "b \<notin> CX4h H (insert a (insert b A))"
apply (case_tac [!] a)
apply (case_tac [!] b)
apply ( (rule CX4h.cases[of "(H, A)"], auto simp: insert_commute)[1] )+
done
qed
text\<open>
Therefore while @{const "substitutes"} supports the monotonicity argument
that underpins their deferred-acceptance algorithm (see
\S\ref{sec:contracts-algorithmics}), it is not enough to rescue
Theorem~1. One way forward is to constrain the hospitals'
choice functions, which we discuss in the next section.
\<close>
subsubsection\<open> Theorem 1 holds with @{emph \<open>independence of rejected contracts\<close>} \label{sec:contracts-irc} \<close>
text\<open>
\citet{AygunSonmez:2012-WP2} propose to rectify this issue by
requiring hospitals' choices to satisfy @{const "irc"}
(\S\ref{sec:cf-irc}). Reassuringly their counterexample fails to
satisfy it:
\<close>
lemma
shows "\<not>irc (CX4h H)"
by %invisible (fastforce simp: insert_commute dest: irc_onD[where a="Xd2" and B="{Xd1, Xd1'}"])
text\<open>
We adopt this hypothesis by extending the @{const "Contracts"} locale:
\<close>
locale ContractsWithIRC = Contracts +
assumes Ch_irc: "\<forall>h. irc (Ch h)"
begin
text\<open>
This property requires that \<open>Ch\<close> behave, for example, as
follows:
\<close>
lemma Ch_domain:
shows "Ch h (A \<inter> {x. Xh x = h}) = Ch h A"
using %invisible irc_on_discard[OF spec[OF Ch_irc, of h], where B="A \<inter> {x. Xh x = h}" and C="A - {x. Xh x = h}"]
by (fastforce simp: Un_Diff_Int ac_simps dest: Ch_range')
lemma %invisible CH_domain:
shows "CH A \<inter> {x. Xh x = h} = Ch h (A \<inter> {x. Xh x = h})"
unfolding CH_def using Ch_range by (auto simp: Ch_domain)
lemma %invisible stable_pair_on_Ch:
assumes "stable_pair_on ds XD_XH"
shows "Ch h (snd XD_XH) = match XD_XH \<inter> {x. Xh x = h}"
using stable_pair_on_CH[OF assms] CH_domain Ch_domain by simp
lemmas %invisible Ch_consistency = irc_on_consistency_on[OF spec[OF Ch_irc], simplified, of h] for h
lemmas Ch_irc_idem = consistency_on_f_idem[OF Ch_f_range Ch_consistency, simplified]
lemma CH_irc_idem:
shows "CH (CH A) = CH A"
unfolding %invisible CH_def by (metis CH_def CH_domain Ch_domain Ch_irc_idem)
lemma Ch_CH_irc_idem:
shows "Ch h (CH A) = Ch h A"
using %invisible CH_domain CH_irc_idem Ch_domain by blast
text\<open>
This suffices to show the left-to-right direction of Theorem~1.
\<close>
lemma stable_pair_on_individually_rational:
assumes "stable_pair_on ds XD_XH"
shows "individually_rational_on ds (match XD_XH)"
by %invisible (metis CD_on_idem CH_irc_idem stable_pair_on_CD_on stable_pair_on_CD_on_CH assms individually_rational_onI)
lemma stable_pair_on_stable_no_blocking_on:
assumes "stable_pair_on ds XD_XH"
shows "stable_no_blocking_on ds (match XD_XH)"
proof(rule stable_no_blocking_onI)
fix h X''
assume C: "X'' = Ch h (match XD_XH \<union> X'')"
assume NE: "X'' \<noteq> Ch h (match XD_XH)"
assume CD: "X'' \<subseteq> CD_on ds (match XD_XH \<union> X'')"
have "X'' \<subseteq> snd XD_XH"
proof -
from CD have "X'' \<subseteq> CD_on ds (CD_on ds (fst XD_XH) \<union> X'')" by (simp only: stable_pair_on_CD_on[OF assms])
then have "X'' \<subseteq> CD_on ds (fst XD_XH \<union> X'')"
using CD_on_path_independent unfolding path_independent_def by (simp add: Un_commute)
moreover have "fst XD_XH \<inter> CD_on ds (fst XD_XH \<union> X'') \<subseteq> CD_on ds (fst XD_XH)"
using CD_on_Chernoff unfolding Chernoff_on_def by (simp add: inf_commute)
ultimately show ?thesis using assms unfolding stable_pair_on_def split_def by blast
qed
then have "Ch h (snd XD_XH) = Ch h (Ch h (snd XD_XH) \<union> X'')"
by (force intro!: consistencyD[OF Ch_consistency] dest: Ch_range')
moreover from NE have "X'' \<noteq> Ch h (snd XD_XH)"
using stable_pair_on_CH[OF assms] CH_domain[of _ h] Ch_domain[of h] by (metis Ch_irc_idem)
ultimately have "X'' \<noteq> Ch h (match XD_XH \<union> X'')"
using stable_pair_on_CH[OF assms] CH_domain[of _ h] Ch_domain[of h]
by (metis (no_types, lifting) inf.right_idem inf_sup_distrib2)
with C show False by blast
qed
theorem stable_pair_on_stable_on:
assumes "stable_pair_on ds XD_XH"
shows "stable_on ds (match XD_XH)"
using %invisible assms stable_pair_on_allocation stable_pair_on_individually_rational stable_pair_on_stable_no_blocking_on
by (blast intro: stable_onI)
end
subsubsection\<open> The converse of Theorem~1 \label{sec:contracts-t1-converse} \<close>
text (in Contracts) \<open>
The forward direction of Theorem~1 gives us a way of finding stable
matches by computing fixed points of a function closely related to
@{const "stable_pair"} (see \S\ref{sec:contracts-algorithmics}). The
converse says that every stable match can be decomposed in this way,
which implies that the stable matches form a lattice (see also
\S\ref{sec:contracts-algorithmics}).
The following proofs assume that the hospitals' choice functions
satisfy @{const "substitutes"} and @{const "irc"}.
\<close>
context ContractsWithIRC
begin
context
fixes ds :: "'b set"
fixes X :: "'a set"
begin
text\<open>
Following \citet[Proof of Theorem~1]{HatfieldMilgrom:2005}, we
partition the set of all contracts into @{term "[X, XD_smallest - X,
XH_largest - X]"} with careful definitions of the two sets @{term
"XD_smallest"} and @{term "XH_largest"}. Specifically @{term
"XH_largest"} contains all contracts ranked at least as good as those
in @{term "X"} by the doctors, considering unemployment and
unacceptable contracts. Similarly @{term "XD_smallest"} contains those
ranked at least as poorly.
\<close>
definition XH_largest :: "'a set" where
"XH_largest =
{y. Xd y \<in> ds
\<and> y \<in> Field (Pd (Xd y))
\<and> (\<forall>x \<in> dX X (Xd y). (x, y) \<in> Pd (Xd y))}"
definition XD_smallest :: "'a set" where
"XD_smallest = - (XH_largest - X)"
context
assumes "stable_on ds X"
begin
lemma Ch_XH_largest_Field:
assumes "x \<in> Ch h XH_largest"
shows "x \<in> Field (Pd (Xd x))"
using assms unfolding XH_largest_def by (blast dest: Ch_range')
lemma Ch_XH_largest_Xd:
assumes "x \<in> Ch h XH_largest"
shows "Xd x \<in> ds"
using assms unfolding XH_largest_def by (blast dest: Ch_range')
lemma X_subseteq_XH_largest:
shows "X \<subseteq> XH_largest"
proof(rule subsetI)
fix x assume "x \<in> X"
then obtain d where "d \<in> ds" "x \<in> Cd d X" using stable_on_CD_on[OF \<open>stable_on ds X\<close>] unfolding CD_on_def by blast
with \<open>stable_on ds X\<close> show "x \<in> XH_largest"
using Pd_linear' Pd_range' Cd_range subset_Image1_Image1_iff[of "Pd d"] stable_on_allocation[of ds X]
unfolding XH_largest_def linear_order_on_def partial_order_on_def stable_on_def inj_on_def dX_def
by simp blast
qed
lemma X_subseteq_XD_smallest:
shows "X \<subseteq> XD_smallest"
unfolding XD_smallest_def by blast
lemma X_XD_smallest_XH_largest:
shows "X = XD_smallest \<inter> XH_largest"
using X_subseteq_XH_largest unfolding XD_smallest_def by blast
text\<open>
The goal of the next few lemmas is to show the constituents of @{term
"stable_pair_on ds (XD_smallest, XH_largest)"}.
Intuitively, if a doctor has a contract @{term "x"} in @{term "X"},
then all of their contracts in @{const "XH_largest"} are at least as
desirable as @{term "x"}, and so the @{const
"stable_no_blocking"} hypothesis guarantees the hospitals choose
@{term "x"} from @{const "XH_largest"}, and similarly the doctors
@{term "x"} from @{const "XD_smallest"}.
\<close>
lemma XH_largestCdXXH_largest:
assumes "x \<in> Ch h XH_largest"
shows "x \<in> Cd (Xd x) (X \<union> Ch h XH_largest)"
proof -
from assms have "(y, x) \<in> Pd (Xd x)" if "Xd y = Xd x" and "y \<in> X" for y
using that by (fastforce simp: XH_largest_def dX_def dest: Ch_range')
with Ch_XH_largest_Field[OF assms] Pd_linear Pd_range show ?thesis
using assms Ch_XH_largest_Field[OF assms]
by (clarsimp simp: Cd_greatest greatest_def)
(metis Ch_singular Pd_range' inj_onD subset_refl underS_incl_iff)
qed
lemma CH_XH_largest:
shows "CH XH_largest = X"
proof -
have "Ch h XH_largest \<subseteq> CD_on ds (X \<union> Ch h XH_largest)" for h
using XH_largestCdXXH_largest Ch_XH_largest_Xd Ch_XH_largest_Field unfolding CD_on_def by blast
from \<open>stable_on ds X\<close> have "Ch h XH_largest = Ch h X" for h
using \<open>Ch h XH_largest \<subseteq> CD_on ds (X \<union> Ch h XH_largest)\<close> X_subseteq_XH_largest
by - (erule stable_on_blocking_onD[where h=h and X''="Ch h XH_largest"];
force intro!: consistencyD[OF Ch_consistency] dest: Ch_range')
with stable_on_CH[OF \<open>stable_on ds X\<close>] show ?thesis unfolding CH_def by simp
qed
lemma Cd_XD_smallest:
assumes "d \<in> ds"
shows "Cd d (XD_smallest \<inter> Field (Pd d)) = Cd d (X \<inter> Field (Pd d))"
proof(cases "X \<inter> Field (Pd d) = {}")
case True
with Pd_range' Cd_range'[where X=X] stable_on_CD_on[OF \<open>stable_on ds X\<close>] mem_CD_on_Cd assms
have "- XH_largest \<inter> Field (Pd d) = {}"
unfolding XH_largest_def dX_def by auto blast
then show ?thesis
unfolding XD_smallest_def by (simp add: Int_Un_distrib2)
next
case False
with Pd_linear'[of d] \<open>stable_on ds X\<close> stable_on_CD_on stable_on_allocation assms
show ?thesis
unfolding XD_smallest_def order_on_defs total_on_def
by (auto 0 0 simp: Int_Un_distrib2 Cd_greatest greatest_def XH_largest_def dX_def)
(metis (mono_tags, lifting) IntI Pd_range' UnCI inj_onD)+
qed
lemma CD_on_XD_smallest:
shows "CD_on ds XD_smallest = X"
using stable_on_CD_on[OF \<open>stable_on ds X\<close>] unfolding CD_on_def2 by (simp add: Cd_XD_smallest)
theorem stable_on_stable_pair_on:
shows "stable_pair_on ds (XD_smallest, XH_largest)"
proof(rule stable_pair_onI, simp_all only: prod.sel)
from CH_XH_largest have "- RH XH_largest = - (XH_largest - X)" by blast
also from X_XD_smallest_XH_largest have "\<dots> = XD_smallest" unfolding XD_smallest_def by blast
finally show "XD_smallest = -RH XH_largest" by blast
next
from CD_on_XD_smallest have "-RD_on ds XD_smallest = -(XD_smallest - X)" by simp
also have "\<dots> = XH_largest" unfolding XD_smallest_def using X_subseteq_XH_largest by blast
finally show "XH_largest = -RD_on ds XD_smallest" by blast
qed
end
end
text\<open>
Our ultimate statement of Theorem~1 of \cite{HatfieldMilgrom:2005} ala
\citet{AygunSonmez:2012-WP2} goes as follows, bearing in mind that we
are working in the @{const "ContractsWithIRC"} locale:
\<close>
theorem T1:
shows "stable_on ds X \<longleftrightarrow> (\<exists>XD_XH. stable_pair_on ds XD_XH \<and> X = match XD_XH)"
using stable_pair_on_stable_on stable_on_stable_pair_on X_XD_smallest_XH_largest by fastforce
end
subsection\<open> Theorem~3: Algorithmics \label{sec:contracts-algorithmics} \<close>
text (in Contracts) \<open>
Having revived Theorem~1, we reformulate @{const "stable_pair"} as a
monotone (aka @{emph \<open>isotone\<close>}) function and exploit the lattice
structure of its fixed points, following \citet[{\S}II,
Theorem~3]{HatfieldMilgrom:2005}. This underpins all of their results
that we formulate here. \citet[\S2]{Fleiner:2002} provides an
intuitive gloss of these definitions.
\<close>
context Contracts
begin
definition F1 :: "'x cfun" where
"F1 X' = - RH X'"
definition F2 :: "'d set \<Rightarrow> 'x cfun" where
"F2 ds X' = - RD_on ds X'"
definition F :: "'d set \<Rightarrow> 'x set \<times> 'x set dual \<Rightarrow> 'x set \<times> 'x set dual" where
"F ds = (\<lambda>(XD, XH). (F1 (undual XH), dual (F2 ds (F1 (undual XH)))))"
text\<open>
We exploit Isabelle/HOL's ordering type classes (over the type
constructors @{typ "'a set"} and @{typ "'a \<times> 'b"}) to define
@{const "F"}. As @{const "F"} is @{const "antimono"} (where @{thm
"antimono_def"} for a lattice order \<open>\<le>\<close>) on its
second argument \<open>XH\<close>, we adopt the dual lattice order
using the type constructor @{typ "'a dual"}, where @{const "dual"} and
@{const "undual"} mediate the isomorphism on values, to satisfy
Isabelle/HOL's @{const "mono"} predicate. Note we work under the
@{const "substitutes"} hypothesis here.
Relating this function to @{const "stable_pair"} is syntactically
awkward but straightforward:
\<close>
lemma fix_F_stable_pair_on:
assumes "X = F ds X"
shows "stable_pair_on ds (map_prod id undual X)"
using %invisible assms
by (cases X) (simp add: F_def F1_def F2_def stable_pair_on_def dual_eq_iff)
lemma stable_pair_on_fix_F:
assumes "stable_pair_on ds X"
shows "map_prod id dual X = F ds (map_prod id dual X)"
using %invisible assms
unfolding F_def F1_def F2_def stable_pair_on_def split_def
by (metis fst_map_prod id_apply prod.collapse snd_map_prod undual_dual)
end
text (in Contracts) \<open>
The function @{const F} is monotonic under @{const substitutes}.
\<close>
locale ContractsWithSubstitutes = Contracts +
assumes Ch_substitutes: "\<forall>h. substitutes (Ch h)"
begin
(*<*)
lemma Rh_mono:
shows "mono (Rh h)"
using %invisible substitutes_on_Rf_mono_on[OF spec[OF Ch_substitutes]] mono_on_mono by (simp add: fun_eq_iff) blast
lemmas Ch_iia = Rh_mono[unfolded Rf_mono_iia]
lemmas Ch_Chernoff = Ch_iia[unfolded Chernoff_on_iia_on[symmetric]]
lemmas Ch_subsitutes_idem = iia_f_idem[OF Ch_f_range Ch_iia, simplified]
lemma RH_mono:
shows "mono RH"
unfolding %invisible CH_def by (rule monoI) (auto dest: monoD[OF Rh_mono])
lemmas CH_iia = RH_mono[unfolded Rf_mono_iia]
lemmas CH_Chernoff = CH_iia[unfolded Chernoff_on_iia_on[symmetric]]
lemmas CH_substitutes_idem = iia_f_idem[OF CH_f_range_on CH_iia, simplified]
(*>*)
text\<open>\<close>
lemma F1_antimono:
shows "antimono F1"
by %invisible (rule antimonoI) (auto simp: F1_def dest: Diff_mono[OF _ monoD[OF RH_mono]])
lemma F2_antimono:
shows "antimono (F2 ds)"
by %invisible (rule antimonoI) (auto simp: F2_def dest: Diff_mono[OF _ monoD[OF RD_on_mono]])
lemma F_mono:
shows "mono (F ds)"
unfolding %invisible F_def using antimonoD[OF F1_antimono] antimonoD[OF F2_antimono]
by (auto intro: monoI simp: less_eq_dual_def)
text\<open>
We define the extremal fixed points using Isabelle/HOL's least and
greatest fixed point operators:
\<close>
definition gfp_F :: "'b set \<Rightarrow> 'a set \<times> 'a set" where
"gfp_F ds = map_prod id undual (gfp (F ds))"
definition lfp_F :: "'b set \<Rightarrow> 'a set \<times> 'a set" where
"lfp_F ds = map_prod id undual (lfp (F ds))"
lemmas gfp_F_stable_pair_on = fix_F_stable_pair_on[OF gfp_unfold[OF F_mono], folded gfp_F_def]
lemmas lfp_F_stable_pair_on = fix_F_stable_pair_on[OF lfp_unfold[OF F_mono], folded lfp_F_def]
text\<open>
These last two lemmas show that the least and greatest fixed points do
satisfy @{const "stable_pair"}.
Using standard fixed-point properties, we can establish:
\<close>
lemma F2_o_F1_mono:
shows "mono (F2 ds \<circ> F1)"
by %invisible (metis F2_antimono F1_antimono antimono_def comp_apply monoI)
lemmas F2_F1_mono = F2_o_F1_mono[unfolded o_def]
lemma gfp_F_lfp_F:
shows "gfp_F ds = (F1 (lfp (F2 ds \<circ> F1)), lfp (F2 ds \<circ> F1))"
proof %invisible -
let ?F' = "dual \<circ> F2 ds \<circ> F1 \<circ> undual"
have "gfp (F ds) = (F1 (undual (gfp ?F')), gfp ?F')"
by (subst gfp_prod[OF F_mono]) (simp add: F_def o_def gfp_const)
also have "gfp ?F' = dual (lfp (F2 ds \<circ> F1))"
by (simp add: lfp_dual_gfp[OF F2_o_F1_mono, simplified o_assoc])
finally show ?thesis unfolding gfp_F_def by simp
qed
end
text\<open>
We need hospital CFs to satisfy both @{const substitutes} and @{const irc}
to relate these fixed points to stable matches.
\<close>
locale ContractsWithSubstitutesAndIRC =
ContractsWithSubstitutes + ContractsWithIRC
begin
lemmas gfp_F_stable_on = stable_pair_on_stable_on[OF gfp_F_stable_pair_on]
lemmas lfp_F_stable_on = stable_pair_on_stable_on[OF lfp_F_stable_pair_on]
end
text\<open>
\label{sec:contracts-codegen-gfp_F}
We demonstrate the effectiveness of our definitions by executing an
example due to \citet[p920]{HatfieldMilgrom:2005} using Isabelle/HOL's
code generator \citep{Haftmann-Nipkow:2010:code}. Note that, while
adequate for this toy instance, the representations used here are
hopelessly n{\"a}ive: sets are represented by lists and operations
typically traverse the entire contract space. It is feasible, with
more effort, to derive efficient algorithms; see, for instance,
\citet{Bijlsma:1991,Bulwahn-et-al:2008:imp_HOL}.
\<close>
context ContractsWithSubstitutes
begin
lemma gfp_F_code[code]:
shows "gfp_F ds = map_prod id undual (while (\<lambda>A. F ds A \<noteq> A) (F ds) top)"
using %invisible gfp_F_def gfp_while_lattice[OF F_mono] by simp
lemma lfp_F_code[code]:
shows "lfp_F ds = map_prod id undual (while (\<lambda>A. F ds A \<noteq> A) (F ds) bot)"
using %invisible lfp_F_def lfp_while_lattice[OF F_mono] by simp
end
text\<open>
There are two hospitals and two doctors.
\<close>
datatype H2 = H1 | H2
text\<open>
The contract space is simply the Cartesian product @{typ "D2 \<times>
H2"}.
\<close>
type_synonym X_D2_H2 = "D2 \<times> H2"
text\<open>
Doctor @{const "D1"} prefers \<open>H1 \<succ> H2\<close>, doctor @{const
"D2"} the same \<open>H1 \<succ> H2\<close> (but over different
contracts).
\<close>
primrec P_D2_H2_d :: "D2 \<Rightarrow> X_D2_H2 rel" where
"P_D2_H2_d D1 = linord_of_list [(D1, H1), (D1, H2)]"
| "P_D2_H2_d D2 = linord_of_list [(D2, H1), (D2, H2)]"
text\<open>
Hospital @{const "H1"} prefers \<open>{D1} \<succ> {D2} \<succ>
\<emptyset>\<close>, and hospital @{const "H2"} \<open>{D1, D2}
\<succ> {D1} \<succ> {D2} \<succ> \<emptyset>\<close>. We interpret
these constraints as follows:
\<close>
definition P_D2_H2_H1 :: "X_D2_H2 cfun" where
"P_D2_H2_H1 A = (if (D1, H1) \<in> A then {(D1, H1)} else if (D2, H1) \<in> A then {(D2, H1)} else {})"
definition P_D2_H2_H2 :: "X_D2_H2 cfun" where
"P_D2_H2_H2 A =
(if {(D1, H2), (D2, H2)} \<subseteq> A then {(D1, H2), (D2, H2)} else
if (D1, H2) \<in> A then {(D1, H2)} else if (D2, H2) \<in> A then {(D2, H2)} else {})"
primrec P_D2_H2_h :: "H2 \<Rightarrow> X_D2_H2 cfun" where
"P_D2_H2_h H1 = P_D2_H2_H1"
| "P_D2_H2_h H2 = P_D2_H2_H2"
(*<*)
lemma H2_UNIV:
shows "UNIV = set [H1, H2]"
using H2.exhaust by auto
instantiation H2 :: enum
begin
definition "enum_class.enum = [H1, H2]"
definition "enum_class.enum_all P = (P H1 \<and> P H2)"
definition "enum_class.enum_ex P = (P H1 \<or> P H2)"
instance
by standard (simp_all add: enum_H2_def enum_all_H2_def enum_ex_H2_def H2_UNIV)
end
lemma H2_ALL [simp]:
shows "(\<forall>h. P h) = (\<forall>h\<in>{H1, H2}. P h)"
using H2_UNIV by auto
lemma H2_UNION:
shows "(\<Union>h. P h) = (\<Union>h\<in>{H1, H2}. P h)"
using H2_UNIV by auto
lemma P_D2_H2_d_linear:
shows "Linear_order (P_D2_H2_d d)"
by (cases d) (simp_all add: linord_of_list_Linear_order)
lemma P_D2_H2_d_range:
shows "Field (P_D2_H2_d d) \<subseteq> {x. fst x = d}"
by (cases d) simp_all
lemma P_D2_H2_h_substitutes:
shows "substitutes (P_D2_H2_h h)"
by %invisible (cases h) (auto intro!: substitutes_onI simp: P_D2_H2_H1_def P_D2_H2_H2_def split: if_splits)
(*>*)
text\<open>
Isabelle's code generator requires us to hoist the relevant
definitions from the locale to the top-level (see the \verb!codegen!
documentation, \S7.3).
\<close>
global_interpretation P920_example:
ContractsWithSubstitutes fst snd P_D2_H2_d P_D2_H2_h
defines P920_example_gfp_F = P920_example.gfp_F
and P920_example_lfp_F = P920_example.lfp_F
and P920_example_F = P920_example.F
and P920_example_F1 = P920_example.F1
and P920_example_F2 = P920_example.F2
and P920_example_maxR = P920_example.maxR
and P920_example_MaxR_f = P920_example.MaxR_f
and P920_example_Cd = P920_example.Cd
and P920_example_CD_on = P920_example.CD_on
and P920_example_CH = P920_example.CH
using %invisible P_D2_H2_d_linear P_D2_H2_h_substitutes
by %invisible unfold_locales (simp_all, simp_all add: D2_ALL P_D2_H2_H1_def P_D2_H2_H2_def)
(*<*)
(*
Codegen hackery: avoid the CoSet constructor as some operations do not
handle it.
*)
declare UNIV_coset[code del]
declare UNIV_enum[code]
declare compl_set[code del] compl_coset[code del]
declare Compl_eq_Diff_UNIV[code]
(*
code_thms P920_example_gfp_F
export_code P920_example_gfp_F in SML module_name F file "F.sml"
value "P920_example_gfp_F UNIV"
*)
lemma P920_example_gfp_F_value:
shows "P920_example_gfp_F UNIV = ({(D1, H1), (D1, H2), (D2, H2)}, {(D1, H1), (D2, H1), (D2, H2)})"
by eval
lemma P920_example_gfp_F_match_value:
shows "P920_example.match (P920_example_gfp_F UNIV) = {(D1, H1), (D2, H2)}"
unfolding P920_example_gfp_F_value by simp
lemma P920_example_lfp_F_value:
shows "P920_example_lfp_F UNIV = ({(D1, H1), (D1, H2), (D2, H2)}, {(D1, H1), (D2, H1), (D2, H2)})"
by eval
(*>*)
text\<open>
We can now evaluate the @{const "gfp"} of @{const "P920_example.F"}
(i.e., \<open>F\<close> specialized to the above constants) using
Isabelle's \verb!value! antiquotation or \verb!eval! method. This
yields the \<open>(XD, XH)\<close> pair:
\begin{center}
@{thm (rhs) "P920_example_gfp_F_value"}
\end{center}
The stable match is therefore @{thm (rhs) "P920_example_gfp_F_match_value"}.
The @{const "lfp"} of @{const "P920_example.F"} is identical to the
@{const "gfp"}:
\begin{center}
@{thm (rhs) "P920_example_lfp_F_value"}
\end{center}
This implies that there is only one stable match in this scenario.
\<close>
subsection\<open> Theorem~4: Optimality \label{sec:contracts-optimality} \<close>
text (in ContractsWithSubstitutes) \<open>
\citet[Theorem~4]{HatfieldMilgrom:2005} assert that the greatest fixed
point @{const "gfp_F"} of @{const "F"} yields the stable match most
preferred by the doctors in the following sense:
\<close>
context Contracts
begin
definition doctor_optimal_match :: "'d set \<Rightarrow> 'x set \<Rightarrow> bool" where
"doctor_optimal_match ds Y
\<longleftrightarrow> (stable_on ds Y \<and> (\<forall>X. \<forall>x\<in>X. stable_on ds X \<longrightarrow> (\<exists>y \<in> Y. (x, y) \<in> Pd (Xd x))))"
(*<*)
lemmas doctor_optimal_matchI = iffD2[OF doctor_optimal_match_def, unfolded conj_imp_eq_imp_imp, rule_format]
lemmas doctor_optimal_match_stable_on = iffD1[OF doctor_optimal_match_def, THEN conjunct1]
lemmas doctor_optimal_match_optimal = iffD1[OF doctor_optimal_match_def, THEN conjunct2, rule_format]
lemma doctor_optimal_match_unique:
assumes "doctor_optimal_match ds X"
assumes "doctor_optimal_match ds Y"
shows "X = Y"
proof(rule iffD2[OF set_eq_iff, rule_format])
fix x
from Pd_linear'[where d="Xd x"] Pd_Xd[where d="Xd x"]
stable_on_allocation[OF doctor_optimal_match_stable_on[OF assms(1)]]
stable_on_allocation[OF doctor_optimal_match_stable_on[OF assms(2)]]
assms
show "x \<in> X \<longleftrightarrow> x \<in> Y"
unfolding doctor_optimal_match_def order_on_defs
by - (rule iffI; metis antisymD inj_on_eq_iff)
qed
(*>*)
end
text (in ContractsWithSubstitutes) \<open>
In a similar sense, @{const "lfp_F"} is the doctor-pessimal match.
We state a basic doctor-optimality result in terms of @{const
"stable_pair"} in the @{const "ContractsWithSubstitutes"} locale for
generality; we can establish @{const "doctor_optimal_match"} only
under additional constraints on hospital choice functions (see
\S\ref{sec:contracts-irc}).
\<close>
context ContractsWithSubstitutes
begin
context
fixes XD_XH :: "'a set \<times> 'a set"
fixes ds :: "'b set"
assumes "stable_pair_on ds XD_XH"
begin
lemma gfp_F_upperbound:
shows "(fst XD_XH, dual (snd XD_XH)) \<le> gfp (F ds)"
proof %invisible -
have "(fst XD_XH, dual (snd XD_XH)) = F ds (fst XD_XH, dual (snd XD_XH))"
using stable_pair_on_fix_F[OF \<open>stable_pair_on ds XD_XH\<close>] by (metis id_apply map_prod_simp prod.collapse)
then show ?thesis by (fastforce intro: gfp_upperbound)
qed
lemma XD_XH_gfp_F:
shows "fst XD_XH \<subseteq> fst (gfp_F ds)"
and "snd (gfp_F ds) \<subseteq> snd XD_XH"
using %invisible gfp_F_upperbound
unfolding gfp_F_def by (simp_all add: less_eq_dual_def less_eq_prod_def)
lemma lfp_F_upperbound:
shows "lfp (F ds) \<le> (fst XD_XH, dual (snd XD_XH))"
proof %invisible -
have "(fst XD_XH, dual (snd XD_XH)) = F ds (fst XD_XH, dual (snd XD_XH))"
using stable_pair_on_fix_F[OF \<open>stable_pair_on ds XD_XH\<close>] by (metis id_apply map_prod_simp prod.collapse)
then show ?thesis by (fastforce intro: lfp_lowerbound)
qed
lemma XD_XH_lfp_F:
shows "fst (lfp_F ds) \<subseteq> fst XD_XH"
and "snd XD_XH \<subseteq> snd (lfp_F ds)"
using %invisible lfp_F_upperbound
unfolding lfp_F_def by (simp_all add: less_eq_dual_def less_eq_prod_def)
text\<open>
We appeal to the doctors' linear preferences to show the optimality
(pessimality) of @{const "gfp_F"} (@{const "lfp_F"}) for doctors.
\<close>
theorem gfp_f_doctor_optimal:
assumes "x \<in> match XD_XH"
shows "\<exists>y \<in> match (gfp_F ds). (x, y) \<in> Pd (Xd x)"
using %invisible assms gfp_F_stable_pair_on[where ds=ds] \<open>stable_pair_on ds XD_XH\<close>
stable_pair_on_CD_on stable_pair_on_Xd Cd_Xd mem_CD_on_Cd
XD_XH_gfp_F(1) Cd_mono[where d="Xd x" and x=x and X="fst XD_XH" and Y="fst (gfp_F ds)"]
by (metis sup.absorb_iff2)
theorem lfp_f_doctor_pessimal:
assumes "x \<in> match (lfp_F ds)"
shows "\<exists>y \<in> match XD_XH. (x, y) \<in> Pd (Xd x)"
using %invisible assms lfp_F_stable_pair_on[where ds=ds] \<open>stable_pair_on ds XD_XH\<close>
stable_pair_on_CD_on stable_pair_on_Xd Cd_Xd mem_CD_on_Cd
XD_XH_lfp_F(1) Cd_mono[where d="Xd x" and x=x and X="fst (lfp_F ds)" and Y="fst XD_XH"]
by (metis sup.absorb_iff2)
end
end
theorem (in ContractsWithSubstitutesAndIRC) gfp_F_doctor_optimal_match:
shows "doctor_optimal_match ds (match (gfp_F ds))"
by %invisible (rule doctor_optimal_matchI[OF gfp_F_stable_on]) (auto simp: T1 elim: gfp_f_doctor_optimal)
text (in ContractsWithSubstitutesAndIRC) \<open>
Conversely @{const "lfp_F"} is most preferred by the hospitals in a
revealed-preference sense, and @{const "gfp_F"} least preferred. These
results depend on @{thm [source] Ch_domain} and hence the @{const
"irc"} hypothesis on hospital choice functions.
\<close>
context ContractsWithSubstitutesAndIRC
begin
theorem lfp_f_hospital_optimal:
assumes "stable_pair_on ds XD_XH"
assumes "x \<in> Ch h (match (lfp_F ds))"
shows "x \<in> Ch h (match (lfp_F ds) \<union> match XD_XH)"
proof %invisible -
from \<open>stable_pair_on ds XD_XH\<close> have "match (lfp_F ds) \<union> match XD_XH \<subseteq> snd (lfp_F ds)"
by (simp add: XD_XH_lfp_F(2) le_infI2)
with \<open>x \<in> Ch h (match (lfp_F ds))\<close> lfp_F_stable_pair_on stable_pair_on_Ch Ch_range show ?thesis
by - (rule iia_onD[OF Ch_iia[where h=h], where B="snd (lfp_F ds)", simplified]; blast)
qed
theorem gfp_f_hospital_pessimal:
assumes "stable_pair_on ds XD_XH"
assumes "x \<in> Ch h (match XD_XH)"
shows "x \<in> Ch h (match (gfp_F ds) \<union> match XD_XH)"
proof %invisible -
from \<open>stable_pair_on ds XD_XH\<close> have "match (gfp_F ds) \<union> match XD_XH \<subseteq> snd XD_XH"
by (simp add: XD_XH_gfp_F(2) le_infI2)
with assms lfp_F_stable_pair_on stable_pair_on_Ch Ch_range show ?thesis
by - (rule iia_onD[OF Ch_iia[where h=h], where B="snd XD_XH", simplified]; blast+)
qed
end
text\<open>
The general lattice-theoretic results of e.g. \citet{Fleiner:2002}
depend on the full Tarski-Knaster fixed point theorem, which is
difficult to state in the present type class-based setting. (The
theorem itself is available in the Isabelle/HOL distribution but
requires working with less convenient machinery.)
\<close>
subsection\<open> Theorem~5 does not hold \citep{HatfieldKojima:2008} \<close>
text (in Contracts) \<open>
\citet[Theorem~5]{HatfieldMilgrom:2005} claim that:
\begin{quote}
Suppose \<open>H\<close> contains at least two hospitals, which we
denote by \<open>h\<close> and \<open>h'\<close>. Further suppose that
@{term "Rh h"} is not isotone, that is, contracts are not @{const
"substitutes"} for \<open>h\<close>. Then there exist preference
orderings for the doctors in set \<open>D\<close>, a preference
ordering for a hospital \<open>h'\<close> with a single job opening
such that, regardless of the preferences of the other hospitals, no
stable set of contracts exists.
\end{quote}
\citet[Observation~1]{HatfieldKojima:2008} show this is not true:
there can be stable matches even if hospital choice functions violate
@{const "substitutes"}. This motivates looking for conditions weaker
than @{const "substitutes"} that still guarantee stable matches, a
project taken up by \citet{HatfieldKojima:2010}; see
\S\ref{sec:cop}. We omit their counterexample to this incorrect claim.
\<close>
subsection\<open> Theorem~6: ``Vacancy chain'' dynamics \label{sec:contracts-vacancy-chain} \<close>
text (in ContractsWithSubstitutesAndIRC) \<open>
\citet[II(C), p923]{HatfieldMilgrom:2005} propose a model for updating
a stable match @{term "X"} when a doctor @{term "d'"}
retires. Intuitively the contracts mentioning @{term "d'"} are
discarded and a modified algorithm run from the @{const "XH_largest"}
and @{const "XD_smallest"} sets determined from @{term "X"}. The
result is another stable match where the remaining doctors @{term "ds
- {d'}"} are (weakly) better off and the hospitals (weakly) worse off
than they were in the initial state. The proofs are essentially the
same as for optimality (\S\ref{sec:contracts-optimality}).
\<close>
context ContractsWithSubstitutesAndIRC
begin
context
fixes X :: "'a set"
fixes d' :: "'b"
fixes ds :: "'b set"
begin
text\<open>
\citeauthor{HatfieldMilgrom:2005} do not motivate why the process uses
this functional and not @{const "F"}.
\<close>
definition F' :: "'a set \<times> 'a set dual \<Rightarrow> 'a set \<times> 'a set dual" where
"F' = (\<lambda>(XD, XH). (- RH (undual XH), dual (- RD_on (ds-{d'}) XD)))"
lemma F'_apply:
"F' (XD, XH) = (- RH (undual XH), dual (- RD_on (ds - {d'}) XD))"
by (simp add: F'_def)
lemma %invisible F1'_antimono:
shows "antimono (\<lambda>XH. - RH XH)"
by %invisible (rule antimonoI) (auto simp: F1_def dest: Diff_mono[OF _ monoD[OF RH_mono]])
lemma %invisible F2'_antimono:
shows "antimono (\<lambda>XD. - RD_on (ds-{d'}) XD)"
by %invisible (rule antimonoI) (auto simp: F2_def dest: Diff_mono[OF _ monoD[OF RD_on_mono]])
lemma F'_mono:
shows "mono F'"
unfolding %invisible F'_def using antimonoD[OF F1'_antimono] antimonoD[OF F2'_antimono]
by (auto intro: monoI simp: less_eq_dual_def)
lemma fix_F'_stable_pair_on:
"stable_pair_on (ds - {d'}) (map_prod id undual A)"
if "A = F' A"
proof %invisible -
obtain x y where "A = (x, y)"
by (cases A)
with that have "F' (x, y) = (x, y)"
by simp
then have "- Rf CH (undual y) = x" and
"dual (- Rf (CD_on (ds - {d'})) x) = y"
by (simp_all only: F'_apply prod_eq_iff fst_conv snd_conv)
with \<open>A = (x, y)\<close> show ?thesis
by (simp add: stable_pair_on_def dual_eq_iff)
qed
text\<open>
We model their update process using the @{const "while"} combinator,
as we cannot connect it to the extremal fixed points as we did in
\S\ref{sec:contracts-algorithmics} because we begin computing from the
stable match @{term "X"}.
\<close>
definition F'_iter :: "'a set \<times> 'a set dual" where
"F'_iter = (while (\<lambda>A. F' A \<noteq> A) F' (XD_smallest ds X, dual (XH_largest ds X)))"
abbreviation F'_iter_match :: "'a set" where
"F'_iter_match \<equiv> match (map_prod id undual F'_iter)"
context
assumes "stable_on ds X"
begin
lemma F_start:
shows "F ds (XD_smallest ds X, dual (XH_largest ds X)) = (XD_smallest ds X, dual (XH_largest ds X))"
using %invisible CH_XH_largest[OF \<open>stable_on ds X\<close>] CD_on_XD_smallest[OF \<open>stable_on ds X\<close>] X_subseteq_XH_largest[OF \<open>stable_on ds X\<close>]
unfolding F_def F1_def F2_def XD_smallest_def by (auto simp add: dual_eq_iff)
lemma F'_start:
shows "(XD_smallest ds X, dual (XH_largest ds X)) \<le> F' (XD_smallest ds X, dual (XH_largest ds X))"
using %invisible F_start unfolding F_def F1_def F2_def F'_def
unfolding CD_on_def XD_smallest_def by (auto simp add: dual_eq_iff dual_less_eq_iff)
lemma
shows F'_iter_stable_pair_on: "stable_pair_on (ds-{d'}) (map_prod id undual F'_iter)" (is "?thesis1")
and F'_start_le_F'_iter: "(XD_smallest ds X, dual (XH_largest ds X)) \<le> F'_iter" (is "?thesis2")
proof %invisible -
obtain P where XXX: "while_option (\<lambda>A. F' A \<noteq> A) F' ((XD_smallest ds X), dual (XH_largest ds X)) = Some P"
using while_option_finite_increasing_Some[OF F'_mono _ F'_start, simplified] by blast
with while_option_stop2[OF XXX] fix_F'_stable_pair_on[where A=P]
show ?thesis1 and ?thesis2
using funpow_mono2[OF F'_mono _ order.refl F'_start, where i=0]
unfolding F'_iter_def while_def by auto
qed
lemma F'_iter_match_stable_on:
shows "stable_on (ds-{d'}) F'_iter_match"
by %invisible (rule stable_pair_on_stable_on) (metis F'_iter_stable_pair_on)
theorem F'_iter_match_doctors_weakly_better_off:
assumes "x \<in> Cd d X"
assumes "d \<noteq> d'"
shows "\<exists>y \<in> Cd d F'_iter_match. (x, y) \<in> Pd d"
proof %invisible -
from \<open>stable_on ds X\<close> assms
have "d \<in> ds" by (blast dest: Cd_Xd Cd_range' stable_on_Xd)
with assms \<open>stable_on ds X\<close> stable_on_stable_pair_on[OF \<open>stable_on ds X\<close>]
have "\<exists>y\<in>Cd d (XD_smallest ds X \<union> fst F'_iter). (x, y) \<in> Pd d"
by - (rule Cd_mono; fastforce dest: X_XD_smallest_XH_largest stable_pair_on_Cd_match)
with F'_iter_stable_pair_on F'_start_le_F'_iter
have "\<exists>y\<in>Cd d (fst F'_iter). (x, y) \<in> Pd d"
by (metis Pair_le Un_absorb1 prod.collapse)
with \<open>d \<noteq> d'\<close> \<open>d \<in> ds\<close>
show ?thesis
using stable_pair_on_Cd[OF F'_iter_stable_pair_on, symmetric, of d]
by (subst Cd_domain[symmetric]) (simp add: Cd_idem)
qed
theorem F'_iter_match_hospitals_weakly_worse_off:
assumes "x \<in> Ch h X"
shows "x \<in> Ch h (F'_iter_match \<union> X)"
proof %invisible -
from F'_iter_stable_pair_on F'_start_le_F'_iter stable_on_stable_pair_on[OF \<open>stable_on ds X\<close>] X_subseteq_XH_largest[OF \<open>stable_on ds X\<close>]
have "F'_iter_match \<union> X \<subseteq> XH_largest ds X"
by (simp add: less_eq_prod_def le_infI2 less_eq_dual_def)
with assms Ch_range \<open>stable_on ds X\<close> show ?thesis
by - (rule iia_onD[OF Ch_iia, where B="XH_largest ds X"], auto, metis Ch_CH_irc_idem CH_XH_largest)
qed
text\<open>
\citeauthor{HatfieldMilgrom:2005} observe but do not prove that
@{const "F'_iter_match"} is not necessarily doctor-optimal wrt the new
set of doctors, even if @{term "X"} was.
These results seem incomplete. One might expect that the process of
reacting to a doctor's retirement would involve considering new
entrants to the workforce and allowing the set of possible contracts
to be refined. There are also the questions of hospitals opening and
closing.
\<close>
end
end
end
subsection\<open> Theorems~8~and~9: A ``rural hospitals'' theorem \label{sec:contracts-rh} \<close>
text\<open>
Given that some hospitals are less desirable than others, the question
arises of whether there is a mechanism that can redistribute doctors
to under-resourced hospitals while retaining the stability of the
match. Roth's @{emph \<open>rural hospitals theorem\<close>}
\citep[Theorem~5.12]{RothSotomayor:1990} resolves this in the negative
by showing that each doctor and hospital signs the same number of
contracts in every stable match. In the context of contracts the
theorem relies on the further hypothesis that hospital choices satisfy
the law of aggregate demand (\S\ref{sec:cf-lad}).
\<close>
locale ContractsWithLAD = Contracts +
assumes Ch_lad: "\<forall>h. lad (Ch h)"
locale ContractsWithSubstitutesAndLAD =
ContractsWithSubstitutes + ContractsWithLAD
text\<open>
We can use results that hold under @{const "irc"} by discharging that
hypothesis against @{const "lad"} using the @{thm [source]
"lad_on_substitutes_on_irc_on"} lemma. This is the effect of the
following \<open>sublocale\<close> command:
\<close>
sublocale ContractsWithSubstitutesAndLAD < ContractsWithSubstitutesAndIRC
using Ch_range Ch_substitutes Ch_lad by unfold_locales (blast intro: lad_on_substitutes_on_irc_on f_range_onI)
context ContractsWithSubstitutesAndLAD
begin
text\<open>
The following results lead to \citet[Theorem~8]{HatfieldMilgrom:2005},
and the proofs go as they say. Again we state these with respect to an
arbitrary solution to @{const "stable_pair"}.
\<close>
context
fixes XD_XH :: "'a set \<times> 'a set"
fixes ds :: "'b set"
assumes "stable_pair_on ds XD_XH"
begin
lemma Cd_XD_gfp_F_card:
assumes "d \<in> ds"
shows "card (Cd d (fst XD_XH)) \<le> card (Cd d (fst (gfp_F ds)))"
using %invisible assms Cd_lad XD_XH_gfp_F(1)[OF \<open>stable_pair_on ds XD_XH\<close>]
unfolding lad_on_def by blast
lemma Ch_gfp_F_XH_card:
shows "card (Ch h (snd (gfp_F ds))) \<le> card (Ch h (snd XD_XH))"
using %invisible Ch_lad XD_XH_gfp_F(2)[OF \<open>stable_pair_on ds XD_XH\<close>]
unfolding lad_on_def by blast
theorem Theorem_8:
shows "d \<in> ds \<Longrightarrow> card (Cd d (fst XD_XH)) = card (Cd d (fst (gfp_F ds)))"
and "card (Ch h (snd XD_XH)) = card (Ch h (snd (gfp_F ds)))"
proof %invisible -
let ?Sum_Cd_gfp = "\<Sum>d\<in>ds. card (Cd d (fst (gfp_F ds)))"
let ?Sum_Ch_gfp = "\<Sum>h\<in>UNIV. card (Ch h (snd (gfp_F ds)))"
let ?Sum_Cd_XD = "\<Sum>d\<in>ds. card (Cd d (fst XD_XH))"
let ?Sum_Ch_XH = "\<Sum>h\<in>UNIV. card (Ch h (snd XD_XH))"
have "?Sum_Cd_gfp = ?Sum_Ch_gfp"
using stable_pair_on_CD_on_CH[OF gfp_F_stable_pair_on] CD_on_card[symmetric] CH_card[symmetric] by simp
also have "\<dots> \<le> ?Sum_Ch_XH"
using Ch_gfp_F_XH_card by (simp add: sum_mono)
also have "\<dots> = ?Sum_Cd_XD"
using stable_pair_on_CD_on_CH[OF \<open>stable_pair_on ds XD_XH\<close>] CD_on_card[symmetric] CH_card[symmetric] by simp
finally have "?Sum_Cd_XD = ?Sum_Cd_gfp"
using Cd_XD_gfp_F_card by (simp add: eq_iff sum_mono)
with Cd_XD_gfp_F_card show "d \<in> ds \<Longrightarrow> card (Cd d (fst XD_XH)) = card (Cd d (fst (gfp_F ds)))"
by (fastforce elim: sum_mono_inv)
have "?Sum_Ch_XH = ?Sum_Cd_XD"
using stable_pair_on_CD_on_CH[OF \<open>stable_pair_on ds XD_XH\<close>] CD_on_card[symmetric] CH_card[symmetric] by simp
also have "\<dots> \<le> ?Sum_Cd_gfp"
using Cd_XD_gfp_F_card by (simp add: sum_mono)
also have "\<dots> = ?Sum_Ch_gfp"
using stable_pair_on_CD_on_CH[OF gfp_F_stable_pair_on] CD_on_card[symmetric] CH_card[symmetric] by simp
finally have "?Sum_Ch_gfp = ?Sum_Ch_XH"
using Ch_gfp_F_XH_card by (simp add: eq_iff sum_mono)
with Ch_gfp_F_XH_card show "card (Ch h (snd XD_XH)) = card (Ch h (snd (gfp_F ds)))"
by (fastforce elim: sym[OF sum_mono_inv])
qed
end
text\<open>
Their result may be more easily understood when phrased in terms of
arbitrary stable matches:
\<close>
corollary rural_hospitals_theorem:
assumes "stable_on ds X"
assumes "stable_on ds Y"
shows "d \<in> ds \<Longrightarrow> card (Cd d X) = card (Cd d Y)"
and "card (Ch h X) = card (Ch h Y)"
using %invisible assms T1[of ds X] T1[of ds Y] Theorem_8 stable_pair_on_Cd_match Ch_CH_irc_idem stable_pair_on_CH
by force+
end
text\<open>
\citet[Theorem~9]{HatfieldMilgrom:2005} show that without @{const
"lad"}, the rural hospitals theorem does not hold. Their proof does
not seem to justify the theorem as stated (for instance, the contracts
\<open>x'\<close>, \<open>y'\<close> and \<open>z'\<close> need not
exist), and so we instead simply provide a counterexample (discovered
by \verb!nitpick!) to the same effect.
\<close>
lemma (in ContractsWithSubstitutesAndIRC) Theorem_9_counterexample:
assumes "stable_on ds Y"
assumes "stable_on ds Z"
shows "card (Ch h Y) = card (Ch h Z)"
oops
datatype X3 = Xd1 | Xd1' | Xd2
(*<*)
lemma X3_UNIV:
shows "UNIV = set [Xd1, Xd1', Xd2]"
using X3.exhaust by auto
lemmas X3_pow = subset_subseqs[OF subset_trans[OF subset_UNIV Set.equalityD1[OF X3_UNIV]]]
instance X3 :: finite
by standard (simp add: X3_UNIV)
lemma X3_all_pow:
shows "(\<forall>X''. P X'') \<longleftrightarrow> (\<forall>X''\<in>set ` set (subseqs [Xd1, Xd1', Xd2]). P X'')"
using X3_pow by blast
(*>*)
primrec X3d :: "X3 \<Rightarrow> D2" where
"X3d Xd1 = D1"
| "X3d Xd1' = D1"
| "X3d Xd2 = D2"
abbreviation X3h :: "X3 \<Rightarrow> H1" where
"X3h _ \<equiv> H"
primrec PX3d :: "D2 \<Rightarrow> X3 rel" where
"PX3d D1 = linord_of_list [Xd1, Xd1']"
| "PX3d D2 = linord_of_list [Xd2]"
function CX3h :: "H1 \<Rightarrow> X3 set \<Rightarrow> X3 set" where
"CX3h _ {Xd1} = {Xd1}"
| "CX3h _ {Xd1'} = {Xd1'}"
| "CX3h _ {Xd2} = {Xd2}"
| "CX3h _ {Xd1, Xd1'} = {Xd1'}"
| "CX3h _ {Xd1, Xd2} = {Xd1, Xd2}"
| "CX3h _ {Xd1', Xd2} = {Xd1'}"
| "CX3h _ {Xd1, Xd1', Xd2} = {Xd1'}"
| "CX3h _ {} = {}"
apply %invisible (case_tac x)
apply (cut_tac X=b in X3_pow)
apply auto
done
(*<*)
termination by lexicographic_order
lemma PX3d_linear:
shows "Linear_order (PX3d d)"
by (cases d) (simp_all add: linord_of_list_Linear_order)
lemma PX3d_range:
shows "Field (PX3d d) \<subseteq> {x. X3d x = d}"
by (cases d) simp_all
lemma CX3h_range:
shows "CX3h h X \<subseteq> {x\<in>X. X3h x = h}"
by (cases "(h, X)" rule: CX3h.cases; simp; metis (mono_tags) H1.exhaust)
lemma CX3h_singular:
shows "inj_on X3d (CX3h h X)"
by (cases "(h, X)" rule: CX3h.cases) auto
lemma CX3h_substitutes:
shows "substitutes (CX3h h)"
apply (rule substitutes_onI)
apply (cases h)
apply (cut_tac X=B in X3_pow)
apply (case_tac b)
apply (case_tac [!] a)
apply (auto simp: insert_commute)
done
lemma CX3h_irc:
shows "irc (CX3h h)"
apply (rule ircI)
apply (cases h)
apply (cut_tac X=B in X3_pow)
apply (case_tac a)
apply (auto simp: insert_commute)
done
(*>*)
interpretation Theorem_9: ContractsWithSubstitutesAndIRC X3d X3h PX3d CX3h
using %invisible PX3d_linear PX3d_range CX3h_range CX3h_singular CX3h_substitutes CX3h_irc
by unfold_locales blast+
lemma Theorem_9_stable_Xd1':
shows "Theorem_9.stable_on UNIV {Xd1'}"
proof %invisible (rule Theorem_9.stable_onI)
note image_cong_simp [cong del] note INF_cong_simp [cong] note SUP_cong_simp [cong]
show "Theorem_9.individually_rational_on UNIV {Xd1'}"
by (rule Theorem_9.individually_rational_onI)
(simp_all add: D2_UNION Theorem_9.CD_on_def Theorem_9.CH_def)
show "Theorem_9.stable_no_blocking_on UNIV {Xd1'}"
apply (rule Theorem_9.stable_no_blocking_onI)
apply (case_tac h)
apply (cut_tac X=X'' in X3_pow)
apply simp
apply safe
apply (simp_all add: insert_commute)
done
qed
lemma Theorem_9_stable_Xd1_Xd2:
shows "Theorem_9.stable_on UNIV {Xd1, Xd2}"
proof %invisible (rule Theorem_9.stable_onI)
note image_cong_simp [cong del] note INF_cong_simp [cong] note SUP_cong_simp [cong]
show "Theorem_9.individually_rational_on UNIV {Xd1, Xd2}"
by (rule Theorem_9.individually_rational_onI)
(simp_all add: D2_UNION Theorem_9.CD_on_def Theorem_9.CH_def insert_commute)
show "Theorem_9.stable_no_blocking_on UNIV {Xd1, Xd2}"
apply (rule Theorem_9.stable_no_blocking_onI)
apply (case_tac h)
apply (cut_tac X=X'' in X3_pow)
apply simp
apply safe
apply (simp_all add: D2_UNION Theorem_9.CD_on_def Theorem_9.maxR_def linord_of_list_linord_of_listP insert_commute)
done
qed
text \<open>
This violates the rural hospitals theorem:
\<close>
theorem
shows "card (Theorem_9.CH {Xd1'}) \<noteq> card (Theorem_9.CH {Xd1, Xd2})"
using %invisible Theorem_9_stable_Xd1' Theorem_9_stable_Xd1_Xd2 Theorem_9.stable_on_CH by simp
text\<open>
{\ldots}which is attributed to the failure of the hospitals' choice
functions to satisfy @{const "lad"}:
\<close>
lemma CX3h_not_lad:
shows "\<not>lad (CX3h h)"
unfolding %invisible lad_on_def
apply (cases h)
apply clarsimp
apply (rule exI[where x="{Xd1, Xd1', Xd2}"])
apply (rule exI[where x="{Xd1, Xd2}"])
apply simp
done
text\<open>
\citet{CiupanHatfieldKominers:2016} discuss an alternative approach to
this result in a marriage market.
\<close>
subsection\<open> Theorems~15 and 16: Cumulative Offer Processes \label{sec:contracts-cop} \<close>
text\<open>
The goal of \citet[{\S}V]{HatfieldMilgrom:2005} is to connect this
theory of contracts with matching to earlier work on auctions by the
first of the authors, in particular by eliminating the @{const
"substitutes"} hypothesis. They do so by defining a @{emph \<open>cumulative
offer process\<close>} (COP):
\<close>
context Contracts
begin
definition cop_F_HM :: "'d set \<Rightarrow> 'x set \<times> 'x set \<Rightarrow> 'x set \<times> 'x set" where
"cop_F_HM ds = (\<lambda>(XD, XH). (- RH XH, XH \<union> CD_on ds (- RH XH)))"
text\<open>
Intuitively all of the doctors simultaneously offer their most
preferred contracts that have yet to be rejected by the hospitals, and
the hospitals choose amongst these and all that have been offered
previously. Asking hospital choice functions to satisfy the @{const
"substitutes"} condition effectively forces hospitals to consider only
the contracts they have previously not rejected.
This definition is neither monotonic nor increasing (i.e., it is not
the case that @{term "\<forall>x. x \<le> cop_F_HM ds x"}). We rectify
this by focusing on the second part of the definition.
\<close>
definition cop_F :: "'d set \<Rightarrow> 'x set \<Rightarrow> 'x set" where
"cop_F ds XH = XH \<union> CD_on ds (- RH XH)"
lemma cop_F_HM_cop_F:
shows "cop_F_HM ds XD_XH = (- RH (snd XD_XH), cop_F ds (snd XD_XH))"
unfolding cop_F_HM_def cop_F_def split_def by simp
lemma cop_F_increasing:
shows "x \<le> cop_F ds x"
unfolding %invisible cop_F_def by simp
text\<open>
We have the following straightforward case distinction principles:
\<close>
lemma cop_F_cases:
assumes "x \<in> cop_F ds fp"
obtains (fp) "x \<in> fp" | (CD_on) "x \<in> CD_on ds (-RH fp) - fp"
using assms unfolding cop_F_def by blast
lemma CH_cop_F_cases:
assumes "x \<in> CH (cop_F ds fp)"
obtains (CH) "x \<in> CH fp" | (RH_fp) "x \<in> RH fp" | (CD_on) "x \<in> CD_on ds (-RH fp) - fp"
using assms CH_range cop_F_def by auto
text\<open>
The existence of fixed points for our earlier definitions
(\S\ref{sec:contracts-algorithmics}) was guaranteed by the
Tarski-Knaster theorem, which relies on the monotonicity of the
defining functional. As @{const "cop_F"} lacks this property, we
appeal instead to the Bourbaki-Witt theorem for increasing
functions.
\<close>
interpretation COP: bourbaki_witt_fixpoint Sup "{(x, y). x \<le> y}" "cop_F ds" for ds
by %invisible (rule bourbaki_witt_fixpoint_complete_latticeI[OF cop_F_increasing])
definition fp_cop_F :: "'d set \<Rightarrow> 'x set" where
"fp_cop_F ds = COP.fixp_above ds {}"
abbreviation "cop ds \<equiv> CH (fp_cop_F ds)"
(*<*)
lemmas fp_cop_F_unfold = COP.fixp_above_unfold[where a="{}", folded fp_cop_F_def, simplified Field_def, simplified]
lemmas fp_cop_F_code = COP.fixp_above_conv_while[where a="{}", folded fp_cop_F_def, simplified Field_def, simplified]
(*>*)
text\<open>
Given that the set of contracts is finite, we avoid continuity and
admissibility issues; we have the following straightforward induction
principle:
\<close>
lemma fp_cop_F_induct[case_names base step]:
assumes "P {}"
assumes "\<And>fp. P fp \<Longrightarrow> P (cop_F ds fp)"
shows "P (fp_cop_F ds)"
using %invisible assms
by (induct rule: COP.fixp_above_induct[where a="{}", folded fp_cop_F_def])
(fastforce intro: admissible_chfin)+
text\<open>
An alternative is to use the @{const "while"} combinator, which is
equivalent to the above by @{thm [source] COP.fixp_above_conv_while}.
In any case, invariant reasoning is essential to verifying the
properties of the COP, no matter how we phrase it. We develop a small
program logic to ease the reuse of the invariants we
prove.
\<close>
definition
valid :: "'d set \<Rightarrow> ('d set \<Rightarrow> 'x set \<Rightarrow> bool) \<Rightarrow> ('d set \<Rightarrow> 'x set \<Rightarrow> bool) \<Rightarrow> bool"
where
"valid ds P Q = (Q ds {} \<and> (\<forall>fp. P ds fp \<and> Q ds fp \<longrightarrow> Q ds (cop_F ds fp)))"
abbreviation
invariant :: "'d set \<Rightarrow> ('d set \<Rightarrow> 'x set \<Rightarrow> bool) \<Rightarrow> bool"
where
"invariant ds P \<equiv> valid ds (\<lambda>_ _. True) P"
text\<open>
Intuitively @{term "valid ds P Q"} asserts that the COP satisfies
@{term "Q"} assuming that it satisfies @{term "P"}. This allows us to
decompose our invariant proofs. By setting the precondition to @{term
"True"}, @{term "invariant ds P"} captures the proof obligations of
@{term "fp_cop_F_induct"} exactly.
The following lemmas ease the syntactic manipulation of these facts.
\<close>
lemma validI[case_names base step]:
assumes "Q ds {}"
assumes "\<And>fp. \<lbrakk>P ds fp; Q ds fp\<rbrakk> \<Longrightarrow> Q ds (cop_F ds fp)"
shows "valid ds P Q"
using %invisible assms unfolding valid_def by blast
lemma invariant_cop_FD:
assumes "invariant ds P"
assumes "P ds fp"
shows "P ds (cop_F ds fp)"
using %invisible assms unfolding valid_def by blast
lemma invariantD:
assumes "invariant ds P"
shows "P ds (fp_cop_F ds)"
using %invisible assms fp_cop_F_induct unfolding valid_def by blast
lemma valid_pre:
assumes "valid ds P' Q"
assumes "\<And>fp. P ds fp \<Longrightarrow> P' ds fp"
shows "valid ds P Q"
using %invisible assms unfolding valid_def by blast
lemma valid_invariant:
assumes "valid ds P Q"
assumes "invariant ds P"
shows "invariant ds (\<lambda> ds fp. P ds fp \<and> Q ds fp)"
using %invisible assms unfolding valid_def by blast
lemma valid_conj:
assumes "valid ds (\<lambda>ds fp. R ds fp \<and> P ds fp \<and> Q ds fp) P"
assumes "valid ds (\<lambda>ds fp. R ds fp \<and> P ds fp \<and> Q ds fp) Q"
shows "valid ds R (\<lambda> ds fp. P ds fp \<and> Q ds fp)"
using %invisible assms unfolding valid_def by blast
end
text (in ContractsWithSubstitutes) \<open>
\citet[Theorem~15]{HatfieldMilgrom:2005} assert that @{const
"fp_cop_F"} is equivalent to the doctor-offering algorithm @{const
"gfp_F"}, assuming @{const "substitutes"}. (Note that the fixed points
generated by increasing functions do not necessarily form a lattice,
so there is not necessarily a hospital-optimal match, and indeed in
general these do not exist.) Our proof is eased by the decomposition
lemma @{thm [source] gfp_F_lfp_F} and the standard properties of fixed
points in a lattice.
\<close>
context ContractsWithSubstitutes
begin
lemma lfp_F2_o_F1_fp_cop_F:
shows "lfp (F2 ds \<circ> F1) = fp_cop_F ds"
proof(rule antisym)
have "(F2 ds \<circ> F1) (fp_cop_F ds) \<subseteq> cop_F ds (fp_cop_F ds)"
by (clarsimp simp: F2_def F1_def cop_F_def)
then show "lfp (F2 ds \<circ> F1) \<subseteq> fp_cop_F ds"
by (simp add: lfp_lowerbound fp_cop_F_unfold[symmetric])
next
show "fp_cop_F ds \<subseteq> lfp (F2 ds \<circ> F1)"
proof(induct rule: fp_cop_F_induct)
case base then show ?case by simp
next
case (step fp) note IH = \<open>fp \<subseteq> lfp (F2 ds \<circ> F1)\<close>
then have "CD_on ds (- RH fp) \<subseteq> lfp (F2 ds \<circ> F1)"
by (subst lfp_unfold[OF F2_o_F1_mono])
(metis (no_types, lifting) Compl_Diff_eq F1_antimono F2_antimono F1_def F2_def Un_subset_iff antimonoD comp_apply)
with IH show ?case
unfolding cop_F_def by blast
qed
qed
theorem Theorem_15:
shows "gfp_F ds = (- RH (fp_cop_F ds), fp_cop_F ds)"
using lfp_F2_o_F1_fp_cop_F unfolding gfp_F_lfp_F F1_def by simp
theorem Theorem_15_match:
shows "match (gfp_F ds) = CH (fp_cop_F ds)"
using Theorem_15 by (fastforce dest: subsetD[OF CH_range])
end
text\<open>
\label{sec:contracts-codegen-fp_cop_F}
With some auxiliary definitions, we can evaluate the COP on the
example from \S\ref{sec:contracts-codegen-gfp_F}.
\<close>
(*<*)
definition "P920_example_cop_F = P920_example.cop_F"
definition "P920_example_fp_cop_F = P920_example.fp_cop_F"
lemmas P920_example_cop_F_code[code] = P920_example.cop_F_def[folded P920_example_cop_F_def]
lemmas P920_example_fp_cop_F_code[code] = P920_example.fp_cop_F_code[folded P920_example_fp_cop_F_def P920_example_cop_F_def]
(*>*)
lemma P920_example_fp_cop_F_value:
shows "P920_example_CH (P920_example_fp_cop_F UNIV) = {(D1, H1), (D2, H2)}"
by eval
text\<open>
\citet[Theorem~16]{HatfieldMilgrom:2005} assert that this process
yields a stable match when we have a single hospital (now called an
auctioneer) with unrestricted preferences. As before, this holds
provided the auctioneer's preferences satisfy @{const "irc"}.
We begin by establishing two obvious invariants of the COP that
hold in general.
\<close>
context Contracts
begin
lemma %invisible CH_Ch_singular:
assumes "(UNIV::'h set) = {h}"
shows "CH A = Ch h A"
unfolding CH_def using assms by auto
definition cop_F_range_inv :: "'d set \<Rightarrow> 'x set \<Rightarrow> bool" where
"cop_F_range_inv ds fp \<longleftrightarrow> (\<forall>x\<in>fp. x \<in> Field (Pd (Xd x)) \<and> Xd x \<in> ds)"
definition cop_F_closed_inv :: "'d set \<Rightarrow> 'x set \<Rightarrow> bool" where
"cop_F_closed_inv ds fp \<longleftrightarrow> (\<forall>x\<in>fp. above (Pd (Xd x)) x \<subseteq> fp)"
text\<open>
The first, @{const "cop_F_range_inv"}, simply states that the result
of the COP respects the structural conditions for doctors. The second
@{const "cop_F_closed_inv"} states that the COP is upwards-closed with
respect to the doctors' preferences.
\<close>
lemma cop_F_range_inv:
shows "invariant ds cop_F_range_inv"
unfolding valid_def cop_F_range_inv_def cop_F_def by (fastforce simp: mem_CD_on_Cd dest: Cd_range')
lemma cop_F_closed_inv:
shows "invariant ds cop_F_closed_inv"
unfolding valid_def cop_F_closed_inv_def cop_F_def above_def
by (clarsimp simp: subset_iff) (metis Cd_preferred ComplI Un_upper1 mem_CD_on_Cd subsetCE)
lemmas fp_cop_F_range_inv = invariantD[OF cop_F_range_inv]
lemmas fp_cop_F_range_inv' = fp_cop_F_range_inv[unfolded cop_F_range_inv_def, rule_format]
lemmas fp_cop_F_closed_inv = invariantD[OF cop_F_closed_inv]
lemmas fp_cop_F_closed_inv' = subsetD[OF bspec[OF invariantD[OF cop_F_closed_inv, unfolded cop_F_closed_inv_def, simplified]]]
text\<open>
The only challenge in showing that the COP yields a stable match is in
establishing @{const "stable_no_blocking_on"}. Our key lemma states
that that if @{const "CH"} rejects all contracts for doctor
\<open>d\<close> in @{const "fp_cop_F"}, then all contracts for
\<open>d\<close> are in @{const "fp_cop_F"}.
\<close>
lemma cop_F_RH:
assumes "d \<in> ds"
assumes "x \<in> Field (Pd d)"
assumes "aboveS (Pd d) x \<subseteq> RH fp"
shows "x \<in> cop_F ds fp"
using %invisible assms Pd_linear unfolding cop_F_def
by (clarsimp simp: mem_CD_on_Cd Cd_greatest greatest_def aboveS_def order_on_defs total_on_def subset_iff)
(metis Compl_Diff_eq Compl_iff Diff_iff IntE Pd_Xd refl_onD)
lemma fp_cop_F_all:
assumes "d \<in> ds"
assumes "d \<notin> Xd ` CH (fp_cop_F ds)"
shows "Field (Pd d) \<subseteq> fp_cop_F ds"
proof %invisible (rule subsetI)
fix x assume "x \<in> Field (Pd d)"
from spec[OF Pd_linear] this finite[of "Pd d"] show "x \<in> fp_cop_F ds"
proof(induct rule: finite_Linear_order_induct)
case (step x)
with assms Pd_range Pd_Xd cop_F_RH[of d ds _ "fp_cop_F ds", unfolded fp_cop_F_unfold[symmetric]]
show ?case unfolding aboveS_def by (fastforce iff: image_iff)
qed
qed
text\<open>
\citet{AygunSonmez:2012-WP2} observe that any blocking contract must
be weakly preferred by its doctor to anything in the outcome of the
@{const "fp_cop_F"}:
\<close>
lemma fp_cop_F_preferred:
assumes "y \<in> CD_on ds (CH (fp_cop_F ds) \<union> X'')"
assumes "x \<in> CH (fp_cop_F ds)"
assumes "Xd x = Xd y"
shows "(x, y) \<in> Pd (Xd x)"
using %invisible assms fp_cop_F_range_inv'[OF CH_range'[OF assms(2)]] Pd_Xd Pd_linear
by (clarsimp simp: CD_on_def Cd_greatest greatest_def) (metis Int_iff Un_iff subset_refl underS_incl_iff)
text\<open>
The headline lemma cobbles these results together.
\<close>
lemma X''_closed:
assumes "X'' \<subseteq> CD_on ds (CH (fp_cop_F ds) \<union> X'')"
shows "X'' \<subseteq> fp_cop_F ds"
proof(rule subsetI)
fix x assume "x \<in> X''"
show "x \<in> fp_cop_F ds"
proof(cases "Xd x \<in> Xd ` CH (fp_cop_F ds)")
case True
then obtain y where "Xd y = Xd x" and "y \<in> CH (fp_cop_F ds)" by clarsimp
with assms \<open>x \<in> X''\<close> show ?thesis
using CH_range fp_cop_F_closed_inv' fp_cop_F_preferred unfolding above_def by blast
next
case False with assms \<open>x \<in> X''\<close> show ?thesis
by (meson Cd_range' IntD2 fp_cop_F_all mem_CD_on_Cd rev_subsetD)
qed
qed
text (in Contracts) \<open>
The @{const "irc"} constraint on the auctioneer's preferences is
needed for @{const "stable_no_blocking"} and their part of @{const
"individually_rational"}.
\<close>
end
context ContractsWithIRC
begin
lemma cop_stable_no_blocking_on:
shows "stable_no_blocking_on ds (cop ds)"
proof(rule stable_no_blocking_onI)
fix h X''
assume C: "X'' = Ch h (CH (fp_cop_F ds) \<union> X'')"
assume NE: "X'' \<noteq> Ch h (CH (fp_cop_F ds))"
assume CD: "X'' \<subseteq> CD_on ds (CH (fp_cop_F ds) \<union> X'')"
from CD have "X'' \<subseteq> fp_cop_F ds" by (rule X''_closed)
then have X: "CH (fp_cop_F ds) \<union> X'' \<subseteq> fp_cop_F ds" using CH_range by simp
from C NE Ch_CH_irc_idem[of h] show False
using consistency_onD[OF Ch_consistency _ X] CH_domain Ch_domain by blast
qed
theorem Theorem_16:
assumes h: "(UNIV::'c set) = {h}"
shows "stable_on ds (cop ds)" (is "stable_on ds ?fp")
proof(rule stable_onI)
show "individually_rational_on ds ?fp"
proof(rule individually_rational_onI)
from h have "allocation ?fp" by (simp add: Ch_singular CH_Ch_singular)
then show "CD_on ds ?fp = ?fp"
by (rule CD_on_closed) (blast dest: CH_range' fp_cop_F_range_inv')
show "CH (CH (fp_cop_F ds)) = CH (fp_cop_F ds)" by (simp add: CH_irc_idem)
qed
show "stable_no_blocking_on ds ?fp" by (rule cop_stable_no_blocking_on)
qed
end
subsection\<open> Concluding remarks \<close>
text\<open>
From \citet{HatfieldMilgrom:2005}, we have not shown Theorems~2, 7, 13
and~14, all of which are intended to position their results against
prior work in this space. We delay establishing their strategic
results (Theorems~10, 11 and~12) to \S\ref{sec:strategic}, after we
have developed more useful invariants for the COP.
By assuming \isa{irc}, \citet{AygunSonmez:2012-WP2} are essentially
trading on Plott's path independence condition
(\S\ref{sec:cf-path-independence}), as observed by
\citet{ChambersYenmez:2013}. The latter show that these results
generalize naturally to many-to-many matches, where doctors also use
path-independent choice functions; see also \citet{Fleiner:2003}.
For many applications, however, @{const "substitutes"} proves to be
too strong a condition. The COP of \S\ref{sec:contracts-cop} provides
a way forward, as we discuss in the next section.
\<close>
(*<*)
end
(*>*)
diff --git a/thys/Tail_Recursive_Functions/CaseStudy2.thy b/thys/Tail_Recursive_Functions/CaseStudy2.thy
--- a/thys/Tail_Recursive_Functions/CaseStudy2.thy
+++ b/thys/Tail_Recursive_Functions/CaseStudy2.thy
@@ -1,797 +1,802 @@
(* Title: A General Method for the Proof of Theorems on Tail-recursive Functions
Author: Pasquale Noce
Security Certification Specialist at Arjo Systems - Gep S.p.A.
pasquale dot noce dot lavoro at gmail dot com
pasquale dot noce at arjowiggins-it dot com
*)
section "Case study 2"
theory CaseStudy2
imports Main "HOL-Library.Multiset"
begin
text \<open>
\null
In the second case study, the problem will be examined of defining a function
\<open>t_ins\<close> performing item insertion into binary search trees (admitting value
repetitions) of elements of a linear order, and then proving the correctness of
this definition, i.e. that the trees output by the function still be sorted if
the input ones are and contain one more occurrence of the inserted value, the
number of occurrences of any other value being left unaltered.
Here below is a naive tail-recursive definition of such function:
\null
\<close>
datatype 'a bintree = Leaf | Branch 'a "'a bintree" "'a bintree"
function (sequential) t_ins_naive ::
"bool \<Rightarrow> 'a::linorder \<Rightarrow> 'a bintree list \<Rightarrow> 'a bintree"
where
"t_ins_naive False x (Branch y yl yr # ts) = (if x \<le> y
then t_ins_naive False x (yl # Branch y yl yr # ts)
else t_ins_naive False x (yr # Branch y yl yr # ts))" |
"t_ins_naive False x (Leaf # ts) =
t_ins_naive True x (Branch x Leaf Leaf # ts)" |
"t_ins_naive True x (xt # Branch y yl yr # ts) = (if x \<le> y
then t_ins_naive True x (Branch y xt yr # ts)
else t_ins_naive True x (Branch y yl xt # ts))" |
"t_ins_naive True x [xt] = xt"
by pat_completeness auto
text \<open>
\null
The list appearing as the third argument, deputed to initially contain the sole
tree into which the second argument has to be inserted, is used to unfold all the
involved subtrees until a leaf is reached; then, such leaf is replaced with a branch
whose root value matches the second argument, and the subtree list is folded again.
The information on whether unfolding or folding is taking place is conveyed by the
first argument, whose value will respectively be \<open>False\<close> or \<open>True\<close>.
According to this plan, the computation is meant to terminate in correspondence
with pattern \<open>True\<close>, \<open>_\<close>, \<open>[_]\<close>. Hence, the above naive
definition comprises a non-recursive equation for this pattern only, so that the
residual ones \<open>True\<close>, \<open>_\<close>, \<open>_ # Leaf # _\<close> and \<open>_\<close>,
\<open>_\<close>, \<open>[]\<close> are not covered by any equation.
That which decreases in recursive calls is the size of the head of the subtree
list during unfolding, and the length of the list during folding. Furthermore,
unfolding precedes folding in the recursive call pipeline, viz. there is a
recursive equation switching from unfolding to folding, but no one carrying out
the opposite transition. These considerations suggest that a measure function
suitable to prove the termination of function \<open>t_ins_naive\<close> should roughly
match the sum of the length of the list and the size of the list head during
unfolding, and the length of the list alone during folding.
This idea can be refined by observing that the length of the list increases by one
at each recursive call during unfolding, and does not change in the recursive call
leading from unfolding to folding, at which the size of the input list head (a
leaf) equals zero. Therefore, in order that the measure function value be strictly
decreasing in each recursive call, the size of the list head has to be counted more
than once during unfolding -- e.g. twice --, and the length of the list has to be
decremented by one during folding -- no more than that, as otherwise the function
value would not change in the passage from a two-item to a one-item list.
As a result, a suitable measure function and the corresponding termination proof
are as follows:
\null
\<close>
fun t_ins_naive_measure :: "bool \<times> 'a \<times> 'a bintree list \<Rightarrow> nat" where
"t_ins_naive_measure (b, x, ts) = (if b
then length ts - 1
else length ts + 2 * size (hd ts))"
termination t_ins_naive
by (relation "measure t_ins_naive_measure", simp_all)
text \<open>
\null
Some further functions are needed to express the aforesaid correctness
properties of function \<open>t_ins_naive\<close>:
\null
\<close>
primrec t_set :: "'a bintree \<Rightarrow> 'a set" where
"t_set Leaf = {}" |
"t_set (Branch x xl xr) = {x} \<union> t_set xl \<union> t_set xr"
primrec t_multiset :: "'a bintree \<Rightarrow> 'a multiset" where
"t_multiset Leaf = {#}" |
"t_multiset (Branch x xl xr) = {#x#} + t_multiset xl + t_multiset xr"
lemma t_set_multiset: "t_set xt = set_mset (t_multiset xt)"
by (induction, simp_all)
primrec t_sorted :: "'a::linorder bintree \<Rightarrow> bool" where
"t_sorted Leaf = True" |
"t_sorted (Branch x xl xr) =
((\<forall>y \<in> t_set xl. y \<le> x) \<and> (\<forall>y \<in> t_set xr. x < y) \<and> t_sorted xl \<and> t_sorted xr)"
definition t_count :: "'a \<Rightarrow> 'a bintree \<Rightarrow> nat" where
"t_count x xt \<equiv> count (t_multiset xt) x"
text \<open>
\null
Functions \<open>t_set\<close> and \<open>t_multiset\<close> return the set and the multiset,
respectively, of the items of the input tree; the connection between them
expressed by lemma \<open>t_set_multiset\<close> will be used in step 9.
The target correctness theorems can then be enunciated as follows:
\null
\<open>t_sorted xt \<longrightarrow> t_sorted (t_ins_naive False x [xt])\<close>
\null
\<open>t_count y (t_ins_naive False x [xt]) =\<close>
\<open>(if y = x then Suc else id) (t_count y xt)\<close>
\<close>
subsection "Step 1"
text \<open>
This time, the Cartesian product of the input types will be implemented as a
record type. The second command instructs the system to regard such type as a
datatype, thus enabling record patterns:
\null
\<close>
record 'a t_type =
folding :: bool
item :: 'a
subtrees :: "'a bintree list"
function (sequential) t_ins_aux :: "'a::linorder t_type \<Rightarrow> 'a t_type" where
"t_ins_aux \<lparr>folding = False, item = x, subtrees = Branch y yl yr # ts\<rparr> =
(if x \<le> y
then t_ins_aux \<lparr>folding = False, item = x,
subtrees = yl # Branch y yl yr # ts\<rparr>
else t_ins_aux \<lparr>folding = False, item = x,
subtrees = yr # Branch y yl yr # ts\<rparr>)" |
"t_ins_aux \<lparr>folding = False, item = x, subtrees = Leaf # ts\<rparr> =
t_ins_aux \<lparr>folding = True, item = x, subtrees = Branch x Leaf Leaf # ts\<rparr>" |
"t_ins_aux \<lparr>folding = True, item = x, subtrees = xt # Branch y yl yr # ts\<rparr> =
(if x \<le> y
then t_ins_aux \<lparr>folding = True, item = x, subtrees = Branch y xt yr # ts\<rparr>
else t_ins_aux \<lparr>folding = True, item = x, subtrees = Branch y yl xt # ts\<rparr>)" |
"t_ins_aux X = X"
by pat_completeness auto
text \<open>
\null
Observe that the pattern appearing in the non-recursive equation matches any
one of the residual patterns
\<open>\<lparr>folding = True, item = _, subtrees = [_]\<rparr>\<close>,
\<open>\<lparr>folding = True, item = _, subtrees = _ # Leaf # _\<rparr>\<close>,
\<open>\<lparr>folding = _, item = _, subtrees = []\<rparr>\<close>, thus complying with the
requirement that the definition of function \<open>t_ins_aux\<close> be total.
Since the arguments of recursive calls in the definition of function
\<open>t_ins_aux\<close> are the same as those of function \<open>t_ins_naive\<close>,
the termination proof developed for the latter can be applied to the former
as well by just turning the input product type of the previous measure
function into the input record type of function \<open>t_ins_aux\<close>.
\null
\<close>
fun t_ins_aux_measure :: "'a t_type \<Rightarrow> nat" where
"t_ins_aux_measure \<lparr>folding = b, item = x, subtrees = ts\<rparr> = (if b
then length ts - 1
else length ts + 2 * size (hd ts))"
termination t_ins_aux
by (relation "measure t_ins_aux_measure", simp_all)
subsection "Step 2"
definition t_ins_in :: "'a \<Rightarrow> 'a bintree \<Rightarrow> 'a t_type" where
"t_ins_in x xt \<equiv> \<lparr>folding = False, item = x, subtrees = [xt]\<rparr>"
definition t_ins_out :: "'a t_type \<Rightarrow> 'a bintree" where
"t_ins_out X \<equiv> hd (subtrees X)"
definition t_ins :: "'a::linorder \<Rightarrow> 'a bintree \<Rightarrow> 'a bintree" where
"t_ins x xt \<equiv> t_ins_out (t_ins_aux (t_ins_in x xt))"
text \<open>
\null
Since the significant inputs of function \<open>t_ins_naive\<close> match pattern
\<open>False\<close>, \<open>_\<close>, \<open>[_]\<close>, those of function \<open>t_ins_aux\<close>
match pattern \<open>\<lparr>folding = False, item = _, subtrees = [_]\<rparr>\<close>, thus
being in a one-to-one correspondence with the Cartesian product of the types
of the second and the third component.
Then, the target correctness theorems can be put into the following equivalent
form:
\null
\<open>t_sorted xt \<longrightarrow> t_sorted (t_ins x xt)\<close>
\null
\<open>t_count y (t_ins x xt) =\<close>
\<open>(if y = x then Suc else id) (t_count y xt)\<close>
\<close>
subsection "Step 3"
inductive_set t_ins_set :: "'a::linorder t_type \<Rightarrow> 'a t_type set"
for X :: "'a t_type" where
R0: "X \<in> t_ins_set X" |
R1: "\<lbrakk>\<lparr>folding = False, item = x, subtrees = Branch y yl yr # ts\<rparr> \<in> t_ins_set X;
x \<le> y\<rbrakk> \<Longrightarrow>
\<lparr>folding = False, item = x, subtrees = yl # Branch y yl yr # ts\<rparr>
\<in> t_ins_set X" |
R2: "\<lbrakk>\<lparr>folding = False, item = x, subtrees = Branch y yl yr # ts\<rparr> \<in> t_ins_set X;
\<not> x \<le> y\<rbrakk> \<Longrightarrow>
\<lparr>folding = False, item = x, subtrees = yr # Branch y yl yr # ts\<rparr>
\<in> t_ins_set X" |
R3: "\<lparr>folding = False, item = x, subtrees = Leaf # ts\<rparr> \<in> t_ins_set X \<Longrightarrow>
\<lparr>folding = True, item = x, subtrees = Branch x Leaf Leaf # ts\<rparr>
\<in> t_ins_set X" |
R4: "\<lbrakk>\<lparr>folding = True, item = x, subtrees = xt # Branch y yl yr # ts\<rparr>
\<in> t_ins_set X; x \<le> y\<rbrakk> \<Longrightarrow>
\<lparr>folding = True, item = x, subtrees = Branch y xt yr # ts\<rparr> \<in> t_ins_set X" |
R5: "\<lbrakk>\<lparr>folding = True, item = x, subtrees = xt # Branch y yl yr # ts\<rparr>
\<in> t_ins_set X; \<not> x \<le> y\<rbrakk> \<Longrightarrow>
\<lparr>folding = True, item = x, subtrees = Branch y yl xt # ts\<rparr> \<in> t_ins_set X"
subsection "Step 4"
lemma t_ins_subset:
assumes XY: "Y \<in> t_ins_set X"
shows "t_ins_set Y \<subseteq> t_ins_set X"
proof (rule subsetI, erule t_ins_set.induct)
show "Y \<in> t_ins_set X" using XY .
next
fix x y yl yr ts
assume
"\<lparr>folding = False, item = x, subtrees = Branch y yl yr # ts\<rparr> \<in> t_ins_set X"
and "x \<le> y"
thus "\<lparr>folding = False, item = x, subtrees = yl # Branch y yl yr # ts\<rparr>
\<in> t_ins_set X" by (rule R1)
next
fix x y yl yr ts
assume
"\<lparr>folding = False, item = x, subtrees = Branch y yl yr # ts\<rparr> \<in> t_ins_set X"
and "\<not> x \<le> y"
thus "\<lparr>folding = False, item = x, subtrees = yr # Branch y yl yr # ts\<rparr>
\<in> t_ins_set X" by (rule R2)
next
fix x ts
assume "\<lparr>folding = False, item = x, subtrees = Leaf # ts\<rparr> \<in> t_ins_set X"
thus "\<lparr>folding = True, item = x, subtrees = Branch x Leaf Leaf # ts\<rparr>
\<in> t_ins_set X" by (rule R3)
next
fix x xt y yl yr ts
assume
"\<lparr>folding = True, item = x, subtrees = xt # Branch y yl yr # ts\<rparr> \<in> t_ins_set X"
and "x \<le> y"
thus "\<lparr>folding = True, item = x, subtrees = Branch y xt yr # ts\<rparr> \<in> t_ins_set X"
by (rule R4)
next
fix x xt y yl yr ts
assume
"\<lparr>folding = True, item = x, subtrees = xt # Branch y yl yr # ts\<rparr> \<in> t_ins_set X"
and "\<not> x \<le> y"
thus "\<lparr>folding = True, item = x, subtrees = Branch y yl xt # ts\<rparr> \<in> t_ins_set X"
by (rule R5)
qed
lemma t_ins_aux_set: "t_ins_aux X \<in> t_ins_set X"
proof (induction rule: t_ins_aux.induct,
simp_all add: R0 del: t_ins_aux.simps(1, 3))
fix x :: 'a and y yl yr ts
let
?X = "\<lparr>folding = False, item = x, subtrees = Branch y yl yr # ts\<rparr>" and
?X' = "\<lparr>folding = False, item = x, subtrees = yl # Branch y yl yr # ts\<rparr>" and
?X'' = "\<lparr>folding = False, item = x, subtrees = yr # Branch y yl yr # ts\<rparr>"
assume
case1: "x \<le> y \<Longrightarrow> t_ins_aux ?X' \<in> t_ins_set ?X'" and
case2: "\<not> x \<le> y \<Longrightarrow> t_ins_aux ?X'' \<in> t_ins_set ?X''"
have 0: "?X \<in> t_ins_set ?X" by (rule R0)
show "t_ins_aux ?X \<in> t_ins_set ?X"
proof (cases "x \<le> y", simp_all)
assume "x \<le> y"
with 0 have "?X' \<in> t_ins_set ?X" by (rule R1)
hence "t_ins_set ?X' \<subseteq> t_ins_set ?X" by (rule t_ins_subset)
moreover have "t_ins_aux ?X' \<in> t_ins_set ?X'"
using case1 and \<open>x \<le> y\<close> by simp
ultimately show "t_ins_aux ?X' \<in> t_ins_set ?X" by (rule subsetD)
next
assume "\<not> x \<le> y"
with 0 have "?X'' \<in> t_ins_set ?X" by (rule R2)
hence "t_ins_set ?X'' \<subseteq> t_ins_set ?X" by (rule t_ins_subset)
moreover have "t_ins_aux ?X'' \<in> t_ins_set ?X''"
using case2 and \<open>\<not> x \<le> y\<close> by simp
ultimately show "t_ins_aux ?X'' \<in> t_ins_set ?X" by (rule subsetD)
qed
next
fix x :: 'a and ts
let
?X = "\<lparr>folding = False, item = x, subtrees = Leaf # ts\<rparr>" and
?X' = "\<lparr>folding = True, item = x, subtrees = Branch x Leaf Leaf # ts\<rparr>"
have "?X \<in> t_ins_set ?X" by (rule R0)
hence "?X' \<in> t_ins_set ?X" by (rule R3)
hence "t_ins_set ?X' \<subseteq> t_ins_set ?X" by (rule t_ins_subset)
moreover assume "t_ins_aux ?X' \<in> t_ins_set ?X'"
ultimately show "t_ins_aux ?X' \<in> t_ins_set ?X" by (rule subsetD)
next
fix x :: 'a and xt y yl yr ts
let
?X = "\<lparr>folding = True, item = x, subtrees = xt # Branch y yl yr # ts\<rparr>" and
?X' = "\<lparr>folding = True, item = x, subtrees = Branch y xt yr # ts\<rparr>" and
?X'' = "\<lparr>folding = True, item = x, subtrees = Branch y yl xt # ts\<rparr>"
assume
case1: "x \<le> y \<Longrightarrow> t_ins_aux ?X' \<in> t_ins_set ?X'" and
case2: "\<not> x \<le> y \<Longrightarrow> t_ins_aux ?X'' \<in> t_ins_set ?X''"
have 0: "?X \<in> t_ins_set ?X" by (rule R0)
show "t_ins_aux ?X \<in> t_ins_set ?X"
proof (cases "x \<le> y", simp_all)
assume "x \<le> y"
with 0 have "?X' \<in> t_ins_set ?X" by (rule R4)
hence "t_ins_set ?X' \<subseteq> t_ins_set ?X" by (rule t_ins_subset)
moreover have "t_ins_aux ?X' \<in> t_ins_set ?X'"
using case1 and \<open>x \<le> y\<close> by simp
ultimately show "t_ins_aux ?X' \<in> t_ins_set ?X" by (rule subsetD)
next
assume "\<not> x \<le> y"
with 0 have "?X'' \<in> t_ins_set ?X" by (rule R5)
hence "t_ins_set ?X'' \<subseteq> t_ins_set ?X" by (rule t_ins_subset)
moreover have "t_ins_aux ?X'' \<in> t_ins_set ?X''"
using case2 and \<open>\<not> x \<le> y\<close> by simp
ultimately show "t_ins_aux ?X'' \<in> t_ins_set ?X" by (rule subsetD)
qed
qed
subsection "Step 5"
primrec t_val :: "'a bintree \<Rightarrow> 'a" where
"t_val (Branch x xl xr) = x"
primrec t_left :: "'a bintree \<Rightarrow> 'a bintree" where
"t_left (Branch x xl xr) = xl"
primrec t_right :: "'a bintree \<Rightarrow> 'a bintree" where
"t_right (Branch x xl xr) = xr"
text \<open>
\null
The partiality of the definition of the previous functions, which merely return
the root value and either subtree of the input branch, does not matter as they
will be applied to branches only.
These functions are used to define the following invariant -- this time, a single
invariant for both of the target correctness theorems:
\null
\<close>
fun t_ins_inv :: "'a::linorder \<Rightarrow> 'a bintree \<Rightarrow> 'a t_type \<Rightarrow> bool" where
"t_ins_inv x xt \<lparr>folding = b, item = y, subtrees = ts\<rparr> =
(y = x \<and>
(\<forall>n \<in> {..<length ts}.
(t_sorted xt \<longrightarrow> t_sorted (ts ! n)) \<and>
(0 < n \<longrightarrow> (\<exists>y yl yr. ts ! n = Branch y yl yr)) \<and>
(let ts' = ts @ [Branch x xt Leaf] in t_multiset (ts ! n) =
(if b \<and> n = 0 then {#x#} else {#}) +
(if x \<le> t_val (ts' ! Suc n)
then t_multiset (t_left (ts' ! Suc n))
else t_multiset (t_right (ts' ! Suc n))))))"
text \<open>
\null
More precisely, the invariant, whose type has to match \<open>'a t_type \<Rightarrow> bool\<close>
according to the method specification, shall be comprised of function
\<open>t_ins_inv x xt\<close>, where \<open>x\<close>, \<open>xt\<close> are the free variables
appearing in the target theorems as the arguments of function \<open>t_ins\<close>.
\<close>
subsection "Step 6"
lemma t_ins_input: "t_ins_inv x xt \<lparr>folding = False, item = x, subtrees = [xt]\<rparr>"
by simp
subsection "Step 7"
fun t_ins_form :: "'a t_type \<Rightarrow> bool" where
"t_ins_form \<lparr>folding = True, item = _, subtrees = [_]\<rparr> = True" |
"t_ins_form \<lparr>folding = True, item = _, subtrees = _ # Leaf # _\<rparr> = True" |
"t_ins_form _ = False"
lemma t_ins_intro_1:
"\<lbrakk>t_ins_inv x xt X; t_ins_form X\<rbrakk> \<Longrightarrow>
t_sorted xt \<longrightarrow> t_sorted (t_ins_out X)"
-proof (rule t_ins_form.cases [of X], simp_all add: t_ins_out_def)
-qed (erule conjE, drule_tac x = "Suc 0" in bspec, simp_all)
+ apply (rule t_ins_form.cases [of X])
+ apply (auto simp add: t_ins_out_def)
+ apply force
+ done
lemma t_ins_intro_2:
"\<lbrakk>t_ins_inv x xt X; t_ins_form X\<rbrakk> \<Longrightarrow>
t_count y (t_ins_out X) = (if y = x then Suc else id) (t_count y xt)"
-proof (rule t_ins_form.cases [of X], simp_all add: t_ins_out_def t_count_def)
-qed (erule conjE, drule_tac x = "Suc 0" in bspec, simp_all)
+ apply (rule t_ins_form.cases [of X])
+ apply (auto simp add: t_ins_out_def t_count_def)
+ apply force
+ apply force
+ done
text \<open>
\null
Defining predicate \<open>t_ins_form\<close> by means of pattern matching rather than
quantifiers permits a faster proof of the introduction rules through a case
distinction followed by simplification. These steps leave the subgoal
corresponding to pattern
\<open>\<lparr>folding = True, item = _, subtrees = _ # Leaf # _\<rparr>\<close> to be proven, which
can be done \emph{ad absurdum} as this pattern is incompatible with the invariant,
stating that all the subtrees in the list except for its head are branches.
The reason why this pattern, unlike
\<open>\<lparr>folding = _, item = _, subtrees = []\<rparr>\<close>, is not filtered by predicate
\<open>t_ins_form\<close>, is that the lack of its occurrences in recursive calls in
correspondence with significant inputs cannot be proven by rule inversion,
being it compatible with the patterns introduced by rules \<open>R3\<close>,
\<open>R4\<close>, and \<open>R5\<close>.
\<close>
subsection "Step 8"
text \<open>
This step will be accomplished by first proving by recursion induction that
the outputs of function \<open>t_ins_aux\<close> match either of the patterns
satisfying predicate \<open>t_ins_form\<close> or else the residual one
\<open>\<lparr>folding = _, item = _, subtrees = []\<rparr>\<close>, and then proving by rule
inversion that the last pattern may not occur in recursive calls in
correspondence with significant inputs.
\null
\<close>
definition t_ins_form_all :: "'a t_type \<Rightarrow> bool" where
"t_ins_form_all X \<equiv> t_ins_form X \<or> subtrees X = []"
lemma t_ins_form_aux_all: "t_ins_form_all (t_ins_aux X)"
by (rule t_ins_aux.induct [of "\<lambda>X. t_ins_form_all (t_ins_aux X)"],
simp_all add: t_ins_form_all_def)
lemma t_ins_form_aux:
"t_ins_form (t_ins_aux \<lparr>folding = False, item = x, subtrees = [xt]\<rparr>)"
(is "_ (t_ins_aux ?X)")
using t_ins_aux_set [of ?X]
proof (rule t_ins_set.cases, insert t_ins_form_aux_all [of ?X])
qed (simp_all add: t_ins_form_all_def)
subsection "Step 9"
lemma t_ins_invariance:
assumes XY: "Y \<in> t_ins_set X" and X: "t_ins_inv x xt X"
shows "t_ins_inv x xt Y"
-using XY
+using XY [[simproc del: defined_all]]
proof (rule t_ins_set.induct, simp_all split del: if_split)
show "t_ins_inv x xt X" using X .
next
fix z :: "'a::linorder" and y yl yr ts
assume "z = x \<and>
(\<forall>n \<in> {..<Suc (length ts)}.
(t_sorted xt \<longrightarrow> t_sorted ((Branch y yl yr # ts) ! n)) \<and>
(0 < n \<longrightarrow> (\<exists>y' yl' yr'. ts ! (n - Suc 0) = Branch y' yl' yr')) \<and>
(let ts' = Branch y yl yr # ts @ [Branch x xt Leaf]
in t_multiset ((Branch y yl yr # ts) ! n) =
(if x \<le> t_val ((ts @ [Branch x xt Leaf]) ! n)
then t_multiset (t_left (ts' ! Suc n))
else t_multiset (t_right (ts' ! Suc n)))))"
(is "_ \<and> (\<forall>n \<in> {..<Suc (length ts)}. ?P n)")
hence I: "\<forall>n \<in> {..<Suc (length ts)}. ?P n" ..
assume xy: "x \<le> y"
show
"\<forall>n \<in> {..<Suc (Suc (length ts))}.
(t_sorted xt \<longrightarrow> t_sorted ((yl # Branch y yl yr # ts) ! n)) \<and>
(0 < n \<longrightarrow> (\<exists>y' yl' yr'. (Branch y yl yr # ts) ! (n - Suc 0) =
Branch y' yl' yr')) \<and>
(let ts' = yl # Branch y yl yr # ts @ [Branch x xt Leaf]
in t_multiset ((yl # Branch y yl yr # ts) ! n) =
(if x \<le> t_val ((Branch y yl yr # ts @ [Branch x xt Leaf]) ! n)
then t_multiset (t_left (ts' ! Suc n))
else t_multiset (t_right (ts' ! Suc n))))"
(is "\<forall>n \<in> {..<Suc (Suc (length ts))}. ?Q n")
proof
fix n
assume n: "n \<in> {..<Suc (Suc (length ts))}"
show "?Q n"
proof (cases n)
case 0
have "0 \<in> {..<Suc (length ts)}" by simp
with I have "?P 0" ..
thus ?thesis by (simp add: Let_def xy 0)
next
case (Suc m)
hence "m \<in> {..<Suc (length ts)}" using n by simp
with I have "?P m" ..
thus ?thesis
proof (simp add: Let_def Suc)
qed (cases m, simp_all)
qed
qed
next
fix z :: "'a::linorder" and y yl yr ts
assume "z = x \<and>
(\<forall>n \<in> {..<Suc (length ts)}.
(t_sorted xt \<longrightarrow> t_sorted ((Branch y yl yr # ts) ! n)) \<and>
(0 < n \<longrightarrow> (\<exists>y' yl' yr'. ts ! (n - Suc 0) = Branch y' yl' yr')) \<and>
(let ts' = Branch y yl yr # ts @ [Branch x xt Leaf]
in t_multiset ((Branch y yl yr # ts) ! n) =
(if x \<le> t_val ((ts @ [Branch x xt Leaf]) ! n)
then t_multiset (t_left (ts' ! Suc n))
else t_multiset (t_right (ts' ! Suc n)))))"
(is "_ \<and> (\<forall>n \<in> {..<Suc (length ts)}. ?P n)")
hence I: "\<forall>n \<in> {..<Suc (length ts)}. ?P n" ..
assume xy: "\<not> x \<le> y"
show
"\<forall>n \<in> {..<Suc (Suc (length ts))}.
(t_sorted xt \<longrightarrow> t_sorted ((yr # Branch y yl yr # ts) ! n)) \<and>
(0 < n \<longrightarrow> (\<exists>y' yl' yr'. (Branch y yl yr # ts) ! (n - Suc 0) =
Branch y' yl' yr')) \<and>
(let ts' = yr # Branch y yl yr # ts @ [Branch x xt Leaf]
in t_multiset ((yr # Branch y yl yr # ts) ! n) =
(if x \<le> t_val ((Branch y yl yr # ts @ [Branch x xt Leaf]) ! n)
then t_multiset (t_left (ts' ! Suc n))
else t_multiset (t_right (ts' ! Suc n))))"
(is "\<forall>n \<in> {..<Suc (Suc (length ts))}. ?Q n")
proof
fix n
assume n: "n \<in> {..<Suc (Suc (length ts))}"
show "?Q n"
proof (cases n)
case 0
have "0 \<in> {..<Suc (length ts)}" by simp
with I have "?P 0" ..
thus ?thesis by (simp add: Let_def xy 0)
next
case (Suc m)
hence "m \<in> {..<Suc (length ts)}" using n by simp
with I have "?P m" ..
thus ?thesis
proof (simp add: Let_def Suc)
qed (cases m, simp_all)
qed
qed
next
fix z :: 'a and ts
assume "z = x \<and>
(\<forall>n \<in> {..<Suc (length ts)}.
(t_sorted xt \<longrightarrow> t_sorted ((Leaf # ts) ! n)) \<and>
(0 < n \<longrightarrow> (\<exists>y yl yr. ts ! (n - Suc 0) = Branch y yl yr)) \<and>
(let ts' = Leaf # ts @ [Branch x xt Leaf]
in t_multiset ((Leaf # ts) ! n) =
(if x \<le> t_val ((ts @ [Branch x xt Leaf]) ! n)
then t_multiset (t_left (ts' ! Suc n))
else t_multiset (t_right (ts' ! Suc n)))))"
(is "_ \<and> (\<forall>n \<in> {..<Suc (length ts)}. ?P n)")
hence I: "\<forall>n \<in> {..<Suc (length ts)}. ?P n" ..
show
"\<forall>n \<in> {..<Suc (length ts)}.
(t_sorted xt \<longrightarrow> t_sorted ((Branch x Leaf Leaf # ts) ! n)) \<and>
(let ts' = Branch x Leaf Leaf # ts @ [Branch x xt Leaf]
in t_multiset ((Branch x Leaf Leaf # ts) ! n) =
(if n = 0 then {#x#} else {#}) +
(if x \<le> t_val ((ts @ [Branch x xt Leaf]) ! n)
then t_multiset (t_left (ts' ! Suc n))
else t_multiset (t_right (ts' ! Suc n))))"
(is "\<forall>n \<in> {..<Suc (length ts)}. ?Q n")
proof
fix n
assume n: "n \<in> {..<Suc (length ts)}"
show "?Q n"
proof (cases n)
case 0
have "0 \<in> {..<Suc (length ts)}" by simp
with I have "?P 0" ..
thus ?thesis by (simp add: Let_def 0 split: if_split_asm)
next
case (Suc m)
have "?P n" using I and n ..
thus ?thesis by (simp add: Let_def Suc)
qed
qed
next
fix z :: 'a and zt y yl yr ts
assume "z = x \<and>
(\<forall>n \<in> {..<Suc (Suc (length ts))}.
(t_sorted xt \<longrightarrow> t_sorted ((zt # Branch y yl yr # ts) ! n)) \<and>
(0 < n \<longrightarrow> (\<exists>y' yl' yr'. (Branch y yl yr # ts) ! (n - Suc 0) =
Branch y' yl' yr')) \<and>
(let ts' = zt # Branch y yl yr # ts @ [Branch x xt Leaf]
in t_multiset ((zt # Branch y yl yr # ts) ! n) =
(if n = 0 then {#x#} else {#}) +
(if x \<le> t_val ((Branch y yl yr # ts @ [Branch x xt Leaf]) ! n)
then t_multiset (t_left (ts' ! Suc n))
else t_multiset (t_right (ts' ! Suc n)))))"
(is "_ \<and> (\<forall>n \<in> {..<Suc (Suc (length ts))}. ?P n)")
hence I: "\<forall>n \<in> {..<Suc (Suc (length ts))}. ?P n" ..
assume xy: "x \<le> y"
show
"\<forall>n \<in> {..<Suc (length ts)}.
(t_sorted xt \<longrightarrow> t_sorted ((Branch y zt yr # ts) ! n)) \<and>
(0 < n \<longrightarrow> (\<exists>y' yl' yr'. ts ! (n - Suc 0) = Branch y' yl' yr')) \<and>
(let ts' = Branch y zt yr # ts @ [Branch x xt Leaf]
in t_multiset ((Branch y zt yr # ts) ! n) =
(if n = 0 then {#x#} else {#}) +
(if x \<le> t_val ((ts @ [Branch x xt Leaf]) ! n)
then t_multiset (t_left (ts' ! Suc n))
else t_multiset (t_right (ts' ! Suc n))))"
(is "\<forall>n \<in> {..<Suc (length ts)}. ?Q n")
proof
fix n
assume n: "n \<in> {..<Suc (length ts)}"
show "?Q n"
proof (cases n)
case 0
have "0 \<in> {..<Suc (Suc (length ts))}" by simp
with I have "?P 0" ..
hence I0: "(t_sorted xt \<longrightarrow> t_sorted zt) \<and>
t_multiset zt = {#x#} + t_multiset yl"
by (simp add: Let_def xy)
have "Suc 0 \<in> {..<Suc (Suc (length ts))}" by simp
with I have "?P (Suc 0)" ..
hence I1: "(t_sorted xt \<longrightarrow> t_sorted (Branch y yl yr)) \<and>
t_multiset (Branch y yl yr) =
(if x \<le> t_val ((ts @ [Branch x xt Leaf]) ! 0)
then t_multiset (t_left ((ts @ [Branch x xt Leaf]) ! 0))
else t_multiset (t_right ((ts @ [Branch x xt Leaf]) ! 0)))"
by (simp add: Let_def)
show ?thesis
proof (simp add: Let_def 0 del: t_sorted.simps split del: if_split,
rule conjI, simp_all add: Let_def 0 del: t_sorted.simps,
rule_tac [2] conjI, rule_tac [!] impI)
assume s: "t_sorted xt"
hence "t_sorted zt" using I0 by simp
moreover have "t_sorted (Branch y yl yr)" using I1 and s by simp
moreover have "t_set zt = {x} \<union> t_set yl" using I0
by (simp add: t_set_multiset)
ultimately show "t_sorted (Branch y zt yr)" using xy by simp
next
assume "x \<le> t_val ((ts @ [Branch x xt Leaf]) ! 0)"
hence "t_multiset (t_left ((ts @ [Branch x xt Leaf]) ! 0)) =
t_multiset (Branch y yl yr)" using I1 by simp
thus "add_mset y (t_multiset zt + t_multiset yr) =
add_mset x (t_multiset (t_left ((ts @ [Branch x xt Leaf]) ! 0)))" using I0
by simp
next
assume "\<not> x \<le> t_val ((ts @ [Branch x xt Leaf]) ! 0)"
hence "t_multiset (t_right ((ts @ [Branch x xt Leaf]) ! 0)) =
t_multiset (Branch y yl yr)" using I1 by simp
thus "add_mset y (t_multiset zt + t_multiset yr) =
add_mset x (t_multiset (t_right ((ts @ [Branch x xt Leaf]) ! 0)))" using I0
by simp
qed
next
case (Suc m)
have "Suc n \<in> {..<Suc (Suc (length ts))}" using n by simp
with I have "?P (Suc n)" ..
thus ?thesis by (simp add: Let_def Suc)
qed
qed
next
fix z :: 'a and zt y yl yr ts
assume "z = x \<and>
(\<forall>n \<in> {..<Suc (Suc (length ts))}.
(t_sorted xt \<longrightarrow> t_sorted ((zt # Branch y yl yr # ts) ! n)) \<and>
(0 < n \<longrightarrow> (\<exists>y' yl' yr'. (Branch y yl yr # ts) ! (n - Suc 0) =
Branch y' yl' yr')) \<and>
(let ts' = zt # Branch y yl yr # ts @ [Branch x xt Leaf]
in t_multiset ((zt # Branch y yl yr # ts) ! n) =
(if n = 0 then {#x#} else {#}) +
(if x \<le> t_val ((Branch y yl yr # ts @ [Branch x xt Leaf]) ! n)
then t_multiset (t_left (ts' ! Suc n))
else t_multiset (t_right (ts' ! Suc n)))))"
(is "_ \<and> (\<forall>n \<in> {..<Suc (Suc (length ts))}. ?P n)")
hence I: "\<forall>n \<in> {..<Suc (Suc (length ts))}. ?P n" ..
assume xy: "\<not> x \<le> y"
show
"\<forall>n \<in> {..<Suc (length ts)}.
(t_sorted xt \<longrightarrow> t_sorted ((Branch y yl zt # ts) ! n)) \<and>
(0 < n \<longrightarrow> (\<exists>y' yl' yr'. ts ! (n - Suc 0) = Branch y' yl' yr')) \<and>
(let ts' = Branch y yl zt # ts @ [Branch x xt Leaf]
in t_multiset ((Branch y yl zt # ts) ! n) =
(if n = 0 then {#x#} else {#}) +
(if x \<le> t_val ((ts @ [Branch x xt Leaf]) ! n)
then t_multiset (t_left (ts' ! Suc n))
else t_multiset (t_right (ts' ! Suc n))))"
(is "\<forall>n \<in> {..<Suc (length ts)}. ?Q n")
proof
fix n
assume n: "n \<in> {..<Suc (length ts)}"
show "?Q n"
proof (cases n)
case 0
have "0 \<in> {..<Suc (Suc (length ts))}" by simp
with I have "?P 0" ..
hence I0: "(t_sorted xt \<longrightarrow> t_sorted zt) \<and>
t_multiset zt = {#x#} + t_multiset yr"
by (simp add: Let_def xy)
have "Suc 0 \<in> {..<Suc (Suc (length ts))}" by simp
with I have "?P (Suc 0)" ..
hence I1: "(t_sorted xt \<longrightarrow> t_sorted (Branch y yl yr)) \<and>
t_multiset (Branch y yl yr) =
(if x \<le> t_val ((ts @ [Branch x xt Leaf]) ! 0)
then t_multiset (t_left ((ts @ [Branch x xt Leaf]) ! 0))
else t_multiset (t_right ((ts @ [Branch x xt Leaf]) ! 0)))"
by (simp add: Let_def)
show ?thesis
proof (simp add: Let_def 0 del: t_sorted.simps split del: if_split,
rule conjI, simp_all add: Let_def 0 del: t_sorted.simps,
rule_tac [2] conjI, rule_tac [!] impI)
assume s: "t_sorted xt"
hence "t_sorted zt" using I0 by simp
moreover have "t_sorted (Branch y yl yr)" using I1 and s by simp
moreover have "t_set zt = {x} \<union> t_set yr" using I0
by (simp add: t_set_multiset)
ultimately show "t_sorted (Branch y yl zt)" using xy by simp
next
assume "x \<le> t_val ((ts @ [Branch x xt Leaf]) ! 0)"
hence "t_multiset (t_left ((ts @ [Branch x xt Leaf]) ! 0)) =
t_multiset (Branch y yl yr)" using I1 by simp
thus "add_mset y (t_multiset yl + t_multiset zt) =
add_mset x (t_multiset (t_left ((ts @ [Branch x xt Leaf]) ! 0)))" using I0
by simp
next
assume "\<not> x \<le> t_val ((ts @ [Branch x xt Leaf]) ! 0)"
hence "t_multiset (t_right ((ts @ [Branch x xt Leaf]) ! 0)) =
t_multiset (Branch y yl yr)" using I1 by simp
thus "add_mset y (t_multiset yl + t_multiset zt) =
add_mset x (t_multiset (t_right ((ts @ [Branch x xt Leaf]) ! 0)))" using I0
by simp
qed
next
case (Suc m)
have "Suc n \<in> {..<Suc (Suc (length ts))}" using n by simp
with I have "?P (Suc n)" ..
thus ?thesis by (simp add: Let_def Suc)
qed
qed
qed
subsection "Step 10"
theorem "t_sorted xt \<longrightarrow> t_sorted (t_ins x xt)"
proof -
let ?X = "\<lparr>folding = False, item = x, subtrees = [xt]\<rparr>"
have "t_ins_aux ?X \<in> t_ins_set ?X" by (rule t_ins_aux_set)
moreover have "t_ins_inv x xt ?X" by (rule t_ins_input)
ultimately have "t_ins_inv x xt (t_ins_aux ?X)" by (rule t_ins_invariance)
moreover have "t_ins_form (t_ins_aux ?X)" by (rule t_ins_form_aux)
ultimately have "t_sorted xt \<longrightarrow> t_sorted (t_ins_out (t_ins_aux ?X))"
by (rule t_ins_intro_1)
moreover have "?X = t_ins_in x xt" by (simp add: t_ins_in_def)
ultimately show ?thesis by (simp add: t_ins_def)
qed
theorem "t_count y (t_ins x xt) = (if y = x then Suc else id) (t_count y xt)"
proof -
let ?X = "\<lparr>folding = False, item = x, subtrees = [xt]\<rparr>"
have "t_ins_aux ?X \<in> t_ins_set ?X" by (rule t_ins_aux_set)
moreover have "t_ins_inv x xt ?X" by (rule t_ins_input)
ultimately have "t_ins_inv x xt (t_ins_aux ?X)" by (rule t_ins_invariance)
moreover have "t_ins_form (t_ins_aux ?X)" by (rule t_ins_form_aux)
ultimately have "t_count y (t_ins_out (t_ins_aux ?X)) =
(if y = x then Suc else id) (t_count y xt)"
by (rule t_ins_intro_2)
moreover have "?X = t_ins_in x xt" by (simp add: t_ins_in_def)
ultimately show ?thesis by (simp add: t_ins_def)
qed
end
diff --git a/thys/Universal_Turing_Machine/Abacus.thy b/thys/Universal_Turing_Machine/Abacus.thy
--- a/thys/Universal_Turing_Machine/Abacus.thy
+++ b/thys/Universal_Turing_Machine/Abacus.thy
@@ -1,3327 +1,3330 @@
(* Title: thys/Abacus.thy
Author: Jian Xu, Xingyuan Zhang, and Christian Urban
Modifications: Sebastiaan Joosten
*)
chapter \<open>Abacus Machines\<close>
theory Abacus
imports Turing_Hoare Abacus_Mopup
begin
declare replicate_Suc[simp add]
(* Abacus instructions *)
datatype abc_inst =
Inc nat
| Dec nat nat
| Goto nat
type_synonym abc_prog = "abc_inst list"
type_synonym abc_state = nat
text \<open>
The memory of Abacus machine is defined as a list of contents, with
every units addressed by index into the list.
\<close>
type_synonym abc_lm = "nat list"
text \<open>
Fetching contents out of memory. Units not represented by list elements are considered
as having content \<open>0\<close>.
\<close>
fun abc_lm_v :: "abc_lm \<Rightarrow> nat \<Rightarrow> nat"
where
"abc_lm_v lm n = (if (n < length lm) then (lm!n) else 0)"
text \<open>
Set the content of memory unit \<open>n\<close> to value \<open>v\<close>.
\<open>am\<close> is the Abacus memory before setting.
If address \<open>n\<close> is outside to scope of \<open>am\<close>, \<open>am\<close>
is extended so that \<open>n\<close> becomes in scope.
\<close>
fun abc_lm_s :: "abc_lm \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_lm"
where
"abc_lm_s am n v = (if (n < length am) then (am[n:=v]) else
am@ (replicate (n - length am) 0) @ [v])"
text \<open>
The configuration of Abaucs machines consists of its current state and its
current memory:
\<close>
type_synonym abc_conf = "abc_state \<times> abc_lm"
text \<open>
Fetch instruction out of Abacus program:
\<close>
fun abc_fetch :: "nat \<Rightarrow> abc_prog \<Rightarrow> abc_inst option"
where
"abc_fetch s p = (if (s < length p) then Some (p ! s) else None)"
text \<open>
Single step execution of Abacus machine. If no instruction is feteched,
configuration does not change.
\<close>
fun abc_step_l :: "abc_conf \<Rightarrow> abc_inst option \<Rightarrow> abc_conf"
where
"abc_step_l (s, lm) a = (case a of
None \<Rightarrow> (s, lm) |
Some (Inc n) \<Rightarrow> (let nv = abc_lm_v lm n in
(s + 1, abc_lm_s lm n (nv + 1))) |
Some (Dec n e) \<Rightarrow> (let nv = abc_lm_v lm n in
if (nv = 0) then (e, abc_lm_s lm n 0)
else (s + 1, abc_lm_s lm n (nv - 1))) |
Some (Goto n) \<Rightarrow> (n, lm)
)"
text \<open>
Multi-step execution of Abacus machine.
\<close>
fun abc_steps_l :: "abc_conf \<Rightarrow> abc_prog \<Rightarrow> nat \<Rightarrow> abc_conf"
where
"abc_steps_l (s, lm) p 0 = (s, lm)" |
"abc_steps_l (s, lm) p (Suc n) =
abc_steps_l (abc_step_l (s, lm) (abc_fetch s p)) p n"
section \<open>
Compiling Abacus machines into Turing machines
\<close>
subsection \<open>
Compiling functions
\<close>
text \<open>
\<open>findnth n\<close> returns the TM which locates the represention of
memory cell \<open>n\<close> on the tape and changes representation of zero
on the way.
\<close>
fun findnth :: "nat \<Rightarrow> instr list"
where
"findnth 0 = []" |
"findnth (Suc n) = (findnth n @ [(W1, 2 * n + 1),
(R, 2 * n + 2), (R, 2 * n + 3), (R, 2 * n + 2)])"
text \<open>
\<open>tinc_b\<close> returns the TM which increments the representation
of the memory cell under rw-head by one and move the representation
of cells afterwards to the right accordingly.
\<close>
definition tinc_b :: "instr list"
where
"tinc_b \<equiv> [(W1, 1), (R, 2), (W1, 3), (R, 2), (W1, 3), (R, 4),
(L, 7), (W0, 5), (R, 6), (W0, 5), (W1, 3), (R, 6),
(L, 8), (L, 7), (R, 9), (L, 7), (R, 10), (W0, 9)]"
text \<open>
\<open>tinc ss n\<close> returns the TM which simulates the execution of
Abacus instruction \<open>Inc n\<close>, assuming that TM is located at
location \<open>ss\<close> in the final TM complied from the whole
Abacus program.
\<close>
fun tinc :: "nat \<Rightarrow> nat \<Rightarrow> instr list"
where
"tinc ss n = shift (findnth n @ shift tinc_b (2 * n)) (ss - 1)"
text \<open>
\<open>tinc_b\<close> returns the TM which decrements the representation
of the memory cell under rw-head by one and move the representation
of cells afterwards to the left accordingly.
\<close>
definition tdec_b :: "instr list"
where
"tdec_b \<equiv> [(W1, 1), (R, 2), (L, 14), (R, 3), (L, 4), (R, 3),
(R, 5), (W0, 4), (R, 6), (W0, 5), (L, 7), (L, 8),
(L, 11), (W0, 7), (W1, 8), (R, 9), (L, 10), (R, 9),
(R, 5), (W0, 10), (L, 12), (L, 11), (R, 13), (L, 11),
(R, 17), (W0, 13), (L, 15), (L, 14), (R, 16), (L, 14),
(R, 0), (W0, 16)]"
text \<open>
\<open>tdec ss n label\<close> returns the TM which simulates the execution of
Abacus instruction \<open>Dec n label\<close>, assuming that TM is located at
location \<open>ss\<close> in the final TM complied from the whole
Abacus program.
\<close>
fun tdec :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> instr list"
where
"tdec ss n e = shift (findnth n) (ss - 1) @ adjust (shift (shift tdec_b (2 * n)) (ss - 1)) e"
text \<open>
\<open>tgoto f(label)\<close> returns the TM simulating the execution of Abacus instruction
\<open>Goto label\<close>, where \<open>f(label)\<close> is the corresponding location of
\<open>label\<close> in the final TM compiled from the overall Abacus program.
\<close>
fun tgoto :: "nat \<Rightarrow> instr list"
where
"tgoto n = [(Nop, n), (Nop, n)]"
text \<open>
The layout of the final TM compiled from an Abacus program is represented
as a list of natural numbers, where the list element at index \<open>n\<close> represents the
starting state of the TM simulating the execution of \<open>n\<close>-th instruction
in the Abacus program.
\<close>
type_synonym layout = "nat list"
text \<open>
\<open>length_of i\<close> is the length of the
TM simulating the Abacus instruction \<open>i\<close>.
\<close>
fun length_of :: "abc_inst \<Rightarrow> nat"
where
"length_of i = (case i of
Inc n \<Rightarrow> 2 * n + 9 |
Dec n e \<Rightarrow> 2 * n + 16 |
Goto n \<Rightarrow> 1)"
text \<open>
\<open>layout_of ap\<close> returns the layout of Abacus program \<open>ap\<close>.
\<close>
fun layout_of :: "abc_prog \<Rightarrow> layout"
where "layout_of ap = map length_of ap"
text \<open>
\<open>start_of layout n\<close> looks out the starting state of \<open>n\<close>-th
TM in the finall TM.
\<close>
fun start_of :: "nat list \<Rightarrow> nat \<Rightarrow> nat"
where
"start_of ly x = (Suc (sum_list (take x ly))) "
text \<open>
\<open>ci lo ss i\<close> complies Abacus instruction \<open>i\<close>
assuming the TM of \<open>i\<close> starts from state \<open>ss\<close>
within the overal layout \<open>lo\<close>.
\<close>
fun ci :: "layout \<Rightarrow> nat \<Rightarrow> abc_inst \<Rightarrow> instr list"
where
"ci ly ss (Inc n) = tinc ss n"
| "ci ly ss (Dec n e) = tdec ss n (start_of ly e)"
| "ci ly ss (Goto n) = tgoto (start_of ly n)"
text \<open>
\<open>tpairs_of ap\<close> transfroms Abacus program \<open>ap\<close> pairing
every instruction with its starting state.
\<close>
fun tpairs_of :: "abc_prog \<Rightarrow> (nat \<times> abc_inst) list"
where "tpairs_of ap = (zip (map (start_of (layout_of ap))
[0..<(length ap)]) ap)"
text \<open>
\<open>tms_of ap\<close> returns the list of TMs, where every one of them simulates
the corresponding Abacus intruction in \<open>ap\<close>.
\<close>
fun tms_of :: "abc_prog \<Rightarrow> (instr list) list"
where "tms_of ap = map (\<lambda> (n, tm). ci (layout_of ap) n tm)
(tpairs_of ap)"
text \<open>
\<open>tm_of ap\<close> returns the final TM machine compiled from Abacus program \<open>ap\<close>.
\<close>
fun tm_of :: "abc_prog \<Rightarrow> instr list"
where "tm_of ap = concat (tms_of ap)"
lemma length_findnth:
"length (findnth n) = 4 * n"
by (induct n, auto)
lemma ci_length : "length (ci ns n ai) div 2 = length_of ai"
apply(auto simp: ci.simps tinc_b_def tdec_b_def length_findnth
split: abc_inst.splits simp del: adjust.simps)
done
subsection \<open>Representation of Abacus memory by TM tapes\<close>
text \<open>
\<open>crsp acf tcf\<close> meams the abacus configuration \<open>acf\<close>
is corretly represented by the TM configuration \<open>tcf\<close>.
\<close>
fun crsp :: "layout \<Rightarrow> abc_conf \<Rightarrow> config \<Rightarrow> cell list \<Rightarrow> bool"
where
"crsp ly (as, lm) (s, l, r) inres =
(s = start_of ly as \<and> (\<exists> x. r = <lm> @ Bk\<up>x) \<and>
l = Bk # Bk # inres)"
declare crsp.simps[simp del]
text \<open>
The type of invarints expressing correspondence between
Abacus configuration and TM configuration.
\<close>
type_synonym inc_inv_t = "abc_conf \<Rightarrow> config \<Rightarrow> cell list \<Rightarrow> bool"
declare tms_of.simps[simp del] tm_of.simps[simp del]
layout_of.simps[simp del] abc_fetch.simps [simp del]
tpairs_of.simps[simp del] start_of.simps[simp del]
ci.simps [simp del] length_of.simps[simp del]
layout_of.simps[simp del]
text \<open>
The lemmas in this section lead to the correctness of
the compilation of \<open>Inc n\<close> instruction.
\<close>
declare abc_step_l.simps[simp del] abc_steps_l.simps[simp del]
lemma start_of_nonzero[simp]: "start_of ly as > 0" "(start_of ly as = 0) = False"
apply(auto simp: start_of.simps)
done
lemma abc_steps_l_0: "abc_steps_l ac ap 0 = ac"
by(cases ac, simp add: abc_steps_l.simps)
lemma abc_step_red:
"abc_steps_l (as, am) ap stp = (bs, bm) \<Longrightarrow>
abc_steps_l (as, am) ap (Suc stp) = abc_step_l (bs, bm) (abc_fetch bs ap) "
proof(induct stp arbitrary: as am bs bm)
case 0
thus "?case"
by(simp add: abc_steps_l.simps abc_steps_l_0)
next
case (Suc stp as am bs bm)
have ind: "\<And>as am bs bm. abc_steps_l (as, am) ap stp = (bs, bm) \<Longrightarrow>
abc_steps_l (as, am) ap (Suc stp) = abc_step_l (bs, bm) (abc_fetch bs ap)"
by fact
have h:" abc_steps_l (as, am) ap (Suc stp) = (bs, bm)" by fact
obtain as' am' where g: "abc_step_l (as, am) (abc_fetch as ap) = (as', am')"
by(cases "abc_step_l (as, am) (abc_fetch as ap)", auto)
then have "abc_steps_l (as', am') ap (Suc stp) = abc_step_l (bs, bm) (abc_fetch bs ap)"
using h
by(intro ind, simp add: abc_steps_l.simps)
thus "?case"
using g
by(simp add: abc_steps_l.simps)
qed
lemma tm_shift_fetch:
"\<lbrakk>fetch A s b = (ac, ns); ns \<noteq> 0 \<rbrakk>
\<Longrightarrow> fetch (shift A off) s b = (ac, ns + off)"
apply(cases b;cases s)
apply(auto simp: fetch.simps shift.simps)
done
lemma tm_shift_eq_step:
assumes exec: "step (s, l, r) (A, 0) = (s', l', r')"
and notfinal: "s' \<noteq> 0"
shows "step (s + off, l, r) (shift A off, off) = (s' + off, l', r')"
using assms
apply(simp add: step.simps)
apply(cases "fetch A s (read r)", auto)
apply(drule_tac [!] off = off in tm_shift_fetch, simp_all)
done
declare step.simps[simp del] steps.simps[simp del] shift.simps[simp del]
lemma tm_shift_eq_steps:
assumes exec: "steps (s, l, r) (A, 0) stp = (s', l', r')"
and notfinal: "s' \<noteq> 0"
shows "steps (s + off, l, r) (shift A off, off) stp = (s' + off, l', r')"
using exec notfinal
proof(induct stp arbitrary: s' l' r', simp add: steps.simps)
fix stp s' l' r'
assume ind: "\<And>s' l' r'. \<lbrakk>steps (s, l, r) (A, 0) stp = (s', l', r'); s' \<noteq> 0\<rbrakk>
\<Longrightarrow> steps (s + off, l, r) (shift A off, off) stp = (s' + off, l', r')"
and h: " steps (s, l, r) (A, 0) (Suc stp) = (s', l', r')" "s' \<noteq> 0"
obtain s1 l1 r1 where g: "steps (s, l, r) (A, 0) stp = (s1, l1, r1)"
apply(cases "steps (s, l, r) (A, 0) stp") by blast
moreover then have "s1 \<noteq> 0"
using h
apply(simp add: step_red)
apply(cases "0 < s1", auto)
done
ultimately have "steps (s + off, l, r) (shift A off, off) stp =
(s1 + off, l1, r1)"
apply(intro ind, simp_all)
done
thus "steps (s + off, l, r) (shift A off, off) (Suc stp) = (s' + off, l', r')"
using h g assms
apply(simp add: step_red)
apply(intro tm_shift_eq_step, auto)
done
qed
lemma startof_ge1[simp]: "Suc 0 \<le> start_of ly as"
apply(simp add: start_of.simps)
done
lemma start_of_Suc1: "\<lbrakk>ly = layout_of ap;
abc_fetch as ap = Some (Inc n)\<rbrakk>
\<Longrightarrow> start_of ly (Suc as) = start_of ly as + 2 * n + 9"
apply(auto simp: start_of.simps layout_of.simps
length_of.simps abc_fetch.simps
take_Suc_conv_app_nth split: if_splits)
done
lemma start_of_Suc2:
"\<lbrakk>ly = layout_of ap;
abc_fetch as ap = Some (Dec n e)\<rbrakk> \<Longrightarrow>
start_of ly (Suc as) =
start_of ly as + 2 * n + 16"
apply(auto simp: start_of.simps layout_of.simps
length_of.simps abc_fetch.simps
take_Suc_conv_app_nth split: if_splits)
done
lemma start_of_Suc3:
"\<lbrakk>ly = layout_of ap;
abc_fetch as ap = Some (Goto n)\<rbrakk> \<Longrightarrow>
start_of ly (Suc as) = start_of ly as + 1"
apply(auto simp: start_of.simps layout_of.simps
length_of.simps abc_fetch.simps
take_Suc_conv_app_nth split: if_splits)
done
lemma length_ci_inc:
"length (ci ly ss (Inc n)) = 4*n + 18"
apply(auto simp: ci.simps length_findnth tinc_b_def)
done
lemma length_ci_dec:
"length (ci ly ss (Dec n e)) = 4*n + 32"
apply(auto simp: ci.simps length_findnth tdec_b_def)
done
lemma length_ci_goto:
"length (ci ly ss (Goto n )) = 2"
apply(auto simp: ci.simps length_findnth tdec_b_def)
done
lemma take_Suc_last[elim]: "Suc as \<le> length xs \<Longrightarrow>
take (Suc as) xs = take as xs @ [xs ! as]"
proof(induct xs arbitrary: as)
case (Cons a xs)
then show ?case by ( simp, cases as;simp)
qed simp
lemma concat_suc: "Suc as \<le> length xs \<Longrightarrow>
concat (take (Suc as) xs) = concat (take as xs) @ xs! as"
apply(subgoal_tac "take (Suc as) xs = take as xs @ [xs ! as]", simp)
by auto
lemma concat_drop_suc_iff:
"Suc n < length tps \<Longrightarrow> concat (drop (Suc n) tps) =
tps ! Suc n @ concat (drop (Suc (Suc n)) tps)"
proof(induct tps arbitrary: n)
case (Cons a tps)
then show ?case
apply(cases tps, simp, simp)
apply(cases n, simp, simp)
done
qed simp
declare append_assoc[simp del]
lemma tm_append:
"\<lbrakk>n < length tps; tp = tps ! n\<rbrakk> \<Longrightarrow>
\<exists> tp1 tp2. concat tps = tp1 @ tp @ tp2 \<and> tp1 =
concat (take n tps) \<and> tp2 = concat (drop (Suc n) tps)"
apply(rule_tac x = "concat (take n tps)" in exI)
apply(rule_tac x = "concat (drop (Suc n) tps)" in exI)
apply(auto)
proof(induct n)
case 0
then show ?case by(cases tps; simp)
next
case (Suc n)
then show ?case
apply(subgoal_tac "concat (take n tps) @ (tps ! n) =
concat (take (Suc n) tps)")
apply(simp only: append_assoc[THEN sym], simp only: append_assoc)
apply(subgoal_tac " concat (drop (Suc n) tps) = tps ! Suc n @
concat (drop (Suc (Suc n)) tps)")
apply (metis append_take_drop_id concat_append)
apply(rule concat_drop_suc_iff,force)
by (simp add: concat_suc)
qed
declare append_assoc[simp]
lemma length_tms_of[simp]: "length (tms_of aprog) = length aprog"
apply(auto simp: tms_of.simps tpairs_of.simps)
done
lemma ci_nth:
"\<lbrakk>ly = layout_of aprog;
abc_fetch as aprog = Some ins\<rbrakk>
\<Longrightarrow> ci ly (start_of ly as) ins = tms_of aprog ! as"
apply(simp add: tms_of.simps tpairs_of.simps
abc_fetch.simps del: map_append split: if_splits)
done
lemma t_split:"\<lbrakk>
ly = layout_of aprog;
abc_fetch as aprog = Some ins\<rbrakk>
\<Longrightarrow> \<exists> tp1 tp2. concat (tms_of aprog) =
tp1 @ (ci ly (start_of ly as) ins) @ tp2
\<and> tp1 = concat (take as (tms_of aprog)) \<and>
tp2 = concat (drop (Suc as) (tms_of aprog))"
apply(insert tm_append[of "as" "tms_of aprog"
"ci ly (start_of ly as) ins"], simp)
apply(subgoal_tac "ci ly (start_of ly as) ins = (tms_of aprog) ! as")
apply(subgoal_tac "length (tms_of aprog) = length aprog")
apply(simp add: abc_fetch.simps split: if_splits, simp)
apply(intro ci_nth, auto)
done
lemma div_apart: "\<lbrakk>x mod (2::nat) = 0; y mod 2 = 0\<rbrakk>
\<Longrightarrow> (x + y) div 2 = x div 2 + y div 2"
by(auto)
lemma length_layout_of[simp]: "length (layout_of aprog) = length aprog"
by(auto simp: layout_of.simps)
lemma length_tms_of_elem_even[intro]: "n < length ap \<Longrightarrow> length (tms_of ap ! n) mod 2 = 0"
apply(cases "ap ! n")
by (auto simp: tms_of.simps tpairs_of.simps ci.simps length_findnth tinc_b_def tdec_b_def)
lemma compile_mod2: "length (concat (take n (tms_of ap))) mod 2 = 0"
proof(induct n)
case 0
then show ?case by (auto simp add: take_Suc_conv_app_nth)
next
case (Suc n)
hence "n < length (tms_of ap) \<Longrightarrow> is_even (length (concat (take (Suc n) (tms_of ap))))"
unfolding take_Suc_conv_app_nth by fastforce
with Suc show ?case by(cases "n < length (tms_of ap)", auto)
qed
lemma tpa_states:
"\<lbrakk>tp = concat (take as (tms_of ap));
as \<le> length ap\<rbrakk> \<Longrightarrow>
start_of (layout_of ap) as = Suc (length tp div 2)"
proof(induct as arbitrary: tp)
case 0
thus "?case"
by(simp add: start_of.simps)
next
case (Suc as tp)
have ind: "\<And>tp. \<lbrakk>tp = concat (take as (tms_of ap)); as \<le> length ap\<rbrakk> \<Longrightarrow>
start_of (layout_of ap) as = Suc (length tp div 2)" by fact
have tp: "tp = concat (take (Suc as) (tms_of ap))" by fact
have le: "Suc as \<le> length ap" by fact
have a: "start_of (layout_of ap) as = Suc (length (concat (take as (tms_of ap))) div 2)"
using le
by(intro ind, simp_all)
from a tp le show "?case"
apply(simp add: start_of.simps take_Suc_conv_app_nth)
apply(subgoal_tac "length (concat (take as (tms_of ap))) mod 2= 0")
apply(subgoal_tac " length (tms_of ap ! as) mod 2 = 0")
apply(simp add: Abacus.div_apart)
apply(simp add: layout_of.simps ci_length tms_of.simps tpairs_of.simps)
apply(auto intro: compile_mod2)
done
qed
declare fetch.simps[simp]
lemma append_append_fetch:
"\<lbrakk>length tp1 mod 2 = 0; length tp mod 2 = 0;
length tp1 div 2 < a \<and> a \<le> length tp1 div 2 + length tp div 2\<rbrakk>
\<Longrightarrow>fetch (tp1 @ tp @ tp2) a b = fetch tp (a - length tp1 div 2) b "
apply(subgoal_tac "\<exists> x. a = length tp1 div 2 + x", erule exE)
apply(rename_tac x)
apply(case_tac x, simp)
apply(subgoal_tac "length tp1 div 2 + Suc nat =
Suc (length tp1 div 2 + nat)")
apply(simp only: fetch.simps nth_of.simps, auto)
apply(cases b, simp)
apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp)
apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, simp)
apply(subgoal_tac "2 * (length tp1 div 2) = length tp1", simp)
apply(subgoal_tac "2 * nat < length tp", simp add: nth_append, auto)
apply(auto simp: nth_append)
apply(rule_tac x = "a - length tp1 div 2" in exI, simp)
done
lemma step_eq_fetch':
assumes layout: "ly = layout_of ap"
and compile: "tp = tm_of ap"
and fetch: "abc_fetch as ap = Some ins"
and range1: "s \<ge> start_of ly as"
and range2: "s < start_of ly (Suc as)"
shows "fetch tp s b = fetch (ci ly (start_of ly as) ins)
(Suc s - start_of ly as) b "
proof -
have "\<exists>tp1 tp2. concat (tms_of ap) = tp1 @ ci ly (start_of ly as) ins @ tp2 \<and>
tp1 = concat (take as (tms_of ap)) \<and> tp2 = concat (drop (Suc as) (tms_of ap))"
using assms
by(intro t_split, simp_all)
then obtain tp1 tp2 where a: "concat (tms_of ap) = tp1 @ ci ly (start_of ly as) ins @ tp2 \<and>
tp1 = concat (take as (tms_of ap)) \<and> tp2 = concat (drop (Suc as) (tms_of ap))" by blast
then have b: "start_of (layout_of ap) as = Suc (length tp1 div 2)"
using fetch
by(intro tpa_states, simp, simp add: abc_fetch.simps split: if_splits)
have "fetch (tp1 @ (ci ly (start_of ly as) ins) @ tp2) s b =
fetch (ci ly (start_of ly as) ins) (s - length tp1 div 2) b"
proof(intro append_append_fetch)
show "length tp1 mod 2 = 0"
using a
by(auto, rule_tac compile_mod2)
next
show "length (ci ly (start_of ly as) ins) mod 2 = 0"
by(cases ins, auto simp: ci.simps length_findnth tinc_b_def tdec_b_def)
next
show "length tp1 div 2 < s \<and> s \<le>
length tp1 div 2 + length (ci ly (start_of ly as) ins) div 2"
proof -
have "length (ci ly (start_of ly as) ins) div 2 = length_of ins"
using ci_length by simp
moreover have "start_of ly (Suc as) = start_of ly as + length_of ins"
using fetch layout
apply(simp add: start_of.simps abc_fetch.simps List.take_Suc_conv_app_nth
split: if_splits)
apply(simp add: layout_of.simps)
done
ultimately show "?thesis"
using b layout range1 range2
apply(simp)
done
qed
qed
thus "?thesis"
using b layout a compile
apply(simp add: tm_of.simps)
done
qed
lemma step_eq_fetch:
assumes layout: "ly = layout_of ap"
and compile: "tp = tm_of ap"
and abc_fetch: "abc_fetch as ap = Some ins"
and fetch: "fetch (ci ly (start_of ly as) ins)
(Suc s - start_of ly as) b = (ac, ns)"
and notfinal: "ns \<noteq> 0"
shows "fetch tp s b = (ac, ns)"
proof -
have "s \<ge> start_of ly as"
proof(cases "s \<ge> start_of ly as")
case True thus "?thesis" by simp
next
case False
have "\<not> start_of ly as \<le> s" by fact
then have "Suc s - start_of ly as = 0"
by arith
then have "fetch (ci ly (start_of ly as) ins)
(Suc s - start_of ly as) b = (Nop, 0)"
by(simp add: fetch.simps)
with notfinal fetch show "?thesis"
by(simp)
qed
moreover have "s < start_of ly (Suc as)"
proof(cases "s < start_of ly (Suc as)")
case True thus "?thesis" by simp
next
case False
have h: "\<not> s < start_of ly (Suc as)"
by fact
then have "s > start_of ly as"
using abc_fetch layout
apply(simp add: start_of.simps abc_fetch.simps split: if_splits)
apply(simp add: List.take_Suc_conv_app_nth, auto)
apply(subgoal_tac "layout_of ap ! as > 0")
apply arith
apply(simp add: layout_of.simps)
apply(cases "ap!as", auto simp: length_of.simps)
done
from this and h have "fetch (ci ly (start_of ly as) ins) (Suc s - start_of ly as) b = (Nop, 0)"
using abc_fetch layout
apply(cases b;cases ins)
apply(simp_all add:Suc_diff_le start_of_Suc2 start_of_Suc1 start_of_Suc3)
by (simp_all only: length_ci_inc length_ci_dec length_ci_goto, auto)
from fetch and notfinal this show "?thesis"by simp
qed
ultimately show "?thesis"
using assms
by(drule_tac b= b and ins = ins in step_eq_fetch', auto)
qed
lemma step_eq_in:
assumes layout: "ly = layout_of ap"
and compile: "tp = tm_of ap"
and fetch: "abc_fetch as ap = Some ins"
and exec: "step (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1)
= (s', l', r')"
and notfinal: "s' \<noteq> 0"
shows "step (s, l, r) (tp, 0) = (s', l', r')"
using assms
apply(simp add: step.simps)
apply(cases "fetch (ci (layout_of ap) (start_of (layout_of ap) as) ins)
(Suc s - start_of (layout_of ap) as) (read r)", simp)
using layout
apply(drule_tac s = s and b = "read r" and ac = a in step_eq_fetch, auto)
done
lemma steps_eq_in:
assumes layout: "ly = layout_of ap"
and compile: "tp = tm_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
and fetch: "abc_fetch as ap = Some ins"
and exec: "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp
= (s', l', r')"
and notfinal: "s' \<noteq> 0"
shows "steps (s, l, r) (tp, 0) stp = (s', l', r')"
using exec notfinal
proof(induct stp arbitrary: s' l' r', simp add: steps.simps)
fix stp s' l' r'
assume ind:
"\<And>s' l' r'. \<lbrakk>steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp = (s', l', r'); s' \<noteq> 0\<rbrakk>
\<Longrightarrow> steps (s, l, r) (tp, 0) stp = (s', l', r')"
and h: "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) (Suc stp) = (s', l', r')" "s' \<noteq> 0"
obtain s1 l1 r1 where g: "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp =
(s1, l1, r1)"
apply(cases "steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp") by blast
moreover hence "s1 \<noteq> 0"
using h
apply(simp add: step_red)
apply(cases "0 < s1", simp_all)
done
ultimately have "steps (s, l, r) (tp, 0) stp = (s1, l1, r1)"
apply(rule_tac ind, auto)
done
thus "steps (s, l, r) (tp, 0) (Suc stp) = (s', l', r')"
using h g assms
apply(simp add: step_red)
apply(rule_tac step_eq_in, auto)
done
qed
lemma tm_append_fetch_first:
"\<lbrakk>fetch A s b = (ac, ns); ns \<noteq> 0\<rbrakk> \<Longrightarrow>
fetch (A @ B) s b = (ac, ns)"
by(cases b;cases s;force simp: fetch.simps nth_append split: if_splits)
lemma tm_append_first_step_eq:
assumes "step (s, l, r) (A, off) = (s', l', r')"
and "s' \<noteq> 0"
shows "step (s, l, r) (A @ B, off) = (s', l', r')"
using assms
apply(simp add: step.simps)
apply(cases "fetch A (s - off) (read r)")
apply(frule_tac B = B and b = "read r" in tm_append_fetch_first, auto)
done
lemma tm_append_first_steps_eq:
assumes "steps (s, l, r) (A, off) stp = (s', l', r')"
and "s' \<noteq> 0"
shows "steps (s, l, r) (A @ B, off) stp = (s', l', r')"
using assms
proof(induct stp arbitrary: s' l' r', simp add: steps.simps)
fix stp s' l' r'
assume ind: "\<And>s' l' r'. \<lbrakk>steps (s, l, r) (A, off) stp = (s', l', r'); s' \<noteq> 0\<rbrakk>
\<Longrightarrow> steps (s, l, r) (A @ B, off) stp = (s', l', r')"
and h: "steps (s, l, r) (A, off) (Suc stp) = (s', l', r')" "s' \<noteq> 0"
obtain sa la ra where a: "steps (s, l, r) (A, off) stp = (sa, la, ra)"
apply(cases "steps (s, l, r) (A, off) stp") by blast
hence "steps (s, l, r) (A @ B, off) stp = (sa, la, ra) \<and> sa \<noteq> 0"
using h ind[of sa la ra]
apply(cases sa, simp_all)
done
thus "steps (s, l, r) (A @ B, off) (Suc stp) = (s', l', r')"
using h a
apply(simp add: step_red)
apply(intro tm_append_first_step_eq, simp_all)
done
qed
lemma tm_append_second_fetch_eq:
assumes
even: "length A mod 2 = 0"
and off: "off = length A div 2"
and fetch: "fetch B s b = (ac, ns)"
and notfinal: "ns \<noteq> 0"
shows "fetch (A @ shift B off) (s + off) b = (ac, ns + off)"
using assms
by(cases b;cases s,auto simp: nth_append shift.simps split: if_splits)
lemma tm_append_second_step_eq:
assumes
exec: "step0 (s, l, r) B = (s', l', r')"
and notfinal: "s' \<noteq> 0"
and off: "off = length A div 2"
and even: "length A mod 2 = 0"
shows "step0 (s + off, l, r) (A @ shift B off) = (s' + off, l', r')"
using assms
apply(simp add: step.simps)
apply(cases "fetch B s (read r)")
apply(frule_tac tm_append_second_fetch_eq, simp_all, auto)
done
lemma tm_append_second_steps_eq:
assumes
exec: "steps (s, l, r) (B, 0) stp = (s', l', r')"
and notfinal: "s' \<noteq> 0"
and off: "off = length A div 2"
and even: "length A mod 2 = 0"
shows "steps (s + off, l, r) (A @ shift B off, 0) stp = (s' + off, l', r')"
using exec notfinal
proof(induct stp arbitrary: s' l' r')
case 0
thus "steps0 (s + off, l, r) (A @ shift B off) 0 = (s' + off, l', r')"
by(simp add: steps.simps)
next
case (Suc stp s' l' r')
have ind: "\<And>s' l' r'. \<lbrakk>steps0 (s, l, r) B stp = (s', l', r'); s' \<noteq> 0\<rbrakk> \<Longrightarrow>
steps0 (s + off, l, r) (A @ shift B off) stp = (s' + off, l', r')"
by fact
have h: "steps0 (s, l, r) B (Suc stp) = (s', l', r')" by fact
have k: "s' \<noteq> 0" by fact
obtain s'' l'' r'' where a: "steps0 (s, l, r) B stp = (s'', l'', r'')"
by (metis prod_cases3)
then have b: "s'' \<noteq> 0"
using h k
by(intro notI, auto)
from a b have c: "steps0 (s + off, l, r) (A @ shift B off) stp = (s'' + off, l'', r'')"
by(erule_tac ind, simp)
from c b h a k assms show "?case"
by(auto intro:tm_append_second_step_eq)
qed
lemma tm_append_second_fetch0_eq:
assumes
even: "length A mod 2 = 0"
and off: "off = length A div 2"
and fetch: "fetch B s b = (ac, 0)"
and notfinal: "s \<noteq> 0"
shows "fetch (A @ shift B off) (s + off) b = (ac, 0)"
using assms
apply(cases b;cases s)
apply(auto simp: fetch.simps nth_append shift.simps split: if_splits)
done
lemma tm_append_second_halt_eq:
assumes
exec: "steps (Suc 0, l, r) (B, 0) stp = (0, l', r')"
and wf_B: "tm_wf (B, 0)"
and off: "off = length A div 2"
and even: "length A mod 2 = 0"
shows "steps (Suc off, l, r) (A @ shift B off, 0) stp = (0, l', r')"
proof -
have "\<exists>n. \<not> is_final (steps0 (1, l, r) B n) \<and> steps0 (1, l, r) B (Suc n) = (0, l', r')"
using exec by(rule_tac before_final, simp)
then obtain n where a:
"\<not> is_final (steps0 (1, l, r) B n) \<and> steps0 (1, l, r) B (Suc n) = (0, l', r')" ..
obtain s'' l'' r'' where b: "steps0 (1, l, r) B n = (s'', l'', r'') \<and> s'' >0"
using a
by(cases "steps0 (1, l, r) B n", auto)
have c: "steps (Suc 0 + off, l, r) (A @ shift B off, 0) n = (s'' + off, l'', r'')"
using a b assms
by(rule_tac tm_append_second_steps_eq, simp_all)
obtain ac where d: "fetch B s'' (read r'') = (ac, 0)"
using b a
by(cases "fetch B s'' (read r'')", auto simp: step_red step.simps)
then have "fetch (A @ shift B off) (s'' + off) (read r'') = (ac, 0)"
using assms b
by(rule_tac tm_append_second_fetch0_eq, simp_all)
then have e: "steps (Suc 0 + off, l, r) (A @ shift B off, 0) (Suc n) = (0, l', r')"
using a b assms c d
by(simp add: step_red step.simps)
from a have "n < stp"
using exec
proof(cases "n < stp")
case True thus "?thesis" by simp
next
case False
have "\<not> n < stp" by fact
then obtain d where "n = stp + d"
by (metis add.comm_neutral less_imp_add_positive nat_neq_iff)
thus "?thesis"
using a e exec
by(simp)
qed
then obtain d where "stp = Suc n + d"
by(metis add_Suc less_iff_Suc_add)
thus "?thesis"
using e
by(simp only: steps_add, simp)
qed
lemma tm_append_steps:
assumes
aexec: "steps (s, l, r) (A, 0) stpa = (Suc (length A div 2), la, ra)"
and bexec: "steps (Suc 0, la, ra) (B, 0) stpb = (sb, lb, rb)"
and notfinal: "sb \<noteq> 0"
and off: "off = length A div 2"
and even: "length A mod 2 = 0"
shows "steps (s, l, r) (A @ shift B off, 0) (stpa + stpb) = (sb + off, lb, rb)"
proof -
have "steps (s, l, r) (A@shift B off, 0) stpa = (Suc (length A div 2), la, ra)"
apply(intro tm_append_first_steps_eq)
apply(auto simp: assms)
done
moreover have "steps (1 + off, la, ra) (A @ shift B off, 0) stpb = (sb + off, lb, rb)"
apply(intro tm_append_second_steps_eq)
apply(auto simp: assms bexec)
done
ultimately show "steps (s, l, r) (A @ shift B off, 0) (stpa + stpb) = (sb + off, lb, rb)"
apply(simp add: steps_add off)
done
qed
subsection \<open>Crsp of Inc\<close>
fun at_begin_fst_bwtn :: "inc_inv_t"
where
"at_begin_fst_bwtn (as, lm) (s, l, r) ires =
(\<exists> lm1 tn rn. lm1 = (lm @ 0\<up>tn) \<and> length lm1 = s \<and>
(if lm1 = [] then l = Bk # Bk # ires
else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = Bk\<up>rn)"
fun at_begin_fst_awtn :: "inc_inv_t"
where
"at_begin_fst_awtn (as, lm) (s, l, r) ires =
(\<exists> lm1 tn rn. lm1 = (lm @ 0\<up>tn) \<and> length lm1 = s \<and>
(if lm1 = [] then l = Bk # Bk # ires
else l = [Bk]@<rev lm1>@Bk#Bk#ires) \<and> r = [Oc]@Bk\<up>rn)"
fun at_begin_norm :: "inc_inv_t"
where
"at_begin_norm (as, lm) (s, l, r) ires=
(\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and> length lm1 = s \<and>
(if lm1 = [] then l = Bk # Bk # ires
else l = Bk # <rev lm1> @ Bk # Bk # ires ) \<and> r = <lm2>@Bk\<up>rn)"
fun in_middle :: "inc_inv_t"
where
"in_middle (as, lm) (s, l, r) ires =
(\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<up>tn = lm1 @ [m] @ lm2
\<and> length lm1 = s \<and> m + 1 = ml + mr \<and>
ml \<noteq> 0 \<and> tn = s + 1 - length lm \<and>
(if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
else l = Oc\<up>ml@[Bk]@<rev lm1>@
Bk # Bk # ires) \<and> (r = Oc\<up>mr @ [Bk] @ <lm2>@ Bk\<up>rn \<or>
(lm2 = [] \<and> r = Oc\<up>mr))
)"
fun inv_locate_a :: "inc_inv_t"
where "inv_locate_a (as, lm) (s, l, r) ires =
(at_begin_norm (as, lm) (s, l, r) ires \<or>
at_begin_fst_bwtn (as, lm) (s, l, r) ires \<or>
at_begin_fst_awtn (as, lm) (s, l, r) ires
)"
fun inv_locate_b :: "inc_inv_t"
where "inv_locate_b (as, lm) (s, l, r) ires =
(in_middle (as, lm) (s, l, r)) ires "
fun inv_after_write :: "inc_inv_t"
where "inv_after_write (as, lm) (s, l, r) ires =
(\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and>
(if lm1 = [] then l = Oc\<up>m @ Bk # Bk # ires
else Oc # l = Oc\<up>Suc m@ Bk # <rev lm1> @
Bk # Bk # ires) \<and> r = [Oc] @ <lm2> @ Bk\<up>rn)"
fun inv_after_move :: "inc_inv_t"
where "inv_after_move (as, lm) (s, l, r) ires =
(\<exists> rn m lm1 lm2. lm = lm1 @ m # lm2 \<and>
(if lm1 = [] then l = Oc\<up>Suc m @ Bk # Bk # ires
else l = Oc\<up>Suc m@ Bk # <rev lm1> @ Bk # Bk # ires) \<and>
r = <lm2> @ Bk\<up>rn)"
fun inv_after_clear :: "inc_inv_t"
where "inv_after_clear (as, lm) (s, l, r) ires =
(\<exists> rn m lm1 lm2 r'. lm = lm1 @ m # lm2 \<and>
(if lm1 = [] then l = Oc\<up>Suc m @ Bk # Bk # ires
else l = Oc\<up>Suc m @ Bk # <rev lm1> @ Bk # Bk # ires) \<and>
r = Bk # r' \<and> Oc # r' = <lm2> @ Bk\<up>rn)"
fun inv_on_right_moving :: "inc_inv_t"
where "inv_on_right_moving (as, lm) (s, l, r) ires =
(\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>
ml + mr = m \<and>
(if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
else l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>
((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or>
(r = Oc\<up>mr \<and> lm2 = [])))"
fun inv_on_left_moving_norm :: "inc_inv_t"
where "inv_on_left_moving_norm (as, lm) (s, l, r) ires =
(\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>
ml + mr = Suc m \<and> mr > 0 \<and> (if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
else l = Oc\<up>ml @ Bk # <rev lm1> @ Bk # Bk # ires)
\<and> (r = Oc\<up>mr @ Bk # <lm2> @ Bk\<up>rn \<or>
(lm2 = [] \<and> r = Oc\<up>mr)))"
fun inv_on_left_moving_in_middle_B:: "inc_inv_t"
where "inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires =
(\<exists> lm1 lm2 rn. lm = lm1 @ lm2 \<and>
(if lm1 = [] then l = Bk # ires
else l = <rev lm1> @ Bk # Bk # ires) \<and>
r = Bk # <lm2> @ Bk\<up>rn)"
fun inv_on_left_moving :: "inc_inv_t"
where "inv_on_left_moving (as, lm) (s, l, r) ires =
(inv_on_left_moving_norm (as, lm) (s, l, r) ires \<or>
inv_on_left_moving_in_middle_B (as, lm) (s, l, r) ires)"
fun inv_check_left_moving_on_leftmost :: "inc_inv_t"
where "inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires =
(\<exists> rn. l = ires \<and> r = [Bk, Bk] @ <lm> @ Bk\<up>rn)"
fun inv_check_left_moving_in_middle :: "inc_inv_t"
where "inv_check_left_moving_in_middle (as, lm) (s, l, r) ires =
(\<exists> lm1 lm2 r' rn. lm = lm1 @ lm2 \<and>
(Oc # l = <rev lm1> @ Bk # Bk # ires) \<and> r = Oc # Bk # r' \<and>
r' = <lm2> @ Bk\<up>rn)"
fun inv_check_left_moving :: "inc_inv_t"
where "inv_check_left_moving (as, lm) (s, l, r) ires =
(inv_check_left_moving_on_leftmost (as, lm) (s, l, r) ires \<or>
inv_check_left_moving_in_middle (as, lm) (s, l, r) ires)"
fun inv_after_left_moving :: "inc_inv_t"
where "inv_after_left_moving (as, lm) (s, l, r) ires=
(\<exists> rn. l = Bk # ires \<and> r = Bk # <lm> @ Bk\<up>rn)"
fun inv_stop :: "inc_inv_t"
where "inv_stop (as, lm) (s, l, r) ires=
(\<exists> rn. l = Bk # Bk # ires \<and> r = <lm> @ Bk\<up>rn)"
lemma halt_lemma2':
"\<lbrakk>wf LE; \<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow>
(Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE)); Q (f 0)\<rbrakk>
\<Longrightarrow> \<exists> n. P (f n)"
apply(intro exCI, simp)
apply(subgoal_tac "\<forall> n. Q (f n)")
apply(drule_tac f = f in wf_inv_image)
apply(erule wf_induct)
apply(auto)
apply(rename_tac n,induct_tac n; simp)
done
lemma halt_lemma2'':
"\<lbrakk>P (f n); \<not> P (f (0::nat))\<rbrakk> \<Longrightarrow>
\<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))"
apply(induct n rule: nat_less_induct, auto)
done
lemma halt_lemma2''':
"\<lbrakk>\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> LE;
Q (f 0); \<forall>i<na. \<not> P (f i)\<rbrakk> \<Longrightarrow> Q (f na)"
apply(induct na, simp, simp)
done
lemma halt_lemma2:
"\<lbrakk>wf LE;
Q (f 0); \<not> P (f 0);
\<forall> n. ((\<not> P (f n) \<and> Q (f n)) \<longrightarrow> (Q (f (Suc n)) \<and> (f (Suc n), (f n)) \<in> LE))\<rbrakk>
\<Longrightarrow> \<exists> n. P (f n) \<and> Q (f n)"
apply(insert halt_lemma2' [of LE P f Q], simp, erule_tac exE)
apply(subgoal_tac "\<exists> n. (P (f n) \<and> (\<forall> i < n. \<not> P (f i)))")
apply(erule_tac exE)+
apply(rename_tac n na)
apply(rule_tac x = na in exI, auto)
apply(rule halt_lemma2''', simp, simp, simp)
apply(erule_tac halt_lemma2'', simp)
done
fun findnth_inv :: "layout \<Rightarrow> nat \<Rightarrow> inc_inv_t"
where
"findnth_inv ly n (as, lm) (s, l, r) ires =
(if s = 0 then False
else if s \<le> Suc (2*n) then
if s mod 2 = 1 then inv_locate_a (as, lm) ((s - 1) div 2, l, r) ires
else inv_locate_b (as, lm) ((s - 1) div 2, l, r) ires
else False)"
fun findnth_state :: "config \<Rightarrow> nat \<Rightarrow> nat"
where
"findnth_state (s, l, r) n = (Suc (2*n) - s)"
fun findnth_step :: "config \<Rightarrow> nat \<Rightarrow> nat"
where
"findnth_step (s, l, r) n =
(if s mod 2 = 1 then
(if (r \<noteq> [] \<and> hd r = Oc) then 0
else 1)
else length r)"
fun findnth_measure :: "config \<times> nat \<Rightarrow> nat \<times> nat"
where
"findnth_measure (c, n) =
(findnth_state c n, findnth_step c n)"
definition lex_pair :: "((nat \<times> nat) \<times> nat \<times> nat) set"
where
"lex_pair \<equiv> less_than <*lex*> less_than"
definition findnth_LE :: "((config \<times> nat) \<times> (config \<times> nat)) set"
where
"findnth_LE \<equiv> (inv_image lex_pair findnth_measure)"
lemma wf_findnth_LE: "wf findnth_LE"
by(auto simp: findnth_LE_def lex_pair_def)
declare findnth_inv.simps[simp del]
lemma x_is_2n_arith[simp]:
"\<lbrakk>x < Suc (Suc (2 * n)); Suc x mod 2 = Suc 0; \<not> x < 2 * n\<rbrakk>
\<Longrightarrow> x = 2*n"
by arith
lemma between_sucs:"x < Suc n \<Longrightarrow> \<not> x < n \<Longrightarrow> x = n" by auto
lemma fetch_findnth[simp]:
"\<lbrakk>0 < a; a < Suc (2 * n); a mod 2 = Suc 0\<rbrakk> \<Longrightarrow> fetch (findnth n) a Oc = (R, Suc a)"
"\<lbrakk>0 < a; a < Suc (2 * n); a mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> fetch (findnth n) a Oc = (R, a)"
"\<lbrakk>0 < a; a < Suc (2 * n); a mod 2 \<noteq> Suc 0\<rbrakk> \<Longrightarrow> fetch (findnth n) a Bk = (R, Suc a)"
"\<lbrakk>0 < a; a < Suc (2 * n); a mod 2 = Suc 0\<rbrakk> \<Longrightarrow> fetch (findnth n) a Bk = (W1, a)"
by(cases a;induct n;force simp: length_findnth nth_append dest!:between_sucs)+
declare at_begin_norm.simps[simp del] at_begin_fst_bwtn.simps[simp del]
at_begin_fst_awtn.simps[simp del] in_middle.simps[simp del]
abc_lm_s.simps[simp del] abc_lm_v.simps[simp del]
ci.simps[simp del] inv_after_move.simps[simp del]
inv_on_left_moving_norm.simps[simp del]
inv_on_left_moving_in_middle_B.simps[simp del]
inv_after_clear.simps[simp del]
inv_after_write.simps[simp del] inv_on_left_moving.simps[simp del]
inv_on_right_moving.simps[simp del]
inv_check_left_moving.simps[simp del]
inv_check_left_moving_in_middle.simps[simp del]
inv_check_left_moving_on_leftmost.simps[simp del]
inv_after_left_moving.simps[simp del]
inv_stop.simps[simp del] inv_locate_a.simps[simp del]
inv_locate_b.simps[simp del]
lemma replicate_once[intro]: "\<exists>rn. [Bk] = Bk \<up> rn"
by (metis replicate.simps)
lemma at_begin_norm_Bk[intro]: "at_begin_norm (as, am) (q, aaa, []) ires
\<Longrightarrow> at_begin_norm (as, am) (q, aaa, [Bk]) ires"
apply(simp add: at_begin_norm.simps)
by fastforce
lemma at_begin_fst_bwtn_Bk[intro]: "at_begin_fst_bwtn (as, am) (q, aaa, []) ires
\<Longrightarrow> at_begin_fst_bwtn (as, am) (q, aaa, [Bk]) ires"
apply(simp only: at_begin_fst_bwtn.simps)
using replicate_once by blast
lemma at_begin_fst_awtn_Bk[intro]: "at_begin_fst_awtn (as, am) (q, aaa, []) ires
\<Longrightarrow> at_begin_fst_awtn (as, am) (q, aaa, [Bk]) ires"
apply(auto simp: at_begin_fst_awtn.simps)
done
lemma inv_locate_a_Bk[intro]: "inv_locate_a (as, am) (q, aaa, []) ires
\<Longrightarrow> inv_locate_a (as, am) (q, aaa, [Bk]) ires"
apply(simp only: inv_locate_a.simps)
apply(erule disj_forward)
defer
apply(erule disj_forward, auto)
done
lemma locate_a_2_locate_a[simp]: "inv_locate_a (as, am) (q, aaa, Bk # xs) ires
\<Longrightarrow> inv_locate_a (as, am) (q, aaa, Oc # xs) ires"
apply(simp only: inv_locate_a.simps at_begin_norm.simps
at_begin_fst_bwtn.simps at_begin_fst_awtn.simps)
apply(erule_tac disjE, erule exE, erule exE, erule exE,
rule disjI2, rule disjI2)
defer
apply(erule_tac disjE, erule exE, erule exE,
erule exE, rule disjI2, rule disjI2)
prefer 2
apply(simp)
proof-
fix lm1 tn rn
assume k: "lm1 = am @ 0\<up>tn \<and> length lm1 = q \<and> (if lm1 = [] then aaa = Bk # Bk #
ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Bk # xs = Bk\<up>rn"
thus "\<exists>lm1 tn rn. lm1 = am @ 0 \<up> tn \<and> length lm1 = q \<and>
(if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and> Oc # xs = [Oc] @ Bk \<up> rn"
(is "\<exists>lm1 tn rn. ?P lm1 tn rn")
proof -
from k have "?P lm1 tn (rn - 1)"
by (auto simp: Cons_replicate_eq)
thus ?thesis by blast
qed
next
fix lm1 lm2 rn
assume h1: "am = lm1 @ lm2 \<and> length lm1 = q \<and> (if lm1 = []
then aaa = Bk # Bk # ires else aaa = Bk # <rev lm1> @ Bk # Bk # ires) \<and>
Bk # xs = <lm2> @ Bk\<up>rn"
from h1 have h2: "lm2 = []"
apply(auto split: if_splits;cases lm2;simp add: tape_of_nl_cons split: if_splits)
done
from h1 and h2 show "\<exists>lm1 tn rn. lm1 = am @ 0\<up>tn \<and> length lm1 = q \<and>
(if lm1 = [] then aaa = Bk # Bk # ires else aaa = [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>
Oc # xs = [Oc] @ Bk\<up>rn"
(is "\<exists>lm1 tn rn. ?P lm1 tn rn")
proof -
from h1 and h2 have "?P lm1 0 (rn - 1)"
apply(auto simp:tape_of_nat_def)
by(cases rn, simp, simp)
thus ?thesis by blast
qed
qed
lemma inv_locate_a[simp]: "inv_locate_a (as, am) (q, aaa, []) ires \<Longrightarrow>
inv_locate_a (as, am) (q, aaa, [Oc]) ires"
apply(insert locate_a_2_locate_a [of as am q aaa "[]"])
apply(subgoal_tac "inv_locate_a (as, am) (q, aaa, [Bk]) ires", auto)
done
(*inv: from locate_b to locate_b*)
lemma inv_locate_b[simp]: "inv_locate_b (as, am) (q, aaa, Oc # xs) ires
\<Longrightarrow> inv_locate_b (as, am) (q, Oc # aaa, xs) ires"
apply(simp only: inv_locate_b.simps in_middle.simps)
apply(erule exE)+
apply(rename_tac lm1 lm2 tn m ml mr rn)
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
rule_tac x = tn in exI, rule_tac x = m in exI)
apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - 1" in exI,
rule_tac x = rn in exI)
- apply(case_tac mr, simp_all, auto)
+ apply(case_tac mr)
+ apply simp_all
done
lemma tape_nat[simp]: "<[x::nat]> = Oc\<up>(Suc x)"
apply(simp add: tape_of_nat_def tape_of_list_def)
done
lemma inv_locate[simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires; \<exists>n. xs = Bk\<up>n\<rbrakk>
\<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
apply(simp add: inv_locate_b.simps inv_locate_a.simps)
apply(rule_tac disjI2, rule_tac disjI1)
apply(simp only: in_middle.simps at_begin_fst_bwtn.simps)
apply(erule_tac exE)+
apply(rename_tac lm1 n lm2 tn m ml mr rn)
apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = tn in exI, simp split: if_splits)
apply(case_tac mr, simp_all)
- apply(cases "length am", simp_all, case_tac tn, simp_all)
+ apply(cases "length am", simp_all)
apply(case_tac lm2, simp_all add: tape_of_nl_cons split: if_splits)
apply(cases am, simp_all)
apply(case_tac n, simp_all)
apply(case_tac n, simp_all)
apply(case_tac mr, simp_all)
apply(case_tac lm2, simp_all add: tape_of_nl_cons split: if_splits, auto)
apply(case_tac [!] n, simp_all)
done
lemma repeat_Bk_no_Oc[simp]: "(Oc # r = Bk \<up> rn) = False"
apply(cases rn, simp_all)
done
lemma repeat_Bk[simp]: "(\<exists>rna. Bk \<up> rn = Bk # Bk \<up> rna) \<or> rn = 0"
apply(cases rn, auto)
done
lemma inv_locate_b_Oc_via_a[simp]:
assumes "inv_locate_a (as, lm) (q, l, Oc # r) ires"
shows "inv_locate_b (as, lm) (q, Oc # l, r) ires"
proof -
show ?thesis using assms unfolding inv_locate_a.simps inv_locate_b.simps
at_begin_norm.simps at_begin_fst_bwtn.simps at_begin_fst_awtn.simps
apply(simp only:in_middle.simps)
apply(erule disjE, erule exE, erule exE, erule exE)
apply(rename_tac Lm1 Lm2 Rn)
apply(rule_tac x = Lm1 in exI, rule_tac x = "tl Lm2" in exI)
apply(rule_tac x = 0 in exI, rule_tac x = "hd Lm2" in exI)
apply(rule_tac x = 1 in exI, rule_tac x = "hd Lm2" in exI)
apply(case_tac Lm2, force simp: tape_of_nl_cons )
apply(case_tac "tl Lm2", simp_all)
apply(case_tac Rn, auto simp: tape_of_nl_cons )
apply(rename_tac tn rn)
apply(rule_tac x = "lm @ replicate tn 0" in exI,
rule_tac x = "[]" in exI,
rule_tac x = "Suc tn" in exI,
rule_tac x = 0 in exI, auto simp add: replicate_append_same)
apply(rule_tac x = "Suc 0" in exI, auto)
done
qed
lemma length_equal: "xs = ys \<Longrightarrow> length xs = length ys"
by auto
lemma inv_locate_a_Bk_via_b[simp]: "\<lbrakk>inv_locate_b (as, am) (q, aaa, Bk # xs) ires;
\<not> (\<exists>n. xs = Bk\<up>n)\<rbrakk>
\<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
+ supply [[simproc del: defined_all]]
apply(simp add: inv_locate_b.simps inv_locate_a.simps)
apply(rule_tac disjI1)
apply(simp only: in_middle.simps at_begin_norm.simps)
apply(erule_tac exE)+
apply(rename_tac lm1 lm2 tn m ml mr rn)
apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = lm2 in exI, simp)
apply(subgoal_tac "tn = 0", simp , auto split: if_splits)
apply(simp add: tape_of_nl_cons)
apply(drule_tac length_equal, simp)
apply(cases "length am", simp_all, erule_tac x = rn in allE, simp)
apply(drule_tac length_equal, simp)
apply(case_tac "(Suc (length lm1) - length am)", simp_all)
apply(case_tac lm2, simp, simp)
done
lemma locate_b_2_a[intro]:
"inv_locate_b (as, am) (q, aaa, Bk # xs) ires
\<Longrightarrow> inv_locate_a (as, am) (Suc q, Bk # aaa, xs) ires"
apply(cases "\<exists> n. xs = Bk\<up>n", simp, simp)
done
lemma inv_locate_b_Bk[simp]: "inv_locate_b (as, am) (q, l, []) ires
\<Longrightarrow> inv_locate_b (as, am) (q, l, [Bk]) ires"
by(force simp add: inv_locate_b.simps in_middle.simps)
(*inv: from locate_b to after_write*)
lemma div_rounding_down[simp]: "(2*q - Suc 0) div 2 = (q - 1)" "(Suc (2*q)) div 2 = q"
by arith+
lemma even_plus_one_odd[simp]: "x mod 2 = 0 \<Longrightarrow> Suc x mod 2 = Suc 0"
by arith
lemma odd_plus_one_even[simp]: "x mod 2 = Suc 0 \<Longrightarrow> Suc x mod 2 = 0"
by arith
lemma locate_b_2_locate_a[simp]:
"\<lbrakk>q > 0; inv_locate_b (as, am) (q - Suc 0, aaa, Bk # xs) ires\<rbrakk>
\<Longrightarrow> inv_locate_a (as, am) (q, Bk # aaa, xs) ires"
apply(insert locate_b_2_a [of as am "q - 1" aaa xs ires], simp)
done
(*inv: from locate_b to after_write*)
lemma findnth_inv_layout_of_via_crsp[simp]:
"crsp (layout_of ap) (as, lm) (s, l, r) ires
\<Longrightarrow> findnth_inv (layout_of ap) n (as, lm) (Suc 0, l, r) ires"
by(auto simp: crsp.simps findnth_inv.simps inv_locate_a.simps
at_begin_norm.simps at_begin_fst_awtn.simps at_begin_fst_bwtn.simps)
lemma findnth_correct_pre:
assumes layout: "ly = layout_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
and not0: "n > 0"
and f: "f = (\<lambda> stp. (steps (Suc 0, l, r) (findnth n, 0) stp, n))"
and P: "P = (\<lambda> ((s, l, r), n). s = Suc (2 * n))"
and Q: "Q = (\<lambda> ((s, l, r), n). findnth_inv ly n (as, lm) (s, l, r) ires)"
shows "\<exists> stp. P (f stp) \<and> Q (f stp)"
proof(rule_tac LE = findnth_LE in halt_lemma2)
show "wf findnth_LE" by(intro wf_findnth_LE)
next
show "Q (f 0)"
using crsp layout
apply(simp add: f P Q steps.simps)
done
next
show "\<not> P (f 0)"
using not0
apply(simp add: f P steps.simps)
done
next
have "\<not> P (f na) \<and> Q (f na) \<Longrightarrow> Q (f (Suc na)) \<and> (f (Suc na), f na)
\<in> findnth_LE" for na
proof(simp add: f,
cases "steps (Suc 0, l, r) (findnth n, 0) na", simp add: P)
fix na a b c
assume "a \<noteq> Suc (2 * n) \<and> Q ((a, b, c), n)"
thus "Q (step (a, b, c) (findnth n, 0), n) \<and>
((step (a, b, c) (findnth n, 0), n), (a, b, c), n) \<in> findnth_LE"
apply(cases c, case_tac [2] "hd c")
apply(simp_all add: step.simps findnth_LE_def Q findnth_inv.simps mod_2 lex_pair_def split: if_splits)
apply(auto simp: mod_ex1 mod_ex2)
done
qed
thus "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow>
Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> findnth_LE" by blast
qed
lemma inv_locate_a_via_crsp[simp]:
"crsp ly (as, lm) (s, l, r) ires \<Longrightarrow> inv_locate_a (as, lm) (0, l, r) ires"
apply(auto simp: crsp.simps inv_locate_a.simps at_begin_norm.simps)
done
lemma findnth_correct:
assumes layout: "ly = layout_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
shows "\<exists> stp l' r'. steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r')
\<and> inv_locate_a (as, lm) (n, l', r') ires"
using crsp
apply(cases "n = 0")
apply(rule_tac x = 0 in exI, auto simp: steps.simps)
using assms
apply(drule_tac findnth_correct_pre, auto)
using findnth_inv.simps by auto
fun inc_inv :: "nat \<Rightarrow> inc_inv_t"
where
"inc_inv n (as, lm) (s, l, r) ires =
(let lm' = abc_lm_s lm n (Suc (abc_lm_v lm n)) in
if s = 0 then False
else if s = 1 then
inv_locate_a (as, lm) (n, l, r) ires
else if s = 2 then
inv_locate_b (as, lm) (n, l, r) ires
else if s = 3 then
inv_after_write (as, lm') (s, l, r) ires
else if s = Suc 3 then
inv_after_move (as, lm') (s, l, r) ires
else if s = Suc 4 then
inv_after_clear (as, lm') (s, l, r) ires
else if s = Suc (Suc 4) then
inv_on_right_moving (as, lm') (s, l, r) ires
else if s = Suc (Suc 5) then
inv_on_left_moving (as, lm') (s, l, r) ires
else if s = Suc (Suc (Suc 5)) then
inv_check_left_moving (as, lm') (s, l, r) ires
else if s = Suc (Suc (Suc (Suc 5))) then
inv_after_left_moving (as, lm') (s, l, r) ires
else if s = Suc (Suc (Suc (Suc (Suc 5)))) then
inv_stop (as, lm') (s, l, r) ires
else False)"
fun abc_inc_stage1 :: "config \<Rightarrow> nat"
where
"abc_inc_stage1 (s, l, r) =
(if s = 0 then 0
else if s \<le> 2 then 5
else if s \<le> 6 then 4
else if s \<le> 8 then 3
else if s = 9 then 2
else 1)"
fun abc_inc_stage2 :: "config \<Rightarrow> nat"
where
"abc_inc_stage2 (s, l, r) =
(if s = 1 then 2
else if s = 2 then 1
else if s = 3 then length r
else if s = 4 then length r
else if s = 5 then length r
else if s = 6 then
if r \<noteq> [] then length r
else 1
else if s = 7 then length l
else if s = 8 then length l
else 0)"
fun abc_inc_stage3 :: "config \<Rightarrow> nat"
where
"abc_inc_stage3 (s, l, r) = (
if s = 4 then 4
else if s = 5 then 3
else if s = 6 then
if r \<noteq> [] \<and> hd r = Oc then 2
else 1
else if s = 3 then 0
else if s = 2 then length r
else if s = 1 then
if (r \<noteq> [] \<and> hd r = Oc) then 0
else 1
else 10 - s)"
definition inc_measure :: "config \<Rightarrow> nat \<times> nat \<times> nat"
where
"inc_measure c =
(abc_inc_stage1 c, abc_inc_stage2 c, abc_inc_stage3 c)"
definition lex_triple ::
"((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set"
where "lex_triple \<equiv> less_than <*lex*> lex_pair"
definition inc_LE :: "(config \<times> config) set"
where
"inc_LE \<equiv> (inv_image lex_triple inc_measure)"
declare inc_inv.simps[simp del]
lemma wf_inc_le[intro]: "wf inc_LE"
by(auto simp: inc_LE_def lex_triple_def lex_pair_def)
lemma inv_locate_b_2_after_write[simp]:
assumes "inv_locate_b (as, am) (n, aaa, Bk # xs) ires"
shows "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n))) (s, aaa, Oc # xs) ires"
proof -
from assms show ?thesis
apply(auto simp: in_middle.simps inv_after_write.simps
abc_lm_v.simps abc_lm_s.simps inv_locate_b.simps simp del:split_head_repeat)
apply(rename_tac lm1 lm2 m ml mr rn)
apply(case_tac [!] mr, auto split: if_splits)
apply(rename_tac lm1 lm2 m rn)
apply(rule_tac x = rn in exI, rule_tac x = "Suc m" in exI,
rule_tac x = "lm1" in exI, simp)
apply(rule_tac x = "lm2" in exI)
apply(simp only: Suc_diff_le exp_ind)
by(subgoal_tac "lm2 = []"; force dest:length_equal)
qed
(*inv: from after_write to after_move*)
lemma inv_after_move_Oc_via_write[simp]: "inv_after_write (as, lm) (x, l, Oc # r) ires
\<Longrightarrow> inv_after_move (as, lm) (y, Oc # l, r) ires"
apply(auto simp:inv_after_move.simps inv_after_write.simps split: if_splits)
done
lemma inv_after_write_Suc[simp]: "inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)
)) (x, aaa, Bk # xs) ires = False"
"inv_after_write (as, abc_lm_s am n (Suc (abc_lm_v am n)))
(x, aaa, []) ires = False"
apply(auto simp: inv_after_write.simps )
done
(*inv: from after_move to after_clear*)
lemma inv_after_clear_Bk_via_Oc[simp]: "inv_after_move (as, lm) (s, l, Oc # r) ires
\<Longrightarrow> inv_after_clear (as, lm) (s', l, Bk # r) ires"
apply(auto simp: inv_after_move.simps inv_after_clear.simps split: if_splits)
done
lemma inv_after_move_2_inv_on_left_moving[simp]:
assumes "inv_after_move (as, lm) (s, l, Bk # r) ires"
shows "(l = [] \<longrightarrow>
inv_on_left_moving (as, lm) (s', [], Bk # Bk # r) ires) \<and>
(l \<noteq> [] \<longrightarrow>
inv_on_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)"
proof (cases l)
case (Cons a list)
from assms Cons show ?thesis
apply(simp only: inv_after_move.simps inv_on_left_moving.simps)
apply(rule conjI, force, rule impI, rule disjI1, simp only: inv_on_left_moving_norm.simps)
apply(erule exE)+
apply(rename_tac rn m lm1 lm2)
apply(subgoal_tac "lm2 = []")
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
rule_tac x = m in exI, rule_tac x = m in exI,
rule_tac x = 1 in exI,
rule_tac x = "rn - 1" in exI)
apply (auto split:if_splits)
apply(case_tac [1-2] rn, simp_all)
by(case_tac [!] lm2, simp_all add: tape_of_nl_cons split: if_splits)
next
case Nil thus ?thesis using assms
unfolding inv_after_move.simps inv_on_left_moving.simps
by (auto split:if_splits)
qed
lemma inv_after_move_2_inv_on_left_moving_B[simp]:
"inv_after_move (as, lm) (s, l, []) ires
\<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], [Bk]) ires) \<and>
(l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, [hd l]) ires)"
apply(simp only: inv_after_move.simps inv_on_left_moving.simps)
apply(subgoal_tac "l \<noteq> []", rule conjI, simp, rule impI, rule disjI1,
simp only: inv_on_left_moving_norm.simps)
apply(erule exE)+
apply(rename_tac rn m lm1 lm2)
apply(subgoal_tac "lm2 = []")
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
rule_tac x = m in exI, rule_tac x = m in exI,
rule_tac x = 1 in exI, rule_tac x = "rn - 1" in exI, force)
apply(metis append_Cons list.distinct(1) list.exhaust replicate_Suc tape_of_nl_cons)
apply(metis append_Cons list.distinct(1) replicate_Suc)
done
lemma inv_after_clear_2_inv_on_right_moving[simp]:
"inv_after_clear (as, lm) (x, l, Bk # r) ires
\<Longrightarrow> inv_on_right_moving (as, lm) (y, Bk # l, r) ires"
apply(auto simp: inv_after_clear.simps inv_on_right_moving.simps simp del:split_head_repeat)
apply(rename_tac rn m lm1 lm2)
apply(subgoal_tac "lm2 \<noteq> []")
apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI,
rule_tac x = "hd lm2" in exI, simp del:split_head_repeat)
apply(rule_tac x = 0 in exI, rule_tac x = "hd lm2" in exI)
apply(simp, rule conjI)
apply(case_tac [!] "lm2::nat list", auto)
apply(case_tac rn, auto split: if_splits simp: tape_of_nl_cons)
apply(case_tac [!] rn, simp_all)
done
(*inv: from on_right_moving to on_right_moving*)
lemma inv_on_right_moving_Oc[simp]: "inv_on_right_moving (as, lm) (x, l, Oc # r) ires
\<Longrightarrow> inv_on_right_moving (as, lm) (y, Oc # l, r) ires"
apply(auto simp: inv_on_right_moving.simps)
apply(rename_tac lm1 lm2 ml mr rn)
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
rule_tac x = "ml + mr" in exI, simp)
apply(rule_tac x = "Suc ml" in exI,
rule_tac x = "mr - 1" in exI, simp)
apply (metis One_nat_def Suc_pred cell.distinct(1) empty_replicate list.inject
list.sel(3) neq0_conv self_append_conv2 tl_append2 tl_replicate)
apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI,
rule_tac x = "ml + mr" in exI, simp)
apply(rule_tac x = "Suc ml" in exI,
rule_tac x = "mr - 1" in exI)
apply (auto simp add: Cons_replicate_eq)
done
lemma inv_on_right_moving_2_inv_on_right_moving[simp]:
"inv_on_right_moving (as, lm) (x, l, Bk # r) ires
\<Longrightarrow> inv_after_write (as, lm) (y, l, Oc # r) ires"
apply(auto simp: inv_on_right_moving.simps inv_after_write.simps)
by (metis append.left_neutral append_Cons )
lemma inv_on_right_moving_singleton_Bk[simp]: "inv_on_right_moving (as, lm) (x, l, []) ires\<Longrightarrow>
inv_on_right_moving (as, lm) (y, l, [Bk]) ires"
apply(auto simp: inv_on_right_moving.simps)
by fastforce
(*inv: from on_left_moving to on_left_moving*)
lemma no_inv_on_left_moving_in_middle_B_Oc[simp]: "inv_on_left_moving_in_middle_B (as, lm)
(s, l, Oc # r) ires = False"
by(auto simp: inv_on_left_moving_in_middle_B.simps )
lemma no_inv_on_left_moving_norm_Bk[simp]: "inv_on_left_moving_norm (as, lm) (s, l, Bk # r) ires
= False"
by(auto simp: inv_on_left_moving_norm.simps)
lemma inv_on_left_moving_in_middle_B_Bk[simp]:
"\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires;
hd l = Bk; l \<noteq> []\<rbrakk> \<Longrightarrow>
inv_on_left_moving_in_middle_B (as, lm) (s, tl l, Bk # Oc # r) ires"
apply(cases l, simp, simp)
apply(simp only: inv_on_left_moving_norm.simps
inv_on_left_moving_in_middle_B.simps)
apply(erule_tac exE)+ unfolding tape_of_nl_cons
apply(rename_tac a list lm1 lm2 m ml mr rn)
apply(rule_tac x = lm1 in exI, rule_tac x = "m # lm2" in exI, auto)
apply(auto simp: tape_of_nl_cons split: if_splits)
done
lemma inv_on_left_moving_norm_Oc_Oc[simp]: "\<lbrakk>inv_on_left_moving_norm (as, lm) (s, l, Oc # r) ires;
hd l = Oc; l \<noteq> []\<rbrakk>
\<Longrightarrow> inv_on_left_moving_norm (as, lm)
(s, tl l, Oc # Oc # r) ires"
apply(simp only: inv_on_left_moving_norm.simps)
apply(erule exE)+
apply(rename_tac lm1 lm2 m ml mr rn)
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
rule_tac x = m in exI, rule_tac x = "ml - 1" in exI,
rule_tac x = "Suc mr" in exI, rule_tac x = rn in exI, simp)
apply(case_tac ml, auto simp: split: if_splits)
done
lemma inv_on_left_moving_in_middle_B_Bk_Oc[simp]: "inv_on_left_moving_norm (as, lm) (s, [], Oc # r) ires
\<Longrightarrow> inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # Oc # r) ires"
by(auto simp: inv_on_left_moving_norm.simps
inv_on_left_moving_in_middle_B.simps split: if_splits)
lemma inv_on_left_moving_Oc_cases[simp]:"inv_on_left_moving (as, lm) (s, l, Oc # r) ires
\<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s, [], Bk # Oc # r) ires)
\<and> (l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s, tl l, hd l # Oc # r) ires)"
apply(simp add: inv_on_left_moving.simps)
apply(cases "l \<noteq> []", rule conjI, simp, simp)
apply(cases "hd l", simp, simp, simp)
done
lemma from_on_left_moving_to_check_left_moving[simp]: "inv_on_left_moving_in_middle_B (as, lm)
(s, Bk # list, Bk # r) ires
\<Longrightarrow> inv_check_left_moving_on_leftmost (as, lm)
(s', list, Bk # Bk # r) ires"
apply(simp only: inv_on_left_moving_in_middle_B.simps inv_check_left_moving_on_leftmost.simps)
apply(erule_tac exE)+
apply(rename_tac lm1 lm2 rn)
apply(case_tac "rev lm1", simp_all)
apply(case_tac "tl (rev lm1)", simp_all add: tape_of_nat_def tape_of_list_def)
done
lemma inv_check_left_moving_in_middle_no_Bk[simp]:
"inv_check_left_moving_in_middle (as, lm) (s, l, Bk # r) ires= False"
by(auto simp: inv_check_left_moving_in_middle.simps )
lemma inv_check_left_moving_on_leftmost_Bk_Bk[simp]:
"inv_on_left_moving_in_middle_B (as, lm) (s, [], Bk # r) ires\<Longrightarrow>
inv_check_left_moving_on_leftmost (as, lm) (s', [], Bk # Bk # r) ires"
apply(auto simp: inv_on_left_moving_in_middle_B.simps
inv_check_left_moving_on_leftmost.simps split: if_splits)
done
lemma inv_check_left_moving_on_leftmost_no_Oc[simp]: "inv_check_left_moving_on_leftmost (as, lm)
(s, list, Oc # r) ires= False"
by(auto simp: inv_check_left_moving_on_leftmost.simps split: if_splits)
lemma inv_check_left_moving_in_middle_Oc_Bk[simp]: "inv_on_left_moving_in_middle_B (as, lm)
(s, Oc # list, Bk # r) ires
\<Longrightarrow> inv_check_left_moving_in_middle (as, lm) (s', list, Oc # Bk # r) ires"
apply(auto simp: inv_on_left_moving_in_middle_B.simps
inv_check_left_moving_in_middle.simps split: if_splits)
done
lemma inv_on_left_moving_2_check_left_moving[simp]:
"inv_on_left_moving (as, lm) (s, l, Bk # r) ires
\<Longrightarrow> (l = [] \<longrightarrow> inv_check_left_moving (as, lm) (s', [], Bk # Bk # r) ires)
\<and> (l \<noteq> [] \<longrightarrow>
inv_check_left_moving (as, lm) (s', tl l, hd l # Bk # r) ires)"
by (cases l;cases "hd l", auto simp: inv_on_left_moving.simps inv_check_left_moving.simps)
lemma inv_on_left_moving_norm_no_empty[simp]: "inv_on_left_moving_norm (as, lm) (s, l, []) ires = False"
apply(auto simp: inv_on_left_moving_norm.simps)
done
lemma inv_on_left_moving_no_empty[simp]: "inv_on_left_moving (as, lm) (s, l, []) ires = False"
apply(simp add: inv_on_left_moving.simps)
apply(simp add: inv_on_left_moving_in_middle_B.simps)
done
lemma
inv_check_left_moving_in_middle_2_on_left_moving_in_middle_B[simp]:
assumes "inv_check_left_moving_in_middle (as, lm) (s, Bk # list, Oc # r) ires"
shows "inv_on_left_moving_in_middle_B (as, lm) (s', list, Bk # Oc # r) ires"
using assms
apply(simp only: inv_check_left_moving_in_middle.simps
inv_on_left_moving_in_middle_B.simps)
apply(erule_tac exE)+
apply(rename_tac lm1 lm2 r' rn)
apply(rule_tac x = "rev (tl (rev lm1))" in exI,
rule_tac x = "[hd (rev lm1)] @ lm2" in exI, auto)
apply(case_tac [!] "rev lm1",case_tac [!] "tl (rev lm1)")
apply(simp_all add: tape_of_nat_def tape_of_list_def tape_of_nat_list.simps)
apply(case_tac [1] lm2, auto simp:tape_of_nat_def)
apply(case_tac lm2, auto simp:tape_of_nat_def)
done
lemma inv_check_left_moving_in_middle_Bk_Oc[simp]:
"inv_check_left_moving_in_middle (as, lm) (s, [], Oc # r) ires\<Longrightarrow>
inv_check_left_moving_in_middle (as, lm) (s', [Bk], Oc # r) ires"
apply(auto simp: inv_check_left_moving_in_middle.simps )
done
lemma inv_on_left_moving_norm_Oc_Oc_via_middle[simp]: "inv_check_left_moving_in_middle (as, lm)
(s, Oc # list, Oc # r) ires
\<Longrightarrow> inv_on_left_moving_norm (as, lm) (s', list, Oc # Oc # r) ires"
apply(auto simp: inv_check_left_moving_in_middle.simps
inv_on_left_moving_norm.simps)
apply(rename_tac lm1 lm2 rn)
apply(rule_tac x = "rev (tl (rev lm1))" in exI,
rule_tac x = lm2 in exI, rule_tac x = "hd (rev lm1)" in exI)
apply(rule_tac conjI)
apply(case_tac "rev lm1", simp, simp)
apply(rule_tac x = "hd (rev lm1) - 1" in exI, auto)
apply(rule_tac [!] x = "Suc (Suc 0)" in exI, simp)
apply(case_tac [!] "rev lm1", simp_all)
apply(case_tac [!] "last lm1", simp_all add: tape_of_nl_cons split: if_splits)
done
lemma inv_check_left_moving_Oc_cases[simp]: "inv_check_left_moving (as, lm) (s, l, Oc # r) ires
\<Longrightarrow> (l = [] \<longrightarrow> inv_on_left_moving (as, lm) (s', [], Bk # Oc # r) ires) \<and>
(l \<noteq> [] \<longrightarrow> inv_on_left_moving (as, lm) (s', tl l, hd l # Oc # r) ires)"
apply(cases l;cases "hd l", auto simp: inv_check_left_moving.simps inv_on_left_moving.simps)
done
(*inv: check_left_moving to after_left_moving*)
lemma inv_after_left_moving_Bk_via_check[simp]: "inv_check_left_moving (as, lm) (s, l, Bk # r) ires
\<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, r) ires"
apply(auto simp: inv_check_left_moving.simps
inv_check_left_moving_on_leftmost.simps inv_after_left_moving.simps)
done
lemma inv_after_left_moving_Bk_empty_via_check[simp]:"inv_check_left_moving (as, lm) (s, l, []) ires
\<Longrightarrow> inv_after_left_moving (as, lm) (s', Bk # l, []) ires"
by(simp add: inv_check_left_moving.simps
inv_check_left_moving_in_middle.simps
inv_check_left_moving_on_leftmost.simps)
(*inv: after_left_moving to inv_stop*)
lemma inv_stop_Bk_move[simp]: "inv_after_left_moving (as, lm) (s, l, Bk # r) ires
\<Longrightarrow> inv_stop (as, lm) (s', Bk # l, r) ires"
apply(auto simp: inv_after_left_moving.simps inv_stop.simps)
done
lemma inv_stop_Bk_empty[simp]: "inv_after_left_moving (as, lm) (s, l, []) ires
\<Longrightarrow> inv_stop (as, lm) (s', Bk # l, []) ires"
by(auto simp: inv_after_left_moving.simps)
(*inv: stop to stop*)
lemma inv_stop_indep_fst[simp]: "inv_stop (as, lm) (x, l, r) ires \<Longrightarrow>
inv_stop (as, lm) (y, l, r) ires"
apply(simp add: inv_stop.simps)
done
lemma inv_after_clear_no_Oc[simp]: "inv_after_clear (as, lm) (s, aaa, Oc # xs) ires= False"
apply(auto simp: inv_after_clear.simps )
done
lemma inv_after_left_moving_no_Oc[simp]:
"inv_after_left_moving (as, lm) (s, aaa, Oc # xs) ires = False"
by(auto simp: inv_after_left_moving.simps )
lemma inv_after_clear_Suc_nonempty[simp]:
"inv_after_clear (as, abc_lm_s lm n (Suc (abc_lm_v lm n))) (s, b, []) ires = False"
apply(auto simp: inv_after_clear.simps)
done
lemma inv_on_left_moving_Suc_nonempty[simp]: "inv_on_left_moving (as, abc_lm_s lm n (Suc (abc_lm_v lm n)))
(s, b, Oc # list) ires \<Longrightarrow> b \<noteq> []"
apply(auto simp: inv_on_left_moving.simps inv_on_left_moving_norm.simps split: if_splits)
done
lemma inv_check_left_moving_Suc_nonempty[simp]:
"inv_check_left_moving (as, abc_lm_s lm n (Suc (abc_lm_v lm n))) (s, b, Oc # list) ires \<Longrightarrow> b \<noteq> []"
apply(auto simp: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps split: if_splits)
done
lemma tinc_correct_pre:
assumes layout: "ly = layout_of ap"
and inv_start: "inv_locate_a (as, lm) (n, l, r) ires"
and lm': "lm' = abc_lm_s lm n (Suc (abc_lm_v lm n))"
and f: "f = steps (Suc 0, l, r) (tinc_b, 0)"
and P: "P = (\<lambda> (s, l, r). s = 10)"
and Q: "Q = (\<lambda> (s, l, r). inc_inv n (as, lm) (s, l, r) ires)"
shows "\<exists> stp. P (f stp) \<and> Q (f stp)"
proof(rule_tac LE = inc_LE in halt_lemma2)
show "wf inc_LE" by(auto)
next
show "Q (f 0)"
using inv_start
apply(simp add: f P Q steps.simps inc_inv.simps)
done
next
show "\<not> P (f 0)"
apply(simp add: f P steps.simps)
done
next
have "\<not> P (f n) \<and> Q (f n) \<Longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n)
\<in> inc_LE" for n
proof(simp add: f,
cases "steps (Suc 0, l, r) (tinc_b, 0) n", simp add: P)
fix n a b c
assume "a \<noteq> 10 \<and> Q (a, b, c)"
thus "Q (step (a, b, c) (tinc_b, 0)) \<and> (step (a, b, c) (tinc_b, 0), a, b, c) \<in> inc_LE"
apply(simp add:Q)
apply(simp add: inc_inv.simps)
apply(cases c; cases "hd c")
apply(auto simp: Let_def step.simps tinc_b_def split: if_splits) (* ~ 12 sec *)
apply(simp_all add: inc_inv.simps inc_LE_def lex_triple_def lex_pair_def
inc_measure_def numeral)
done
qed
thus "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> inc_LE" by blast
qed
lemma tinc_correct:
assumes layout: "ly = layout_of ap"
and inv_start: "inv_locate_a (as, lm) (n, l, r) ires"
and lm': "lm' = abc_lm_s lm n (Suc (abc_lm_v lm n))"
shows "\<exists> stp l' r'. steps (Suc 0, l, r) (tinc_b, 0) stp = (10, l', r')
\<and> inv_stop (as, lm') (10, l', r') ires"
using assms
apply(drule_tac tinc_correct_pre, auto)
apply(rule_tac x = stp in exI, simp)
apply(simp add: inc_inv.simps)
done
lemma is_even_4[simp]: "(4::nat) * n mod 2 = 0"
apply(arith)
done
lemma crsp_step_inc_pre:
assumes layout: "ly = layout_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
and aexec: "abc_step_l (as, lm) (Some (Inc n)) = (asa, lma)"
shows "\<exists> stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp
= (2*n + 10, Bk # Bk # ires, <lma> @ Bk\<up>k) \<and> stp > 0"
proof -
have "\<exists> stp l' r'. steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r')
\<and> inv_locate_a (as, lm) (n, l', r') ires"
using assms
apply(rule_tac findnth_correct, simp_all add: crsp layout)
done
from this obtain stp l' r' where a:
"steps (Suc 0, l, r) (findnth n, 0) stp = (Suc (2 * n), l', r')
\<and> inv_locate_a (as, lm) (n, l', r') ires" by blast
moreover have
"\<exists> stp la ra. steps (Suc 0, l', r') (tinc_b, 0) stp = (10, la, ra)
\<and> inv_stop (as, lma) (10, la, ra) ires"
using assms a
proof(rule_tac lm' = lma and n = n and lm = lm and ly = ly and ap = ap in tinc_correct,
simp, simp)
show "lma = abc_lm_s lm n (Suc (abc_lm_v lm n))"
using aexec
apply(simp add: abc_step_l.simps)
done
qed
from this obtain stpa la ra where b:
"steps (Suc 0, l', r') (tinc_b, 0) stpa = (10, la, ra)
\<and> inv_stop (as, lma) (10, la, ra) ires" by blast
from a b show "\<exists>stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp
= (2 * n + 10, Bk # Bk # ires, <lma> @ Bk \<up> k) \<and> stp > 0"
apply(rule_tac x = "stp + stpa" in exI)
using tm_append_steps[of "Suc 0" l r "findnth n" stp l' r' tinc_b stpa 10 la ra "length (findnth n) div 2"]
apply(simp add: length_findnth inv_stop.simps)
apply(cases stpa, simp_all add: steps.simps)
done
qed
lemma crsp_step_inc:
assumes layout: "ly = layout_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
and fetch: "abc_fetch as ap = Some (Inc n)"
shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Inc n)))
(steps (s, l, r) (ci ly (start_of ly as) (Inc n), start_of ly as - Suc 0) stp) ires"
proof(cases "(abc_step_l (as, lm) (Some (Inc n)))")
fix a b
assume aexec: "abc_step_l (as, lm) (Some (Inc n)) = (a, b)"
then have "\<exists> stp k. steps (Suc 0, l, r) (findnth n @ shift tinc_b (2 * n), 0) stp
= (2*n + 10, Bk # Bk # ires, <b> @ Bk\<up>k) \<and> stp > 0"
using assms
apply(rule_tac crsp_step_inc_pre, simp_all)
done
thus "?thesis"
using assms aexec
apply(erule_tac exE)
apply(erule_tac exE)
apply(erule_tac conjE)
apply(rename_tac stp k)
apply(rule_tac x = stp in exI, simp add: ci.simps tm_shift_eq_steps)
apply(drule_tac off = "(start_of (layout_of ap) as - Suc 0)" in tm_shift_eq_steps)
apply(auto simp: crsp.simps abc_step_l.simps fetch start_of_Suc1)
done
qed
subsection\<open>Crsp of Dec n e\<close>
type_synonym dec_inv_t = "(nat * nat list) \<Rightarrow> config \<Rightarrow> cell list \<Rightarrow> bool"
fun dec_first_on_right_moving :: "nat \<Rightarrow> dec_inv_t"
where
"dec_first_on_right_moving n (as, lm) (s, l, r) ires =
(\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>
ml + mr = Suc m \<and> length lm1 = n \<and> ml > 0 \<and> m > 0 \<and>
(if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
else l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>
((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or> (r = Oc\<up>mr \<and> lm2 = [])))"
fun dec_on_right_moving :: "dec_inv_t"
where
"dec_on_right_moving (as, lm) (s, l, r) ires =
(\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>
ml + mr = Suc (Suc m) \<and>
(if lm1 = [] then l = Oc\<up>ml@ Bk # Bk # ires
else l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>
((r = Oc\<up>mr @ [Bk] @ <lm2> @ Bk\<up>rn) \<or> (r = Oc\<up>mr \<and> lm2 = [])))"
fun dec_after_clear :: "dec_inv_t"
where
"dec_after_clear (as, lm) (s, l, r) ires =
(\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>
ml + mr = Suc m \<and> ml = Suc m \<and> r \<noteq> [] \<and> r \<noteq> [] \<and>
(if lm1 = [] then l = Oc\<up>ml@ Bk # Bk # ires
else l = Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>
(tl r = Bk # <lm2> @ Bk\<up>rn \<or> tl r = [] \<and> lm2 = []))"
fun dec_after_write :: "dec_inv_t"
where
"dec_after_write (as, lm) (s, l, r) ires =
(\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>
ml + mr = Suc m \<and> ml = Suc m \<and> lm2 \<noteq> [] \<and>
(if lm1 = [] then l = Bk # Oc\<up>ml @ Bk # Bk # ires
else l = Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>
tl r = <lm2> @ Bk\<up>rn)"
fun dec_right_move :: "dec_inv_t"
where
"dec_right_move (as, lm) (s, l, r) ires =
(\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2
\<and> ml = Suc m \<and> mr = (0::nat) \<and>
(if lm1 = [] then l = Bk # Oc\<up>ml @ Bk # Bk # ires
else l = Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires)
\<and> (r = Bk # <lm2> @ Bk\<up>rn \<or> r = [] \<and> lm2 = []))"
fun dec_check_right_move :: "dec_inv_t"
where
"dec_check_right_move (as, lm) (s, l, r) ires =
(\<exists> lm1 lm2 m ml mr rn. lm = lm1 @ [m] @ lm2 \<and>
ml = Suc m \<and> mr = (0::nat) \<and>
(if lm1 = [] then l = Bk # Bk # Oc\<up>ml @ Bk # Bk # ires
else l = Bk # Bk # Oc\<up>ml @ [Bk] @ <rev lm1> @ Bk # Bk # ires) \<and>
r = <lm2> @ Bk\<up>rn)"
fun dec_left_move :: "dec_inv_t"
where
"dec_left_move (as, lm) (s, l, r) ires =
(\<exists> lm1 m rn. (lm::nat list) = lm1 @ [m::nat] \<and>
rn > 0 \<and>
(if lm1 = [] then l = Bk # Oc\<up>Suc m @ Bk # Bk # ires
else l = Bk # Oc\<up>Suc m @ Bk # <rev lm1> @ Bk # Bk # ires) \<and> r = Bk\<up>rn)"
declare
dec_on_right_moving.simps[simp del] dec_after_clear.simps[simp del]
dec_after_write.simps[simp del] dec_left_move.simps[simp del]
dec_check_right_move.simps[simp del] dec_right_move.simps[simp del]
dec_first_on_right_moving.simps[simp del]
fun inv_locate_n_b :: "inc_inv_t"
where
"inv_locate_n_b (as, lm) (s, l, r) ires=
(\<exists> lm1 lm2 tn m ml mr rn. lm @ 0\<up>tn = lm1 @ [m] @ lm2 \<and>
length lm1 = s \<and> m + 1 = ml + mr \<and>
ml = 1 \<and> tn = s + 1 - length lm \<and>
(if lm1 = [] then l = Oc\<up>ml @ Bk # Bk # ires
else l = Oc\<up>ml @ Bk # <rev lm1> @ Bk # Bk # ires) \<and>
(r = Oc\<up>mr @ [Bk] @ <lm2>@ Bk\<up>rn \<or> (lm2 = [] \<and> r = Oc\<up>mr))
)"
fun dec_inv_1 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
where
"dec_inv_1 ly n e (as, am) (s, l, r) ires =
(let ss = start_of ly as in
let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
let am'' = abc_lm_s am n (abc_lm_v am n) in
if s = start_of ly e then inv_stop (as, am'') (s, l, r) ires
else if s = ss then False
else if s = ss + 2 * n + 1 then
inv_locate_b (as, am) (n, l, r) ires
else if s = ss + 2 * n + 13 then
inv_on_left_moving (as, am'') (s, l, r) ires
else if s = ss + 2 * n + 14 then
inv_check_left_moving (as, am'') (s, l, r) ires
else if s = ss + 2 * n + 15 then
inv_after_left_moving (as, am'') (s, l, r) ires
else False)"
declare fetch.simps[simp del]
lemma x_plus_helpers:
"x + 4 = Suc (x + 3)"
"x + 5 = Suc (x + 4)"
"x + 6 = Suc (x + 5)"
"x + 7 = Suc (x + 6)"
"x + 8 = Suc (x + 7)"
"x + 9 = Suc (x + 8)"
"x + 10 = Suc (x + 9)"
"x + 11 = Suc (x + 10)"
"x + 12 = Suc (x + 11)"
"x + 13 = Suc (x + 12)"
"14 + x = Suc (x + 13)"
"15 + x = Suc (x + 14)"
"16 + x = Suc (x + 15)"
by auto
lemma fetch_Dec[simp]:
"fetch (ci ly (start_of ly as) (Dec n e)) (Suc (2 * n)) Bk = (W1, start_of ly as + 2 *n)"
"fetch (ci ly (start_of ly as) (Dec n e)) (Suc (2 * n)) Oc = (R, Suc (start_of ly as) + 2 *n)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Oc
= (R, start_of ly as + 2*n + 2)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (Suc (Suc (2 * n))) Bk
= (L, start_of ly as + 2*n + 13)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (Suc (Suc (Suc (2 * n)))) Oc
= (R, start_of ly as + 2*n + 2)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (Suc (Suc (Suc (2 * n)))) Bk
= (L, start_of ly as + 2*n + 3)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 4) Oc = (W0, start_of ly as + 2*n + 3)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 4) Bk = (R, start_of ly as + 2*n + 4)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 5) Bk = (R, start_of ly as + 2*n + 5)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 6) Bk = (L, start_of ly as + 2*n + 6)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 6) Oc = (L, start_of ly as + 2*n + 7)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 7) Bk = (L, start_of ly as + 2*n + 10)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 8) Bk = (W1, start_of ly as + 2*n + 7)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 8) Oc = (R, start_of ly as + 2*n + 8)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 9) Bk = (L, start_of ly as + 2*n + 9)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 9) Oc = (R, start_of ly as + 2*n + 8)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 10) Bk = (R, start_of ly as + 2*n + 4)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 10) Oc = (W0, start_of ly as + 2*n + 9)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 11) Oc = (L, start_of ly as + 2*n + 10)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 11) Bk = (L, start_of ly as + 2*n + 11)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 12) Oc = (L, start_of ly as + 2*n + 10)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 12) Bk = (R, start_of ly as + 2*n + 12)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (2 * n + 13) Bk = (R, start_of ly as + 2*n + 16)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (14 + 2 * n) Oc = (L, start_of ly as + 2*n + 13)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (14 + 2 * n) Bk = (L, start_of ly as + 2*n + 14)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (15 + 2 * n) Oc = (L, start_of ly as + 2*n + 13)"
"fetch (ci (ly) (start_of ly as) (Dec n e)) (15 + 2 * n) Bk = (R, start_of ly as + 2*n + 15)"
"fetch (ci (ly) (start_of (ly) as) (Dec n e)) (16 + 2 * n) Bk = (R, start_of (ly) e)"
unfolding x_plus_helpers fetch.simps
by(auto simp: ci.simps shift.simps nth_append tdec_b_def length_findnth adjust.simps)
lemma steps_start_of_invb_inv_locate_a1[simp]:
"\<lbrakk>r = [] \<or> hd r = Bk; inv_locate_a (as, lm) (n, l, r) ires\<rbrakk>
\<Longrightarrow> \<exists>stp la ra.
steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e),
start_of ly as - Suc 0) stp = (Suc (start_of ly as + 2 * n), la, ra) \<and>
inv_locate_b (as, lm) (n, la, ra) ires"
apply(rule_tac x = "Suc (Suc 0)" in exI)
apply(auto simp: steps.simps step.simps length_ci_dec)
apply(cases r, simp_all)
done
lemma steps_start_of_invb_inv_locate_a2[simp]:
"\<lbrakk>inv_locate_a (as, lm) (n, l, r) ires; r \<noteq> [] \<and> hd r \<noteq> Bk\<rbrakk>
\<Longrightarrow> \<exists>stp la ra.
steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e),
start_of ly as - Suc 0) stp = (Suc (start_of ly as + 2 * n), la, ra) \<and>
inv_locate_b (as, lm) (n, la, ra) ires"
apply(rule_tac x = "(Suc 0)" in exI, cases "hd r", simp_all)
apply(auto simp: steps.simps step.simps length_ci_dec)
apply(cases r, simp_all)
done
fun abc_dec_1_stage1:: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"abc_dec_1_stage1 (s, l, r) ss n =
(if s > ss \<and> s \<le> ss + 2*n + 1 then 4
else if s = ss + 2 * n + 13 \<or> s = ss + 2*n + 14 then 3
else if s = ss + 2*n + 15 then 2
else 0)"
fun abc_dec_1_stage2:: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"abc_dec_1_stage2 (s, l, r) ss n =
(if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s)
else if s = ss + 2*n + 13 then length l
else if s = ss + 2*n + 14 then length l
else 0)"
fun abc_dec_1_stage3 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"abc_dec_1_stage3 (s, l, r) ss n =
(if s \<le> ss + 2*n + 1 then
if (s - ss) mod 2 = 0 then
if r \<noteq> [] \<and> hd r = Oc then 0 else 1
else length r
else if s = ss + 2 * n + 13 then
if r \<noteq> [] \<and> hd r = Oc then 2
else 1
else if s = ss + 2 * n + 14 then
if r \<noteq> [] \<and> hd r = Oc then 3 else 0
else 0)"
fun abc_dec_1_measure :: "(config \<times> nat \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
where
"abc_dec_1_measure (c, ss, n) = (abc_dec_1_stage1 c ss n,
abc_dec_1_stage2 c ss n, abc_dec_1_stage3 c ss n)"
definition abc_dec_1_LE ::
"((config \<times> nat \<times>
nat) \<times> (config \<times> nat \<times> nat)) set"
where "abc_dec_1_LE \<equiv> (inv_image lex_triple abc_dec_1_measure)"
lemma wf_dec_le: "wf abc_dec_1_LE"
by(auto intro:wf_inv_image simp:abc_dec_1_LE_def lex_triple_def lex_pair_def)
lemma startof_Suc2:
"abc_fetch as ap = Some (Dec n e) \<Longrightarrow>
start_of (layout_of ap) (Suc as) =
start_of (layout_of ap) as + 2 * n + 16"
apply(auto simp: start_of.simps layout_of.simps
length_of.simps abc_fetch.simps
take_Suc_conv_app_nth split: if_splits)
done
lemma start_of_less_2:
"start_of ly e \<le> start_of ly (Suc e)"
apply(cases "e < length ly")
apply(auto simp: start_of.simps take_Suc take_Suc_conv_app_nth)
done
lemma start_of_less_1: "start_of ly e \<le> start_of ly (e + d)"
proof(induct d)
case 0 thus "?case" by simp
next
case (Suc d)
have "start_of ly e \<le> start_of ly (e + d)" by fact
moreover have "start_of ly (e + d) \<le> start_of ly (Suc (e + d))"
by(rule_tac start_of_less_2)
ultimately show"?case"
by(simp)
qed
lemma start_of_less:
assumes "e < as"
shows "start_of ly e \<le> start_of ly as"
proof -
obtain d where " as = e + d"
using assms by (metis less_imp_add_positive)
thus "?thesis"
by(simp add: start_of_less_1)
qed
lemma start_of_ge:
assumes fetch: "abc_fetch as ap = Some (Dec n e)"
and layout: "ly = layout_of ap"
and great: "e > as"
shows "start_of ly e \<ge> start_of ly as + 2*n + 16"
proof(cases "e = Suc as")
case True
have "e = Suc as" by fact
moreover hence "start_of ly (Suc as) = start_of ly as + 2*n + 16"
using layout fetch
by(simp add: startof_Suc2)
ultimately show "?thesis" by (simp)
next
case False
have "e \<noteq> Suc as" by fact
then have "e > Suc as" using great by arith
then have "start_of ly (Suc as) \<le> start_of ly e"
by(simp add: start_of_less)
moreover have "start_of ly (Suc as) = start_of ly as + 2*n + 16"
using layout fetch
by(simp add: startof_Suc2)
ultimately show "?thesis"
by arith
qed
declare dec_inv_1.simps[simp del]
lemma start_of_ineq1[simp]:
"\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk>
\<Longrightarrow> (start_of ly e \<noteq> Suc (start_of ly as + 2 * n) \<and>
start_of ly e \<noteq> Suc (Suc (start_of ly as + 2 * n)) \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 3 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 4 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 5 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 6 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 7 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 8 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 9 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 10 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 11 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 12 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 13 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 14 \<and>
start_of ly e \<noteq> start_of ly as + 2 * n + 15)"
using start_of_ge[of as aprog n e ly] start_of_less[of e as ly]
apply(cases "e < as", simp)
apply(cases "e = as", simp, simp)
done
lemma start_of_ineq2[simp]: "\<lbrakk>abc_fetch as aprog = Some (Dec n e); ly = layout_of aprog\<rbrakk>
\<Longrightarrow> (Suc (start_of ly as + 2 * n) \<noteq> start_of ly e \<and>
Suc (Suc (start_of ly as + 2 * n)) \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 3 \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 4 \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 5 \<noteq>start_of ly e \<and>
start_of ly as + 2 * n + 6 \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 7 \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 8 \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 9 \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 10 \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 11 \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 12 \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 13 \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 14 \<noteq> start_of ly e \<and>
start_of ly as + 2 * n + 15 \<noteq> start_of ly e)"
using start_of_ge[of as aprog n e ly] start_of_less[of e as ly]
apply(cases "e < as", simp, simp)
apply(cases "e = as", simp, simp)
done
lemma inv_locate_b_nonempty[simp]: "inv_locate_b (as, lm) (n, [], []) ires = False"
apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits)
done
lemma inv_locate_b_no_Bk[simp]: "inv_locate_b (as, lm) (n, [], Bk # list) ires = False"
apply(auto simp: inv_locate_b.simps in_middle.simps split: if_splits)
done
lemma dec_first_on_right_moving_Oc[simp]:
"\<lbrakk>dec_first_on_right_moving n (as, am) (s, aaa, Oc # xs) ires\<rbrakk>
\<Longrightarrow> dec_first_on_right_moving n (as, am) (s', Oc # aaa, xs) ires"
apply(simp only: dec_first_on_right_moving.simps)
apply(erule exE)+
apply(rename_tac lm1 lm2 m ml mr rn)
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
rule_tac x = m in exI, rule_tac x = "Suc ml" in exI,
rule_tac x = "mr - 1" in exI)
apply(case_tac [!] mr, auto)
done
lemma dec_first_on_right_moving_Bk_nonempty[simp]:
"dec_first_on_right_moving n (as, am) (s, l, Bk # xs) ires \<Longrightarrow> l \<noteq> []"
apply(auto simp: dec_first_on_right_moving.simps split: if_splits)
done
lemma replicateE:
"\<lbrakk>\<not> length lm1 < length am;
am @ replicate (length lm1 - length am) 0 @ [0::nat] =
lm1 @ m # lm2;
0 < m\<rbrakk>
\<Longrightarrow> RR"
apply(subgoal_tac "lm2 = []", simp)
apply(drule_tac length_equal, simp)
done
lemma dec_after_clear_Bk_strip_hd[simp]:
"\<lbrakk>dec_first_on_right_moving n (as,
abc_lm_s am n (abc_lm_v am n)) (s, l, Bk # xs) ires\<rbrakk>
\<Longrightarrow> dec_after_clear (as, abc_lm_s am n
(abc_lm_v am n - Suc 0)) (s', tl l, hd l # Bk # xs) ires"
apply(simp only: dec_first_on_right_moving.simps
dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps)
apply(erule_tac exE)+
apply(rename_tac lm1 lm2 m ml mr rn)
apply(cases "n < length am")
by(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
rule_tac x = "m - 1" in exI, auto elim:replicateE)
lemma dec_first_on_right_moving_dec_after_clear_cases[simp]:
"\<lbrakk>dec_first_on_right_moving n (as,
abc_lm_s am n (abc_lm_v am n)) (s, l, []) ires\<rbrakk>
\<Longrightarrow> (l = [] \<longrightarrow> dec_after_clear (as,
abc_lm_s am n (abc_lm_v am n - Suc 0)) (s', [], [Bk]) ires) \<and>
(l \<noteq> [] \<longrightarrow> dec_after_clear (as, abc_lm_s am n
(abc_lm_v am n - Suc 0)) (s', tl l, [hd l]) ires)"
apply(subgoal_tac "l \<noteq> []",
simp only: dec_first_on_right_moving.simps
dec_after_clear.simps abc_lm_s.simps abc_lm_v.simps)
apply(erule_tac exE)+
apply(rename_tac lm1 lm2 m ml mr rn)
apply(cases "n < length am", simp)
apply(rule_tac x = lm1 in exI, rule_tac x = "m - 1" in exI, auto)
apply(case_tac [1-2] m, auto)
apply(auto simp: dec_first_on_right_moving.simps split: if_splits)
done
lemma dec_after_clear_Bk_via_Oc[simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Oc # r) ires\<rbrakk>
\<Longrightarrow> dec_after_clear (as, am) (s', l, Bk # r) ires"
apply(auto simp: dec_after_clear.simps)
done
lemma dec_right_move_Bk_via_clear_Bk[simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, Bk # r) ires\<rbrakk>
\<Longrightarrow> dec_right_move (as, am) (s', Bk # l, r) ires"
apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits)
done
lemma dec_right_move_Bk_Bk_via_clear[simp]: "\<lbrakk>dec_after_clear (as, am) (s, l, []) ires\<rbrakk>
\<Longrightarrow> dec_right_move (as, am) (s', Bk # l, [Bk]) ires"
apply(auto simp: dec_after_clear.simps dec_right_move.simps split: if_splits)
done
lemma dec_right_move_no_Oc[simp]:"dec_right_move (as, am) (s, l, Oc # r) ires = False"
apply(auto simp: dec_right_move.simps)
done
lemma dec_right_move_2_check_right_move[simp]:
"\<lbrakk>dec_right_move (as, am) (s, l, Bk # r) ires\<rbrakk>
\<Longrightarrow> dec_check_right_move (as, am) (s', Bk # l, r) ires"
apply(auto simp: dec_right_move.simps dec_check_right_move.simps split: if_splits)
done
lemma lm_iff_empty[simp]: "(<lm::nat list> = []) = (lm = [])"
apply(cases lm, simp_all add: tape_of_nl_cons)
done
lemma dec_right_move_asif_Bk_singleton[simp]:
"dec_right_move (as, am) (s, l, []) ires=
dec_right_move (as, am) (s, l, [Bk]) ires"
apply(simp add: dec_right_move.simps)
done
lemma dec_check_right_move_nonempty[simp]: "dec_check_right_move (as, am) (s, l, r) ires\<Longrightarrow> l \<noteq> []"
apply(auto simp: dec_check_right_move.simps split: if_splits)
done
lemma dec_check_right_move_Oc_tail[simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Oc # r) ires\<rbrakk>
\<Longrightarrow> dec_after_write (as, am) (s', tl l, hd l # Oc # r) ires"
apply(auto simp: dec_check_right_move.simps dec_after_write.simps)
apply(rename_tac lm1 lm2 m rn)
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI, rule_tac x = m in exI, auto)
done
lemma dec_left_move_Bk_tail[simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, Bk # r) ires\<rbrakk>
\<Longrightarrow> dec_left_move (as, am) (s', tl l, hd l # Bk # r) ires"
apply(auto simp: dec_check_right_move.simps dec_left_move.simps inv_after_move.simps)
apply(rename_tac lm1 lm2 m rn)
apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto split: if_splits)
apply(case_tac [!] lm2, simp_all add: tape_of_nl_cons split: if_splits)
apply(rule_tac [!] x = "(Suc rn)" in exI, simp_all)
done
lemma dec_left_move_tail[simp]: "\<lbrakk>dec_check_right_move (as, am) (s, l, []) ires\<rbrakk>
\<Longrightarrow> dec_left_move (as, am) (s', tl l, [hd l]) ires"
apply(auto simp: dec_check_right_move.simps dec_left_move.simps inv_after_move.simps)
apply(rename_tac lm1 m)
apply(rule_tac x = lm1 in exI, rule_tac x = m in exI, auto)
done
lemma dec_left_move_no_Oc[simp]: "dec_left_move (as, am) (s, aaa, Oc # xs) ires = False"
apply(auto simp: dec_left_move.simps inv_after_move.simps)
done
lemma dec_left_move_nonempty[simp]: "dec_left_move (as, am) (s, l, r) ires
\<Longrightarrow> l \<noteq> []"
apply(auto simp: dec_left_move.simps split: if_splits)
done
lemma inv_on_left_moving_in_middle_B_Oc_Bk_Bks[simp]: "inv_on_left_moving_in_middle_B (as, [m])
(s', Oc # Oc\<up>m @ Bk # Bk # ires, Bk # Bk\<up>rn) ires"
apply(simp add: inv_on_left_moving_in_middle_B.simps)
apply(rule_tac x = "[m]" in exI, auto)
done
lemma inv_on_left_moving_in_middle_B_Oc_Bk_Bks_rev[simp]: "lm1 \<noteq> [] \<Longrightarrow>
inv_on_left_moving_in_middle_B (as, lm1 @ [m]) (s',
Oc # Oc\<up>m @ Bk # <rev lm1> @ Bk # Bk # ires, Bk # Bk\<up>rn) ires"
apply(simp only: inv_on_left_moving_in_middle_B.simps)
apply(rule_tac x = "lm1 @ [m ]" in exI, rule_tac x = "[]" in exI, simp)
apply(simp add: tape_of_nl_cons split: if_splits)
done
lemma inv_on_left_moving_Bk_tail[simp]: "dec_left_move (as, am) (s, l, Bk # r) ires
\<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, hd l # Bk # r) ires"
apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits)
done
lemma inv_on_left_moving_tail[simp]: "dec_left_move (as, am) (s, l, []) ires
\<Longrightarrow> inv_on_left_moving (as, am) (s', tl l, [hd l]) ires"
apply(auto simp: dec_left_move.simps inv_on_left_moving.simps split: if_splits)
done
lemma dec_on_right_moving_Oc_mv[simp]: "dec_after_write (as, am) (s, l, Oc # r) ires
\<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires"
apply(auto simp: dec_after_write.simps dec_on_right_moving.simps)
apply(rename_tac lm1 lm2 m rn)
apply(rule_tac x = "lm1 @ [m]" in exI, rule_tac x = "tl lm2" in exI,
rule_tac x = "hd lm2" in exI, simp)
apply(rule_tac x = "Suc 0" in exI,rule_tac x = "Suc (hd lm2)" in exI)
apply(case_tac lm2, auto split: if_splits simp: tape_of_nl_cons)
done
lemma dec_after_write_Oc_via_Bk[simp]: "dec_after_write (as, am) (s, l, Bk # r) ires
\<Longrightarrow> dec_after_write (as, am) (s', l, Oc # r) ires"
apply(auto simp: dec_after_write.simps)
done
lemma dec_after_write_Oc_empty[simp]: "dec_after_write (as, am) (s, aaa, []) ires
\<Longrightarrow> dec_after_write (as, am) (s', aaa, [Oc]) ires"
apply(auto simp: dec_after_write.simps)
done
lemma dec_on_right_moving_Oc_move[simp]: "dec_on_right_moving (as, am) (s, l, Oc # r) ires
\<Longrightarrow> dec_on_right_moving (as, am) (s', Oc # l, r) ires"
apply(simp only: dec_on_right_moving.simps)
apply(erule_tac exE)+
apply(rename_tac lm1 lm2 m ml mr rn)
apply(erule conjE)+
apply(rule_tac x = lm1 in exI, rule_tac x = lm2 in exI,
rule_tac x = "m" in exI, rule_tac x = "Suc ml" in exI,
rule_tac x = "mr - 1" in exI, simp)
apply(case_tac mr, auto)
done
lemma dec_on_right_moving_nonempty[simp]: "dec_on_right_moving (as, am) (s, l, r) ires\<Longrightarrow> l \<noteq> []"
apply(auto simp: dec_on_right_moving.simps split: if_splits)
done
lemma dec_after_clear_Bk_tail[simp]: "dec_on_right_moving (as, am) (s, l, Bk # r) ires
\<Longrightarrow> dec_after_clear (as, am) (s', tl l, hd l # Bk # r) ires"
apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps simp del:split_head_repeat)
apply(rename_tac lm1 lm2 m ml mr rn)
apply(case_tac mr, auto split: if_splits)
done
lemma dec_after_clear_tail[simp]: "dec_on_right_moving (as, am) (s, l, []) ires
\<Longrightarrow> dec_after_clear (as, am) (s', tl l, [hd l]) ires"
apply(auto simp: dec_on_right_moving.simps dec_after_clear.simps)
apply(simp_all split: if_splits)
apply(rule_tac x = lm1 in exI, simp)
done
lemma dec_false_1[simp]:
"\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
\<Longrightarrow> False"
apply(auto simp: inv_locate_b.simps in_middle.simps)
apply(rename_tac lm1 lm2 m ml Mr rn)
apply(case_tac "length lm1 \<ge> length am", auto)
apply(subgoal_tac "lm2 = []", simp, subgoal_tac "m = 0", simp)
apply(case_tac Mr, auto simp: )
apply(subgoal_tac "Suc (length lm1) - length am =
Suc (length lm1 - length am)",
simp add: exp_ind del: replicate.simps, simp)
apply(drule_tac xs = "am @ replicate (Suc (length lm1) - length am) 0"
and ys = "lm1 @ m # lm2" in length_equal, simp)
apply(case_tac Mr, auto simp: abc_lm_v.simps)
apply(rename_tac lm1 m ml Mr)
apply(case_tac "Mr = 0", simp_all split: if_splits)
apply(subgoal_tac "Suc (length lm1) - length am =
Suc (length lm1 - length am)",
simp add: exp_ind del: replicate.simps, simp)
done
lemma inv_on_left_moving_Bk_tl[simp]:
"\<lbrakk>inv_locate_b (as, am) (n, aaa, Bk # xs) ires;
abc_lm_v am n = 0\<rbrakk>
\<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0)
(s, tl aaa, hd aaa # Bk # xs) ires"
apply(simp add: inv_on_left_moving.simps)
apply(simp only: inv_locate_b.simps in_middle.simps)
apply(erule_tac exE)+
apply(rename_tac Lm1 Lm2 tn M ml Mr rn)
apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B
(as, abc_lm_s am n 0) (s, tl aaa, hd aaa # Bk # xs) ires", simp)
apply(simp only: inv_on_left_moving_norm.simps)
apply(erule_tac conjE)+
apply(rule_tac x = Lm1 in exI, rule_tac x = Lm2 in exI,
rule_tac x = M in exI, rule_tac x = M in exI,
rule_tac x = "Suc 0" in exI, simp add: abc_lm_s.simps)
apply(case_tac Mr, auto simp: abc_lm_v.simps)
apply(simp only: exp_ind[THEN sym] replicate_Suc Nat.Suc_diff_le)
apply(auto simp: inv_on_left_moving_in_middle_B.simps split: if_splits)
done
lemma inv_on_left_moving_tl[simp]:
"\<lbrakk>abc_lm_v am n = 0; inv_locate_b (as, am) (n, aaa, []) ires\<rbrakk>
\<Longrightarrow> inv_on_left_moving (as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires"
+ supply [[simproc del: defined_all]]
apply(simp add: inv_on_left_moving.simps)
apply(simp only: inv_locate_b.simps in_middle.simps)
apply(erule_tac exE)+
apply(rename_tac Lm1 Lm2 tn M ml Mr rn)
apply(simp add: inv_on_left_moving.simps)
apply(subgoal_tac "\<not> inv_on_left_moving_in_middle_B
(as, abc_lm_s am n 0) (s, tl aaa, [hd aaa]) ires", simp)
apply(simp only: inv_on_left_moving_norm.simps)
apply(erule_tac conjE)+
apply(rule_tac x = Lm1 in exI, rule_tac x = Lm2 in exI,
rule_tac x = M in exI, rule_tac x = M in exI,
rule_tac x = "Suc 0" in exI, simp add: abc_lm_s.simps)
apply(case_tac Mr, simp_all, auto simp: abc_lm_v.simps)
apply(simp_all only: exp_ind Nat.Suc_diff_le del: replicate_Suc, simp_all)
apply(auto simp: inv_on_left_moving_in_middle_B.simps split: if_splits)
apply(case_tac [!] M, simp_all)
done
declare dec_inv_1.simps[simp del]
declare inv_locate_n_b.simps [simp del]
lemma dec_first_on_right_moving_Oc_via_inv_locate_n_b[simp]:
"\<lbrakk>inv_locate_n_b (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
\<Longrightarrow> dec_first_on_right_moving n (as, abc_lm_s am n (abc_lm_v am n))
(s, Oc # aaa, xs) ires"
apply(auto simp: inv_locate_n_b.simps dec_first_on_right_moving.simps
abc_lm_s.simps abc_lm_v.simps)
apply(rename_tac Lm1 Lm2 m rn)
apply(rule_tac x = Lm1 in exI, rule_tac x = Lm2 in exI,
rule_tac x = m in exI, simp)
apply(rule_tac x = "Suc (Suc 0)" in exI,
rule_tac x = "m - 1" in exI, simp)
apply (metis One_nat_def Suc_pred cell.distinct(1) empty_replicate list.inject list.sel(3)
neq0_conv self_append_conv2 tl_append2 tl_replicate)
apply(rename_tac Lm1 Lm2 m rn)
apply(rule_tac x = Lm1 in exI, rule_tac x = Lm2 in exI,
rule_tac x = m in exI,
simp add: Suc_diff_le exp_ind del: replicate.simps)
apply(rule_tac x = "Suc (Suc 0)" in exI,
rule_tac x = "m - 1" in exI, simp)
apply (metis cell.distinct(1) empty_replicate gr_zeroI list.inject replicateE self_append_conv2)
apply(rename_tac Lm1 m)
apply(rule_tac x = Lm1 in exI, rule_tac x = "[]" in exI,
rule_tac x = m in exI, simp)
apply(rule_tac x = "Suc (Suc 0)" in exI,
rule_tac x = "m - 1" in exI, simp)
apply(case_tac m, auto)
apply(rename_tac Lm1 m)
apply(rule_tac x = Lm1 in exI, rule_tac x = "[]" in exI, rule_tac x = m in exI,
simp add: Suc_diff_le exp_ind del: replicate.simps, simp)
done
lemma inv_on_left_moving_nonempty[simp]: "inv_on_left_moving (as, am) (s, [], r) ires
= False"
apply(simp add: inv_on_left_moving.simps inv_on_left_moving_norm.simps
inv_on_left_moving_in_middle_B.simps)
done
lemma inv_check_left_moving_startof_nonempty[simp]:
"inv_check_left_moving (as, abc_lm_s am n 0)
(start_of (layout_of aprog) as + 2 * n + 14, [], Oc # xs) ires
= False"
apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
done
lemma start_of_lessE[elim]: "\<lbrakk>abc_fetch as ap = Some (Dec n e);
start_of (layout_of ap) as < start_of (layout_of ap) e;
start_of (layout_of ap) e \<le> Suc (start_of (layout_of ap) as + 2 * n)\<rbrakk>
\<Longrightarrow> RR"
using start_of_less[of e as "layout_of ap"] start_of_ge[of as ap n e "layout_of ap"]
apply(cases "as < e", simp)
apply(cases "as = e", simp, simp)
done
lemma crsp_step_dec_b_e_pre':
assumes layout: "ly = layout_of ap"
and inv_start: "inv_locate_b (as, lm) (n, la, ra) ires"
and fetch: "abc_fetch as ap = Some (Dec n e)"
and dec_0: "abc_lm_v lm n = 0"
and f: "f = (\<lambda> stp. (steps (Suc (start_of ly as) + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e),
start_of ly as - Suc 0) stp, start_of ly as, n))"
and P: "P = (\<lambda> ((s, l, r), ss, x). s = start_of ly e)"
and Q: "Q = (\<lambda> ((s, l, r), ss, x). dec_inv_1 ly x e (as, lm) (s, l, r) ires)"
shows "\<exists> stp. P (f stp) \<and> Q (f stp)"
proof(rule_tac LE = abc_dec_1_LE in halt_lemma2)
show "wf abc_dec_1_LE" by(intro wf_dec_le)
next
show "Q (f 0)"
using layout fetch
apply(simp add: f steps.simps Q dec_inv_1.simps)
apply(subgoal_tac "e > as \<or> e = as \<or> e < as")
apply(auto simp: inv_start)
done
next
show "\<not> P (f 0)"
using layout fetch
apply(simp add: f steps.simps P)
done
next
show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> abc_dec_1_LE"
using fetch
proof(rule_tac allI, rule_tac impI)
fix na
assume "\<not> P (f na) \<and> Q (f na)"
thus "Q (f (Suc na)) \<and> (f (Suc na), f na) \<in> abc_dec_1_LE"
apply(simp add: f)
apply(cases "steps (Suc (start_of ly as + 2 * n), la, ra)
(ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) na", simp)
proof -
fix a b c
assume "\<not> P ((a, b, c), start_of ly as, n) \<and> Q ((a, b, c), start_of ly as, n)"
thus "Q (step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n) \<and>
((step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n),
(a, b, c), start_of ly as, n) \<in> abc_dec_1_LE"
apply(simp add: Q)
apply(cases c;cases "hd c")
apply(simp_all add: dec_inv_1.simps Let_def split: if_splits)
using fetch layout dec_0
apply(auto simp: step.simps P dec_inv_1.simps Let_def abc_dec_1_LE_def
lex_triple_def lex_pair_def)
using dec_0
apply(drule_tac dec_false_1, simp_all)
done
qed
qed
qed
lemma crsp_step_dec_b_e_pre:
assumes "ly = layout_of ap"
and inv_start: "inv_locate_b (as, lm) (n, la, ra) ires"
and dec_0: "abc_lm_v lm n = 0"
and fetch: "abc_fetch as ap = Some (Dec n e)"
shows "\<exists>stp lb rb.
steps (Suc (start_of ly as) + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e),
start_of ly as - Suc 0) stp = (start_of ly e, lb, rb) \<and>
dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires"
using assms
apply(drule_tac crsp_step_dec_b_e_pre', auto)
apply(rename_tac stp a b)
apply(rule_tac x = stp in exI, simp)
done
lemma crsp_abc_step_via_stop[simp]:
"\<lbrakk>abc_lm_v lm n = 0;
inv_stop (as, abc_lm_s lm n (abc_lm_v lm n)) (start_of ly e, lb, rb) ires\<rbrakk>
\<Longrightarrow> crsp ly (abc_step_l (as, lm) (Some (Dec n e))) (start_of ly e, lb, rb) ires"
apply(auto simp: crsp.simps abc_step_l.simps inv_stop.simps)
done
lemma crsp_step_dec_b_e:
assumes layout: "ly = layout_of ap"
and inv_start: "inv_locate_a (as, lm) (n, l, r) ires"
and dec_0: "abc_lm_v lm n = 0"
and fetch: "abc_fetch as ap = Some (Dec n e)"
shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
(steps (start_of ly as + 2 * n, l, r) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
proof -
let ?P = "ci ly (start_of ly as) (Dec n e)"
let ?off = "start_of ly as - Suc 0"
have "\<exists> stp la ra. steps (start_of ly as + 2 * n, l, r) (?P, ?off) stp = (Suc (start_of ly as) + 2*n, la, ra)
\<and> inv_locate_b (as, lm) (n, la, ra) ires"
using inv_start
apply(cases "r = [] \<or> hd r = Bk", simp_all)
done
from this obtain stpa la ra where a:
"steps (start_of ly as + 2 * n, l, r) (?P, ?off) stpa = (Suc (start_of ly as) + 2*n, la, ra)
\<and> inv_locate_b (as, lm) (n, la, ra) ires" by blast
have "\<exists> stp lb rb. steps (Suc (start_of ly as) + 2 * n, la, ra) (?P, ?off) stp = (start_of ly e, lb, rb)
\<and> dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires"
using assms a
apply(rule_tac crsp_step_dec_b_e_pre, auto)
done
from this obtain stpb lb rb where b:
"steps (Suc (start_of ly as) + 2 * n, la, ra) (?P, ?off) stpb = (start_of ly e, lb, rb)
\<and> dec_inv_1 ly n e (as, lm) (start_of ly e, lb, rb) ires" by blast
from a b show "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
(steps (start_of ly as + 2 * n, l, r) (?P, ?off) stp) ires"
apply(rule_tac x = "stpa + stpb" in exI)
using dec_0
apply(simp add: dec_inv_1.simps )
apply(cases stpa, simp_all add: steps.simps)
done
qed
fun dec_inv_2 :: "layout \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> dec_inv_t"
where
"dec_inv_2 ly n e (as, am) (s, l, r) ires =
(let ss = start_of ly as in
let am' = abc_lm_s am n (abc_lm_v am n - Suc 0) in
let am'' = abc_lm_s am n (abc_lm_v am n) in
if s = 0 then False
else if s = ss + 2 * n then
inv_locate_a (as, am) (n, l, r) ires
else if s = ss + 2 * n + 1 then
inv_locate_n_b (as, am) (n, l, r) ires
else if s = ss + 2 * n + 2 then
dec_first_on_right_moving n (as, am'') (s, l, r) ires
else if s = ss + 2 * n + 3 then
dec_after_clear (as, am') (s, l, r) ires
else if s = ss + 2 * n + 4 then
dec_right_move (as, am') (s, l, r) ires
else if s = ss + 2 * n + 5 then
dec_check_right_move (as, am') (s, l, r) ires
else if s = ss + 2 * n + 6 then
dec_left_move (as, am') (s, l, r) ires
else if s = ss + 2 * n + 7 then
dec_after_write (as, am') (s, l, r) ires
else if s = ss + 2 * n + 8 then
dec_on_right_moving (as, am') (s, l, r) ires
else if s = ss + 2 * n + 9 then
dec_after_clear (as, am') (s, l, r) ires
else if s = ss + 2 * n + 10 then
inv_on_left_moving (as, am') (s, l, r) ires
else if s = ss + 2 * n + 11 then
inv_check_left_moving (as, am') (s, l, r) ires
else if s = ss + 2 * n + 12 then
inv_after_left_moving (as, am') (s, l, r) ires
else if s = ss + 2 * n + 16 then
inv_stop (as, am') (s, l, r) ires
else False)"
declare dec_inv_2.simps[simp del]
fun abc_dec_2_stage1 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"abc_dec_2_stage1 (s, l, r) ss n =
(if s \<le> ss + 2*n + 1 then 7
else if s = ss + 2*n + 2 then 6
else if s = ss + 2*n + 3 then 5
else if s \<ge> ss + 2*n + 4 \<and> s \<le> ss + 2*n + 9 then 4
else if s = ss + 2*n + 6 then 3
else if s = ss + 2*n + 10 \<or> s = ss + 2*n + 11 then 2
else if s = ss + 2*n + 12 then 1
else 0)"
fun abc_dec_2_stage2 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"abc_dec_2_stage2 (s, l, r) ss n =
(if s \<le> ss + 2 * n + 1 then (ss + 2 * n + 16 - s)
else if s = ss + 2*n + 10 then length l
else if s = ss + 2*n + 11 then length l
else if s = ss + 2*n + 4 then length r - 1
else if s = ss + 2*n + 5 then length r
else if s = ss + 2*n + 7 then length r - 1
else if s = ss + 2*n + 8 then
length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1
else if s = ss + 2*n + 9 then
length r + length (takeWhile (\<lambda> a. a = Oc) l) - 1
else 0)"
fun abc_dec_2_stage3 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"abc_dec_2_stage3 (s, l, r) ss n =
(if s \<le> ss + 2*n + 1 then
if (s - ss) mod 2 = 0 then if r \<noteq> [] \<and>
hd r = Oc then 0 else 1
else length r
else if s = ss + 2 * n + 10 then
if r \<noteq> [] \<and> hd r = Oc then 2
else 1
else if s = ss + 2 * n + 11 then
if r \<noteq> [] \<and> hd r = Oc then 3
else 0
else (ss + 2 * n + 16 - s))"
fun abc_dec_2_stage4 :: "config \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"abc_dec_2_stage4 (s, l, r) ss n =
(if s = ss + 2*n + 2 then length r
else if s = ss + 2*n + 8 then length r
else if s = ss + 2*n + 3 then
if r \<noteq> [] \<and> hd r = Oc then 1
else 0
else if s = ss + 2*n + 7 then
if r \<noteq> [] \<and> hd r = Oc then 0
else 1
else if s = ss + 2*n + 9 then
if r \<noteq> [] \<and> hd r = Oc then 1
else 0
else 0)"
fun abc_dec_2_measure :: "(config \<times> nat \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat \<times> nat)"
where
"abc_dec_2_measure (c, ss, n) =
(abc_dec_2_stage1 c ss n,
abc_dec_2_stage2 c ss n, abc_dec_2_stage3 c ss n, abc_dec_2_stage4 c ss n)"
definition lex_square::
"((nat \<times> nat \<times> nat \<times> nat) \<times> (nat \<times> nat \<times> nat \<times> nat)) set"
where "lex_square \<equiv> less_than <*lex*> lex_triple"
definition abc_dec_2_LE ::
"((config \<times> nat \<times>
nat) \<times> (config \<times> nat \<times> nat)) set"
where "abc_dec_2_LE \<equiv> (inv_image lex_square abc_dec_2_measure)"
lemma wf_dec2_le: "wf abc_dec_2_LE"
by(auto simp:abc_dec_2_LE_def lex_square_def lex_triple_def lex_pair_def)
lemma fix_add: "fetch ap ((x::nat) + 2*n) b = fetch ap (2*n + x) b"
using Suc_1 add.commute by metis
lemma inv_locate_n_b_Bk_elim[elim]:
"\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am) (n, aaa, Bk # xs) ires\<rbrakk>
\<Longrightarrow> RR"
by(auto simp:gr0_conv_Suc inv_locate_n_b.simps abc_lm_v.simps split: if_splits)
lemma inv_locate_n_b_nonemptyE[elim]:
"\<lbrakk>0 < abc_lm_v am n; inv_locate_n_b (as, am)
(n, aaa, []) ires\<rbrakk> \<Longrightarrow> RR"
apply(auto simp: inv_locate_n_b.simps abc_lm_v.simps split: if_splits)
done
lemma no_Ocs_dec_after_write[simp]: "dec_after_write (as, am) (s, aa, r) ires
\<Longrightarrow> takeWhile (\<lambda>a. a = Oc) aa = []"
apply(simp only : dec_after_write.simps)
apply(erule exE)+
apply(erule_tac conjE)+
apply(cases aa, simp)
apply(cases "hd aa", simp only: takeWhile.simps , simp_all split: if_splits)
done
lemma fewer_Ocs_dec_on_right_moving[simp]:
"\<lbrakk>dec_on_right_moving (as, lm) (s, aa, []) ires;
length (takeWhile (\<lambda>a. a = Oc) (tl aa))
\<noteq> length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0\<rbrakk>
\<Longrightarrow> length (takeWhile (\<lambda>a. a = Oc) (tl aa)) <
length (takeWhile (\<lambda>a. a = Oc) aa) - Suc 0"
apply(simp only: dec_on_right_moving.simps)
apply(erule_tac exE)+
apply(erule_tac conjE)+
apply(rename_tac lm1 lm2 m ml Mr rn)
apply(case_tac Mr, auto split: if_splits)
done
lemma more_Ocs_dec_after_clear[simp]:
"dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0))
(start_of (layout_of aprog) as + 2 * n + 9, aa, Bk # xs) ires
\<Longrightarrow> length xs - Suc 0 < length xs +
length (takeWhile (\<lambda>a. a = Oc) aa)"
apply(simp only: dec_after_clear.simps)
apply(erule_tac exE)+
apply(erule conjE)+
apply(simp split: if_splits )
done
lemma more_Ocs_dec_after_clear2[simp]:
"\<lbrakk>dec_after_clear (as, abc_lm_s am n (abc_lm_v am n - Suc 0))
(start_of (layout_of aprog) as + 2 * n + 9, aa, []) ires\<rbrakk>
\<Longrightarrow> Suc 0 < length (takeWhile (\<lambda>a. a = Oc) aa)"
apply(simp add: dec_after_clear.simps split: if_splits)
done
lemma inv_check_left_moving_nonemptyE[elim]:
"inv_check_left_moving (as, lm) (s, [], Oc # xs) ires
\<Longrightarrow> RR"
apply(simp add: inv_check_left_moving.simps inv_check_left_moving_in_middle.simps)
done
lemma inv_locate_n_b_Oc_via_at_begin_norm[simp]:
"\<lbrakk>0 < abc_lm_v am n;
at_begin_norm (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
\<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
apply(simp only: at_begin_norm.simps inv_locate_n_b.simps)
apply(erule_tac exE)+
apply(rename_tac lm1 lm2 rn)
apply(rule_tac x = lm1 in exI, simp)
apply(case_tac "length lm2", simp)
apply(case_tac "lm2", simp, simp)
apply(case_tac "lm2", auto simp: tape_of_nl_cons split: if_splits)
done
lemma inv_locate_n_b_Oc_via_at_begin_fst_awtn[simp]:
"\<lbrakk>0 < abc_lm_v am n;
at_begin_fst_awtn (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
\<Longrightarrow> inv_locate_n_b (as, am) (n, Oc # aaa, xs) ires"
apply(simp only: at_begin_fst_awtn.simps inv_locate_n_b.simps )
apply(erule exE)+
apply(rename_tac lm1 tn rn)
apply(erule conjE)+
apply(rule_tac x = lm1 in exI, rule_tac x = "[]" in exI,
rule_tac x = "Suc tn" in exI, rule_tac x = 0 in exI)
apply(simp add: exp_ind del: replicate.simps)
apply(rule conjI)+
apply(auto)
done
lemma inv_locate_n_b_Oc_via_inv_locate_n_a[simp]:
"\<lbrakk>0 < abc_lm_v am n; inv_locate_a (as, am) (n, aaa, Oc # xs) ires\<rbrakk>
\<Longrightarrow> inv_locate_n_b (as, am) (n, Oc#aaa, xs) ires"
apply(auto simp: inv_locate_a.simps at_begin_fst_bwtn.simps)
done
lemma more_Oc_dec_on_right_moving[simp]:
"\<lbrakk>dec_on_right_moving (as, am) (s, aa, Bk # xs) ires;
Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa)))
\<noteq> length (takeWhile (\<lambda>a. a = Oc) aa)\<rbrakk>
\<Longrightarrow> Suc (length (takeWhile (\<lambda>a. a = Oc) (tl aa)))
< length (takeWhile (\<lambda>a. a = Oc) aa)"
apply(simp only: dec_on_right_moving.simps)
apply(erule exE)+
apply(rename_tac ml mr rn)
apply(case_tac ml, auto split: if_splits )
done
lemma crsp_step_dec_b_suc_pre:
assumes layout: "ly = layout_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires"
and fetch: "abc_fetch as ap = Some (Dec n e)"
and dec_suc: "0 < abc_lm_v lm n"
and f: "f = (\<lambda> stp. (steps (start_of ly as + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e),
start_of ly as - Suc 0) stp, start_of ly as, n))"
and P: "P = (\<lambda> ((s, l, r), ss, x). s = start_of ly as + 2*n + 16)"
and Q: "Q = (\<lambda> ((s, l, r), ss, x). dec_inv_2 ly x e (as, lm) (s, l, r) ires)"
shows "\<exists> stp. P (f stp) \<and> Q(f stp)"
proof(rule_tac LE = abc_dec_2_LE in halt_lemma2)
show "wf abc_dec_2_LE" by(intro wf_dec2_le)
next
show "Q (f 0)"
using layout fetch inv_start
apply(simp add: f steps.simps Q)
apply(simp only: dec_inv_2.simps)
apply(auto simp: Let_def start_of_ge start_of_less inv_start dec_inv_2.simps)
done
next
show "\<not> P (f 0)"
using layout fetch
apply(simp add: f steps.simps P)
done
next
show "\<forall>n. \<not> P (f n) \<and> Q (f n) \<longrightarrow> Q (f (Suc n)) \<and> (f (Suc n), f n) \<in> abc_dec_2_LE"
using fetch
proof(rule_tac allI, rule_tac impI)
fix na
assume "\<not> P (f na) \<and> Q (f na)"
thus "Q (f (Suc na)) \<and> (f (Suc na), f na) \<in> abc_dec_2_LE"
apply(simp add: f)
apply(cases "steps ((start_of ly as + 2 * n), la, ra)
(ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) na", simp)
proof -
fix a b c
assume "\<not> P ((a, b, c), start_of ly as, n) \<and> Q ((a, b, c), start_of ly as, n)"
thus "Q (step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n) \<and>
((step (a, b, c) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0), start_of ly as, n),
(a, b, c), start_of ly as, n) \<in> abc_dec_2_LE"
apply(simp add: Q)
apply(erule_tac conjE)
apply(cases c; cases "hd c")
apply(simp_all add: dec_inv_2.simps Let_def)
apply(simp_all split: if_splits)
using fetch layout dec_suc
apply(auto simp: step.simps P dec_inv_2.simps Let_def abc_dec_2_LE_def lex_triple_def lex_pair_def lex_square_def
fix_add numeral_3_eq_3)
done
qed
qed
qed
lemma crsp_abc_step_l_start_of[simp]:
"\<lbrakk>inv_stop (as, abc_lm_s lm n (abc_lm_v lm n - Suc 0))
(start_of (layout_of ap) as + 2 * n + 16, a, b) ires;
abc_lm_v lm n > 0;
abc_fetch as ap = Some (Dec n e)\<rbrakk>
\<Longrightarrow> crsp (layout_of ap) (abc_step_l (as, lm) (Some (Dec n e)))
(start_of (layout_of ap) as + 2 * n + 16, a, b) ires"
by(auto simp: inv_stop.simps crsp.simps abc_step_l.simps startof_Suc2)
lemma crsp_step_dec_b_suc:
assumes layout: "ly = layout_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires"
and fetch: "abc_fetch as ap = Some (Dec n e)"
and dec_suc: "0 < abc_lm_v lm n"
shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
(steps (start_of ly as + 2 * n, la, ra) (ci (layout_of ap)
(start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
using assms
apply(drule_tac crsp_step_dec_b_suc_pre, auto)
apply(rename_tac stp a b)
apply(rule_tac x = stp in exI)
apply(case_tac stp, simp_all add: steps.simps dec_inv_2.simps)
done
lemma crsp_step_dec_b:
assumes layout: "ly = layout_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
and inv_start: "inv_locate_a (as, lm) (n, la, ra) ires"
and fetch: "abc_fetch as ap = Some (Dec n e)"
shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
(steps (start_of ly as + 2 * n, la, ra) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
using assms
apply(cases "abc_lm_v lm n = 0")
apply(rule_tac crsp_step_dec_b_e, simp_all)
apply(rule_tac crsp_step_dec_b_suc, simp_all)
done
lemma crsp_step_dec:
assumes layout: "ly = layout_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
and fetch: "abc_fetch as ap = Some (Dec n e)"
shows "\<exists>stp > 0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
(steps (s, l, r) (ci ly (start_of ly as) (Dec n e), start_of ly as - Suc 0) stp) ires"
proof(simp add: ci.simps)
let ?off = "start_of ly as - Suc 0"
let ?A = "findnth n"
let ?B = "adjust (shift (shift tdec_b (2 * n)) ?off) (start_of ly e)"
have "\<exists> stp la ra. steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp = (start_of ly as + 2*n, la, ra)
\<and> inv_locate_a (as, lm) (n, la, ra) ires"
proof -
have "\<exists>stp l' r'. steps (Suc 0, l, r) (?A, 0) stp = (Suc (2 * n), l', r') \<and>
inv_locate_a (as, lm) (n, l', r') ires"
using assms
apply(rule_tac findnth_correct, simp_all)
done
then obtain stp l' r' where a:
"steps (Suc 0, l, r) (?A, 0) stp = (Suc (2 * n), l', r') \<and>
inv_locate_a (as, lm) (n, l', r') ires" by blast
then have "steps (Suc 0 + ?off, l, r) (shift ?A ?off, ?off) stp = (Suc (2 * n) + ?off, l', r')"
apply(rule_tac tm_shift_eq_steps, simp_all)
done
moreover have "s = start_of ly as"
using crsp
apply(auto simp: crsp.simps)
done
ultimately show "\<exists> stp la ra. steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp = (start_of ly as + 2*n, la, ra)
\<and> inv_locate_a (as, lm) (n, la, ra) ires"
using a
apply(drule_tac B = ?B in tm_append_first_steps_eq, auto)
apply(rule_tac x = stp in exI, simp)
done
qed
from this obtain stpa la ra where a:
"steps (s, l, r) (shift ?A ?off @ ?B, ?off) stpa = (start_of ly as + 2*n, la, ra)
\<and> inv_locate_a (as, lm) (n, la, ra) ires" by blast
have "\<exists>stp. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
(steps (start_of ly as + 2*n, la, ra) (shift ?A ?off @ ?B, ?off) stp) ires \<and> stp > 0"
using assms a
apply(drule_tac crsp_step_dec_b, auto)
apply(rename_tac stp)
apply(rule_tac x = stp in exI, simp add: ci.simps)
done
then obtain stpb where b:
"crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
(steps (start_of ly as + 2*n, la, ra) (shift ?A ?off @ ?B, ?off) stpb) ires \<and> stpb > 0" ..
from a b show "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some (Dec n e)))
(steps (s, l, r) (shift ?A ?off @ ?B, ?off) stp) ires"
apply(rule_tac x = "stpa + stpb" in exI)
apply(simp)
done
qed
subsection\<open>Crsp of Goto\<close>
lemma crsp_step_goto:
assumes layout: "ly = layout_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
shows "\<exists>stp>0. crsp ly (abc_step_l (as, lm) (Some (Goto n)))
(steps (s, l, r) (ci ly (start_of ly as) (Goto n),
start_of ly as - Suc 0) stp) ires"
using crsp
apply(rule_tac x = "Suc 0" in exI)
apply(cases r;cases "hd r")
apply(simp_all add: ci.simps steps.simps step.simps crsp.simps fetch.simps abc_step_l.simps)
done
lemma crsp_step_in:
assumes layout: "ly = layout_of ap"
and compile: "tp = tm_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
and fetch: "abc_fetch as ap = Some ins"
shows "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some ins))
(steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) ires"
using assms
apply(cases ins, simp_all)
apply(rule crsp_step_inc, simp_all)
apply(rule crsp_step_dec, simp_all)
apply(rule_tac crsp_step_goto, simp_all)
done
lemma crsp_step:
assumes layout: "ly = layout_of ap"
and compile: "tp = tm_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
and fetch: "abc_fetch as ap = Some ins"
shows "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some ins))
(steps (s, l, r) (tp, 0) stp) ires"
proof -
have "\<exists> stp>0. crsp ly (abc_step_l (as, lm) (Some ins))
(steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) ires"
using assms
apply(rule_tac crsp_step_in, simp_all)
done
from this obtain stp where d: "stp > 0 \<and> crsp ly (abc_step_l (as, lm) (Some ins))
(steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) ires" ..
obtain s' l' r' where e:
"(steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp) = (s', l', r')"
apply(cases "(steps (s, l, r) (ci ly (start_of ly as) ins, start_of ly as - 1) stp)")
by blast
then have "steps (s, l, r) (tp, 0) stp = (s', l', r')"
using assms d
apply(rule_tac steps_eq_in)
apply(simp_all)
apply(cases "(abc_step_l (as, lm) (Some ins))", simp add: crsp.simps)
done
thus " \<exists>stp>0. crsp ly (abc_step_l (as, lm) (Some ins)) (steps (s, l, r) (tp, 0) stp) ires"
using d e
apply(rule_tac x = stp in exI, simp)
done
qed
lemma crsp_steps:
assumes layout: "ly = layout_of ap"
and compile: "tp = tm_of ap"
and crsp: "crsp ly (as, lm) (s, l, r) ires"
shows "\<exists> stp. crsp ly (abc_steps_l (as, lm) ap n)
(steps (s, l, r) (tp, 0) stp) ires"
using crsp
proof(induct n)
case 0
then show ?case apply(rule_tac x = 0 in exI)
by(simp add: steps.simps abc_steps_l.simps)
next
case (Suc n)
then obtain stp where "crsp ly (abc_steps_l (as, lm) ap n) (steps0 (s, l, r) tp stp) ires"
by blast
thus ?case
apply(cases "(abc_steps_l (as, lm) ap n)", auto)
apply(frule_tac abc_step_red, simp)
apply(cases "abc_fetch (fst (abc_steps_l (as, lm) ap n)) ap", simp add: abc_step_l.simps, auto)
apply(cases "steps (s, l, r) (tp, 0) stp", simp)
using assms
apply(drule_tac s = "fst (steps0 (s, l, r) (tm_of ap) stp)"
and l = "fst (snd (steps0 (s, l, r) (tm_of ap) stp))"
and r = "snd (snd (steps0 (s, l, r) (tm_of ap) stp))" in crsp_step, auto)
by (metis steps_add)
qed
lemma tp_correct':
assumes layout: "ly = layout_of ap"
and compile: "tp = tm_of ap"
and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
and abc_halt: "abc_steps_l (0, lm) ap stp = (length ap, am)"
shows "\<exists> stp k. steps (Suc 0, l, r) (tp, 0) stp = (start_of ly (length ap), Bk # Bk # ires, <am> @ Bk\<up>k)"
using assms
apply(drule_tac n = stp in crsp_steps, auto)
apply(rename_tac stpA)
apply(rule_tac x = stpA in exI)
apply(case_tac "steps (Suc 0, l, r) (tm_of ap, 0) stpA", simp add: crsp.simps)
done
text\<open>The tp @ [(Nop, 0), (Nop, 0)] is nomoral turing machines, so we can use Hoare\_plus when composing with Mop machine\<close>
lemma layout_id_cons: "layout_of (ap @ [p]) = layout_of ap @ [length_of p]"
apply(simp add: layout_of.simps)
done
lemma map_start_of_layout[simp]:
"map (start_of (layout_of xs @ [length_of x])) [0..<length xs] = (map (start_of (layout_of xs)) [0..<length xs])"
apply(auto)
apply(simp add: layout_of.simps start_of.simps)
done
lemma tpairs_id_cons:
"tpairs_of (xs @ [x]) = tpairs_of xs @ [(start_of (layout_of (xs @ [x])) (length xs), x)]"
apply(auto simp: tpairs_of.simps layout_id_cons )
done
lemma map_length_ci:
"(map (length \<circ> (\<lambda>(xa, y). ci (layout_of xs @ [length_of x]) xa y)) (tpairs_of xs)) =
(map (length \<circ> (\<lambda>(x, y). ci (layout_of xs) x y)) (tpairs_of xs)) "
apply(auto simp: ci.simps adjust.simps) apply(rename_tac A B)
apply(case_tac B, auto simp: ci.simps adjust.simps)
done
lemma length_tp'[simp]:
"\<lbrakk>ly = layout_of ap; tp = tm_of ap\<rbrakk> \<Longrightarrow>
length tp = 2 * sum_list (take (length ap) (layout_of ap))"
proof(induct ap arbitrary: ly tp rule: rev_induct)
case Nil
thus "?case"
by(simp add: tms_of.simps tm_of.simps tpairs_of.simps)
next
fix x xs ly tp
assume ind: "\<And>ly tp. \<lbrakk>ly = layout_of xs; tp = tm_of xs\<rbrakk> \<Longrightarrow>
length tp = 2 * sum_list (take (length xs) (layout_of xs))"
and layout: "ly = layout_of (xs @ [x])"
and tp: "tp = tm_of (xs @ [x])"
obtain ly' where a: "ly' = layout_of xs"
by metis
obtain tp' where b: "tp' = tm_of xs"
by metis
have c: "length tp' = 2 * sum_list (take (length xs) (layout_of xs))"
using a b
by(erule_tac ind, simp)
thus "length tp = 2 *
sum_list (take (length (xs @ [x])) (layout_of (xs @ [x])))"
using tp b
apply(auto simp: layout_id_cons tm_of.simps tms_of.simps length_concat tpairs_id_cons map_length_ci)
apply(cases x)
apply(auto simp: ci.simps tinc_b_def tdec_b_def length_findnth adjust.simps length_of.simps
split: abc_inst.splits)
done
qed
lemma length_tp:
"\<lbrakk>ly = layout_of ap; tp = tm_of ap\<rbrakk> \<Longrightarrow>
start_of ly (length ap) = Suc (length tp div 2)"
apply(frule_tac length_tp', simp_all)
apply(simp add: start_of.simps)
done
lemma compile_correct_halt:
assumes layout: "ly = layout_of ap"
and compile: "tp = tm_of ap"
and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
and abc_halt: "abc_steps_l (0, lm) ap stp = (length ap, am)"
and rs_loc: "n < length am"
and rs: "abc_lm_v am n = rs"
and off: "off = length tp div 2"
shows "\<exists> stp i j. steps (Suc 0, l, r) (tp @ shift (mopup n) off, 0) stp = (0, Bk\<up>i @ Bk # Bk # ires, Oc\<up>Suc rs @ Bk\<up>j)"
proof -
have "\<exists> stp k. steps (Suc 0, l, r) (tp, 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)"
using assms tp_correct'[of ly ap tp lm l r ires stp am]
by(simp add: length_tp)
then obtain stp k where "steps (Suc 0, l, r) (tp, 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)"
by blast
then have a: "steps (Suc 0, l, r) (tp@shift (mopup n) off , 0) stp = (Suc off, Bk # Bk # ires, <am> @ Bk\<up>k)"
using assms
by(auto intro: tm_append_first_steps_eq)
have "\<exists> stp i j. (steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stp)
= (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)"
using assms
by(rule_tac mopup_correct, auto simp: abc_lm_v.simps)
then obtain stpb i j where
"steps (Suc 0, Bk # Bk # ires, <am> @ Bk \<up> k) (mopup_a n @ shift mopup_b (2 * n), 0) stpb
= (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)" by blast
then have b: "steps (Suc 0 + off, Bk # Bk # ires, <am> @ Bk \<up> k) (tp @ shift (mopup n) off, 0) stpb
= (0, Bk\<up>i @ Bk # Bk # ires, Oc # Oc\<up> rs @ Bk\<up>j)"
using assms wf_mopup
apply(drule_tac tm_append_second_halt_eq, auto)
done
from a b show "?thesis"
by(rule_tac x = "stp + stpb" in exI, simp add: steps_add)
qed
declare mopup.simps[simp del]
lemma abc_step_red2:
"abc_steps_l (s, lm) p (Suc n) = (let (as', am') = abc_steps_l (s, lm) p n in
abc_step_l (as', am') (abc_fetch as' p))"
apply(cases "abc_steps_l (s, lm) p n", simp)
apply(drule_tac abc_step_red, simp)
done
lemma crsp_steps2:
assumes
layout: "ly = layout_of ap"
and compile: "tp = tm_of ap"
and crsp: "crsp ly (0, lm) (Suc 0, l, r) ires"
and nothalt: "as < length ap"
and aexec: "abc_steps_l (0, lm) ap stp = (as, am)"
shows "\<exists>stpa\<ge>stp. crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires"
using nothalt aexec
proof(induct stp arbitrary: as am)
case 0
thus "?case"
using crsp
by(rule_tac x = 0 in exI, auto simp: abc_steps_l.simps steps.simps crsp)
next
case (Suc stp as am)
have ind:
"\<And> as am. \<lbrakk>as < length ap; abc_steps_l (0, lm) ap stp = (as, am)\<rbrakk>
\<Longrightarrow> \<exists>stpa\<ge>stp. crsp ly (as, am) (steps (Suc 0, l, r) (tp, 0) stpa) ires" by fact
have a: "as < length ap" by fact
have b: "abc_steps_l (0, lm) ap (Suc stp) = (as, am)" by fact
obtain as' am' where c: "abc_steps_l (0, lm) ap stp = (as', am')"
by(cases "abc_steps_l (0, lm) ap stp", auto)
then have d: "as' < length ap"
using a b
by(simp add: abc_step_red2, cases "as' < length ap", simp,
simp add: abc_fetch.simps abc_steps_l.simps abc_step_l.simps)
have "\<exists>stpa\<ge>stp. crsp ly (as', am') (steps (Suc 0, l, r) (tp, 0) stpa) ires"
using d c ind by simp
from this obtain stpa where e:
"stpa \<ge> stp \<and> crsp ly (as', am') (steps (Suc 0, l, r) (tp, 0) stpa) ires"
by blast
obtain s' l' r' where f: "steps (Suc 0, l, r) (tp, 0) stpa = (s', l', r')"
by(cases "steps (Suc 0, l, r) (tp, 0) stpa", auto)
obtain ins where g: "abc_fetch as' ap = Some ins" using d
by(cases "abc_fetch as' ap",auto simp: abc_fetch.simps)
then have "\<exists>stp> (0::nat). crsp ly (abc_step_l (as', am') (Some ins))
(steps (s', l', r') (tp, 0) stp) ires "
using layout compile e f
by(rule_tac crsp_step, simp_all)
then obtain stpb where "stpb > 0 \<and> crsp ly (abc_step_l (as', am') (Some ins))
(steps (s', l', r') (tp, 0) stpb) ires" ..
from this show "?case" using b e g f c
by(rule_tac x = "stpa + stpb" in exI, simp add: steps_add abc_step_red2)
qed
lemma compile_correct_unhalt:
assumes layout: "ly = layout_of ap"
and compile: "tp = tm_of ap"
and crsp: "crsp ly (0, lm) (1, l, r) ires"
and off: "off = length tp div 2"
and abc_unhalt: "\<forall> stp. (\<lambda> (as, am). as < length ap) (abc_steps_l (0, lm) ap stp)"
shows "\<forall> stp.\<not> is_final (steps (1, l, r) (tp @ shift (mopup n) off, 0) stp)"
using assms
proof(rule_tac allI, rule_tac notI)
fix stp
assume h: "is_final (steps (1, l, r) (tp @ shift (mopup n) off, 0) stp)"
obtain as am where a: "abc_steps_l (0, lm) ap stp = (as, am)"
by(cases "abc_steps_l (0, lm) ap stp", auto)
then have b: "as < length ap"
using abc_unhalt
by(erule_tac x = stp in allE, simp)
have "\<exists> stpa\<ge>stp. crsp ly (as, am) (steps (1, l, r) (tp, 0) stpa) ires "
using assms b a
apply(simp add: numeral)
apply(rule_tac crsp_steps2)
apply(simp_all)
done
then obtain stpa where
"stpa\<ge>stp \<and> crsp ly (as, am) (steps (1, l, r) (tp, 0) stpa) ires" ..
then obtain s' l' r' where b: "(steps (1, l, r) (tp, 0) stpa) = (s', l', r') \<and>
stpa\<ge>stp \<and> crsp ly (as, am) (s', l', r') ires"
by(cases "steps (1, l, r) (tp, 0) stpa", auto)
hence c:
"(steps (1, l, r) (tp @ shift (mopup n) off, 0) stpa) = (s', l', r')"
by(rule_tac tm_append_first_steps_eq, simp_all add: crsp.simps)
from b have d: "s' > 0 \<and> stpa \<ge> stp"
by(simp add: crsp.simps)
then obtain diff where e: "stpa = stp + diff" by (metis le_iff_add)
obtain s'' l'' r'' where f:
"steps (1, l, r) (tp @ shift (mopup n) off, 0) stp = (s'', l'', r'') \<and> is_final (s'', l'', r'')"
using h
by(cases "steps (1, l, r) (tp @ shift (mopup n) off, 0) stp", auto)
then have "is_final (steps (s'', l'', r'') (tp @ shift (mopup n) off, 0) diff)"
by(auto intro: after_is_final)
then have "is_final (steps (1, l, r) (tp @ shift (mopup n) off, 0) stpa)"
using e f by simp
from this and c d show "False" by simp
qed
end
diff --git a/thys/Universal_Turing_Machine/Recursive.thy b/thys/Universal_Turing_Machine/Recursive.thy
--- a/thys/Universal_Turing_Machine/Recursive.thy
+++ b/thys/Universal_Turing_Machine/Recursive.thy
@@ -1,2830 +1,2830 @@
(* Title: thys/Recursive.thy
Author: Jian Xu, Xingyuan Zhang, and Christian Urban
Modifications: Sebastiaan Joosten
*)
theory Recursive
imports Abacus Rec_Def Abacus_Hoare
begin
fun addition :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_prog"
where
"addition m n p = [Dec m 4, Inc n, Inc p, Goto 0, Dec p 7, Inc m, Goto 4]"
fun mv_box :: "nat \<Rightarrow> nat \<Rightarrow> abc_prog"
where
"mv_box m n = [Dec m 3, Inc n, Goto 0]"
text \<open>The compilation of \<open>z\<close>-operator.\<close>
definition rec_ci_z :: "abc_inst list"
where
"rec_ci_z \<equiv> [Goto 1]"
text \<open>The compilation of \<open>s\<close>-operator.\<close>
definition rec_ci_s :: "abc_inst list"
where
"rec_ci_s \<equiv> (addition 0 1 2 [+] [Inc 1])"
text \<open>The compilation of \<open>id i j\<close>-operator\<close>
fun rec_ci_id :: "nat \<Rightarrow> nat \<Rightarrow> abc_inst list"
where
"rec_ci_id i j = addition j i (i + 1)"
fun mv_boxes :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> abc_inst list"
where
"mv_boxes ab bb 0 = []" |
"mv_boxes ab bb (Suc n) = mv_boxes ab bb n [+] mv_box (ab + n) (bb + n)"
fun empty_boxes :: "nat \<Rightarrow> abc_inst list"
where
"empty_boxes 0 = []" |
"empty_boxes (Suc n) = empty_boxes n [+] [Dec n 2, Goto 0]"
fun cn_merge_gs ::
"(abc_inst list \<times> nat \<times> nat) list \<Rightarrow> nat \<Rightarrow> abc_inst list"
where
"cn_merge_gs [] p = []" |
"cn_merge_gs (g # gs) p =
(let (gprog, gpara, gn) = g in
gprog [+] mv_box gpara p [+] cn_merge_gs gs (Suc p))"
text \<open>
The compiler of recursive functions, where \<open>rec_ci recf\<close> return
\<open>(ap, arity, fp)\<close>, where \<open>ap\<close> is the Abacus program, \<open>arity\<close> is the
arity of the recursive function \<open>recf\<close>,
\<open>fp\<close> is the amount of memory which is going to be
used by \<open>ap\<close> for its execution.
\<close>
fun rec_ci :: "recf \<Rightarrow> abc_inst list \<times> nat \<times> nat"
where
"rec_ci z = (rec_ci_z, 1, 2)" |
"rec_ci s = (rec_ci_s, 1, 3)" |
"rec_ci (id m n) = (rec_ci_id m n, m, m + 2)" |
"rec_ci (Cn n f gs) =
(let cied_gs = map (\<lambda> g. rec_ci g) gs in
let (fprog, fpara, fn) = rec_ci f in
let pstr = Max (set (Suc n # fn # (map (\<lambda> (aprog, p, n). n) cied_gs))) in
let qstr = pstr + Suc (length gs) in
(cn_merge_gs cied_gs pstr [+] mv_boxes 0 qstr n [+]
mv_boxes pstr 0 (length gs) [+] fprog [+]
mv_box fpara pstr [+] empty_boxes (length gs) [+]
mv_box pstr n [+] mv_boxes qstr 0 n, n, qstr + n))" |
"rec_ci (Pr n f g) =
(let (fprog, fpara, fn) = rec_ci f in
let (gprog, gpara, gn) = rec_ci g in
let p = Max (set ([n + 3, fn, gn])) in
let e = length gprog + 7 in
(mv_box n p [+] fprog [+] mv_box n (Suc n) [+]
(([Dec p e] [+] gprog [+]
[Inc n, Dec (Suc n) 3, Goto 1]) @
[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gprog + 4)]),
Suc n, p + 1))" |
"rec_ci (Mn n f) =
(let (fprog, fpara, fn) = rec_ci f in
let len = length (fprog) in
(fprog @ [Dec (Suc n) (len + 5), Dec (Suc n) (len + 3),
Goto (len + 1), Inc n, Goto 0], n, max (Suc n) fn))"
declare rec_ci.simps [simp del] rec_ci_s_def[simp del]
rec_ci_z_def[simp del] rec_ci_id.simps[simp del]
mv_boxes.simps[simp del]
mv_box.simps[simp del] addition.simps[simp del]
declare abc_steps_l.simps[simp del] abc_fetch.simps[simp del]
abc_step_l.simps[simp del]
inductive_cases terminate_pr_reverse: "terminate (Pr n f g) xs"
inductive_cases terminate_z_reverse[elim!]: "terminate z xs"
inductive_cases terminate_s_reverse[elim!]: "terminate s xs"
inductive_cases terminate_id_reverse[elim!]: "terminate (id m n) xs"
inductive_cases terminate_cn_reverse[elim!]: "terminate (Cn n f gs) xs"
inductive_cases terminate_mn_reverse[elim!]:"terminate (Mn n f) xs"
fun addition_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow>
nat list \<Rightarrow> bool"
where
"addition_inv (as, lm') m n p lm =
(let sn = lm ! n in
let sm = lm ! m in
lm ! p = 0 \<and>
(if as = 0 then \<exists> x. x \<le> lm ! m \<and> lm' = lm[m := x,
n := (sn + sm - x), p := (sm - x)]
else if as = 1 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
n := (sn + sm - x - 1), p := (sm - x - 1)]
else if as = 2 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
n := (sn + sm - x), p := (sm - x - 1)]
else if as = 3 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
n := (sn + sm - x), p := (sm - x)]
else if as = 4 then \<exists> x. x \<le> lm ! m \<and> lm' = lm[m := x,
n := (sn + sm), p := (sm - x)]
else if as = 5 then \<exists> x. x < lm ! m \<and> lm' = lm[m := x,
n := (sn + sm), p := (sm - x - 1)]
else if as = 6 then \<exists> x. x < lm ! m \<and> lm' =
lm[m := Suc x, n := (sn + sm), p := (sm - x - 1)]
else if as = 7 then lm' = lm[m := sm, n := (sn + sm)]
else False))"
fun addition_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"addition_stage1 (as, lm) m p =
(if as = 0 \<or> as = 1 \<or> as = 2 \<or> as = 3 then 2
else if as = 4 \<or> as = 5 \<or> as = 6 then 1
else 0)"
fun addition_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"addition_stage2 (as, lm) m p =
(if 0 \<le> as \<and> as \<le> 3 then lm ! m
else if 4 \<le> as \<and> as \<le> 6 then lm ! p
else 0)"
fun addition_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"addition_stage3 (as, lm) m p =
(if as = 1 then 4
else if as = 2 then 3
else if as = 3 then 2
else if as = 0 then 1
else if as = 5 then 2
else if as = 6 then 1
else if as = 4 then 0
else 0)"
fun addition_measure :: "((nat \<times> nat list) \<times> nat \<times> nat) \<Rightarrow>
(nat \<times> nat \<times> nat)"
where
"addition_measure ((as, lm), m, p) =
(addition_stage1 (as, lm) m p,
addition_stage2 (as, lm) m p,
addition_stage3 (as, lm) m p)"
definition addition_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times>
((nat \<times> nat list) \<times> nat \<times> nat)) set"
where "addition_LE \<equiv> (inv_image lex_triple addition_measure)"
lemma wf_additon_LE[simp]: "wf addition_LE"
by(auto simp: addition_LE_def lex_triple_def lex_pair_def)
declare addition_inv.simps[simp del]
lemma update_zero_to_zero[simp]: "\<lbrakk>am ! n = (0::nat); n < length am\<rbrakk> \<Longrightarrow> am[n := 0] = am"
apply(simp add: list_update_same_conv)
done
lemma addition_inv_init:
"\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
addition_inv (0, lm) m n p lm"
apply(simp add: addition_inv.simps Let_def )
apply(rule_tac x = "lm ! m" in exI, simp)
done
lemma abs_fetch[simp]:
"abc_fetch 0 (addition m n p) = Some (Dec m 4)"
"abc_fetch (Suc 0) (addition m n p) = Some (Inc n)"
"abc_fetch 2 (addition m n p) = Some (Inc p)"
"abc_fetch 3 (addition m n p) = Some (Goto 0)"
"abc_fetch 4 (addition m n p) = Some (Dec p 7)"
"abc_fetch 5 (addition m n p) = Some (Inc m)"
"abc_fetch 6 (addition m n p) = Some (Goto 4)"
by(simp_all add: abc_fetch.simps addition.simps)
lemma exists_small_list_elem1[simp]:
"\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x \<le> lm ! m; 0 < x\<rbrakk>
\<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - x,
p := lm ! m - x, m := x - Suc 0] =
lm[m := xa, n := lm ! n + lm ! m - Suc xa,
p := lm ! m - Suc xa]"
apply(cases x, simp, simp)
apply(rule_tac x = "x - 1" in exI, simp add: list_update_swap
list_update_overwrite)
done
lemma exists_small_list_elem2[simp]:
"\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
\<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - Suc x,
p := lm ! m - Suc x, n := lm ! n + lm ! m - x]
= lm[m := xa, n := lm ! n + lm ! m - xa,
p := lm ! m - Suc xa]"
apply(rule_tac x = x in exI,
simp add: list_update_swap list_update_overwrite)
done
lemma exists_small_list_elem3[simp]:
"\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
\<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m - x,
p := lm ! m - Suc x, p := lm ! m - x]
= lm[m := xa, n := lm ! n + lm ! m - xa,
p := lm ! m - xa]"
apply(rule_tac x = x in exI, simp add: list_update_overwrite)
done
lemma exists_small_list_elem4[simp]:
"\<lbrakk>m \<noteq> n; p < length lm; lm ! p = (0::nat); m < p; n < p; x < lm ! m\<rbrakk>
\<Longrightarrow> \<exists>xa\<le>lm ! m. lm[m := x, n := lm ! n + lm ! m - x,
p := lm ! m - x] =
lm[m := xa, n := lm ! n + lm ! m - xa,
p := lm ! m - xa]"
apply(rule_tac x = x in exI, simp)
done
lemma exists_small_list_elem5[simp]:
"\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p;
x \<le> lm ! m; lm ! m \<noteq> x\<rbrakk>
\<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m,
p := lm ! m - x, p := lm ! m - Suc x]
= lm[m := xa, n := lm ! n + lm ! m,
p := lm ! m - Suc xa]"
apply(rule_tac x = x in exI, simp add: list_update_overwrite)
done
lemma exists_small_list_elem6[simp]:
"\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
\<Longrightarrow> \<exists>xa<lm ! m. lm[m := x, n := lm ! n + lm ! m,
p := lm ! m - Suc x, m := Suc x]
= lm[m := Suc xa, n := lm ! n + lm ! m,
p := lm ! m - Suc xa]"
apply(rule_tac x = x in exI,
simp add: list_update_swap list_update_overwrite)
done
lemma exists_small_list_elem7[simp]:
"\<lbrakk>m \<noteq> n; p < length lm; lm ! p = 0; m < p; n < p; x < lm ! m\<rbrakk>
\<Longrightarrow> \<exists>xa\<le>lm ! m. lm[m := Suc x, n := lm ! n + lm ! m,
p := lm ! m - Suc x]
= lm[m := xa, n := lm ! n + lm ! m, p := lm ! m - xa]"
apply(rule_tac x = "Suc x" in exI, simp)
done
lemma abc_steps_zero: "abc_steps_l asm ap 0 = asm"
apply(cases asm, simp add: abc_steps_l.simps)
done
lemma list_double_update_2:
"lm[a := x, b := y, a := z] = lm[b := y, a:=z]"
by (metis list_update_overwrite list_update_swap)
declare Let_def[simp]
lemma addition_halt_lemma:
"\<lbrakk>m \<noteq> n; max m n < p; length lm > p\<rbrakk> \<Longrightarrow>
\<forall>na. \<not> (\<lambda>(as, lm') (m, p). as = 7)
(abc_steps_l (0, lm) (addition m n p) na) (m, p) \<and>
addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm
\<longrightarrow> addition_inv (abc_steps_l (0, lm) (addition m n p)
(Suc na)) m n p lm
\<and> ((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p),
abc_steps_l (0, lm) (addition m n p) na, m, p) \<in> addition_LE"
proof -
assume assms:"m\<noteq>n" "max m n < p" "length lm > p"
{ fix na
obtain a b where ab:"abc_steps_l (0, lm) (addition m n p) na = (a, b)" by force
assume assms2: "\<not> (\<lambda>(as, lm') (m, p). as = 7)
(abc_steps_l (0, lm) (addition m n p) na) (m, p)"
"addition_inv (abc_steps_l (0, lm) (addition m n p) na) m n p lm"
have r1:"addition_inv (abc_steps_l (0, lm) (addition m n p)
(Suc na)) m n p lm" using assms(1-3) assms2
unfolding abc_step_red2 ab abc_step_l.simps abc_lm_v.simps abc_lm_s.simps
addition_inv.simps
by (auto split:if_splits simp add: addition_inv.simps Suc_diff_Suc)
have r2:"((abc_steps_l (0, lm) (addition m n p) (Suc na), m, p),
abc_steps_l (0, lm) (addition m n p) na, m, p) \<in> addition_LE" using assms(1-3) assms2
unfolding abc_step_red2 ab
apply(auto split:if_splits simp add: addition_inv.simps abc_steps_zero)
by(auto simp add: addition_LE_def lex_triple_def lex_pair_def
abc_step_l.simps abc_lm_v.simps abc_lm_s.simps split: if_splits)
note r1 r2
}
thus ?thesis by auto
qed
lemma addition_correct':
"\<lbrakk>m \<noteq> n; max m n < p; length lm > p; lm ! p = 0\<rbrakk> \<Longrightarrow>
\<exists> stp. (\<lambda> (as, lm'). as = 7 \<and> addition_inv (as, lm') m n p lm)
(abc_steps_l (0, lm) (addition m n p) stp)"
apply(insert halt_lemma2[of addition_LE
"\<lambda> ((as, lm'), m, p). addition_inv (as, lm') m n p lm"
"\<lambda> stp. (abc_steps_l (0, lm) (addition m n p) stp, m, p)"
"\<lambda> ((as, lm'), m, p). as = 7"],
simp add: abc_steps_zero addition_inv_init)
apply(drule_tac addition_halt_lemma,force,force)
apply (simp,erule_tac exE)
apply(rename_tac na)
apply(rule_tac x = na in exI)
apply(auto)
done
lemma length_addition[simp]: "length (addition a b c) = 7"
by(auto simp: addition.simps)
lemma addition_correct:
assumes "m \<noteq> n" "max m n < p" "length lm > p" "lm ! p = 0"
shows "{\<lambda> a. a = lm} (addition m n p) {\<lambda> nl. addition_inv (7, nl) m n p lm}"
using assms
proof(rule_tac abc_Hoare_haltI, simp)
fix lma
assume "m \<noteq> n" "m < p \<and> n < p" "p < length lm" "lm ! p = 0"
then have "\<exists> stp. (\<lambda> (as, lm'). as = 7 \<and> addition_inv (as, lm') m n p lm)
(abc_steps_l (0, lm) (addition m n p) stp)"
by(rule_tac addition_correct', auto simp: addition_inv.simps)
then obtain stp where "(\<lambda> (as, lm'). as = 7 \<and> addition_inv (as, lm') m n p lm)
(abc_steps_l (0, lm) (addition m n p) stp)"
using exE by presburger
thus "\<exists>na. abc_final (abc_steps_l (0, lm) (addition m n p) na) (addition m n p) \<and>
(\<lambda>nl. addition_inv (7, nl) m n p lm) abc_holds_for abc_steps_l (0, lm) (addition m n p) na"
by(auto intro:exI[of _ stp])
qed
lemma compile_s_correct':
"{\<lambda>nl. nl = n # 0 \<up> 2 @ anything} addition 0 (Suc 0) 2 [+] [Inc (Suc 0)] {\<lambda>nl. nl = n # Suc n # 0 # anything}"
proof(rule_tac abc_Hoare_plus_halt)
show "{\<lambda>nl. nl = n # 0 \<up> 2 @ anything} addition 0 (Suc 0) 2 {\<lambda> nl. addition_inv (7, nl) 0 (Suc 0) 2 (n # 0 \<up> 2 @ anything)}"
by(rule_tac addition_correct, auto simp: numeral_2_eq_2)
next
show "{\<lambda>nl. addition_inv (7, nl) 0 (Suc 0) 2 (n # 0 \<up> 2 @ anything)} [Inc (Suc 0)] {\<lambda>nl. nl = n # Suc n # 0 # anything}"
by(rule_tac abc_Hoare_haltI, rule_tac x = 1 in exI, auto simp: addition_inv.simps
abc_steps_l.simps abc_step_l.simps abc_fetch.simps numeral_2_eq_2 abc_lm_s.simps abc_lm_v.simps)
qed
declare rec_exec.simps[simp del]
lemma abc_comp_commute: "(A [+] B) [+] C = A [+] (B [+] C)"
apply(auto simp: abc_comp.simps abc_shift.simps)
apply(rename_tac x)
apply(case_tac x, auto)
done
lemma compile_z_correct:
"\<lbrakk>rec_ci z = (ap, arity, fp); rec_exec z [n] = r\<rbrakk> \<Longrightarrow>
{\<lambda>nl. nl = n # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = n # r # 0 \<up> (fp - Suc arity) @ anything}"
apply(rule_tac abc_Hoare_haltI)
apply(rule_tac x = 1 in exI)
apply(auto simp: abc_steps_l.simps rec_ci.simps rec_ci_z_def
numeral_2_eq_2 abc_fetch.simps abc_step_l.simps rec_exec.simps)
done
lemma compile_s_correct:
"\<lbrakk>rec_ci s = (ap, arity, fp); rec_exec s [n] = r\<rbrakk> \<Longrightarrow>
{\<lambda>nl. nl = n # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = n # r # 0 \<up> (fp - Suc arity) @ anything}"
apply(auto simp: rec_ci.simps rec_ci_s_def compile_s_correct' rec_exec.simps)
done
lemma compile_id_correct':
assumes "n < length args"
shows "{\<lambda>nl. nl = args @ 0 \<up> 2 @ anything} addition n (length args) (Suc (length args))
{\<lambda>nl. nl = args @ rec_exec (recf.id (length args) n) args # 0 # anything}"
proof -
have "{\<lambda>nl. nl = args @ 0 \<up> 2 @ anything} addition n (length args) (Suc (length args))
{\<lambda>nl. addition_inv (7, nl) n (length args) (Suc (length args)) (args @ 0 \<up> 2 @ anything)}"
using assms
by(rule_tac addition_correct, auto simp: numeral_2_eq_2 nth_append)
thus "?thesis"
using assms
by(simp add: addition_inv.simps rec_exec.simps
nth_append numeral_2_eq_2 list_update_append)
qed
lemma compile_id_correct:
"\<lbrakk>n < m; length xs = m; rec_ci (recf.id m n) = (ap, arity, fp); rec_exec (recf.id m n) xs = r\<rbrakk>
\<Longrightarrow> {\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ r # 0 \<up> (fp - Suc arity) @ anything}"
apply(auto simp: rec_ci.simps rec_ci_id.simps compile_id_correct')
done
lemma cn_merge_gs_tl_app:
"cn_merge_gs (gs @ [g]) pstr =
cn_merge_gs gs pstr [+] cn_merge_gs [g] (pstr + length gs)"
apply(induct gs arbitrary: pstr, simp add: cn_merge_gs.simps, auto)
apply(simp add: abc_comp_commute)
done
lemma footprint_ge:
"rec_ci a = (p, arity, fp) \<Longrightarrow> arity < fp"
proof(induct a)
case (Cn x1 a x3)
then show ?case by(cases "rec_ci a", auto simp:rec_ci.simps)
next
case (Pr x1 a1 a2)
then show ?case by(cases "rec_ci a1";cases "rec_ci a2", auto simp:rec_ci.simps)
next
case (Mn x1 a)
then show ?case by(cases "rec_ci a", auto simp:rec_ci.simps)
qed (auto simp: rec_ci.simps)
lemma param_pattern:
"\<lbrakk>terminate f xs; rec_ci f = (p, arity, fp)\<rbrakk> \<Longrightarrow> length xs = arity"
proof(induct arbitrary: p arity fp rule: terminate.induct)
case (termi_cn f xs gs n) thus ?case
by(cases "rec_ci f", (auto simp: rec_ci.simps))
next
case (termi_pr x g xs n f) thus ?case
by (cases "rec_ci f", cases "rec_ci g", auto simp: rec_ci.simps)
next
case (termi_mn xs n f r) thus ?case
by (cases "rec_ci f", auto simp: rec_ci.simps)
qed (auto simp: rec_ci.simps)
lemma replicate_merge_anywhere:
"x\<up>a @ x\<up>b @ ys = x\<up>(a+b) @ ys"
by(simp add:replicate_add)
fun mv_box_inv :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool"
where
"mv_box_inv (as, lm) m n initlm =
(let plus = initlm ! m + initlm ! n in
length initlm > max m n \<and> m \<noteq> n \<and>
(if as = 0 then \<exists> k l. lm = initlm[m := k, n := l] \<and>
k + l = plus \<and> k \<le> initlm ! m
else if as = 1 then \<exists> k l. lm = initlm[m := k, n := l]
\<and> k + l + 1 = plus \<and> k < initlm ! m
else if as = 2 then \<exists> k l. lm = initlm[m := k, n := l]
\<and> k + l = plus \<and> k \<le> initlm ! m
else if as = 3 then lm = initlm[m := 0, n := plus]
else False))"
fun mv_box_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
where
"mv_box_stage1 (as, lm) m =
(if as = 3 then 0
else 1)"
fun mv_box_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
where
"mv_box_stage2 (as, lm) m = (lm ! m)"
fun mv_box_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat"
where
"mv_box_stage3 (as, lm) m = (if as = 1 then 3
else if as = 2 then 2
else if as = 0 then 1
else 0)"
fun mv_box_measure :: "((nat \<times> nat list) \<times> nat) \<Rightarrow> (nat \<times> nat \<times> nat)"
where
"mv_box_measure ((as, lm), m) =
(mv_box_stage1 (as, lm) m, mv_box_stage2 (as, lm) m,
mv_box_stage3 (as, lm) m)"
definition lex_pair :: "((nat \<times> nat) \<times> nat \<times> nat) set"
where
"lex_pair = less_than <*lex*> less_than"
definition lex_triple ::
"((nat \<times> (nat \<times> nat)) \<times> (nat \<times> (nat \<times> nat))) set"
where
"lex_triple \<equiv> less_than <*lex*> lex_pair"
definition mv_box_LE ::
"(((nat \<times> nat list) \<times> nat) \<times> ((nat \<times> nat list) \<times> nat)) set"
where
"mv_box_LE \<equiv> (inv_image lex_triple mv_box_measure)"
lemma wf_lex_triple: "wf lex_triple"
by (auto simp:lex_triple_def lex_pair_def)
lemma wf_mv_box_le[intro]: "wf mv_box_LE"
by(auto intro:wf_lex_triple simp: mv_box_LE_def)
declare mv_box_inv.simps[simp del]
lemma mv_box_inv_init:
"\<lbrakk>m < length initlm; n < length initlm; m \<noteq> n\<rbrakk> \<Longrightarrow>
mv_box_inv (0, initlm) m n initlm"
apply(simp add: abc_steps_l.simps mv_box_inv.simps)
apply(rule_tac x = "initlm ! m" in exI,
rule_tac x = "initlm ! n" in exI, simp)
done
lemma abc_fetch[simp]:
"abc_fetch 0 (mv_box m n) = Some (Dec m 3)"
"abc_fetch (Suc 0) (mv_box m n) = Some (Inc n)"
"abc_fetch 2 (mv_box m n) = Some (Goto 0)"
"abc_fetch 3 (mv_box m n) = None"
apply(simp_all add: mv_box.simps abc_fetch.simps)
done
lemma replicate_Suc_iff_anywhere: "x # x\<up>b @ ys = x\<up>(Suc b) @ ys"
by simp
lemma exists_smaller_in_list0[simp]:
"\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm;
k + l = initlm ! m + initlm ! n; k \<le> initlm ! m; 0 < k\<rbrakk>
\<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, m := k - Suc 0] =
initlm[m := ka, n := la] \<and>
Suc (ka + la) = initlm ! m + initlm ! n \<and>
ka < initlm ! m"
apply(rule_tac x = "k - Suc 0" in exI, rule_tac x = l in exI, auto)
apply(subgoal_tac
"initlm[m := k, n := l, m := k - Suc 0] =
initlm[n := l, m := k, m := k - Suc 0]",force intro:list_update_swap)
by(simp add: list_update_swap)
lemma exists_smaller_in_list1[simp]:
"\<lbrakk>m \<noteq> n; m < length initlm; n < length initlm;
Suc (k + l) = initlm ! m + initlm ! n;
k < initlm ! m\<rbrakk>
\<Longrightarrow> \<exists>ka la. initlm[m := k, n := l, n := Suc l] =
initlm[m := ka, n := la] \<and>
ka + la = initlm ! m + initlm ! n \<and>
ka \<le> initlm ! m"
apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, auto)
done
lemma abc_steps_prop[simp]:
"\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow>
\<not> (\<lambda>(as, lm) m. as = 3)
(abc_steps_l (0, initlm) (mv_box m n) na) m \<and>
mv_box_inv (abc_steps_l (0, initlm)
(mv_box m n) na) m n initlm \<longrightarrow>
mv_box_inv (abc_steps_l (0, initlm)
(mv_box m n) (Suc na)) m n initlm \<and>
((abc_steps_l (0, initlm) (mv_box m n) (Suc na), m),
abc_steps_l (0, initlm) (mv_box m n) na, m) \<in> mv_box_LE"
apply(rule impI, simp add: abc_step_red2)
apply(cases "(abc_steps_l (0, initlm) (mv_box m n) na)",
simp)
apply(auto split:if_splits simp add:abc_steps_l.simps mv_box_inv.simps)
apply(auto simp add: mv_box_LE_def lex_triple_def lex_pair_def
abc_step_l.simps abc_steps_l.simps
mv_box_inv.simps abc_lm_v.simps abc_lm_s.simps
split: if_splits )
apply(rule_tac x = k in exI, rule_tac x = "Suc l" in exI, simp)
done
lemma mv_box_inv_halt:
"\<lbrakk>length initlm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow>
\<exists> stp. (\<lambda> (as, lm). as = 3 \<and>
mv_box_inv (as, lm) m n initlm)
(abc_steps_l (0::nat, initlm) (mv_box m n) stp)"
apply(insert halt_lemma2[of mv_box_LE
"\<lambda> ((as, lm), m). mv_box_inv (as, lm) m n initlm"
"\<lambda> stp. (abc_steps_l (0, initlm) (mv_box m n) stp, m)"
"\<lambda> ((as, lm), m). as = (3::nat)"
])
apply(insert wf_mv_box_le)
apply(simp add: mv_box_inv_init abc_steps_zero)
apply(erule_tac exE)
by (metis (no_types, lifting) case_prodE' case_prodI)
lemma mv_box_halt_cond:
"\<lbrakk>m \<noteq> n; mv_box_inv (a, b) m n lm; a = 3\<rbrakk> \<Longrightarrow>
b = lm[n := lm ! m + lm ! n, m := 0]"
apply(simp add: mv_box_inv.simps, auto)
apply(simp add: list_update_swap)
done
lemma mv_box_correct':
"\<lbrakk>length lm > max m n; m \<noteq> n\<rbrakk> \<Longrightarrow>
\<exists> stp. abc_steps_l (0::nat, lm) (mv_box m n) stp
= (3, (lm[n := (lm ! m + lm ! n)])[m := 0::nat])"
by(drule mv_box_inv_halt, auto dest:mv_box_halt_cond)
lemma length_mvbox[simp]: "length (mv_box m n) = 3"
by(simp add: mv_box.simps)
lemma mv_box_correct:
"\<lbrakk>length lm > max m n; m\<noteq>n\<rbrakk>
\<Longrightarrow> {\<lambda> nl. nl = lm} mv_box m n {\<lambda> nl. nl = lm[n := (lm ! m + lm ! n), m:=0]}"
apply(drule_tac mv_box_correct', simp)
apply(auto simp: abc_Hoare_halt_def)
by (metis abc_final.simps abc_holds_for.simps length_mvbox)
declare list_update.simps(2)[simp del]
lemma zero_case_rec_exec[simp]:
"\<lbrakk>length xs < gf; gf \<le> ft; n < length gs\<rbrakk>
\<Longrightarrow> (rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything)
[ft + n - length xs := rec_exec (gs ! n) xs, 0 := 0] =
0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything"
using list_update_append[of "rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) (take n gs)"
"0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything" "ft + n - length xs" "rec_exec (gs ! n) xs"]
apply(auto)
apply(cases "length gs - n", simp, simp add: list_update.simps replicate_Suc_iff_anywhere Suc_diff_Suc del: replicate_Suc)
apply(simp add: list_update.simps)
done
lemma compile_cn_gs_correct':
assumes
g_cond: "\<forall>g\<in>set (take n gs). terminate g xs \<and>
(\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> (\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc}))"
and ft: "ft = max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
shows
"{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything}
cn_merge_gs (map rec_ci (take n gs)) ft
{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @
map (\<lambda>i. rec_exec i xs) (take n gs) @ 0\<up>(length gs - n) @ 0 \<up> Suc (length xs) @ anything}"
using g_cond
proof(induct n)
case 0
have "ft > length xs"
using ft
by simp
thus "?case"
apply(rule_tac abc_Hoare_haltI)
apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps replicate_add[THEN sym]
replicate_Suc[THEN sym] del: replicate_Suc)
done
next
case (Suc n)
have ind': "\<forall>g\<in>set (take n gs).
terminate g xs \<and> (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow>
(\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc})) \<Longrightarrow>
{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft
{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything}"
by fact
have g_newcond: "\<forall>g\<in>set (take (Suc n) gs).
terminate g xs \<and> (\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> (\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc}))"
by fact
from g_newcond have ind:
"{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft
{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything}"
apply(rule_tac ind', rule_tac ballI, erule_tac x = g in ballE, simp_all add: take_Suc)
by(cases "n < length gs", simp add:take_Suc_conv_app_nth, simp)
show "?case"
proof(cases "n < length gs")
case True
have h: "n < length gs" by fact
thus "?thesis"
proof(simp add: take_Suc_conv_app_nth cn_merge_gs_tl_app)
obtain gp ga gf where a: "rec_ci (gs!n) = (gp, ga, gf)"
by (metis prod_cases3)
moreover have "min (length gs) n = n"
using h by simp
moreover have
"{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything}
cn_merge_gs (map rec_ci (take n gs)) ft [+] (gp [+] mv_box ga (ft + n))
{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @
rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}"
proof(rule_tac abc_Hoare_plus_halt)
show "{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything} cn_merge_gs (map rec_ci (take n gs)) ft
{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything}"
using ind by simp
next
have x: "gs!n \<in> set (take (Suc n) gs)"
using h
by(simp add: take_Suc_conv_app_nth)
have b: "terminate (gs!n) xs"
using a g_newcond h x
by(erule_tac x = "gs!n" in ballE, simp_all)
hence c: "length xs = ga"
using a param_pattern by metis
have d: "gf > ga" using footprint_ge a by simp
have e: "ft \<ge> gf"
using ft a h Max_ge image_eqI
by(simp, rule_tac max.coboundedI2, rule_tac Max_ge, simp,
rule_tac insertI2,
rule_tac f = "(\<lambda>(aprog, p, n). n)" and x = "rec_ci (gs!n)" in image_eqI, simp,
rule_tac x = "gs!n" in image_eqI, simp, simp)
show "{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @
map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything} gp [+] mv_box ga (ft + n)
{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs)
(take n gs) @ rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}"
proof(rule_tac abc_Hoare_plus_halt)
show "{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything} gp
{\<lambda>nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs)
(take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything}"
proof -
have
"({\<lambda>nl. nl = xs @ 0 \<up> (gf - ga) @ 0\<up>(ft - gf)@map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything}
gp {\<lambda>nl. nl = xs @ (rec_exec (gs!n) xs) # 0 \<up> (gf - Suc ga) @
0\<up>(ft - gf)@map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 \<up> Suc (length xs) @ anything})"
using a g_newcond h x
apply(erule_tac x = "gs!n" in ballE)
apply(simp, simp)
done
thus "?thesis"
using a b c d e
by(simp add: replicate_merge_anywhere)
qed
next
show
"{\<lambda>nl. nl = xs @ rec_exec (gs ! n) xs #
0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything}
mv_box ga (ft + n)
{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @
rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}"
proof -
have "{\<lambda>nl. nl = xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @
map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything}
mv_box ga (ft + n) {\<lambda>nl. nl = (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @
map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything)
[ft + n := (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) (take n gs) @
0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! ga +
(xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @
map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) !
(ft + n), ga := 0]}"
using a c d e h
apply(rule_tac mv_box_correct)
apply(simp, arith, arith)
done
moreover have "(xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @
map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything)
[ft + n := (xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @ map (\<lambda>i. rec_exec i xs) (take n gs) @
0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) ! ga +
(xs @ rec_exec (gs ! n) xs # 0 \<up> (ft - Suc (length xs)) @
map (\<lambda>i. rec_exec i xs) (take n gs) @ 0 \<up> (length gs - n) @ 0 # 0 \<up> length xs @ anything) !
(ft + n), ga := 0]=
xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything"
using a c d e h
by(simp add: list_update_append nth_append length_replicate split: if_splits del: list_update.simps(2), auto)
ultimately show "?thesis"
by(simp)
qed
qed
qed
ultimately show
"{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything}
cn_merge_gs (map rec_ci (take n gs)) ft [+] (case rec_ci (gs ! n) of (gprog, gpara, gn) \<Rightarrow>
gprog [+] mv_box gpara (ft + min (length gs) n))
{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) (take n gs) @ rec_exec (gs ! n) xs # 0 \<up> (length gs - Suc n) @ 0 # 0 \<up> length xs @ anything}"
by simp
qed
next
case False
have h: "\<not> n < length gs" by fact
hence ind':
"{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything} cn_merge_gs (map rec_ci gs) ft
{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @ map (\<lambda>i. rec_exec i xs) gs @ 0 \<up> Suc (length xs) @ anything}"
using ind
by simp
thus "?thesis"
using h
by(simp)
qed
qed
lemma compile_cn_gs_correct:
assumes
g_cond: "\<forall>g\<in>set gs. terminate g xs \<and>
(\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> (\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc}))"
and ft: "ft = max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
shows
"{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft + length gs) @ anything}
cn_merge_gs (map rec_ci gs) ft
{\<lambda>nl. nl = xs @ 0 \<up> (ft - length xs) @
map (\<lambda>i. rec_exec i xs) gs @ 0 \<up> Suc (length xs) @ anything}"
using assms
using compile_cn_gs_correct'[of "length gs" gs xs ft ffp anything ]
apply(auto)
done
lemma length_mvboxes[simp]: "length (mv_boxes aa ba n) = 3*n"
by(induct n, auto simp: mv_boxes.simps)
lemma exp_suc: "a\<up>Suc b = a\<up>b @ [a]"
by(simp add: exp_ind del: replicate.simps)
lemma last_0[simp]:
"\<lbrakk>Suc n \<le> ba - aa; length lm2 = Suc n;
length lm3 = ba - Suc (aa + n)\<rbrakk>
\<Longrightarrow> (last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = (0::nat)"
proof -
assume h: "Suc n \<le> ba - aa"
and g: "length lm2 = Suc n" "length lm3 = ba - Suc (aa + n)"
from h and g have k: "ba - aa = Suc (length lm3 + n)"
by arith
from k show
"(last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba - aa) = 0"
apply(simp, insert g)
apply(simp add: nth_append)
done
qed
lemma butlast_last[simp]: "length lm1 = aa \<Longrightarrow>
(lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (aa + n) = last lm2"
apply(simp add: nth_append)
done
lemma arith_as_simp[simp]: "\<lbrakk>Suc n \<le> ba - aa; aa < ba\<rbrakk> \<Longrightarrow>
(ba < Suc (aa + (ba - Suc (aa + n) + n))) = False"
apply arith
done
lemma butlast_elem[simp]: "\<lbrakk>Suc n \<le> ba - aa; aa < ba; length lm1 = aa;
length lm2 = Suc n; length lm3 = ba - Suc (aa + n)\<rbrakk>
\<Longrightarrow> (lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4) ! (ba + n) = 0"
using nth_append[of "lm1 @ (0::'a)\<up>n @ last lm2 # lm3 @ butlast lm2"
"(0::'a) # lm4" "ba + n"]
apply(simp)
done
lemma update_butlast_eq0[simp]:
"\<lbrakk>Suc n \<le> ba - aa; aa < ba; length lm1 = aa; length lm2 = Suc n;
length lm3 = ba - Suc (aa + n)\<rbrakk>
\<Longrightarrow> (lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ (0::nat) # lm4)
[ba + n := last lm2, aa + n := 0] =
lm1 @ 0 # 0\<up>n @ lm3 @ lm2 @ lm4"
using list_update_append[of "lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2" "0 # lm4"
"ba + n" "last lm2"]
apply(simp add: list_update_append list_update.simps(2-) replicate_Suc_iff_anywhere exp_suc
del: replicate_Suc)
apply(cases lm2, simp, simp)
done
lemma update_butlast_eq1[simp]:
"\<lbrakk>Suc (length lm1 + n) \<le> ba; length lm2 = Suc n; length lm3 = ba - Suc (length lm1 + n);
\<not> ba - Suc (length lm1) < ba - Suc (length lm1 + n); \<not> ba + n - length lm1 < n\<rbrakk>
\<Longrightarrow> (0::nat) \<up> n @ (last lm2 # lm3 @ butlast lm2 @ 0 # lm4)[ba - length lm1 := last lm2, 0 := 0] =
0 # 0 \<up> n @ lm3 @ lm2 @ lm4"
apply(subgoal_tac "ba - length lm1 = Suc n + length lm3", simp add: list_update.simps(2-) list_update_append)
apply(simp add: replicate_Suc_iff_anywhere exp_suc del: replicate_Suc)
apply(cases lm2, simp, simp)
apply(auto)
done
lemma mv_boxes_correct:
"\<lbrakk>aa + n \<le> ba; ba > aa; length lm1 = aa; length lm2 = n; length lm3 = ba - aa - n\<rbrakk>
\<Longrightarrow> {\<lambda> nl. nl = lm1 @ lm2 @ lm3 @ 0\<up>n @ lm4} (mv_boxes aa ba n)
{\<lambda> nl. nl = lm1 @ 0\<up>n @ lm3 @ lm2 @ lm4}"
proof(induct n arbitrary: lm2 lm3 lm4)
case 0
thus "?case"
by(simp add: mv_boxes.simps abc_Hoare_halt_def, rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
next
case (Suc n)
have ind:
"\<And>lm2 lm3 lm4.
\<lbrakk>aa + n \<le> ba; aa < ba; length lm1 = aa; length lm2 = n; length lm3 = ba - aa - n\<rbrakk>
\<Longrightarrow> {\<lambda>nl. nl = lm1 @ lm2 @ lm3 @ 0 \<up> n @ lm4} mv_boxes aa ba n {\<lambda>nl. nl = lm1 @ 0 \<up> n @ lm3 @ lm2 @ lm4}"
by fact
have h1: "aa + Suc n \<le> ba" by fact
have h2: "aa < ba" by fact
have h3: "length lm1 = aa" by fact
have h4: "length lm2 = Suc n" by fact
have h5: "length lm3 = ba - aa - Suc n" by fact
have "{\<lambda>nl. nl = lm1 @ lm2 @ lm3 @ 0 \<up> Suc n @ lm4} mv_boxes aa ba n [+] mv_box (aa + n) (ba + n)
{\<lambda>nl. nl = lm1 @ 0 \<up> Suc n @ lm3 @ lm2 @ lm4}"
proof(rule_tac abc_Hoare_plus_halt)
have "{\<lambda>nl. nl = lm1 @ butlast lm2 @ (last lm2 # lm3) @ 0 \<up> n @ (0 # lm4)} mv_boxes aa ba n
{\<lambda> nl. nl = lm1 @ 0\<up>n @ (last lm2 # lm3) @ butlast lm2 @ (0 # lm4)}"
using h1 h2 h3 h4 h5
by(rule_tac ind, simp_all)
moreover have " lm1 @ butlast lm2 @ (last lm2 # lm3) @ 0 \<up> n @ (0 # lm4)
= lm1 @ lm2 @ lm3 @ 0 \<up> Suc n @ lm4"
using h4
by(simp add: replicate_Suc[THEN sym] exp_suc del: replicate_Suc,
cases lm2, simp_all)
ultimately show "{\<lambda>nl. nl = lm1 @ lm2 @ lm3 @ 0 \<up> Suc n @ lm4} mv_boxes aa ba n
{\<lambda> nl. nl = lm1 @ 0\<up>n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4}"
by (metis append_Cons)
next
let ?lm = "lm1 @ 0 \<up> n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4"
have "{\<lambda>nl. nl = ?lm} mv_box (aa + n) (ba + n)
{\<lambda> nl. nl = ?lm[(ba + n) := ?lm!(aa+n) + ?lm!(ba+n), (aa+n):=0]}"
using h1 h2 h3 h4 h5
by(rule_tac mv_box_correct, simp_all)
moreover have "?lm[(ba + n) := ?lm!(aa+n) + ?lm!(ba+n), (aa+n):=0]
= lm1 @ 0 \<up> Suc n @ lm3 @ lm2 @ lm4"
using h1 h2 h3 h4 h5
by(auto simp: nth_append list_update_append split: if_splits)
ultimately show "{\<lambda>nl. nl = lm1 @ 0 \<up> n @ last lm2 # lm3 @ butlast lm2 @ 0 # lm4} mv_box (aa + n) (ba + n)
{\<lambda>nl. nl = lm1 @ 0 \<up> Suc n @ lm3 @ lm2 @ lm4}"
by simp
qed
thus "?case"
by(simp add: mv_boxes.simps)
qed
lemma update_butlast_eq2[simp]:
"\<lbrakk>Suc n \<le> aa - length lm1; length lm1 < aa;
length lm2 = aa - Suc (length lm1 + n);
length lm3 = Suc n;
\<not> aa - Suc (length lm1) < aa - Suc (length lm1 + n);
\<not> aa + n - length lm1 < n\<rbrakk>
\<Longrightarrow> butlast lm3 @ ((0::nat) # lm2 @ 0 \<up> n @ last lm3 # lm4)[0 := last lm3, aa - length lm1 := 0] = lm3 @ lm2 @ 0 # 0 \<up> n @ lm4"
apply(subgoal_tac "aa - length lm1 = length lm2 + Suc n")
apply(simp add: list_update.simps list_update_append)
apply(simp add: replicate_Suc[THEN sym] exp_suc del: replicate_Suc)
apply(cases lm3, simp, simp)
apply(auto)
done
lemma mv_boxes_correct2:
"\<lbrakk>n \<le> aa - ba;
ba < aa;
length (lm1::nat list) = ba;
length (lm2::nat list) = aa - ba - n;
length (lm3::nat list) = n\<rbrakk>
\<Longrightarrow>{\<lambda> nl. nl = lm1 @ 0\<up>n @ lm2 @ lm3 @ lm4}
(mv_boxes aa ba n)
{\<lambda> nl. nl = lm1 @ lm3 @ lm2 @ 0\<up>n @ lm4}"
proof(induct n arbitrary: lm2 lm3 lm4)
case 0
thus "?case"
by(simp add: mv_boxes.simps abc_Hoare_halt_def, rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
next
case (Suc n)
have ind:
"\<And>lm2 lm3 lm4.
\<lbrakk>n \<le> aa - ba; ba < aa; length lm1 = ba; length lm2 = aa - ba - n; length lm3 = n\<rbrakk>
\<Longrightarrow> {\<lambda>nl. nl = lm1 @ 0 \<up> n @ lm2 @ lm3 @ lm4} mv_boxes aa ba n {\<lambda>nl. nl = lm1 @ lm3 @ lm2 @ 0 \<up> n @ lm4}"
by fact
have h1: "Suc n \<le> aa - ba" by fact
have h2: "ba < aa" by fact
have h3: "length lm1 = ba" by fact
have h4: "length lm2 = aa - ba - Suc n" by fact
have h5: "length lm3 = Suc n" by fact
have "{\<lambda>nl. nl = lm1 @ 0 \<up> Suc n @ lm2 @ lm3 @ lm4} mv_boxes aa ba n [+] mv_box (aa + n) (ba + n)
{\<lambda>nl. nl = lm1 @ lm3 @ lm2 @ 0 \<up> Suc n @ lm4}"
proof(rule_tac abc_Hoare_plus_halt)
have "{\<lambda> nl. nl = lm1 @ 0 \<up> n @ (0 # lm2) @ (butlast lm3) @ (last lm3 # lm4)} mv_boxes aa ba n
{\<lambda> nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0\<up>n @ (last lm3 # lm4)}"
using h1 h2 h3 h4 h5
by(rule_tac ind, simp_all)
moreover have "lm1 @ 0 \<up> n @ (0 # lm2) @ (butlast lm3) @ (last lm3 # lm4)
= lm1 @ 0 \<up> Suc n @ lm2 @ lm3 @ lm4"
using h5
by(simp add: replicate_Suc_iff_anywhere exp_suc
del: replicate_Suc, cases lm3, simp_all)
ultimately show "{\<lambda>nl. nl = lm1 @ 0 \<up> Suc n @ lm2 @ lm3 @ lm4} mv_boxes aa ba n
{\<lambda> nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0\<up>n @ (last lm3 # lm4)}"
by metis
next
thm mv_box_correct
let ?lm = "lm1 @ butlast lm3 @ (0 # lm2) @ 0 \<up> n @ last lm3 # lm4"
have "{\<lambda>nl. nl = ?lm} mv_box (aa + n) (ba + n)
{\<lambda>nl. nl = ?lm[ba+n := ?lm!(aa+n)+?lm!(ba+n), (aa+n):=0]}"
using h1 h2 h3 h4 h5
by(rule_tac mv_box_correct, simp_all)
moreover have "?lm[ba+n := ?lm!(aa+n)+?lm!(ba+n), (aa+n):=0]
= lm1 @ lm3 @ lm2 @ 0 \<up> Suc n @ lm4"
using h1 h2 h3 h4 h5
by(auto simp: nth_append list_update_append split: if_splits)
ultimately show "{\<lambda>nl. nl = lm1 @ butlast lm3 @ (0 # lm2) @ 0 \<up> n @ last lm3 # lm4} mv_box (aa + n) (ba + n)
{\<lambda>nl. nl = lm1 @ lm3 @ lm2 @ 0 \<up> Suc n @ lm4}"
by simp
qed
thus "?case"
by(simp add: mv_boxes.simps)
qed
lemma save_paras:
"{\<lambda>nl. nl = xs @ 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length xs) @
map (\<lambda>i. rec_exec i xs) gs @ 0 \<up> Suc (length xs) @ anything}
mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs)
{\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ map (\<lambda>i. rec_exec i xs) gs @ 0 # xs @ anything}"
proof -
let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
have "{\<lambda>nl. nl = [] @ xs @ (0\<up>(?ft - length xs) @ map (\<lambda>i. rec_exec i xs) gs @ [0]) @
0 \<up> (length xs) @ anything} mv_boxes 0 (Suc ?ft + length gs) (length xs)
{\<lambda>nl. nl = [] @ 0 \<up> (length xs) @ (0\<up>(?ft - length xs) @ map (\<lambda>i. rec_exec i xs) gs @ [0]) @ xs @ anything}"
by(rule_tac mv_boxes_correct, auto)
thus "?thesis"
by(simp add: replicate_merge_anywhere)
qed
lemma length_le_max_insert_rec_ci[intro]:
"length gs \<le> ffp \<Longrightarrow> length gs \<le> max x1 (Max (insert ffp (x2 ` x3 ` set gs)))"
apply(rule_tac max.coboundedI2)
apply(simp add: Max_ge_iff)
done
lemma restore_new_paras:
"ffp \<ge> length gs
\<Longrightarrow> {\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ map (\<lambda>i. rec_exec i xs) gs @ 0 # xs @ anything}
mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs)
{\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @ 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ 0 # xs @ anything}"
proof -
let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
assume j: "ffp \<ge> length gs"
hence "{\<lambda> nl. nl = [] @ 0\<up>length gs @ 0\<up>(?ft - length gs) @ map (\<lambda>i. rec_exec i xs) gs @ ((0 # xs) @ anything)}
mv_boxes ?ft 0 (length gs)
{\<lambda> nl. nl = [] @ map (\<lambda>i. rec_exec i xs) gs @ 0\<up>(?ft - length gs) @ 0\<up>length gs @ ((0 # xs) @ anything)}"
by(rule_tac mv_boxes_correct2, auto)
moreover have "?ft \<ge> length gs"
using j
by(auto)
ultimately show "?thesis"
using j
by(simp add: replicate_merge_anywhere le_add_diff_inverse)
qed
lemma le_max_insert[intro]: "ffp \<le> max x0 (Max (insert ffp (x1 ` x2 ` set gs)))"
by (rule max.coboundedI2) auto
declare max_less_iff_conj[simp del]
lemma save_rs:
"\<lbrakk>far = length gs;
ffp \<le> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)));
far < ffp\<rbrakk>
\<Longrightarrow> {\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @
rec_exec (Cn (length xs) f gs) xs # 0 \<up> max (Suc (length xs))
(Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @ xs @ anything}
mv_box far (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))))
{\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @
0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @
rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}"
proof -
let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
thm mv_box_correct
let ?lm= " map (\<lambda>i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> ?ft @ xs @ anything"
assume h: "far = length gs" "ffp \<le> ?ft" "far < ffp"
hence "{\<lambda> nl. nl = ?lm} mv_box far ?ft {\<lambda> nl. nl = ?lm[?ft := ?lm!far + ?lm!?ft, far := 0]}"
apply(rule_tac mv_box_correct)
by( auto)
moreover have "?lm[?ft := ?lm!far + ?lm!?ft, far := 0]
= map (\<lambda>i. rec_exec i xs) gs @
0 \<up> (?ft - length gs) @
rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything"
using h
apply(simp add: nth_append)
using list_update_length[of "map (\<lambda>i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs #
0 \<up> (?ft - Suc (length gs))" 0 "0 \<up> length gs @ xs @ anything" "rec_exec (Cn (length xs) f gs) xs"]
apply(simp add: replicate_merge_anywhere replicate_Suc_iff_anywhere del: replicate_Suc)
by(simp add: list_update_append list_update.simps replicate_Suc_iff_anywhere del: replicate_Suc)
ultimately show "?thesis"
by(simp)
qed
lemma length_empty_boxes[simp]: "length (empty_boxes n) = 2*n"
apply(induct n, simp, simp)
done
lemma empty_one_box_correct:
"{\<lambda>nl. nl = 0 \<up> n @ x # lm} [Dec n 2, Goto 0] {\<lambda>nl. nl = 0 # 0 \<up> n @ lm}"
proof(induct x)
case 0
thus "?case"
by(simp add: abc_Hoare_halt_def,
rule_tac x = 1 in exI, simp add: abc_steps_l.simps
abc_step_l.simps abc_fetch.simps abc_lm_v.simps nth_append abc_lm_s.simps
replicate_Suc[THEN sym] exp_suc del: replicate_Suc)
next
case (Suc x)
have "{\<lambda>nl. nl = 0 \<up> n @ x # lm} [Dec n 2, Goto 0] {\<lambda>nl. nl = 0 # 0 \<up> n @ lm}"
by fact
then obtain stp where "abc_steps_l (0, 0 \<up> n @ x # lm) [Dec n 2, Goto 0] stp
= (Suc (Suc 0), 0 # 0 \<up> n @ lm)"
apply(auto simp: abc_Hoare_halt_def)
by (smt abc_final.simps abc_holds_for.elims(2) length_Cons list.size(3))
moreover have "abc_steps_l (0, 0\<up>n @ Suc x # lm) [Dec n 2, Goto 0] (Suc (Suc 0))
= (0, 0 \<up> n @ x # lm)"
by(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps abc_lm_v.simps
nth_append abc_lm_s.simps list_update.simps list_update_append)
ultimately have "abc_steps_l (0, 0\<up>n @ Suc x # lm) [Dec n 2, Goto 0] (Suc (Suc 0) + stp)
= (Suc (Suc 0), 0 # 0\<up>n @ lm)"
by(simp only: abc_steps_add)
thus "?case"
apply(simp add: abc_Hoare_halt_def)
apply(rule_tac x = "Suc (Suc stp)" in exI, simp)
done
qed
lemma empty_boxes_correct:
"length lm \<ge> n \<Longrightarrow>
{\<lambda> nl. nl = lm} empty_boxes n {\<lambda> nl. nl = 0\<up>n @ drop n lm}"
proof(induct n)
case 0
thus "?case"
by(simp add: empty_boxes.simps abc_Hoare_halt_def,
rule_tac x = 0 in exI, simp add: abc_steps_l.simps)
next
case (Suc n)
have ind: "n \<le> length lm \<Longrightarrow> {\<lambda>nl. nl = lm} empty_boxes n {\<lambda>nl. nl = 0 \<up> n @ drop n lm}" by fact
have h: "Suc n \<le> length lm" by fact
have "{\<lambda>nl. nl = lm} empty_boxes n [+] [Dec n 2, Goto 0] {\<lambda>nl. nl = 0 # 0 \<up> n @ drop (Suc n) lm}"
proof(rule_tac abc_Hoare_plus_halt)
show "{\<lambda>nl. nl = lm} empty_boxes n {\<lambda>nl. nl = 0 \<up> n @ drop n lm}"
using h
by(rule_tac ind, simp)
next
show "{\<lambda>nl. nl = 0 \<up> n @ drop n lm} [Dec n 2, Goto 0] {\<lambda>nl. nl = 0 # 0 \<up> n @ drop (Suc n) lm}"
using empty_one_box_correct[of n "lm ! n" "drop (Suc n) lm"]
using h
by(simp add: Cons_nth_drop_Suc)
qed
thus "?case"
by(simp add: empty_boxes.simps)
qed
lemma insert_dominated[simp]: "length gs \<le> ffp \<Longrightarrow>
length gs + (max xs (Max (insert ffp (x1 ` x2 ` set gs))) - length gs) =
max xs (Max (insert ffp (x1 ` x2 ` set gs)))"
apply(rule_tac le_add_diff_inverse)
apply(rule_tac max.coboundedI2)
apply(simp add: Max_ge_iff)
done
lemma clean_paras:
"ffp \<ge> length gs \<Longrightarrow>
{\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @
0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length gs) @
rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}
empty_boxes (length gs)
{\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @
rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}"
proof-
let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
assume h: "length gs \<le> ffp"
let ?lm = "map (\<lambda>i. rec_exec i xs) gs @ 0 \<up> (?ft - length gs) @
rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything"
have "{\<lambda> nl. nl = ?lm} empty_boxes (length gs) {\<lambda> nl. nl = 0\<up>length gs @ drop (length gs) ?lm}"
by(rule_tac empty_boxes_correct, simp)
moreover have "0\<up>length gs @ drop (length gs) ?lm
= 0 \<up> ?ft @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything"
using h
by(simp add: replicate_merge_anywhere)
ultimately show "?thesis"
by metis
qed
lemma restore_rs:
"{\<lambda>nl. nl = 0 \<up> max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) @
rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything}
mv_box (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs)
{\<lambda>nl. nl = 0 \<up> length xs @
rec_exec (Cn (length xs) f gs) xs #
0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - (length xs)) @
0 \<up> length gs @ xs @ anything}"
proof -
let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
let ?lm = "0\<up>(length xs) @ 0\<up>(?ft - (length xs)) @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> length gs @ xs @ anything"
thm mv_box_correct
have "{\<lambda> nl. nl = ?lm} mv_box ?ft (length xs) {\<lambda> nl. nl = ?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0]}"
by(rule_tac mv_box_correct, simp, simp)
moreover have "?lm[length xs := ?lm!?ft + ?lm!(length xs), ?ft := 0]
= 0 \<up> length xs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> (?ft - (length xs)) @ 0 \<up> length gs @ xs @ anything"
apply(auto simp: list_update_append nth_append) (* ~ 7 sec *)
apply(cases ?ft, simp_all add: Suc_diff_le list_update.simps)
apply(simp add: exp_suc replicate_Suc[THEN sym] del: replicate_Suc)
done
ultimately show "?thesis"
by(simp add: replicate_merge_anywhere)
qed
lemma restore_orgin_paras:
"{\<lambda>nl. nl = 0 \<up> length xs @
rec_exec (Cn (length xs) f gs) xs #
0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) - length xs) @ 0 \<up> length gs @ xs @ anything}
mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs)
{\<lambda>nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up>
(max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}"
proof -
let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
thm mv_boxes_correct2
have "{\<lambda> nl. nl = [] @ 0\<up>(length xs) @ (rec_exec (Cn (length xs) f gs) xs # 0 \<up> (?ft - length xs) @ 0 \<up> length gs) @ xs @ anything}
mv_boxes (Suc ?ft + length gs) 0 (length xs)
{\<lambda> nl. nl = [] @ xs @ (rec_exec (Cn (length xs) f gs) xs # 0 \<up> (?ft - length xs) @ 0 \<up> length gs) @ 0\<up>length xs @ anything}"
by(rule_tac mv_boxes_correct2, auto)
thus "?thesis"
by(simp add: replicate_merge_anywhere)
qed
lemma compile_cn_correct':
assumes f_ind:
"\<And> anything r. rec_exec f (map (\<lambda>g. rec_exec g xs) gs) = rec_exec (Cn (length xs) f gs) xs \<Longrightarrow>
{\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ 0 \<up> (ffp - far) @ anything} fap
{\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> (ffp - Suc far) @ anything}"
and compile: "rec_ci f = (fap, far, ffp)"
and term_f: "terminate f (map (\<lambda>g. rec_exec g xs) gs)"
and g_cond: "\<forall>g\<in>set gs. terminate g xs \<and>
(\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow>
(\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc}))"
shows
"{\<lambda>nl. nl = xs @ 0 # 0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}
cn_merge_gs (map rec_ci gs) (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) [+]
(mv_boxes 0 (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) (length xs) [+]
(mv_boxes (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) 0 (length gs) [+]
(fap [+] (mv_box far (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) [+]
(empty_boxes (length gs) [+]
(mv_box (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))) (length xs) [+]
mv_boxes (Suc (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs)) 0 (length xs)))))))
{\<lambda>nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs #
0 \<up> (max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs))) + length gs) @ anything}"
proof -
let ?ft = "max (Suc (length xs)) (Max (insert ffp ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
let ?A = "cn_merge_gs (map rec_ci gs) ?ft"
let ?B = "mv_boxes 0 (Suc (?ft+length gs)) (length xs)"
let ?C = "mv_boxes ?ft 0 (length gs)"
let ?D = fap
let ?E = "mv_box far ?ft"
let ?F = "empty_boxes (length gs)"
let ?G = "mv_box ?ft (length xs)"
let ?H = "mv_boxes (Suc (?ft + length gs)) 0 (length xs)"
let ?P1 = "\<lambda>nl. nl = xs @ 0 # 0 \<up> (?ft + length gs) @ anything"
let ?S = "\<lambda>nl. nl = xs @ rec_exec (Cn (length xs) f gs) xs # 0 \<up> (?ft + length gs) @ anything"
let ?Q1 = "\<lambda> nl. nl = xs @ 0\<up>(?ft - length xs) @ map (\<lambda> i. rec_exec i xs) gs @ 0\<up>(Suc (length xs)) @ anything"
show "{?P1} (?A [+] (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))))) {?S}"
proof(rule_tac abc_Hoare_plus_halt)
show "{?P1} ?A {?Q1}"
using g_cond
by(rule_tac compile_cn_gs_correct, auto)
next
let ?Q2 = "\<lambda>nl. nl = 0 \<up> ?ft @
map (\<lambda>i. rec_exec i xs) gs @ 0 # xs @ anything"
show "{?Q1} (?B [+] (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H)))))) {?S}"
proof(rule_tac abc_Hoare_plus_halt)
show "{?Q1} ?B {?Q2}"
by(rule_tac save_paras)
next
let ?Q3 = "\<lambda> nl. nl = map (\<lambda>i. rec_exec i xs) gs @ 0\<up>?ft @ 0 # xs @ anything"
show "{?Q2} (?C [+] (?D [+] (?E [+] (?F [+] (?G [+] ?H))))) {?S}"
proof(rule_tac abc_Hoare_plus_halt)
have "ffp \<ge> length gs"
using compile term_f
apply(subgoal_tac "length gs = far")
apply(drule_tac footprint_ge, simp)
by(drule_tac param_pattern, auto)
thus "{?Q2} ?C {?Q3}"
by(erule_tac restore_new_paras)
next
let ?Q4 = "\<lambda> nl. nl = map (\<lambda>i. rec_exec i xs) gs @ rec_exec (Cn (length xs) f gs) xs # 0\<up>?ft @ xs @ anything"
have a: "far = length gs"
using compile term_f
by(drule_tac param_pattern, auto)
have b:"?ft \<ge> ffp"
by auto
have c: "ffp > far"
using compile
by(erule_tac footprint_ge)
show "{?Q3} (?D [+] (?E [+] (?F [+] (?G [+] ?H)))) {?S}"
proof(rule_tac abc_Hoare_plus_halt)
have "{\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ 0 \<up> (ffp - far) @ 0\<up>(?ft - ffp + far) @ 0 # xs @ anything} fap
{\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ rec_exec (Cn (length xs) f gs) xs #
0 \<up> (ffp - Suc far) @ 0\<up>(?ft - ffp + far) @ 0 # xs @ anything}"
by(rule_tac f_ind, simp add: rec_exec.simps)
thus "{?Q3} fap {?Q4}"
using a b c
by(simp add: replicate_merge_anywhere,
cases "?ft", simp_all add: exp_suc del: replicate_Suc)
next
let ?Q5 = "\<lambda>nl. nl = map (\<lambda>i. rec_exec i xs) gs @
0\<up>(?ft - length gs) @ rec_exec (Cn (length xs) f gs) xs # 0\<up>(length gs)@ xs @ anything"
show "{?Q4} (?E [+] (?F [+] (?G [+] ?H))) {?S}"
proof(rule_tac abc_Hoare_plus_halt)
from a b c show "{?Q4} ?E {?Q5}"
by(erule_tac save_rs, simp_all)
next
let ?Q6 = "\<lambda>nl. nl = 0\<up>?ft @ rec_exec (Cn (length xs) f gs) xs # 0\<up>(length gs)@ xs @ anything"
show "{?Q5} (?F [+] (?G [+] ?H)) {?S}"
proof(rule_tac abc_Hoare_plus_halt)
have "length gs \<le> ffp" using a b c
by simp
thus "{?Q5} ?F {?Q6}"
by(erule_tac clean_paras)
next
let ?Q7 = "\<lambda>nl. nl = 0\<up>length xs @ rec_exec (Cn (length xs) f gs) xs # 0\<up>(?ft - (length xs)) @ 0\<up>(length gs)@ xs @ anything"
show "{?Q6} (?G [+] ?H) {?S}"
proof(rule_tac abc_Hoare_plus_halt)
show "{?Q6} ?G {?Q7}"
by(rule_tac restore_rs)
next
show "{?Q7} ?H {?S}"
by(rule_tac restore_orgin_paras)
qed
qed
qed
qed
qed
qed
qed
qed
lemma compile_cn_correct:
assumes termi_f: "terminate f (map (\<lambda>g. rec_exec g xs) gs)"
and f_ind: "\<And>ap arity fp anything.
rec_ci f = (ap, arity, fp)
\<Longrightarrow> {\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ 0 \<up> (fp - arity) @ anything} ap
{\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ rec_exec f (map (\<lambda>g. rec_exec g xs) gs) # 0 \<up> (fp - Suc arity) @ anything}"
and g_cond:
"\<forall>g\<in>set gs. terminate g xs \<and>
(\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow> (\<forall>xc. {\<lambda>nl. nl = xs @ 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ rec_exec g xs # 0 \<up> (xb - Suc xa) @ xc}))"
and compile: "rec_ci (Cn n f gs) = (ap, arity, fp)"
and len: "length xs = n"
shows "{\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ rec_exec (Cn n f gs) xs # 0 \<up> (fp - Suc arity) @ anything}"
proof(cases "rec_ci f")
fix fap far ffp
assume h: "rec_ci f = (fap, far, ffp)"
then have f_newind: "\<And> anything .{\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ 0 \<up> (ffp - far) @ anything} fap
{\<lambda>nl. nl = map (\<lambda>g. rec_exec g xs) gs @ rec_exec f (map (\<lambda>g. rec_exec g xs) gs) # 0 \<up> (ffp - Suc far) @ anything}"
by(rule_tac f_ind, simp_all)
thus "{\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ rec_exec (Cn n f gs) xs # 0 \<up> (fp - Suc arity) @ anything}"
using compile len h termi_f g_cond
apply(auto simp: rec_ci.simps abc_comp_commute)
apply(rule_tac compile_cn_correct', simp_all)
done
qed
lemma mv_box_correct_simp[simp]:
"\<lbrakk>length xs = n; ft = max (n+3) (max fft gft)\<rbrakk>
\<Longrightarrow> {\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft - n) @ anything} mv_box n ft
{\<lambda>nl. nl = xs @ 0 # 0 \<up> (ft - n) @ anything}"
using mv_box_correct[of n ft "xs @ 0 # 0 \<up> (ft - n) @ anything"]
by(auto)
lemma length_under_max[simp]: "length xs < max (length xs + 3) fft"
by auto
lemma save_init_rs:
"\<lbrakk>length xs = n; ft = max (n+3) (max fft gft)\<rbrakk>
\<Longrightarrow> {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (ft - n) @ anything} mv_box n (Suc n)
{\<lambda>nl. nl = xs @ 0 # rec_exec f xs # 0 \<up> (ft - Suc n) @ anything}"
using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \<up> (ft - n) @ anything"]
apply(auto simp: list_update_append list_update.simps nth_append split: if_splits)
apply(cases "(max (length xs + 3) (max fft gft))", simp_all add: list_update.simps Suc_diff_le)
done
lemma less_then_max_plus2[simp]: "n + (2::nat) < max (n + 3) x"
by auto
lemma less_then_max_plus3[simp]: "n < max (n + (3::nat)) x"
by auto
lemma mv_box_max_plus_3_correct[simp]:
"length xs = n \<Longrightarrow>
{\<lambda>nl. nl = xs @ x # 0 \<up> (max (n + (3::nat)) (max fft gft) - n) @ anything} mv_box n (max (n + 3) (max fft gft))
{\<lambda>nl. nl = xs @ 0 \<up> (max (n + 3) (max fft gft) - n) @ x # anything}"
proof -
assume h: "length xs = n"
let ?ft = "max (n+3) (max fft gft)"
let ?lm = "xs @ x # 0\<up>(?ft - Suc n) @ 0 # anything"
have g: "?ft > n + 2"
by simp
thm mv_box_correct
have a: "{\<lambda> nl. nl = ?lm} mv_box n ?ft {\<lambda> nl. nl = ?lm[?ft := ?lm!n + ?lm!?ft, n := 0]}"
using h
by(rule_tac mv_box_correct, auto)
have b:"?lm = xs @ x # 0 \<up> (max (n + 3) (max fft gft) - n) @ anything"
by(cases ?ft, simp_all add: Suc_diff_le exp_suc del: replicate_Suc)
have c: "?lm[?ft := ?lm!n + ?lm!?ft, n := 0] = xs @ 0 \<up> (max (n + 3) (max fft gft) - n) @ x # anything"
using h g
apply(auto simp: nth_append list_update_append split: if_splits)
using list_update_append[of "x # 0 \<up> (max (length xs + 3) (max fft gft) - Suc (length xs))" "0 # anything"
"max (length xs + 3) (max fft gft) - length xs" "x"]
apply(auto simp: if_splits)
apply(simp add: list_update.simps replicate_Suc[THEN sym] del: replicate_Suc)
done
from a c show "?thesis"
using h
apply(simp)
using b
by simp
qed
lemma max_less_suc_suc[simp]: "max n (Suc n) < Suc (Suc (max (n + 3) x + anything - Suc 0))"
by arith
lemma suc_less_plus_3[simp]: "Suc n < max (n + 3) x"
by arith
lemma mv_box_ok_suc_simp[simp]:
"length xs = n
\<Longrightarrow> {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ x # anything} mv_box n (Suc n)
{\<lambda>nl. nl = xs @ 0 # rec_exec f xs # 0 \<up> (max (n + 3) (max fft gft) - Suc (Suc n)) @ x # anything}"
using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ x # anything"]
apply(simp add: nth_append list_update_append list_update.simps)
apply(cases "max (n + 3) (max fft gft)", simp_all)
apply(cases "max (n + 3) (max fft gft) - 1", simp_all add: Suc_diff_le list_update.simps(2))
done
lemma abc_append_frist_steps_eq_pre:
assumes notfinal: "abc_notfinal (abc_steps_l (0, lm) A n) A"
and notnull: "A \<noteq> []"
shows "abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n"
using notfinal
proof(induct n)
case 0
thus "?case"
by(simp add: abc_steps_l.simps)
next
case (Suc n)
have ind: "abc_notfinal (abc_steps_l (0, lm) A n) A \<Longrightarrow> abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n"
by fact
have h: "abc_notfinal (abc_steps_l (0, lm) A (Suc n)) A" by fact
then have a: "abc_notfinal (abc_steps_l (0, lm) A n) A"
by(simp add: notfinal_Suc)
then have b: "abc_steps_l (0, lm) (A @ B) n = abc_steps_l (0, lm) A n"
using ind by simp
obtain s lm' where c: "abc_steps_l (0, lm) A n = (s, lm')"
by (metis prod.exhaust)
then have d: "s < length A \<and> abc_steps_l (0, lm) (A @ B) n = (s, lm')"
using a b by simp
thus "?case"
using c
by(simp add: abc_step_red2 abc_fetch.simps abc_step_l.simps nth_append)
qed
lemma abc_append_first_step_eq_pre:
"st < length A
\<Longrightarrow> abc_step_l (st, lm) (abc_fetch st (A @ B)) =
abc_step_l (st, lm) (abc_fetch st A)"
by(simp add: abc_step_l.simps abc_fetch.simps nth_append)
lemma abc_append_frist_steps_halt_eq':
assumes final: "abc_steps_l (0, lm) A n = (length A, lm')"
and notnull: "A \<noteq> []"
shows "\<exists> n'. abc_steps_l (0, lm) (A @ B) n' = (length A, lm')"
proof -
have "\<exists> n'. abc_notfinal (abc_steps_l (0, lm) A n') A \<and>
abc_final (abc_steps_l (0, lm) A (Suc n')) A"
using assms
by(rule_tac n = n in abc_before_final, simp_all)
then obtain na where a:
"abc_notfinal (abc_steps_l (0, lm) A na) A \<and>
abc_final (abc_steps_l (0, lm) A (Suc na)) A" ..
obtain sa lma where b: "abc_steps_l (0, lm) A na = (sa, lma)"
by (metis prod.exhaust)
then have c: "abc_steps_l (0, lm) (A @ B) na = (sa, lma)"
using a abc_append_frist_steps_eq_pre[of lm A na B] assms
by simp
have d: "sa < length A" using b a by simp
then have e: "abc_step_l (sa, lma) (abc_fetch sa (A @ B)) =
abc_step_l (sa, lma) (abc_fetch sa A)"
by(rule_tac abc_append_first_step_eq_pre)
from a have "abc_steps_l (0, lm) A (Suc na) = (length A, lm')"
using final equal_when_halt
by(cases "abc_steps_l (0, lm) A (Suc na)" , simp)
then have "abc_steps_l (0, lm) (A @ B) (Suc na) = (length A, lm')"
using a b c e
by(simp add: abc_step_red2)
thus "?thesis"
by blast
qed
lemma abc_append_frist_steps_halt_eq:
assumes final: "abc_steps_l (0, lm) A n = (length A, lm')"
shows "\<exists> n'. abc_steps_l (0, lm) (A @ B) n' = (length A, lm')"
using final
apply(cases "A = []")
apply(rule_tac x = 0 in exI, simp add: abc_steps_l.simps abc_exec_null)
apply(rule_tac abc_append_frist_steps_halt_eq', simp_all)
done
lemma suc_suc_max_simp[simp]: "Suc (Suc (max (xs + 3) fft - Suc (Suc ( xs))))
= max ( xs + 3) fft - ( xs)"
by arith
lemma contract_dec_ft_length_plus_7[simp]: "\<lbrakk>ft = max (n + 3) (max fft gft); length xs = n\<rbrakk> \<Longrightarrow>
{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything}
[Dec ft (length gap + 7)]
{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything}"
apply(simp add: abc_Hoare_halt_def)
apply(rule_tac x = 1 in exI)
apply(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append
abc_lm_v.simps abc_lm_s.simps list_update_append)
using list_update_length
[of "(x - Suc y) # rec_exec (Pr (length xs) f g) (xs @ [x - Suc y]) #
0 \<up> (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" "Suc y" anything y]
apply(simp)
done
lemma adjust_paras':
"length xs = n \<Longrightarrow> {\<lambda>nl. nl = xs @ x # y # anything} [Inc n] [+] [Dec (Suc n) 2, Goto 0]
{\<lambda>nl. nl = xs @ Suc x # 0 # anything}"
proof(rule_tac abc_Hoare_plus_halt)
assume "length xs = n"
thus "{\<lambda>nl. nl = xs @ x # y # anything} [Inc n] {\<lambda> nl. nl = xs @ Suc x # y # anything}"
apply(simp add: abc_Hoare_halt_def)
apply(rule_tac x = 1 in exI, force simp add: abc_steps_l.simps abc_step_l.simps
abc_fetch.simps abc_comp.simps
abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps(2))
done
next
assume h: "length xs = n"
thus "{\<lambda>nl. nl = xs @ Suc x # y # anything} [Dec (Suc n) 2, Goto 0] {\<lambda>nl. nl = xs @ Suc x # 0 # anything}"
proof(induct y)
case 0
thus "?case"
apply(simp add: abc_Hoare_halt_def)
apply(rule_tac x = 1 in exI, simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps
abc_comp.simps
abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps(2))
done
next
case (Suc y)
have "length xs = n \<Longrightarrow>
{\<lambda>nl. nl = xs @ Suc x # y # anything} [Dec (Suc n) 2, Goto 0] {\<lambda>nl. nl = xs @ Suc x # 0 # anything}" by fact
then obtain stp where
"abc_steps_l (0, xs @ Suc x # y # anything) [Dec (Suc n) 2, Goto 0] stp = (2, xs @ Suc x # 0 # anything)"
using h
apply(auto simp: abc_Hoare_halt_def numeral_2_eq_2)
by (metis (mono_tags, lifting) abc_final.simps abc_holds_for.elims(2) length_Cons list.size(3))
moreover have "abc_steps_l (0, xs @ Suc x # Suc y # anything) [Dec (Suc n) 2, Goto 0] 2 =
(0, xs @ Suc x # y # anything)"
using h
by(simp add: abc_steps_l.simps numeral_2_eq_2 abc_step_l.simps abc_fetch.simps
abc_lm_v.simps abc_lm_s.simps nth_append list_update_append list_update.simps(2))
ultimately show "?case"
apply(simp add: abc_Hoare_halt_def)
by(rule exI[of _ "2 + stp"], simp only: abc_steps_add, simp)
qed
qed
lemma adjust_paras:
"length xs = n \<Longrightarrow> {\<lambda>nl. nl = xs @ x # y # anything} [Inc n, Dec (Suc n) 3, Goto (Suc 0)]
{\<lambda>nl. nl = xs @ Suc x # 0 # anything}"
using adjust_paras'[of xs n x y anything]
by(simp add: abc_comp.simps abc_shift.simps numeral_2_eq_2 numeral_3_eq_3)
lemma rec_ci_SucSuc_n[simp]: "\<lbrakk>rec_ci g = (gap, gar, gft); \<forall>y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]);
length xs = n; Suc y\<le>x\<rbrakk> \<Longrightarrow> gar = Suc (Suc n)"
by(auto dest:param_pattern elim!:allE[of _ y])
lemma loop_back':
assumes h: "length A = length gap + 4" "length xs = n"
and le: "y \<ge> x"
shows "\<exists> stp. abc_steps_l (length A, xs @ m # (y - x) # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp
= (length A, xs @ m # y # 0 # anything)"
using le
proof(induct x)
case 0
thus "?case"
using h
by(rule_tac x = 0 in exI,
auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_s.simps abc_lm_v.simps)
next
case (Suc x)
have "x \<le> y \<Longrightarrow> \<exists>stp. abc_steps_l (length A, xs @ m # (y - x) # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp =
(length A, xs @ m # y # 0 # anything)" by fact
moreover have "Suc x \<le> y" by fact
moreover then have "\<exists> stp. abc_steps_l (length A, xs @ m # (y - Suc x) # Suc x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp
= (length A, xs @ m # (y - x) # x # anything)"
using h
apply(rule_tac x = 3 in exI)
by(simp add: abc_steps_l.simps numeral_3_eq_3 abc_step_l.simps abc_fetch.simps nth_append
abc_lm_v.simps abc_lm_s.simps list_update_append list_update.simps(2))
ultimately show "?case"
apply(auto simp add: abc_steps_add)
by (metis abc_steps_add)
qed
lemma loop_back:
assumes h: "length A = length gap + 4" "length xs = n"
shows "\<exists> stp. abc_steps_l (length A, xs @ m # 0 # x # anything) (A @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]) stp
= (0, xs @ m # x # 0 # anything)"
using loop_back'[of A gap xs n x x m anything] assms
apply(auto) apply(rename_tac stp)
apply(rule_tac x = "stp + 1" in exI)
apply(simp only: abc_steps_add, simp)
apply(simp add: abc_steps_l.simps abc_step_l.simps abc_fetch.simps nth_append abc_lm_v.simps
abc_lm_s.simps)
done
lemma rec_exec_pr_0_simps: "rec_exec (Pr n f g) (xs @ [0]) = rec_exec f xs"
by(simp add: rec_exec.simps)
lemma rec_exec_pr_Suc_simps: "rec_exec (Pr n f g) (xs @ [Suc y])
= rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])"
apply(induct y)
apply(simp add: rec_exec.simps)
apply(simp add: rec_exec.simps)
done
lemma suc_max_simp[simp]: "Suc (max (n + 3) fft - Suc (Suc (Suc n))) = max (n + 3) fft - Suc (Suc n)"
by arith
lemma pr_loop:
assumes code: "code = ([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @
[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]"
and len: "length xs = n"
and g_ind: "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap
{\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})"
and compile_g: "rec_ci g = (gap, gar, gft)"
and termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])"
and ft: "ft = max (n + 3) (max fft gft)"
and less: "Suc y \<le> x"
shows
"\<exists>stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything)
code stp = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything)"
proof -
let ?A = "[Dec ft (length gap + 7)]"
let ?B = "gap"
let ?C = "[Inc n, Dec (Suc n) 3, Goto (Suc 0)]"
let ?D = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]"
have "\<exists> stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything)
((?A [+] (?B [+] ?C)) @ ?D) stp = (length (?A [+] (?B [+] ?C)),
xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])])
# 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)"
proof -
have "\<exists> stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything)
((?A [+] (?B [+] ?C))) stp = (length (?A [+] (?B [+] ?C)), xs @ (x - y) # 0 #
rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)"
proof -
have "{\<lambda> nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything}
(?A [+] (?B [+] ?C))
{\<lambda> nl. nl = xs @ (x - y) # 0 #
rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}"
proof(rule_tac abc_Hoare_plus_halt)
show "{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything}
[Dec ft (length gap + 7)]
{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything}"
using ft len
by(simp)
next
show
"{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything}
?B [+] ?C
{\<lambda>nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}"
proof(rule_tac abc_Hoare_plus_halt)
have a: "gar = Suc (Suc n)"
using compile_g termi_g len less
by simp
have b: "gft > gar"
using compile_g
by(erule_tac footprint_ge)
show "{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ y # anything} gap
{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) #
rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}"
proof -
have
"{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (gft - gar) @ 0\<up>(ft - gft) @ y # anything} gap
{\<lambda>nl. nl = xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) #
rec_exec g (xs @ [(x - Suc y), rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (gft - Suc gar) @ 0\<up>(ft - gft) @ y # anything}"
using g_ind less by simp
thus "?thesis"
using a b ft
by(simp add: replicate_merge_anywhere numeral_3_eq_3)
qed
next
show "{\<lambda>nl. nl = xs @ (x - Suc y) #
rec_exec (Pr n f g) (xs @ [x - Suc y]) #
rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}
[Inc n, Dec (Suc n) 3, Goto (Suc 0)]
{\<lambda>nl. nl = xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g)
(xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything}"
using len less
using adjust_paras[of xs n "x - Suc y" " rec_exec (Pr n f g) (xs @ [x - Suc y])"
" rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) #
0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything"]
by(simp add: Suc_diff_Suc)
qed
qed
thus "?thesis"
apply(simp add: abc_Hoare_halt_def, auto)
apply(rename_tac na)
apply(rule_tac x = na in exI, case_tac "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) #
0 \<up> (ft - Suc (Suc n)) @ Suc y # anything)
([Dec ft (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) na", simp)
done
qed
then obtain stpa where "abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (ft - Suc (Suc n)) @ Suc y # anything)
((?A [+] (?B [+] ?C))) stpa = (length (?A [+] (?B [+] ?C)),
xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])])
# 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)" ..
thus "?thesis"
by(erule_tac abc_append_frist_steps_halt_eq)
qed
moreover have
"\<exists> stp. abc_steps_l (length (?A [+] (?B [+] ?C)),
xs @ (x - y) # 0 # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)
((?A [+] (?B [+] ?C)) @ ?D) stp = (0, xs @ (x - y) # rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) #
0 # 0 \<up> (ft - Suc (Suc (Suc n))) @ y # anything)"
using len
by(rule_tac loop_back, simp_all)
moreover have "rec_exec g (xs @ [x - Suc y, rec_exec (Pr n f g) (xs @ [x - Suc y])]) = rec_exec (Pr n f g) (xs @ [x - y])"
using less
apply(cases "x - y", simp_all add: rec_exec_pr_Suc_simps)
apply(rename_tac nat)
by(subgoal_tac "nat = x - Suc y", simp, arith)
ultimately show "?thesis"
using code ft
apply (auto simp add: abc_steps_add replicate_Suc_iff_anywhere)
apply(rename_tac stp stpa)
apply(rule_tac x = "stp + stpa" in exI)
by (simp add: abc_steps_add replicate_Suc_iff_anywhere del: replicate_Suc)
qed
lemma abc_lm_s_simp0[simp]:
"length xs = n \<Longrightarrow> abc_lm_s (xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3)
(max fft gft) - Suc (Suc n)) @ 0 # anything) (max (n + 3) (max fft gft)) 0 =
xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything"
apply(simp add: abc_lm_s.simps)
using list_update_length[of "xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc (Suc n))"
0 anything 0]
apply(auto simp: Suc_diff_Suc)
apply(simp add: exp_suc[THEN sym] Suc_diff_Suc del: replicate_Suc)
done
lemma index_at_zero_elem[simp]:
"(xs @ x # re # 0 \<up> (max (length xs + 3)
(max fft gft) - Suc (Suc (length xs))) @ 0 # anything) !
max (length xs + 3) (max fft gft) = 0"
using nth_append_length[of "xs @ x # re #
0 \<up> (max (length xs + 3) (max fft gft) - Suc (Suc (length xs)))" 0 anything]
by(simp)
lemma pr_loop_correct:
assumes less: "y \<le> x"
and len: "length xs = n"
and compile_g: "rec_ci g = (gap, gar, gft)"
and termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])"
and g_ind: "\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap
{\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})"
shows "{\<lambda>nl. nl = xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (max (n + 3) (max fft gft) - Suc (Suc n)) @ y # anything}
([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)])) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]
{\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything}"
using less
proof(induct y)
case 0
thus "?case"
using len
apply(simp add: abc_Hoare_halt_def)
apply(rule_tac x = 1 in exI)
by(auto simp: abc_steps_l.simps abc_step_l.simps abc_fetch.simps
nth_append abc_comp.simps abc_shift.simps, simp add: abc_lm_v.simps)
next
case (Suc y)
let ?ft = "max (n + 3) (max fft gft)"
let ?C = "[Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+]
[Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @ [Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]"
have ind: "y \<le> x \<Longrightarrow>
{\<lambda>nl. nl = xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything}
?C {\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything}" by fact
have less: "Suc y \<le> x" by fact
have stp1:
"\<exists> stp. abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (?ft - Suc (Suc n)) @ Suc y # anything)
?C stp = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)"
using assms less
by(rule_tac pr_loop, auto)
then obtain stp1 where a:
"abc_steps_l (0, xs @ (x - Suc y) # rec_exec (Pr n f g) (xs @ [x - Suc y]) # 0 \<up> (?ft - Suc (Suc n)) @ Suc y # anything)
?C stp1 = (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)" ..
moreover have
"\<exists> stp. abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)
?C stp = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything)"
using ind less
apply(auto simp: abc_Hoare_halt_def)
apply(rename_tac na,case_tac "abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g)
(xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything) ?C na", rule_tac x = na in exI)
by simp
then obtain stp2 where b:
"abc_steps_l (0, xs @ (x - y) # rec_exec (Pr n f g) (xs @ [x - y]) # 0 \<up> (?ft - Suc (Suc n)) @ y # anything)
?C stp2 = (length ?C, xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything)" ..
from a b show "?case"
apply(simp add: abc_Hoare_halt_def)
apply(rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add).
qed
lemma compile_pr_correct':
assumes termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])"
and g_ind:
"\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap
{\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})"
and termi_f: "terminate f xs"
and f_ind: "\<And> anything. {\<lambda>nl. nl = xs @ 0 \<up> (fft - far) @ anything} fap {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (fft - Suc far) @ anything}"
and len: "length xs = n"
and compile1: "rec_ci f = (fap, far, fft)"
and compile2: "rec_ci g = (gap, gar, gft)"
shows
"{\<lambda>nl. nl = xs @ x # 0 \<up> (max (n + 3) (max fft gft) - n) @ anything}
mv_box n (max (n + 3) (max fft gft)) [+]
(fap [+] (mv_box n (Suc n) [+]
([Dec (max (n + 3) (max fft gft)) (length gap + 7)] [+] (gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]) @
[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)])))
{\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ anything}"
proof -
let ?ft = "max (n+3) (max fft gft)"
let ?A = "mv_box n ?ft"
let ?B = "fap"
let ?C = "mv_box n (Suc n)"
let ?D = "[Dec ?ft (length gap + 7)]"
let ?E = "gap [+] [Inc n, Dec (Suc n) 3, Goto (Suc 0)]"
let ?F = "[Dec (Suc (Suc n)) 0, Inc (Suc n), Goto (length gap + 4)]"
let ?P = "\<lambda>nl. nl = xs @ x # 0 \<up> (?ft - n) @ anything"
let ?S = "\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (?ft - Suc n) @ anything"
let ?Q1 = "\<lambda>nl. nl = xs @ 0 \<up> (?ft - n) @ x # anything"
show "{?P} (?A [+] (?B [+] (?C [+] (?D [+] ?E @ ?F)))) {?S}"
proof(rule_tac abc_Hoare_plus_halt)
show "{?P} ?A {?Q1}"
using len by simp
next
let ?Q2 = "\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (?ft - Suc n) @ x # anything"
have a: "?ft \<ge> fft"
by arith
have b: "far = n"
using compile1 termi_f len
by(drule_tac param_pattern, auto)
have c: "fft > far"
using compile1
by(simp add: footprint_ge)
show "{?Q1} (?B [+] (?C [+] (?D [+] ?E @ ?F))) {?S}"
proof(rule_tac abc_Hoare_plus_halt)
have "{\<lambda>nl. nl = xs @ 0 \<up> (fft - far) @ 0\<up>(?ft - fft) @ x # anything} fap
{\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (fft - Suc far) @ 0\<up>(?ft - fft) @ x # anything}"
by(rule_tac f_ind)
moreover have "fft - far + ?ft - fft = ?ft - far"
using a b c by arith
moreover have "fft - Suc n + ?ft - fft = ?ft - Suc n"
using a b c by arith
ultimately show "{?Q1} ?B {?Q2}"
using b
by(simp add: replicate_merge_anywhere)
next
let ?Q3 = "\<lambda> nl. nl = xs @ 0 # rec_exec f xs # 0\<up>(?ft - Suc (Suc n)) @ x # anything"
show "{?Q2} (?C [+] (?D [+] ?E @ ?F)) {?S}"
proof(rule_tac abc_Hoare_plus_halt)
show "{?Q2} (?C) {?Q3}"
using mv_box_correct[of n "Suc n" "xs @ rec_exec f xs # 0 \<up> (max (n + 3) (max fft gft) - Suc n) @ x # anything"]
using len
by(auto)
next
show "{?Q3} (?D [+] ?E @ ?F) {?S}"
using pr_loop_correct[of x x xs n g gap gar gft f fft anything] assms
by(simp add: rec_exec_pr_0_simps)
qed
qed
qed
qed
lemma compile_pr_correct:
assumes g_ind: "\<forall>y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) \<and>
(\<forall>x xa xb. rec_ci g = (x, xa, xb) \<longrightarrow>
(\<forall>xc. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (xb - xa) @ xc} x
{\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (xb - Suc xa) @ xc}))"
and termi_f: "terminate f xs"
and f_ind:
"\<And>ap arity fp anything.
rec_ci f = (ap, arity, fp) \<Longrightarrow> {\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (fp - Suc arity) @ anything}"
and len: "length xs = n"
and compile: "rec_ci (Pr n f g) = (ap, arity, fp)"
shows "{\<lambda>nl. nl = xs @ x # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ x # rec_exec (Pr n f g) (xs @ [x]) # 0 \<up> (fp - Suc arity) @ anything}"
proof(cases "rec_ci f", cases "rec_ci g")
fix fap far fft gap gar gft
assume h: "rec_ci f = (fap, far, fft)" "rec_ci g = (gap, gar, gft)"
have g:
"\<forall>y<x. (terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) \<and>
(\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap
{\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything}))"
using g_ind h
by(auto)
hence termi_g: "\<forall> y<x. terminate g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])])"
by simp
from g have g_newind:
"\<forall> y<x. (\<forall>anything. {\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # 0 \<up> (gft - gar) @ anything} gap
{\<lambda>nl. nl = xs @ y # rec_exec (Pr n f g) (xs @ [y]) # rec_exec g (xs @ [y, rec_exec (Pr n f g) (xs @ [y])]) # 0 \<up> (gft - Suc gar) @ anything})"
by auto
have f_newind: "\<And> anything. {\<lambda>nl. nl = xs @ 0 \<up> (fft - far) @ anything} fap {\<lambda>nl. nl = xs @ rec_exec f xs # 0 \<up> (fft - Suc far) @ anything}"
using h
by(rule_tac f_ind, simp)
show "?thesis"
using termi_f termi_g h compile
apply(simp add: rec_ci.simps abc_comp_commute, auto)
using g_newind f_newind len
by(rule_tac compile_pr_correct', simp_all)
qed
fun mn_ind_inv ::
"nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> nat list \<Rightarrow> bool"
where
"mn_ind_inv (as, lm') ss x rsx suf_lm lm =
(if as = ss then lm' = lm @ x # rsx # suf_lm
else if as = ss + 1 then
\<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
else if as = ss + 2 then
\<exists>y. (lm' = lm @ x # y # suf_lm) \<and> y \<le> rsx
else if as = ss + 3 then lm' = lm @ x # 0 # suf_lm
else if as = ss + 4 then lm' = lm @ Suc x # 0 # suf_lm
else if as = 0 then lm' = lm @ Suc x # 0 # suf_lm
else False
)"
fun mn_stage1 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"mn_stage1 (as, lm) ss n =
(if as = 0 then 0
else if as = ss + 4 then 1
else if as = ss + 3 then 2
else if as = ss + 2 \<or> as = ss + 1 then 3
else if as = ss then 4
else 0
)"
fun mn_stage2 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"mn_stage2 (as, lm) ss n =
(if as = ss + 1 \<or> as = ss + 2 then (lm ! (Suc n))
else 0)"
fun mn_stage3 :: "nat \<times> nat list \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat"
where
"mn_stage3 (as, lm) ss n = (if as = ss + 2 then 1 else 0)"
fun mn_measure :: "((nat \<times> nat list) \<times> nat \<times> nat) \<Rightarrow>
(nat \<times> nat \<times> nat)"
where
"mn_measure ((as, lm), ss, n) =
(mn_stage1 (as, lm) ss n, mn_stage2 (as, lm) ss n,
mn_stage3 (as, lm) ss n)"
definition mn_LE :: "(((nat \<times> nat list) \<times> nat \<times> nat) \<times>
((nat \<times> nat list) \<times> nat \<times> nat)) set"
where "mn_LE \<equiv> (inv_image lex_triple mn_measure)"
lemma wf_mn_le[intro]: "wf mn_LE"
by(auto intro:wf_inv_image wf_lex_triple simp: mn_LE_def)
declare mn_ind_inv.simps[simp del]
lemma put_in_tape_small_enough0[simp]:
"0 < rsx \<Longrightarrow>
\<exists>y. (xs @ x # rsx # anything)[Suc (length xs) := rsx - Suc 0] = xs @ x # y # anything \<and> y \<le> rsx"
apply(rule_tac x = "rsx - 1" in exI)
apply(simp add: list_update_append list_update.simps)
done
lemma put_in_tape_small_enough1[simp]:
"\<lbrakk>y \<le> rsx; 0 < y\<rbrakk>
\<Longrightarrow> \<exists>ya. (xs @ x # y # anything)[Suc (length xs) := y - Suc 0] = xs @ x # ya # anything \<and> ya \<le> rsx"
apply(rule_tac x = "y - 1" in exI)
apply(simp add: list_update_append list_update.simps)
done
lemma abc_comp_null[simp]: "(A [+] B = []) = (A = [] \<and> B = [])"
by(auto simp: abc_comp.simps abc_shift.simps)
lemma rec_ci_not_null[simp]: "(rec_ci f \<noteq> ([], a, b))"
proof(cases f)
case (Cn x41 x42 x43)
then show ?thesis
by(cases "rec_ci x42", auto simp: mv_box.simps rec_ci.simps rec_ci_id.simps)
next
case (Pr x51 x52 x53)
then show ?thesis
apply(cases "rec_ci x52", cases "rec_ci x53")
by (auto simp: mv_box.simps rec_ci.simps rec_ci_id.simps)
next
case (Mn x61 x62)
then show ?thesis
by(cases "rec_ci x62") (auto simp: rec_ci.simps rec_ci_id.simps)
qed (auto simp: rec_ci_z_def rec_ci_s_def rec_ci.simps addition.simps rec_ci_id.simps)
lemma mn_correct:
assumes compile: "rec_ci rf = (fap, far, fft)"
and ge: "0 < rsx"
and len: "length xs = arity"
and B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]"
and f: "f = (\<lambda> stp. (abc_steps_l (length fap, xs @ x # rsx # anything) (fap @ B) stp, (length fap), arity)) "
and P: "P =(\<lambda> ((as, lm), ss, arity). as = 0)"
and Q: "Q = (\<lambda> ((as, lm), ss, arity). mn_ind_inv (as, lm) (length fap) x rsx anything xs)"
shows "\<exists> stp. P (f stp) \<and> Q (f stp)"
proof(rule_tac halt_lemma2)
show "wf mn_LE"
using wf_mn_le by simp
next
show "Q (f 0)"
by(auto simp: Q f abc_steps_l.simps mn_ind_inv.simps)
next
have "fap \<noteq> []"
using compile by auto
thus "\<not> P (f 0)"
by(auto simp: f P abc_steps_l.simps)
next
have "fap \<noteq> []"
using compile by auto
then have "\<lbrakk>\<not> P (f stp); Q (f stp)\<rbrakk> \<Longrightarrow> Q (f (Suc stp)) \<and> (f (Suc stp), f stp) \<in> mn_LE" for stp
using ge len
apply(cases "(abc_steps_l (length fap, xs @ x # rsx # anything) (fap @ B) stp)")
apply(simp add: abc_step_red2 B f P Q)
apply(auto split:if_splits simp add:abc_steps_l.simps mn_ind_inv.simps abc_steps_zero B abc_fetch.simps nth_append)
by(auto simp: mn_LE_def lex_triple_def lex_pair_def
abc_step_l.simps abc_steps_l.simps mn_ind_inv.simps
abc_lm_v.simps abc_lm_s.simps nth_append abc_fetch.simps
split: if_splits)
thus "\<forall>stp. \<not> P (f stp) \<and> Q (f stp) \<longrightarrow> Q (f (Suc stp)) \<and> (f (Suc stp), f stp) \<in> mn_LE"
by(auto)
qed
lemma abc_Hoare_haltE:
"{\<lambda> nl. nl = lm1} p {\<lambda> nl. nl = lm2}
\<Longrightarrow> \<exists> stp. abc_steps_l (0, lm1) p stp = (length p, lm2)"
by(auto simp:abc_Hoare_halt_def elim!: abc_holds_for.elims)
lemma mn_loop:
assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]"
and ft: "ft = max (Suc arity) fft"
and len: "length xs = arity"
and far: "far = Suc arity"
and ind: " (\<forall>xc. ({\<lambda>nl. nl = xs @ x # 0 \<up> (fft - far) @ xc} fap
{\<lambda>nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (fft - Suc far) @ xc}))"
and exec_less: "rec_exec f (xs @ [x]) > 0"
and compile: "rec_ci f = (fap, far, fft)"
shows "\<exists> stp > 0. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
(0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)"
proof -
have "\<exists> stp. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
(length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
proof -
have "\<exists> stp. abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) fap stp =
(length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
proof -
have "{\<lambda>nl. nl = xs @ x # 0 \<up> (fft - far) @ 0\<up>(ft - fft) @ anything} fap
{\<lambda>nl. nl = xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (fft - Suc far) @ 0\<up>(ft - fft) @ anything}"
using ind by simp
moreover have "fft > far"
using compile
by(erule_tac footprint_ge)
ultimately show "?thesis"
using ft far
apply(drule_tac abc_Hoare_haltE)
by(simp add: replicate_merge_anywhere max_absorb2)
qed
then obtain stp where "abc_steps_l (0, xs @ x # 0 \<up> (ft - Suc arity) @ anything) fap stp =
(length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" ..
thus "?thesis"
by(erule_tac abc_append_frist_steps_halt_eq)
qed
moreover have
"\<exists> stp > 0. abc_steps_l (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything) (fap @ B) stp =
(0, xs @ Suc x # 0 # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
using mn_correct[of f fap far fft "rec_exec f (xs @ [x])" xs arity B
"(\<lambda>stp. (abc_steps_l (length fap, xs @ x # rec_exec f (xs @ [x]) # 0 \<up> (ft - Suc (Suc arity)) @ anything) (fap @ B) stp, length fap, arity))"
x "0 \<up> (ft - Suc (Suc arity)) @ anything" "(\<lambda>((as, lm), ss, arity). as = 0)"
"(\<lambda>((as, lm), ss, aritya). mn_ind_inv (as, lm) (length fap) x (rec_exec f (xs @ [x])) (0 \<up> (ft - Suc (Suc arity)) @ anything) xs) "]
B compile exec_less len
apply(subgoal_tac "fap \<noteq> []", auto)
apply(rename_tac stp y)
apply(rule_tac x = stp in exI, auto simp: mn_ind_inv.simps)
by(case_tac "stp", simp_all add: abc_steps_l.simps)
moreover have "fft > far"
using compile
by(erule_tac footprint_ge)
ultimately show "?thesis"
using ft far
apply(auto) apply(rename_tac stp1 stp2)
by(rule_tac x = "stp1 + stp2" in exI,
simp add: abc_steps_add replicate_Suc[THEN sym] diff_Suc_Suc del: replicate_Suc)
qed
lemma mn_loop_correct':
assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]"
and ft: "ft = max (Suc arity) fft"
and len: "length xs = arity"
and ind_all: "\<forall>i\<le>x. (\<forall>xc. ({\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap
{\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}))"
and exec_ge: "\<forall> i\<le>x. rec_exec f (xs @ [i]) > 0"
and compile: "rec_ci f = (fap, far, fft)"
and far: "far = Suc arity"
shows "\<exists> stp > x. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
(0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)"
using ind_all exec_ge
proof(induct x)
case 0
thus "?case"
using assms
by(rule_tac mn_loop, simp_all)
next
case (Suc x)
have ind': "\<lbrakk>\<forall>i\<le>x. \<forall>xc. {\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap {\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc};
\<forall>i\<le>x. 0 < rec_exec f (xs @ [i])\<rbrakk> \<Longrightarrow>
\<exists>stp > x. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)" by fact
have exec_ge: "\<forall>i\<le>Suc x. 0 < rec_exec f (xs @ [i])" by fact
have ind_all: "\<forall>i\<le>Suc x. \<forall>xc. {\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap
{\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}" by fact
have ind: "\<exists>stp > x. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
(0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)" using ind' exec_ge ind_all by simp
have stp: "\<exists> stp > 0. abc_steps_l (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
(0, xs @ Suc (Suc x) # 0 \<up> (ft - Suc arity) @ anything)"
using ind_all exec_ge B ft len far compile
by(rule_tac mn_loop, simp_all)
from ind stp show "?case"
apply(auto simp add: abc_steps_add)
apply(rename_tac stp1 stp2)
by(rule_tac x = "stp1 + stp2" in exI, simp add: abc_steps_add)
qed
lemma mn_loop_correct:
assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]"
and ft: "ft = max (Suc arity) fft"
and len: "length xs = arity"
and ind_all: "\<forall>i\<le>x. (\<forall>xc. ({\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap
{\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}))"
and exec_ge: "\<forall> i\<le>x. rec_exec f (xs @ [i]) > 0"
and compile: "rec_ci f = (fap, far, fft)"
and far: "far = Suc arity"
shows "\<exists> stp. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
(0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)"
proof -
have "\<exists>stp>x. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp = (0, xs @ Suc x # 0 \<up> (ft - Suc arity) @ anything)"
using assms
by(rule_tac mn_loop_correct', simp_all)
thus "?thesis"
by(auto)
qed
lemma compile_mn_correct':
assumes B: "B = [Dec (Suc arity) (length fap + 5), Dec (Suc arity) (length fap + 3), Goto (Suc (length fap)), Inc arity, Goto 0]"
and ft: "ft = max (Suc arity) fft"
and len: "length xs = arity"
and termi_f: "terminate f (xs @ [r])"
and f_ind: "\<And>anything. {\<lambda>nl. nl = xs @ r # 0 \<up> (fft - far) @ anything} fap
{\<lambda>nl. nl = xs @ r # 0 # 0 \<up> (fft - Suc far) @ anything}"
and ind_all: "\<forall>i < r. (\<forall>xc. ({\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap
{\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}))"
and exec_less: "\<forall> i<r. rec_exec f (xs @ [i]) > 0"
and exec: "rec_exec f (xs @ [r]) = 0"
and compile: "rec_ci f = (fap, far, fft)"
shows "{\<lambda>nl. nl = xs @ 0 \<up> (max (Suc arity) fft - arity) @ anything}
fap @ B
{\<lambda>nl. nl = xs @ rec_exec (Mn arity f) xs # 0 \<up> (max (Suc arity) fft - Suc arity) @ anything}"
proof -
have a: "far = Suc arity"
using len compile termi_f
by(drule_tac param_pattern, auto)
have b: "fft > far"
using compile
by(erule_tac footprint_ge)
have "\<exists> stp. abc_steps_l (0, xs @ 0 # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
(0, xs @ r # 0 \<up> (ft - Suc arity) @ anything)"
using assms a
apply(cases r, rule_tac x = 0 in exI, simp add: abc_steps_l.simps, simp)
by(rule_tac mn_loop_correct, auto)
moreover have
"\<exists> stp. abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) (fap @ B) stp =
(length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
proof -
have "\<exists> stp. abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) fap stp =
(length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
proof -
have "{\<lambda>nl. nl = xs @ r # 0 \<up> (fft - far) @ 0\<up>(ft - fft) @ anything} fap
{\<lambda>nl. nl = xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (fft - Suc far) @ 0\<up>(ft - fft) @ anything}"
using f_ind exec by simp
thus "?thesis"
using ft a b
apply(drule_tac abc_Hoare_haltE)
by(simp add: replicate_merge_anywhere max_absorb2)
qed
then obtain stp where "abc_steps_l (0, xs @ r # 0 \<up> (ft - Suc arity) @ anything) fap stp =
(length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)" ..
thus "?thesis"
by(erule_tac abc_append_frist_steps_halt_eq)
qed
moreover have
"\<exists> stp. abc_steps_l (length fap, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything) (fap @ B) stp =
(length fap + 5, xs @ r # rec_exec f (xs @ [r]) # 0 \<up> (ft - Suc (Suc arity)) @ anything)"
using ft a b len B exec
apply(rule_tac x = 1 in exI, auto)
by(auto simp: abc_steps_l.simps B abc_step_l.simps
abc_fetch.simps nth_append max_absorb2 abc_lm_v.simps abc_lm_s.simps)
moreover have "rec_exec (Mn (length xs) f) xs = r"
using exec exec_less
apply(auto simp: rec_exec.simps Least_def)
thm the_equality
apply(rule_tac the_equality, auto)
apply(metis exec_less less_not_refl3 linorder_not_less)
by (metis le_neq_implies_less less_not_refl3)
ultimately show "?thesis"
using ft a b len B exec
apply(auto simp: abc_Hoare_halt_def)
apply(rename_tac stp1 stp2 stp3)
apply(rule_tac x = "stp1 + stp2 + stp3" in exI)
by(simp add: abc_steps_add replicate_Suc_iff_anywhere max_absorb2 Suc_diff_Suc del: replicate_Suc)
qed
lemma compile_mn_correct:
assumes len: "length xs = n"
and termi_f: "terminate f (xs @ [r])"
and f_ind: "\<And>ap arity fp anything. rec_ci f = (ap, arity, fp) \<Longrightarrow>
{\<lambda>nl. nl = xs @ r # 0 \<up> (fp - arity) @ anything} ap {\<lambda>nl. nl = xs @ r # 0 # 0 \<up> (fp - Suc arity) @ anything}"
and exec: "rec_exec f (xs @ [r]) = 0"
and ind_all:
"\<forall>i<r. terminate f (xs @ [i]) \<and>
(\<forall>x xa xb. rec_ci f = (x, xa, xb) \<longrightarrow>
(\<forall>xc. {\<lambda>nl. nl = xs @ i # 0 \<up> (xb - xa) @ xc} x {\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (xb - Suc xa) @ xc})) \<and>
0 < rec_exec f (xs @ [i])"
and compile: "rec_ci (Mn n f) = (ap, arity, fp)"
shows "{\<lambda>nl. nl = xs @ 0 \<up> (fp - arity) @ anything} ap
{\<lambda>nl. nl = xs @ rec_exec (Mn n f) xs # 0 \<up> (fp - Suc arity) @ anything}"
proof(cases "rec_ci f")
fix fap far fft
assume h: "rec_ci f = (fap, far, fft)"
hence f_newind:
"\<And>anything. {\<lambda>nl. nl = xs @ r # 0 \<up> (fft - far) @ anything} fap
{\<lambda>nl. nl = xs @ r # 0 # 0 \<up> (fft - Suc far) @ anything}"
by(rule_tac f_ind, simp)
have newind_all:
"\<forall>i < r. (\<forall>xc. ({\<lambda>nl. nl = xs @ i # 0 \<up> (fft - far) @ xc} fap
{\<lambda>nl. nl = xs @ i # rec_exec f (xs @ [i]) # 0 \<up> (fft - Suc far) @ xc}))"
using ind_all h
by(auto)
have all_less: "\<forall> i<r. rec_exec f (xs @ [i]) > 0"
using ind_all by auto
show "?thesis"
using h compile f_newind newind_all all_less len termi_f exec
apply(auto simp: rec_ci.simps)
by(rule_tac compile_mn_correct', auto)
qed
lemma recursive_compile_correct:
"\<lbrakk>terminate recf args; rec_ci recf = (ap, arity, fp)\<rbrakk>
\<Longrightarrow> {\<lambda> nl. nl = args @ 0\<up>(fp - arity) @ anything} ap
{\<lambda> nl. nl = args@ rec_exec recf args # 0\<up>(fp - Suc arity) @ anything}"
apply(induct arbitrary: ap arity fp anything rule: terminate.induct)
apply(simp_all add: compile_s_correct compile_z_correct compile_id_correct
compile_cn_correct compile_pr_correct compile_mn_correct)
done
definition dummy_abc :: "nat \<Rightarrow> abc_inst list"
where
"dummy_abc k = [Inc k, Dec k 0, Goto 3]"
definition abc_list_crsp:: "nat list \<Rightarrow> nat list \<Rightarrow> bool"
where
"abc_list_crsp xs ys = (\<exists> n. xs = ys @ 0\<up>n \<or> ys = xs @ 0\<up>n)"
lemma abc_list_crsp_simp1[intro]: "abc_list_crsp (lm @ 0\<up>m) lm"
by(auto simp: abc_list_crsp_def)
lemma abc_list_crsp_lm_v:
"abc_list_crsp lma lmb \<Longrightarrow> abc_lm_v lma n = abc_lm_v lmb n"
by(auto simp: abc_list_crsp_def abc_lm_v.simps
nth_append)
lemma abc_list_crsp_elim:
"\<lbrakk>abc_list_crsp lma lmb; \<exists> n. lma = lmb @ 0\<up>n \<or> lmb = lma @ 0\<up>n \<Longrightarrow> P \<rbrakk> \<Longrightarrow> P"
by(auto simp: abc_list_crsp_def)
lemma abc_list_crsp_simp[simp]:
"\<lbrakk>abc_list_crsp lma lmb; m < length lma; m < length lmb\<rbrakk> \<Longrightarrow>
abc_list_crsp (lma[m := n]) (lmb[m := n])"
by(auto simp: abc_list_crsp_def list_update_append)
lemma abc_list_crsp_simp2[simp]:
"\<lbrakk>abc_list_crsp lma lmb; m < length lma; \<not> m < length lmb\<rbrakk> \<Longrightarrow>
abc_list_crsp (lma[m := n]) (lmb @ 0 \<up> (m - length lmb) @ [n])"
apply(auto simp: abc_list_crsp_def list_update_append)
apply(rename_tac N)
apply(rule_tac x = "N + length lmb - Suc m" in exI)
apply(rule_tac disjI1)
apply(simp add: upd_conv_take_nth_drop min_absorb1)
done
lemma abc_list_crsp_simp3[simp]:
"\<lbrakk>abc_list_crsp lma lmb; \<not> m < length lma; m < length lmb\<rbrakk> \<Longrightarrow>
abc_list_crsp (lma @ 0 \<up> (m - length lma) @ [n]) (lmb[m := n])"
apply(auto simp: abc_list_crsp_def list_update_append)
apply(rename_tac N)
apply(rule_tac x = "N + length lma - Suc m" in exI)
apply(rule_tac disjI2)
apply(simp add: upd_conv_take_nth_drop min_absorb1)
done
lemma abc_list_crsp_simp4[simp]: "\<lbrakk>abc_list_crsp lma lmb; \<not> m < length lma; \<not> m < length lmb\<rbrakk> \<Longrightarrow>
abc_list_crsp (lma @ 0 \<up> (m - length lma) @ [n]) (lmb @ 0 \<up> (m - length lmb) @ [n])"
by(auto simp: abc_list_crsp_def list_update_append replicate_merge_anywhere)
lemma abc_list_crsp_lm_s:
"abc_list_crsp lma lmb \<Longrightarrow>
abc_list_crsp (abc_lm_s lma m n) (abc_lm_s lmb m n)"
by(auto simp: abc_lm_s.simps)
lemma abc_list_crsp_step:
"\<lbrakk>abc_list_crsp lma lmb; abc_step_l (aa, lma) i = (a, lma');
abc_step_l (aa, lmb) i = (a', lmb')\<rbrakk>
\<Longrightarrow> a' = a \<and> abc_list_crsp lma' lmb'"
apply(cases i, auto simp: abc_step_l.simps
abc_list_crsp_lm_s abc_list_crsp_lm_v
split: abc_inst.splits if_splits)
done
lemma abc_list_crsp_steps:
"\<lbrakk>abc_steps_l (0, lm @ 0\<up>m) aprog stp = (a, lm'); aprog \<noteq> []\<rbrakk>
\<Longrightarrow> \<exists> lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and>
abc_list_crsp lm' lma"
proof(induct stp arbitrary: a lm')
case (Suc stp)
- then show ?case apply(cases "abc_steps_l (0, lm @ 0\<up>m) aprog stp", simp add: abc_step_red)
+ then show ?case using [[simproc del: defined_all]] apply(cases "abc_steps_l (0, lm @ 0\<up>m) aprog stp", simp add: abc_step_red)
proof -
fix stp a lm' aa b
assume ind:
"\<And>a lm'. aa = a \<and> b = lm' \<Longrightarrow>
\<exists>lma. abc_steps_l (0, lm) aprog stp = (a, lma) \<and>
abc_list_crsp lm' lma"
and h: "abc_step_l (aa, b) (abc_fetch aa aprog) = (a, lm')"
"abc_steps_l (0, lm @ 0\<up>m) aprog stp = (aa, b)"
"aprog \<noteq> []"
have "\<exists>lma. abc_steps_l (0, lm) aprog stp = (aa, lma) \<and>
abc_list_crsp b lma"
apply(rule_tac ind, simp)
done
from this obtain lma where g2:
"abc_steps_l (0, lm) aprog stp = (aa, lma) \<and>
abc_list_crsp b lma" ..
hence g3: "abc_steps_l (0, lm) aprog (Suc stp)
= abc_step_l (aa, lma) (abc_fetch aa aprog)"
apply(rule_tac abc_step_red, simp)
done
show "\<exists>lma. abc_steps_l (0, lm) aprog (Suc stp) = (a, lma) \<and> abc_list_crsp lm' lma"
using g2 g3 h
apply(auto)
apply(cases "abc_step_l (aa, b) (abc_fetch aa aprog)",
cases "abc_step_l (aa, lma) (abc_fetch aa aprog)", simp)
apply(rule_tac abc_list_crsp_step, auto)
done
qed
qed (force simp add: abc_steps_l.simps)
lemma list_crsp_simp2: "abc_list_crsp (lm1 @ 0\<up>n) lm2 \<Longrightarrow> abc_list_crsp lm1 lm2"
proof(induct n)
case 0
thus "?case"
by(auto simp: abc_list_crsp_def)
next
case (Suc n)
have ind: "abc_list_crsp (lm1 @ 0 \<up> n) lm2 \<Longrightarrow> abc_list_crsp lm1 lm2" by fact
have h: "abc_list_crsp (lm1 @ 0 \<up> Suc n) lm2" by fact
then have "abc_list_crsp (lm1 @ 0 \<up> n) lm2"
apply(auto simp only: exp_suc abc_list_crsp_def del: replicate_Suc)
apply (metis Suc_pred append_eq_append_conv
append_eq_append_conv2 butlast_append butlast_snoc length_replicate list.distinct(1)
neq0_conv replicate_Suc replicate_Suc_iff_anywhere replicate_app_Cons_same
replicate_empty self_append_conv self_append_conv2)
apply (auto,metis replicate_Suc)
.
thus "?case"
using ind
by auto
qed
lemma recursive_compile_correct_norm':
"\<lbrakk>rec_ci f = (ap, arity, ft);
terminate f args\<rbrakk>
\<Longrightarrow> \<exists> stp rl. (abc_steps_l (0, args) ap stp) = (length ap, rl) \<and> abc_list_crsp (args @ [rec_exec f args]) rl"
using recursive_compile_correct[of f args ap arity ft "[]"]
apply(auto simp: abc_Hoare_halt_def)
apply(rename_tac n)
apply(rule_tac x = n in exI)
apply(case_tac "abc_steps_l (0, args @ 0 \<up> (ft - arity)) ap n", auto)
apply(drule_tac abc_list_crsp_steps, auto)
apply(rule_tac list_crsp_simp2, auto)
done
lemma find_exponent_rec_exec[simp]:
assumes a: "args @ [rec_exec f args] = lm @ 0 \<up> n"
and b: "length args < length lm"
shows "\<exists>m. lm = args @ rec_exec f args # 0 \<up> m"
using assms
apply(cases n, simp)
apply(rule_tac x = 0 in exI, simp)
apply(drule_tac length_equal, simp)
done
lemma find_exponent_complex[simp]:
"\<lbrakk>args @ [rec_exec f args] = lm @ 0 \<up> n; \<not> length args < length lm\<rbrakk>
\<Longrightarrow> \<exists>m. (lm @ 0 \<up> (length args - length lm) @ [Suc 0])[length args := 0] =
args @ rec_exec f args # 0 \<up> m"
apply(cases n, simp_all add: exp_suc list_update_append list_update.simps del: replicate_Suc)
apply(subgoal_tac "length args = Suc (length lm)", simp)
apply(rule_tac x = "Suc (Suc 0)" in exI, simp)
apply(drule_tac length_equal, simp, auto)
done
lemma compile_append_dummy_correct:
assumes compile: "rec_ci f = (ap, ary, fp)"
and termi: "terminate f args"
shows "{\<lambda> nl. nl = args} (ap [+] dummy_abc (length args)) {\<lambda> nl. (\<exists> m. nl = args @ rec_exec f args # 0\<up>m)}"
proof(rule_tac abc_Hoare_plus_halt)
show "{\<lambda>nl. nl = args} ap {\<lambda> nl. abc_list_crsp (args @ [rec_exec f args]) nl}"
using compile termi recursive_compile_correct_norm'[of f ap ary fp args]
apply(auto simp: abc_Hoare_halt_def)
by (metis abc_final.simps abc_holds_for.simps)
next
show "{abc_list_crsp (args @ [rec_exec f args])} dummy_abc (length args)
{\<lambda>nl. \<exists>m. nl = args @ rec_exec f args # 0 \<up> m}"
apply(auto simp: dummy_abc_def abc_Hoare_halt_def)
apply(rule_tac x = 3 in exI)
by(force simp: abc_steps_l.simps abc_list_crsp_def abc_step_l.simps numeral_3_eq_3 abc_fetch.simps
abc_lm_v.simps nth_append abc_lm_s.simps)
qed
lemma cn_merge_gs_split:
"\<lbrakk>i < length gs; rec_ci (gs!i) = (ga, gb, gc)\<rbrakk> \<Longrightarrow>
cn_merge_gs (map rec_ci gs) p = cn_merge_gs (map rec_ci (take i gs)) p [+] (ga [+]
mv_box gb (p + i)) [+] cn_merge_gs (map rec_ci (drop (Suc i) gs)) (p + Suc i)"
proof(induct i arbitrary: gs p)
case 0
then show ?case by(cases gs; simp)
next
case (Suc i)
then show ?case
by(cases gs, simp, cases "rec_ci (hd gs)",
simp add: abc_comp_commute[THEN sym])
qed
lemma cn_unhalt_case:
assumes compile1: "rec_ci (Cn n f gs) = (ap, ar, ft) \<and> length args = ar"
and g: "i < length gs"
and compile2: "rec_ci (gs!i) = (gap, gar, gft) \<and> gar = length args"
and g_unhalt: "\<And> anything. {\<lambda> nl. nl = args @ 0\<up>(gft - gar) @ anything} gap \<up>"
and g_ind: "\<And> apj arj ftj j anything. \<lbrakk>j < i; rec_ci (gs!j) = (apj, arj, ftj)\<rbrakk>
\<Longrightarrow> {\<lambda> nl. nl = args @ 0\<up>(ftj - arj) @ anything} apj {\<lambda> nl. nl = args @ rec_exec (gs!j) args # 0\<up>(ftj - Suc arj) @ anything}"
and all_termi: "\<forall> j<i. terminate (gs!j) args"
shows "{\<lambda> nl. nl = args @ 0\<up>(ft - ar) @ anything} ap \<up>"
using compile1
apply(cases "rec_ci f", auto simp: rec_ci.simps abc_comp_commute)
proof(rule_tac abc_Hoare_plus_unhalt1)
fix fap far fft
let ?ft = "max (Suc (length args)) (Max (insert fft ((\<lambda>(aprog, p, n). n) ` rec_ci ` set gs)))"
let ?Q = "\<lambda>nl. nl = args @ 0\<up> (?ft - length args) @ map (\<lambda>i. rec_exec i args) (take i gs) @
0\<up>(length gs - i) @ 0\<up> Suc (length args) @ anything"
have "cn_merge_gs (map rec_ci gs) ?ft =
cn_merge_gs (map rec_ci (take i gs)) ?ft [+] (gap [+]
mv_box gar (?ft + i)) [+] cn_merge_gs (map rec_ci (drop (Suc i) gs)) (?ft + Suc i)"
using g compile2 cn_merge_gs_split by simp
thus "{\<lambda>nl. nl = args @ 0 # 0 \<up> (?ft + length gs) @ anything} (cn_merge_gs (map rec_ci gs) ?ft) \<up>"
proof(simp, rule_tac abc_Hoare_plus_unhalt1, rule_tac abc_Hoare_plus_unhalt2,
rule_tac abc_Hoare_plus_unhalt1)
let ?Q_tmp = "\<lambda>nl. nl = args @ 0\<up> (gft - gar) @ 0\<up>(?ft - (length args) - (gft -gar)) @ map (\<lambda>i. rec_exec i args) (take i gs) @
0\<up>(length gs - i) @ 0\<up> Suc (length args) @ anything"
have a: "{?Q_tmp} gap \<up>"
using g_unhalt[of "0 \<up> (?ft - (length args) - (gft - gar)) @
map (\<lambda>i. rec_exec i args) (take i gs) @ 0 \<up> (length gs - i) @ 0 \<up> Suc (length args) @ anything"]
by simp
moreover have "?ft \<ge> gft"
using g compile2
apply(rule_tac max.coboundedI2, rule_tac Max_ge, simp, rule_tac insertI2)
apply(rule_tac x = "rec_ci (gs ! i)" in image_eqI, simp)
by(rule_tac x = "gs!i" in image_eqI, simp, simp)
then have b:"?Q_tmp = ?Q"
using compile2
apply(rule_tac arg_cong)
by(simp add: replicate_merge_anywhere)
thus "{?Q} gap \<up>"
using a by simp
next
show "{\<lambda>nl. nl = args @ 0 # 0 \<up> (?ft + length gs) @ anything}
cn_merge_gs (map rec_ci (take i gs)) ?ft
{\<lambda>nl. nl = args @ 0 \<up> (?ft - length args) @
map (\<lambda>i. rec_exec i args) (take i gs) @ 0 \<up> (length gs - i) @ 0 \<up> Suc (length args) @ anything}"
using all_termi
by(rule_tac compile_cn_gs_correct', auto simp: set_conv_nth intro:g_ind)
qed
qed
lemma mn_unhalt_case':
assumes compile: "rec_ci f = (a, b, c)"
and all_termi: "\<forall>i. terminate f (args @ [i]) \<and> 0 < rec_exec f (args @ [i])"
and B: "B = [Dec (Suc (length args)) (length a + 5), Dec (Suc (length args)) (length a + 3),
Goto (Suc (length a)), Inc (length args), Goto 0]"
shows "{\<lambda>nl. nl = args @ 0 \<up> (max (Suc (length args)) c - length args) @ anything}
a @ B \<up>"
proof(rule_tac abc_Hoare_unhaltI, auto)
fix n
have a: "b = Suc (length args)"
using all_termi compile
apply(erule_tac x = 0 in allE)
by(auto, drule_tac param_pattern,auto)
moreover have b: "c > b"
using compile by(elim footprint_ge)
ultimately have c: "max (Suc (length args)) c = c" by arith
have "\<exists> stp > n. abc_steps_l (0, args @ 0 # 0\<up>(c - Suc (length args)) @ anything) (a @ B) stp
= (0, args @ Suc n # 0\<up>(c - Suc (length args)) @ anything)"
using assms a b c
proof(rule_tac mn_loop_correct', auto)
fix i xc
show "{\<lambda>nl. nl = args @ i # 0 \<up> (c - Suc (length args)) @ xc} a
{\<lambda>nl. nl = args @ i # rec_exec f (args @ [i]) # 0 \<up> (c - Suc (Suc (length args))) @ xc}"
using all_termi recursive_compile_correct[of f "args @ [i]" a b c xc] compile a
by(simp)
qed
then obtain stp where d: "stp > n \<and> abc_steps_l (0, args @ 0 # 0\<up>(c - Suc (length args)) @ anything) (a @ B) stp
= (0, args @ Suc n # 0\<up>(c - Suc (length args)) @ anything)" ..
then obtain d where e: "stp = n + Suc d"
by (metis add_Suc_right less_iff_Suc_add)
obtain s nl where f: "abc_steps_l (0, args @ 0 # 0\<up>(c - Suc (length args)) @ anything) (a @ B) n = (s, nl)"
by (metis prod.exhaust)
have g: "s < length (a @ B)"
using d e f
apply(rule_tac classical, simp only: abc_steps_add)
by(simp add: halt_steps2 leI)
from f g show "abc_notfinal (abc_steps_l (0, args @ 0 \<up>
(max (Suc (length args)) c - length args) @ anything) (a @ B) n) (a @ B)"
using c b a
by(simp add: replicate_Suc_iff_anywhere Suc_diff_Suc del: replicate_Suc)
qed
lemma mn_unhalt_case:
assumes compile: "rec_ci (Mn n f) = (ap, ar, ft) \<and> length args = ar"
and all_term: "\<forall> i. terminate f (args @ [i]) \<and> rec_exec f (args @ [i]) > 0"
shows "{\<lambda> nl. nl = args @ 0\<up>(ft - ar) @ anything} ap \<up> "
using assms
apply(cases "rec_ci f", auto simp: rec_ci.simps abc_comp_commute)
by(rule_tac mn_unhalt_case', simp_all)
fun tm_of_rec :: "recf \<Rightarrow> instr list"
where "tm_of_rec recf = (let (ap, k, fp) = rec_ci recf in
let tp = tm_of (ap [+] dummy_abc k) in
tp @ (shift (mopup k) (length tp div 2)))"
lemma recursive_compile_to_tm_correct1:
assumes compile: "rec_ci recf = (ap, ary, fp)"
and termi: " terminate recf args"
and tp: "tp = tm_of (ap [+] dummy_abc (length args))"
shows "\<exists> stp m l. steps0 (Suc 0, Bk # Bk # ires, <args> @ Bk\<up>rn)
(tp @ shift (mopup (length args)) (length tp div 2)) stp = (0, Bk\<up>m @ Bk # Bk # ires, Oc\<up>Suc (rec_exec recf args) @ Bk\<up>l)"
proof -
have "{\<lambda>nl. nl = args} ap [+] dummy_abc (length args) {\<lambda>nl. \<exists>m. nl = args @ rec_exec recf args # 0 \<up> m}"
using compile termi compile
by(rule_tac compile_append_dummy_correct, auto)
then obtain stp m where h: "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) stp =
(length (ap [+] dummy_abc (length args)), args @ rec_exec recf args # 0\<up>m) "
apply(simp add: abc_Hoare_halt_def, auto)
apply(rename_tac n)
by(case_tac "abc_steps_l (0, args) (ap [+] dummy_abc (length args)) n", auto)
thus "?thesis"
using assms tp compile_correct_halt[OF refl refl _ h _ _ refl]
by(auto simp: crsp.simps start_of.simps abc_lm_v.simps)
qed
lemma recursive_compile_to_tm_correct2:
assumes termi: " terminate recf args"
shows "\<exists> stp m l. steps0 (Suc 0, [Bk, Bk], <args>) (tm_of_rec recf) stp =
(0, Bk\<up>Suc (Suc m), Oc\<up>Suc (rec_exec recf args) @ Bk\<up>l)"
proof(cases "rec_ci recf", simp add: tm_of_rec.simps)
fix ap ar fp
assume "rec_ci recf = (ap, ar, fp)"
thus "\<exists>stp m l. steps0 (Suc 0, [Bk, Bk], <args>)
(tm_of (ap [+] dummy_abc ar) @ shift (mopup ar) (sum_list (layout_of (ap [+] dummy_abc ar)))) stp =
(0, Bk # Bk # Bk \<up> m, Oc # Oc \<up> rec_exec recf args @ Bk \<up> l)"
using recursive_compile_to_tm_correct1[of recf ap ar fp args "tm_of (ap [+] dummy_abc (length args))" "[]" 0]
assms param_pattern[of recf args ap ar fp]
by(simp add: replicate_Suc[THEN sym] replicate_Suc_iff_anywhere del: replicate_Suc,
simp add: exp_suc del: replicate_Suc)
qed
lemma recursive_compile_to_tm_correct3:
assumes termi: "terminate recf args"
shows "{\<lambda> tp. tp =([Bk, Bk], <args>)} (tm_of_rec recf)
{\<lambda> tp. \<exists> k l. tp = (Bk\<up> k, <rec_exec recf args> @ Bk \<up> l)}"
using recursive_compile_to_tm_correct2[OF assms]
apply(auto simp add: Hoare_halt_def ) apply(rename_tac stp M l)
apply(rule_tac x = stp in exI)
apply(auto simp add: tape_of_nat_def)
apply(rule_tac x = "Suc (Suc M)" in exI)
apply(simp)
done
lemma list_all_suc_many[simp]:
"list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))) xs \<Longrightarrow>
list_all (\<lambda>(acn, s). s \<le> Suc (Suc (Suc (Suc (Suc (Suc (Suc (Suc (2 * n))))))))) xs"
proof(induct xs)
case (Cons a xs)
then show ?case by(cases a, simp)
qed simp
lemma shift_append: "shift (xs @ ys) n = shift xs n @ shift ys n"
apply(simp add: shift.simps)
done
lemma length_shift_mopup[simp]: "length (shift (mopup n) ss) = 4 * n + 12"
apply(auto simp: mopup.simps shift_append mopup_b_def)
done
lemma length_tm_even[intro]: "length (tm_of ap) mod 2 = 0"
apply(simp add: tm_of.simps)
done
lemma tms_of_at_index[simp]: "k < length ap \<Longrightarrow> tms_of ap ! k =
ci (layout_of ap) (start_of (layout_of ap) k) (ap ! k)"
apply(simp add: tms_of.simps tpairs_of.simps)
done
lemma start_of_suc_inc:
"\<lbrakk>k < length ap; ap ! k = Inc n\<rbrakk> \<Longrightarrow> start_of (layout_of ap) (Suc k) =
start_of (layout_of ap) k + 2 * n + 9"
apply(rule_tac start_of_Suc1, auto simp: abc_fetch.simps)
done
lemma start_of_suc_dec:
"\<lbrakk>k < length ap; ap ! k = (Dec n e)\<rbrakk> \<Longrightarrow> start_of (layout_of ap) (Suc k) =
start_of (layout_of ap) k + 2 * n + 16"
apply(rule_tac start_of_Suc2, auto simp: abc_fetch.simps)
done
lemma inc_state_all_le:
"\<lbrakk>k < length ap; ap ! k = Inc n;
(a, b) \<in> set (shift (shift tinc_b (2 * n))
(start_of (layout_of ap) k - Suc 0))\<rbrakk>
\<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
apply(subgoal_tac "b \<le> start_of (layout_of ap) (Suc k)")
apply(subgoal_tac "start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap) ")
apply(arith)
apply(cases "Suc k = length ap", simp)
apply(rule_tac start_of_less, simp)
apply(auto simp: tinc_b_def shift.simps start_of_suc_inc length_of.simps )
done
lemma findnth_le[elim]:
"(a, b) \<in> set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))
\<Longrightarrow> b \<le> Suc (start_of (layout_of ap) k + 2 * n)"
apply(induct n, force simp add: shift.simps)
apply(simp add: shift_append, auto)
apply(auto simp: shift.simps)
done
lemma findnth_state_all_le1:
"\<lbrakk>k < length ap; ap ! k = Inc n;
(a, b) \<in>
set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\<rbrakk>
\<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
apply(subgoal_tac "b \<le> start_of (layout_of ap) (Suc k)")
apply(subgoal_tac "start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap) ")
apply(arith)
apply(cases "Suc k = length ap", simp)
apply(rule_tac start_of_less, simp)
apply(subgoal_tac "b \<le> start_of (layout_of ap) k + 2*n + 1 \<and>
start_of (layout_of ap) k + 2*n + 1 \<le> start_of (layout_of ap) (Suc k)", auto)
apply(auto simp: tinc_b_def shift.simps length_of.simps start_of_suc_inc)
done
lemma start_of_eq: "length ap < as \<Longrightarrow> start_of (layout_of ap) as = start_of (layout_of ap) (length ap)"
proof(induct as)
case (Suc as)
then show ?case
apply(cases "length ap < as", simp add: start_of.simps)
apply(subgoal_tac "as = length ap")
apply(simp add: start_of.simps)
apply arith
done
qed simp
lemma start_of_all_le: "start_of (layout_of ap) as \<le> start_of (layout_of ap) (length ap)"
apply(subgoal_tac "as > length ap \<or> as = length ap \<or> as < length ap",
auto simp: start_of_eq start_of_less)
done
lemma findnth_state_all_le2:
"\<lbrakk>k < length ap;
ap ! k = Dec n e;
(a, b) \<in> set (shift (findnth n) (start_of (layout_of ap) k - Suc 0))\<rbrakk>
\<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
apply(subgoal_tac "b \<le> start_of (layout_of ap) k + 2*n + 1 \<and>
start_of (layout_of ap) k + 2*n + 1 \<le> start_of (layout_of ap) (Suc k) \<and>
start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap)", auto)
apply(subgoal_tac "start_of (layout_of ap) (Suc k) =
start_of (layout_of ap) k + 2*n + 16", simp)
apply(simp add: start_of_suc_dec)
apply(rule_tac start_of_all_le)
done
lemma dec_state_all_le[simp]:
"\<lbrakk>k < length ap; ap ! k = Dec n e;
(a, b) \<in> set (shift (shift tdec_b (2 * n))
(start_of (layout_of ap) k - Suc 0))\<rbrakk>
\<Longrightarrow> b \<le> start_of (layout_of ap) (length ap)"
apply(subgoal_tac "2*n + start_of (layout_of ap) k + 16 \<le> start_of (layout_of ap) (length ap) \<and> start_of (layout_of ap) k > 0")
prefer 2
apply(subgoal_tac "start_of (layout_of ap) (Suc k) = start_of (layout_of ap) k + 2*n + 16
\<and> start_of (layout_of ap) (Suc k) \<le> start_of (layout_of ap) (length ap)")
apply(simp, rule_tac conjI)
apply(simp add: start_of_suc_dec)
apply(rule_tac start_of_all_le)
apply(auto simp: tdec_b_def shift.simps)
done
lemma tms_any_less:
"\<lbrakk>k < length ap; (a, b) \<in> set (tms_of ap ! k)\<rbrakk> \<Longrightarrow>
b \<le> start_of (layout_of ap) (length ap)"
apply(cases "ap!k", auto simp: tms_of.simps tpairs_of.simps ci.simps shift_append adjust.simps)
apply(erule_tac findnth_state_all_le1, simp_all)
apply(erule_tac inc_state_all_le, simp_all)
apply(erule_tac findnth_state_all_le2, simp_all)
apply(rule_tac start_of_all_le)
apply(rule_tac start_of_all_le)
done
lemma concat_in: "i < length (concat xs) \<Longrightarrow>
\<exists>k < length xs. concat xs ! i \<in> set (xs ! k)"
proof(induct xs rule: rev_induct)
case (snoc x xs)
then show ?case
apply(cases "i < length (concat xs)", simp)
apply(erule_tac exE, rule_tac x = k in exI)
apply(simp add: nth_append)
apply(rule_tac x = "length xs" in exI, simp)
apply(simp add: nth_append)
done
qed auto
declare length_concat[simp]
lemma in_tms: "i < length (tm_of ap) \<Longrightarrow> \<exists> k < length ap. (tm_of ap ! i) \<in> set (tms_of ap ! k)"
apply(simp only: tm_of.simps)
using concat_in[of i "tms_of ap"]
apply(auto)
done
lemma all_le_start_of: "list_all (\<lambda>(acn, s).
s \<le> start_of (layout_of ap) (length ap)) (tm_of ap)"
apply(simp only: list_all_length)
apply(rule_tac allI, rule_tac impI)
apply(drule_tac in_tms, auto elim: tms_any_less)
done
lemma length_ci:
"\<lbrakk>k < length ap; length (ci ly y (ap ! k)) = 2 * qa\<rbrakk>
\<Longrightarrow> layout_of ap ! k = qa"
apply(cases "ap ! k")
apply(auto simp: layout_of.simps ci.simps
length_of.simps tinc_b_def tdec_b_def length_findnth adjust.simps)
done
lemma ci_even[intro]: "length (ci ly y i) mod 2 = 0"
apply(cases i, auto simp: ci.simps length_findnth
tinc_b_def adjust.simps tdec_b_def)
done
lemma sum_list_ci_even[intro]: "sum_list (map (length \<circ> (\<lambda>(x, y). ci ly x y)) zs) mod 2 = 0"
proof(induct zs rule: rev_induct)
case (snoc x xs)
then show ?case
apply(cases x, simp)
apply(subgoal_tac "length (ci ly (fst x) (snd x)) mod 2 = 0")
apply(auto)
done
qed (simp)
lemma zip_pre:
"(length ys) \<le> length ap \<Longrightarrow>
zip ys ap = zip ys (take (length ys) (ap::'a list))"
proof(induct ys arbitrary: ap)
case (Cons a ys)
from Cons(2) have z:"ap = aa # list \<Longrightarrow> zip (a # ys) ap = zip (a # ys) (take (length (a # ys)) ap)"
for aa list using Cons(1)[of list] by simp
thus ?case by (cases ap;simp)
qed simp
lemma length_start_of_tm: "start_of (layout_of ap) (length ap) = Suc (length (tm_of ap) div 2)"
using tpa_states[of "tm_of ap" "length ap" ap]
by(simp add: tm_of.simps)
lemma list_all_add_6E[elim]: "list_all (\<lambda>(acn, s). s \<le> Suc q) xs
\<Longrightarrow> list_all (\<lambda>(acn, s). s \<le> q + (2 * n + 6)) xs"
by(auto simp: list_all_length)
lemma mopup_b_12[simp]: "length mopup_b = 12"
by(simp add: mopup_b_def)
lemma mp_up_all_le: "list_all (\<lambda>(acn, s). s \<le> q + (2 * n + 6))
[(R, Suc (Suc (2 * n + q))), (R, Suc (2 * n + q)),
(L, 5 + 2 * n + q), (W0, Suc (Suc (Suc (2 * n + q)))), (R, 4 + 2 * n + q),
(W0, Suc (Suc (Suc (2 * n + q)))), (R, Suc (Suc (2 * n + q))),
(W0, Suc (Suc (Suc (2 * n + q)))), (L, 5 + 2 * n + q),
(L, 6 + 2 * n + q), (R, 0), (L, 6 + 2 * n + q)]"
by(auto)
lemma mopup_le6[simp]: "(a, b) \<in> set (mopup_a n) \<Longrightarrow> b \<le> 2 * n + 6"
by(induct n, auto simp: mopup_a.simps)
lemma shift_le2[simp]: "(a, b) \<in> set (shift (mopup n) x)
\<Longrightarrow> b \<le> (2 * x + length (mopup n)) div 2"
apply(auto simp: mopup.simps shift_append shift.simps)
apply(auto simp: mopup_b_def)
done
lemma mopup_ge2[intro]: " 2 \<le> x + length (mopup n)"
apply(simp add: mopup.simps)
done
lemma mopup_even[intro]: " (2 * x + length (mopup n)) mod 2 = 0"
by(auto simp: mopup.simps)
lemma mopup_div_2[simp]: "b \<le> Suc x
\<Longrightarrow> b \<le> (2 * x + length (mopup n)) div 2"
by(auto simp: mopup.simps)
lemma wf_tm_from_abacus: assumes "tp = tm_of ap"
shows "tm_wf0 (tp @ shift (mopup n) (length tp div 2))"
proof -
have "is_even (length (mopup n))" for n using tm_wf.simps by blast
moreover have "(aa, ba) \<in> set (mopup n) \<Longrightarrow> ba \<le> length (mopup n) div 2" for aa ba
by (metis (no_types, lifting) add_cancel_left_right case_prodD tm_wf.simps wf_mopup)
moreover have "(\<forall>x\<in>set (tm_of ap). case x of (acn, s) \<Rightarrow> s \<le> Suc (sum_list (layout_of ap))) \<Longrightarrow>
(a, b) \<in> set (tm_of ap) \<Longrightarrow> b \<le> sum_list (layout_of ap) + length (mopup n) div 2"
for a b s
by (metis (no_types, lifting) add_Suc add_cancel_left_right case_prodD div_mult_mod_eq le_SucE mult_2_right not_numeral_le_zero tm_wf.simps trans_le_add1 wf_mopup)
ultimately show ?thesis unfolding assms
using length_start_of_tm[of ap] all_le_start_of[of ap] tm_wf.simps
by(auto simp: List.list_all_iff shift.simps)
qed
lemma wf_tm_from_recf:
assumes compile: "tp = tm_of_rec recf"
shows "tm_wf0 tp"
proof -
obtain a b c where "rec_ci recf = (a, b, c)"
by (metis prod_cases3)
thus "?thesis"
using compile
using wf_tm_from_abacus[of "tm_of (a [+] dummy_abc b)" "(a [+] dummy_abc b)" b]
by simp
qed
end
\ No newline at end of file
diff --git a/thys/Universal_Turing_Machine/UTM.thy b/thys/Universal_Turing_Machine/UTM.thy
--- a/thys/Universal_Turing_Machine/UTM.thy
+++ b/thys/Universal_Turing_Machine/UTM.thy
@@ -1,4708 +1,4706 @@
(* Title: thys/UTM.thy
Author: Jian Xu, Xingyuan Zhang, and Christian Urban
Modifications: Sebastiaan Joosten
*)
chapter \<open>Construction of a Universal Turing Machine\<close>
theory UTM
imports Recursive Abacus UF HOL.GCD Turing_Hoare
begin
section \<open>Wang coding of input arguments\<close>
text \<open>
The direct compilation of the universal function \<open>rec_F\<close> can
not give us UTM, because \<open>rec_F\<close> is of arity 2, where the
first argument represents the Godel coding of the TM being simulated
and the second argument represents the right number (in Wang's
coding) of the TM tape. (Notice, left number is always \<open>0\<close>
at the very beginning). However, UTM needs to simulate the execution
of any TM which may very well take many input arguments. Therefore,
a initialization TM needs to run before the TM compiled from \<open>rec_F\<close>, and the sequential composition of these two TMs will give
rise to the UTM we are seeking. The purpose of this initialization
TM is to transform the multiple input arguments of the TM being
simulated into Wang's coding, so that it can be consumed by the TM
compiled from \<open>rec_F\<close> as the second argument.
However, this initialization TM (named \<open>t_wcode\<close>) can not be
constructed by compiling from any recursive function, because every
recursive function takes a fixed number of input arguments, while
\<open>t_wcode\<close> needs to take varying number of arguments and
tranform them into Wang's coding. Therefore, this section give a
direct construction of \<open>t_wcode\<close> with just some parts being
obtained from recursive functions.
\newlength{\basewidth}
\settowidth{\basewidth}{xxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxxx}
\newlength{\baseheight}
\settoheight{\baseheight}{$B:R$}
\newcommand{\vsep}{5\baseheight}
The TM used to generate the Wang's code of input arguments is divided into three TMs
executed sequentially, namely $prepare$, $mainwork$ and $adjust$.
According to the
convention, the start state of ever TM is fixed to state $1$ while the final state is
fixed to $0$.
The input and output of $prepare$ are illustrated respectively by Figure
\ref{prepare_input} and \ref{prepare_output}.
\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
[tbox/.style = {draw, thick, inner sep = 5pt}]
\node (0) {};
\node (1) [tbox, text height = 3.5pt, right = -0.9pt of 0] {$m$};
\node (2) [tbox, right = -0.9pt of 1] {$0$};
\node (3) [tbox, right = -0.9pt of 2] {$a_1$};
\node (4) [tbox, right = -0.9pt of 3] {$0$};
\node (5) [tbox, right = -0.9pt of 4] {$a_2$};
\node (6) [right = -0.9pt of 5] {\ldots \ldots};
\node (7) [tbox, right = -0.9pt of 6] {$a_n$};
\draw [->, >=latex, thick] (1)+(0, -4\baseheight) -- (1);
\end{tikzpicture}}
\caption{The input of TM $prepare$} \label{prepare_input}
\end{figure}
\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
\node (0) {};
\node (1) [draw, text height = 3.5pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$};
\node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$};
\node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$};
\node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$};
\node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$};
\node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$};
\node (7) [right = -0.9pt of 6] {\ldots \ldots};
\node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_n$};
\node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$0$};
\node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$};
\node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$1$};
\draw [->, >=latex, thick] (10)+(0, -4\baseheight) -- (10);
\end{tikzpicture}}
\caption{The output of TM $prepare$} \label{prepare_output}
\end{figure}
As shown in Figure \ref{prepare_input}, the input of $prepare$ is the
same as the the input of UTM, where $m$ is the Godel coding of the TM
being interpreted and $a_1$ through $a_n$ are the $n$ input arguments
of the TM under interpretation. The purpose of $purpose$ is to
transform this initial tape layout to the one shown in Figure
\ref{prepare_output}, which is convenient for the generation of Wang's
codding of $a_1, \ldots, a_n$. The coding procedure starts from $a_n$
and ends after $a_1$ is encoded. The coding result is stored in an
accumulator at the end of the tape (initially represented by the $1$
two blanks right to $a_n$ in Figure \ref{prepare_output}). In Figure
\ref{prepare_output}, arguments $a_1, \ldots, a_n$ are separated by
two blanks on both ends with the rest so that movement conditions can
be implemented conveniently in subsequent TMs, because, by convention,
two consecutive blanks are usually used to signal the end or start of
a large chunk of data. The diagram of $prepare$ is given in Figure
\ref{prepare_diag}.
\begin{figure}[h!]
\centering
\scalebox{0.9}{
\begin{tikzpicture}
\node[circle,draw] (1) {$1$};
\node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$};
\node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$};
\node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$};
\node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$};
\node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$};
\node[circle,draw] (7) at ($(6)+(0.3\basewidth, 0)$) {$7$};
\node[circle,draw] (8) at ($(7)+(0.3\basewidth, 0)$) {$0$};
\draw [->, >=latex] (1) edge [loop above] node[above] {$S_1:L$} (1)
;
\draw [->, >=latex] (1) -- node[above] {$S_0:S_1$} (2)
;
\draw [->, >=latex] (2) edge [loop above] node[above] {$S_1:R$} (2)
;
\draw [->, >=latex] (2) -- node[above] {$S_0:L$} (3)
;
\draw [->, >=latex] (3) edge[loop above] node[above] {$S_1:S_0$} (3)
;
\draw [->, >=latex] (3) -- node[above] {$S_0:R$} (4)
;
\draw [->, >=latex] (4) edge[loop above] node[above] {$S_0:R$} (4)
;
\draw [->, >=latex] (4) -- node[above] {$S_0:R$} (5)
;
\draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:R$} (5)
;
\draw [->, >=latex] (5) -- node[above] {$S_0:R$} (6)
;
\draw [->, >=latex] (6) edge[bend left = 50] node[below] {$S_1:R$} (5)
;
\draw [->, >=latex] (6) -- node[above] {$S_0:R$} (7)
;
\draw [->, >=latex] (7) edge[loop above] node[above] {$S_0:S_1$} (7)
;
\draw [->, >=latex] (7) -- node[above] {$S_1:L$} (8)
;
\end{tikzpicture}}
\caption{The diagram of TM $prepare$} \label{prepare_diag}
\end{figure}
The purpose of TM $mainwork$ is to compute the Wang's encoding of $a_1, \ldots, a_n$. Every bit of $a_1, \ldots, a_n$, including the separating bits, is processed from left to right.
In order to detect the termination condition when the left most bit of $a_1$ is reached,
TM $mainwork$ needs to look ahead and consider three different situations at the start of
every iteration:
\begin{enumerate}
\item The TM configuration for the first situation is shown in Figure \ref{mainwork_case_one_input},
where the accumulator is stored in $r$, both of the next two bits
to be encoded are $1$. The configuration at the end of the iteration
is shown in Figure \ref{mainwork_case_one_output}, where the first 1-bit has been
encoded and cleared. Notice that the accumulator has been changed to
$(r+1) \times 2$ to reflect the encoded bit.
\item The TM configuration for the second situation is shown in Figure
\ref{mainwork_case_two_input},
where the accumulator is stored in $r$, the next two bits
to be encoded are $1$ and $0$. After the first
$1$-bit was encoded and cleared, the second $0$-bit is difficult to detect
and process. To solve this problem, these two consecutive bits are
encoded in one iteration. In this situation, only the first $1$-bit needs
to be cleared since the second one is cleared by definition.
The configuration at the end of the iteration
is shown in Figure \ref{mainwork_case_two_output}.
Notice that the accumulator has been changed to
$(r+1) \times 4$ to reflect the two encoded bits.
\item The third situation corresponds to the case when the last bit of $a_1$ is reached.
The TM configurations at the start and end of the iteration are shown in
Figure \ref{mainwork_case_three_input} and \ref{mainwork_case_three_output}
respectively. For this situation, only the read write head needs to be moved to
the left to prepare a initial configuration for TM $adjust$ to start with.
\end{enumerate}
The diagram of $mainwork$ is given in Figure \ref{mainwork_diag}. The two rectangular nodes
labeled with $2 \times x$ and $4 \times x$ are two TMs compiling from recursive functions
so that we do not have to design and verify two quite complicated TMs.
\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
\node (0) {};
\node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$};
\node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$};
\node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$};
\node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$};
\node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$};
\node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$};
\node (7) [right = -0.9pt of 6] {\ldots \ldots};
\node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$};
\node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$};
\node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$1$};
\node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$0$};
\node (12) [right = -0.9pt of 11] {\ldots \ldots};
\node (13) [draw, right = -0.9pt of 12, thick, inner sep = 5pt] {$0$};
\node (14) [draw, text height = 3.9pt, right = -0.9pt of 13, thick, inner sep = 5pt] {$r$};
\draw [->, >=latex, thick] (13)+(0, -4\baseheight) -- (13);
\end{tikzpicture}}
\caption{The first situation for TM $mainwork$ to consider} \label{mainwork_case_one_input}
\end{figure}
\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
\node (0) {};
\node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$};
\node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$};
\node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$};
\node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$};
\node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$};
\node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$};
\node (7) [right = -0.9pt of 6] {\ldots \ldots};
\node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$};
\node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$};
\node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$};
\node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$0$};
\node (12) [right = -0.9pt of 11] {\ldots \ldots};
\node (13) [draw, right = -0.9pt of 12, thick, inner sep = 5pt] {$0$};
\node (14) [draw, text height = 2.7pt, right = -0.9pt of 13, thick, inner sep = 5pt] {$(r+1) \times 2$};
\draw [->, >=latex, thick] (13)+(0, -4\baseheight) -- (13);
\end{tikzpicture}}
\caption{The output for the first case of TM $mainwork$'s processing}
\label{mainwork_case_one_output}
\end{figure}
\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
\node (0) {};
\node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$};
\node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$};
\node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$};
\node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$};
\node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$};
\node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$};
\node (7) [right = -0.9pt of 6] {\ldots \ldots};
\node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$};
\node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$};
\node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$};
\node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$1$};
\node (12) [draw, right = -0.9pt of 11, thick, inner sep = 5pt] {$0$};
\node (13) [right = -0.9pt of 12] {\ldots \ldots};
\node (14) [draw, right = -0.9pt of 13, thick, inner sep = 5pt] {$0$};
\node (15) [draw, text height = 3.9pt, right = -0.9pt of 14, thick, inner sep = 5pt] {$r$};
\draw [->, >=latex, thick] (14)+(0, -4\baseheight) -- (14);
\end{tikzpicture}}
\caption{The second situation for TM $mainwork$ to consider} \label{mainwork_case_two_input}
\end{figure}
\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
\node (0) {};
\node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$};
\node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$};
\node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$};
\node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$a_1$};
\node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$};
\node (6) [draw, right = -0.9pt of 5, thick, inner sep = 5pt] {$a_2$};
\node (7) [right = -0.9pt of 6] {\ldots \ldots};
\node (8) [draw, right = -0.9pt of 7, thick, inner sep = 5pt] {$a_i$};
\node (9) [draw, right = -0.9pt of 8, thick, inner sep = 5pt] {$1$};
\node (10) [draw, right = -0.9pt of 9, thick, inner sep = 5pt] {$0$};
\node (11) [draw, right = -0.9pt of 10, thick, inner sep = 5pt] {$0$};
\node (12) [draw, right = -0.9pt of 11, thick, inner sep = 5pt] {$0$};
\node (13) [right = -0.9pt of 12] {\ldots \ldots};
\node (14) [draw, right = -0.9pt of 13, thick, inner sep = 5pt] {$0$};
\node (15) [draw, text height = 2.7pt, right = -0.9pt of 14, thick, inner sep = 5pt] {$(r+1) \times 4$};
\draw [->, >=latex, thick] (14)+(0, -4\baseheight) -- (14);
\end{tikzpicture}}
\caption{The output for the second case of TM $mainwork$'s processing}
\label{mainwork_case_two_output}
\end{figure}
\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
\node (0) {};
\node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$};
\node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$};
\node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$};
\node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$1$};
\node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$};
\node (6) [right = -0.9pt of 5] {\ldots \ldots};
\node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {$0$};
\node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {$r$};
\draw [->, >=latex, thick] (7)+(0, -4\baseheight) -- (7);
\end{tikzpicture}}
\caption{The third situation for TM $mainwork$ to consider} \label{mainwork_case_three_input}
\end{figure}
\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
\node (0) {};
\node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$};
\node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$};
\node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$};
\node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$1$};
\node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$};
\node (6) [right = -0.9pt of 5] {\ldots \ldots};
\node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {$0$};
\node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {$r$};
\draw [->, >=latex, thick] (3)+(0, -4\baseheight) -- (3);
\end{tikzpicture}}
\caption{The output for the third case of TM $mainwork$'s processing}
\label{mainwork_case_three_output}
\end{figure}
\begin{figure}[h!]
\centering
\scalebox{0.9}{
\begin{tikzpicture}
\node[circle,draw] (1) {$1$};
\node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$};
\node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$};
\node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$};
\node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$};
\node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$};
\node[circle,draw] (7) at ($(2)+(0, -7\baseheight)$) {$7$};
\node[circle,draw] (8) at ($(7)+(0, -7\baseheight)$) {$8$};
\node[circle,draw] (9) at ($(8)+(0.3\basewidth, 0)$) {$9$};
\node[circle,draw] (10) at ($(9)+(0.3\basewidth, 0)$) {$10$};
\node[circle,draw] (11) at ($(10)+(0.3\basewidth, 0)$) {$11$};
\node[circle,draw] (12) at ($(11)+(0.3\basewidth, 0)$) {$12$};
\node[draw] (13) at ($(6)+(0.3\basewidth, 0)$) {$2 \times x$};
\node[circle,draw] (14) at ($(13)+(0.3\basewidth, 0)$) {$j_1$};
\node[draw] (15) at ($(12)+(0.3\basewidth, 0)$) {$4 \times x$};
\node[draw] (16) at ($(15)+(0.3\basewidth, 0)$) {$j_2$};
\node[draw] (17) at ($(7)+(0.3\basewidth, 0)$) {$0$};
\draw [->, >=latex] (1) edge[loop left] node[above] {$S_0:L$} (1)
;
\draw [->, >=latex] (1) -- node[above] {$S_1:L$} (2)
;
\draw [->, >=latex] (2) -- node[above] {$S_1:R$} (3)
;
\draw [->, >=latex] (2) -- node[left] {$S_1:L$} (7)
;
\draw [->, >=latex] (3) edge[loop above] node[above] {$S_1:S_0$} (3)
;
\draw [->, >=latex] (3) -- node[above] {$S_0:R$} (4)
;
\draw [->, >=latex] (4) edge[loop above] node[above] {$S_0:R$} (4)
;
\draw [->, >=latex] (4) -- node[above] {$S_1:R$} (5)
;
\draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:R$} (5)
;
\draw [->, >=latex] (5) -- node[above] {$S_0:S_1$} (6)
;
\draw [->, >=latex] (6) edge[loop above] node[above] {$S_1:L$} (6)
;
\draw [->, >=latex] (6) -- node[above] {$S_0:R$} (13)
;
\draw [->, >=latex] (13) -- (14)
;
\draw (14) -- ($(14)+(0, 6\baseheight)$) -- ($(1) + (0, 6\baseheight)$) node [above,midway] {$S_1:L$}
;
\draw [->, >=latex] ($(1) + (0, 6\baseheight)$) -- (1)
;
\draw [->, >=latex] (7) -- node[above] {$S_0:R$} (17)
;
\draw [->, >=latex] (7) -- node[left] {$S_1:R$} (8)
;
\draw [->, >=latex] (8) -- node[above] {$S_0:R$} (9)
;
\draw [->, >=latex] (9) -- node[above] {$S_0:R$} (10)
;
\draw [->, >=latex] (10) -- node[above] {$S_1:R$} (11)
;
\draw [->, >=latex] (10) edge[loop above] node[above] {$S_0:R$} (10)
;
\draw [->, >=latex] (11) edge[loop above] node[above] {$S_1:R$} (11)
;
\draw [->, >=latex] (11) -- node[above] {$S_0:S_1$} (12)
;
\draw [->, >=latex] (12) -- node[above] {$S_0:R$} (15)
;
\draw [->, >=latex] (12) edge[loop above] node[above] {$S_1:L$} (12)
;
\draw [->, >=latex] (15) -- (16)
;
\draw (16) -- ($(16)+(0, -4\baseheight)$) -- ($(1) + (0, -18\baseheight)$) node [below,midway] {$S_1:L$}
;
\draw [->, >=latex] ($(1) + (0, -18\baseheight)$) -- (1)
;
\end{tikzpicture}}
\caption{The diagram of TM $mainwork$} \label{mainwork_diag}
\end{figure}
The purpose of TM $adjust$ is to encode the last bit of $a_1$. The initial and final configuration
of this TM are shown in Figure \ref{adjust_initial} and \ref{adjust_final} respectively.
The diagram of TM $adjust$ is shown in Figure \ref{adjust_diag}.
\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
\node (0) {};
\node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$};
\node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$};
\node (3) [draw, right = -0.9pt of 2, thick, inner sep = 5pt] {$0$};
\node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$1$};
\node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$};
\node (6) [right = -0.9pt of 5] {\ldots \ldots};
\node (7) [draw, right = -0.9pt of 6, thick, inner sep = 5pt] {$0$};
\node (8) [draw, text height = 3.9pt, right = -0.9pt of 7, thick, inner sep = 5pt] {$r$};
\draw [->, >=latex, thick] (3)+(0, -4\baseheight) -- (3);
\end{tikzpicture}}
\caption{Initial configuration of TM $adjust$} \label{adjust_initial}
\end{figure}
\begin{figure}[h!]
\centering
\scalebox{1.2}{
\begin{tikzpicture}
\node (0) {};
\node (1) [draw, text height = 3.9pt, right = -0.9pt of 0, thick, inner sep = 5pt] {$m$};
\node (2) [draw, right = -0.9pt of 1, thick, inner sep = 5pt] {$0$};
\node (3) [draw, text height = 2.9pt, right = -0.9pt of 2, thick, inner sep = 5pt] {$r+1$};
\node (4) [draw, right = -0.9pt of 3, thick, inner sep = 5pt] {$0$};
\node (5) [draw, right = -0.9pt of 4, thick, inner sep = 5pt] {$0$};
\node (6) [right = -0.9pt of 5] {\ldots \ldots};
\draw [->, >=latex, thick] (1)+(0, -4\baseheight) -- (1);
\end{tikzpicture}}
\caption{Final configuration of TM $adjust$} \label{adjust_final}
\end{figure}
\begin{figure}[h!]
\centering
\scalebox{0.9}{
\begin{tikzpicture}
\node[circle,draw] (1) {$1$};
\node[circle,draw] (2) at ($(1)+(0.3\basewidth, 0)$) {$2$};
\node[circle,draw] (3) at ($(2)+(0.3\basewidth, 0)$) {$3$};
\node[circle,draw] (4) at ($(3)+(0.3\basewidth, 0)$) {$4$};
\node[circle,draw] (5) at ($(4)+(0.3\basewidth, 0)$) {$5$};
\node[circle,draw] (6) at ($(5)+(0.3\basewidth, 0)$) {$6$};
\node[circle,draw] (7) at ($(6)+(0.3\basewidth, 0)$) {$7$};
\node[circle,draw] (8) at ($(4)+(0, -7\baseheight)$) {$8$};
\node[circle,draw] (9) at ($(8)+(0.3\basewidth, 0)$) {$9$};
\node[circle,draw] (10) at ($(9)+(0.3\basewidth, 0)$) {$10$};
\node[circle,draw] (11) at ($(10)+(0.3\basewidth, 0)$) {$11$};
\node[circle,draw] (12) at ($(11)+(0.3\basewidth, 0)$) {$0$};
\draw [->, >=latex] (1) -- node[above] {$S_1:R$} (2)
;
\draw [->, >=latex] (1) edge[loop above] node[above] {$S_0:S_1$} (1)
;
\draw [->, >=latex] (2) -- node[above] {$S_1:R$} (3)
;
\draw [->, >=latex] (3) edge[loop above] node[above] {$S_0:R$} (3)
;
\draw [->, >=latex] (3) -- node[above] {$S_1:R$} (4)
;
\draw [->, >=latex] (4) -- node[above] {$S_1:L$} (5)
;
\draw [->, >=latex] (4) -- node[right] {$S_0:L$} (8)
;
\draw [->, >=latex] (5) -- node[above] {$S_0:L$} (6)
;
\draw [->, >=latex] (5) edge[loop above] node[above] {$S_1:S_0$} (5)
;
\draw [->, >=latex] (6) -- node[above] {$S_1:R$} (7)
;
\draw [->, >=latex] (6) edge[loop above] node[above] {$S_0:L$} (6)
;
\draw (7) -- ($(7)+(0, 6\baseheight)$) -- ($(2) + (0, 6\baseheight)$) node [above,midway] {$S_0:S_1$}
;
\draw [->, >=latex] ($(2) + (0, 6\baseheight)$) -- (2)
;
\draw [->, >=latex] (8) edge[loop left] node[left] {$S_1:S_0$} (8)
;
\draw [->, >=latex] (8) -- node[above] {$S_0:L$} (9)
;
\draw [->, >=latex] (9) edge[loop above] node[above] {$S_0:L$} (9)
;
\draw [->, >=latex] (9) -- node[above] {$S_1:L$} (10)
;
\draw [->, >=latex] (10) edge[loop above] node[above] {$S_0:L$} (10)
;
\draw [->, >=latex] (10) -- node[above] {$S_0:L$} (11)
;
\draw [->, >=latex] (11) edge[loop above] node[above] {$S_1:L$} (11)
;
\draw [->, >=latex] (11) -- node[above] {$S_0:R$} (12)
;
\end{tikzpicture}}
\caption{Diagram of TM $adjust$} \label{adjust_diag}
\end{figure}
\<close>
definition rec_twice :: "recf"
where
"rec_twice = Cn 1 rec_mult [id 1 0, constn 2]"
definition rec_fourtimes :: "recf"
where
"rec_fourtimes = Cn 1 rec_mult [id 1 0, constn 4]"
definition abc_twice :: "abc_prog"
where
"abc_twice = (let (aprog, ary, fp) = rec_ci rec_twice in
aprog [+] dummy_abc ((Suc 0)))"
definition abc_fourtimes :: "abc_prog"
where
"abc_fourtimes = (let (aprog, ary, fp) = rec_ci rec_fourtimes in
aprog [+] dummy_abc ((Suc 0)))"
definition twice_ly :: "nat list"
where
"twice_ly = layout_of abc_twice"
definition fourtimes_ly :: "nat list"
where
"fourtimes_ly = layout_of abc_fourtimes"
definition t_twice_compile :: "instr list"
where
"t_twice_compile= (tm_of abc_twice @ (shift (mopup 1) (length (tm_of abc_twice) div 2)))"
definition t_twice :: "instr list"
where
"t_twice = adjust0 t_twice_compile"
definition t_fourtimes_compile :: "instr list"
where
"t_fourtimes_compile= (tm_of abc_fourtimes @ (shift (mopup 1) (length (tm_of abc_fourtimes) div 2)))"
definition t_fourtimes :: "instr list"
where
"t_fourtimes = adjust0 t_fourtimes_compile"
definition t_twice_len :: "nat"
where
"t_twice_len = length t_twice div 2"
definition t_wcode_main_first_part:: "instr list"
where
"t_wcode_main_first_part \<equiv>
[(L, 1), (L, 2), (L, 7), (R, 3),
(R, 4), (W0, 3), (R, 4), (R, 5),
(W1, 6), (R, 5), (R, 13), (L, 6),
(R, 0), (R, 8), (R, 9), (Nop, 8),
(R, 10), (W0, 9), (R, 10), (R, 11),
(W1, 12), (R, 11), (R, t_twice_len + 14), (L, 12)]"
definition t_wcode_main :: "instr list"
where
"t_wcode_main = (t_wcode_main_first_part @ shift t_twice 12 @ [(L, 1), (L, 1)]
@ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])"
fun bl_bin :: "cell list \<Rightarrow> nat"
where
"bl_bin [] = 0"
| "bl_bin (Bk # xs) = 2 * bl_bin xs"
| "bl_bin (Oc # xs) = Suc (2 * bl_bin xs)"
declare bl_bin.simps[simp del]
type_synonym bin_inv_t = "cell list \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
fun wcode_before_double :: "bin_inv_t"
where
"wcode_before_double ires rs (l, r) =
(\<exists> ln rn. l = Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and>
r = Oc\<up>((Suc (Suc rs))) @ Bk\<up>(rn ))"
declare wcode_before_double.simps[simp del]
fun wcode_after_double :: "bin_inv_t"
where
"wcode_after_double ires rs (l, r) =
(\<exists> ln rn. l = Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and>
r = Oc\<up>(Suc (Suc (Suc 2*rs))) @ Bk\<up>(rn))"
declare wcode_after_double.simps[simp del]
fun wcode_on_left_moving_1_B :: "bin_inv_t"
where
"wcode_on_left_moving_1_B ires rs (l, r) =
(\<exists> ml mr rn. l = Bk\<up>(ml) @ Oc # Oc # ires \<and>
r = Bk\<up>(mr) @ Oc\<up>(Suc rs) @ Bk\<up>(rn) \<and>
ml + mr > Suc 0 \<and> mr > 0)"
declare wcode_on_left_moving_1_B.simps[simp del]
fun wcode_on_left_moving_1_O :: "bin_inv_t"
where
"wcode_on_left_moving_1_O ires rs (l, r) =
(\<exists> ln rn.
l = Oc # ires \<and>
r = Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
declare wcode_on_left_moving_1_O.simps[simp del]
fun wcode_on_left_moving_1 :: "bin_inv_t"
where
"wcode_on_left_moving_1 ires rs (l, r) =
(wcode_on_left_moving_1_B ires rs (l, r) \<or> wcode_on_left_moving_1_O ires rs (l, r))"
declare wcode_on_left_moving_1.simps[simp del]
fun wcode_on_checking_1 :: "bin_inv_t"
where
"wcode_on_checking_1 ires rs (l, r) =
(\<exists> ln rn. l = ires \<and>
r = Oc # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
fun wcode_erase1 :: "bin_inv_t"
where
"wcode_erase1 ires rs (l, r) =
(\<exists> ln rn. l = Oc # ires \<and>
tl r = Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
declare wcode_erase1.simps [simp del]
fun wcode_on_right_moving_1 :: "bin_inv_t"
where
"wcode_on_right_moving_1 ires rs (l, r) =
(\<exists> ml mr rn.
l = Bk\<up>(ml) @ Oc # ires \<and>
r = Bk\<up>(mr) @ Oc\<up>(Suc rs) @ Bk\<up>(rn) \<and>
ml + mr > Suc 0)"
declare wcode_on_right_moving_1.simps [simp del]
declare wcode_on_right_moving_1.simps[simp del]
fun wcode_goon_right_moving_1 :: "bin_inv_t"
where
"wcode_goon_right_moving_1 ires rs (l, r) =
(\<exists> ml mr ln rn.
l = Oc\<up>(ml) @ Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and>
r = Oc\<up>(mr) @ Bk\<up>(rn) \<and>
ml + mr = Suc rs)"
declare wcode_goon_right_moving_1.simps[simp del]
fun wcode_backto_standard_pos_B :: "bin_inv_t"
where
"wcode_backto_standard_pos_B ires rs (l, r) =
(\<exists> ln rn. l = Bk # Bk\<up>(ln) @ Oc # ires \<and>
r = Bk # Oc\<up>((Suc (Suc rs))) @ Bk\<up>(rn ))"
declare wcode_backto_standard_pos_B.simps[simp del]
fun wcode_backto_standard_pos_O :: "bin_inv_t"
where
"wcode_backto_standard_pos_O ires rs (l, r) =
(\<exists> ml mr ln rn.
l = Oc\<up>(ml) @ Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and>
r = Oc\<up>(mr) @ Bk\<up>(rn) \<and>
ml + mr = Suc (Suc rs) \<and> mr > 0)"
declare wcode_backto_standard_pos_O.simps[simp del]
fun wcode_backto_standard_pos :: "bin_inv_t"
where
"wcode_backto_standard_pos ires rs (l, r) = (wcode_backto_standard_pos_B ires rs (l, r) \<or>
wcode_backto_standard_pos_O ires rs (l, r))"
declare wcode_backto_standard_pos.simps[simp del]
lemma bin_wc_eq: "bl_bin xs = bl2wc xs"
proof(induct xs)
show " bl_bin [] = bl2wc []"
apply(simp add: bl_bin.simps)
done
next
fix a xs
assume "bl_bin xs = bl2wc xs"
thus " bl_bin (a # xs) = bl2wc (a # xs)"
apply(case_tac a, simp_all add: bl_bin.simps bl2wc.simps)
apply(simp_all add: bl2nat.simps bl2nat_double)
done
qed
lemma tape_of_nl_append_one: "lm \<noteq> [] \<Longrightarrow> <lm @ [a]> = <lm> @ Bk # Oc\<up>Suc a"
apply(induct lm, auto simp: tape_of_nl_cons split:if_splits)
done
lemma tape_of_nl_rev: "rev (<lm::nat list>) = (<rev lm>)"
apply(induct lm, simp, auto)
apply(auto simp: tape_of_nl_cons tape_of_nl_append_one split: if_splits)
apply(simp add: exp_ind[THEN sym])
done
lemma exp_1[simp]: "a\<up>(Suc 0) = [a]"
by(simp)
lemma tape_of_nl_cons_app1: "(<a # xs @ [b]>) = (Oc\<up>(Suc a) @ Bk # (<xs@ [b]>))"
apply(case_tac xs; simp add: tape_of_list_def tape_of_nat_list.simps tape_of_nat_def)
done
lemma bl_bin_bk_oc[simp]:
"bl_bin (xs @ [Bk, Oc]) =
bl_bin xs + 2*2^(length xs)"
apply(simp add: bin_wc_eq)
using bl2nat_cons_oc[of "xs @ [Bk]"]
apply(simp add: bl2nat_cons_bk bl2wc.simps)
done
lemma tape_of_nat[simp]: "(<a::nat>) = Oc\<up>(Suc a)"
apply(simp add: tape_of_nat_def)
done
lemma tape_of_nl_cons_app2: "(<c # xs @ [b]>) = (<c # xs> @ Bk # Oc\<up>(Suc b))"
proof(induct "length xs" arbitrary: xs c, simp add: tape_of_list_def)
fix x xs c
assume ind: "\<And>xs c. x = length xs \<Longrightarrow> <c # xs @ [b]> =
<c # xs> @ Bk # Oc\<up>(Suc b)"
and h: "Suc x = length (xs::nat list)"
show "<c # xs @ [b]> = <c # xs> @ Bk # Oc\<up>(Suc b)"
proof(cases xs, simp add: tape_of_list_def)
fix a list
assume g: "xs = a # list"
hence k: "<a # list @ [b]> = <a # list> @ Bk # Oc\<up>(Suc b)"
apply(rule_tac ind)
using h
apply(simp)
done
from g and k show "<c # xs @ [b]> = <c # xs> @ Bk # Oc\<up>(Suc b)"
apply(simp add: tape_of_list_def)
done
qed
qed
lemma length_2_elems[simp]: "length (<aa # a # list>) = Suc (Suc aa) + length (<a # list>)"
apply(simp add: tape_of_list_def)
done
lemma bl_bin_addition[simp]: "bl_bin (Oc\<up>(Suc aa) @ Bk # tape_of_nat_list (a # lista) @ [Bk, Oc]) =
bl_bin (Oc\<up>(Suc aa) @ Bk # tape_of_nat_list (a # lista)) +
2* 2^(length (Oc\<up>(Suc aa) @ Bk # tape_of_nat_list (a # lista)))"
using bl_bin_bk_oc[of "Oc\<up>(Suc aa) @ Bk # tape_of_nat_list (a # lista)"]
apply(simp)
done
declare replicate_Suc[simp del]
lemma bl_bin_2[simp]:
"bl_bin (<aa # list>) + (4 * rs + 4) * 2 ^ (length (<aa # list>) - Suc 0)
= bl_bin (Oc\<up>(Suc aa) @ Bk # <list @ [0]>) + rs * (2 * 2 ^ (aa + length (<list @ [0]>)))"
apply(case_tac "list", simp add: add_mult_distrib)
apply(simp add: tape_of_nl_cons_app2 add_mult_distrib)
apply(simp add: tape_of_list_def)
done
lemma tape_of_nl_app_Suc: "((<list @ [Suc ab]>)) = (<list @ [ab]>) @ [Oc]"
proof(induct list)
case (Cons a list)
then show ?case by(cases list;simp_all add:tape_of_list_def exp_ind)
qed (simp add: tape_of_list_def exp_ind)
lemma bl_bin_3[simp]: "bl_bin (Oc # Oc\<up>(aa) @ Bk # <list @ [ab]> @ [Oc])
= bl_bin (Oc # Oc\<up>(aa) @ Bk # <list @ [ab]>) +
2^(length (Oc # Oc\<up>(aa) @ Bk # <list @ [ab]>))"
apply(simp add: bin_wc_eq)
apply(simp add: bl2nat_cons_oc bl2wc.simps)
using bl2nat_cons_oc[of "Oc # Oc\<up>(aa) @ Bk # <list @ [ab]>"]
apply(simp)
done
lemma bl_bin_4[simp]: "bl_bin (Oc # Oc\<up>(aa) @ Bk # <list @ [ab]>) + (4 * 2 ^ (aa + length (<list @ [ab]>)) +
4 * (rs * 2 ^ (aa + length (<list @ [ab]>)))) =
bl_bin (Oc # Oc\<up>(aa) @ Bk # <list @ [Suc ab]>) +
rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>)))"
apply(simp add: tape_of_nl_app_Suc)
done
declare tape_of_nat[simp del]
fun wcode_double_case_inv :: "nat \<Rightarrow> bin_inv_t"
where
"wcode_double_case_inv st ires rs (l, r) =
(if st = Suc 0 then wcode_on_left_moving_1 ires rs (l, r)
else if st = Suc (Suc 0) then wcode_on_checking_1 ires rs (l, r)
else if st = 3 then wcode_erase1 ires rs (l, r)
else if st = 4 then wcode_on_right_moving_1 ires rs (l, r)
else if st = 5 then wcode_goon_right_moving_1 ires rs (l, r)
else if st = 6 then wcode_backto_standard_pos ires rs (l, r)
else if st = 13 then wcode_before_double ires rs (l, r)
else False)"
declare wcode_double_case_inv.simps[simp del]
fun wcode_double_case_state :: "config \<Rightarrow> nat"
where
"wcode_double_case_state (st, l, r) =
13 - st"
fun wcode_double_case_step :: "config \<Rightarrow> nat"
where
"wcode_double_case_step (st, l, r) =
(if st = Suc 0 then (length l)
else if st = Suc (Suc 0) then (length r)
else if st = 3 then
if hd r = Oc then 1 else 0
else if st = 4 then (length r)
else if st = 5 then (length r)
else if st = 6 then (length l)
else 0)"
fun wcode_double_case_measure :: "config \<Rightarrow> nat \<times> nat"
where
"wcode_double_case_measure (st, l, r) =
(wcode_double_case_state (st, l, r),
wcode_double_case_step (st, l, r))"
definition wcode_double_case_le :: "(config \<times> config) set"
where "wcode_double_case_le \<equiv> (inv_image lex_pair wcode_double_case_measure)"
lemma wf_lex_pair[intro]: "wf lex_pair"
by(auto intro:wf_lex_prod simp:lex_pair_def)
lemma wf_wcode_double_case_le[intro]: "wf wcode_double_case_le"
by(auto intro:wf_inv_image simp: wcode_double_case_le_def )
lemma fetch_t_wcode_main[simp]:
"fetch t_wcode_main (Suc 0) Bk = (L, Suc 0)"
"fetch t_wcode_main (Suc 0) Oc = (L, Suc (Suc 0))"
"fetch t_wcode_main (Suc (Suc 0)) Oc = (R, 3)"
"fetch t_wcode_main (Suc (Suc 0)) Bk = (L, 7)"
"fetch t_wcode_main (Suc (Suc (Suc 0))) Bk = (R, 4)"
"fetch t_wcode_main (Suc (Suc (Suc 0))) Oc = (W0, 3)"
"fetch t_wcode_main 4 Bk = (R, 4)"
"fetch t_wcode_main 4 Oc = (R, 5)"
"fetch t_wcode_main 5 Oc = (R, 5)"
"fetch t_wcode_main 5 Bk = (W1, 6)"
"fetch t_wcode_main 6 Bk = (R, 13)"
"fetch t_wcode_main 6 Oc = (L, 6)"
"fetch t_wcode_main 7 Oc = (R, 8)"
"fetch t_wcode_main 7 Bk = (R, 0)"
"fetch t_wcode_main 8 Bk = (R, 9)"
"fetch t_wcode_main 9 Bk = (R, 10)"
"fetch t_wcode_main 9 Oc = (W0, 9)"
"fetch t_wcode_main 10 Bk = (R, 10)"
"fetch t_wcode_main 10 Oc = (R, 11)"
"fetch t_wcode_main 11 Bk = (W1, 12)"
"fetch t_wcode_main 11 Oc = (R, 11)"
"fetch t_wcode_main 12 Oc = (L, 12)"
"fetch t_wcode_main 12 Bk = (R, t_twice_len + 14)"
by(auto simp: t_wcode_main_def t_wcode_main_first_part_def fetch.simps numeral)
declare wcode_on_checking_1.simps[simp del]
lemmas wcode_double_case_inv_simps =
wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps
wcode_on_left_moving_1_B.simps wcode_on_checking_1.simps
wcode_erase1.simps wcode_on_right_moving_1.simps
wcode_goon_right_moving_1.simps wcode_backto_standard_pos.simps
lemma wcode_on_left_moving_1[simp]:
"wcode_on_left_moving_1 ires rs (b, []) = False"
"wcode_on_left_moving_1 ires rs (b, r) \<Longrightarrow> b \<noteq> []"
by(auto simp: wcode_on_left_moving_1.simps wcode_on_left_moving_1_B.simps
wcode_on_left_moving_1_O.simps)
lemma wcode_on_left_moving_1E[elim]: "\<lbrakk>wcode_on_left_moving_1 ires rs (b, Bk # list);
tl b = aa \<and> hd b # Bk # list = ba\<rbrakk> \<Longrightarrow>
wcode_on_left_moving_1 ires rs (aa, ba)"
apply(simp only: wcode_on_left_moving_1.simps wcode_on_left_moving_1_O.simps
wcode_on_left_moving_1_B.simps)
apply(erule_tac disjE)
apply(erule_tac exE)+
apply(rename_tac ml mr rn)
apply(case_tac ml, simp)
apply(rule_tac x = "mr - Suc (Suc 0)" in exI, rule_tac x = rn in exI)
apply (smt One_nat_def Suc_diff_Suc append_Cons empty_replicate list.sel(3) neq0_conv replicate_Suc replicate_app_Cons_same tl_append2 tl_replicate)
apply(rule_tac disjI1)
apply (metis add_Suc_shift less_SucI list.exhaust_sel list.inject list.simps(3) replicate_Suc_iff_anywhere)
by simp
declare replicate_Suc[simp]
lemma wcode_on_moving_1_Elim[elim]:
"\<lbrakk>wcode_on_left_moving_1 ires rs (b, Oc # list); tl b = aa \<and> hd b # Oc # list = ba\<rbrakk>
\<Longrightarrow> wcode_on_checking_1 ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac disjE)
apply (metis cell.distinct(1) empty_replicate hd_append2 hd_replicate list.sel(1) not_gr_zero)
apply force.
lemma wcode_on_checking_1_Elim[elim]: "\<lbrakk>wcode_on_checking_1 ires rs (b, Oc # ba);Oc # b = aa \<and> list = ba\<rbrakk>
\<Longrightarrow> wcode_erase1 ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac exE)+ by auto
lemma wcode_on_checking_1_simp[simp]:
"wcode_on_checking_1 ires rs (b, []) = False"
"wcode_on_checking_1 ires rs (b, Bk # list) = False"
by(auto simp: wcode_double_case_inv_simps)
lemma wcode_erase1_nonempty_snd[simp]: "wcode_erase1 ires rs (b, []) = False"
apply(simp add: wcode_double_case_inv_simps)
done
lemma wcode_on_right_moving_1_nonempty_snd[simp]: "wcode_on_right_moving_1 ires rs (b, []) = False"
apply(simp add: wcode_double_case_inv_simps)
done
lemma wcode_on_right_moving_1_BkE[elim]:
"\<lbrakk>wcode_on_right_moving_1 ires rs (b, Bk # ba); Bk # b = aa \<and> list = b\<rbrakk> \<Longrightarrow>
wcode_on_right_moving_1 ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac exE)+
apply(rename_tac ml mr rn)
apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI,
rule_tac x = rn in exI)
apply(simp)
apply(case_tac mr, simp, simp)
done
lemma wcode_on_right_moving_1_OcE[elim]:
"\<lbrakk>wcode_on_right_moving_1 ires rs (b, Oc # ba); Oc # b = aa \<and> list = ba\<rbrakk>
\<Longrightarrow> wcode_goon_right_moving_1 ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps)
apply(erule_tac exE)+
apply(rename_tac ml mr rn)
apply(rule_tac x = "Suc 0" in exI, rule_tac x = "rs" in exI,
rule_tac x = "ml - Suc (Suc 0)" in exI, rule_tac x = rn in exI)
apply(case_tac mr, simp_all)
apply(case_tac ml, simp, case_tac nat, simp, simp)
done
lemma wcode_erase1_BkE[elim]:
assumes "wcode_erase1 ires rs (b, Bk # ba)" "Bk # b = aa \<and> list = ba" "c = Bk # ba"
shows "wcode_on_right_moving_1 ires rs (aa, ba)"
proof -
from assms obtain rn ln where "b = Oc # ires"
"tl (Bk # ba) = Bk \<up> ln @ Bk # Bk # Oc \<up> Suc rs @ Bk \<up> rn"
unfolding wcode_double_case_inv_simps by auto
thus ?thesis using assms(2-) unfolding wcode_double_case_inv_simps
apply(rule_tac x = "Suc 0" in exI, rule_tac x = "Suc (Suc ln)" in exI,
rule_tac x = rn in exI, simp add: exp_ind del: replicate_Suc)
done
qed
lemma wcode_erase1_OcE[elim]: "\<lbrakk>wcode_erase1 ires rs (aa, Oc # list); b = aa \<and> Bk # list = ba\<rbrakk> \<Longrightarrow>
wcode_erase1 ires rs (aa, ba)"
unfolding wcode_double_case_inv_simps
by auto auto
lemma wcode_goon_right_moving_1_emptyE[elim]:
assumes "wcode_goon_right_moving_1 ires rs (aa, [])" "b = aa \<and> [Oc] = ba"
shows "wcode_backto_standard_pos ires rs (aa, ba)"
proof -
from assms obtain ml ln rn mr where "aa = Oc \<up> ml @ Bk # Bk # Bk \<up> ln @ Oc # ires"
"[] = Oc \<up> mr @ Bk \<up> rn" "ml + mr = Suc rs"
by(auto simp:wcode_double_case_inv_simps)
thus ?thesis using assms(2)
apply(simp only: wcode_double_case_inv_simps)
apply(rule_tac disjI2)
apply(simp only:wcode_backto_standard_pos_O.simps)
apply(rule_tac x = ml in exI, rule_tac x = "Suc 0" in exI, rule_tac x = ln in exI,
rule_tac x = rn in exI, simp)
done
qed
lemma wcode_goon_right_moving_1_BkE[elim]:
assumes "wcode_goon_right_moving_1 ires rs (aa, Bk # list)" "b = aa \<and> Oc # list = ba"
shows "wcode_backto_standard_pos ires rs (aa, ba)"
proof -
from assms obtain ln rn where "aa = Oc \<up> Suc rs @ Bk \<up> Suc (Suc ln) @ Oc # ires"
"Bk # list = Bk \<up> rn" "b = Oc \<up> Suc rs @ Bk \<up> Suc (Suc ln) @ Oc # ires" "ba = Oc # list"
by(auto simp:wcode_double_case_inv_simps)
thus ?thesis using assms(2)
apply(simp only: wcode_double_case_inv_simps wcode_backto_standard_pos_O.simps)
apply(rule_tac disjI2)
apply(rule exI[of _ "Suc rs"], rule exI[of _ "Suc 0"], rule_tac x = ln in exI,
rule_tac x = "rn - Suc 0" in exI, simp)
apply(cases rn;auto)
done
qed
lemma wcode_goon_right_moving_1_OcE[elim]:
assumes "wcode_goon_right_moving_1 ires rs (b, Oc # ba)" "Oc # b = aa \<and> list = ba"
shows "wcode_goon_right_moving_1 ires rs (aa, ba)"
proof -
from assms obtain ml mr ln rn where
"b = Oc \<up> ml @ Bk # Bk # Bk \<up> ln @ Oc # ires \<and>
Oc # ba = Oc \<up> mr @ Bk \<up> rn \<and> ml + mr = Suc rs"
unfolding wcode_double_case_inv_simps by auto
with assms(2) show ?thesis unfolding wcode_double_case_inv_simps
apply(rule_tac x = "Suc ml" in exI, rule_tac x = "mr - Suc 0" in exI,
rule_tac x = ln in exI, rule_tac x = rn in exI)
apply(simp)
apply(case_tac mr, simp, case_tac rn, simp_all)
done
qed
lemma wcode_backto_standard_pos_BkE[elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, Bk # ba); Bk # b = aa \<and> list = ba\<rbrakk>
\<Longrightarrow> wcode_before_double ires rs (aa, ba)"
apply(simp only: wcode_double_case_inv_simps wcode_backto_standard_pos_B.simps
wcode_backto_standard_pos_O.simps wcode_before_double.simps)
apply(erule_tac disjE)
apply(erule_tac exE)+
by auto
lemma wcode_backto_standard_pos_no_Oc[simp]: "wcode_backto_standard_pos ires rs ([], Oc # list) = False"
apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
wcode_backto_standard_pos_O.simps)
done
lemma wcode_backto_standard_pos_nonempty_snd[simp]: "wcode_backto_standard_pos ires rs (b, []) = False"
apply(auto simp: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
wcode_backto_standard_pos_O.simps)
done
lemma wcode_backto_standard_pos_OcE[elim]: "\<lbrakk>wcode_backto_standard_pos ires rs (b, Oc # list); tl b = aa; hd b # Oc # list = ba\<rbrakk>
\<Longrightarrow> wcode_backto_standard_pos ires rs (aa, ba)"
apply(simp only: wcode_backto_standard_pos.simps wcode_backto_standard_pos_B.simps
wcode_backto_standard_pos_O.simps)
apply(erule_tac disjE)
apply(simp)
apply(erule_tac exE)+
apply(simp)
apply (rename_tac ml mr ln rn)
apply(case_tac ml)
apply(rule_tac disjI1, rule_tac conjI)
apply(rule_tac x = ln in exI, force, rule_tac x = rn in exI, force, force).
declare nth_of.simps[simp del] fetch.simps[simp del]
lemma wcode_double_case_first_correctness:
"let P = (\<lambda> (st, l, r). st = 13) in
let Q = (\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r)) in
let f = (\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp) in
\<exists> n .P (f n) \<and> Q (f (n::nat))"
proof -
let ?P = "(\<lambda> (st, l, r). st = 13)"
let ?Q = "(\<lambda> (st, l, r). wcode_double_case_inv st ires rs (l, r))"
let ?f = "(\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp)"
have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))"
proof(rule_tac halt_lemma2)
show "wf wcode_double_case_le"
by auto
next
show "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_double_case_le"
proof(rule_tac allI, case_tac "?f na", simp)
fix na a b c
show "a \<noteq> 13 \<and> wcode_double_case_inv a ires rs (b, c) \<longrightarrow>
(case step0 (a, b, c) t_wcode_main of (st, x) \<Rightarrow>
wcode_double_case_inv st ires rs x) \<and>
(step0 (a, b, c) t_wcode_main, a, b, c) \<in> wcode_double_case_le"
apply(rule_tac impI, simp add: wcode_double_case_inv.simps)
apply(auto split: if_splits simp: step.simps,
case_tac [!] c, simp_all, case_tac [!] "(c::cell list)!0")
apply(simp_all add: wcode_double_case_inv.simps wcode_double_case_le_def
lex_pair_def)
apply(auto split: if_splits)
done
qed
next
show "?Q (?f 0)"
apply(simp add: steps.simps wcode_double_case_inv.simps
wcode_on_left_moving_1.simps
wcode_on_left_moving_1_B.simps)
apply(rule_tac disjI1)
apply(rule_tac x = "Suc m" in exI, simp)
apply(rule_tac x = "Suc 0" in exI, simp)
done
next
show "\<not> ?P (?f 0)"
apply(simp add: steps.simps)
done
qed
thus "let P = \<lambda>(st, l, r). st = 13;
Q = \<lambda>(st, l, r). wcode_double_case_inv st ires rs (l, r);
f = steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main
in \<exists>n. P (f n) \<and> Q (f n)"
apply(simp)
done
qed
lemma tm_append_shift_append_steps:
"\<lbrakk>steps0 (st, l, r) tp stp = (st', l', r');
0 < st';
length tp1 mod 2 = 0
\<rbrakk>
\<Longrightarrow> steps0 (st + length tp1 div 2, l, r) (tp1 @ shift tp (length tp1 div 2) @ tp2) stp
= (st' + length tp1 div 2, l', r')"
proof -
assume h:
"steps0 (st, l, r) tp stp = (st', l', r')"
"0 < st'"
"length tp1 mod 2 = 0 "
from h have
"steps (st + length tp1 div 2, l, r) (tp1 @ shift tp (length tp1 div 2), 0) stp =
(st' + length tp1 div 2, l', r')"
by(rule_tac tm_append_second_steps_eq, simp_all)
then have "steps (st + length tp1 div 2, l, r) ((tp1 @ shift tp (length tp1 div 2)) @ tp2, 0) stp =
(st' + length tp1 div 2, l', r')"
using h
apply(rule_tac tm_append_first_steps_eq, simp_all)
done
thus "?thesis"
by simp
qed
declare start_of.simps[simp del]
lemma twice_lemma: "rec_exec rec_twice [rs] = 2*rs"
by(auto simp: rec_twice_def rec_exec.simps)
lemma t_twice_correct:
"\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(tm_of abc_twice @ shift (mopup (Suc 0)) ((length (tm_of abc_twice) div 2))) stp =
(0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
proof(case_tac "rec_ci rec_twice")
fix a b c
assume h: "rec_ci rec_twice = (a, b, c)"
have "\<exists>stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<up>(n)) (tm_of abc_twice @ shift (mopup (length [rs]))
(length (tm_of abc_twice) div 2)) stp = (0, Bk\<up>(m) @ Bk # Bk # ires, Oc\<up>(Suc (rec_exec rec_twice [rs])) @ Bk\<up>(l))"
thm recursive_compile_to_tm_correct1
proof(rule_tac recursive_compile_to_tm_correct1)
show "rec_ci rec_twice = (a, b, c)" by (simp add: h)
next
show "terminate rec_twice [rs]"
apply(rule_tac primerec_terminate, auto)
apply(simp add: rec_twice_def, auto simp: constn.simps numeral_2_eq_2)
by(auto)
next
show "tm_of abc_twice = tm_of (a [+] dummy_abc (length [rs]))"
using h
by(simp add: abc_twice_def)
qed
thus "?thesis"
apply(simp add: tape_of_list_def tape_of_nat_def rec_exec.simps twice_lemma)
done
qed
declare adjust.simps[simp]
lemma adjust_fetch0:
"\<lbrakk>0 < a; a \<le> length ap div 2; fetch ap a b = (aa, 0)\<rbrakk>
\<Longrightarrow> fetch (adjust0 ap) a b = (aa, Suc (length ap div 2))"
apply(case_tac b, auto simp: fetch.simps nth_of.simps nth_map
split: if_splits)
apply(case_tac [!] a, auto simp: fetch.simps nth_of.simps)
done
lemma adjust_fetch_norm:
"\<lbrakk>st > 0; st \<le> length tp div 2; fetch ap st b = (aa, ns); ns \<noteq> 0\<rbrakk>
\<Longrightarrow> fetch (adjust0 ap) st b = (aa, ns)"
apply(case_tac b, auto simp: fetch.simps nth_of.simps nth_map
split: if_splits)
apply(case_tac [!] st, auto simp: fetch.simps nth_of.simps)
done
declare adjust.simps[simp del]
lemma adjust_step_eq:
assumes exec: "step0 (st,l,r) ap = (st', l', r')"
and wf_tm: "tm_wf (ap, 0)"
and notfinal: "st' > 0"
shows "step0 (st, l, r) (adjust0 ap) = (st', l', r')"
using assms
proof -
have "st > 0"
using assms
by(case_tac st, simp_all add: step.simps fetch.simps)
moreover hence "st \<le> (length ap) div 2"
using assms
apply(case_tac "st \<le> (length ap) div 2", simp)
apply(case_tac st, auto simp: step.simps fetch.simps)
apply(case_tac "read r", simp_all add: fetch.simps
nth_of.simps adjust.simps tm_wf.simps split: if_splits)
apply(auto simp: mod_ex2)
done
ultimately have "fetch (adjust0 ap) st (read r) = fetch ap st (read r)"
using assms
apply(case_tac "fetch ap st (read r)")
apply(drule_tac adjust_fetch_norm, simp_all)
apply(simp add: step.simps)
done
thus "?thesis"
using exec
by(simp add: step.simps)
qed
declare adjust.simps[simp del]
lemma adjust_steps_eq:
assumes exec: "steps0 (st,l,r) ap stp = (st', l', r')"
and wf_tm: "tm_wf (ap, 0)"
and notfinal: "st' > 0"
shows "steps0 (st, l, r) (adjust0 ap) stp = (st', l', r')"
using exec notfinal
proof(induct stp arbitrary: st' l' r')
case 0
thus "?case"
by(simp add: steps.simps)
next
case (Suc stp st' l' r')
have ind: "\<And>st' l' r'. \<lbrakk>steps0 (st, l, r) ap stp = (st', l', r'); 0 < st'\<rbrakk>
\<Longrightarrow> steps0 (st, l, r) (adjust0 ap) stp = (st', l', r')" by fact
have h: "steps0 (st, l, r) ap (Suc stp) = (st', l', r')" by fact
have g: "0 < st'" by fact
obtain st'' l'' r'' where a: "steps0 (st, l, r) ap stp = (st'', l'', r'')"
by (metis prod_cases3)
hence c:"0 < st''"
using h g
apply(simp add: step_red)
apply(case_tac st'', auto)
done
hence b: "steps0 (st, l, r) (adjust0 ap) stp = (st'', l'', r'')"
using a
by(rule_tac ind, simp_all)
thus "?case"
using assms a b h g
apply(simp add: step_red)
apply(rule_tac adjust_step_eq, simp_all)
done
qed
lemma adjust_halt_eq:
assumes exec: "steps0 (1, l, r) ap stp = (0, l', r')"
and tm_wf: "tm_wf (ap, 0)"
shows "\<exists> stp. steps0 (Suc 0, l, r) (adjust0 ap) stp =
(Suc (length ap div 2), l', r')"
proof -
have "\<exists> stp. \<not> is_final (steps0 (1, l, r) ap stp) \<and> (steps0 (1, l, r) ap (Suc stp) = (0, l', r'))"
using exec
by(erule_tac before_final)
then obtain stpa where a:
"\<not> is_final (steps0 (1, l, r) ap stpa) \<and> (steps0 (1, l, r) ap (Suc stpa) = (0, l', r'))" ..
obtain sa la ra where b:"steps0 (1, l, r) ap stpa = (sa, la, ra)" by (metis prod_cases3)
hence c: "steps0 (Suc 0, l, r) (adjust0 ap) stpa = (sa, la, ra)"
using assms a
apply(rule_tac adjust_steps_eq, simp_all)
done
have d: "sa \<le> length ap div 2"
using steps_in_range[of "(l, r)" ap stpa] a tm_wf b
by(simp)
obtain ac ns where e: "fetch ap sa (read ra) = (ac, ns)"
by (metis prod.exhaust)
hence f: "ns = 0"
using b a
apply(simp add: step_red step.simps)
done
have k: "fetch (adjust0 ap) sa (read ra) = (ac, Suc (length ap div 2))"
using a b c d e f
apply(rule_tac adjust_fetch0, simp_all)
done
from a b e f k and c show "?thesis"
apply(rule_tac x = "Suc stpa" in exI)
apply(simp add: step_red, auto)
apply(simp add: step.simps)
done
qed
declare tm_wf.simps[simp del]
lemma tm_wf_t_twice_compile [simp]: "tm_wf (t_twice_compile, 0)"
apply(simp only: t_twice_compile_def)
apply(rule_tac wf_tm_from_abacus, simp)
done
lemma t_twice_change_term_state:
"\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_twice stp
= (Suc t_twice_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
proof -
have "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(tm_of abc_twice @ shift (mopup (Suc 0)) ((length (tm_of abc_twice) div 2))) stp =
(0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
by(rule_tac t_twice_correct)
then obtain stp ln rn where " steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(tm_of abc_twice @ shift (mopup (Suc 0)) ((length (tm_of abc_twice) div 2))) stp =
(0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" by blast
hence "\<exists> stp. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(adjust0 t_twice_compile) stp
= (Suc (length t_twice_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
apply(rule_tac stp = stp in adjust_halt_eq)
apply(simp add: t_twice_compile_def, auto)
done
then obtain stpb where
"steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(adjust0 t_twice_compile) stpb
= (Suc (length t_twice_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))" ..
thus "?thesis"
apply(simp add: t_twice_def t_twice_len_def)
by metis
qed
lemma length_t_wcode_main_first_part_even[intro]: "length t_wcode_main_first_part mod 2 = 0"
apply(auto simp: t_wcode_main_first_part_def)
done
lemma t_twice_append_pre:
"steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_twice stp
= (Suc t_twice_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))
\<Longrightarrow> steps0 (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @
([(L, 1), (L, 1)] @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp
= (Suc (t_twice_len) + length t_wcode_main_first_part div 2,
Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
by(rule_tac tm_append_shift_append_steps, auto)
lemma t_twice_append:
"\<exists> stp ln rn. steps0 (Suc 0 + length t_wcode_main_first_part div 2, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @
([(L, 1), (L, 1)] @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)])) stp
= (Suc (t_twice_len) + length t_wcode_main_first_part div 2, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
using t_twice_change_term_state[of ires rs n]
apply(erule_tac exE)
apply(erule_tac exE)
apply(erule_tac exE)
apply(drule_tac t_twice_append_pre)
apply(rename_tac stp ln rn)
apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
apply(simp)
done
lemma mopup_mod2: "length (mopup k) mod 2 = 0"
by(auto simp: mopup.simps)
lemma fetch_t_wcode_main_Oc[simp]: "fetch t_wcode_main (Suc (t_twice_len + length t_wcode_main_first_part div 2)) Oc
= (L, Suc 0)"
apply(subgoal_tac "length (t_twice) mod 2 = 0")
apply(simp add: t_wcode_main_def nth_append fetch.simps t_wcode_main_first_part_def
nth_of.simps t_twice_len_def, auto)
apply(simp add: t_twice_def t_twice_compile_def)
using mopup_mod2[of 1]
apply(simp)
done
lemma wcode_jump1:
"\<exists> stp ln rn. steps0 (Suc (t_twice_len) + length t_wcode_main_first_part div 2,
Bk\<up>(m) @ Bk # Bk # ires, Oc\<up>(Suc (2 * rs)) @ Bk\<up>(n))
t_wcode_main stp
= (Suc 0, Bk\<up>(ln) @ Bk # ires, Bk # Oc\<up>(Suc (2 * rs)) @ Bk\<up>(rn))"
apply(rule_tac x = "Suc 0" in exI, rule_tac x = "m" in exI, rule_tac x = n in exI)
apply(simp add: steps.simps step.simps exp_ind)
apply(case_tac m, simp_all)
apply(simp add: exp_ind[THEN sym])
done
lemma wcode_main_first_part_len[simp]:
"length t_wcode_main_first_part = 24"
apply(simp add: t_wcode_main_first_part_def)
done
lemma wcode_double_case:
shows "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rn))"
(is "\<exists>stp ln rn. ?tm stp ln rn")
proof -
from wcode_double_case_first_correctness[of ires rs m n] obtain na ln rn where
"steps0 (Suc 0, Bk # Bk \<up> m @ Oc # Oc # ires, Bk # Oc # Oc \<up> rs @ Bk \<up> n) t_wcode_main na
= (13, Bk # Bk # Bk \<up> ln @ Oc # ires, Oc # Oc # Oc \<up> rs @ Bk \<up> rn)"
by(auto simp: wcode_double_case_inv.simps wcode_before_double.simps)
hence "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(13, Bk # Bk # Bk\<up>(ln) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
by(case_tac "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires,
Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main na", auto)
from this obtain stpa lna rna where stp1:
"steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa =
(13, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rna))" by blast
from t_twice_append[of "Bk\<up>(lna) @ Oc # ires" "Suc rs" rna] obtain stp ln rn
where "steps0 (Suc 0 + length t_wcode_main_first_part div 2,
Bk # Bk # Bk \<up> lna @ Oc # ires, Oc \<up> Suc (Suc rs) @ Bk \<up> rna)
(t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @
[(L, 1), (L, 1)] @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)]) stp =
(Suc t_twice_len + length t_wcode_main_first_part div 2,
Bk \<up> ln @ Bk # Bk # Bk \<up> lna @ Oc # ires, Oc \<up> Suc (2 * Suc rs) @ Bk \<up> rn)" by blast
hence "\<exists> stp ln rn. steps0 (13, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rna)) t_wcode_main stp =
(13 + t_twice_len, Bk # Bk # Bk\<up>(ln) @ Oc # ires, Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rn))"
using t_twice_append[of "Bk\<up>(lna) @ Oc # ires" "Suc rs" rna]
apply(simp)
apply(rule_tac x = stp in exI, rule_tac x = "ln + lna" in exI,
rule_tac x = rn in exI)
apply(simp add: t_wcode_main_def)
apply(simp add: replicate_Suc[THEN sym] replicate_add [THEN sym] del: replicate_Suc)
done
from this obtain stpb lnb rnb where stp2:
"steps0 (13, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rna)) t_wcode_main stpb =
(13 + t_twice_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnb))" by blast
from wcode_jump1[of lnb "Oc # ires" "Suc rs" rnb] obtain stp ln rn where
"steps0 (Suc t_twice_len + length t_wcode_main_first_part div 2,
Bk \<up> lnb @ Bk # Bk # Oc # ires, Oc \<up> Suc (2 * Suc rs) @ Bk \<up> rnb) t_wcode_main stp =
(Suc 0, Bk \<up> ln @ Bk # Oc # ires, Bk # Oc \<up> Suc (2 * Suc rs) @ Bk \<up> rn)" by metis
hence "steps0 (13 + t_twice_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires,
Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnb)) t_wcode_main stp =
(Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rn))"
apply(auto simp add: t_wcode_main_def)
apply(subgoal_tac "Bk\<up>(lnb) @ Bk # Bk # Oc # ires = Bk # Bk # Bk\<up>(lnb) @ Oc # ires", simp)
apply(simp add: replicate_Suc[THEN sym] exp_ind[THEN sym] del: replicate_Suc)
apply(simp)
apply(simp add: replicate_Suc[THEN sym] exp_ind del: replicate_Suc)
done
hence "\<exists>stp ln rn. steps0 (13 + t_twice_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires,
Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnb)) t_wcode_main stp =
(Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rn))"
by blast
from this obtain stpc lnc rnc where stp3:
"steps0 (13 + t_twice_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires,
Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnb)) t_wcode_main stpc =
(Suc 0, Bk # Bk\<up>(lnc) @ Oc # ires, Bk # Oc\<up>(Suc (Suc (Suc (2 *rs)))) @ Bk\<up>(rnc))"
by blast
from stp1 stp2 stp3 have "?tm (stpa + stpb + stpc) lnc rnc" by simp
thus "?thesis" by blast
qed
(* Begin: fourtime_case*)
fun wcode_on_left_moving_2_B :: "bin_inv_t"
where
"wcode_on_left_moving_2_B ires rs (l, r) =
(\<exists> ml mr rn. l = Bk\<up>(ml) @ Oc # Bk # Oc # ires \<and>
r = Bk\<up>(mr) @ Oc\<up>(Suc rs) @ Bk\<up>(rn) \<and>
ml + mr > Suc 0 \<and> mr > 0)"
fun wcode_on_left_moving_2_O :: "bin_inv_t"
where
"wcode_on_left_moving_2_O ires rs (l, r) =
(\<exists> ln rn. l = Bk # Oc # ires \<and>
r = Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
fun wcode_on_left_moving_2 :: "bin_inv_t"
where
"wcode_on_left_moving_2 ires rs (l, r) =
(wcode_on_left_moving_2_B ires rs (l, r) \<or>
wcode_on_left_moving_2_O ires rs (l, r))"
fun wcode_on_checking_2 :: "bin_inv_t"
where
"wcode_on_checking_2 ires rs (l, r) =
(\<exists> ln rn. l = Oc#ires \<and>
r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
fun wcode_goon_checking :: "bin_inv_t"
where
"wcode_goon_checking ires rs (l, r) =
(\<exists> ln rn. l = ires \<and>
r = Oc # Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
fun wcode_right_move :: "bin_inv_t"
where
"wcode_right_move ires rs (l, r) =
(\<exists> ln rn. l = Oc # ires \<and>
r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
fun wcode_erase2 :: "bin_inv_t"
where
"wcode_erase2 ires rs (l, r) =
(\<exists> ln rn. l = Bk # Oc # ires \<and>
tl r = Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
fun wcode_on_right_moving_2 :: "bin_inv_t"
where
"wcode_on_right_moving_2 ires rs (l, r) =
(\<exists> ml mr rn. l = Bk\<up>(ml) @ Oc # ires \<and>
r = Bk\<up>(mr) @ Oc\<up>(Suc rs) @ Bk\<up>(rn) \<and> ml + mr > Suc 0)"
fun wcode_goon_right_moving_2 :: "bin_inv_t"
where
"wcode_goon_right_moving_2 ires rs (l, r) =
(\<exists> ml mr ln rn. l = Oc\<up>(ml) @ Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and>
r = Oc\<up>(mr) @ Bk\<up>(rn) \<and> ml + mr = Suc rs)"
fun wcode_backto_standard_pos_2_B :: "bin_inv_t"
where
"wcode_backto_standard_pos_2_B ires rs (l, r) =
(\<exists> ln rn. l = Bk # Bk\<up>(ln) @ Oc # ires \<and>
r = Bk # Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
fun wcode_backto_standard_pos_2_O :: "bin_inv_t"
where
"wcode_backto_standard_pos_2_O ires rs (l, r) =
(\<exists> ml mr ln rn. l = Oc\<up>(ml )@ Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and>
r = Oc\<up>(mr) @ Bk\<up>(rn) \<and>
ml + mr = (Suc (Suc rs)) \<and> mr > 0)"
fun wcode_backto_standard_pos_2 :: "bin_inv_t"
where
"wcode_backto_standard_pos_2 ires rs (l, r) =
(wcode_backto_standard_pos_2_O ires rs (l, r) \<or>
wcode_backto_standard_pos_2_B ires rs (l, r))"
fun wcode_before_fourtimes :: "bin_inv_t"
where
"wcode_before_fourtimes ires rs (l, r) =
(\<exists> ln rn. l = Bk # Bk # Bk\<up>(ln) @ Oc # ires \<and>
r = Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
declare wcode_on_left_moving_2_B.simps[simp del] wcode_on_left_moving_2.simps[simp del]
wcode_on_left_moving_2_O.simps[simp del] wcode_on_checking_2.simps[simp del]
wcode_goon_checking.simps[simp del] wcode_right_move.simps[simp del]
wcode_erase2.simps[simp del]
wcode_on_right_moving_2.simps[simp del] wcode_goon_right_moving_2.simps[simp del]
wcode_backto_standard_pos_2_B.simps[simp del] wcode_backto_standard_pos_2_O.simps[simp del]
wcode_backto_standard_pos_2.simps[simp del]
lemmas wcode_fourtimes_invs =
wcode_on_left_moving_2_B.simps wcode_on_left_moving_2.simps
wcode_on_left_moving_2_O.simps wcode_on_checking_2.simps
wcode_goon_checking.simps wcode_right_move.simps
wcode_erase2.simps
wcode_on_right_moving_2.simps wcode_goon_right_moving_2.simps
wcode_backto_standard_pos_2_B.simps wcode_backto_standard_pos_2_O.simps
wcode_backto_standard_pos_2.simps
fun wcode_fourtimes_case_inv :: "nat \<Rightarrow> bin_inv_t"
where
"wcode_fourtimes_case_inv st ires rs (l, r) =
(if st = Suc 0 then wcode_on_left_moving_2 ires rs (l, r)
else if st = Suc (Suc 0) then wcode_on_checking_2 ires rs (l, r)
else if st = 7 then wcode_goon_checking ires rs (l, r)
else if st = 8 then wcode_right_move ires rs (l, r)
else if st = 9 then wcode_erase2 ires rs (l, r)
else if st = 10 then wcode_on_right_moving_2 ires rs (l, r)
else if st = 11 then wcode_goon_right_moving_2 ires rs (l, r)
else if st = 12 then wcode_backto_standard_pos_2 ires rs (l, r)
else if st = t_twice_len + 14 then wcode_before_fourtimes ires rs (l, r)
else False)"
declare wcode_fourtimes_case_inv.simps[simp del]
fun wcode_fourtimes_case_state :: "config \<Rightarrow> nat"
where
"wcode_fourtimes_case_state (st, l, r) = 13 - st"
fun wcode_fourtimes_case_step :: "config \<Rightarrow> nat"
where
"wcode_fourtimes_case_step (st, l, r) =
(if st = Suc 0 then length l
else if st = 9 then
(if hd r = Oc then 1
else 0)
else if st = 10 then length r
else if st = 11 then length r
else if st = 12 then length l
else 0)"
fun wcode_fourtimes_case_measure :: "config \<Rightarrow> nat \<times> nat"
where
"wcode_fourtimes_case_measure (st, l, r) =
(wcode_fourtimes_case_state (st, l, r),
wcode_fourtimes_case_step (st, l, r))"
definition wcode_fourtimes_case_le :: "(config \<times> config) set"
where "wcode_fourtimes_case_le \<equiv> (inv_image lex_pair wcode_fourtimes_case_measure)"
lemma wf_wcode_fourtimes_case_le[intro]: "wf wcode_fourtimes_case_le"
by(auto simp: wcode_fourtimes_case_le_def)
lemma nonempty_snd [simp]:
"wcode_on_left_moving_2 ires rs (b, []) = False"
"wcode_on_checking_2 ires rs (b, []) = False"
"wcode_goon_checking ires rs (b, []) = False"
"wcode_right_move ires rs (b, []) = False"
"wcode_erase2 ires rs (b, []) = False"
"wcode_on_right_moving_2 ires rs (b, []) = False"
"wcode_backto_standard_pos_2 ires rs (b, []) = False"
"wcode_on_checking_2 ires rs (b, Oc # list) = False"
by(auto simp: wcode_fourtimes_invs)
lemma wcode_on_left_moving_2[simp]:
"wcode_on_left_moving_2 ires rs (b, Bk # list) \<Longrightarrow> wcode_on_left_moving_2 ires rs (tl b, hd b # Bk # list)"
apply(simp only: wcode_fourtimes_invs)
apply(erule_tac disjE)
apply(erule_tac exE)+
apply(simp add: gr1_conv_Suc exp_ind replicate_app_Cons_same split:hd_repeat_cases)
apply (auto simp add: gr0_conv_Suc[symmetric] replicate_app_Cons_same split:hd_repeat_cases)
by force+
lemma wcode_goon_checking_via_2 [simp]: "wcode_on_checking_2 ires rs (b, Bk # list)
\<Longrightarrow> wcode_goon_checking ires rs (tl b, hd b # Bk # list)"
unfolding wcode_fourtimes_invs by auto
lemma wcode_erase2_via_move [simp]: "wcode_right_move ires rs (b, Bk # list) \<Longrightarrow> wcode_erase2 ires rs (Bk # b, list)"
by (auto simp:wcode_fourtimes_invs ) auto
lemma wcode_on_right_moving_2_via_erase2[simp]:
"wcode_erase2 ires rs (b, Bk # list) \<Longrightarrow> wcode_on_right_moving_2 ires rs (Bk # b, list)"
apply(auto simp:wcode_fourtimes_invs )
apply(rule_tac x = "Suc (Suc 0)" in exI, simp add: exp_ind)
by (metis replicate_Suc_iff_anywhere replicate_app_Cons_same)
lemma wcode_on_right_moving_2_move_Bk[simp]: "wcode_on_right_moving_2 ires rs (b, Bk # list)
\<Longrightarrow> wcode_on_right_moving_2 ires rs (Bk # b, list)"
apply(auto simp: wcode_fourtimes_invs) apply(rename_tac ml mr rn)
apply(rule_tac x = "Suc ml" in exI, simp)
apply(rule_tac x = "mr - 1" in exI, case_tac mr,auto)
done
lemma wcode_backto_standard_pos_2_via_right[simp]:
"wcode_goon_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow>
wcode_backto_standard_pos_2 ires rs (b, Oc # list)"
apply(simp add: wcode_fourtimes_invs, auto)
by (metis add.right_neutral add_Suc_shift append_Cons list.sel(3)
replicate.simps(1) replicate_Suc replicate_Suc_iff_anywhere self_append_conv2
tl_replicate zero_less_Suc)
lemma wcode_on_checking_2_via_left[simp]: "wcode_on_left_moving_2 ires rs (b, Oc # list) \<Longrightarrow>
wcode_on_checking_2 ires rs (tl b, hd b # Oc # list)"
by(auto simp: wcode_fourtimes_invs)
lemma wcode_backto_standard_pos_2_empty_via_right[simp]:
"wcode_goon_right_moving_2 ires rs (b, []) \<Longrightarrow>
wcode_backto_standard_pos_2 ires rs (b, [Oc])"
- apply(simp only: wcode_fourtimes_invs)
- apply(erule_tac exE)+
- by(rule_tac disjI1,auto)
+ by (auto simp add: wcode_fourtimes_invs) force
lemma wcode_goon_checking_cases[simp]: "wcode_goon_checking ires rs (b, Oc # list) \<Longrightarrow>
(b = [] \<longrightarrow> wcode_right_move ires rs ([Oc], list)) \<and>
(b \<noteq> [] \<longrightarrow> wcode_right_move ires rs (Oc # b, list))"
apply(simp only: wcode_fourtimes_invs)
apply(erule_tac exE)+
apply(auto)
done
lemma wcode_right_move_no_Oc[simp]: "wcode_right_move ires rs (b, Oc # list) = False"
apply(auto simp: wcode_fourtimes_invs)
done
lemma wcode_erase2_Bk_via_Oc[simp]: "wcode_erase2 ires rs (b, Oc # list)
\<Longrightarrow> wcode_erase2 ires rs (b, Bk # list)"
apply(auto simp: wcode_fourtimes_invs)
done
lemma wcode_goon_right_moving_2_Oc_move[simp]:
"wcode_on_right_moving_2 ires rs (b, Oc # list)
\<Longrightarrow> wcode_goon_right_moving_2 ires rs (Oc # b, list)"
apply(auto simp: wcode_fourtimes_invs)
apply(rule_tac x = "Suc 0" in exI, auto)
apply(rule_tac x = "ml - 2" in exI)
apply(case_tac ml, simp, case_tac "ml - 1", simp_all)
done
lemma wcode_backto_standard_pos_2_exists[simp]: "wcode_backto_standard_pos_2 ires rs (b, Bk # list)
\<Longrightarrow> (\<exists>ln. b = Bk # Bk\<up>(ln) @ Oc # ires) \<and> (\<exists>rn. list = Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
by(simp add: wcode_fourtimes_invs)
lemma wcode_goon_right_moving_2_move_Oc[simp]: "wcode_goon_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow>
wcode_goon_right_moving_2 ires rs (Oc # b, list)"
apply(simp only:wcode_fourtimes_invs, auto)
apply(rename_tac ml ln mr rn)
apply(case_tac mr;force)
done
lemma wcode_backto_standard_pos_2_Oc_mv_hd[simp]:
"wcode_backto_standard_pos_2 ires rs (b, Oc # list)
\<Longrightarrow> wcode_backto_standard_pos_2 ires rs (tl b, hd b # Oc # list)"
apply(simp only: wcode_fourtimes_invs)
apply(erule_tac disjE)
apply(erule_tac exE)+ apply(rename_tac ml mr ln rn)
by (case_tac ml, force,force,force)
lemma nonempty_fst[simp]:
"wcode_on_left_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
"wcode_on_checking_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
"wcode_goon_checking ires rs (b, Bk # list) = False"
"wcode_right_move ires rs (b, Bk # list) \<Longrightarrow> b\<noteq> []"
"wcode_erase2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
"wcode_on_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
"wcode_goon_right_moving_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
"wcode_backto_standard_pos_2 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
"wcode_on_left_moving_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
"wcode_goon_right_moving_2 ires rs (b, []) \<Longrightarrow> b \<noteq> []"
"wcode_erase2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
"wcode_on_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
"wcode_goon_right_moving_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
"wcode_backto_standard_pos_2 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
by(auto simp: wcode_fourtimes_invs)
lemma wcode_fourtimes_case_first_correctness:
shows "let P = (\<lambda> (st, l, r). st = t_twice_len + 14) in
let Q = (\<lambda> (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r)) in
let f = (\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp) in
\<exists> n .P (f n) \<and> Q (f (n::nat))"
proof -
let ?P = "(\<lambda> (st, l, r). st = t_twice_len + 14)"
let ?Q = "(\<lambda> (st, l, r). wcode_fourtimes_case_inv st ires rs (l, r))"
let ?f = "(\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp)"
have "\<exists> n . ?P (?f n) \<and> ?Q (?f (n::nat))"
proof(rule_tac halt_lemma2)
show "wf wcode_fourtimes_case_le"
by auto
next
have "\<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_fourtimes_case_le" for na
apply(cases "?f na", rule_tac impI)
apply(simp add: step_red step.simps)
apply(case_tac "snd (snd (?f na))", simp, case_tac [2] "hd (snd (snd (?f na)))", simp_all)
apply(simp_all add: wcode_fourtimes_case_inv.simps
wcode_fourtimes_case_le_def lex_pair_def split: if_splits)
by(auto simp: wcode_backto_standard_pos_2.simps wcode_backto_standard_pos_2_O.simps
wcode_backto_standard_pos_2_B.simps gr0_conv_Suc)
thus "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_fourtimes_case_le" by auto
next
show "?Q (?f 0)"
apply(simp add: steps.simps wcode_fourtimes_case_inv.simps)
apply(simp add: wcode_on_left_moving_2.simps wcode_on_left_moving_2_B.simps
wcode_on_left_moving_2_O.simps)
apply(rule_tac x = "Suc m" in exI, simp )
apply(rule_tac x ="Suc 0" in exI, auto)
done
next
show "\<not> ?P (?f 0)"
apply(simp add: steps.simps)
done
qed
thus "?thesis"
apply(erule_tac exE, simp)
done
qed
definition t_fourtimes_len :: "nat"
where
"t_fourtimes_len = (length t_fourtimes div 2)"
lemma primerec_rec_fourtimes_1[intro]: "primerec rec_fourtimes (Suc 0)"
apply(auto simp: rec_fourtimes_def numeral_4_eq_4 constn.simps)
by auto
lemma fourtimes_lemma: "rec_exec rec_fourtimes [rs] = 4 * rs"
by(simp add: rec_exec.simps rec_fourtimes_def)
lemma t_fourtimes_correct:
"\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(tm_of abc_fourtimes @ shift (mopup 1) (length (tm_of abc_fourtimes) div 2)) stp =
(0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
proof(case_tac "rec_ci rec_fourtimes")
fix a b c
assume h: "rec_ci rec_fourtimes = (a, b, c)"
have "\<exists>stp m l. steps0 (Suc 0, Bk # Bk # ires, <[rs]> @ Bk\<up>(n)) (tm_of abc_fourtimes @ shift (mopup (length [rs]))
(length (tm_of abc_fourtimes) div 2)) stp = (0, Bk\<up>(m) @ Bk # Bk # ires, Oc\<up>(Suc (rec_exec rec_fourtimes [rs])) @ Bk\<up>(l))"
thm recursive_compile_to_tm_correct1
proof(rule_tac recursive_compile_to_tm_correct1)
show "rec_ci rec_fourtimes = (a, b, c)" by (simp add: h)
next
show "terminate rec_fourtimes [rs]"
apply(rule_tac primerec_terminate)
by auto
next
show "tm_of abc_fourtimes = tm_of (a [+] dummy_abc (length [rs]))"
using h
by(simp add: abc_fourtimes_def)
qed
thus "?thesis"
apply(simp add: tape_of_list_def tape_of_nat_def fourtimes_lemma)
done
qed
lemma wf_fourtimes[intro]: "tm_wf (t_fourtimes_compile, 0)"
apply(simp only: t_fourtimes_compile_def)
apply(rule_tac wf_tm_from_abacus, simp)
done
lemma t_fourtimes_change_term_state:
"\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_fourtimes stp
= (Suc t_fourtimes_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
proof -
have "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(tm_of abc_fourtimes @ shift (mopup 1) ((length (tm_of abc_fourtimes) div 2))) stp =
(0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
by(rule_tac t_fourtimes_correct)
then obtain stp ln rn where
"steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(tm_of abc_fourtimes @ shift (mopup 1) ((length (tm_of abc_fourtimes) div 2))) stp =
(0, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" by blast
hence "\<exists> stp. steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(adjust0 t_fourtimes_compile) stp
= (Suc (length t_fourtimes_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
apply(rule_tac stp = stp in adjust_halt_eq)
apply(simp add: t_fourtimes_compile_def, auto)
done
then obtain stpb where
"steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
(adjust0 t_fourtimes_compile) stpb
= (Suc (length t_fourtimes_compile div 2), Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))" ..
thus "?thesis"
apply(simp add: t_fourtimes_def t_fourtimes_len_def)
by metis
qed
lemma length_t_twice_even[intro]: "is_even (length t_twice)"
by(auto simp: t_twice_def t_twice_compile_def intro!:mopup_mod2)
lemma t_fourtimes_append_pre:
"steps0 (Suc 0, Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n)) t_fourtimes stp
= (Suc t_fourtimes_len, Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))
\<Longrightarrow> steps0 (Suc 0 + length (t_wcode_main_first_part @
shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2,
Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
((t_wcode_main_first_part @
shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) @
shift t_fourtimes (length (t_wcode_main_first_part @
shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2) @ ([(L, 1), (L, 1)])) stp
= ((Suc t_fourtimes_len) + length (t_wcode_main_first_part @
shift t_twice (length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2,
Bk\<up>(ln) @ Bk # Bk # ires, Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
using length_t_twice_even
by(intro tm_append_shift_append_steps, auto)
lemma split_26_even[simp]: "(26 + l::nat) div 2 = l div 2 + 13" by(simp)
lemma t_twice_len_plust_14[simp]: "t_twice_len + 14 = 14 + length (shift t_twice 12) div 2"
apply(simp add: t_twice_def t_twice_len_def)
done
lemma t_fourtimes_append:
"\<exists> stp ln rn.
steps0 (Suc 0 + length (t_wcode_main_first_part @ shift t_twice
(length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2,
Bk # Bk # ires, Oc\<up>(Suc rs) @ Bk\<up>(n))
((t_wcode_main_first_part @ shift t_twice (length t_wcode_main_first_part div 2) @
[(L, 1), (L, 1)]) @ shift t_fourtimes (t_twice_len + 13) @ [(L, 1), (L, 1)]) stp
= (Suc t_fourtimes_len + length (t_wcode_main_first_part @ shift t_twice
(length t_wcode_main_first_part div 2) @ [(L, 1), (L, 1)]) div 2, Bk\<up>(ln) @ Bk # Bk # ires,
Oc\<up>(Suc (4 * rs)) @ Bk\<up>(rn))"
using t_fourtimes_change_term_state[of ires rs n]
apply(erule_tac exE)
apply(erule_tac exE)
apply(erule_tac exE)
apply(drule_tac t_fourtimes_append_pre)
apply(rule_tac x = stp in exI, rule_tac x = ln in exI, rule_tac x = rn in exI)
apply(simp add: t_twice_len_def)
done
lemma even_fourtimes_len: "length t_fourtimes mod 2 = 0"
apply(auto simp: t_fourtimes_def t_fourtimes_compile_def)
by (metis mopup_mod2)
lemma t_twice_even[simp]: "2 * (length t_twice div 2) = length t_twice"
using length_t_twice_even by arith
lemma t_fourtimes_even[simp]: "2 * (length t_fourtimes div 2) = length t_fourtimes"
using even_fourtimes_len
by arith
lemma fetch_t_wcode_14_Oc: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) Oc
= (L, Suc 0)"
apply(subgoal_tac "14 = Suc 13")
apply(simp only: fetch.simps add_Suc nth_of.simps t_wcode_main_def)
apply(simp add:length_t_twice_even t_fourtimes_len_def nth_append)
by arith
lemma fetch_t_wcode_14_Bk: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) Bk
= (L, Suc 0)"
apply(subgoal_tac "14 = Suc 13")
apply(simp only: fetch.simps add_Suc nth_of.simps t_wcode_main_def)
apply(simp add:length_t_twice_even t_fourtimes_len_def nth_append)
by arith
lemma fetch_t_wcode_14 [simp]: "fetch t_wcode_main (14 + length t_twice div 2 + t_fourtimes_len) b
= (L, Suc 0)"
apply(case_tac b, simp_all add:fetch_t_wcode_14_Bk fetch_t_wcode_14_Oc)
done
lemma wcode_jump2:
"\<exists> stp ln rn. steps0 (t_twice_len + 14 + t_fourtimes_len
, Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (4 * rs + 4)) @ Bk\<up>(rnb)) t_wcode_main stp =
(Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (4 * rs + 4)) @ Bk\<up>(rn))"
apply(rule_tac x = "Suc 0" in exI)
apply(simp add: steps.simps)
apply(rule_tac x = lnb in exI, rule_tac x = rnb in exI)
apply(simp add: step.simps)
done
lemma wcode_fourtimes_case:
shows "\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))"
proof -
have "\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(t_twice_len + 14, Bk # Bk # Bk\<up>(ln) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rn))"
using wcode_fourtimes_case_first_correctness[of ires rs m n]
by (auto simp add: wcode_fourtimes_case_inv.simps) auto
from this obtain stpa lna rna where stp1:
"steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Oc # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa =
(t_twice_len + 14, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rna))" by blast
have "\<exists>stp ln rn. steps0 (t_twice_len + 14, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rna))
t_wcode_main stp =
(t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<up>(ln) @ Oc # ires, Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))"
using t_fourtimes_append[of " Bk\<up>(lna) @ Oc # ires" "rs + 1" rna]
apply(erule_tac exE)
apply(erule_tac exE)
apply(erule_tac exE)
apply(simp add: t_wcode_main_def) apply(rename_tac stp ln rn)
apply(rule_tac x = stp in exI,
rule_tac x = "ln + lna" in exI,
rule_tac x = rn in exI, simp)
apply(simp add: replicate_Suc[THEN sym] replicate_add[THEN sym] del: replicate_Suc)
done
from this obtain stpb lnb rnb where stp2:
"steps0 (t_twice_len + 14, Bk # Bk # Bk\<up>(lna) @ Oc # ires, Oc\<up>(Suc (rs + 1)) @ Bk\<up>(rna))
t_wcode_main stpb =
(t_twice_len + 14 + t_fourtimes_len, Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnb))"
by blast
have "\<exists>stp ln rn. steps0 (t_twice_len + 14 + t_fourtimes_len,
Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnb))
t_wcode_main stp =
(Suc 0, Bk # Bk\<up>(ln) @ Oc # ires, Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))"
apply(rule wcode_jump2)
done
from this obtain stpc lnc rnc where stp3:
"steps0 (t_twice_len + 14 + t_fourtimes_len,
Bk # Bk # Bk\<up>(lnb) @ Oc # ires, Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnb))
t_wcode_main stpc =
(Suc 0, Bk # Bk\<up>(lnc) @ Oc # ires, Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rnc))"
by blast
from stp1 stp2 stp3 show "?thesis"
apply(rule_tac x = "stpa + stpb + stpc" in exI,
rule_tac x = lnc in exI, rule_tac x = rnc in exI)
apply(simp)
done
qed
fun wcode_on_left_moving_3_B :: "bin_inv_t"
where
"wcode_on_left_moving_3_B ires rs (l, r) =
(\<exists> ml mr rn. l = Bk\<up>(ml) @ Oc # Bk # Bk # ires \<and>
r = Bk\<up>(mr) @ Oc\<up>(Suc rs) @ Bk\<up>(rn) \<and>
ml + mr > Suc 0 \<and> mr > 0 )"
fun wcode_on_left_moving_3_O :: "bin_inv_t"
where
"wcode_on_left_moving_3_O ires rs (l, r) =
(\<exists> ln rn. l = Bk # Bk # ires \<and>
r = Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
fun wcode_on_left_moving_3 :: "bin_inv_t"
where
"wcode_on_left_moving_3 ires rs (l, r) =
(wcode_on_left_moving_3_B ires rs (l, r) \<or>
wcode_on_left_moving_3_O ires rs (l, r))"
fun wcode_on_checking_3 :: "bin_inv_t"
where
"wcode_on_checking_3 ires rs (l, r) =
(\<exists> ln rn. l = Bk # ires \<and>
r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
fun wcode_goon_checking_3 :: "bin_inv_t"
where
"wcode_goon_checking_3 ires rs (l, r) =
(\<exists> ln rn. l = ires \<and>
r = Bk # Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
fun wcode_stop :: "bin_inv_t"
where
"wcode_stop ires rs (l, r) =
(\<exists> ln rn. l = Bk # ires \<and>
r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
fun wcode_halt_case_inv :: "nat \<Rightarrow> bin_inv_t"
where
"wcode_halt_case_inv st ires rs (l, r) =
(if st = 0 then wcode_stop ires rs (l, r)
else if st = Suc 0 then wcode_on_left_moving_3 ires rs (l, r)
else if st = Suc (Suc 0) then wcode_on_checking_3 ires rs (l, r)
else if st = 7 then wcode_goon_checking_3 ires rs (l, r)
else False)"
fun wcode_halt_case_state :: "config \<Rightarrow> nat"
where
"wcode_halt_case_state (st, l, r) =
(if st = 1 then 5
else if st = Suc (Suc 0) then 4
else if st = 7 then 3
else 0)"
fun wcode_halt_case_step :: "config \<Rightarrow> nat"
where
"wcode_halt_case_step (st, l, r) =
(if st = 1 then length l
else 0)"
fun wcode_halt_case_measure :: "config \<Rightarrow> nat \<times> nat"
where
"wcode_halt_case_measure (st, l, r) =
(wcode_halt_case_state (st, l, r),
wcode_halt_case_step (st, l, r))"
definition wcode_halt_case_le :: "(config \<times> config) set"
where "wcode_halt_case_le \<equiv> (inv_image lex_pair wcode_halt_case_measure)"
lemma wf_wcode_halt_case_le[intro]: "wf wcode_halt_case_le"
by(auto simp: wcode_halt_case_le_def)
declare wcode_on_left_moving_3_B.simps[simp del] wcode_on_left_moving_3_O.simps[simp del]
wcode_on_checking_3.simps[simp del] wcode_goon_checking_3.simps[simp del]
wcode_on_left_moving_3.simps[simp del] wcode_stop.simps[simp del]
lemmas wcode_halt_invs =
wcode_on_left_moving_3_B.simps wcode_on_left_moving_3_O.simps
wcode_on_checking_3.simps wcode_goon_checking_3.simps
wcode_on_left_moving_3.simps wcode_stop.simps
lemma wcode_on_left_moving_3_mv_Bk[simp]: "wcode_on_left_moving_3 ires rs (b, Bk # list)
\<Longrightarrow> wcode_on_left_moving_3 ires rs (tl b, hd b # Bk # list)"
apply(simp only: wcode_halt_invs)
apply(erule_tac disjE)
apply(erule_tac exE)+ apply(rename_tac ml mr rn)
apply(case_tac ml, simp)
apply(rule_tac x = "mr - 2" in exI, rule_tac x = rn in exI)
apply(case_tac mr, force, simp add: exp_ind del: replicate_Suc)
apply(case_tac "mr - 1", force, simp add: exp_ind del: replicate_Suc)
apply force
apply force
done
lemma wcode_goon_checking_3_cases[simp]: "wcode_goon_checking_3 ires rs (b, Bk # list) \<Longrightarrow>
(b = [] \<longrightarrow> wcode_stop ires rs ([Bk], list)) \<and>
(b \<noteq> [] \<longrightarrow> wcode_stop ires rs (Bk # b, list))"
apply(auto simp: wcode_halt_invs)
done
lemma wcode_on_checking_3_mv_Oc[simp]: "wcode_on_left_moving_3 ires rs (b, Oc # list) \<Longrightarrow>
wcode_on_checking_3 ires rs (tl b, hd b # Oc # list)"
by(simp add:wcode_halt_invs)
lemma wcode_3_nonempty[simp]:
"wcode_on_left_moving_3 ires rs (b, []) = False"
"wcode_on_checking_3 ires rs (b, []) = False"
"wcode_goon_checking_3 ires rs (b, []) = False"
"wcode_on_left_moving_3 ires rs (b, Oc # list) \<Longrightarrow> b \<noteq> []"
"wcode_on_checking_3 ires rs (b, Oc # list) = False"
"wcode_on_left_moving_3 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
"wcode_on_checking_3 ires rs (b, Bk # list) \<Longrightarrow> b \<noteq> []"
"wcode_goon_checking_3 ires rs (b, Oc # list) = False"
by(auto simp: wcode_halt_invs)
lemma wcode_goon_checking_3_mv_Bk[simp]: "wcode_on_checking_3 ires rs (b, Bk # list) \<Longrightarrow>
wcode_goon_checking_3 ires rs (tl b, hd b # Bk # list)"
apply(auto simp: wcode_halt_invs)
done
lemma t_halt_case_correctness:
shows "let P = (\<lambda> (st, l, r). st = 0) in
let Q = (\<lambda> (st, l, r). wcode_halt_case_inv st ires rs (l, r)) in
let f = (\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp) in
\<exists> n .P (f n) \<and> Q (f (n::nat))"
proof -
let ?P = "(\<lambda> (st, l, r). st = 0)"
let ?Q = "(\<lambda> (st, l, r). wcode_halt_case_inv st ires rs (l, r))"
let ?f = "(\<lambda> stp. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp)"
have "\<exists> n. ?P (?f n) \<and> ?Q (?f (n::nat))"
proof(rule_tac halt_lemma2)
show "wf wcode_halt_case_le" by auto
next
{ fix na
obtain a b c where abc:"?f na = (a,b,c)" by(cases "?f na",auto)
hence "\<not> ?P (?f na) \<and> ?Q (?f na) \<Longrightarrow>
?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_halt_case_le"
apply(simp add: step.simps)
apply(cases c;cases "hd c")
apply(auto simp: wcode_halt_case_le_def lex_pair_def split: if_splits)
done
}
thus "\<forall> na. \<not> ?P (?f na) \<and> ?Q (?f na) \<longrightarrow>
?Q (?f (Suc na)) \<and> (?f (Suc na), ?f na) \<in> wcode_halt_case_le" by blast
next
show "?Q (?f 0)"
apply(simp add: steps.simps wcode_halt_invs)
apply(rule_tac x = "Suc m" in exI, simp)
apply(rule_tac x = "Suc 0" in exI, auto)
done
next
show "\<not> ?P (?f 0)"
apply(simp add: steps.simps)
done
qed
thus "?thesis"
apply(auto)
done
qed
declare wcode_halt_case_inv.simps[simp del]
lemma leading_Oc[intro]: "\<exists> xs. (<rev list @ [aa::nat]> :: cell list) = Oc # xs"
apply(case_tac "rev list", simp)
apply(simp add: tape_of_nl_cons)
done
lemma wcode_halt_case:
"\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n))
t_wcode_main stp = (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
proof -
let ?P = "\<lambda>(st, l, r). st = 0"
let ?Q = "\<lambda>(st, l, r). wcode_halt_case_inv st ires rs (l, r)"
let ?f = "steps0 (Suc 0, Bk # Bk \<up> m @ Oc # Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main"
from t_halt_case_correctness[of ires rs m n] obtain n where "?P (?f n) \<and> ?Q (?f n)" by metis
thus ?thesis
apply(simp add: wcode_halt_case_inv.simps wcode_stop.simps)
apply(case_tac "steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # Bk # ires,
Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main n")
apply(auto simp: wcode_halt_case_inv.simps wcode_stop.simps)
by auto
qed
lemma bl_bin_one[simp]: "bl_bin [Oc] = 1"
apply(simp add: bl_bin.simps)
done
lemma twice_power[intro]: "2 * 2 ^ a = Suc (Suc (2 * bl_bin (Oc \<up> a)))"
apply(induct a, auto simp: bl_bin.simps)
done
declare replicate_Suc[simp del]
lemma t_wcode_main_lemma_pre:
"\<lbrakk>args \<noteq> []; lm = <args::nat list>\<rbrakk> \<Longrightarrow>
\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main
stp
= (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2^(length lm - 1) ) @ Bk\<up>(rn))"
proof(induct "length args" arbitrary: args lm rs m n, simp)
fix x args lm rs m n
assume ind:
"\<And>args lm rs m n.
\<lbrakk>x = length args; (args::nat list) \<noteq> []; lm = <args>\<rbrakk>
\<Longrightarrow> \<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\<up>(rn))"
and h: "Suc x = length args" "(args::nat list) \<noteq> []" "lm = <args>"
from h have "\<exists> (a::nat) xs. args = xs @ [a]"
apply(rule_tac x = "last args" in exI)
apply(rule_tac x = "butlast args" in exI, auto)
done
from this obtain a xs where "args = xs @ [a]" by blast
from h and this show
"\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\<up>(rn))"
proof(case_tac "xs::nat list", simp)
show "\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk \<up> m @ Oc \<up> Suc a @ Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp =
(0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> (bl_bin (Oc \<up> Suc a) + rs * 2 ^ a) @ Bk \<up> rn)"
proof(induct "a" arbitrary: m n rs ires, simp)
fix m n rs ires
show "\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk \<up> m @ Oc # Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp =
(0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> Suc rs @ Bk \<up> rn)"
apply(rule_tac wcode_halt_case)
done
next
fix a m n rs ires
assume ind2:
"\<And>m n rs ires.
\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk \<up> m @ Oc \<up> Suc a @ Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp =
(0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> (bl_bin (Oc \<up> Suc a) + rs * 2 ^ a) @ Bk \<up> rn)"
show " \<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk \<up> m @ Oc \<up> Suc (Suc a) @ Bk # Bk # ires, Bk # Oc \<up> Suc rs @ Bk \<up> n) t_wcode_main stp =
(0, Bk # ires, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> (bl_bin (Oc \<up> Suc (Suc a)) + rs * 2 ^ Suc a) @ Bk \<up> rn)"
proof -
have "\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(Suc 0, Bk # Bk\<up>(ln) @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rn))"
apply(simp add: tape_of_nat)
using wcode_double_case[of m "Oc\<up>(a) @ Bk # Bk # ires" rs n]
apply(simp add: replicate_Suc)
done
from this obtain stpa lna rna where stp1:
"steps0 (Suc 0, Bk # Bk\<up>(m) @ rev (<Suc a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa =
(Suc 0, Bk # Bk\<up>(lna) @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rna))" by blast
moreover have
"\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(lna) @ rev (<a::nat>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rna)) t_wcode_main stp =
(0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<a>) + (2*rs + 2) * 2 ^ a) @ Bk\<up>(rn))"
using ind2[of lna ires "2*rs + 2" rna] by(simp add: tape_of_list_def tape_of_nat_def)
from this obtain stpb lnb rnb where stp2:
"steps0 (Suc 0, Bk # Bk\<up>(lna) @ rev (<a>) @ Bk # Bk # ires, Bk # Oc\<up>(Suc (2 * rs + 2)) @ Bk\<up>(rna)) t_wcode_main stpb =
(0, Bk # ires, Bk # Oc # Bk\<up>(lnb) @ Bk # Bk # Oc\<up>(bl_bin (<a>) + (2*rs + 2) * 2 ^ a) @ Bk\<up>(rnb))"
by blast
from stp1 and stp2 show "?thesis"
apply(rule_tac x = "stpa + stpb" in exI,
rule_tac x = lnb in exI, rule_tac x = rnb in exI, simp add: tape_of_nat_def)
apply(simp add: bl_bin.simps replicate_Suc)
apply(auto)
done
qed
qed
next
fix aa list
assume g: "Suc x = length args" "args \<noteq> []" "lm = <args>" "args = xs @ [a::nat]" "xs = (aa::nat) # list"
thus "\<exists>stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(m) @ rev lm @ Bk # Bk # ires, Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin lm + rs * 2 ^ (length lm - 1)) @ Bk\<up>(rn))"
proof(induct a arbitrary: m n rs args lm, simp_all add: tape_of_nl_rev del: subst_all,
simp only: tape_of_nl_cons_app1, simp del: subst_all)
fix m n rs args lm
have "\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # rev (<(aa::nat) # list>) @ Bk # Bk # ires,
Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(Suc 0, Bk # Bk\<up>(ln) @ rev (<aa # list>) @ Bk # Bk # ires,
Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rn))"
proof(simp add: tape_of_nl_rev)
have "\<exists> xs. (<rev list @ [aa]>) = Oc # xs" by auto
from this obtain xs where "(<rev list @ [aa]>) = Oc # xs" ..
thus "\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires,
Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(Suc 0, Bk # Bk\<up>(ln) @ <rev list @ [aa]> @ Bk # Bk # ires, Bk # Oc\<up>(5 + 4 * rs) @ Bk\<up>(rn))"
apply(simp)
using wcode_fourtimes_case[of m "xs @ Bk # Bk # ires" rs n]
apply(simp)
done
qed
from this obtain stpa lna rna where stp1:
"steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # rev (<aa # list>) @ Bk # Bk # ires,
Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stpa =
(Suc 0, Bk # Bk\<up>(lna) @ rev (<aa # list>) @ Bk # Bk # ires,
Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rna))" by blast
from g have
"\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(lna) @ rev (<(aa::nat) # list>) @ Bk # Bk # ires,
Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rna)) t_wcode_main stp = (0, Bk # ires,
Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) ) @ Bk\<up>(rn))"
apply(rule_tac args = "(aa::nat)#list" in ind, simp_all)
done
from this obtain stpb lnb rnb where stp2:
"steps0 (Suc 0, Bk # Bk\<up>(lna) @ rev (<(aa::nat) # list>) @ Bk # Bk # ires,
Bk # Oc\<up>(Suc (4*rs + 4)) @ Bk\<up>(rna)) t_wcode_main stpb = (0, Bk # ires,
Bk # Oc # Bk\<up>(lnb) @ Bk # Bk # Oc\<up>(bl_bin (<aa#list>)+ (4*rs + 4) * 2^(length (<aa#list>) - 1) ) @ Bk\<up>(rnb))"
by blast
from stp1 and stp2 and h
show "\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc # Bk # <rev list @ [aa]> @ Bk # Bk # ires,
Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk #
Bk # Oc\<up>(bl_bin (Oc\<up>(Suc aa) @ Bk # <list @ [0]>) + rs * (2 * 2 ^ (aa + length (<list @ [0]>)))) @ Bk\<up>(rn))"
apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI,
rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_rev)
done
next
fix ab m n rs args lm
assume ind2:
"\<And> m n rs args lm.
\<lbrakk>lm = <aa # list @ [ab]>; args = aa # list @ [ab]\<rbrakk>
\<Longrightarrow> \<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ <ab # rev list @ [aa]> @ Bk # Bk # ires,
Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk #
Bk # Oc\<up>(bl_bin (<aa # list @ [ab]>) + rs * 2 ^ (length (<aa # list @ [ab]>) - Suc 0)) @ Bk\<up>(rn))"
and k: "args = aa # list @ [Suc ab]" "lm = <aa # list @ [Suc ab]>"
show "\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ <Suc ab # rev list @ [aa]> @ Bk # Bk # ires,
Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(0, Bk # ires,Bk # Oc # Bk\<up>(ln) @ Bk #
Bk # Oc\<up>(bl_bin (<aa # list @ [Suc ab]>) + rs * 2 ^ (length (<aa # list @ [Suc ab]>) - Suc 0)) @ Bk\<up>(rn))"
proof(simp add: tape_of_nl_cons_app1)
have "\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc\<up>(Suc (Suc ab)) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
Bk # Oc # Oc\<up>(rs) @ Bk\<up>(n)) t_wcode_main stp
= (Suc 0, Bk # Bk\<up>(ln) @ Oc\<up>(Suc ab) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rn))"
using wcode_double_case[of m "Oc\<up>(ab) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires"
rs n]
apply(simp add: replicate_Suc)
done
from this obtain stpa lna rna where stp1:
"steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc\<up>(Suc (Suc ab)) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
Bk # Oc # Oc\<up>(rs) @ Bk\<up>(n)) t_wcode_main stpa
= (Suc 0, Bk # Bk\<up>(lna) @ Oc\<up>(Suc ab) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rna))" by blast
from k have
"\<exists> stp ln rn. steps0 (Suc 0, Bk # Bk\<up>(lna) @ <ab # rev list @ [aa]> @ Bk # Bk # ires,
Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rna)) t_wcode_main stp
= (0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk #
Bk # Oc\<up>(bl_bin (<aa # list @ [ab]> ) + (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)) @ Bk\<up>(rn))"
apply(rule_tac ind2, simp_all)
done
from this obtain stpb lnb rnb where stp2:
"steps0 (Suc 0, Bk # Bk\<up>(lna) @ <ab # rev list @ [aa]> @ Bk # Bk # ires,
Bk # Oc\<up>(Suc (2*rs + 2)) @ Bk\<up>(rna)) t_wcode_main stpb
= (0, Bk # ires, Bk # Oc # Bk\<up>(lnb) @ Bk #
Bk # Oc\<up>(bl_bin (<aa # list @ [ab]> ) + (2*rs + 2)* 2^(length (<aa # list @ [ab]>) - Suc 0)) @ Bk\<up>(rnb))"
by blast
from stp1 and stp2 show
"\<exists>stp ln rn.
steps0 (Suc 0, Bk # Bk\<up>(m) @ Oc\<up>(Suc (Suc ab)) @ Bk # <rev list @ [aa]> @ Bk # Bk # ires,
Bk # Oc\<up>(Suc rs) @ Bk\<up>(n)) t_wcode_main stp =
(0, Bk # ires, Bk # Oc # Bk\<up>(ln) @ Bk # Bk #
Oc\<up>(bl_bin (Oc\<up>(Suc aa) @ Bk # <list @ [Suc ab]>) + rs * (2 * 2 ^ (aa + length (<list @ [Suc ab]>))))
@ Bk\<up>(rn))"
apply(rule_tac x = "stpa + stpb" in exI, rule_tac x = lnb in exI,
rule_tac x = rnb in exI, simp add: steps_add tape_of_nl_cons_app1 replicate_Suc)
done
qed
qed
qed
qed
definition t_wcode_prepare :: "instr list"
where
"t_wcode_prepare \<equiv>
[(W1, 2), (L, 1), (L, 3), (R, 2), (R, 4), (W0, 3),
(R, 4), (R, 5), (R, 6), (R, 5), (R, 7), (R, 5),
(W1, 7), (L, 0)]"
fun wprepare_add_one :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_add_one m lm (l, r) =
(\<exists> rn. l = [] \<and>
(r = <m # lm> @ Bk\<up>(rn) \<or>
r = Bk # <m # lm> @ Bk\<up>(rn)))"
fun wprepare_goto_first_end :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_goto_first_end m lm (l, r) =
(\<exists> ml mr rn. l = Oc\<up>(ml) \<and>
r = Oc\<up>(mr) @ Bk # <lm> @ Bk\<up>(rn) \<and>
ml + mr = Suc (Suc m))"
fun wprepare_erase :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_erase m lm (l, r) =
(\<exists> rn. l = Oc\<up>(Suc m) \<and>
tl r = Bk # <lm> @ Bk\<up>(rn))"
fun wprepare_goto_start_pos_B :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_goto_start_pos_B m lm (l, r) =
(\<exists> rn. l = Bk # Oc\<up>(Suc m) \<and>
r = Bk # <lm> @ Bk\<up>(rn))"
fun wprepare_goto_start_pos_O :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_goto_start_pos_O m lm (l, r) =
(\<exists> rn. l = Bk # Bk # Oc\<up>(Suc m) \<and>
r = <lm> @ Bk\<up>(rn))"
fun wprepare_goto_start_pos :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_goto_start_pos m lm (l, r) =
(wprepare_goto_start_pos_B m lm (l, r) \<or>
wprepare_goto_start_pos_O m lm (l, r))"
fun wprepare_loop_start_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_loop_start_on_rightmost m lm (l, r) =
(\<exists> rn mr. rev l @ r = Oc\<up>(Suc m) @ Bk # Bk # <lm> @ Bk\<up>(rn) \<and> l \<noteq> [] \<and>
r = Oc\<up>(mr) @ Bk\<up>(rn))"
fun wprepare_loop_start_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_loop_start_in_middle m lm (l, r) =
(\<exists> rn (mr:: nat) (lm1::nat list).
rev l @ r = Oc\<up>(Suc m) @ Bk # Bk # <lm> @ Bk\<up>(rn) \<and> l \<noteq> [] \<and>
r = Oc\<up>(mr) @ Bk # <lm1> @ Bk\<up>(rn) \<and> lm1 \<noteq> [])"
fun wprepare_loop_start :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_loop_start m lm (l, r) = (wprepare_loop_start_on_rightmost m lm (l, r) \<or>
wprepare_loop_start_in_middle m lm (l, r))"
fun wprepare_loop_goon_on_rightmost :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_loop_goon_on_rightmost m lm (l, r) =
(\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<up>(Suc m) \<and>
r = Bk\<up>(rn))"
fun wprepare_loop_goon_in_middle :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_loop_goon_in_middle m lm (l, r) =
(\<exists> rn (mr:: nat) (lm1::nat list).
rev l @ r = Oc\<up>(Suc m) @ Bk # Bk # <lm> @ Bk\<up>(rn) \<and> l \<noteq> [] \<and>
(if lm1 = [] then r = Oc\<up>(mr) @ Bk\<up>(rn)
else r = Oc\<up>(mr) @ Bk # <lm1> @ Bk\<up>(rn)) \<and> mr > 0)"
fun wprepare_loop_goon :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_loop_goon m lm (l, r) =
(wprepare_loop_goon_in_middle m lm (l, r) \<or>
wprepare_loop_goon_on_rightmost m lm (l, r))"
fun wprepare_add_one2 :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_add_one2 m lm (l, r) =
(\<exists> rn. l = Bk # Bk # <rev lm> @ Bk # Bk # Oc\<up>(Suc m) \<and>
(r = [] \<or> tl r = Bk\<up>(rn)))"
fun wprepare_stop :: "nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_stop m lm (l, r) =
(\<exists> rn. l = Bk # <rev lm> @ Bk # Bk # Oc\<up>(Suc m) \<and>
r = Bk # Oc # Bk\<up>(rn))"
fun wprepare_inv :: "nat \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> tape \<Rightarrow> bool"
where
"wprepare_inv st m lm (l, r) =
(if st = 0 then wprepare_stop m lm (l, r)
else if st = Suc 0 then wprepare_add_one m lm (l, r)
else if st = Suc (Suc 0) then wprepare_goto_first_end m lm (l, r)
else if st = Suc (Suc (Suc 0)) then wprepare_erase m lm (l, r)
else if st = 4 then wprepare_goto_start_pos m lm (l, r)
else if st = 5 then wprepare_loop_start m lm (l, r)
else if st = 6 then wprepare_loop_goon m lm (l, r)
else if st = 7 then wprepare_add_one2 m lm (l, r)
else False)"
fun wprepare_stage :: "config \<Rightarrow> nat"
where
"wprepare_stage (st, l, r) =
(if st \<ge> 1 \<and> st \<le> 4 then 3
else if st = 5 \<or> st = 6 then 2
else 1)"
fun wprepare_state :: "config \<Rightarrow> nat"
where
"wprepare_state (st, l, r) =
(if st = 1 then 4
else if st = Suc (Suc 0) then 3
else if st = Suc (Suc (Suc 0)) then 2
else if st = 4 then 1
else if st = 7 then 2
else 0)"
fun wprepare_step :: "config \<Rightarrow> nat"
where
"wprepare_step (st, l, r) =
(if st = 1 then (if hd r = Oc then Suc (length l)
else 0)
else if st = Suc (Suc 0) then length r
else if st = Suc (Suc (Suc 0)) then (if hd r = Oc then 1
else 0)
else if st = 4 then length r
else if st = 5 then Suc (length r)
else if st = 6 then (if r = [] then 0 else Suc (length r))
else if st = 7 then (if (r \<noteq> [] \<and> hd r = Oc) then 0
else 1)
else 0)"
fun wcode_prepare_measure :: "config \<Rightarrow> nat \<times> nat \<times> nat"
where
"wcode_prepare_measure (st, l, r) =
(wprepare_stage (st, l, r),
wprepare_state (st, l, r),
wprepare_step (st, l, r))"
definition wcode_prepare_le :: "(config \<times> config) set"
where "wcode_prepare_le \<equiv> (inv_image lex_triple wcode_prepare_measure)"
lemma wf_wcode_prepare_le[intro]: "wf wcode_prepare_le"
by(auto intro:wf_inv_image simp: wcode_prepare_le_def
lex_triple_def)
declare wprepare_add_one.simps[simp del] wprepare_goto_first_end.simps[simp del]
wprepare_erase.simps[simp del] wprepare_goto_start_pos.simps[simp del]
wprepare_loop_start.simps[simp del] wprepare_loop_goon.simps[simp del]
wprepare_add_one2.simps[simp del]
lemmas wprepare_invs = wprepare_add_one.simps wprepare_goto_first_end.simps
wprepare_erase.simps wprepare_goto_start_pos.simps
wprepare_loop_start.simps wprepare_loop_goon.simps
wprepare_add_one2.simps
declare wprepare_inv.simps[simp del]
lemma fetch_t_wcode_prepare[simp]:
"fetch t_wcode_prepare (Suc 0) Bk = (W1, 2)"
"fetch t_wcode_prepare (Suc 0) Oc = (L, 1)"
"fetch t_wcode_prepare (Suc (Suc 0)) Bk = (L, 3)"
"fetch t_wcode_prepare (Suc (Suc 0)) Oc = (R, 2)"
"fetch t_wcode_prepare (Suc (Suc (Suc 0))) Bk = (R, 4)"
"fetch t_wcode_prepare (Suc (Suc (Suc 0))) Oc = (W0, 3)"
"fetch t_wcode_prepare 4 Bk = (R, 4)"
"fetch t_wcode_prepare 4 Oc = (R, 5)"
"fetch t_wcode_prepare 5 Oc = (R, 5)"
"fetch t_wcode_prepare 5 Bk = (R, 6)"
"fetch t_wcode_prepare 6 Oc = (R, 5)"
"fetch t_wcode_prepare 6 Bk = (R, 7)"
"fetch t_wcode_prepare 7 Oc = (L, 0)"
"fetch t_wcode_prepare 7 Bk = (W1, 7)"
unfolding fetch.simps t_wcode_prepare_def nth_of.simps
numeral by auto
lemma wprepare_add_one_nonempty_snd[simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_add_one m lm (b, []) = False"
apply(simp add: wprepare_invs)
done
lemma wprepare_goto_first_end_nonempty_snd[simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_first_end m lm (b, []) = False"
apply(simp add: wprepare_invs)
done
lemma wprepare_erase_nonempty_snd[simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_erase m lm (b, []) = False"
apply(simp add: wprepare_invs)
done
lemma wprepare_goto_start_pos_nonempty_snd[simp]: "lm \<noteq> [] \<Longrightarrow> wprepare_goto_start_pos m lm (b, []) = False"
apply(simp add: wprepare_invs)
done
lemma wprepare_loop_start_empty_nonempty_fst[simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []"
apply(simp add: wprepare_invs)
done
lemma rev_eq: "rev xs = rev ys \<Longrightarrow> xs = ys" by auto
lemma wprepare_loop_goon_Bk_empty_snd[simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [])\<rbrakk> \<Longrightarrow>
wprepare_loop_goon m lm (Bk # b, [])"
apply(simp only: wprepare_invs)
apply(erule_tac disjE)
apply(rule_tac disjI2)
apply(simp add: wprepare_loop_start_on_rightmost.simps
wprepare_loop_goon_on_rightmost.simps, auto)
apply(rule_tac rev_eq, simp add: tape_of_nl_rev)
done
lemma wprepare_loop_goon_nonempty_fst[simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs, auto)
done
lemma wprepare_add_one2_Bk_empty[simp]:"\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, [])\<rbrakk> \<Longrightarrow>
wprepare_add_one2 m lm (Bk # b, [])"
apply(simp only: wprepare_invs, auto split: if_splits)
done
lemma wprepare_add_one2_nonempty_fst[simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs, auto)
done
lemma wprepare_add_one2_Oc[simp]: "wprepare_add_one2 m lm (b, []) \<Longrightarrow> wprepare_add_one2 m lm (b, [Oc])"
apply(simp only: wprepare_invs, auto)
done
lemma Bk_not_tape_start[simp]: "(Bk # list = <(m::nat) # lm> @ ys) = False"
apply(case_tac lm, auto simp: tape_of_nl_cons replicate_Suc)
done
lemma wprepare_goto_first_end_cases[simp]:
"\<lbrakk>lm \<noteq> []; wprepare_add_one m lm (b, Bk # list)\<rbrakk>
\<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([], Oc # list)) \<and>
(b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (b, Oc # list))"
apply(simp only: wprepare_invs)
apply(auto simp: tape_of_nl_cons split: if_splits)
apply(cases lm, auto simp add:tape_of_list_def replicate_Suc)
done
lemma wprepare_goto_first_end_Bk_nonempty_fst[simp]:
"wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs , auto simp: replicate_Suc)
done
declare replicate_Suc[simp]
lemma wprepare_erase_elem_Bk_rest[simp]: "wprepare_goto_first_end m lm (b, Bk # list) \<Longrightarrow>
wprepare_erase m lm (tl b, hd b # Bk # list)"
by(simp add: wprepare_invs)
lemma wprepare_erase_Bk_nonempty_fst[simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
by(simp add: wprepare_invs)
lemma wprepare_goto_start_pos_Bk[simp]: "wprepare_erase m lm (b, Bk # list) \<Longrightarrow>
wprepare_goto_start_pos m lm (Bk # b, list)"
apply(simp only: wprepare_invs, auto)
done
lemma wprepare_add_one_Bk_nonempty_snd[simp]: "\<lbrakk>wprepare_add_one m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
apply(simp only: wprepare_invs)
apply(case_tac lm, simp_all add: tape_of_list_def tape_of_nat_def, auto)
done
lemma wprepare_goto_first_end_nonempty_snd_tl[simp]:
"\<lbrakk>lm \<noteq> []; wprepare_goto_first_end m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
by(simp only: wprepare_invs, auto)
lemma wprepare_erase_Bk_nonempty_list[simp]: "\<lbrakk>lm \<noteq> []; wprepare_erase m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
apply(simp only: wprepare_invs, auto)
done
lemma wprepare_goto_start_pos_Bk_nonempty[simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> list \<noteq> []"
by(cases lm;cases list;simp only: wprepare_invs, auto)
lemma wprepare_goto_start_pos_Bk_nonempty_fst[simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs)
apply(auto)
done
lemma wprepare_loop_goon_Bk_nonempty[simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, Bk # list)\<rbrakk> \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs, auto)
done
lemma wprepare_loop_goon_wprepare_add_one2_cases[simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_goon m lm (b, Bk # list)\<rbrakk> \<Longrightarrow>
(list = [] \<longrightarrow> wprepare_add_one2 m lm (Bk # b, [])) \<and>
(list \<noteq> [] \<longrightarrow> wprepare_add_one2 m lm (Bk # b, list))"
unfolding wprepare_invs
apply(cases list;auto split:nat.split if_splits)
by (metis list.sel(3) tl_replicate)
lemma wprepare_add_one2_nonempty[simp]: "wprepare_add_one2 m lm (b, Bk # list) \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs, simp)
done
lemma wprepare_add_one2_cases[simp]: "wprepare_add_one2 m lm (b, Bk # list) \<Longrightarrow>
(list = [] \<longrightarrow> wprepare_add_one2 m lm (b, [Oc])) \<and>
(list \<noteq> [] \<longrightarrow> wprepare_add_one2 m lm (b, Oc # list))"
apply(simp only: wprepare_invs, auto)
done
lemma wprepare_goto_first_end_cases_Oc[simp]: "wprepare_goto_first_end m lm (b, Oc # list)
\<Longrightarrow> (b = [] \<longrightarrow> wprepare_goto_first_end m lm ([Oc], list)) \<and>
(b \<noteq> [] \<longrightarrow> wprepare_goto_first_end m lm (Oc # b, list))"
apply(simp only: wprepare_invs, auto)
apply(rule_tac x = 1 in exI, auto) apply(rename_tac ml mr rn)
apply(case_tac mr, simp_all add: )
apply(case_tac ml, simp_all add: )
apply(rule_tac x = "Suc ml" in exI, simp_all add: )
apply(rule_tac x = "mr - 1" in exI, simp)
done
lemma wprepare_erase_nonempty[simp]: "wprepare_erase m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp only: wprepare_invs, auto simp: )
done
lemma wprepare_erase_Bk[simp]: "wprepare_erase m lm (b, Oc # list)
\<Longrightarrow> wprepare_erase m lm (b, Bk # list)"
apply(simp only:wprepare_invs, auto simp: )
done
lemma wprepare_goto_start_pos_Bk_move[simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Bk # list)\<rbrakk>
\<Longrightarrow> wprepare_goto_start_pos m lm (Bk # b, list)"
apply(simp only:wprepare_invs, auto)
apply(case_tac [!] lm, simp, simp_all)
done
lemma wprepare_loop_start_b_nonempty[simp]: "wprepare_loop_start m lm (b, aa) \<Longrightarrow> b \<noteq> []"
apply(simp only:wprepare_invs, auto)
done
lemma exists_exp_of_Bk[elim]: "Bk # list = Oc\<up>(mr) @ Bk\<up>(rn) \<Longrightarrow> \<exists>rn. list = Bk\<up>(rn)"
apply(case_tac mr, simp_all)
apply(case_tac rn, simp_all)
done
lemma wprepare_loop_start_in_middle_Bk_False[simp]: "wprepare_loop_start_in_middle m lm (b, [Bk]) = False"
by(auto)
declare wprepare_loop_start_in_middle.simps[simp del]
declare wprepare_loop_start_on_rightmost.simps[simp del]
wprepare_loop_goon_in_middle.simps[simp del]
wprepare_loop_goon_on_rightmost.simps[simp del]
lemma wprepare_loop_goon_in_middle_Bk_False[simp]: "wprepare_loop_goon_in_middle m lm (Bk # b, []) = False"
apply(simp add: wprepare_loop_goon_in_middle.simps, auto)
done
lemma wprepare_loop_goon_Bk[simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, [Bk])\<rbrakk> \<Longrightarrow>
wprepare_loop_goon m lm (Bk # b, [])"
unfolding wprepare_invs
apply(auto simp add: wprepare_loop_goon_on_rightmost.simps
wprepare_loop_start_on_rightmost.simps)
apply(rule_tac rev_eq)
apply(simp add: tape_of_nl_rev)
apply(simp add: exp_ind replicate_Suc[THEN sym] del: replicate_Suc)
done
lemma wprepare_loop_goon_in_middle_Bk_False2[simp]: "wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)
\<Longrightarrow> wprepare_loop_goon_in_middle m lm (Bk # b, a # lista) = False"
apply(auto simp: wprepare_loop_start_on_rightmost.simps
wprepare_loop_goon_in_middle.simps)
done
lemma wprepare_loop_goon_on_rightbmost_Bk_False[simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start_on_rightmost m lm (b, Bk # a # lista)\<rbrakk>
\<Longrightarrow> wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista)"
apply(simp only: wprepare_loop_start_on_rightmost.simps
wprepare_loop_goon_on_rightmost.simps, auto simp: tape_of_nl_rev)
apply(simp add: replicate_Suc[THEN sym] exp_ind tape_of_nl_rev del: replicate_Suc)
by (meson Cons_replicate_eq)
lemma wprepare_loop_goon_in_middle_Bk_False3[simp]:
assumes "lm \<noteq> []" "wprepare_loop_start_in_middle m lm (b, Bk # a # lista)"
shows "wprepare_loop_goon_in_middle m lm (Bk # b, a # lista)" (is "?t1")
and "wprepare_loop_goon_on_rightmost m lm (Bk # b, a # lista) = False" (is ?t2)
proof -
from assms obtain rn mr lm1 where *:"rev b @ Oc \<up> mr @ Bk # <lm1> = Oc # Oc \<up> m @ Bk # Bk # <lm>"
"b \<noteq> []" "Bk # a # lista = Oc \<up> mr @ Bk # <lm1::nat list> @ Bk \<up> rn" "lm1 \<noteq> []"
by(auto simp add: wprepare_loop_start_in_middle.simps)
thus ?t1 apply(simp add: wprepare_loop_start_in_middle.simps
wprepare_loop_goon_in_middle.simps, auto)
apply(rule_tac x = rn in exI, simp)
apply(case_tac mr, simp_all add: )
apply(case_tac lm1, simp)
apply(rule_tac x = "Suc (hd lm1)" in exI, simp)
apply(rule_tac x = "tl lm1" in exI)
apply(case_tac "tl lm1", simp_all add: tape_of_list_def tape_of_nat_def)
done
from * show ?t2
apply(simp add: wprepare_loop_start_in_middle.simps
wprepare_loop_goon_on_rightmost.simps del:split_head_repeat, auto simp del:split_head_repeat)
apply(case_tac mr)
apply(case_tac "lm1::nat list", simp_all, case_tac "tl lm1", simp_all)
apply(auto simp add: tape_of_list_def )
apply(case_tac [!] rna, simp_all add: )
apply(case_tac mr, simp_all add: )
apply(case_tac lm1, simp, case_tac list, simp)
apply(simp_all add: tape_of_nat_def)
by (metis Bk_not_tape_start tape_of_list_def tape_of_nat_list.elims)
qed
lemma wprepare_loop_goon_Bk2[simp]: "\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, Bk # a # lista)\<rbrakk> \<Longrightarrow>
wprepare_loop_goon m lm (Bk # b, a # lista)"
apply(simp add: wprepare_loop_start.simps
wprepare_loop_goon.simps)
apply(erule_tac disjE, simp, auto)
done
lemma start_2_goon:
"\<lbrakk>lm \<noteq> []; wprepare_loop_start m lm (b, Bk # list)\<rbrakk> \<Longrightarrow>
(list = [] \<longrightarrow> wprepare_loop_goon m lm (Bk # b, [])) \<and>
(list \<noteq> [] \<longrightarrow> wprepare_loop_goon m lm (Bk # b, list))"
apply(case_tac list, auto)
done
lemma add_one_2_add_one: "wprepare_add_one m lm (b, Oc # list)
\<Longrightarrow> (hd b = Oc \<longrightarrow> (b = [] \<longrightarrow> wprepare_add_one m lm ([], Bk # Oc # list)) \<and>
(b \<noteq> [] \<longrightarrow> wprepare_add_one m lm (tl b, Oc # Oc # list))) \<and>
(hd b \<noteq> Oc \<longrightarrow> (b = [] \<longrightarrow> wprepare_add_one m lm ([], Bk # Oc # list)) \<and>
(b \<noteq> [] \<longrightarrow> wprepare_add_one m lm (tl b, hd b # Oc # list)))"
unfolding wprepare_add_one.simps by auto
lemma wprepare_loop_start_on_rightmost_Oc[simp]: "wprepare_loop_start_on_rightmost m lm (b, Oc # list) \<Longrightarrow>
wprepare_loop_start_on_rightmost m lm (Oc # b, list)"
apply(simp add: wprepare_loop_start_on_rightmost.simps)
by (metis Cons_replicate_eq cell.distinct(1) list.sel(3) self_append_conv2 tl_append2 tl_replicate)
lemma wprepare_loop_start_in_middle_Oc[simp]:
assumes "wprepare_loop_start_in_middle m lm (b, Oc # list)"
shows "wprepare_loop_start_in_middle m lm (Oc # b, list)"
proof -
from assms obtain mr lm1 rn
where "rev b @ Oc \<up> mr @ Bk # <lm1::nat list> = Oc # Oc \<up> m @ Bk # Bk # <lm>"
"Oc # list = Oc \<up> mr @ Bk # <lm1> @ Bk \<up> rn" "lm1 \<noteq> []"
by(auto simp add: wprepare_loop_start_in_middle.simps)
thus ?thesis
apply(auto simp add: wprepare_loop_start_in_middle.simps)
apply(rule_tac x = rn in exI, auto)
apply(case_tac mr, simp, simp add: )
apply(rule_tac x = "mr - 1" in exI, simp)
apply(rule_tac x = lm1 in exI, simp)
done
qed
lemma start_2_start: "wprepare_loop_start m lm (b, Oc # list) \<Longrightarrow>
wprepare_loop_start m lm (Oc # b, list)"
apply(simp add: wprepare_loop_start.simps)
apply(erule_tac disjE, simp_all )
done
lemma wprepare_loop_goon_Oc_nonempty[simp]: "wprepare_loop_goon m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp add: wprepare_loop_goon.simps
wprepare_loop_goon_in_middle.simps
wprepare_loop_goon_on_rightmost.simps)
apply(auto)
done
lemma wprepare_goto_start_pos_Oc_nonempty[simp]: "wprepare_goto_start_pos m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(simp add: wprepare_goto_start_pos.simps)
done
lemma wprepare_loop_goon_on_rightmost_Oc_False[simp]: "wprepare_loop_goon_on_rightmost m lm (b, Oc # list) = False"
apply(simp add: wprepare_loop_goon_on_rightmost.simps)
done
lemma wprepare_loop1: "\<lbrakk>rev b @ Oc\<up>(mr) = Oc\<up>(Suc m) @ Bk # Bk # <lm>;
b \<noteq> []; 0 < mr; Oc # list = Oc\<up>(mr) @ Bk\<up>(rn)\<rbrakk>
\<Longrightarrow> wprepare_loop_start_on_rightmost m lm (Oc # b, list)"
apply(simp add: wprepare_loop_start_on_rightmost.simps)
apply(rule_tac x = rn in exI, simp)
apply(case_tac mr, simp, simp)
done
lemma wprepare_loop2: "\<lbrakk>rev b @ Oc\<up>(mr) @ Bk # <a # lista> = Oc\<up>(Suc m) @ Bk # Bk # <lm>;
b \<noteq> []; Oc # list = Oc\<up>(mr) @ Bk # <(a::nat) # lista> @ Bk\<up>(rn)\<rbrakk>
\<Longrightarrow> wprepare_loop_start_in_middle m lm (Oc # b, list)"
apply(simp add: wprepare_loop_start_in_middle.simps)
apply(rule_tac x = rn in exI, simp)
apply(case_tac mr, simp_all add: ) apply(rename_tac nat)
apply(rule_tac x = nat in exI, simp)
apply(rule_tac x = "a#lista" in exI, simp)
done
lemma wprepare_loop_goon_in_middle_cases[simp]: "wprepare_loop_goon_in_middle m lm (b, Oc # list) \<Longrightarrow>
wprepare_loop_start_on_rightmost m lm (Oc # b, list) \<or>
wprepare_loop_start_in_middle m lm (Oc # b, list)"
apply(simp add: wprepare_loop_goon_in_middle.simps split: if_splits) apply(rename_tac lm1)
apply(case_tac lm1, simp_all add: wprepare_loop1 wprepare_loop2)
done
lemma wprepare_add_one_b[simp]: "wprepare_add_one m lm (b, Oc # list)
\<Longrightarrow> b = [] \<longrightarrow> wprepare_add_one m lm ([], Bk # Oc # list)"
"wprepare_loop_goon m lm (b, Oc # list)
\<Longrightarrow> wprepare_loop_start m lm (Oc # b, list)"
apply(auto simp add: wprepare_add_one.simps wprepare_loop_goon.simps
wprepare_loop_start.simps)
done
lemma wprepare_loop_start_on_rightmost_Oc2[simp]: "wprepare_goto_start_pos m [a] (b, Oc # list)
\<Longrightarrow> wprepare_loop_start_on_rightmost m [a] (Oc # b, list) "
apply(auto simp: wprepare_goto_start_pos.simps
wprepare_loop_start_on_rightmost.simps) apply(rename_tac rn)
apply(rule_tac x = rn in exI, simp)
apply(simp add: replicate_Suc[THEN sym] exp_ind del: replicate_Suc)
done
lemma wprepare_loop_start_in_middle_2_Oc[simp]: "wprepare_goto_start_pos m (a # aa # listaa) (b, Oc # list)
\<Longrightarrow>wprepare_loop_start_in_middle m (a # aa # listaa) (Oc # b, list)"
apply(auto simp: wprepare_goto_start_pos.simps
wprepare_loop_start_in_middle.simps) apply(rename_tac rn)
apply(rule_tac x = rn in exI, simp)
apply(simp add: exp_ind[THEN sym])
apply(rule_tac x = a in exI, rule_tac x = "aa#listaa" in exI, simp)
apply(simp add: tape_of_nl_cons)
done
lemma wprepare_loop_start_Oc2[simp]: "\<lbrakk>lm \<noteq> []; wprepare_goto_start_pos m lm (b, Oc # list)\<rbrakk>
\<Longrightarrow> wprepare_loop_start m lm (Oc # b, list)"
by(cases lm;cases "tl lm", auto simp add: wprepare_loop_start.simps)
lemma wprepare_add_one2_Oc_nonempty[simp]: "wprepare_add_one2 m lm (b, Oc # list) \<Longrightarrow> b \<noteq> []"
apply(auto simp: wprepare_add_one2.simps)
done
lemma add_one_2_stop:
"wprepare_add_one2 m lm (b, Oc # list)
\<Longrightarrow> wprepare_stop m lm (tl b, hd b # Oc # list)"
apply(simp add: wprepare_add_one2.simps)
done
declare wprepare_stop.simps[simp del]
lemma wprepare_correctness:
assumes h: "lm \<noteq> []"
shows "let P = (\<lambda> (st, l, r). st = 0) in
let Q = (\<lambda> (st, l, r). wprepare_inv st m lm (l, r)) in
let f = (\<lambda> stp. steps0 (Suc 0, [], (<m # lm>)) t_wcode_prepare stp) in
\<exists> n .P (f n) \<and> Q (f n)"
proof -
let ?P = "(\<lambda> (st, l, r). st = 0)"
let ?Q = "(\<lambda> (st, l, r). wprepare_inv st m lm (l, r))"
let ?f = "(\<lambda> stp. steps0 (Suc 0, [], (<m # lm>)) t_wcode_prepare stp)"
have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)"
proof(rule_tac halt_lemma2)
show "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow>
?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wcode_prepare_le"
using h
apply(rule_tac allI, rule_tac impI) apply(rename_tac n)
apply(case_tac "?f n", simp add: step.simps) apply(rename_tac c)
apply(case_tac c, simp, case_tac [2] aa)
apply(simp_all add: wprepare_inv.simps wcode_prepare_le_def lex_triple_def lex_pair_def
split: if_splits) (* slow *)
apply(simp_all add: start_2_goon start_2_start
add_one_2_add_one add_one_2_stop)
apply(auto simp: wprepare_add_one2.simps)
done
qed (auto simp add: steps.simps wprepare_inv.simps wprepare_invs)
thus "?thesis"
apply(auto)
done
qed
lemma tm_wf_t_wcode_prepare[intro]: "tm_wf (t_wcode_prepare, 0)"
apply(simp add:tm_wf.simps t_wcode_prepare_def)
done
lemma is_28_even[intro]: "(28 + (length t_twice_compile + length t_fourtimes_compile)) mod 2 = 0"
by(auto simp: t_twice_compile_def t_fourtimes_compile_def)
lemma b_le_28[elim]: "(a, b) \<in> set t_wcode_main_first_part \<Longrightarrow>
b \<le> (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2"
apply(auto simp: t_wcode_main_first_part_def t_twice_def)
done
lemma tm_wf_change_termi:
assumes "tm_wf (tp, 0)"
shows "list_all (\<lambda>(acn, st). (st \<le> Suc (length tp div 2))) (adjust0 tp)"
proof -
{ fix acn st n
assume "tp ! n = (acn, st)" "n < length tp" "0 < st"
hence "(acn, st)\<in>set tp" by (metis nth_mem)
with assms tm_wf.simps have "st \<le> length tp div 2 + 0" by auto
hence "st \<le> Suc (length tp div 2)" by auto
}
thus ?thesis
by(auto simp: tm_wf.simps List.list_all_length adjust.simps split: if_splits prod.split)
qed
lemma tm_wf_shift:
assumes "list_all (\<lambda>(acn, st). (st \<le> y)) tp"
shows "list_all (\<lambda>(acn, st). (st \<le> y + off)) (shift tp off)"
proof -
have [dest!]:"\<And> P Q n. \<forall>n. Q n \<longrightarrow> P n \<Longrightarrow> Q n \<Longrightarrow> P n" by metis
from assms show ?thesis by(auto simp: tm_wf.simps List.list_all_length shift.simps)
qed
declare length_tp'[simp del]
lemma length_mopup_1[simp]: "length (mopup (Suc 0)) = 16"
apply(auto simp: mopup.simps)
done
lemma twice_plus_28_elim[elim]: "(a, b) \<in> set (shift (adjust0 t_twice_compile) 12) \<Longrightarrow>
b \<le> (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2"
apply(simp add: t_twice_compile_def t_fourtimes_compile_def)
proof -
assume g: "(a, b)
\<in> set (shift
(adjust
(tm_of abc_twice @
shift (mopup (Suc 0)) (length (tm_of abc_twice) div 2))
(Suc ((length (tm_of abc_twice) + 16) div 2)))
12)"
moreover have "length (tm_of abc_twice) mod 2 = 0" by auto
moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto
ultimately have "list_all (\<lambda>(acn, st). (st \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2))
(shift (adjust0 t_twice_compile) 12)"
proof(auto simp add: mod_ex1 del: adjust.simps)
assume "even (length (tm_of abc_twice))"
then obtain q where q:"length (tm_of abc_twice) = 2 * q" by auto
assume "even (length (tm_of abc_fourtimes))"
then obtain qa where qa:"length (tm_of abc_fourtimes) = 2 * qa" by auto
note h = q qa
hence "list_all (\<lambda>(acn, st). st \<le> (18 + (q + qa)) + 12) (shift (adjust0 t_twice_compile) 12)"
proof(rule_tac tm_wf_shift t_twice_compile_def)
have "list_all (\<lambda>(acn, st). st \<le> Suc (length t_twice_compile div 2)) (adjust0 t_twice_compile)"
by(rule_tac tm_wf_change_termi, auto)
thus "list_all (\<lambda>(acn, st). st \<le> 18 + (q + qa)) (adjust0 t_twice_compile)"
using h
apply(simp add: t_twice_compile_def, auto simp: List.list_all_length)
done
qed
thus "list_all (\<lambda>(acn, st). st \<le> 30 + (length (tm_of abc_twice) div 2 + length (tm_of abc_fourtimes) div 2))
(shift (adjust0 t_twice_compile) 12)" using h
by simp
qed
thus "b \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2"
using g
apply(auto simp:t_twice_compile_def)
apply(simp add: Ball_set[THEN sym])
apply(erule_tac x = "(a, b)" in ballE, simp, simp)
done
qed
lemma length_plus_28_elim2[elim]: "(a, b) \<in> set (shift (adjust0 t_fourtimes_compile) (t_twice_len + 13))
\<Longrightarrow> b \<le> (28 + (length t_twice_compile + length t_fourtimes_compile)) div 2"
apply(simp add: t_twice_compile_def t_fourtimes_compile_def t_twice_len_def)
proof -
assume g: "(a, b)
\<in> set (shift
(adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) (length (tm_of abc_fourtimes) div 2))
(Suc ((length (tm_of abc_fourtimes) + 16) div 2)))
(length t_twice div 2 + 13))"
moreover have "length (tm_of abc_twice) mod 2 = 0" by auto
moreover have "length (tm_of abc_fourtimes) mod 2 = 0" by auto
ultimately have "list_all (\<lambda>(acn, st). (st \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2))
(shift (adjust0 (tm_of abc_fourtimes @ shift (mopup (Suc 0))
(length (tm_of abc_fourtimes) div 2))) (length t_twice div 2 + 13))"
proof(auto simp: mod_ex1 t_twice_def t_twice_compile_def)
assume "even (length (tm_of abc_twice))"
then obtain q where q:"length (tm_of abc_twice) = 2 * q" by auto
assume "even (length (tm_of abc_fourtimes))"
then obtain qa where qa:"length (tm_of abc_fourtimes) = 2 * qa" by auto
note h = q qa
hence "list_all (\<lambda>(acn, st). st \<le> (9 + qa + (21 + q)))
(shift (adjust0 (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)) (21 + q))"
proof(rule_tac tm_wf_shift t_twice_compile_def)
have "list_all (\<lambda>(acn, st). st \<le> Suc (length (tm_of abc_fourtimes @ shift
(mopup (Suc 0)) qa) div 2)) (adjust0 (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa))"
apply(rule_tac tm_wf_change_termi)
using wf_fourtimes h
apply(simp add: t_fourtimes_compile_def)
done
thus "list_all (\<lambda>(acn, st). st \<le> 9 + qa)
(adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa)
(Suc (length (tm_of abc_fourtimes @ shift (mopup (Suc 0)) qa) div
2)))"
using h
apply(simp)
done
qed
thus "list_all
(\<lambda>(acn, st). st \<le> 30 + (length (tm_of abc_twice) div 2 + length (tm_of abc_fourtimes) div 2))
(shift
(adjust (tm_of abc_fourtimes @ shift (mopup (Suc 0)) (length (tm_of abc_fourtimes) div 2))
(9 + length (tm_of abc_fourtimes) div 2))
(21 + length (tm_of abc_twice) div 2))"
apply(subgoal_tac "qa + q = q + qa")
apply(simp add: h)
apply(simp)
done
qed
thus "b \<le> (60 + (length (tm_of abc_twice) + length (tm_of abc_fourtimes))) div 2"
using g
apply(simp add: Ball_set[THEN sym])
apply(erule_tac x = "(a, b)" in ballE, simp, simp)
done
qed
lemma tm_wf_t_wcode_main[intro]: "tm_wf (t_wcode_main, 0)"
by(auto simp: t_wcode_main_def tm_wf.simps
t_twice_def t_fourtimes_def del: List.list_all_iff)
declare tm_comp.simps[simp del]
lemma prepare_mainpart_lemma:
"args \<noteq> [] \<Longrightarrow>
\<exists> stp ln rn. steps0 (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) stp
= (0, Bk # Oc\<up>(Suc m), Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<args>)) @ Bk\<up>(rn))"
proof -
let ?P1 = "(\<lambda> (l, r). (l::cell list) = [] \<and> r = <m # args>)"
let ?Q1 = "(\<lambda> (l, r). wprepare_stop m args (l, r))"
let ?P2 = ?Q1
let ?Q2 = "(\<lambda> (l, r). (\<exists> ln rn. l = Bk # Oc\<up>(Suc m) \<and>
r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<args>)) @ Bk\<up>(rn)))"
let ?P3 = "\<lambda> tp. False"
assume h: "args \<noteq> []"
have "{?P1} t_wcode_prepare |+| t_wcode_main {?Q2}"
proof(rule_tac Hoare_plus_halt)
show "{?P1} t_wcode_prepare {?Q1}"
proof(rule_tac Hoare_haltI, auto)
show "\<exists>n. is_final (steps0 (Suc 0, [], <m # args>) t_wcode_prepare n) \<and>
wprepare_stop m args holds_for steps0 (Suc 0, [], <m # args>) t_wcode_prepare n"
using wprepare_correctness[of args m,OF h]
apply(auto simp add: wprepare_inv.simps)
by (metis holds_for.simps is_finalI)
qed
next
show "{?P2} t_wcode_main {?Q2}"
proof(rule_tac Hoare_haltI, auto)
fix l r
assume "wprepare_stop m args (l, r)"
thus "\<exists>n. is_final (steps0 (Suc 0, l, r) t_wcode_main n) \<and>
(\<lambda>(l, r). l = Bk # Oc # Oc \<up> m \<and> (\<exists>ln rn. r = Bk # Oc # Bk \<up> ln @
Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn)) holds_for steps0 (Suc 0, l, r) t_wcode_main n"
proof(auto simp: wprepare_stop.simps)
fix rn
show " \<exists>n. is_final (steps0 (Suc 0, Bk # <rev args> @ Bk # Bk # Oc # Oc \<up> m, Bk # Oc # Bk \<up> rn) t_wcode_main n) \<and>
(\<lambda>(l, r). l = Bk # Oc # Oc \<up> m \<and>
(\<exists>ln rn. r = Bk # Oc # Bk \<up> ln @
Bk # Bk # Oc \<up> bl_bin (<args>) @
Bk \<up> rn)) holds_for steps0 (Suc 0, Bk # <rev args> @ Bk # Bk # Oc # Oc \<up> m, Bk # Oc # Bk \<up> rn) t_wcode_main n"
using t_wcode_main_lemma_pre[of "args" "<args>" 0 "Oc\<up>(Suc m)" 0 rn,OF h refl]
apply(auto simp: tape_of_nl_rev)
apply(rename_tac stp ln rna)
apply(rule_tac x = stp in exI, auto)
done
qed
qed
next
show "tm_wf0 t_wcode_prepare"
by auto
qed
then obtain n
where "\<And> tp. (case tp of (l, r) \<Rightarrow> l = [] \<and> r = <m # args>) \<longrightarrow>
(is_final (steps0 (1, tp) (t_wcode_prepare |+| t_wcode_main) n) \<and>
(\<lambda>(l, r).
\<exists>ln rn.
l = Bk # Oc \<up> Suc m \<and>
r = Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) holds_for steps0 (1, tp) (t_wcode_prepare |+| t_wcode_main) n)"
unfolding Hoare_halt_def by auto
thus "?thesis"
apply(rule_tac x = n in exI)
apply(case_tac "(steps0 (Suc 0, [], <m # args>)
(adjust0 t_wcode_prepare @ shift t_wcode_main (length t_wcode_prepare div 2)) n)")
apply(auto simp: tm_comp.simps)
done
qed
definition tinres :: "cell list \<Rightarrow> cell list \<Rightarrow> bool"
where
"tinres xs ys = (\<exists>n. xs = ys @ Bk \<up> n \<or> ys = xs @ Bk \<up> n)"
lemma tinres_fetch_congr[simp]: "tinres r r' \<Longrightarrow>
fetch t ss (read r) =
fetch t ss (read r')"
apply(simp add: fetch.simps, auto split: if_splits simp: tinres_def)
using hd_replicate apply fastforce
using hd_replicate apply fastforce
done
lemma nonempty_hd_tinres[simp]: "\<lbrakk>tinres r r'; r \<noteq> []; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r = hd r'"
apply(auto simp: tinres_def)
done
lemma tinres_nonempty[simp]:
"\<lbrakk>tinres r []; r \<noteq> []\<rbrakk> \<Longrightarrow> hd r = Bk"
"\<lbrakk>tinres [] r'; r' \<noteq> []\<rbrakk> \<Longrightarrow> hd r' = Bk"
"\<lbrakk>tinres r []; r \<noteq> []\<rbrakk> \<Longrightarrow> tinres (tl r) []"
"tinres r r' \<Longrightarrow> tinres (b # r) (b # r')"
by(auto simp: tinres_def)
lemma ex_move_tl[intro]: "\<exists>na. tl r = tl (r @ Bk\<up>(n)) @ Bk\<up>(na) \<or> tl (r @ Bk\<up>(n)) = tl r @ Bk\<up>(na)"
apply(case_tac r, simp)
by(case_tac n, auto)
lemma tinres_tails[simp]: "tinres r r' \<Longrightarrow> tinres (tl r) (tl r')"
apply(auto simp: tinres_def)
by(case_tac r', auto)
lemma tinres_empty[simp]:
"\<lbrakk>tinres [] r'\<rbrakk> \<Longrightarrow> tinres [] (tl r')"
"tinres r [] \<Longrightarrow> tinres (Bk # tl r) [Bk]"
"tinres r [] \<Longrightarrow> tinres (Oc # tl r) [Oc]"
by(auto simp: tinres_def)
lemma tinres_step2:
assumes "tinres r r'" "step0 (ss, l, r) t = (sa, la, ra)" "step0 (ss, l, r') t = (sb, lb, rb)"
shows "la = lb \<and> tinres ra rb \<and> sa = sb"
proof (cases "fetch t ss (read r')")
case (Pair a b)
have sa:"sa = sb" using assms Pair by(force simp: step.simps)
have "la = lb \<and> tinres ra rb" using assms Pair
by(cases a, auto simp: step.simps split: if_splits)
thus ?thesis using sa by auto
qed
lemma tinres_steps2:
"\<lbrakk>tinres r r'; steps0 (ss, l, r) t stp = (sa, la, ra); steps0 (ss, l, r') t stp = (sb, lb, rb)\<rbrakk>
\<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
proof(induct stp arbitrary: sa la ra sb lb rb)
case (Suc stp sa la ra sb lb rb)
then show ?case
apply(simp)
apply(case_tac "(steps0 (ss, l, r) t stp)")
apply(case_tac "(steps0 (ss, l, r') t stp)")
proof -
fix stp a b c aa ba ca
assume ind: "\<And>sa la ra sb lb rb. \<lbrakk>steps0 (ss, l, r) t stp = (sa, la, ra);
steps0 (ss, l, r') t stp = (sb, lb, rb)\<rbrakk> \<Longrightarrow> la = lb \<and> tinres ra rb \<and> sa = sb"
and h: " tinres r r'" "step0 (steps0 (ss, l, r) t stp) t = (sa, la, ra)"
"step0 (steps0 (ss, l, r') t stp) t = (sb, lb, rb)" "steps0 (ss, l, r) t stp = (a, b, c)"
"steps0 (ss, l, r') t stp = (aa, ba, ca)"
have "b = ba \<and> tinres c ca \<and> a = aa"
apply(rule_tac ind, simp_all add: h)
done
thus "la = lb \<and> tinres ra rb \<and> sa = sb"
apply(rule_tac l = b and r = c and ss = a and r' = ca
and t = t in tinres_step2)
using h
apply(simp, simp, simp)
done
qed
qed (simp add: steps.simps)
definition t_wcode_adjust :: "instr list"
where
"t_wcode_adjust = [(W1, 1), (R, 2), (Nop, 2), (R, 3), (R, 3), (R, 4),
(L, 8), (L, 5), (L, 6), (W0, 5), (L, 6), (R, 7),
(W1, 2), (Nop, 7), (L, 9), (W0, 8), (L, 9), (L, 10),
(L, 11), (L, 10), (R, 0), (L, 11)]"
lemma fetch_t_wcode_adjust[simp]:
"fetch t_wcode_adjust (Suc 0) Bk = (W1, 1)"
"fetch t_wcode_adjust (Suc 0) Oc = (R, 2)"
"fetch t_wcode_adjust (Suc (Suc 0)) Oc = (R, 3)"
"fetch t_wcode_adjust (Suc (Suc (Suc 0))) Oc = (R, 4)"
"fetch t_wcode_adjust (Suc (Suc (Suc 0))) Bk = (R, 3)"
"fetch t_wcode_adjust 4 Bk = (L, 8)"
"fetch t_wcode_adjust 4 Oc = (L, 5)"
"fetch t_wcode_adjust 5 Oc = (W0, 5)"
"fetch t_wcode_adjust 5 Bk = (L, 6)"
"fetch t_wcode_adjust 6 Oc = (R, 7)"
"fetch t_wcode_adjust 6 Bk = (L, 6)"
"fetch t_wcode_adjust 7 Bk = (W1, 2)"
"fetch t_wcode_adjust 8 Bk = (L, 9)"
"fetch t_wcode_adjust 8 Oc = (W0, 8)"
"fetch t_wcode_adjust 9 Oc = (L, 10)"
"fetch t_wcode_adjust 9 Bk = (L, 9)"
"fetch t_wcode_adjust 10 Bk = (L, 11)"
"fetch t_wcode_adjust 10 Oc = (L, 10)"
"fetch t_wcode_adjust 11 Oc = (L, 11)"
"fetch t_wcode_adjust 11 Bk = (R, 0)"
by(auto simp: fetch.simps t_wcode_adjust_def nth_of.simps numeral)
fun wadjust_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_start m rs (l, r) =
(\<exists> ln rn. l = Bk # Oc\<up>(Suc m) \<and>
tl r = Oc # Bk\<up>(ln) @ Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn))"
fun wadjust_loop_start :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_loop_start m rs (l, r) =
(\<exists> ln rn ml mr. l = Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
r = Oc # Bk\<up>(ln) @ Bk # Oc\<up>(mr) @ Bk\<up>(rn) \<and>
ml + mr = Suc (Suc rs) \<and> mr > 0)"
fun wadjust_loop_right_move :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_loop_right_move m rs (l, r) =
(\<exists> ml mr nl nr rn. l = Bk\<up>(nl) @ Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
r = Bk\<up>(nr) @ Oc\<up>(mr) @ Bk\<up>(rn) \<and>
ml + mr = Suc (Suc rs) \<and> mr > 0 \<and>
nl + nr > 0)"
fun wadjust_loop_check :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_loop_check m rs (l, r) =
(\<exists> ml mr ln rn. l = Oc # Bk\<up>(ln) @ Bk # Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
r = Oc\<up>(mr) @ Bk\<up>(rn) \<and> ml + mr = (Suc rs))"
fun wadjust_loop_erase :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_loop_erase m rs (l, r) =
(\<exists> ml mr ln rn. l = Bk\<up>(ln) @ Bk # Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
tl r = Oc\<up>(mr) @ Bk\<up>(rn) \<and> ml + mr = (Suc rs) \<and> mr > 0)"
fun wadjust_loop_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_loop_on_left_moving_O m rs (l, r) =
(\<exists> ml mr ln rn. l = Oc\<up>(ml) @ Bk # Oc\<up>(Suc m )\<and>
r = Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(mr) @ Bk\<up>(rn) \<and>
ml + mr = Suc rs \<and> mr > 0)"
fun wadjust_loop_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_loop_on_left_moving_B m rs (l, r) =
(\<exists> ml mr nl nr rn. l = Bk\<up>(nl) @ Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
r = Bk\<up>(nr) @ Bk # Bk # Oc\<up>(mr) @ Bk\<up>(rn) \<and>
ml + mr = Suc rs \<and> mr > 0)"
fun wadjust_loop_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_loop_on_left_moving m rs (l, r) =
(wadjust_loop_on_left_moving_O m rs (l, r) \<or>
wadjust_loop_on_left_moving_B m rs (l, r))"
fun wadjust_loop_right_move2 :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_loop_right_move2 m rs (l, r) =
(\<exists> ml mr ln rn. l = Oc # Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
r = Bk\<up>(ln) @ Bk # Bk # Oc\<up>(mr) @ Bk\<up>(rn) \<and>
ml + mr = Suc rs \<and> mr > 0)"
fun wadjust_erase2 :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_erase2 m rs (l, r) =
(\<exists> ln rn. l = Bk\<up>(ln) @ Bk # Oc # Oc\<up>(Suc rs) @ Bk # Oc\<up>(Suc m) \<and>
tl r = Bk\<up>(rn))"
fun wadjust_on_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_on_left_moving_O m rs (l, r) =
(\<exists> rn. l = Oc\<up>(Suc rs) @ Bk # Oc\<up>(Suc m) \<and>
r = Oc # Bk\<up>(rn))"
fun wadjust_on_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_on_left_moving_B m rs (l, r) =
(\<exists> ln rn. l = Bk\<up>(ln) @ Oc # Oc\<up>(Suc rs) @ Bk # Oc\<up>(Suc m) \<and>
r = Bk\<up>(rn))"
fun wadjust_on_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_on_left_moving m rs (l, r) =
(wadjust_on_left_moving_O m rs (l, r) \<or>
wadjust_on_left_moving_B m rs (l, r))"
fun wadjust_goon_left_moving_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_goon_left_moving_B m rs (l, r) =
(\<exists> rn. l = Oc\<up>(Suc m) \<and>
r = Bk # Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
fun wadjust_goon_left_moving_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_goon_left_moving_O m rs (l, r) =
(\<exists> ml mr rn. l = Oc\<up>(ml) @ Bk # Oc\<up>(Suc m) \<and>
r = Oc\<up>(mr) @ Bk\<up>(rn) \<and>
ml + mr = Suc (Suc rs) \<and> mr > 0)"
fun wadjust_goon_left_moving :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_goon_left_moving m rs (l, r) =
(wadjust_goon_left_moving_B m rs (l, r) \<or>
wadjust_goon_left_moving_O m rs (l, r))"
fun wadjust_backto_standard_pos_B :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_backto_standard_pos_B m rs (l, r) =
(\<exists> rn. l = [] \<and>
r = Bk # Oc\<up>(Suc m )@ Bk # Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
fun wadjust_backto_standard_pos_O :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_backto_standard_pos_O m rs (l, r) =
(\<exists> ml mr rn. l = Oc\<up>(ml) \<and>
r = Oc\<up>(mr) @ Bk # Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn) \<and>
ml + mr = Suc m \<and> mr > 0)"
fun wadjust_backto_standard_pos :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_backto_standard_pos m rs (l, r) =
(wadjust_backto_standard_pos_B m rs (l, r) \<or>
wadjust_backto_standard_pos_O m rs (l, r))"
fun wadjust_stop :: "nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_stop m rs (l, r) =
(\<exists> rn. l = [Bk] \<and>
r = Oc\<up>(Suc m )@ Bk # Oc\<up>(Suc (Suc rs)) @ Bk\<up>(rn))"
declare wadjust_start.simps[simp del] wadjust_loop_start.simps[simp del]
wadjust_loop_right_move.simps[simp del] wadjust_loop_check.simps[simp del]
wadjust_loop_erase.simps[simp del] wadjust_loop_on_left_moving.simps[simp del]
wadjust_loop_right_move2.simps[simp del] wadjust_erase2.simps[simp del]
wadjust_on_left_moving_O.simps[simp del] wadjust_on_left_moving_B.simps[simp del]
wadjust_on_left_moving.simps[simp del] wadjust_goon_left_moving_B.simps[simp del]
wadjust_goon_left_moving_O.simps[simp del] wadjust_goon_left_moving.simps[simp del]
wadjust_backto_standard_pos.simps[simp del] wadjust_backto_standard_pos_B.simps[simp del]
wadjust_backto_standard_pos_O.simps[simp del] wadjust_stop.simps[simp del]
fun wadjust_inv :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> tape \<Rightarrow> bool"
where
"wadjust_inv st m rs (l, r) =
(if st = Suc 0 then wadjust_start m rs (l, r)
else if st = Suc (Suc 0) then wadjust_loop_start m rs (l, r)
else if st = Suc (Suc (Suc 0)) then wadjust_loop_right_move m rs (l, r)
else if st = 4 then wadjust_loop_check m rs (l, r)
else if st = 5 then wadjust_loop_erase m rs (l, r)
else if st = 6 then wadjust_loop_on_left_moving m rs (l, r)
else if st = 7 then wadjust_loop_right_move2 m rs (l, r)
else if st = 8 then wadjust_erase2 m rs (l, r)
else if st = 9 then wadjust_on_left_moving m rs (l, r)
else if st = 10 then wadjust_goon_left_moving m rs (l, r)
else if st = 11 then wadjust_backto_standard_pos m rs (l, r)
else if st = 0 then wadjust_stop m rs (l, r)
else False
)"
declare wadjust_inv.simps[simp del]
fun wadjust_phase :: "nat \<Rightarrow> config \<Rightarrow> nat"
where
"wadjust_phase rs (st, l, r) =
(if st = 1 then 3
else if st \<ge> 2 \<and> st \<le> 7 then 2
else if st \<ge> 8 \<and> st \<le> 11 then 1
else 0)"
fun wadjust_stage :: "nat \<Rightarrow> config \<Rightarrow> nat"
where
"wadjust_stage rs (st, l, r) =
(if st \<ge> 2 \<and> st \<le> 7 then
rs - length (takeWhile (\<lambda> a. a = Oc)
(tl (dropWhile (\<lambda> a. a = Oc) (rev l @ r))))
else 0)"
fun wadjust_state :: "nat \<Rightarrow> config \<Rightarrow> nat"
where
"wadjust_state rs (st, l, r) =
(if st \<ge> 2 \<and> st \<le> 7 then 8 - st
else if st \<ge> 8 \<and> st \<le> 11 then 12 - st
else 0)"
fun wadjust_step :: "nat \<Rightarrow> config \<Rightarrow> nat"
where
"wadjust_step rs (st, l, r) =
(if st = 1 then (if hd r = Bk then 1
else 0)
else if st = 3 then length r
else if st = 5 then (if hd r = Oc then 1
else 0)
else if st = 6 then length l
else if st = 8 then (if hd r = Oc then 1
else 0)
else if st = 9 then length l
else if st = 10 then length l
else if st = 11 then (if hd r = Bk then 0
else Suc (length l))
else 0)"
fun wadjust_measure :: "(nat \<times> config) \<Rightarrow> nat \<times> nat \<times> nat \<times> nat"
where
"wadjust_measure (rs, (st, l, r)) =
(wadjust_phase rs (st, l, r),
wadjust_stage rs (st, l, r),
wadjust_state rs (st, l, r),
wadjust_step rs (st, l, r))"
definition wadjust_le :: "((nat \<times> config) \<times> nat \<times> config) set"
where "wadjust_le \<equiv> (inv_image lex_square wadjust_measure)"
lemma wf_lex_square[intro]: "wf lex_square"
by(auto intro:wf_lex_prod simp: Abacus.lex_pair_def lex_square_def
Abacus.lex_triple_def)
lemma wf_wadjust_le[intro]: "wf wadjust_le"
by(auto intro:wf_inv_image simp: wadjust_le_def
Abacus.lex_triple_def Abacus.lex_pair_def)
lemma wadjust_start_snd_nonempty[simp]: "wadjust_start m rs (c, []) = False"
apply(auto simp: wadjust_start.simps)
done
lemma wadjust_loop_right_move_fst_nonempty[simp]: "wadjust_loop_right_move m rs (c, []) \<Longrightarrow> c \<noteq> []"
apply(auto simp: wadjust_loop_right_move.simps)
done
lemma wadjust_loop_check_fst_nonempty[simp]: "wadjust_loop_check m rs (c, []) \<Longrightarrow> c \<noteq> []"
apply(simp only: wadjust_loop_check.simps, auto)
done
lemma wadjust_loop_start_snd_nonempty[simp]: "wadjust_loop_start m rs (c, []) = False"
apply(simp add: wadjust_loop_start.simps)
done
lemma wadjust_erase2_singleton[simp]: "wadjust_loop_check m rs (c, []) \<Longrightarrow> wadjust_erase2 m rs (tl c, [hd c])"
apply(simp only: wadjust_loop_check.simps wadjust_erase2.simps, auto)
done
lemma wadjust_loop_on_left_moving_snd_nonempty[simp]:
"wadjust_loop_on_left_moving m rs (c, []) = False"
"wadjust_loop_right_move2 m rs (c, []) = False"
"wadjust_erase2 m rs ([], []) = False"
by(auto simp: wadjust_loop_on_left_moving.simps
wadjust_loop_right_move2.simps
wadjust_erase2.simps)
lemma wadjust_on_left_moving_B_Bk1[simp]: "wadjust_on_left_moving_B m rs
(Oc # Oc # Oc\<up>(rs) @ Bk # Oc # Oc\<up>(m), [Bk])"
apply(simp add: wadjust_on_left_moving_B.simps, auto)
done
lemma wadjust_on_left_moving_B_Bk2[simp]: "wadjust_on_left_moving_B m rs
(Bk\<up>(n) @ Bk # Oc # Oc # Oc\<up>(rs) @ Bk # Oc # Oc\<up>(m), [Bk])"
apply(simp add: wadjust_on_left_moving_B.simps , auto)
apply(rule_tac x = "Suc n" in exI, simp add: exp_ind del: replicate_Suc)
done
lemma wadjust_on_left_moving_singleton[simp]: "\<lbrakk>wadjust_erase2 m rs (c, []); c \<noteq> []\<rbrakk> \<Longrightarrow>
wadjust_on_left_moving m rs (tl c, [hd c])" unfolding wadjust_erase2.simps
apply(auto simp add: wadjust_on_left_moving.simps)
apply (metis (no_types, lifting) empty_replicate hd_append hd_replicate list.sel(1) list.sel(3)
self_append_conv2 tl_append2 tl_replicate
wadjust_on_left_moving_B_Bk1 wadjust_on_left_moving_B_Bk2)+
done
lemma wadjust_erase2_cases[simp]: "wadjust_erase2 m rs (c, [])
\<Longrightarrow> (c = [] \<longrightarrow> wadjust_on_left_moving m rs ([], [Bk])) \<and>
(c \<noteq> [] \<longrightarrow> wadjust_on_left_moving m rs (tl c, [hd c]))"
apply(auto)
done
lemma wadjust_on_left_moving_nonempty[simp]:
"wadjust_on_left_moving m rs ([], []) = False"
"wadjust_on_left_moving_O m rs (c, []) = False"
apply(auto simp: wadjust_on_left_moving.simps
wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps)
done
lemma wadjust_on_left_moving_B_singleton_Bk[simp]:
" \<lbrakk>wadjust_on_left_moving_B m rs (c, []); c \<noteq> []; hd c = Bk\<rbrakk> \<Longrightarrow>
wadjust_on_left_moving_B m rs (tl c, [Bk])"
apply(auto simp add: wadjust_on_left_moving_B.simps hd_append)
by (metis cell.distinct(1) empty_replicate list.sel(1) tl_append2 tl_replicate)
lemma wadjust_on_left_moving_B_singleton_Oc[simp]:
"\<lbrakk>wadjust_on_left_moving_B m rs (c, []); c \<noteq> []; hd c = Oc\<rbrakk> \<Longrightarrow>
wadjust_on_left_moving_O m rs (tl c, [Oc])"
apply(auto simp add: wadjust_on_left_moving_B.simps wadjust_on_left_moving_O.simps hd_append)
apply (metis cell.distinct(1) empty_replicate hd_replicate list.sel(3) self_append_conv2)+
done
lemma wadjust_on_left_moving_singleton2[simp]:
"\<lbrakk>wadjust_on_left_moving m rs (c, []); c \<noteq> []\<rbrakk> \<Longrightarrow>
wadjust_on_left_moving m rs (tl c, [hd c])"
apply(simp add: wadjust_on_left_moving.simps)
apply(case_tac "hd c", simp_all)
done
lemma wadjust_nonempty[simp]: "wadjust_goon_left_moving m rs (c, []) = False"
"wadjust_backto_standard_pos m rs (c, []) = False"
by(auto simp: wadjust_goon_left_moving.simps wadjust_goon_left_moving_B.simps
wadjust_goon_left_moving_O.simps wadjust_backto_standard_pos.simps
wadjust_backto_standard_pos_B.simps wadjust_backto_standard_pos_O.simps)
lemma wadjust_loop_start_no_Bk[simp]: "wadjust_loop_start m rs (c, Bk # list) = False"
apply(auto simp: wadjust_loop_start.simps)
done
lemma wadjust_loop_check_nonempty[simp]: "wadjust_loop_check m rs (c, b) \<Longrightarrow> c \<noteq> []"
apply(simp only: wadjust_loop_check.simps, auto)
done
lemma wadjust_erase2_via_loop_check_Bk[simp]: "wadjust_loop_check m rs (c, Bk # list)
\<Longrightarrow> wadjust_erase2 m rs (tl c, hd c # Bk # list)"
by (auto simp: wadjust_loop_check.simps wadjust_erase2.simps)
declare wadjust_loop_on_left_moving_O.simps[simp del]
wadjust_loop_on_left_moving_B.simps[simp del]
lemma wadjust_loop_on_left_moving_B_via_erase[simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); hd c = Bk\<rbrakk>
\<Longrightarrow> wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)"
unfolding wadjust_loop_erase.simps wadjust_loop_on_left_moving_B.simps
apply(erule_tac exE)+
apply(rename_tac ml mr ln rn)
apply(rule_tac x = ml in exI, rule_tac x = mr in exI,
rule_tac x = ln in exI, rule_tac x = 0 in exI)
apply(case_tac ln, auto)
apply(simp add: exp_ind [THEN sym])
done
lemma wadjust_loop_on_left_moving_O_Bk_via_erase[simp]:
"\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); c \<noteq> []; hd c = Oc\<rbrakk> \<Longrightarrow>
wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)"
apply(auto simp: wadjust_loop_erase.simps wadjust_loop_on_left_moving_O.simps)
by (metis cell.distinct(1) empty_replicate hd_append hd_replicate list.sel(1))
lemma wadjust_loop_on_left_moving_Bk_via_erase[simp]: "\<lbrakk>wadjust_loop_erase m rs (c, Bk # list); c \<noteq> []\<rbrakk> \<Longrightarrow>
wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)"
apply(case_tac "hd c", simp_all add:wadjust_loop_on_left_moving.simps)
done
lemma wadjust_loop_on_left_moving_B_Bk_move[simp]:
"\<lbrakk>wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
\<Longrightarrow> wadjust_loop_on_left_moving_B m rs (tl c, Bk # Bk # list)"
apply(simp only: wadjust_loop_on_left_moving_B.simps)
apply(erule_tac exE)+
by (metis (no_types, lifting) cell.distinct(1) list.sel(1)
replicate_Suc_iff_anywhere self_append_conv2 tl_append2 tl_replicate)
lemma wadjust_loop_on_left_moving_O_Oc_move[simp]:
"\<lbrakk>wadjust_loop_on_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
\<Longrightarrow> wadjust_loop_on_left_moving_O m rs (tl c, Oc # Bk # list)"
apply(simp only: wadjust_loop_on_left_moving_O.simps
wadjust_loop_on_left_moving_B.simps)
by (metis cell.distinct(1) empty_replicate hd_append hd_replicate list.sel(3) self_append_conv2)
lemma wadjust_loop_erase_nonempty[simp]: "wadjust_loop_erase m rs (c, b) \<Longrightarrow> c \<noteq> []"
"wadjust_loop_on_left_moving m rs (c, b) \<Longrightarrow> c \<noteq> []"
"wadjust_loop_right_move2 m rs (c, b) \<Longrightarrow> c \<noteq> []"
"wadjust_erase2 m rs (c, Bk # list) \<Longrightarrow> c \<noteq> []"
"wadjust_on_left_moving m rs (c,b) \<Longrightarrow> c \<noteq> []"
"wadjust_on_left_moving_O m rs (c, Bk # list) = False"
"wadjust_goon_left_moving m rs (c, b) \<Longrightarrow> c \<noteq> []"
"wadjust_loop_on_left_moving_O m rs (c, Bk # list) = False"
by(auto simp: wadjust_loop_erase.simps wadjust_loop_on_left_moving.simps
wadjust_loop_on_left_moving_O.simps wadjust_loop_on_left_moving_B.simps
wadjust_loop_right_move2.simps wadjust_erase2.simps
wadjust_on_left_moving.simps
wadjust_on_left_moving_O.simps
wadjust_on_left_moving_B.simps wadjust_goon_left_moving.simps
wadjust_goon_left_moving_B.simps
wadjust_goon_left_moving_O.simps)
lemma wadjust_loop_on_left_moving_Bk_move[simp]:
"wadjust_loop_on_left_moving m rs (c, Bk # list)
\<Longrightarrow> wadjust_loop_on_left_moving m rs (tl c, hd c # Bk # list)"
apply(simp add: wadjust_loop_on_left_moving.simps)
apply(case_tac "hd c", simp_all)
done
lemma wadjust_loop_start_Oc_via_Bk_move[simp]:
"wadjust_loop_right_move2 m rs (c, Bk # list) \<Longrightarrow> wadjust_loop_start m rs (c, Oc # list)"
apply(auto simp: wadjust_loop_right_move2.simps wadjust_loop_start.simps replicate_app_Cons_same)
by (metis add_Suc replicate_Suc)
lemma wadjust_on_left_moving_Bk_via_erase[simp]: "wadjust_erase2 m rs (c, Bk # list) \<Longrightarrow>
wadjust_on_left_moving m rs (tl c, hd c # Bk # list)"
apply(auto simp: wadjust_erase2.simps wadjust_on_left_moving.simps replicate_app_Cons_same
wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps)
apply (metis exp_ind replicate_append_same)+
done
lemma wadjust_on_left_moving_B_Bk_drop_one: "\<lbrakk>wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Bk\<rbrakk>
\<Longrightarrow> wadjust_on_left_moving_B m rs (tl c, Bk # Bk # list)"
apply(auto simp: wadjust_on_left_moving_B.simps)
by (metis cell.distinct(1) hd_append list.sel(1) tl_append2 tl_replicate)
lemma wadjust_on_left_moving_B_Bk_drop_Oc: "\<lbrakk>wadjust_on_left_moving_B m rs (c, Bk # list); hd c = Oc\<rbrakk>
\<Longrightarrow> wadjust_on_left_moving_O m rs (tl c, Oc # Bk # list)"
apply(auto simp: wadjust_on_left_moving_O.simps wadjust_on_left_moving_B.simps)
by (metis cell.distinct(1) empty_replicate hd_append hd_replicate list.sel(3) self_append_conv2)
lemma wadjust_on_left_moving_B_drop[simp]: "wadjust_on_left_moving m rs (c, Bk # list) \<Longrightarrow>
wadjust_on_left_moving m rs (tl c, hd c # Bk # list)"
by(cases "hd c", auto simp:wadjust_on_left_moving.simps wadjust_on_left_moving_B_Bk_drop_one
wadjust_on_left_moving_B_Bk_drop_Oc)
lemma wadjust_goon_left_moving_O_no_Bk[simp]: "wadjust_goon_left_moving_O m rs (c, Bk # list) = False"
by (auto simp add: wadjust_goon_left_moving_O.simps)
lemma wadjust_backto_standard_pos_via_left_Bk[simp]:
"wadjust_goon_left_moving m rs (c, Bk # list) \<Longrightarrow>
wadjust_backto_standard_pos m rs (tl c, hd c # Bk # list)"
by(case_tac "hd c", simp_all add: wadjust_backto_standard_pos.simps wadjust_goon_left_moving.simps
wadjust_goon_left_moving_B.simps wadjust_backto_standard_pos_O.simps)
lemma wadjust_loop_right_move_Oc[simp]:
"wadjust_loop_start m rs (c, Oc # list) \<Longrightarrow> wadjust_loop_right_move m rs (Oc # c, list)"
apply(auto simp add: wadjust_loop_start.simps wadjust_loop_right_move.simps
simp del:split_head_repeat)
apply(rename_tac ln rn ml mr)
apply(rule_tac x = ml in exI, rule_tac x = mr in exI,
rule_tac x = 0 in exI, simp)
apply(rule_tac x = "Suc ln" in exI, simp add: exp_ind del: replicate_Suc)
done
lemma wadjust_loop_check_Oc[simp]:
assumes "wadjust_loop_right_move m rs (c, Oc # list)"
shows "wadjust_loop_check m rs (Oc # c, list)"
proof -
from assms obtain ml mr nl nr rn
where "c = Bk \<up> nl @ Oc # Oc \<up> ml @ Bk # Oc \<up> m @ [Oc]"
"Oc # list = Bk \<up> nr @ Oc \<up> mr @ Bk \<up> rn"
"ml + mr = Suc (Suc rs)" "0 < mr" "0 < nl + nr"
unfolding wadjust_loop_right_move.simps exp_ind
wadjust_loop_check.simps by auto
hence "\<exists>ln. Oc # c = Oc # Bk \<up> ln @ Bk # Oc # Oc \<up> ml @ Bk # Oc \<up> Suc m"
"\<exists>rn. list = Oc \<up> (mr - 1) @ Bk \<up> rn" "ml + (mr - 1) = Suc rs"
by(cases nl;cases nr;cases mr;force simp add: wadjust_loop_right_move.simps exp_ind
wadjust_loop_check.simps replicate_append_same)+
thus ?thesis unfolding wadjust_loop_check.simps by auto
qed
lemma wadjust_loop_erase_move_Oc[simp]: "wadjust_loop_check m rs (c, Oc # list) \<Longrightarrow>
wadjust_loop_erase m rs (tl c, hd c # Oc # list)"
apply(simp only: wadjust_loop_check.simps wadjust_loop_erase.simps)
apply(erule_tac exE)+
using Cons_replicate_eq by fastforce
lemma wadjust_loop_on_move_no_Oc[simp]:
"wadjust_loop_on_left_moving_B m rs (c, Oc # list) = False"
"wadjust_loop_right_move2 m rs (c, Oc # list) = False"
"wadjust_loop_on_left_moving m rs (c, Oc # list)
\<Longrightarrow> wadjust_loop_right_move2 m rs (Oc # c, list)"
"wadjust_on_left_moving_B m rs (c, Oc # list) = False"
"wadjust_loop_erase m rs (c, Oc # list) \<Longrightarrow>
wadjust_loop_erase m rs (c, Bk # list)"
by(auto simp: wadjust_loop_on_left_moving_B.simps wadjust_loop_on_left_moving_O.simps
wadjust_loop_right_move2.simps replicate_app_Cons_same wadjust_loop_on_left_moving.simps
wadjust_on_left_moving_B.simps wadjust_loop_erase.simps)
lemma wadjust_goon_left_moving_B_Bk_Oc: "\<lbrakk>wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Bk\<rbrakk> \<Longrightarrow>
wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)"
apply(auto simp: wadjust_on_left_moving_O.simps
wadjust_goon_left_moving_B.simps )
done
lemma wadjust_goon_left_moving_O_Oc_Oc: "\<lbrakk>wadjust_on_left_moving_O m rs (c, Oc # list); hd c = Oc\<rbrakk>
\<Longrightarrow> wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)"
apply(auto simp: wadjust_on_left_moving_O.simps
wadjust_goon_left_moving_O.simps )
apply(auto simp: numeral_2_eq_2)
done
lemma wadjust_goon_left_moving_Oc[simp]: "wadjust_on_left_moving m rs (c, Oc # list) \<Longrightarrow>
wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
by(cases "hd c"; force simp: wadjust_on_left_moving.simps wadjust_goon_left_moving.simps
wadjust_goon_left_moving_B_Bk_Oc wadjust_goon_left_moving_O_Oc_Oc)+
lemma left_moving_Bk_Oc[simp]: "\<lbrakk>wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Bk\<rbrakk>
\<Longrightarrow> wadjust_goon_left_moving_B m rs (tl c, Bk # Oc # list)"
apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps hd_append
dest!: gr0_implies_Suc)
apply (metis cell.distinct(1) empty_replicate hd_replicate list.sel(3) self_append_conv2)
by (metis add_cancel_right_left cell.distinct(1) hd_replicate replicate_Suc_iff_anywhere)
lemma left_moving_Oc_Oc[simp]: "\<lbrakk>wadjust_goon_left_moving_O m rs (c, Oc # list); hd c = Oc\<rbrakk> \<Longrightarrow>
wadjust_goon_left_moving_O m rs (tl c, Oc # Oc # list)"
apply(auto simp: wadjust_goon_left_moving_O.simps wadjust_goon_left_moving_B.simps)
apply(rename_tac mlx mrx rnx)
apply(rule_tac x = "mlx - 1" in exI, simp)
apply(case_tac mlx, simp_all add: )
apply(rule_tac x = "Suc mrx" in exI, auto simp: )
done
lemma wadjust_goon_left_moving_B_no_Oc[simp]:
"wadjust_goon_left_moving_B m rs (c, Oc # list) = False"
apply(auto simp: wadjust_goon_left_moving_B.simps)
done
lemma wadjust_goon_left_moving_Oc_move[simp]: "wadjust_goon_left_moving m rs (c, Oc # list) \<Longrightarrow>
wadjust_goon_left_moving m rs (tl c, hd c # Oc # list)"
by(cases "hd c",auto simp: wadjust_goon_left_moving.simps)
lemma wadjust_backto_standard_pos_B_no_Oc[simp]:
"wadjust_backto_standard_pos_B m rs (c, Oc # list) = False"
apply(simp add: wadjust_backto_standard_pos_B.simps)
done
lemma wadjust_backto_standard_pos_O_no_Bk[simp]:
"wadjust_backto_standard_pos_O m rs (c, Bk # xs) = False"
by(simp add: wadjust_backto_standard_pos_O.simps)
lemma wadjust_backto_standard_pos_B_Bk_Oc[simp]:
"wadjust_backto_standard_pos_O m rs ([], Oc # list) \<Longrightarrow>
wadjust_backto_standard_pos_B m rs ([], Bk # Oc # list)"
apply(auto simp: wadjust_backto_standard_pos_O.simps
wadjust_backto_standard_pos_B.simps)
done
lemma wadjust_backto_standard_pos_B_Bk_Oc_via_O[simp]:
"\<lbrakk>wadjust_backto_standard_pos_O m rs (c, Oc # list); c \<noteq> []; hd c = Bk\<rbrakk>
\<Longrightarrow> wadjust_backto_standard_pos_B m rs (tl c, Bk # Oc # list)"
apply(simp add:wadjust_backto_standard_pos_O.simps
wadjust_backto_standard_pos_B.simps, auto)
done
lemma wadjust_backto_standard_pos_B_Oc_Oc_via_O[simp]: "\<lbrakk>wadjust_backto_standard_pos_O m rs (c, Oc # list); c \<noteq> []; hd c = Oc\<rbrakk>
\<Longrightarrow> wadjust_backto_standard_pos_O m rs (tl c, Oc # Oc # list)"
apply(simp add: wadjust_backto_standard_pos_O.simps, auto)
by force
lemma wadjust_backto_standard_pos_cases[simp]: "wadjust_backto_standard_pos m rs (c, Oc # list)
\<Longrightarrow> (c = [] \<longrightarrow> wadjust_backto_standard_pos m rs ([], Bk # Oc # list)) \<and>
(c \<noteq> [] \<longrightarrow> wadjust_backto_standard_pos m rs (tl c, hd c # Oc # list))"
apply(auto simp: wadjust_backto_standard_pos.simps)
apply(case_tac "hd c", simp_all)
done
lemma wadjust_loop_right_move_nonempty_snd[simp]: "wadjust_loop_right_move m rs (c, []) = False"
proof -
{fix nl ml mr rn nr
have "(c = Bk \<up> nl @ Oc # Oc \<up> ml @ Bk # Oc \<up> Suc m \<and>
[] = Bk \<up> nr @ Oc \<up> mr @ Bk \<up> rn \<and> ml + mr = Suc (Suc rs) \<and> 0 < mr \<and> 0 < nl + nr) =
False" by auto
} note t=this
thus ?thesis unfolding wadjust_loop_right_move.simps t by blast
qed
lemma wadjust_loop_erase_nonempty_snd[simp]: "wadjust_loop_erase m rs (c, []) = False"
apply(simp only: wadjust_loop_erase.simps, auto)
done
lemma wadjust_loop_erase_cases2[simp]: "\<lbrakk>Suc (Suc rs) = a; wadjust_loop_erase m rs (c, Bk # list)\<rbrakk>
\<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list))))
< a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list)))) \<or>
a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) =
a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
apply(simp only: wadjust_loop_erase.simps)
apply(rule_tac disjI2)
apply(case_tac c, simp, simp)
done
lemma dropWhile_exp1: "dropWhile (\<lambda>a. a = Oc) (Oc\<up>(n) @ xs) = dropWhile (\<lambda>a. a = Oc) xs"
apply(induct n, simp_all add: )
done
lemma takeWhile_exp1: "takeWhile (\<lambda>a. a = Oc) (Oc\<up>(n) @ xs) = Oc\<up>(n) @ takeWhile (\<lambda>a. a = Oc) xs"
apply(induct n, simp_all add: )
done
lemma wadjust_correctness_helper_1:
assumes "Suc (Suc rs) = a" " wadjust_loop_right_move2 m rs (c, Bk # list)"
shows "a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list))))
< a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
proof -
have "ml + mr = Suc rs \<Longrightarrow> 0 < mr \<Longrightarrow>
rs - (ml + length (takeWhile (\<lambda>a. a = Oc) list))
< Suc rs -
(ml +
length
(takeWhile (\<lambda>a. a = Oc)
(Bk \<up> ln @ Bk # Bk # Oc \<up> mr @ Bk \<up> rn))) "
for ml mr ln rn
by(cases ln, auto)
thus ?thesis using assms
by (auto simp: wadjust_loop_right_move2.simps dropWhile_exp1 takeWhile_exp1)
qed
lemma wadjust_correctness_helper_2:
"\<lbrakk>Suc (Suc rs) = a; wadjust_loop_on_left_moving m rs (c, Bk # list)\<rbrakk>
\<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list))))
< a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list)))) \<or>
a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Bk # list)))) =
a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))"
apply(subgoal_tac "c \<noteq> []")
apply(case_tac c, simp_all)
done
lemma wadjust_loop_check_empty_false[simp]: "wadjust_loop_check m rs ([], b) = False"
apply(simp add: wadjust_loop_check.simps)
done
lemma wadjust_loop_check_cases: "\<lbrakk>Suc (Suc rs) = a; wadjust_loop_check m rs (c, Oc # list)\<rbrakk>
\<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Oc # list))))
< a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list)))) \<or>
a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev (tl c) @ hd c # Oc # list)))) =
a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list))))"
apply(case_tac "c", simp_all)
done
lemma wadjust_loop_erase_cases_or:
"\<lbrakk>Suc (Suc rs) = a; wadjust_loop_erase m rs (c, Oc # list)\<rbrakk>
\<Longrightarrow> a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list))))
< a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list)))) \<or>
a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Bk # list)))) =
a - length (takeWhile (\<lambda>a. a = Oc) (tl (dropWhile (\<lambda>a. a = Oc) (rev c @ Oc # list))))"
apply(simp add: wadjust_loop_erase.simps)
apply(rule_tac disjI2)
apply(auto)
apply(simp add: dropWhile_exp1 takeWhile_exp1)
done
lemmas wadjust_correctness_helpers = wadjust_correctness_helper_2 wadjust_correctness_helper_1 wadjust_loop_erase_cases_or wadjust_loop_check_cases
declare numeral_2_eq_2[simp del]
lemma wadjust_start_Oc[simp]: "wadjust_start m rs (c, Bk # list)
\<Longrightarrow> wadjust_start m rs (c, Oc # list)"
apply(auto simp: wadjust_start.simps)
done
lemma wadjust_stop_Bk[simp]: "wadjust_backto_standard_pos m rs (c, Bk # list)
\<Longrightarrow> wadjust_stop m rs (Bk # c, list)"
apply(auto simp: wadjust_backto_standard_pos.simps
wadjust_stop.simps wadjust_backto_standard_pos_B.simps)
done
lemma wadjust_loop_start_Oc[simp]:
assumes "wadjust_start m rs (c, Oc # list)"
shows "wadjust_loop_start m rs (Oc # c, list)"
proof -
from assms[unfolded wadjust_start.simps] obtain ln rn where
"c = Bk # Oc # Oc \<up> m" "list = Oc # Bk \<up> ln @ Bk # Oc # Oc \<up> rs @ Bk \<up> rn"
by(auto)
hence "Oc # c = Oc \<up> 1 @ Bk # Oc \<up> Suc m \<and>
list = Oc # Bk \<up> ln @ Bk # Oc \<up>Suc rs @ Bk \<up> rn \<and> 1 + (Suc rs) = Suc (Suc rs) \<and> 0 < Suc rs"
by auto
thus ?thesis unfolding wadjust_loop_start.simps by blast
qed
lemma erase2_Bk_if_Oc[simp]:" wadjust_erase2 m rs (c, Oc # list)
\<Longrightarrow> wadjust_erase2 m rs (c, Bk # list)"
apply(auto simp: wadjust_erase2.simps)
done
lemma wadjust_loop_right_move_Bk[simp]: "wadjust_loop_right_move m rs (c, Bk # list)
\<Longrightarrow> wadjust_loop_right_move m rs (Bk # c, list)"
apply(simp only: wadjust_loop_right_move.simps)
apply(erule_tac exE)+
apply auto
apply (metis cell.distinct(1) empty_replicate hd_append hd_replicate less_SucI
list.sel(1) list.sel(3) neq0_conv replicate_Suc_iff_anywhere tl_append2 tl_replicate)+
done
lemma wadjust_correctness:
shows "let P = (\<lambda> (len, st, l, r). st = 0) in
let Q = (\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)) in
let f = (\<lambda> stp. (Suc (Suc rs), steps0 (Suc 0, Bk # Oc\<up>(Suc m),
Bk # Oc # Bk\<up>(ln) @ Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn)) t_wcode_adjust stp)) in
\<exists> n .P (f n) \<and> Q (f n)"
proof -
let ?P = "(\<lambda> (len, st, l, r). st = 0)"
let ?Q = "\<lambda> (len, st, l, r). wadjust_inv st m rs (l, r)"
let ?f = "\<lambda> stp. (Suc (Suc rs), steps0 (Suc 0, Bk # Oc\<up>(Suc m),
Bk # Oc # Bk\<up>(ln) @ Bk # Oc\<up>(Suc rs) @ Bk\<up>(rn)) t_wcode_adjust stp)"
have "\<exists> n. ?P (?f n) \<and> ?Q (?f n)"
proof(rule_tac halt_lemma2)
show "wf wadjust_le" by auto
next
{ fix n assume a:"\<not> ?P (?f n) \<and> ?Q (?f n)"
have "?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wadjust_le"
proof(cases "?f n")
case (fields a b c d)
then show ?thesis proof(cases d)
case Nil
then show ?thesis using a fields apply(simp add: step.simps)
apply(simp_all only: wadjust_inv.simps split: if_splits)
apply(simp_all add: wadjust_inv.simps wadjust_le_def
wadjust_correctness_helpers
Abacus.lex_triple_def Abacus.lex_pair_def lex_square_def split: if_splits).
next
case (Cons aa list)
then show ?thesis using a fields Nil Cons
apply((case_tac aa); simp add: step.simps)
apply(simp_all only: wadjust_inv.simps split: if_splits)
apply(simp_all)
apply(simp_all add: wadjust_inv.simps wadjust_le_def
wadjust_correctness_helpers
Abacus.lex_triple_def Abacus.lex_pair_def lex_square_def split: if_splits).
qed
qed
}
thus "\<forall> n. \<not> ?P (?f n) \<and> ?Q (?f n) \<longrightarrow>
?Q (?f (Suc n)) \<and> (?f (Suc n), ?f n) \<in> wadjust_le" by auto
next
show "?Q (?f 0)" by(auto simp add: steps.simps wadjust_inv.simps wadjust_start.simps)
next
show "\<not> ?P (?f 0)" by (simp add: steps.simps)
qed
thus"?thesis" by simp
qed
lemma tm_wf_t_wcode_adjust[intro]: "tm_wf (t_wcode_adjust, 0)"
by(auto simp: t_wcode_adjust_def tm_wf.simps)
lemma bl_bin_nonzero[simp]: "args \<noteq> [] \<Longrightarrow> bl_bin (<args::nat list>) > 0"
by(cases args)
(auto simp: tape_of_nl_cons bl_bin.simps)
lemma wcode_lemma_pre':
"args \<noteq> [] \<Longrightarrow>
\<exists> stp rn. steps0 (Suc 0, [], <m # args>)
((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) stp
= (0, [Bk], Oc\<up>(Suc m) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn))"
proof -
let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <m # args>"
let ?Q1 = "\<lambda>(l, r). l = Bk # Oc\<up>(Suc m) \<and>
(\<exists>ln rn. r = Bk # Oc # Bk\<up>(ln) @ Bk # Bk # Oc\<up>(bl_bin (<args>)) @ Bk\<up>(rn))"
let ?P2 = ?Q1
let ?Q2 = "\<lambda> (l, r). (wadjust_stop m (bl_bin (<args>) - 1) (l, r))"
let ?P3 = "\<lambda> tp. False"
assume h: "args \<noteq> []"
hence a: "bl_bin (<args>) > 0"
using h by simp
hence "{?P1} (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust {?Q2}"
proof(rule_tac Hoare_plus_halt)
next
show "tm_wf (t_wcode_prepare |+| t_wcode_main, 0)"
by(rule_tac tm_comp_wf, auto)
next
show "{?P1} t_wcode_prepare |+| t_wcode_main {?Q1}"
proof(rule_tac Hoare_haltI, auto)
show
"\<exists>n. is_final (steps0 (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) n) \<and>
(\<lambda>(l, r). l = Bk # Oc # Oc \<up> m \<and>
(\<exists>ln rn. r = Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn))
holds_for steps0 (Suc 0, [], <m # args>) (t_wcode_prepare |+| t_wcode_main) n"
using h prepare_mainpart_lemma[of args m]
apply(auto) apply(rename_tac stp ln rn)
apply(rule_tac x = stp in exI, simp)
apply(rule_tac x = ln in exI, auto)
done
qed
next
show "{?P2} t_wcode_adjust {?Q2}"
proof(rule_tac Hoare_haltI, auto del: replicate_Suc)
fix ln rn
obtain n a b where "steps0
(Suc 0, Bk # Oc \<up> m @ [Oc],
Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> (bl_bin (<args>) - Suc 0) @ Oc # Bk \<up> rn)
t_wcode_adjust n = (0, a, b)"
"wadjust_inv 0 m (bl_bin (<args>) - Suc 0) (a, b)"
using wadjust_correctness[of m "bl_bin (<args>) - 1" "Suc ln" rn,unfolded Let_def]
by(simp del: replicate_Suc add: replicate_Suc[THEN sym] exp_ind, auto)
thus "\<exists>n. is_final (steps0 (Suc 0, Bk # Oc # Oc \<up> m,
Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_wcode_adjust n) \<and>
wadjust_stop m (bl_bin (<args>) - Suc 0) holds_for steps0
(Suc 0, Bk # Oc # Oc \<up> m, Bk # Oc # Bk \<up> ln @ Bk # Bk # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_wcode_adjust n"
apply(rule_tac x = n in exI)
using a
apply(case_tac "bl_bin (<args>)", simp, simp del: replicate_Suc add: exp_ind wadjust_inv.simps)
by (simp add: replicate_append_same)
qed
qed
thus "?thesis"
apply(simp add: Hoare_halt_def, auto)
apply(rename_tac n)
apply(case_tac "(steps0 (Suc 0, [], <(m::nat) # args>)
((t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust) n)")
apply(rule_tac x = n in exI, auto simp: wadjust_stop.simps)
using a
apply(case_tac "bl_bin (<args>)", simp_all)
done
qed
text \<open>
The initialization TM \<open>t_wcode\<close>.
\<close>
definition t_wcode :: "instr list"
where
"t_wcode = (t_wcode_prepare |+| t_wcode_main) |+| t_wcode_adjust "
text \<open>
The correctness of \<open>t_wcode\<close>.
\<close>
lemma wcode_lemma_1:
"args \<noteq> [] \<Longrightarrow>
\<exists> stp ln rn. steps0 (Suc 0, [], <m # args>) (t_wcode) stp =
(0, [Bk], Oc\<up>(Suc m) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn))"
apply(simp add: wcode_lemma_pre' t_wcode_def del: replicate_Suc)
done
lemma wcode_lemma:
"args \<noteq> [] \<Longrightarrow>
\<exists> stp ln rn. steps0 (Suc 0, [], <m # args>) (t_wcode) stp =
(0, [Bk], <[m ,bl_bin (<args>)]> @ Bk\<up>(rn))"
using wcode_lemma_1[of args m]
apply(simp add: t_wcode_def tape_of_list_def tape_of_nat_def)
done
section \<open>The universal TM\<close>
text \<open>
This section gives the explicit construction of {\em Universal Turing Machine}, defined as \<open>UTM\<close> and proves its
correctness. It is pretty easy by composing the partial results we have got so far.
\<close>
definition UTM :: "instr list"
where
"UTM = (let (aprog, rs_pos, a_md) = rec_ci rec_F in
let abc_F = aprog [+] dummy_abc (Suc (Suc 0)) in
(t_wcode |+| (tm_of abc_F @ shift (mopup (Suc (Suc 0))) (length (tm_of abc_F) div 2))))"
definition F_aprog :: "abc_prog"
where
"F_aprog \<equiv> (let (aprog, rs_pos, a_md) = rec_ci rec_F in
aprog [+] dummy_abc (Suc (Suc 0)))"
definition F_tprog :: "instr list"
where
"F_tprog = tm_of (F_aprog)"
definition t_utm :: "instr list"
where
"t_utm \<equiv>
F_tprog @ shift (mopup (Suc (Suc 0))) (length F_tprog div 2)"
definition UTM_pre :: "instr list"
where
"UTM_pre = t_wcode |+| t_utm"
lemma tinres_step1:
assumes "tinres l l'" "step (ss, l, r) (t, 0) = (sa, la, ra)"
"step (ss, l', r) (t, 0) = (sb, lb, rb)"
shows "tinres la lb \<and> ra = rb \<and> sa = sb"
proof(cases "r")
case Nil
then show ?thesis using assms
by (cases "(fetch t ss Bk)";cases "fst (fetch t ss Bk)";auto simp:step.simps split:if_splits)
next
case (Cons a list)
then show ?thesis using assms
by (cases "(fetch t ss a)";cases "fst (fetch t ss a)";auto simp:step.simps split:if_splits)
qed
lemma tinres_steps1:
"\<lbrakk>tinres l l'; steps (ss, l, r) (t, 0) stp = (sa, la, ra);
steps (ss, l', r) (t, 0) stp = (sb, lb, rb)\<rbrakk>
\<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb"
proof (induct stp arbitrary: sa la ra sb lb rb)
case (Suc stp)
then show ?case apply simp
apply(case_tac "(steps (ss, l, r) (t, 0) stp)")
apply(case_tac "(steps (ss, l', r) (t, 0) stp)")
proof -
fix stp sa la ra sb lb rb a b c aa ba ca
assume ind: "\<And>sa la ra sb lb rb. \<lbrakk>steps (ss, l, r) (t, 0) stp = (sa, (la::cell list), ra);
steps (ss, l', r) (t, 0) stp = (sb, lb, rb)\<rbrakk> \<Longrightarrow> tinres la lb \<and> ra = rb \<and> sa = sb"
and h: " tinres l l'" "step (steps (ss, l, r) (t, 0) stp) (t, 0) = (sa, la, ra)"
"step (steps (ss, l', r) (t, 0) stp) (t, 0) = (sb, lb, rb)" "steps (ss, l, r) (t, 0) stp = (a, b, c)"
"steps (ss, l', r) (t, 0) stp = (aa, ba, ca)"
have "tinres b ba \<and> c = ca \<and> a = aa"
using ind h by metis
thus "tinres la lb \<and> ra = rb \<and> sa = sb"
using tinres_step1 h by metis
qed
qed (simp add: steps.simps)
lemma tinres_some_exp[simp]:
"tinres (Bk \<up> m @ [Bk, Bk]) la \<Longrightarrow> \<exists>m. la = Bk \<up> m" unfolding tinres_def
proof -
let ?c1 = "\<lambda> n. Bk \<up> m @ [Bk, Bk] = la @ Bk \<up> n"
let ?c2 = "\<lambda> n. la = (Bk \<up> m @ [Bk, Bk]) @ Bk \<up> n"
assume "\<exists>n. ?c1 n \<or> ?c2 n"
then obtain n where "?c1 n \<or> ?c2 n" by auto
then consider "?c1 n" | "?c2 n" by blast
thus ?thesis proof(cases)
case 1
hence "Bk \<up> Suc (Suc m) = la @ Bk \<up> n"
by (metis exp_ind append_Cons append_eq_append_conv2 self_append_conv2)
hence "la = Bk \<up> (Suc (Suc m) - n)"
by (metis replicate_add append_eq_append_conv diff_add_inverse2 length_append length_replicate)
then show ?thesis by auto
next
case 2
hence "la = Bk \<up> (m + Suc (Suc n))"
by (metis append_Cons append_eq_append_conv2 replicate_Suc replicate_add self_append_conv2)
then show ?thesis by blast
qed
qed
lemma t_utm_halt_eq:
assumes tm_wf: "tm_wf (tp, 0)"
and exec: "steps0 (Suc 0, Bk\<up>(l), <lm::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(n))"
and resutl: "0 < rs"
shows "\<exists>stp m n. steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>(i)) t_utm stp =
(0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))"
proof -
obtain ap arity fp where a: "rec_ci rec_F = (ap, arity, fp)"
by (metis prod_cases3)
moreover have b: "rec_exec rec_F [code tp, (bl2wc (<lm>))] = (rs - Suc 0)"
using assms
apply(rule_tac F_correct, simp_all)
done
have "\<exists> stp m l. steps0 (Suc 0, Bk # Bk # [], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i)
(F_tprog @ shift (mopup (length [code tp, bl2wc (<lm>)])) (length F_tprog div 2)) stp
= (0, Bk\<up>m @ Bk # Bk # [], Oc\<up>Suc (rec_exec rec_F [code tp, (bl2wc (<lm>))]) @ Bk\<up>l)"
proof(rule_tac recursive_compile_to_tm_correct1)
show "rec_ci rec_F = (ap, arity, fp)" using a by simp
next
show "terminate rec_F [code tp, bl2wc (<lm>)]"
using assms
by(rule_tac terminate_F, simp_all)
next
show "F_tprog = tm_of (ap [+] dummy_abc (length [code tp, bl2wc (<lm>)]))"
using a
apply(simp add: F_tprog_def F_aprog_def numeral_2_eq_2)
done
qed
then obtain stp m l where
"steps0 (Suc 0, Bk # Bk # [], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i)
(F_tprog @ shift (mopup (length [code tp, (bl2wc (<lm>))])) (length F_tprog div 2)) stp
= (0, Bk\<up>m @ Bk # Bk # [], Oc\<up>Suc (rec_exec rec_F [code tp, (bl2wc (<lm>))]) @ Bk\<up>l)" by blast
hence "\<exists> m. steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i)
(F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp
= (0, Bk\<up>m, Oc\<up>Suc (rs - 1) @ Bk\<up>l)"
proof -
assume g: "steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]> @ Bk \<up> i)
(F_tprog @ shift (mopup (length [code tp, bl2wc (<lm>)])) (length F_tprog div 2)) stp =
(0, Bk \<up> m @ [Bk, Bk], Oc \<up> Suc ((rec_exec rec_F [code tp, bl2wc (<lm>)])) @ Bk \<up> l)"
moreover have "tinres [Bk, Bk] [Bk]"
apply(auto simp: tinres_def)
done
moreover obtain sa la ra where "steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i)
(F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp = (sa, la, ra)"
apply(case_tac "steps0 (Suc 0, [Bk], <[code tp, bl2wc (<lm>)]> @ Bk\<up>i)
(F_tprog @ shift (mopup 2) (length F_tprog div 2)) stp", auto)
done
ultimately show "?thesis"
using b
apply(drule_tac la = "Bk\<up>m @ [Bk, Bk]" in tinres_steps1, auto simp: numeral_2_eq_2)
done
qed
thus "?thesis"
apply(auto)
apply(rule_tac x = stp in exI, simp add: t_utm_def)
using assms
apply(case_tac rs, simp_all add: numeral_2_eq_2)
done
qed
lemma tm_wf_t_wcode[intro]: "tm_wf (t_wcode, 0)"
apply(simp add: t_wcode_def)
apply(rule_tac tm_comp_wf)
apply(rule_tac tm_comp_wf, auto)
done
lemma UTM_halt_lemma_pre:
assumes wf_tm: "tm_wf (tp, 0)"
and result: "0 < rs"
and args: "args \<noteq> []"
and exec: "steps0 (Suc 0, Bk\<up>(i), <args::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(k))"
shows "\<exists>stp m n. steps0 (Suc 0, [], <code tp # args>) UTM_pre stp =
(0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))"
proof -
let ?Q2 = "\<lambda> (l, r). (\<exists> ln rn. l = Bk\<up>(ln) \<and> r = Oc\<up>(rs) @ Bk\<up>(rn))"
let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>"
let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and>
(\<exists> rn. r = Oc\<up>(Suc (code tp)) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn)))"
let ?P2 = ?Q1
let ?P3 = "\<lambda> (l, r). False"
have "{?P1} (t_wcode |+| t_utm) {?Q2}"
proof(rule_tac Hoare_plus_halt)
show "tm_wf (t_wcode, 0)" by auto
next
show "{?P1} t_wcode {?Q1}"
apply(rule_tac Hoare_haltI, auto)
using wcode_lemma_1[of args "code tp"] args
apply(auto)
by (metis (mono_tags, lifting) holds_for.simps is_finalI old.prod.case)
next
show "{?P2} t_utm {?Q2}"
proof(rule_tac Hoare_haltI, auto)
fix rn
show "\<exists>n. is_final (steps0 (Suc 0, [Bk], Oc # Oc \<up> code tp @ Bk # Oc # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_utm n) \<and>
(\<lambda>(l, r). (\<exists>ln. l = Bk \<up> ln) \<and>
(\<exists>rn. r = Oc \<up> rs @ Bk \<up> rn)) holds_for steps0 (Suc 0, [Bk],
Oc # Oc \<up> code tp @ Bk # Oc # Oc \<up> bl_bin (<args>) @ Bk \<up> rn) t_utm n"
using t_utm_halt_eq[of tp i "args" stp m rs k rn] assms
apply(auto simp: bin_wc_eq tape_of_list_def tape_of_nat_def)
apply(rename_tac stpa) apply(rule_tac x = stpa in exI, simp)
done
qed
qed
thus "?thesis"
apply(auto simp: Hoare_halt_def UTM_pre_def)
apply(case_tac "steps0 (Suc 0, [], <code tp # args>) (t_wcode |+| t_utm) n",simp)
by auto
qed
text \<open>
The correctness of \<open>UTM\<close>, the halt case.
\<close>
lemma UTM_halt_lemma':
assumes tm_wf: "tm_wf (tp, 0)"
and result: "0 < rs"
and args: "args \<noteq> []"
and exec: "steps0 (Suc 0, Bk\<up>(i), <args::nat list>) tp stp = (0, Bk\<up>(m), Oc\<up>(rs)@Bk\<up>(k))"
shows "\<exists>stp m n. steps0 (Suc 0, [], <code tp # args>) UTM stp =
(0, Bk\<up>(m), Oc\<up>(rs) @ Bk\<up>(n))"
using UTM_halt_lemma_pre[of tp rs args i stp m k] assms
apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def)
apply(case_tac "rec_ci rec_F", simp)
done
definition TSTD:: "config \<Rightarrow> bool"
where
"TSTD c = (let (st, l, r) = c in
st = 0 \<and> (\<exists> m. l = Bk\<up>(m)) \<and> (\<exists> rs n. r = Oc\<up>(Suc rs) @ Bk\<up>(n)))"
lemma nstd_case1: "0 < a \<Longrightarrow> NSTD (trpl_code (a, b, c))"
by(simp add: NSTD.simps trpl_code.simps)
lemma nonzero_bl2wc[simp]: "\<forall>m. b \<noteq> Bk\<up>(m) \<Longrightarrow> 0 < bl2wc b"
proof -
have "\<forall>m. b \<noteq> Bk \<up> m \<Longrightarrow> bl2wc b = 0 \<Longrightarrow> False" proof(induct b)
case (Cons a b)
then show ?case
apply(simp add: bl2wc.simps, case_tac a, simp_all
add: bl2nat.simps bl2nat_double)
apply(case_tac "\<exists> m. b = Bk\<up>(m)", erule exE)
apply(metis append_Nil2 replicate_Suc_iff_anywhere)
by simp
qed auto
thus "\<forall>m. b \<noteq> Bk\<up>(m) \<Longrightarrow> 0 < bl2wc b" by auto
qed
lemma nstd_case2: "\<forall>m. b \<noteq> Bk\<up>(m) \<Longrightarrow> NSTD (trpl_code (a, b, c))"
apply(simp add: NSTD.simps trpl_code.simps)
done
lemma even_not_odd[elim]: "Suc (2 * x) = 2 * y \<Longrightarrow> RR"
proof(induct x arbitrary: y)
case (Suc x) thus ?case by(cases y;auto)
qed auto
declare replicate_Suc[simp del]
lemma bl2nat_zero_eq[simp]: "(bl2nat c 0 = 0) = (\<exists>n. c = Bk\<up>(n))"
proof(induct c)
case (Cons a c)
then show ?case by (cases a;auto simp: bl2nat.simps bl2nat_double Cons_replicate_eq)
qed (auto simp: bl2nat.simps)
lemma bl2wc_exp_ex:
"\<lbrakk>Suc (bl2wc c) = 2 ^ m\<rbrakk> \<Longrightarrow> \<exists> rs n. c = Oc\<up>(rs) @ Bk\<up>(n)"
proof(induct c arbitrary: m)
case (Cons a c m)
{ fix n
have "Bk # Bk \<up> n = Oc \<up> 0 @ Bk \<up> Suc n" by (auto simp:replicate_Suc)
hence "\<exists>rs na. Bk # Bk \<up> n = Oc \<up> rs @ Bk \<up> na" by blast
}
with Cons show ?case apply(cases a, auto)
apply(case_tac m, simp_all add: bl2wc.simps, auto)
apply(simp add: bl2wc.simps bl2nat.simps bl2nat_double Cons)
apply(case_tac m, simp, simp add: bin_wc_eq bl2wc.simps twice_power )
by (metis Cons.hyps Suc_pred bl2wc.simps neq0_conv power_not_zero
replicate_Suc_iff_anywhere zero_neq_numeral)
qed (simp add: bl2wc.simps bl2nat.simps)
lemma lg_bin:
assumes "\<forall>rs n. c \<noteq> Oc\<up>(Suc rs) @ Bk\<up>(n)"
"bl2wc c = 2 ^ lg (Suc (bl2wc c)) 2 - Suc 0"
shows "bl2wc c = 0"
proof -
from assms obtain rs nat n where *:"2 ^ rs - Suc 0 = nat"
"c = Oc \<up> rs @ Bk \<up> n"
using bl2wc_exp_ex[of c "lg (Suc (bl2wc c)) 2"]
by(case_tac "(2::nat) ^ lg (Suc (bl2wc c)) 2",
simp, simp, erule_tac exE, erule_tac exE, simp)
have r:"bl2wc (Oc \<up> rs) = nat"
by (metis "*"(1) bl2nat_exp_zero bl2wc.elims)
hence "Suc (bl2wc c) = 2^rs" using *
by(case_tac "(2::nat)^rs", auto)
thus ?thesis using * assms(1)
apply(drule_tac bl2wc_exp_ex, simp, erule_tac exE, erule_tac exE)
by(case_tac rs, simp, simp)
qed
lemma nstd_case3:
"\<forall>rs n. c \<noteq> Oc\<up>(Suc rs) @ Bk\<up>(n) \<Longrightarrow> NSTD (trpl_code (a, b, c))"
apply(simp add: NSTD.simps trpl_code.simps)
apply(auto)
apply(drule_tac lg_bin, simp_all)
done
lemma NSTD_1: "\<not> TSTD (a, b, c)
\<Longrightarrow> rec_exec rec_NSTD [trpl_code (a, b, c)] = Suc 0"
using NSTD_lemma1[of "trpl_code (a, b, c)"]
NSTD_lemma2[of "trpl_code (a, b, c)"]
apply(simp add: TSTD_def)
apply(erule_tac disjE, erule_tac nstd_case1)
apply(erule_tac disjE, erule_tac nstd_case2)
apply(erule_tac nstd_case3)
done
lemma nonstop_t_uhalt_eq:
"\<lbrakk>tm_wf (tp, 0);
steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp = (a, b, c);
\<not> TSTD (a, b, c)\<rbrakk>
\<Longrightarrow> rec_exec rec_nonstop [code tp, bl2wc (<lm>), stp] = Suc 0"
apply(simp add: rec_nonstop_def rec_exec.simps)
apply(subgoal_tac
"rec_exec rec_conf [code tp, bl2wc (<lm>), stp] =
trpl_code (a, b, c)", simp)
apply(erule_tac NSTD_1)
using rec_t_eq_steps[of tp l lm stp]
apply(simp)
done
lemma nonstop_true:
"\<lbrakk>tm_wf (tp, 0);
\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp))\<rbrakk>
\<Longrightarrow> \<forall>y. rec_exec rec_nonstop ([code tp, bl2wc (<lm>), y]) = (Suc 0)"
proof fix y
assume a:"tm_wf0 tp" "\<forall>stp. \<not> TSTD (steps0 (Suc 0, Bk \<up> l, <lm>) tp stp)"
hence "\<not> TSTD (steps0 (Suc 0, Bk \<up> l, <lm>) tp y)" by auto
thus "rec_exec rec_nonstop [code tp, bl2wc (<lm>), y] = Suc 0"
by (cases "steps0 (Suc 0, Bk\<up>(l), <lm>) tp y")
(auto intro: nonstop_t_uhalt_eq[OF a(1)])
qed
lemma cn_arity: "rec_ci (Cn n f gs) = (a, b, c) \<Longrightarrow> b = n"
by(case_tac "rec_ci f", simp add: rec_ci.simps)
lemma mn_arity: "rec_ci (Mn n f) = (a, b, c) \<Longrightarrow> b = n"
by(case_tac "rec_ci f", simp add: rec_ci.simps)
lemma F_aprog_uhalt:
assumes wf_tm: "tm_wf (tp,0)"
and unhalt: "\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp))"
and compile: "rec_ci rec_F = (F_ap, rs_pos, a_md)"
shows "{\<lambda> nl. nl = [code tp, bl2wc (<lm>)] @ 0\<up>(a_md - rs_pos ) @ suflm} (F_ap) \<up>"
using compile
proof(simp only: rec_F_def)
assume h: "rec_ci (Cn (Suc (Suc 0)) rec_valu [Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0))
rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt]]]) =
(F_ap, rs_pos, a_md)"
moreover hence "rs_pos = Suc (Suc 0)"
using cn_arity
by simp
moreover obtain ap1 ar1 ft1 where a: "rec_ci
(Cn (Suc (Suc 0)) rec_right
[Cn (Suc (Suc 0)) rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt]]) = (ap1, ar1, ft1)"
by(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_right [Cn (Suc (Suc 0))
rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt]])", auto)
moreover hence b: "ar1 = Suc (Suc 0)"
using cn_arity by simp
ultimately show "?thesis"
proof(rule_tac i = 0 in cn_unhalt_case, auto)
fix anything
obtain ap2 ar2 ft2 where c:
"rec_ci (Cn (Suc (Suc 0)) rec_conf [recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt])
= (ap2, ar2, ft2)"
by(case_tac "rec_ci (Cn (Suc (Suc 0)) rec_conf
[recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), rec_halt])", auto)
moreover hence d:"ar2 = Suc (Suc 0)"
using cn_arity by simp
ultimately have "{\<lambda>nl. nl = [code tp, bl2wc (<lm>)] @ 0 \<up> (ft1 - Suc (Suc 0)) @ anything} ap1 \<up>"
using a b c d
proof(rule_tac i = 0 in cn_unhalt_case, auto)
fix anything
obtain ap3 ar3 ft3 where e: "rec_ci rec_halt = (ap3, ar3, ft3)"
by(case_tac "rec_ci rec_halt", auto)
hence f: "ar3 = Suc (Suc 0)"
using mn_arity
by(simp add: rec_halt_def)
have "{\<lambda>nl. nl = [code tp, bl2wc (<lm>)] @ 0 \<up> (ft2 - Suc (Suc 0)) @ anything} ap2 \<up>"
using c d e f
proof(rule_tac i = 2 in cn_unhalt_case, auto simp: rec_halt_def)
fix anything
have "{\<lambda>nl. nl = [code tp, bl2wc (<lm>)] @ 0 \<up> (ft3 - Suc (Suc 0)) @ anything} ap3 \<up>"
using e f
proof(rule_tac mn_unhalt_case, auto simp: rec_halt_def)
fix i
show "terminate rec_nonstop [code tp, bl2wc (<lm>), i]"
by(rule_tac primerec_terminate, auto)
next
fix i
show "0 < rec_exec rec_nonstop [code tp, bl2wc (<lm>), i]"
using assms
by(drule_tac nonstop_true, auto)
qed
thus "{\<lambda>nl. nl = code tp # bl2wc (<lm>) # 0 \<up> (ft3 - Suc (Suc 0)) @ anything} ap3 \<up>" by simp
next
fix apj arj ftj j anything
assume "j<2" "rec_ci ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j) = (apj, arj, ftj)"
hence "{\<lambda>nl. nl = [code tp, bl2wc (<lm>)] @ 0 \<up> (ftj - arj) @ anything} apj
{\<lambda>nl. nl = [code tp, bl2wc (<lm>)] @
rec_exec ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j) [code tp, bl2wc (<lm>)] #
0 \<up> (ftj - Suc arj) @ anything}"
apply(rule_tac recursive_compile_correct)
apply(case_tac j, auto)
apply(rule_tac [!] primerec_terminate)
by(auto)
thus "{\<lambda>nl. nl = code tp # bl2wc (<lm>) # 0 \<up> (ftj - arj) @ anything} apj
{\<lambda>nl. nl = code tp # bl2wc (<lm>) # rec_exec ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0))
(Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j) [code tp, bl2wc (<lm>)] # 0 \<up> (ftj - Suc arj) @ anything}"
by simp
next
fix j
assume "(j::nat) < 2"
thus "terminate ([recf.id (Suc (Suc 0)) 0, recf.id (Suc (Suc 0)) (Suc 0), Mn (Suc (Suc 0)) rec_nonstop] ! j)
[code tp, bl2wc (<lm>)]"
by(case_tac j, auto intro!: primerec_terminate)
qed
thus "{\<lambda>nl. nl = code tp # bl2wc (<lm>) # 0 \<up> (ft2 - Suc (Suc 0)) @ anything} ap2 \<up>"
by simp
qed
thus "{\<lambda>nl. nl = code tp # bl2wc (<lm>) # 0 \<up> (ft1 - Suc (Suc 0)) @ anything} ap1 \<up>" by simp
qed
qed
lemma uabc_uhalt':
"\<lbrakk>tm_wf (tp, 0);
\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp));
rec_ci rec_F = (ap, pos, md)\<rbrakk>
\<Longrightarrow> {\<lambda> nl. nl = [code tp, bl2wc (<lm>)]} ap \<up>"
proof(frule_tac F_ap = ap and rs_pos = pos and a_md = md
and suflm = "[]" in F_aprog_uhalt, auto simp: abc_Hoare_unhalt_def,
case_tac "abc_steps_l (0, [code tp, bl2wc (<lm>)]) ap n", simp)
fix n a b
assume h:
"\<forall>n. abc_notfinal (abc_steps_l (0, code tp # bl2wc (<lm>) # 0 \<up> (md - pos)) ap n) ap"
"abc_steps_l (0, [code tp, bl2wc (<lm>)]) ap n = (a, b)"
"tm_wf (tp, 0)"
"rec_ci rec_F = (ap, pos, md)"
moreover have a: "ap \<noteq> []"
using h rec_ci_not_null[of "rec_F" pos md] by auto
ultimately show "a < length ap"
proof(erule_tac x = n in allE)
assume g: "abc_notfinal (abc_steps_l (0, code tp # bl2wc (<lm>) # 0 \<up> (md - pos)) ap n) ap"
obtain ss nl where b : "abc_steps_l (0, code tp # bl2wc (<lm>) # 0 \<up> (md - pos)) ap n = (ss, nl)"
by (metis prod.exhaust)
then have c: "ss < length ap"
using g by simp
thus "?thesis"
using a b c
using abc_list_crsp_steps[of "[code tp, bl2wc (<lm>)]"
"md - pos" ap n ss nl] h
by(simp)
qed
qed
lemma uabc_uhalt:
"\<lbrakk>tm_wf (tp, 0);
\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp))\<rbrakk>
\<Longrightarrow> {\<lambda> nl. nl = [code tp, bl2wc (<lm>)]} F_aprog \<up> "
proof -
obtain a b c where abc:"rec_ci rec_F = (a,b,c)" by (cases "rec_ci rec_F") force
assume a:"tm_wf (tp, 0)" "\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <lm>) tp stp))"
from uabc_uhalt'[OF a abc] abc_Hoare_plus_unhalt1
show "{\<lambda> nl. nl = [code tp, bl2wc (<lm>)]} F_aprog \<up>"
by(simp add: F_aprog_def abc)
qed
lemma tutm_uhalt':
assumes tm_wf: "tm_wf (tp,0)"
and unhalt: "\<forall> stp. (\<not> TSTD (steps0 (1, Bk\<up>(l), <lm>) tp stp))"
shows "\<forall> stp. \<not> is_final (steps0 (1, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)"
unfolding t_utm_def
proof(rule_tac compile_correct_unhalt, auto)
show "F_tprog = tm_of F_aprog"
by(simp add: F_tprog_def)
next
show "crsp (layout_of F_aprog) (0, [code tp, bl2wc (<lm>)]) (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) []"
by(auto simp: crsp.simps start_of.simps)
next
fix stp a b
show "abc_steps_l (0, [code tp, bl2wc (<lm>)]) F_aprog stp = (a, b) \<Longrightarrow> a < length F_aprog"
using assms
apply(drule_tac uabc_uhalt, auto simp: abc_Hoare_unhalt_def)
by(erule_tac x = stp in allE, erule_tac x = stp in allE, simp)
qed
lemma tinres_commute: "tinres r r' \<Longrightarrow> tinres r' r"
apply(auto simp: tinres_def)
done
lemma inres_tape:
"\<lbrakk>steps0 (st, l, r) tp stp = (a, b, c); steps0 (st, l', r') tp stp = (a', b', c');
tinres l l'; tinres r r'\<rbrakk>
\<Longrightarrow> a = a' \<and> tinres b b' \<and> tinres c c'"
proof(case_tac "steps0 (st, l', r) tp stp")
fix aa ba ca
assume h: "steps0 (st, l, r) tp stp = (a, b, c)"
"steps0 (st, l', r') tp stp = (a', b', c')"
"tinres l l'" "tinres r r'"
"steps0 (st, l', r) tp stp = (aa, ba, ca)"
have "tinres b ba \<and> c = ca \<and> a = aa"
using h
apply(rule_tac tinres_steps1, auto)
done
moreover have "b' = ba \<and> tinres c' ca \<and> a' = aa"
using h
apply(rule_tac tinres_steps2, auto intro: tinres_commute)
done
ultimately show "?thesis"
apply(auto intro: tinres_commute)
done
qed
lemma tape_normalize:
assumes "\<forall> stp. \<not> is_final(steps0 (Suc 0, [Bk,Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp)"
shows "\<forall> stp. \<not> is_final (steps0 (Suc 0, Bk\<up>(m), <[code tp, bl2wc (<lm>)]> @ Bk\<up>(n)) t_utm stp)"
(is "\<forall> stp. ?P stp")
proof
fix stp
from assms[rule_format,of stp] show "?P stp"
apply(case_tac "steps0 (Suc 0, Bk\<up>(m), <[code tp, bl2wc (<lm>)]> @ Bk\<up>(n)) t_utm stp", simp)
apply(case_tac "steps0 (Suc 0, [Bk, Bk], <[code tp, bl2wc (<lm>)]>) t_utm stp", simp)
apply(drule_tac inres_tape, auto)
apply(auto simp: tinres_def)
apply(case_tac "m > Suc (Suc 0)")
apply(rule_tac x = "m - Suc (Suc 0)" in exI)
apply(case_tac m, simp_all)
apply(metis Suc_lessD Suc_pred replicate_Suc)
apply(rule_tac x = "2 - m" in exI, simp add: replicate_add[THEN sym])
apply(simp only: numeral_2_eq_2, simp add: replicate_Suc)
done
qed
lemma tutm_uhalt:
"\<lbrakk>tm_wf (tp,0);
\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <args>) tp stp))\<rbrakk>
\<Longrightarrow> \<forall> stp. \<not> is_final (steps0 (Suc 0, Bk\<up>(m), <[code tp, bl2wc (<args>)]> @ Bk\<up>(n)) t_utm stp)"
apply(rule_tac tape_normalize)
apply(rule_tac tutm_uhalt'[simplified], simp_all)
done
lemma UTM_uhalt_lemma_pre:
assumes tm_wf: "tm_wf (tp, 0)"
and exec: "\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <args>) tp stp))"
and args: "args \<noteq> []"
shows "\<forall> stp. \<not> is_final (steps0 (Suc 0, [], <code tp # args>) UTM_pre stp)"
proof -
let ?P1 = "\<lambda> (l, r). l = [] \<and> r = <code tp # args>"
let ?Q1 = "\<lambda> (l, r). (l = [Bk] \<and>
(\<exists> rn. r = Oc\<up>(Suc (code tp)) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn)))"
let ?P2 = ?Q1
have "{?P1} (t_wcode |+| t_utm) \<up>"
proof(rule_tac Hoare_plus_unhalt)
show "tm_wf (t_wcode, 0)" by auto
next
show "{?P1} t_wcode {?Q1}"
apply(rule_tac Hoare_haltI, auto)
using wcode_lemma_1[of args "code tp"] args
apply(auto)
by (metis (mono_tags, lifting) holds_for.simps is_finalI old.prod.case)
next
show "{?P2} t_utm \<up>"
proof(rule_tac Hoare_unhaltI, auto)
fix n rn
assume h: "is_final (steps0 (Suc 0, [Bk], Oc \<up> Suc (code tp) @ Bk # Oc \<up> Suc (bl_bin (<args>)) @ Bk \<up> rn) t_utm n)"
have "\<forall> stp. \<not> is_final (steps0 (Suc 0, Bk\<up>(Suc 0), <[code tp, bl2wc (<args>)]> @ Bk\<up>(rn)) t_utm stp)"
using assms
apply(rule_tac tutm_uhalt, simp_all)
done
thus "False"
using h
apply(erule_tac x = n in allE)
apply(simp add: tape_of_list_def bin_wc_eq tape_of_nat_def)
done
qed
qed
thus "?thesis"
apply(simp add: Hoare_unhalt_def UTM_pre_def)
done
qed
text \<open>
The correctness of \<open>UTM\<close>, the unhalt case.
\<close>
lemma UTM_uhalt_lemma':
assumes tm_wf: "tm_wf (tp, 0)"
and unhalt: "\<forall> stp. (\<not> TSTD (steps0 (Suc 0, Bk\<up>(l), <args>) tp stp))"
and args: "args \<noteq> []"
shows " \<forall> stp. \<not> is_final (steps0 (Suc 0, [], <code tp # args>) UTM stp)"
using UTM_uhalt_lemma_pre[of tp l args] assms
apply(simp add: UTM_pre_def t_utm_def UTM_def F_aprog_def F_tprog_def)
apply(case_tac "rec_ci rec_F", simp)
done
lemma UTM_halt_lemma:
assumes tm_wf: "tm_wf (p, 0)"
and resut: "rs > 0"
and args: "(args::nat list) \<noteq> []"
and exec: "{(\<lambda>tp. tp = (Bk\<up>i, <args>))} p {(\<lambda>tp. tp = (Bk\<up>m, Oc\<up>rs @ Bk\<up>k))}"
shows "{(\<lambda>tp. tp = ([], <code p # args>))} UTM {(\<lambda>tp. (\<exists> m n. tp = (Bk\<up>m, Oc\<up>rs @ Bk\<up>n)))}"
proof -
let ?steps0 = "steps0 (Suc 0, [], <code p # args>)"
let ?stepsBk = "steps0 (Suc 0, Bk\<up>i, <args>) p"
from wcode_lemma_1[OF args,of "code p"] obtain stp ln rn where
wcl1:"?steps0 t_wcode stp =
(0, [Bk], Oc \<up> Suc (code p) @ Bk # Oc \<up> Suc (bl_bin (<args>)) @ Bk \<up> rn)" by fast
from exec Hoare_halt_def obtain n where
n:"{\<lambda>tp. tp = (Bk \<up> i, <args>)} p {\<lambda>tp. tp = (Bk \<up> m, Oc \<up> rs @ Bk \<up> k)}"
"is_final (?stepsBk n)"
"(\<lambda>tp. tp = (Bk \<up> m, Oc \<up> rs @ Bk \<up> k)) holds_for steps0 (Suc 0, Bk \<up> i, <args>) p n"
by auto
obtain a where a:"a = fst (rec_ci rec_F)" by blast
have "{(\<lambda> (l, r). l = [] \<and> r = <code p # args>)} (t_wcode |+| t_utm)
{(\<lambda> (l, r). (\<exists> m. l = Bk\<up>m) \<and> (\<exists> n. r = Oc\<up>rs @ Bk\<up>n))}"
proof(rule_tac Hoare_plus_halt)
show "{\<lambda>(l, r). l = [] \<and> r = <code p # args>} t_wcode {\<lambda> (l, r). (l = [Bk] \<and>
(\<exists> rn. r = Oc\<up>(Suc (code p)) @ Bk # Oc\<up>(Suc (bl_bin (<args>))) @ Bk\<up>(rn)))}"
using wcl1 by (auto intro!:Hoare_haltI exI[of _ stp])
next
have "\<exists> stp. (?stepsBk stp = (0, Bk\<up>m, Oc\<up>rs @ Bk\<up>k))"
using n by (case_tac "?stepsBk n", auto)
then obtain stp where k: "steps0 (Suc 0, Bk\<up>i, <args>) p stp = (0, Bk\<up>m, Oc\<up>rs @ Bk\<up>k)"
..
thus "{\<lambda>(l, r). l = [Bk] \<and> (\<exists>rn. r = Oc \<up> Suc (code p) @ Bk # Oc \<up> Suc (bl_bin (<args>)) @ Bk \<up> rn)}
t_utm {\<lambda>(l, r). (\<exists>m. l = Bk \<up> m) \<and> (\<exists>n. r = Oc \<up> rs @ Bk \<up> n)}"
proof(rule_tac Hoare_haltI, auto)
fix rn
from t_utm_halt_eq[OF assms(1) k assms(2),of rn] assms k
have "\<exists> ma n stp. steps0 (Suc 0, [Bk], <[code p, bl2wc (<args>)]> @ Bk \<up> rn) t_utm stp =
(0, Bk \<up> ma, Oc \<up> rs @ Bk \<up> n)" by (auto simp add: bin_wc_eq)
then obtain stpx m' n' where
t:"steps0 (Suc 0, [Bk], <[code p, bl2wc (<args>)]> @ Bk \<up> rn) t_utm stpx =
(0, Bk \<up> m', Oc \<up> rs @ Bk \<up> n')" by auto
show "\<exists>n. is_final (steps0 (Suc 0, [Bk], Oc \<up> Suc (code p) @ Bk # Oc \<up> Suc (bl_bin (<args>)) @ Bk \<up> rn) t_utm n) \<and>
(\<lambda>(l, r). (\<exists>m. l = Bk \<up> m) \<and> (\<exists>n. r = Oc \<up> rs @ Bk \<up> n)) holds_for steps0
(Suc 0, [Bk], Oc \<up> Suc (code p) @ Bk # Oc \<up> Suc (bl_bin (<args>)) @ Bk \<up> rn) t_utm n"
using t
by(auto simp: bin_wc_eq tape_of_list_def tape_of_nat_def intro:exI[of _ stpx])
qed
next
show "tm_wf0 t_wcode" by auto
qed
then obtain n where
"is_final (?steps0 (t_wcode |+| t_utm) n)"
"(\<lambda>(l, r). (\<exists>m. l = Bk \<up> m) \<and>
(\<exists>n. r = Oc \<up> rs @ Bk \<up> n)) holds_for ?steps0 (t_wcode |+| t_utm) n"
by(auto simp add: Hoare_halt_def a)
thus "?thesis"
apply(case_tac "rec_ci rec_F")
apply(auto simp add: UTM_def Hoare_halt_def)
apply(case_tac "(?steps0 (t_wcode |+| t_utm) n)")
apply(rule_tac x="n" in exI)
apply(auto simp add:a t_utm_def F_aprog_def F_tprog_def)
done
qed
lemma UTM_halt_lemma2:
assumes tm_wf: "tm_wf (p, 0)"
and args: "(args::nat list) \<noteq> []"
and exec: "{(\<lambda>tp. tp = ([], <args>))} p {(\<lambda>tp. tp = (Bk\<up>m, <(n::nat)> @ Bk\<up>k))}"
shows "{(\<lambda>tp. tp = ([], <code p # args>))} UTM {(\<lambda>tp. (\<exists> m k. tp = (Bk\<up>m, <n> @ Bk\<up>k)))}"
using UTM_halt_lemma[OF assms(1) _ assms(2), where i="0"]
using assms(3)
by(simp add: tape_of_nat_def)
lemma UTM_unhalt_lemma:
assumes tm_wf: "tm_wf (p, 0)"
and unhalt: "{(\<lambda>tp. tp = (Bk\<up>i, <args>))} p \<up>"
and args: "args \<noteq> []"
shows "{(\<lambda>tp. tp = ([], <code p # args>))} UTM \<up>"
proof -
have "(\<not> TSTD (steps0 (Suc 0, Bk\<up>(i), <args>) p stp))" for stp
(* in unhalt, we substitute inner 'forall' n\<rightarrow>stp *)
using unhalt[unfolded Hoare_unhalt_def,rule_format,OF refl,of stp]
by(cases "steps0 (Suc 0, Bk \<up> i, <args>) p stp",auto simp: Hoare_unhalt_def TSTD_def)
then have "\<forall> stp. \<not> is_final (steps0 (Suc 0, [], <code p # args>) UTM stp)"
using assms by(intro UTM_uhalt_lemma', auto)
thus "?thesis" by(simp add: Hoare_unhalt_def)
qed
lemma UTM_unhalt_lemma2:
assumes tm_wf: "tm_wf (p, 0)"
and unhalt: "{(\<lambda>tp. tp = ([], <args>))} p \<up>"
and args: "args \<noteq> []"
shows "{(\<lambda>tp. tp = ([], <code p # args>))} UTM \<up>"
using UTM_unhalt_lemma[OF assms(1), where i="0"]
using assms(2-3)
by(simp add: tape_of_nat_def)
end
\ No newline at end of file
diff --git a/thys/Verified-Prover/Prover.thy b/thys/Verified-Prover/Prover.thy
--- a/thys/Verified-Prover/Prover.thy
+++ b/thys/Verified-Prover/Prover.thy
@@ -1,1345 +1,1345 @@
theory Prover
imports Main
begin
declare ex_image_cong_iff [simp del]
subsection "Formulas"
type_synonym pred = nat
type_synonym var = nat
datatype form =
PAtom pred "var list"
| NAtom pred "var list"
| FConj form form
| FDisj form form
| FAll form
| FEx form
primrec preSuc :: "nat list => nat list"
where
"preSuc [] = []"
| "preSuc (a#list) = (case a of 0 => preSuc list | Suc n => n#(preSuc list))"
primrec fv :: "form => var list" \<comment> \<open>shouldn't need to be more constructive than this\<close>
where
"fv (PAtom p vs) = vs"
| "fv (NAtom p vs) = vs"
| "fv (FConj f g) = (fv f) @ (fv g)"
| "fv (FDisj f g) = (fv f) @ (fv g)"
| "fv (FAll f) = preSuc (fv f)"
| "fv (FEx f) = preSuc (fv f)"
definition
bump :: "(var => var) => (var => var)" \<comment> \<open>substitute a different var for 0\<close> where
"bump phi y = (case y of 0 => 0 | Suc n => Suc (phi n))"
primrec subst :: "(nat => nat) => form => form"
where
"subst r (PAtom p vs) = (PAtom p (map r vs))"
| "subst r (NAtom p vs) = (NAtom p (map r vs))"
| "subst r (FConj f g) = FConj (subst r f) (subst r g)"
| "subst r (FDisj f g) = FDisj (subst r f) (subst r g)"
| "subst r (FAll f) = FAll (subst (bump r) f)"
| "subst r (FEx f) = FEx (subst (bump r) f)"
lemma size_subst[simp]: "\<forall>m. size (subst m f) = size f"
by (induct f) (force+)
definition
finst :: "form => var => form" where
"finst body w = (subst (% v. case v of 0 => w | Suc n => n) body)"
lemma size_finst[simp]: "size (finst f m) = size f"
by (simp add: finst_def)
type_synonym seq = "form list"
type_synonym nform = "nat * form"
type_synonym nseq = "nform list"
definition
s_of_ns :: "nseq => seq" where
"s_of_ns ns = map snd ns"
definition
ns_of_s :: "seq => nseq" where
"ns_of_s s = map (% x. (0,x)) s"
primrec flatten :: "'a list list => 'a list"
where
"flatten [] = []"
| "flatten (a#list) = (a@(flatten list))"
definition
sfv :: "seq => var list" where
"sfv s = flatten (map fv s)"
lemma sfv_nil: "sfv [] = []" by(force simp: sfv_def)
lemma sfv_cons: "sfv (a#list) = (fv a) @ (sfv list) " by(force simp: sfv_def)
primrec maxvar :: "var list => var"
where
"maxvar [] = 0"
| "maxvar (a#list) = max a (maxvar list)"
lemma maxvar: "\<forall>v \<in> set vs. v \<le> maxvar vs"
by (induct vs) (auto simp: max_def)
definition
newvar :: "var list => var" where
"newvar vs = (if vs = [] then 0 else Suc (maxvar vs))"
\<comment> \<open>note that for newvar to be constructive, need an operation to get a different var from a given set\<close>
lemma newvar: "newvar vs \<notin> (set vs)"
using length_pos_if_in_set maxvar newvar_def by force
(*lemmas newvar_sfv = newvar[of "sfv s"]*)
primrec subs :: "nseq => nseq list"
where
"subs [] = [[]]"
| "subs (x#xs) =
(let (m,f) = x in
case f of
PAtom p vs => if NAtom p vs : set (map snd xs) then [] else [xs@[(0,PAtom p vs)]]
| NAtom p vs => if PAtom p vs : set (map snd xs) then [] else [xs@[(0,NAtom p vs)]]
| FConj f g => [xs@[(0,f)],xs@[(0,g)]]
| FDisj f g => [xs@[(0,f),(0,g)]]
| FAll f => [xs@[(0,finst f (newvar (sfv (s_of_ns (x#xs)))))]]
| FEx f => [xs@[(0,finst f m),(Suc m,FEx f)]]
)"
subsection "Derivations"
primrec is_axiom :: "seq => bool"
where
"is_axiom [] = False"
| "is_axiom (a#list) = ((? p vs. a = PAtom p vs & NAtom p vs : set list) | (? p vs. a = NAtom p vs & PAtom p vs : set list))"
inductive_set
deriv :: "nseq => (nat * nseq) set"
for s :: nseq
where
init: "(0,s) \<in> deriv s"
| step: "(n,x) \<in> deriv s ==> y : set (subs x) ==> (Suc n,y) \<in> deriv s"
\<comment> \<open>the closure of the branch at isaxiom\<close>
(*lemma step': "(n,x) \<in> deriv s ==> y : set (subs x) ==> ~ is_axiom (s_of_ns x) ==> (Suc n,y) \<in> deriv s"*)
lemma patom: "(n,(m,PAtom p vs)#xs) \<in> deriv(nfs) ==> ~is_axiom (s_of_ns ((m,PAtom p vs)#xs)) ==> (Suc n,xs@[(0,PAtom p vs)]) \<in> deriv(nfs)"
and natom: "(n,(m,NAtom p vs)#xs) \<in> deriv(nfs) ==> ~is_axiom (s_of_ns ((m,NAtom p vs)#xs)) ==> (Suc n,xs@[(0,NAtom p vs)]) \<in> deriv(nfs)"
and fconj1: "(n,(m,FConj f g)#xs) \<in> deriv(nfs) ==> ~is_axiom (s_of_ns ((m,FConj f g)#xs)) ==> (Suc n,xs@[(0,f)]) \<in> deriv(nfs)"
and fconj2: "(n,(m,FConj f g)#xs) \<in> deriv(nfs) ==> ~is_axiom (s_of_ns ((m,FConj f g)#xs)) ==> (Suc n,xs@[(0,g)]) \<in> deriv(nfs)"
and fdisj: "(n,(m,FDisj f g)#xs) \<in> deriv(nfs) ==> ~is_axiom (s_of_ns ((m,FDisj f g)#xs)) ==> (Suc n,xs@[(0,f),(0,g)]) \<in> deriv(nfs)"
and fall: "(n,(m,FAll f)#xs) \<in> deriv(nfs) ==> ~is_axiom (s_of_ns ((m,FAll f)#xs)) ==> (Suc n,xs@[(0,finst f (newvar (sfv (s_of_ns ((m,FAll f)#xs)))))]) \<in> deriv(nfs)"
and fex: "(n,(m,FEx f)#xs) \<in> deriv(nfs) ==> ~is_axiom (s_of_ns ((m,FEx f)#xs)) ==> (Suc n,xs@[(0,finst f m),(Suc m,FEx f)]) \<in> deriv(nfs)"
apply(auto intro: step simp add: Let_def s_of_ns_def)
done
lemmas not_is_axiom_subs = patom natom fconj1 fconj2 fdisj fall fex
lemmas[simp] = init
lemmas [intro] = init step
lemma deriv0[simp]: "(0,x) \<in> deriv y = (x = y)"
apply(rule)
apply(erule deriv.cases) apply(simp) apply(simp)
apply(simp)
done
(*
lemma deriv_exists: "(n,x) \<in> deriv s ==> x \<noteq> [] ==> ~ is_axiom (s_of_ns x) ==> (\<exists>y. (Suc n, y) \<in> deriv s & y : set (subs x))"
apply(case_tac x) apply(simp)
apply(case_tac a) apply(case_tac b)
apply(auto simp add: Let_def s_of_ns_def intro: step)
apply(rule_tac x= "list @ [(0, form1)]" in exI)
apply(simp)
apply(rule step) apply(assumption) apply(simp add: Let_def)
done
*)
lemma deriv_upwards: "(n,list) \<in> deriv s ==> ~ is_axiom (s_of_ns (list)) ==> (\<exists>zs. (Suc n, zs) \<in> deriv s & zs : set (subs list))"
apply(case_tac list) apply force
apply(case_tac a) apply(case_tac b)
apply(simp add: Let_def) apply(rule) apply(simp add: s_of_ns_def) apply(force dest: not_is_axiom_subs)
apply(simp add: Let_def) apply(rule) apply(simp add: s_of_ns_def) apply(force dest: not_is_axiom_subs)
apply(simp add: Let_def) apply(force dest: not_is_axiom_subs)
apply(simp add: Let_def) apply(force dest: not_is_axiom_subs)
apply(simp add: Let_def) apply(force dest: not_is_axiom_subs)
apply(simp add: Let_def) apply(force dest: not_is_axiom_subs)
done
lemma deriv_downwards (*derivSucE*): "(Suc n, x) \<in> deriv s ==> \<exists>y. (n,y) \<in> deriv s & x : set (subs y) & ~ is_axiom (s_of_ns y)"
apply(erule deriv.cases)
apply(simp)
apply(simp add: s_of_ns_def Let_def)
apply(rule_tac x=xa in exI) apply(simp)
apply(case_tac xa) apply(simp)
apply(case_tac a) apply(case_tac b)
apply(auto simp add: Let_def)
apply (rename_tac[!] nat lista a)
apply(subgoal_tac "NAtom nat lista \<in> snd ` set list") apply(simp) apply(force)
apply(subgoal_tac "PAtom nat lista \<in> snd ` set list") apply(simp) apply(force)
done
lemma deriv_deriv_child(*derivSuc*)[rule_format]: "\<forall>x y. (Suc n,x) \<in> deriv y = (\<exists>z. z : set (subs y) & ~ is_axiom (s_of_ns y) & (n,x) \<in> deriv z)"
apply(induct n)
apply(rule, rule) apply(rule) apply(frule_tac deriv_downwards) apply(simp)
apply(simp) apply(rule step) apply(simp) apply(simp)
apply(blast dest!: deriv_downwards elim: deriv.cases) \<comment> \<open>blast needs some help with the reasoning, hence derivSucE\<close>
done
(*
lemma deriv_not_nil: "s \<noteq> [] ==> \<forall>t. (n,t) \<in> deriv s --> t \<noteq> []"
apply(induct_tac n)
apply(force)
apply (blast dest: derivSucE deriv_exists);
done
*)
lemma deriv_progress: "(n,a#list) \<in> deriv s ==> ~ is_axiom (s_of_ns (a#list)) ==> (\<exists>zs. (Suc n, list@zs) \<in> deriv s)"
apply(subgoal_tac "a#list \<noteq> []") prefer 2 apply(simp)
apply(case_tac a) apply(case_tac b)
apply(force dest: not_is_axiom_subs)+
done
definition
inc :: "nat * nseq => nat * nseq" where
"inc = (%(n,fs). (Suc n, fs))"
lemma inj_inc[simp]: "inj inc"
apply(simp add: inc_def) apply(simp add: inj_on_def) done
lemma deriv: "deriv y = insert (0,y) (inc ` (Union (deriv ` {w. ~is_axiom (s_of_ns y) & w : set (subs y)})))"
apply(rule set_eqI)
apply(simp add: split_paired_all)
apply(case_tac a)
apply(force simp: inc_def)
apply(force simp: deriv_deriv_child inc_def)
done
lemma deriv_is_axiom: "is_axiom (s_of_ns s) ==> deriv s = {(0,s)}"
apply(rule)
apply(rule)
apply(case_tac x) apply(simp)
apply(erule_tac deriv.induct) apply(force) apply(simp add: s_of_ns_def) apply(case_tac s) apply(simp) apply(case_tac aa) apply(case_tac ba)
apply(simp_all add: Let_def)
done
lemma is_axiom_finite_deriv: "is_axiom (s_of_ns s) ==> finite (deriv s)"
apply(simp add: deriv_is_axiom) done
subsection "Failing path"
primrec failing_path :: "(nat * nseq) set => nat => (nat * nseq)"
where
"failing_path ns 0 = (SOME x. x \<in> ns & fst x = 0 & infinite (deriv (snd x)) & ~ is_axiom (s_of_ns (snd x)))"
| "failing_path ns (Suc n) = (let fn = failing_path ns n in
(SOME fsucn. fsucn \<in> ns & fst fsucn = Suc n & (snd fsucn) : set (subs (snd fn)) & infinite (deriv (snd fsucn)) & ~ is_axiom (s_of_ns (snd fsucn))))"
locale loc1 =
fixes s and f
assumes f: "f = failing_path (deriv s)"
lemma (in loc1) f0: "infinite (deriv s) ==> f 0 \<in> (deriv s) & fst (f 0) = 0 & infinite (deriv (snd (f 0))) & ~ is_axiom (s_of_ns (snd (f 0)))"
apply(simp add: f) apply(rule someI_ex) apply(rule_tac x="(0,s)" in exI) apply(force simp add: deriv_is_axiom) done
lemma infinite_union: "finite Y ==> infinite (Union (f ` Y)) ==> \<exists>y. y \<in> Y & infinite (f y)"
apply(rule ccontr) apply(simp) done
lemma infinite_inj_infinite_image: "inj_on f Z ==> infinite (f ` Z) = infinite Z"
apply(rule ccontr)
apply(simp)
apply(force dest: finite_imageD)
done
lemma inj_inj_on: "inj f ==> inj_on f A"
by(blast intro: subset_inj_on)
lemma collect_disj: "{x. P x | Q x} = {x. P x} Un {x. Q x}" by(force)
lemma t: "finite {w. P w} ==> finite {w. Q w & P w}"
by (rule finite_subset, auto)
lemma finite_subs: "finite {w. ~is_axiom (s_of_ns y) & w : set (subs y)}"
apply(case_tac y) apply(simp add: s_of_ns_def)
apply(case_tac a) apply(case_tac b)
apply(simp_all only: subs.simps)
apply(auto intro: t simp add: collect_disj)
done
lemma (in loc1) fSuc:
shows "[|
f n \<in> deriv s & fst (f n) = n & infinite (deriv (snd (f n))) & ~is_axiom (s_of_ns (snd (f n)))
|] ==> f (Suc n) \<in> deriv s & fst (f (Suc n)) = Suc n & (snd (f (Suc n))) : set (subs (snd (f n))) & infinite (deriv (snd (f (Suc n)))) & ~is_axiom (s_of_ns (snd (f (Suc n))))"
apply(simp add: Let_def f)
apply(rule_tac someI_ex)
apply(simp only: f[symmetric])
apply(drule_tac subst[OF deriv[of "snd (f n)"] ])
apply(simp only: finite_insert) apply(subgoal_tac "infinite (\<Union>(deriv ` {w. ~is_axiom (s_of_ns (snd (f n))) & w : set (subs (snd (f n)))}))")
apply(drule_tac infinite_union[OF finite_subs]) apply(erule exE) apply(rule_tac x="(Suc n, y)" in exI)
apply(clarify) apply(simp) apply(case_tac "f n") apply(simp add: step) apply(force simp add: is_axiom_finite_deriv)
apply(force simp add: infinite_inj_infinite_image inj_inj_on)
done
lemma (in loc1) is_path_f_0: "infinite (deriv s) ==> f 0 = (0,s)"
apply(subgoal_tac "f 0 \<in> deriv s & fst (f 0) = 0")
prefer 2 apply(frule_tac f0) apply(simp)
apply(case_tac "f 0") apply(elim conjE, simp)
done
lemma (in loc1) is_path_f': "infinite (deriv s) ==> \<forall>n. f n \<in> deriv s & fst (f n) = n & infinite (deriv (snd (f n))) & ~ is_axiom (s_of_ns (snd (f n)))"
apply(rule)
apply(induct_tac n)
apply(simp add: f0)
apply(simp add: fSuc)
done
lemma (in loc1) is_path_f: "infinite (deriv s) ==> \<forall>n. f n \<in> deriv s & fst (f n) = n & (snd (f (Suc n))) : set (subs (snd (f n))) & infinite (deriv (snd (f n)))"
apply(rule)
apply(frule_tac is_path_f') apply(simp) apply(simp add: fSuc)
done
subsection "Models"
typedecl U
type_synonym model = "U set * (pred => U list => bool)"
type_synonym env = "var => U"
primrec FEval :: "model => env => form => bool"
where
"FEval MI e (PAtom P vs) = (let IP = (snd MI) P in IP (map e vs))"
| "FEval MI e (NAtom P vs) = (let IP = (snd MI) P in ~ (IP (map e vs)))"
| "FEval MI e (FConj f g) = ((FEval MI e f) & (FEval MI e g))"
| "FEval MI e (FDisj f g) = ((FEval MI e f) | (FEval MI e g))"
| "FEval MI e (FAll f) = (\<forall>m \<in> (fst MI). FEval MI (% y. case y of 0 => m | Suc n => e n) f)"
| "FEval MI e (FEx f) = (\<exists>m \<in> (fst MI). FEval MI (% y. case y of 0 => m | Suc n => e n) f)"
lemma and_lem: "(a=c) ==> (b=d) ==> (a & b) = (c & d)" by simp
lemma or_lem: "(a=c) ==> (b=d) ==> (a | b) = (c | d)" by simp
(*lemma all_eq_all: "\<forall>x. P x = Q x ==> (\<forall>x. P x) = (\<forall>x. Q x)" by blast*)
lemma ball_eq_ball: "\<forall>x \<in> m. P x = Q x ==> (\<forall>x \<in> m. P x) = (\<forall>x \<in> m. Q x)" by blast
lemma bex_eq_bex: "\<forall>x \<in> m. P x = Q x ==> (\<exists>x \<in> m. P x) = (\<exists>x \<in> m. Q x)" by blast
lemma preSuc[simp]:"\<forall>n. Suc n \<in> set A = (n \<in> set (preSuc A))"
apply(induct_tac A) apply(simp) apply(case_tac a, force, force) done
lemma FEval_cong: "\<forall>e1 e2. (\<forall>x. x \<in> set (fv A) --> e1 x = e2 x) --> FEval MI e1 A = FEval MI e2 A"
proof (induction A)
case (PAtom x1 x2)
then show ?case
by (metis FEval.simps(1) fv.simps(1) map_cong)
next
case (NAtom x1 x2)
then show ?case
by simp (metis list.map_cong0)
next
case (FConj A1 A2)
then show ?case
by simp blast
next
case (FDisj A1 A2)
then show ?case
by simp blast
next
case (FAll A)
then show ?case
by (metis (no_types, lifting) FEval.simps(5) Nitpick.case_nat_unfold One_nat_def Suc_pred fv.simps(5) gr0I preSuc)
next
case (FEx A)
then show ?case
by (metis (no_types, lifting) FEval.simps(6) Nitpick.case_nat_unfold One_nat_def Suc_pred fv.simps(6) gr0I preSuc)
qed
primrec SEval :: "model => env => form list => bool"
where
"SEval m e [] = False"
| "SEval m e (x#xs) = (FEval m e x | SEval m e xs)"
lemma SEval_def2: "SEval m e s = (\<exists>f. f \<in> set s & FEval m e f)"
by (induct s) auto
lemma SEval_append: "SEval m e (xs@ys) = ( (SEval m e xs) | (SEval m e ys))"
by (induct xs) auto
lemma all_eq_all: "\<forall>x. P x = Q x ==> (\<forall>x. P x) = (\<forall>x. Q x)" by blast
lemma all_conj: "(\<forall>x. A x & B x) = ((\<forall>x. A x) & (\<forall>x. B x))" by blast
lemma SEval_cong: "(\<forall>x. x \<in> set (sfv s) --> e1 x = e2 x) --> SEval m e1 s = SEval m e2 s"
apply(induct s) apply(simp)
apply(simp) apply(intro allI impI)
apply(rule or_lem) apply(simp add: sfv_cons) apply(simp add: FEval_cong)
apply(simp add: sfv_cons)
done
definition
is_env :: "model => env => bool" where
"is_env MI e = (\<forall>x. e x \<in> (fst MI))"
definition
Svalid :: "form list => bool" where
"Svalid s = (\<forall>MI e. is_env MI e --> SEval MI e s)"
subsection "Soundness"
lemma fold_compose1: "(% x. f (g x)) = (f o g)"
apply(rule ext) apply(force) done
lemma FEval_subst: "\<forall>e f. (FEval MI e (subst f A)) = (FEval MI (e o f) A)"
apply(induct A)
apply(simp add: Let_def) apply(simp only: fold_compose1 map_map) apply(blast)
apply(simp add: Let_def) apply(simp only: fold_compose1 map_map) apply(blast)
apply(simp)
apply(simp)
apply(simp (no_asm_use)) apply(simp) apply(rule,rule) apply(rule ball_eq_ball) apply(rule) apply(simp add: bump_def)
apply(subgoal_tac "(%u. case_nat m e (case u of 0 => 0 | Suc n => Suc (f n))) = (case_nat m (%n. e (f n)))") apply(simp)
apply(rule ext) apply(case_tac u)
apply(simp) apply(simp)
apply(simp (no_asm_use)) apply(simp) apply(rule,rule) apply(rule bex_eq_bex) apply(rule) apply(simp add: bump_def)
apply(subgoal_tac "(%u. case_nat m e (case u of 0 => 0 | Suc n => Suc (f n))) = (case_nat m (%n. e (f n)))") apply(simp)
apply(rule ext) apply(case_tac u)
apply(simp) apply(simp)
done
(*
lemma bump_f_x_0[simp]: "bump f x 0 = x" apply(simp add: bump_def) done
lemma bump_id_suc[simp]: "bump id x (Suc n) = Suc n" apply(simp add: bump_def) done
lemma bump_id_0[simp]: "bump id 0 = id" apply(rule ext) apply(simp add: bump_def) apply(case_tac x) apply(auto) done
*)
lemma FEval_finst: "FEval mo e (finst A u) = FEval mo (case_nat (e u) e) A"
apply(simp add: finst_def)
apply(simp add: FEval_subst)
apply(subgoal_tac "(e o case_nat u (%n. n)) = (case_nat (e u) e)") apply(simp)
apply(rule ext)
apply(case_tac x,auto)
done
lemma ball_maxscope: "(\<forall>x \<in> m. P x | Q) ==> (\<forall>x \<in> m. P x) | Q " by blast
lemma sound_FAll: "u \<notin> set (sfv (FAll f#s)) ==> Svalid (s@[finst f u]) ==> Svalid (FAll f#s)"
apply(simp add: Svalid_def del: SEval.simps)
apply(rule allI)
apply(rule allI)
apply(rename_tac M I)
apply(rule allI) apply(rule)
apply(simp)
apply(simp add: SEval_append)
apply(rule ball_maxscope)
apply(rule)
apply(simp add: FEval_finst)
apply(drule_tac x=M in spec, drule_tac x=I in spec) \<comment> \<open>this is the goal\<close>
apply(drule_tac x="e(u:=m)" in spec) apply(erule impE) apply(simp add: is_env_def) apply(erule disjE)
apply(rule disjI2)
apply(subgoal_tac "SEval (M,I) (e(u :=m)) s = SEval (M,I) e s")
apply(simp)
apply(rule SEval_cong[rule_format]) apply(simp add: sfv_cons) apply(force)
apply(rule disjI1)
apply(simp)
apply(subgoal_tac "FEval (M,I) (case_nat m (e(u :=m))) f = FEval (M,I) (case_nat m e) f")
apply(simp)
apply(rule FEval_cong[rule_format])
apply(case_tac x, simp)
apply(simp)
apply(simp only: preSuc[rule_format, symmetric])
apply(subgoal_tac "nat \<in> set (fv (FAll f))") prefer 2 apply(simp)
apply(force simp: sfv_cons)
done
\<comment> \<open>phew, that really was a bit more difficult than expected\<close>
\<comment> \<open>note that we can avoid maxscoping at the cost of instantiating the hyp twice- an additional time for M\<close>
\<comment> \<open>different proof, instantiating quantifier twice, avoiding maxscoping\<close>
lemma sound_FAll': "u \<notin> set (sfv (FAll f#s)) ==> Svalid (s@[finst f u]) ==> Svalid (FAll f#s)"
apply(simp add: Svalid_def del: SEval.simps)
apply(rule allI)
apply(rule allI)
apply(rename_tac M I)
apply(rule allI) apply(rule)
apply(simp)
apply(simp add: SEval_append)
apply(drule_tac x=M in spec, drule_tac x=I in spec)
apply(frule_tac x="e" in spec) apply(erule impE) apply(simp add: is_env_def)
apply(case_tac "SEval (M, I) e s") apply(simp)
apply(simp)
\<comment> \<open>second instantiation\<close>
apply(simp add: FEval_finst)
apply(rule)
apply(drule_tac x="e(u:=m)" in spec) apply(erule impE) apply(simp add: is_env_def)
apply(erule disjE)
apply(subgoal_tac "SEval (M,I) (e(u :=m)) s = SEval (M,I) e s")
apply(simp)
apply(rule SEval_cong[rule_format]) apply(simp add: sfv_cons) apply(force)
apply(simp)
apply(subgoal_tac "FEval (M,I) (case_nat m (e(u :=m))) f = FEval (M,I) (case_nat m e) f")
apply(simp)
apply(rule FEval_cong[rule_format])
apply(case_tac x, simp)
apply(simp)
apply(simp only: preSuc[rule_format, symmetric])
apply(subgoal_tac "nat \<in> set (fv (FAll f))") prefer 2 apply(simp)
apply(force simp: sfv_cons)
done
\<comment> \<open>not much better, probably slightly worse\<close>
(*
lemma sound_FAll: "u \<notin> sfv ((0,FAll f)#s) ==> SEval m e (map snd (s@[(0,finst u f)])) ==> SEval m e (map snd ((0,FAll f)#s))"
apply(simp)
apply(rule all_maxscope)
apply(rule)
apply(simp add: SEval_append)
apply(erule disjE)
apply(simp)
apply(rule disjI1)
apply(simp add: FEval_finst)
apply(subgoal_tac "FEval m (case_nat (e u) e) f = FEval m (case_nat x e) f") apply(simp)
apply(rule FEval_cong[rule_format])
apply(case_tac xa)
apply(simp)
oops -- "and this isn't provable- the heart of the rules-preserve-validity debate"
*)
lemma sound_FEx: "Svalid (s@[finst f u,FEx f]) ==> Svalid (FEx f#s)"
apply(simp add: Svalid_def del: SEval.simps)
apply(rule allI)
apply(rule allI)
apply(rename_tac ms m)
apply(rule) apply(rule)
apply(simp)
apply(simp add: SEval_append)
apply(simp add: FEval_finst)
apply(drule_tac x=ms in spec, drule_tac x=m in spec)
apply(drule_tac x=e in spec) apply(erule impE) apply(assumption)
apply(erule disjE)
apply(simp)
apply(erule disjE)
apply(rule disjI1)
apply(rule_tac x="e u" in bexI) apply(simp) apply(simp add: is_env_def)
apply(force)
done
lemma max_exists: "finite (X::nat set) ==> X \<noteq> {} --> (\<exists>x. x \<in> X & (\<forall>y. y \<in> X --> y \<le> x))"
apply(erule_tac finite_induct)
apply(force)
apply(rule)
apply(case_tac "F = {}")
apply(force)
apply(erule impE) apply(force)
apply(elim exE conjE)
apply(rule_tac x="max x xa" in exI)
apply(rule) apply(simp add: max_def)
apply(simp add: max_def) apply(force)
done
\<comment> \<open>poor max lemmas in lib\<close>
lemma inj_finite_image_eq_finite: "inj_on f Z ==> finite (f ` Z) = finite Z"
by (blast intro: finite_imageD)
lemma finite_inc: "finite (inc ` X) = finite X"
apply(rule inj_finite_image_eq_finite)
apply(rule_tac B=UNIV in subset_inj_on)
apply(auto)
done
lemma finite_deriv_deriv: "finite (deriv s) ==> finite (deriv ` {w. ~is_axiom (s_of_ns s) & w : set (subs s)})"
by (simp only: deriv) simp
definition
init :: "nseq => bool" where
"init s = (\<forall>x \<in> (set s). fst x = 0)"
definition
is_FEx :: "form => bool" where
"is_FEx f = (case f of
PAtom p vs => False
| NAtom p vs => False
| FConj f g => False
| FDisj f g => False
| FAll f => False
| FEx f => True)"
lemma is_FEx[simp]: "~ is_FEx (PAtom p vs)
& ~ is_FEx (NAtom p vs)
& ~ is_FEx (FConj f g)
& ~ is_FEx (FDisj f g)
& ~ is_FEx (FAll f)"
by(force simp: is_FEx_def)
lemma index0: "init s ==> \<forall>u m A. (n, u) \<in> deriv s --> (m,A) \<in> (set u) --> (~ is_FEx A) --> m = 0"
apply(induct_tac n)
apply(rule,rule,rule,rule,rule,rule) apply(simp) apply(force simp add: init_def)
apply(rule,rule,rule,rule,rule,rule)
\<comment> \<open>inversion on @{term "(Suc n, u) \<in> deriv s"}\<close>
apply(drule_tac deriv_downwards) apply(elim exE conjE)
apply(case_tac y) apply(simp)
apply(case_tac a) apply(case_tac b)
apply(force simp add: Let_def s_of_ns_def)
apply(force simp add: Let_def s_of_ns_def)
apply(force simp add: Let_def s_of_ns_def)
apply(force simp add: Let_def s_of_ns_def)
apply(force simp add: Let_def s_of_ns_def)
apply(force simp add: is_FEx_def Let_def s_of_ns_def)
done
lemma soundness': "init s ==> finite (deriv s) ==> m \<in> (fst ` (deriv s)) ==> \<forall>y u. (y,u) \<in> (deriv s) --> y \<le> m ==> \<forall>n t. h = m - n & (n,t) \<in> deriv s --> Svalid (s_of_ns t)"
apply(induct_tac h)
\<comment> \<open>base case\<close>
apply(simp) apply(rule,rule,rule) apply(elim conjE)
apply(subgoal_tac "n=m") prefer 2 apply(force)
apply(simp)
apply(simp add: Svalid_def) apply(rule,rule) apply(rename_tac gs g) apply(rule) apply(rule) apply(simp add: SEval_def2)
apply(case_tac "is_axiom (s_of_ns t)")
\<comment> \<open>base case, is axiom\<close>
apply(simp add: s_of_ns_def) apply(case_tac t) apply(simp) apply(simp)
apply(erule disjE)
\<comment> \<open>base case, is axiom, starts with PAtom\<close>
apply(elim conjE exE)
apply(subgoal_tac "FEval (gs,g) e (PAtom p vs) | FEval (gs,g) e (NAtom p vs)")
apply(erule disjE) apply(force) apply(rule_tac x="NAtom p vs" in exI) apply(force)
apply(simp add: Let_def)
\<comment> \<open>base case, is axiom, starts with NAtom\<close>
apply(elim conjE exE)
apply(subgoal_tac "FEval (gs,g) e (PAtom p vs) | FEval (gs,g) e (NAtom p vs)")
apply(erule disjE) apply(rule_tac x="PAtom p vs" in exI) apply(force) apply(force)
apply(simp add: Let_def)
\<comment> \<open>base case, not is axiom: if not a satax, then subs holds... but this can't be\<close>
apply(drule_tac deriv_upwards) apply(assumption) apply(elim conjE exE) apply(force)
\<comment> \<open>step case, by case analysis\<close>
apply(intro allI impI) apply(elim exE conjE)
apply(case_tac "is_axiom (s_of_ns t)")
\<comment> \<open>step case, is axiom\<close>
apply(simp add: Svalid_def) apply(rule,rule) apply(rename_tac gs g) apply(rule) apply(rule) apply(simp add: SEval_def2)
apply(simp add: s_of_ns_def) apply(case_tac t) apply(simp) apply(simp)
apply(erule disjE)
apply(elim conjE exE)
apply(subgoal_tac "FEval (gs,g) e (PAtom p vs) | FEval (gs,g) e (NAtom p vs)")
apply(erule disjE) apply(force) apply(rule_tac x="NAtom p vs" in exI) apply(blast)
apply(simp add: Let_def)
apply(elim conjE exE)
apply(subgoal_tac "FEval (gs,g) e (PAtom p vs) | FEval (gs,g) e (NAtom p vs)")
apply(erule disjE) apply(rule_tac x="PAtom p vs" in exI) apply(blast) apply(simp) apply(force)
apply(simp add: Let_def)
\<comment> \<open>we hit FAll/ FEx cases first\<close>
apply(case_tac "\<exists>(a::nat) f list. t = (a,FAll f)#list")
apply(elim exE) apply(simp)
apply(subgoal_tac "a = 0")
prefer 2
apply(rule_tac n=na and u=t and A="FAll f" in index0[rule_format])
apply(assumption) apply(simp) apply(simp) apply(simp)
apply(frule_tac deriv.step) apply(simp add: Let_def) \<comment> \<open>nice use of simp to instantiate\<close>
apply(drule_tac x="Suc na" in spec, drule_tac x="list @ [(0, finst f (newvar (sfv (s_of_ns t))))]" in spec) apply(erule impE, simp)
apply(subgoal_tac "newvar (sfv (s_of_ns t)) \<notin> set (sfv (s_of_ns t))")
prefer 2 apply(rule newvar)
apply(simp)
apply(simp add: s_of_ns_def)
apply(frule_tac sound_FAll) apply(simp) apply(simp)
apply(case_tac "\<exists>a f list. t = (a,FEx f)#list")
apply(elim exE)
apply(frule_tac deriv.step) apply(simp add: Let_def) \<comment> \<open>nice use of simp to instantiate\<close>
apply(drule_tac x="Suc na" in spec, drule_tac x="list @ [(0, finst f a), (Suc a, FEx f)]" in spec) apply(erule impE, assumption)
apply(drule_tac x="Suc na" in spec, drule_tac x="list @ [(0, finst f a), (Suc a, FEx f)]" in spec) apply(erule impE) apply(rule) apply(arith) apply(assumption)
apply(subgoal_tac "Svalid (s_of_ns (list@[(0,finst f a), (Suc a, FEx f)]))")
apply(simp add: s_of_ns_def)
apply(frule_tac sound_FEx) apply(simp) apply(simp)
\<comment> \<open>now for other cases\<close>
\<comment> \<open>case empty list\<close>
apply(case_tac t) apply(simp)
apply(frule_tac step) apply(simp) apply(simp) apply(drule_tac x="Suc na" in spec) back apply(drule_tac x="[]" in spec) apply(erule impE) apply(rule) apply(arith) apply(assumption) apply(assumption)
apply(simp add: Svalid_def) apply(rule,rule) apply(rename_tac gs g) apply(rule) apply(rule) apply(simp add: SEval_def2)
\<comment> \<open>na t in deriv, so too is subs\<close>
\<comment> \<open>if not a satax, then subs holds...\<close>
apply(case_tac a)
apply(case_tac b)
apply(simp del: FEval.simps) apply(frule_tac patom) apply(assumption)
apply(rename_tac nat lista)
apply(frule_tac x="Suc na" in spec, drule_tac x="list @ [(0, PAtom nat lista)]" in spec)
apply(erule impE) apply(arith)
apply(drule_tac x=gs in spec, drule_tac x=g in spec, drule_tac x=e in spec) apply(erule impE) apply(simp add: is_env_def)
apply(elim exE conjE) apply(rule_tac x=f in exI) apply(simp add: s_of_ns_def) \<comment> \<open>weirdly, simp succeeds, force and blast fail\<close>
apply(simp del: FEval.simps) apply(frule_tac natom) apply(assumption)
apply(rename_tac nat lista)
apply(frule_tac x="Suc na" in spec, drule_tac x="list @ [(0, NAtom nat lista)]" in spec)
apply(erule impE) apply(arith)
apply(drule_tac x=gs in spec, drule_tac x=g in spec, drule_tac x=e in spec) apply(erule impE, simp)
apply(elim exE conjE) apply(rule_tac x=f in exI) apply(simp add: s_of_ns_def)
apply(simp del: FEval.simps) apply(frule_tac fconj1) apply(assumption) apply(frule_tac fconj2) apply(assumption)
apply(rename_tac form1 form2)
apply(frule_tac x="Suc na" in spec, drule_tac x="list @ [(0, form1)]" in spec)
apply(erule impE) apply(arith)
apply(drule_tac x=gs in spec, drule_tac x=g in spec, drule_tac x=e in spec) apply(erule impE, simp) apply(elim exE conjE)
apply(drule_tac x="Suc na" in spec, drule_tac x="list @ [(0, form2)]" in spec)
apply(erule impE) apply(arith)
apply(drule_tac x=gs in spec, drule_tac x=g in spec, drule_tac x=e in spec) apply(erule impE, simp) apply(elim exE conjE)
apply(simp only: s_of_ns_def)
apply(simp)
apply(elim disjE)
apply(simp) apply(rule_tac x="FConj form1 form2" in exI) apply(simp)
apply(simp) apply(rule_tac x="fa" in exI) apply(simp)
apply(simp) apply(rule_tac x="f" in exI) apply(simp)
apply(rule_tac x="f" in exI, simp)
apply(simp del: FEval.simps) apply(frule_tac fdisj) apply(assumption)
apply(rename_tac form1 form2)
apply(frule_tac x="Suc na" in spec, drule_tac x="list @ [(0, form1),(0,form2)]" in spec)
apply(erule impE) apply(arith)
apply(drule_tac x=gs in spec, drule_tac x=g in spec, drule_tac x=e in spec) apply(erule impE, simp) apply(elim exE conjE)
apply(simp add: s_of_ns_def)
apply(elim disjE)
apply(simp) apply(rule_tac x="FDisj form1 form2" in exI) apply(simp)
apply(simp) apply(rule_tac x="FDisj form1 form2" in exI) apply(simp)
apply(rule_tac x="f" in exI) apply(simp)
\<comment> \<open>all case\<close>
apply(force)
apply(force)
done
lemma [simp]: "s_of_ns (ns_of_s s) = s"
apply(induct s)
apply(simp add: s_of_ns_def ns_of_s_def)
apply(simp add: s_of_ns_def ns_of_s_def)
done
lemma soundness: "finite (deriv (ns_of_s s)) ==> Svalid s"
apply(subgoal_tac "init (ns_of_s s)")
prefer 2 apply(simp add: init_def ns_of_s_def)
apply(subgoal_tac "finite (fst ` (deriv (ns_of_s s)))") prefer 2 apply(simp)
apply(frule_tac max_exists) apply(erule impE) apply(simp) apply(subgoal_tac "(0,ns_of_s s) \<in> deriv (ns_of_s s)") apply(force) apply(simp)
apply(elim exE conjE)
apply(frule_tac soundness') apply(assumption) apply(assumption) apply(force)
apply(drule_tac x=0 in spec, drule_tac x="ns_of_s s" in spec)
apply(force )
done
subsection "Contains, Considers"
definition
"contains" :: "(nat => (nat*nseq)) => nat => nform => bool" where
"contains f n nf = (nf \<in> set (snd (f n)))"
definition
considers :: "(nat => (nat*nseq)) => nat => nform => bool" where
"considers f n nf = (case snd (f n) of [] => False | (x#xs) => x = nf)"
lemma (in loc1) progress: "infinite (deriv s) ==> snd (f n) = a#list --> (\<exists>zs'. snd (f (Suc n)) = list@zs')"
apply(subgoal_tac "(snd (f (Suc n))) : set (subs (snd (f n)))") defer apply(frule_tac is_path_f) apply(blast)
apply(case_tac a)
apply(case_tac b)
apply(safe)
apply(simp_all add: Let_def split: if_splits)
apply(erule disjE)
apply(simp_all)
done
lemma (in loc1) contains_considers': "infinite (deriv s) ==> \<forall>n y ys. snd (f n) = xs@y#ys --> (\<exists>m zs'. snd (f (n+m)) = y#zs')"
apply(induct_tac xs)
apply(rule,rule,rule,rule) apply(rule_tac x=0 in exI) apply(rule_tac x=ys in exI) apply(force)
apply(rule,rule,rule,rule) apply(simp) apply(frule_tac progress) apply(erule impE) apply(assumption)
apply(erule exE) apply(simp)
apply(drule_tac x="Suc n" in spec)
apply(case_tac y) apply(rename_tac u v)
apply(drule_tac x="u" in spec)
apply(drule_tac x="v" in spec)
apply(erule impE) apply(force)
apply(elim exE)
apply(rule_tac x="Suc m" in exI)
apply(force)
done
lemma list_decomp[rule_format]: "v \<in> set p --> (\<exists> xs ys. p = xs@(v#ys) \<and> v \<notin> set xs)"
apply(induct p)
apply(force)
apply(case_tac "v=a")
apply(force)
apply(auto)
apply(rule_tac x="a#xs" in exI)
apply(auto)
done
lemma (in loc1) contains_considers: "infinite (deriv s) ==> contains f n y ==> (\<exists>m. considers f (n+m) y)"
apply(simp add: contains_def considers_def)
apply(frule_tac list_decomp) apply(elim exE conjE)
apply(frule_tac contains_considers'[rule_format]) apply(assumption) apply(elim exE)
apply(rule_tac x=m in exI)
apply(force)
done
lemma (in loc1) contains_propagates_patoms[rule_format]: "infinite (deriv s) ==> contains f n (0, PAtom p vs) --> contains f (n+q) (0, PAtom p vs)"
apply(induct_tac q) apply(simp)
apply(rule)
apply(erule impE) apply(assumption)
apply(subgoal_tac "~is_axiom (s_of_ns (snd (f (n+na))))") defer
apply(subgoal_tac "infinite (deriv (snd (f (n+na))))") defer
apply(force dest: is_path_f)
defer
apply(rule) apply(simp add: deriv_is_axiom)
apply(simp add: contains_def)
apply(drule_tac p="snd (f (n + na))" in list_decomp)
apply(elim exE conjE)
apply(case_tac xs)
apply(simp)
apply(subgoal_tac "(snd (f (Suc (n + na)))) : set (subs (snd (f (n + na))))")
apply(simp add: Let_def split: if_splits)
apply(frule_tac is_path_f) apply(drule_tac x="n+na" in spec) apply(force)
apply(drule_tac progress)
apply(erule impE) apply(force)
apply(force)
done
lemma (in loc1) contains_propagates_natoms[rule_format]: "infinite (deriv s) ==> contains f n (0, NAtom p vs) --> contains f (n+q) (0, NAtom p vs)"
apply(induct_tac q) apply(simp)
apply(rule)
apply(erule impE) apply(assumption)
apply(subgoal_tac "~is_axiom (s_of_ns (snd (f (n+na))))") defer
apply(subgoal_tac "infinite (deriv (snd (f (n+na))))") defer
apply(force dest: is_path_f)
defer
apply(rule) apply(simp add: deriv_is_axiom)
apply(simp add: contains_def)
apply(drule_tac p = "snd (f (n + na))" in list_decomp)
apply(elim exE conjE)
apply(case_tac xs)
apply(simp)
apply(subgoal_tac "(snd (f (Suc (n + na)))) : set (subs (snd (f (n + na))))")
apply(simp add: Let_def split: if_splits)
apply(frule_tac is_path_f) apply(drule_tac x="n+na" in spec) apply(force)
apply(drule_tac progress)
apply(erule impE) apply(force)
apply(force)
done
lemma (in loc1) contains_propagates_fconj: "infinite (deriv s) ==> contains f n (0, FConj g h) ==> (\<exists>y. contains f (n+y) (0,g) | contains f (n+y) (0,h))"
apply(subgoal_tac "(\<exists>l. considers f (n+l) (0,FConj g h))") defer apply(rule contains_considers) apply(assumption) apply(assumption)
apply(erule exE)
apply(rule_tac x="Suc l" in exI)
apply(simp add: considers_def) apply(case_tac "snd (f (n + l))", simp)
apply(simp)
apply(subgoal_tac "(snd (f (Suc (n + l)))) : set (subs (snd (f (n + l))))")
apply(simp add: contains_def Let_def) apply(force)
apply(frule_tac is_path_f) apply(drule_tac x="n+l" in spec) apply(force)
done
lemma (in loc1) contains_propagates_fdisj: "infinite (deriv s) ==> contains f n (0, FDisj g h) ==> (\<exists>y. contains f (n+y) (0,g) & contains f (n+y) (0,h))"
apply(subgoal_tac "(\<exists>l. considers f (n+l) (0,FDisj g h))") defer apply(rule contains_considers) apply(assumption) apply(assumption)
apply(erule exE)
apply(rule_tac x="Suc l" in exI)
apply(simp add: considers_def) apply(case_tac "snd (f (n + l))", simp)
apply(simp)
apply(subgoal_tac " (snd (f (Suc (n + l)))) : set (subs (snd (f (n + l))))")
apply(simp add: contains_def Let_def)
apply(frule_tac is_path_f) apply(drule_tac x="n+l" in spec) apply(force)
done
lemma (in loc1) contains_propagates_fall: "infinite (deriv s) ==> contains f n (0, FAll g)
==> (\<exists>y. contains f (Suc(n+y)) (0,finst g (newvar (sfv (s_of_ns (snd (f (n+y))))))))" \<comment> \<open>may need constraint on fv\<close>
apply(subgoal_tac "(\<exists>l. considers f (n+l) (0,FAll g))") defer apply(rule contains_considers) apply(assumption) apply(assumption)
apply(erule exE)
apply(rule_tac x="l" in exI)
apply(simp add: considers_def) apply(case_tac "snd (f (n + l))", simp)
apply(simp)
apply(subgoal_tac "(snd (f (Suc (n + l)))) : set (subs (snd (f (n + l))))")
apply(simp add: contains_def Let_def)
apply(frule_tac is_path_f) apply(drule_tac x="n+l" in spec) apply(force)
done
lemma (in loc1) contains_propagates_fex: "infinite (deriv s) ==> contains f n (m, FEx g)
==> (\<exists>y.
(contains f (n+y) (0,finst g m))
& (contains f (n+y) (Suc m,FEx g)))"
apply(subgoal_tac "(\<exists>l. considers f (n+l) (m,FEx g))") defer apply(rule contains_considers) apply(assumption) apply(assumption)
apply(erule exE)
apply(rule_tac x="Suc l" in exI)
apply(simp add: considers_def) apply(case_tac "snd (f (n + l))", simp)
apply(simp)
apply(subgoal_tac " (snd (f (Suc (n + l)))) : set (subs (snd (f (n + l))))")
apply(simp add: contains_def Let_def)
apply(frule_tac is_path_f) apply(drule_tac x="n+l" in spec) apply(force)
done
\<comment> \<open>also need that if contains one, then contained an original at beginning\<close>
\<comment> \<open>existentials: show that for exists formulae, if Suc m is marker, then there must have been m\<close>
\<comment> \<open>show this by showing that nodes are upwardly closed, i.e. if never contains (m,x), then never contains (Suc m, x), by induction on n\<close>
lemma (in loc1) FEx_downward: "infinite (deriv s) ==> init s ==> \<forall>m. (Suc m,FEx g) \<in> set (snd (f n)) --> (\<exists>n'. (m, FEx g) \<in> set (snd (f n')))"
apply(frule_tac is_path_f)
apply(induct_tac n)
apply(drule_tac x="0" in spec) apply(case_tac "f 0") apply(force simp: init_def)
apply(intro allI impI)
apply(frule_tac x="Suc n" in spec, elim conjE) apply(drule_tac x="n" in spec, elim conjE)
apply(thin_tac "(snd (f (Suc (Suc n)))) : set (subs (snd (f (Suc n))))")
apply(case_tac "f n") apply(simp)
apply(case_tac b) apply(simp)
apply(case_tac aa) apply(case_tac ba)
apply(simp add: Let_def split: if_splits)
apply(simp add: Let_def split: if_splits)
apply(force simp add: Let_def)
apply(force simp add: Let_def)
apply(force simp add: Let_def)
apply(rename_tac form)
apply(case_tac "(ab, FEx form) = (m, FEx g)")
apply(rule_tac x=n in exI) apply(force)
apply(force simp add: Let_def)
done
lemma (in loc1) FEx0: "infinite (deriv s) ==> init s ==> \<forall>n. contains f n (m,FEx g) --> (\<exists>n'. contains f n' (0, FEx g))"
apply(simp add: contains_def)
apply(induct_tac m)
apply(force)
apply(intro allI impI) apply(erule exE)
apply(drule_tac FEx_downward[rule_format]) apply(assumption) apply(force)
apply(elim exE conjE)
apply(force)
done
lemma (in loc1) FEx_upward': "infinite (deriv s) ==> init s ==> \<forall>n. contains f n (0, FEx g) --> (\<exists>n'. contains f n' (m, FEx g))"
apply(intro allI impI)
apply(induct_tac m) apply(force)
apply(erule exE)
apply(frule_tac n=n' in contains_considers) apply(assumption)
apply(erule exE)
apply(rule_tac x="Suc(n'+m)" in exI)
apply(frule_tac is_path_f)
apply(simp add: considers_def) apply(case_tac "snd (f (n'+m))") apply(force)
apply(simp)
apply(drule_tac x="n'+m" in spec)
apply(force simp add: contains_def considers_def Let_def)
done
\<comment> \<open>FIXME contains and considers aren't buying us much\<close>
lemma (in loc1) FEx_upward: "infinite (deriv s) ==> init s ==> contains f n (m, FEx g) ==> (\<forall>m'. \<exists>n'. contains f n' (0, finst g m'))"
apply(intro allI impI)
apply(subgoal_tac "\<exists>n'. contains f n' (m', FEx g)") defer
apply(frule_tac m = m and g = g in FEx0) apply(assumption)
apply(drule_tac x=n in spec)
apply(simp)
apply(elim exE)
apply(frule_tac g=g and m=m' in FEx_upward') apply(assumption)
apply (blast dest: contains_propagates_fex intro: elim:)+
done
subsection "Models 2"
axiomatization ntou :: "nat => U"
where ntou: "inj ntou" \<comment> \<open>assume universe set is infinite\<close>
definition uton :: "U => nat" where "uton = inv ntou"
lemma uton_ntou: "uton (ntou x) = x"
apply(simp add: uton_def) apply(simp add: ntou inv_f_f) done
lemma map_uton_ntou[simp]: "map uton (map ntou xs) = xs"
apply(induct xs, auto simp: uton_ntou) done
lemma ntou_uton: "x \<in> range ntou ==> ntou (uton x) = x"
apply(simp add: uton_def)
apply(simp add: f_inv_into_f) done
subsection "Falsifying Model From Failing Path"
definition model :: "nseq => model" where
"model s = (range ntou, % p ms. (let f = failing_path (deriv s) in
(\<forall>n m. ~ contains f n (m,PAtom p (map uton ms)))))"
locale loc2 = loc1 +
fixes mo
assumes mo: "mo = model s"
lemma is_env_model_ntou: "is_env (model s) ntou"
apply(simp add: is_env_def) apply(simp add: model_def) done
lemma (in loc1) [simp]: "infinite (deriv s) ==> init s ==> (contains f n (m,A)) ==> ~ is_FEx A ==> m = 0"
apply(frule_tac n=n in index0)
apply(frule_tac is_path_f) apply(drule_tac x=n in spec) apply(case_tac "f n")
apply(simp)
apply(simp add: contains_def)
apply(force)
done
lemma (in loc2) model':
notes [simp] = FEval_subst
notes [simp del] = is_axiom.simps
shows "infinite (deriv s) ==> init s ==> \<forall>A. size A = h --> (\<forall>m n. contains f n (m,A) --> ~ (FEval mo ntou A))"
-
+ supply [[simproc del: defined_all]]
apply(rule_tac nat_less_induct) apply(rule, rule) apply(case_tac A)
apply(rule,rule,rule) apply(simp add: mo Let_def) apply(simp add: model_def Let_def del: map_map) apply(simp only: f[symmetric]) apply(force)
apply(rule,rule,rule) apply(simp add: mo Let_def) apply(simp add: model_def Let_def del: map_map) apply(simp only: f[symmetric]) apply(rule ccontr) apply(simp del: map_map) apply(elim exE)
apply(subgoal_tac "m = 0 & ma = 0") prefer 2 apply(simp del: map_map)
apply(simp del: map_map)
apply(rename_tac nat list m na nb ma)
apply(subgoal_tac "? y. considers f (nb+na+y) (0, PAtom nat list)") prefer 2 apply(rule contains_considers) apply(assumption)
apply(rule contains_propagates_patoms) apply(assumption) apply(assumption)
apply(erule exE)
apply(subgoal_tac "contains f (na+nb+y) (0, NAtom nat list)")
apply(subgoal_tac "nb+na=na+nb")
apply(simp) apply(subgoal_tac "is_axiom (s_of_ns (snd (f (na+nb+y))))")
apply(drule_tac is_axiom_finite_deriv) apply(force dest: is_path_f)
apply(simp add: contains_def considers_def) apply(case_tac "snd (f (na + nb + y))") apply(simp) apply(simp add: s_of_ns_def is_axiom.simps) apply(force)
apply(force)
apply(force intro: contains_propagates_natoms contains_propagates_patoms)
apply(intro impI allI)
apply(subgoal_tac "m=0") prefer 2 apply(simp) apply(simp del: FEval.simps)
apply(frule_tac contains_propagates_fconj) apply(assumption)
apply(rename_tac form1 form2 m na)
apply(frule_tac x="size form1" in spec) apply(erule impE) apply(force) apply(drule_tac x="form1" in spec) apply(drule_tac x="size form2" in spec) apply(erule impE) apply(force) apply(simp)
apply(force)
apply(intro impI allI)
apply(subgoal_tac "m=0") prefer 2 apply(simp) apply(simp del: FEval.simps)
apply(frule_tac contains_propagates_fdisj) apply(assumption)
apply(rename_tac form1 form2 m na)
apply(frule_tac x="size form1" in spec) apply(erule impE) apply(force) apply(drule_tac x="form1" in spec) apply(drule_tac x="size form2" in spec) apply(erule impE) apply(force) apply(simp)
apply(force)
apply(intro impI allI)
apply(subgoal_tac "m=0") prefer 2 apply(simp) apply(simp del: FEval.simps)
apply(frule_tac contains_propagates_fall) apply(assumption)
apply(erule exE) \<comment> \<open>all case\<close>
apply(rename_tac form m na y)
apply(drule_tac x="size form" in spec) apply(erule impE, force) apply(drule_tac x="finst form (newvar (sfv (s_of_ns (snd (f (na + y))))))" in spec) apply(erule impE, force)
apply(erule impE, force) apply(simp add: FEval_finst) apply(rule_tac x="ntou (newvar (sfv (s_of_ns (snd (f (na + y))))))" in bexI) apply(assumption)
using is_env_model_ntou[of s] apply(simp add: is_env_def mo)
apply(intro impI allI) apply(simp del: FEval.simps)
apply(frule_tac FEx_upward) apply(assumption) apply(assumption)
apply(simp)
apply(rule)
apply(rename_tac form m na ma)
apply(subgoal_tac "\<forall>m'. ~ FEval mo ntou (finst form m')")
prefer 2 apply(rule)
apply(drule_tac x="size form" in spec) apply(erule impE, force)
apply(drule_tac x="finst form m'" in spec) apply(erule impE, force) apply(erule impE) apply(force) apply(simp add: id_def)
apply(simp add: FEval_finst id_def)
apply(drule_tac x="uton ma" and P="%m'. ~ (P m')" for P in spec)
apply(subgoal_tac "ma \<in> range ntou") apply(frule_tac ntou_uton) apply(simp)
apply(simp add: mo model_def)
done
lemma (in loc2) model: "infinite (deriv s) ==> init s ==> (\<forall>A m n. contains f n (m,A) --> ~ (FEval mo ntou A))"
apply(rule)
apply(frule_tac model') apply(auto) done
subsection "Completeness"
lemma (in loc2) completeness': "infinite (deriv s) ==> init s ==> \<forall>mA \<in> set s. ~ FEval mo ntou (snd mA)" \<comment> \<open>FIXME tidy deriv s so that s consists of formulae only?\<close>
apply(frule_tac model) apply(assumption)
apply(rule)
apply(case_tac mA)
apply(drule_tac x="b" in spec) apply(drule_tac x="0" in spec) apply(drule_tac x=0 in spec) apply(erule impE)
apply(simp add: contains_def) apply(frule_tac is_path_f_0) apply(simp)
apply(subgoal_tac "a=0")
prefer 2 apply(simp only: init_def, force)
apply auto
done \<comment> \<open>FIXME very ugly\<close>
thm loc2.completeness'[simplified loc2_def loc2_axioms_def loc1_def]
lemma completeness': "infinite (deriv s) ==> init s ==> \<forall>mA \<in> set s. ~ FEval (model s) ntou (snd mA)"
apply(rule loc2.completeness'[simplified loc2_def loc2_axioms_def loc1_def])
apply(auto) done
lemma completeness'': "infinite (deriv (ns_of_s s)) ==> init (ns_of_s s) ==> \<forall>A. A \<in> set s --> ~ FEval (model (ns_of_s s)) ntou A"
apply(frule_tac completeness') apply(assumption)
apply(simp add: ns_of_s_def)
done
lemma completeness: "infinite (deriv (ns_of_s s)) ==> ~ Svalid s"
apply(subgoal_tac "init (ns_of_s s)")
prefer 2 apply(simp add: init_def ns_of_s_def)
apply(frule_tac completeness'') apply(assumption)
apply(simp add: Svalid_def)
apply(simp add: SEval_def2)
apply(rule_tac x="fst (model (ns_of_s s))" in exI)
apply(rule_tac x="snd (model (ns_of_s s))" in exI)
apply(rule_tac x="ntou" in exI)
apply(simp add: model_def)
apply(simp add: is_env_def)
done
\<comment> \<open>FIXME silly splitting of quantified pairs\<close>
subsection "Sound and Complete"
lemma "Svalid s = finite (deriv (ns_of_s s))"
using soundness completeness by blast
subsection "Algorithm"
primrec iter :: "('a => 'a) => 'a => nat => 'a" \<comment> \<open>fold for nats\<close>
where
"iter g a 0 = a"
| "iter g a (Suc n) = g (iter g a n)"
lemma iter: "\<forall>a. (iter g (g a) n) = (g (iter g a n))"
apply(induct n)
apply(simp)
apply(simp)
done
lemma ex_iter': "(\<exists>n. R (iter g a n)) = (R a | (\<exists>n. R (iter g (g a) n)))"
apply(rule)
apply(erule exE)
apply(case_tac n)
apply(simp)
apply(rule disjI2)
apply(rule_tac x=nat in exI)
apply(simp add: iter)
apply(erule disjE)
apply(rule_tac x=0 in exI, simp)
apply(erule exE) apply(rule_tac x="Suc n" in exI)
apply(simp add: iter)
done
\<comment> \<open>version suitable for computation\<close>
lemma ex_iter: "(\<exists>n. R (iter g a n)) = (if R a then True else (\<exists>n. R (iter g (g a) n)))"
apply (rule trans)
apply (rule ex_iter')
apply (force )
done
definition
f :: "nseq list => nat => nseq list" where
"f s n = iter (% x. flatten (map subs x)) s n"
lemma f_upwards: "f s n = [] ==> f s (n+m) = []"
apply(induct_tac m) apply(simp)
apply(simp add: f_def)
done
lemma flatten_append: "flatten (xs@ys) = ((flatten xs) @ (flatten ys))"
apply(induct xs) by auto
lemma set_flatten: "set (flatten xs) = Union (set ` set xs)"
apply(induct xs) apply(simp)
apply(simp)
done
lemma f: "\<forall>x. ((n,x) \<in> deriv s) = (x \<in> set (f [s] n))"
apply(induct n)
apply(simp) apply(simp add: f_def)
apply(rule)
apply(rule)
apply(drule_tac deriv_downwards)
apply(elim exE conjE)
apply(drule_tac x=y in spec)
apply(simp)
apply(drule_tac list_decomp) apply(elim exE conjE)
apply(simp add: flatten_append f_def Let_def)
apply(simp add: f_def)
apply(simp add: set_flatten)
apply(erule bexE)
apply(drule_tac x=xa in spec)
apply(rule step) apply(auto)
done
lemma deriv_f: "deriv s = (\<Union>x. set (map (Pair x) (f [s] x)))"
using f by (auto simp add: set_eq_iff)
lemma finite_f: "finite (set (f s x))"
by (fact finite_set)
lemma finite_deriv: "finite (deriv s) = (\<exists>m. f [s] m = [])"
apply(rule)
apply(subgoal_tac "finite (fst ` (deriv s))") prefer 2 apply(simp)
apply(frule_tac max_exists) apply(erule impE) apply(simp) apply(subgoal_tac "(0,s) \<in> deriv s") apply(force) apply(simp)
apply(elim exE conjE)
apply(rule_tac x="Suc x" in exI)
apply(simp)
apply(rule ccontr) apply(case_tac "f [s] (Suc x)") apply(simp)
apply(subgoal_tac "(Suc x, a) \<in> deriv s") apply(force)
apply(simp add: f)
apply(erule exE)
apply(subgoal_tac "\<forall>y. f [s] (m+y) = []")
prefer 2 apply(rule) apply(rule f_upwards) apply(assumption)
apply(simp add: deriv_f)
apply(subgoal_tac "(UNIV::nat set) = {y. y < m} Un {y. m \<le> y}")
prefer 2 apply force
apply(erule_tac t="UNIV::nat set" in ssubst)
apply(simp)
apply(subgoal_tac "(UN x:Collect ((\<le>) m). Pair x ` set (f [s] x)) = (UN x:Collect ((\<le>) m). {})") apply(simp only:) apply(force)
apply(rule SUP_cong) apply(force) apply(drule_tac x="x-m" in spec) apply(force)
done
lemma ex_iter_fSucn: "(\<exists>m. iter (% x. flat (map subs x)) l m = []) = (if l = [] then True else (\<exists>n. (iter (% x. flat (map subs x)) ((% x. flat (map subs x)) l) n) = []))"
using ex_iter[of "% x. x = []", of "(% x. flat (map subs x))" l ] apply(force) done
definition prove' :: "nseq list => bool" where
"prove' s = (\<exists>m. iter (% x. flatten (map subs x)) s m = [])"
lemma prove': "prove' l = (if l = [] then True else prove' ((% x. flatten (map subs x)) l))"
apply(simp only: prove'_def)
apply(rule ex_iter_fSucn)
done
\<comment> \<open>this is the main claim for efficiency- we have a tail recursive implementation via this lemma\<close>
definition prove :: "nseq => bool" where "prove s = prove' ([s])"
lemma finite_deriv_prove: "finite (deriv s) = prove s"
by (simp add: finite_deriv prove_def prove'_def f_def)
subsection "Computation"
\<comment> \<open>a sample formula to prove\<close>
lemma "(\<exists>x. A x | B x) --> ( (\<exists>x. B x) | (\<exists>x. A x))" by blast
\<comment> \<open>convert to our form\<close>
lemma "((\<exists>x. A x | B x) --> ( (\<exists>x. B x) | (\<exists>x. A x)))
= ( (\<forall>x. ~ A x & ~ B x) | ( (\<exists>x. B x) | (\<exists>x. A x)))" by blast
definition my_f :: "form" where
"my_f = FDisj
(FAll (FConj (NAtom 0 [0]) (NAtom 1 [0])))
(FDisj (FEx (PAtom 1 [0])) (FEx (PAtom 0 [0])))"
\<comment> \<open>we compute by rewriting\<close>
lemma membership_simps:
"x \<in> set [] \<longleftrightarrow> False"
"x \<in> set (y # ys) \<longleftrightarrow> x = y \<or> x \<in> set ys"
by simp_all
lemmas ss = list.inject if_True if_False flatten.simps list.map
bump_def sfv_def filter.simps is_axiom.simps fst_conv snd_conv
form.simps collect_disj inc_def finst_def ns_of_s_def s_of_ns_def
Let_def newvar_def subs.simps split_beta append_Nil append_Cons
subst.simps nat.simps fv.simps maxvar.simps preSuc.simps simp_thms
membership_simps
lemmas prove'_Nil = prove' [of "[]", simplified]
lemmas prove'_Cons = prove' [of "x#l", simplified] for x l
lemma search: "finite (deriv [(0,my_f)])"
apply(simp add: my_f_def finite_deriv_prove prove_def)
apply(simp only: prove'_Nil prove'_Cons ss)
done
abbreviation Sprove :: "form list \<Rightarrow> bool" where "Sprove \<equiv> prove o ns_of_s"
abbreviation check :: "form \<Rightarrow> bool" where "check formula \<equiv> Sprove [formula]"
abbreviation valid :: "form \<Rightarrow> bool" where "valid formula \<equiv> Svalid [formula]"
theorem "check = valid" using soundness completeness finite_deriv_prove by auto
ML \<open>
fun max x y = if x > y then x else y;
fun flatten [] = []
| flatten (a::list) = a @ (flatten list);
type pred = int;
type var = int;
datatype form =
PAtom of pred * (var list)
| NAtom of pred * (var list)
| FConj of form * form
| FDisj of form * form
| FAll of form
| FEx of form;
fun preSuc [] = []
| preSuc (a::list) = if a = 0 then preSuc list else (a-1)::(preSuc list);
fun fv (PAtom (_,vs)) = vs
| fv (NAtom (_,vs)) = vs
| fv (FConj (f,g)) = (fv f) @ (fv g)
| fv (FDisj (f,g)) = (fv f) @ (fv g)
| fv (FAll f) = preSuc (fv f)
| fv (FEx f) = preSuc (fv f);
fun subst r (PAtom (p,vs)) = PAtom (p,map r vs)
| subst r (NAtom (p,vs)) = NAtom (p,map r vs)
| subst r (FConj (f,g)) = FConj (subst r f,subst r g)
| subst r (FDisj (f,g)) = FDisj (subst r f,subst r g)
| subst r (FAll f) = FAll (subst (fn 0 => 0 | v => (r (v-1))+1) f)
| subst r (FEx f) = FEx (subst (fn 0 => 0 | v => (r (v-1))+1) f);
fun finst body w = subst (fn 0 => w | v => v-1) body;
fun s_of_ns ns = map (fn (_,y) => y) ns;
fun ns_of_s s = map (fn y => (0,y)) s;
fun sfv s = flatten (map fv s);
fun maxvar [] = 0
| maxvar (a::list) = max a (maxvar list);
fun newvar vs = if vs = [] then 0 else (maxvar vs)+1;
fun test [] _ = false
| test ((_,y)::list) z = if y = z then true else test list z;
fun subs [] = [[]]
| subs (x::xs) = let val (n,f') = x in case f' of
PAtom (p,vs) => if test xs (NAtom (p,vs)) then [] else [xs @ [(0,PAtom (p,vs))]]
| NAtom (p,vs) => if test xs (PAtom (p,vs)) then [] else [xs @ [(0,NAtom (p,vs))]]
| FConj (f,g) => [xs @ [(0,f)],xs @ [(0,g)]]
| FDisj (f,g) => [xs @ [(0,f),(0,g)]]
| FAll f => [xs @ [(0,finst f (newvar (sfv (s_of_ns (x::xs)))))]]
| FEx f => [xs @ [(0,finst f n),(n+1,f')]]
end;
fun step s = flatten (map subs s);
fun prove' s = if s = [] then true else prove' (step s);
fun prove s = prove' [s];
fun check f = (prove o ns_of_s) [f];
val my_f = FDisj (
(FAll (FConj ((NAtom (0,[0])), (NAtom (1,[0])))),
(FDisj ((FEx ((PAtom (1,[0])))), (FEx (PAtom (0,[0])))))));
check my_f;
\<close>
end

File Metadata

Mime Type
application/octet-stream
Expires
Mon, Apr 22, 8:38 AM (1 d, 23 h)
Storage Engine
chunks
Storage Format
Chunks
Storage Handle
B4dwdOLUOzEi
Default Alt Text
(5 MB)

Event Timeline