diff --git a/thys/IFC_Tracking/IFC.thy b/thys/IFC_Tracking/IFC.thy new file mode 100644 --- /dev/null +++ b/thys/IFC_Tracking/IFC.thy @@ -0,0 +1,4580 @@ +section \Definitions\ + +text \ +This section contains all necessary definitions of this development. Section~\ref{sec:pm} contains +the structural definition of our program model which includes the security specification as well +as abstractions of control flow and data. Executions of our program model are defined in +section~\ref{sec:ex}. Additional well-formedness properties are defined in section~\ref{sec:wf}. +Our security property is defined in section~\ref{sec:sec}. Our characterisation of how information +is propagated by executions of our program model is defined in section~\ref{sec:char-cp}, for which +the correctness result can be found in section~\ref{sec:cor-cp}. Section~\ref{sec:char-scp} contains +an additional approximation of this characterisation whose correctness result can be found in +section~\ref{sec:cor-scp}. +\ + + +theory IFC +imports Main +begin + +subsection \Program Model\ +text_raw \\label{sec:pm}\ + +text \Our program model contains all necessary components for the remaining development and consists of:\ + +record ('n, 'var, 'val, 'obs) ifc_problem = +\ \A set of nodes representing program locations:\ + nodes :: \'n set\ +\ \An initial node where all executions start:\ + entry :: \'n\ +\ \A final node where executions can terminate:\ + return :: \'n\ +\ \An abstraction of control flow in the form of an edge relation:\ + edges :: \('n \ 'n) set\ +\ \An abstraction of variables written at program locations:\ + writes :: \'n \ 'var set\ +\ \An abstraction of variables read at program locations:\ + reads :: \'n \ 'var set\ +\ \A set of variables containing the confidential information in the initial state:\ + hvars :: \'var set\ +\ \A step function on location state pairs:\ + step :: \('n \ ('var \ 'val)) \ ('n \ ('var \ 'val))\ +\ \An attacker model producing observations based on the reached state at certain locations:\ + att :: \'n \ (('var \ 'val) \ 'obs)\ + +text \We fix a program in the following in order to define the central concepts. +The necessary well-formedness assumptions will be made in section~\ref{sec:wf}.\ +locale IFC_def = +fixes prob :: \('n, 'var, 'val, 'obs) ifc_problem\ +begin + +text \Some short hands to the components of the program which we will utilise exclusively in the following.\ +definition nodes where \nodes = ifc_problem.nodes prob\ +definition entry where \entry = ifc_problem.entry prob\ +definition return where \return = ifc_problem.return prob\ +definition edges where \edges = ifc_problem.edges prob\ +definition writes where \writes = ifc_problem.writes prob\ +definition reads where \reads = ifc_problem.reads prob\ +definition hvars where \hvars = ifc_problem.hvars prob\ +definition step where \step = ifc_problem.step prob\ +definition att where \att = ifc_problem.att prob\ + +text \The components of the step function for convenience.\ +definition suc where \suc n \ = fst (step (n, \))\ +definition sem where \sem n \ = snd (step (n, \))\ + +lemma step_suc_sem: \step (n,\) = (suc n \, sem n \)\ unfolding suc_def sem_def by auto + + +subsubsection \Executions\ +text \\label{sec:ex}\ +text \In order to define what it means for a program to be well-formed, we first require concepts +of executions and program paths.\ + +text \The sequence of nodes visited by the execution corresponding to an input state.\ +definition path where +\path \ k= fst ((step^^k) (entry,\))\ + +text \The sequence of states visited by the execution corresponding to an input state.\ +definition kth_state ( \_\<^bsup>_\<^esup>\ [111,111] 110) where +\\\<^bsup>k\<^esup> = snd ((step^^k) (entry,\))\ + +text \A predicate asserting that a sequence of nodes is a valid program path according to the +control flow graph.\ + +definition is_path where +\is_path \ = (\ n. (\ n, \ (Suc n)) \ edges)\ +end + +subsubsection \Well-formed Programs\ +text_raw \\label{sec:wf}\ + +text \The following assumptions define our notion of valid programs.\ +locale IFC = IFC_def \prob\ for prob:: \('n, 'var, 'val, 'out) ifc_problem\ + +assumes ret_is_node[simp,intro]: \return \ nodes\ +and entry_is_node[simp,intro]: \entry \ nodes\ +and writes: \\ v n. (\\. \ v \ sem n \ v) \ v \ writes n\ +and writes_return: \writes return = {}\ +and uses_writes: \\ n \ \'. (\ v \ reads n. \ v = \' v) \ \ v \ writes n. sem n \ v = sem n \' v\ +and uses_suc: \\ n \ \'. (\ v \ reads n. \ v = \' v) \ suc n \ = suc n \'\ +and uses_att: \\ n f \ \'. att n = Some f \ (\ v \ reads n. \ v = \' v) \ f \ = f \'\ +and edges_complete[intro,simp]: \\m \. m \ nodes \ (m,suc m \) \ edges\ +and edges_return : \\x. (return,x) \ edges \ x = return \ +and edges_nodes: \edges \ nodes \ nodes\ +and reaching_ret: \\ x. x \ nodes \ \ \ n. is_path \ \ \ 0 = x \ \ n = return\ + + +subsection \Security\ +text_raw \\label{sec:sec}\ + +text \We define our notion of security, which corresponds to what Bohannon et al.~\cite{Bohannon:2009:RN:1653662.1653673} +refer to as indistinguishable security. In order to do so we require notions of observations made +by the attacker, termination and equivalence of input states.\ + +context IFC_def +begin + +subsubsection \Observations\ +text_raw \\label{sec:obs}\ + +text \The observation made at a given index within an execution.\ +definition obsp where +\obsp \ k = (case att(path \ k) of Some f \ Some (f (\\<^bsup>k\<^esup>)) | None \ None)\ + +text \The indices within a path where an observation is made.\ +definition obs_ids :: \(nat \ 'n) \ nat set\ where +\obs_ids \ = {k. att (\ k) \ None}\ + +text \A predicate relating an observable index to the number of observations made before.\ +definition is_kth_obs :: \(nat \ 'n) \ nat \ nat \ bool\where +\is_kth_obs \ k i = (card (obs_ids \ \ {.. att (\ i) \ None)\ + +text \The final sequence of observations made for an execution.\ +definition obs where +\obs \ k = (if (\i. is_kth_obs (path \) k i) then obsp \ (THE i. is_kth_obs (path \) k i) else None)\ + +text \Comparability of observations.\ +definition obs_prefix :: \(nat \ 'obs option) \ (nat \ 'obs option) \ bool\ (infix \\\ 50) where +\a \ b \ \ i. a i \ None \ a i = b i\ + +definition obs_comp (infix \\\ 50) where +\a \ b \ a \ b \ b \ a\ + +subsubsection \Low equivalence of input states\ + +definition restrict (infix \\\ 100 ) where +\f\U = (\ n. if n \ U then f n else undefined)\ + +text \Two input states are low equivalent if they coincide on the non high variables.\ +definition loweq (infix \=\<^sub>L\ 50) +where \\ =\<^sub>L \' = (\\(-hvars) = \'\(-hvars))\ + +subsubsection \Termination\ + +text \An execution terminates iff it reaches the terminal node at any point.\ +definition terminates where +\terminates \ \ \ i. path \ i = return\ + + +subsubsection \Security Property\ +text \The fixed program is secure if and only if for all pairs of low equivalent inputs the observation +sequences are comparable and if the execution for an input state terminates then the observation sequence +is not missing any observations.\ + +definition secure where +\secure \ \ \ \'. \ =\<^sub>L \' \ (obs \ \ obs \' \ (terminates \ \ obs \' \ obs \))\ + + + +subsection \Characterisation of Information Flows\ +text \We now define our characterisation of information flows which tracks data and control dependencies +within executions. To do so we first require some additional concepts.\ + +subsubsection \Post Dominance\ +text \We utilise the post dominance relation in order to define control dependence.\ + +text \The basic post dominance relation.\ +definition is_pd (infix \pd\\ 50) where +\y pd\ x \ x \ nodes \ (\ \ n. is_path \ \ \ (0::nat) = x \ \ n = return \ (\k\n. \ k = y))\ + +text \The immediate post dominance relation.\ +definition is_ipd (infix \ipd\\ 50)where +\y ipd\ x \ x \ y \ y pd\ x \ (\ z. z\x \ z pd\ x \ z pd\ y)\ + +definition ipd where +\ipd x = (THE y. y ipd\ x)\ + +text \The post dominance tree.\ +definition pdt where +\pdt = {(x,y). x\y \ y pd\ x}\ + + +subsubsection \Control Dependence\ + +text \An index on an execution path is control dependent upon another if the path does not visit +the immediate post domiator of the node reached by the smaller index.\ +definition is_cdi (\_ cd\<^bsup>_\<^esup>\ _\ [51,51,51]50) where +\i cd\<^bsup>\\<^esup>\ k \ is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ j \ ipd (\ k))\ + +text \The largest control dependency of an index is the immediate control dependency.\ +definition is_icdi (\_ icd\<^bsup>_\<^esup>\ _\ [51,51,51]50) where +\n icd\<^bsup>\\<^esup>\ n' \ is_path \ \ n cd\<^bsup>\\<^esup>\ n' \ (\ m \ {n'<.. n cd\<^bsup>\\<^esup>\ m)\ + +text \For the definition of the control slice, which we will define next, we require the uniqueness +of the immediate control dependency.\ + +lemma icd_uniq: assumes \m icd\<^bsup>\\<^esup>\ n\ \ m icd\<^bsup>\\<^esup>\ n'\ shows \n = n'\ +proof - + { + fix n n' assume *: \m icd\<^bsup>\\<^esup>\ n\ \ m icd\<^bsup>\\<^esup>\ n'\ \n < n'\ + have \n' using * unfolding is_icdi_def is_cdi_def by auto + hence \\ m cd\<^bsup>\\<^esup>\ n'\ using * unfolding is_icdi_def by auto + with *(2) have \False\ unfolding is_icdi_def by auto + } + thus ?thesis using assms by (metis linorder_neqE_nat) +qed + + +subsubsection \Control Slice\ +text \We utilise the control slice, that is the sequence of nodes visited by the control dependencies +of an index, to match indices between executions.\ + +function cs:: \(nat \ 'n) \ nat \ 'n list\ (\cs\<^bsup>_\<^esup> _\ [51,70] 71) where +\cs\<^bsup>\\<^esup> n = (if (\ m. n icd\<^bsup>\\<^esup>\ m) then (cs \ (THE m. n icd\<^bsup>\\<^esup>\ m))@[\ n] else [\ n])\ +by pat_completeness auto +termination \cs\ proof + show \wf (measure snd)\ by simp + fix \ n + define m where \m == (The (is_icdi n \))\ + assume \Ex (is_icdi n \)\ + hence \n icd\<^bsup>\\<^esup>\ m\ unfolding m_def by (metis (full_types) icd_uniq theI') + hence \m < n\ unfolding is_icdi_def is_cdi_def by simp + thus \((\, The (is_icdi n \)), \, n) \ measure snd\ by (metis in_measure m_def snd_conv) +qed + +inductive cs_less (infix \\\ 50) where +\length xs < length ys \ take (length xs) ys = xs \ xs \ ys\ + +definition cs_select (infix \\\ 50) where +\\\xs = (THE k. cs\<^bsup>\\<^esup> k = xs)\ + + +subsubsection \Data Dependence\ + +text \Data dependence is defined straight forward. An index is data dependent upon another, +if the index reads a variable written by the earlier index and the variable in question has not +been written by any index in between.\ +definition is_ddi (\_ dd\<^bsup>_,_\<^esup>\ _\ [51,51,51,51] 50) where +\n dd\<^bsup>\,v\<^esup>\ m \ is_path \ \ m < n \ v \ reads (\ n) \ (writes (\ m)) \ (\ l \ {m<.. writes (\ l))\ + + + +subsubsection \Characterisation via Critical Paths\ +text_raw \\label{sec:char-cp}\ +text \With the above we define the set of critical paths which as we will prove characterise the matching +points in executions where diverging data is read.\ + +inductive_set cp where + +\ \Any pair of low equivalent input states and indices where a diverging high variable is first +read is critical.\ + +\\\ =\<^sub>L \'; + cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'; + h \ reads(path \ n); + (\\<^bsup>n\<^esup>) h \ (\'\<^bsup>n'\<^esup>) h; + \ kwrites(path \ k); + \ k'writes(path \' k') + \ \ ((\,n),(\',n')) \ cp\ | + +\ \If from a pair of critical indices in two executions there exist data dependencies from both +indices to a pair of matching indices where the variable diverges, the later pair of indices is critical.\ + +\\((\,k),(\',k')) \ cp; + n dd\<^bsup>path \,v\<^esup>\ k; + n' dd\<^bsup>path \',v\<^esup>\ k'; + cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'; + (\\<^bsup>n\<^esup>) v \ (\'\<^bsup>n'\<^esup>) v + \ \ ((\,n),(\',n')) \ cp\ | + +\ \If from a pair of critical indices the executions take different branches and one of the critical +indices is a control dependency of an index that is data dependency of a matched index where diverging +data is read and the variable in question is not written by the other execution after the executions +first reached matching indices again, then the later matching pair of indices is critical.\ + +\\((\,k),(\',k')) \ cp; + n dd\<^bsup>path \,v\<^esup>\ l; + l cd\<^bsup>path \\<^esup>\ k; + cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'; + path \ (Suc k) \ path \' (Suc k'); + (\\<^bsup>n\<^esup>) v \ (\'\<^bsup>n'\<^esup>) v; + \j'\{(LEAST i'. k' < i' \ (\i. cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i'))..writes (path \' j') + \ \ ((\,n),(\',n')) \ cp\ | + +\ \The relation is symmetric.\ + +\\((\,k),(\',k')) \ cp\ \ ((\',k'),(\,k)) \ cp\ + + +text \Based on the set of critical paths, the critical observable paths are those that either directly +reach observable nodes or are diverging control dependencies of an observable index.\ + +inductive_set cop where +\\((\,n),(\',n')) \ cp; + path \ n \ dom att + \ \ ((\,n),(\',n')) \ cop\ | + +\\((\,k),(\',k')) \ cp; + n cd\<^bsup>path \\<^esup>\ k; + path \ (Suc k) \ path \' (Suc k'); + path \ n \ dom att + \ \ ((\,n),(\',k')) \ cop\ + + + +subsubsection \Approximation via Single Critical Paths\ +text_raw \\label{sec:char-scp}\ + +text \For applications we also define a single execution approximation.\ + +definition is_dcdi_via (\_ dcd\<^bsup>_,_\<^esup>\ _ via _ _\ [51,51,51,51,51,51] 50) where +\n dcd\<^bsup>\,v\<^esup>\ m via \' m' = (is_path \ \ m < n \ (\ l' n'. cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m' \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n' \ n' dd\<^bsup>\',v\<^esup>\ l' \ l' cd\<^bsup>\'\<^esup>\ m') \ (\ l \ {m.. writes(\ l)))\ + +inductive_set scp where +\\h \ hvars; h \ reads (path \ n); (\ k writes(path \ k))\ \ (path \,n) \ scp\ | +\\(\,m) \ scp; n cd\<^bsup>\\<^esup>\ m\ \ (\,n) \ scp\| +\\(\,m) \ scp; n dd\<^bsup>\,v\<^esup>\ m\ \ (\,n) \ scp\| +\\(\,m) \ scp; (\',m') \ scp; n dcd\<^bsup>\,v\<^esup>\ m via \' m'\ \ (\,n) \ scp\ + +inductive_set scop where +\\(\,n) \ scp; \ n \ dom att\ \ (\,n) \ scop\ + + + +subsubsection \Further Definitions\ +text \The following concepts are utilised by the proofs.\ + +inductive contradicts (infix \\\ 50) where +\\cs\<^bsup>\'\<^esup> k' \ cs\<^bsup>\\<^esup> k ; \ = path \; \' = path \' ; \ (Suc (\\cs\<^bsup>\'\<^esup> k')) \ \' (Suc k')\ \ (\', k') \ (\, k)\| +\\cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k ; \ = path \; \' = path \' ; \\<^bsup>k\<^esup> \ (reads (\ k)) \ \'\<^bsup>k'\<^esup> \ (reads (\ k))\ \ (\',k') \ (\,k)\ + +definition path_shift (infixl \\\ 51) where +[simp]: \\\m = (\ n. \ (m+n))\ + +definition path_append :: \(nat \ 'n) \ nat \ (nat \ 'n) \ (nat \ 'n)\ (\_ @\<^bsup>_\<^esup> _\ [0,0,999] 51) where +[simp]: \\ @\<^bsup>m\<^esup> \' = (\n.(if n \ m then \ n else \' (n-m)))\ + +definition eq_up_to :: \(nat \ 'n) \ nat \ (nat \ 'n) \ bool\ (\_ =\<^bsub>_\<^esub> _\ [55,55,55] 50) where +\\ =\<^bsub>k\<^esub> \' = (\ i \ k. \ i = \' i)\ + +end (* End of locale IFC_def *) + + + + +section \Proofs\ +text_raw \\label{sec:proofs}\ + +subsection \Miscellaneous Facts\ + +lemma option_neq_cases: assumes \x \ y\ obtains (none1) a where \x = None\ \y = Some a\ | (none2) a where \x = Some a\ \y = None\ | (some) a b where \x = Some a\ \y = Some b\ \a \ b\ using assms by fastforce + +lemmas nat_sym_cases[case_names less sym eq] = linorder_less_wlog + +lemma mod_bound_instance: assumes \j < (i::nat)\ obtains j' where \k < j'\ and \j' mod i = j\ proof - + have \k < Suc k * i + j\ using assms less_imp_Suc_add by fastforce + moreover + have \(Suc k * i + j) mod i = j\ by (metis assms mod_less mod_mult_self3) + ultimately show \thesis\ using that by auto +qed + +lemma list_neq_prefix_cases: assumes \ls \ ls'\ and \ls \ Nil\ and \ls' \ Nil\ + obtains (diverge) xs x x' ys ys' where \ls = xs@[x]@ys\ \ls' = xs@[x']@ys'\ \x \ x'\ | + (prefix1) xs where \ls = ls'@xs\ and \xs \ Nil\ | + (prefix2) xs where \ls@xs = ls'\ and \xs \ Nil\ +using assms proof (induct \length ls\ arbitrary: \ls\ \ls'\ rule: less_induct) + case (less ls ls') + obtain z zs z' zs' where + lz: \ls = z#zs\ \ls' = z'#zs'\ by (metis list.exhaust less(6,7)) + show \?case\ proof cases + assume zz: \z = z'\ + hence zsz: \zs \ zs'\ using less(5) lz by auto + have lenz: \length zs < length ls\ using lz by auto + show \?case\ proof(cases \zs = Nil\) + assume zs: \zs = Nil\ + hence \zs' \ Nil\ using zsz by auto + moreover + have \ls@zs' = ls'\ using zs lz zz by auto + ultimately + show \thesis\ using less(4) by blast + next + assume zs: \zs \ Nil\ + show \thesis\ proof (cases \zs' = Nil\) + assume \zs' = Nil\ + hence \ls = ls'@zs\ using lz zz by auto + thus \thesis\ using zs less(3) by blast + next + assume zs': \zs' \ Nil\ + { fix xs x ys x' ys' + assume \zs = xs @ [x] @ ys\ \zs' = xs @ [x'] @ ys'\ and xx: \x \ x'\ + hence \ls = (z#xs) @ [x] @ ys\ \ls' = (z#xs) @ [x'] @ ys'\ using lz zz by auto + hence \thesis\ using less(2) xx by blast + } note * = this + { fix xs + assume \zs = zs' @ xs\ and xs: \xs \ []\ + hence \ls = ls' @ xs\ using lz zz by auto + hence \thesis\ using xs less(3) by blast + } note ** = this + { fix xs + assume \zs@xs = zs'\ and xs: \xs \ []\ + hence \ls@xs = ls'\ using lz zz by auto + hence \thesis\ using xs less(4) by blast + } note *** = this + have \(\xs x ys x' ys'. zs = xs @ [x] @ ys \ zs' = xs @ [x'] @ ys' \ x \ x' \ thesis) \ + (\xs. zs = zs' @ xs \ xs \ [] \ thesis) \ + (\xs. zs @ xs = zs' \ xs \ [] \ thesis) \ thesis\ + using less(1)[OF lenz _ _ _ zsz zs zs' ] . + thus \thesis\ using * ** *** by blast + qed + qed + next + assume \z \ z'\ + moreover + have \ls = []@[z]@zs\ \ls' = []@[z']@zs'\ using lz by auto + ultimately show \thesis\ using less(2) by blast + qed +qed + +lemma three_cases: assumes \A \ B \ C\ obtains \A\ | \B\ | \C\ using assms by auto + +lemma insort_greater: \\ x \ set ls. x < y \ insort y ls = ls@[y]\ by (induction \ls\,auto) + +lemma insort_append_first: assumes \\ y \ set ys. x \ y\ shows \insort x (xs@ys) = insort x xs @ ys\ using assms by (induction \xs\,auto,metis insort_is_Cons) + +lemma sorted_list_of_set_append: assumes \finite xs\ \finite ys\ \\ x \ xs. \ y \ ys. x < y\ shows \sorted_list_of_set (xs \ ys) = sorted_list_of_set xs @ (sorted_list_of_set ys)\ +using assms(1,3) proof (induction \xs\) + case empty thus \?case\ by simp +next + case (insert x xs) + hence iv: \sorted_list_of_set (xs \ ys) = sorted_list_of_set xs @ sorted_list_of_set ys\ by blast + have le: \\ y \ set (sorted_list_of_set ys). x < y\ using insert(4) assms(2) sorted_list_of_set by auto + have \sorted_list_of_set (insert x xs \ ys) = sorted_list_of_set (insert x (xs \ ys))\ by auto + also + have \\ = insort x (sorted_list_of_set (xs \ ys))\ by (metis Un_iff assms(2) finite_Un insert.hyps(1) insert.hyps(2) insert.prems insertI1 less_irrefl sorted_list_of_set.insert) + also + have \\ = insort x (sorted_list_of_set xs @ sorted_list_of_set ys)\ using iv by simp + also + have \\ = insort x (sorted_list_of_set xs) @ sorted_list_of_set ys\ by (metis le insort_append_first less_le_not_le) + also + have \\ = sorted_list_of_set (insert x xs) @ sorted_list_of_set ys\ using sorted_list_of_set_insert[OF insert(1),of \x\] insert(2) by auto + finally + show \?case\ . +qed + +lemma filter_insort: \sorted xs \ filter P (insort x xs) = (if P x then insort x (filter P xs) else filter P xs)\ by (induction \xs\, simp) (metis filter_insort filter_insort_triv map_ident) + +lemma filter_sorted_list_of_set: assumes \finite xs\ shows \filter P (sorted_list_of_set xs) = sorted_list_of_set {x \ xs. P x}\ using assms proof(induction \xs\) + case empty thus \?case\ by simp +next + case (insert x xs) + have *: \set (sorted_list_of_set xs) = xs\ \sorted (sorted_list_of_set xs)\ \distinct (sorted_list_of_set xs)\ by (auto simp add: insert.hyps(1)) + have **: \P x \ {y \ insert x xs. P y} = insert x {y \ xs. P y}\ by auto + have ***: \\ P x \ {y \ insert x xs. P y} = {y \ xs. P y}\ by auto + note filter_insort[OF *(2),of \P\ \x\] sorted_list_of_set_insert[OF insert(1), of \x\] insert(2,3) ** *** + thus \?case\ by (metis (mono_tags) "*"(1) List.finite_set distinct_filter distinct_insort distinct_sorted_list_of_set set_filter sorted_list_of_set.insert) +qed + +lemma unbounded_nat_set_infinite: assumes \\ (i::nat). \ j\i. j \ A\ shows \\ finite A\ using assms +by (metis finite_nat_set_iff_bounded_le not_less_eq_eq) + +lemma infinite_ascending: assumes nf: \\ finite (A::nat set)\ obtains f where \range f = A\ \\ i. f i < f (Suc i)\ proof + let \?f\ = \\ i. (LEAST a. a \ A \ card (A \ {.. + { fix i + obtain a where \a \ A\ \card (A \ {.. + proof (induction \i\ arbitrary: \thesis\) + case 0 + let \?a0\ = \(LEAST a. a \ A)\ + have \?a0 \ A\ by (metis LeastI empty_iff finite.emptyI nf set_eq_iff) + moreover + have \\b. b \ A \ ?a0 \ b\ by (metis Least_le) + hence \card (A \ {.. by force + ultimately + show \?case\ using 0 by blast + next + case (Suc i) + obtain a where aa: \a \ A\ and card: \card (A \ {.. using Suc.IH by metis + have nf': \~ finite (A - {..a})\ using nf by auto + let \?b\ = \LEAST b. b \ A - {..a}\ + have bin: \?b \ A-{..a}\ by (metis LeastI empty_iff finite.emptyI nf' set_eq_iff) + have le: \\c. c \ A-{..a} \ ?b \ c\ by (metis Least_le) + have ab: \a < ?b\ using bin by auto + have \\ c. c \ A \ c < ?b \ c \ a\ using le by force + hence \A \ {.. {.. using bin ab aa by force + hence \card (A \{.. using card by auto + thus \?case\ using Suc.prems bin by auto + qed + note \\ thesis. ((\a. a \ A \ card (A \ {.. thesis) \ thesis)\ + } + note ex = this + + { + fix i + obtain a where a: \a \ A \ card (A \{.. using ex by blast + have ina: \?f i \ A\ and card: \card (A \{.. using LeastI[of \\ a. a \ A \ card (A \{.. \a\, OF a] by auto + obtain b where b: \b \ A \ card (A \{.. using ex by blast + have inab: \?f (Suc i) \ A\ and cardb: \card (A \{.. using LeastI[of \\ a. a \ A \ card (A \{.. \b\, OF b] by auto + have \?f i < ?f (Suc i)\ proof (rule ccontr) + assume \\ ?f i < ?f (Suc i)\ + hence \A \{.. A \{.. by auto + moreover have \finite (A \{.. by auto + ultimately have \card(A \{.. card (A \{.. by (metis (erased, lifting) card_mono) + thus \False\ using card cardb by auto + qed + note this ina + } + note b = this + thus \\ i. ?f i < ?f (Suc i)\ by auto + have *: \range ?f \ A\ using b by auto + moreover + { + fix a assume ina: \a \ A\ + let \?i\ = \card (A \ {.. + obtain b where b: \b \ A \ card (A \{.. using ex by blast + have inab: \?f ?i \ A\ and cardb: \card (A \{.. using LeastI[of \\ a. a \ A \ card (A \{.. \b\, OF b] by auto + have le: \?f ?i \ a\ using Least_le[of \\ a. a \ A \ card (A \{.. \a\] ina by auto + have \a = ?f ?i\ proof (rule ccontr) + have fin: \finite (A \ {.. by auto + assume \a \ ?f ?i\ + hence \?f ?i < a\ using le by simp + hence \?f ?i \ A \ {.. using inab by auto + moreover + have \A \ {.. A \ {.. using le by auto + hence \A \ {.. {.. using cardb card_subset_eq[OF fin] by auto + ultimately + show \False\ by auto + qed + hence \a \ range ?f\ by auto + } + hence \A \ range ?f\ by auto + ultimately show \range ?f = A\ by auto +qed + +lemma mono_ge_id: \\ i. f i < f (Suc i) \ i \ f i\ + apply (induction \i\,auto) + by (metis not_le not_less_eq_eq order_trans) + +lemma insort_map_mono: assumes mono: \\ n m. n < m \ f n < f m\ shows \map f (insort n ns) = insort (f n) (map f ns)\ + apply (induction \ns\) + apply auto + apply (metis not_less not_less_iff_gr_or_eq mono) + apply (metis antisym_conv1 less_imp_le mono) + apply (metis mono not_less) + by (metis mono not_less) + +lemma sorted_list_of_set_map_mono: assumes mono: \\ n m. n < m \ f n < f m\ and fin: \finite A\ +shows \map f (sorted_list_of_set A) = sorted_list_of_set (f`A)\ +using fin proof (induction) + case empty thus \?case\ by simp +next + case (insert x A) + have [simp]:\sorted_list_of_set (insert x A) = insort x (sorted_list_of_set A)\ using insert sorted_list_of_set.insert by simp + have \f ` insert x A = insert (f x) (f ` A)\ by auto + moreover + have \f x \ f`A\ apply (rule ccontr) using insert(2) mono apply auto by (metis insert.hyps(2) mono neq_iff) + ultimately + have \sorted_list_of_set (f ` insert x A) = insort (f x) (sorted_list_of_set (f`A))\ using insert(1) sorted_list_of_set.insert by simp + also + have \\ = insort (f x) (map f (sorted_list_of_set A))\ using insert.IH by auto + also have \\ = map f (insort x (sorted_list_of_set A))\ using insort_map_mono[OF mono] by auto + finally + show \map f (sorted_list_of_set (insert x A)) = sorted_list_of_set (f ` insert x A)\ by simp +qed + +lemma GreatestIB: +fixes n :: \nat\ and P +assumes a:\\k\n. P k\ +shows GreatestBI: \P (GREATEST k. k\n \ P k)\ and GreatestB: \(GREATEST k. k\n \ P k) \ n\ +proof - + show \P (GREATEST k. k\n \ P k)\ using GreatestI_ex_nat[OF assms] by auto + show \(GREATEST k. k\n \ P k) \ n\ using GreatestI_ex_nat[OF assms] by auto +qed + +lemma GreatestB_le: +fixes n :: \nat\ +assumes \x\n\ and \P x\ +shows \x \ (GREATEST k. k\n \ P k)\ +proof - + have *: \\ y. y\n \ P y \ y by auto + then show \x \ (GREATEST k. k\n \ P k)\ using assms by (blast intro: Greatest_le_nat) +qed + +lemma LeastBI_ex: assumes \\k \ n. P k\ shows \P (LEAST k::'c::wellorder. P k)\ and \(LEAST k. P k) \ n\ +proof - + from assms guess k .. + hence k: \k \ n\ \P k\ by auto + thus \P (LEAST k. P k)\ using LeastI[of \P\ \k\] by simp + show \(LEAST k. P k) \ n\ using Least_le[of \P\ \k\] k by auto +qed + +lemma allB_atLeastLessThan_lower: assumes \(i::nat) \ j\ \\ x\{i.. shows \\ x\{j.. proof + fix x assume \x\{j.. hence \x\{i.. using assms(1) by simp + thus \P x\ using assms(2) by auto +qed + + +subsection \Facts about Paths\ + +context IFC +begin + +lemma path0: \path \ 0 = entry\ unfolding path_def by auto + +lemma path_in_nodes[intro]: \path \ k \ nodes\ proof (induction \k\) + case (Suc k) + hence \\ \'. (path \ k, suc (path \ k) \') \ edges\ by auto + hence \(path \ k, path \ (Suc k)) \ edges\ unfolding path_def + by (metis suc_def comp_apply funpow.simps(2) prod.collapse) + thus \?case\ using edges_nodes by force +qed (auto simp add: path_def) + +lemma path_is_path[simp]: \is_path (path \)\ unfolding is_path_def path_def using step_suc_sem apply auto +by (metis path_def suc_def edges_complete path_in_nodes prod.collapse) + +lemma term_path_stable: assumes \is_path \\ \\ i = return\ and le: \i \ j\ shows \\ j = return\ +using le proof (induction \j\) + case (Suc j) + show \?case\ proof cases + assume \i\j\ + hence \\ j = return\ using Suc by simp + hence \(return, \ (Suc j)) \ edges\ using assms(1) unfolding is_path_def by metis + thus \\ (Suc j) = return\ using edges_return by auto + next + assume \\ i \ j\ + hence \Suc j = i\ using Suc by auto + thus \?thesis\ using assms(2) by auto + qed +next + case 0 thus \?case\ using assms by simp +qed + +lemma path_path_shift: assumes \is_path \\ shows \is_path (\\m)\ +using assms unfolding is_path_def by simp + +lemma path_cons: assumes \is_path \\ \is_path \'\ \\ m = \' 0\ shows \is_path (\ @\<^bsup>m\<^esup> \')\ +unfolding is_path_def proof(rule,cases) + fix n assume \m < n\ thus \((\ @\<^bsup>m\<^esup> \') n, (\ @\<^bsup>m\<^esup> \') (Suc n)) \ edges\ + using assms(2) unfolding is_path_def path_append_def + by (auto,metis Suc_diff_Suc diff_Suc_Suc less_SucI) +next + fix n assume *: \\ m < n\ thus \((\ @\<^bsup>m\<^esup> \') n, (\ @\<^bsup>m\<^esup> \') (Suc n)) \ edges\ proof cases + assume [simp]: \n = m\ + thus \?thesis\ using assms unfolding is_path_def path_append_def by force + next + assume \n \ m\ + hence \Suc n \ m\ \n\ m\ using * by auto + with assms(1) show \?thesis\ unfolding is_path_def by auto + qed +qed + +lemma is_path_loop: assumes \is_path \\ \0 < i\ \\ i = \ 0\ shows \is_path (\ n. \ (n mod i))\ unfolding is_path_def proof (rule,cases) + fix n + assume \0 < Suc n mod i\ + hence \Suc n mod i = Suc (n mod i)\ by (metis mod_Suc neq0_conv) + moreover + have \(\ (n mod i), \ (Suc (n mod i))) \ edges\ using assms(1) unfolding is_path_def by auto + ultimately + show \(\ (n mod i), \ (Suc n mod i)) \ edges\ by simp + next + fix n + assume \\ 0 < Suc n mod i\ + hence \Suc n mod i = 0\ by auto + moreover + hence \n mod i = i - 1\ using assms(2) by (metis Zero_neq_Suc diff_Suc_1 mod_Suc) + ultimately + show \(\(n mod i), \ (Suc n mod i)) \ edges\ using assms(1) unfolding is_path_def by (metis assms(3) mod_Suc) +qed + +lemma path_nodes: \is_path \ \ \ k \ nodes\ unfolding is_path_def using edges_nodes by force + +lemma direct_path_return': assumes \is_path \ \ \\ 0 = x\ \x \ return\ \\ n = return\ +obtains \' n' where \is_path \'\ \\' 0 = x\ \\' n' = return\ \\ i> 0. \' i \ x\ +using assms proof (induction \n\ arbitrary: \\\ rule: less_induct) + case (less n \) + hence ih: \\ n' \'. n' < n \ is_path \' \ \' 0 = x \ \' n' = return \ thesis\ using assms by auto + show \thesis\ proof cases + assume \\ i>0. \ i \ x\ thus \thesis\ using less by auto + next + assume \\ (\ i>0. \ i \ x)\ + then obtain i where \0 \\ i = x\ by auto + hence \(\\i) 0 = x\ by auto + moreover + have \i < n\ using less(3,5,6) \\ i = x\ by (metis linorder_neqE_nat term_path_stable less_imp_le) + hence \(\\i) (n-i) = return\ using less(6) by auto + moreover + have \is_path (\\i)\ using less(3) by (metis path_path_shift) + moreover + have \n - i < n\ using \0 \i < n\ by auto + ultimately show \thesis\ using ih by auto + qed +qed + +lemma direct_path_return: assumes \x \ nodes\ \x \ return\ +obtains \ n where \is_path \\ \\ 0 = x\ \\ n = return\ \\ i> 0. \ i \ x\ +using direct_path_return'[of _ \x\] reaching_ret[OF assms(1)] assms(2) by blast + +lemma path_append_eq_up_to: \(\ @\<^bsup>k\<^esup> \') =\<^bsub>k\<^esub> \\ unfolding eq_up_to_def by auto + +lemma eq_up_to_le: assumes \k \ n\ \\ =\<^bsub>n\<^esub> \'\ shows \\ =\<^bsub>k\<^esub> \'\ using assms unfolding eq_up_to_def by auto + +lemma eq_up_to_refl: shows \\ =\<^bsub>k\<^esub> \\ unfolding eq_up_to_def by auto + +lemma eq_up_to_sym: assumes \\ =\<^bsub>k\<^esub> \'\ shows \\' =\<^bsub>k\<^esub> \\ using assms unfolding eq_up_to_def by auto + +lemma eq_up_to_apply: assumes \\ =\<^bsub>k\<^esub> \'\ \j \ k\ shows \\ j = \' j\ using assms unfolding eq_up_to_def by auto + +lemma path_swap_ret: assumes \is_path \\ obtains \' n where \is_path \'\ \\ =\<^bsub>k\<^esub> \'\ \\' n = return\ +proof - + have nd: \\ k \ nodes\ using assms path_nodes by simp + obtain \' n where *: \is_path \'\ \\' 0 = \ k\ \\' n = return\ using reaching_ret[OF nd] by blast + have \\ =\<^bsub>k\<^esub> (\@\<^bsup>k\<^esup> \')\ by (metis eq_up_to_sym path_append_eq_up_to) + moreover + have \is_path (\@\<^bsup>k\<^esup> \')\ using assms * path_cons by metis + moreover + have \(\@\<^bsup>k\<^esup> \') (k + n) = return\ using * by auto + ultimately + show \thesis\ using that by blast +qed + +lemma path_suc: \path \ (Suc k) = fst (step (path \ k, \\<^bsup>k\<^esup>))\ by (induction \k\, auto simp: path_def kth_state_def) + +lemma kth_state_suc: \\\<^bsup>Suc k\<^esup> = snd (step (path \ k, \\<^bsup>k\<^esup>))\ by (induction \k\, auto simp: path_def kth_state_def) + + +subsection \Facts about Post Dominators\ + +lemma pd_trans: assumes 1: \y pd\ x\ and 2: \z pd\y\ shows \z pd\x\ +proof - + { + fix \ n + assume 3[simp]: \is_path \\ \\ 0 = x\ \\ n = return\ + then obtain k where \\ k = y\ and 7: \k \ n\ using 1 unfolding is_pd_def by blast + then have \(\\k) 0 = y\ and \(\\k) (n-k) = return\ by auto + moreover have \is_path (\\k)\ by(metis 3(1) path_path_shift) + ultimately obtain k' where 8: \(\\k) k' = z\ and \k' \ n-k\ using 2 unfolding is_pd_def by blast + hence \k+k'\n\ and \\ (k+ k') = z\ using 7 by auto + hence \\k\n. \ k = z\ using path_nodes by auto + } + thus \?thesis\ using 1 unfolding is_pd_def by blast +qed + +lemma pd_path: assumes \y pd\ x\ +obtains \ n k where \is_path \\ and \\ 0 = x\ and \\ n = return\ and \\ k = y\ and \k \ n\ +using assms unfolding is_pd_def using reaching_ret[of \x\] by blast + +lemma pd_antisym: assumes xpdy: \x pd\ y\ and ypdx: \y pd\ x\ shows \x = y\ +proof - + obtain \ n where path: \is_path \\ and \0: \\ 0 = x\ and \n: \\ n = return\ using pd_path[OF ypdx] by metis + hence kex: \\k\n. \ k = y\ using ypdx unfolding is_pd_def by auto + obtain k where k: \k = (GREATEST k. k\n \ \ k = y)\ by simp + have \k: \\ k = y\ and kn: \k \ n\ using k kex by (auto intro: GreatestIB) + + have kpath: \is_path (\\k)\ by (metis path_path_shift path) + moreover have k0: \(\\k) 0 = y\ using \k by simp + moreover have kreturn: \(\\k) (n-k) = return\ using kn \n by simp + ultimately have ky': \\k'\(n-k).(\\k) k' = x\ using xpdy unfolding is_pd_def by simp + + obtain k' where k': \k' = (GREATEST k'. k'\(n-k) \ (\\k) k' = x)\ by simp + + with ky' have \k': \(\\k) k' = x\ and kn': \k' \ (n-k)\ by (auto intro: GreatestIB) + have k'path: \is_path (\\k\k')\ using kpath by(metis path_path_shift) + moreover have k'0: \(\\k\k') 0 = x\ using \k' by simp + moreover have k'return: \(\\k\k') (n-k-k') = return\ using kn' kreturn by (metis path_shift_def le_add_diff_inverse) + ultimately have ky'': \\k''\(n-k-k').(\\k\k') k'' = y\ using ypdx unfolding is_pd_def by blast + + obtain k'' where k'': \k''= (GREATEST k''. k''\(n-k-k') \ (\\k\k') k'' = y)\ by simp + with ky'' have \k'': \(\\k\k') k'' = y\ and kn'': \k'' \ (n-k-k')\ by (auto intro: GreatestIB) + + from this(1) have \\ (k + k' + k'') = y\ by (metis path_shift_def add.commute add.left_commute) + moreover + have \k + k' +k'' \ n\ using kn'' kn' kn by simp + ultimately have \k + k' + k''\ k\ using k by(auto simp: GreatestB_le) + hence \k' = 0\ by simp + with k0 \k' show \x = y\ by simp +qed + +lemma pd_refl[simp]: \x \ nodes \ x pd\ x\ unfolding is_pd_def by blast + +lemma pdt_trans_in_pdt: \(x,y) \ pdt\<^sup>+ \ (x,y) \ pdt\ +proof (induction rule: trancl_induct) + case base thus \?case\ by simp +next + case (step y z) show \?case\ unfolding pdt_def proof (simp) + have *: \y pd\ x\ \z pd\ y\ using step unfolding pdt_def by auto + hence [simp]: \z pd\ x\ using pd_trans[where x=\x\ and y=\y\ and z=\z\] by simp + have \x\z\ proof + assume \x = z\ + hence \z pd\ y\ \y pd\ z\ using * by auto + hence \z = y\ using pd_antisym by auto + thus \False\ using step(2) unfolding pdt_def by simp + qed + thus \x \ z \ z pd\ x\ by auto + qed +qed + +lemma pdt_trancl_pdt: \pdt\<^sup>+ = pdt\ using pdt_trans_in_pdt by fast + +lemma trans_pdt: \trans pdt\ by (metis pdt_trancl_pdt trans_trancl) + +definition [simp]: \pdt_inv = pdt\\ + +lemma wf_pdt_inv: \wf (pdt_inv)\ proof (rule ccontr) + assume \\ wf (pdt_inv)\ + then obtain f where \\i. (f (Suc i), f i) \ pdt\\ using wf_iff_no_infinite_down_chain by force + hence *: \\ i. (f i, f (Suc i)) \ pdt\ by simp + have **:\\ i. \ j>i. (f i, f j) \ pdt\ proof(rule,rule,rule) + fix i j assume \i < (j::nat)\ thus \(f i, f j) \ pdt\ proof (induction \j\ rule: less_induct) + case (less k) + show \?case\ proof (cases \Suc i < k\) + case True + hence k:\k-1 < k\ \i < k-1\ and sk: \Suc (k-1) = k\ by auto + show \?thesis\ using less(1)[OF k] *[rule_format,of \k-1\,unfolded sk] trans_pdt[unfolded trans_def] by blast + next + case False + hence \Suc i = k\ using less(2) by auto + then show \?thesis\ using * by auto + qed + qed + qed + hence ***:\\ i. \ j > i. f j pd\ f i\ \\ i. \ j > i. f i \ f j\ unfolding pdt_def by auto + hence ****:\\ i>0. f i pd\ f 0\ by simp + hence \f 0 \ nodes\ using * is_pd_def by fastforce + then obtain \ n where \:\is_path \\ \\ 0 = f 0\ \\ n = return\ using reaching_ret by blast + hence \\ i>0. \ k\n. \ k = f i\ using ***(1) \f 0 \ nodes\ unfolding is_pd_def by blast + hence \f:\\ i. \ k\n. \ k = f i\ using \(2) by (metis le0 not_gr_zero) + have \range f \ \ ` {..n}\ proof(rule subsetI) + fix x assume \x \ range f\ + then obtain i where \x = f i\ by auto + then obtain k where \x = \ k\ \k \ n\ using \f by metis + thus \x \ \ ` {..n}\ by simp + qed + hence f:\finite (range f)\ using finite_surj by auto + hence fi:\\ i. infinite {j. f j = f i}\ using pigeonhole_infinite[OF _ f] by auto + obtain i where \infinite {j. f j = f i}\ using fi .. + thus \False\ + by (metis (mono_tags, lifting) "***"(2) bounded_nat_set_is_finite gt_ex mem_Collect_eq nat_neq_iff) +qed + +lemma return_pd: assumes \x \ nodes\ shows \return pd\ x\ unfolding is_pd_def using assms by blast + +lemma pd_total: assumes xz: \x pd\ z\ and yz: \y pd\ z\ shows \x pd\ y \ y pd\x\ +proof - + obtain \ n where path: \is_path \\ and \0: \\ 0 = z\ and \n: \\ n = return\ using xz reaching_ret unfolding is_pd_def by force + have *: \\ k\n. (\ k = x \ \ k = y)\ (is \\ k\n. ?P k\) using path \0 \n xz yz unfolding is_pd_def by auto + obtain k where k: \k = (LEAST k. \ k = x \ \ k = y)\ by simp + hence kn: \k\n\ and \k: \\ k = x \ \ k = y\ using LeastBI_ex[OF *] by auto + note k_le = Least_le[where P = \?P\] + show \?thesis\ proof (cases) + assume kx: \\ k = x\ + have k_min: \\ k'. \ k' = y \ k \ k'\ using k_le unfolding k by auto + { + fix \' + and n' :: \nat\ + assume path': \is_path \'\ and \'0: \\' 0 = x\ and \'n': \\' n' = return\ + have path'': \is_path (\ @\<^bsup>k\<^esup> \')\ using path_cons[OF path path'] kx \'0 by auto + have \''0: \(\ @\<^bsup>k\<^esup> \') 0 = z\ using \0 by simp + have \''n: \(\ @\<^bsup>k\<^esup> \') (k+n') = return\ using \'n' kx \'0 by auto + obtain k' where k': \k' \ k + n'\ \(\ @\<^bsup>k\<^esup> \') k' = y\ using yz path'' \''0 \''n unfolding is_pd_def by blast + have **: \k \ k'\ proof (rule ccontr) + assume \\ k \ k'\ + hence \k' < k\ by simp + moreover + hence \\ k' = y\ using k' by auto + ultimately + show \False\ using k_min by force + qed + hence \\' (k' - k) = y\ using k' \'0 kx by auto + moreover + have \(k' - k) \ n'\ using k' by auto + ultimately + have \\ k\ n'. \' k = y\ by auto + } + hence \y pd\ x\ using kx path_nodes path unfolding is_pd_def by auto + thus \?thesis\ .. + next \ \This is analogous argument\ + assume kx: \\ k \ x\ + hence ky: \\ k = y\ using \k by auto + have k_min: \\ k'. \ k' = x \ k \ k'\ using k_le unfolding k by auto + { + fix \' + and n' :: \nat\ + assume path': \is_path \'\ and \'0: \\' 0 = y\ and \'n': \\' n' = return\ + have path'': \is_path (\ @\<^bsup>k\<^esup> \')\ using path_cons[OF path path'] ky \'0 by auto + have \''0: \(\ @\<^bsup>k\<^esup> \') 0 = z\ using \0 by simp + have \''n: \(\ @\<^bsup>k\<^esup> \') (k+n') = return\ using \'n' ky \'0 by auto + obtain k' where k': \k' \ k + n'\ \(\ @\<^bsup>k\<^esup> \') k' = x\ using xz path'' \''0 \''n unfolding is_pd_def by blast + have **: \k \ k'\ proof (rule ccontr) + assume \\ k \ k'\ + hence \k' < k\ by simp + moreover + hence \\ k' = x\ using k' by auto + ultimately + show \False\ using k_min by force + qed + hence \\' (k' - k) = x\ using k' \'0 ky by auto + moreover + have \(k' - k) \ n'\ using k' by auto + ultimately + have \\ k\ n'. \' k = x\ by auto + } + hence \x pd\ y\ using ky path_nodes path unfolding is_pd_def by auto + thus \?thesis\ .. + qed +qed + +lemma pds_finite: \finite {y . (x,y) \ pdt}\ proof cases + assume \x \ nodes\ + then obtain \ n where \:\is_path \\ \\ 0 = x\ \\ n = return\ using reaching_ret by blast + have *: \\ y \ {y. (x,y)\ pdt}. y pd\ x\ using pdt_def by auto + have \\ y \ {y. (x,y)\ pdt}. \ k \ n. \ k = y\ using * \ is_pd_def by blast + hence \{y. (x,y)\ pdt} \ \ ` {..n}\ by auto + then show \?thesis\ using finite_surj by blast +next + assume \\ x\ nodes\ + hence \{y. (x,y)\pdt} = {}\ unfolding pdt_def is_pd_def using path_nodes reaching_ret by fastforce + then show \?thesis\ by simp +qed + +lemma ipd_exists: assumes node: \x \ nodes\ and not_ret: \x\return\ shows \\y. y ipd\ x\ +proof - + let \?Q\ = \{y. x\y \ y pd\ x}\ + have *: \return \ ?Q\ using assms return_pd by simp + hence **: \\ x. x\ ?Q\ by auto + have fin: \finite ?Q\ using pds_finite unfolding pdt_def by auto + have tot: \\ y z. y\?Q \ z \ ?Q \ z pd\ y \ y pd\ z\ using pd_total by auto + obtain y where ymax: \y\ ?Q\ \\ z\?Q. z = y \ z pd\ y\ using fin ** tot proof (induct) + case empty + then show \?case\ by auto + next + case (insert x F) show \thesis\ proof (cases \F = {}\) + assume \F = {}\ + thus \thesis\ using insert(4)[of \x\] by auto + next + assume \F \ {}\ + hence \\ x. x\ F\ by auto + have \\y. y \ F \ \z\F. z = y \ z pd\ y \ thesis\ proof - + fix y assume a: \y \ F\ \\z\F. z = y \ z pd\ y\ + have \x \ y\ using insert a by auto + have \x pd\ y \ y pd\ x\ using insert(6) a(1) by auto + thus \thesis\ proof + assume \x pd\ y\ + hence \\z\insert x F. z = y \ z pd\ y\ using a(2) by blast + thus \thesis\ using a(1) insert(4) by blast + next + assume \y pd\ x\ + have \\z\insert x F. z = x \ z pd\ x\ proof + fix z assume \z\ insert x F\ thus \z = x \ z pd\ x\ proof(rule,simp) + assume \z\F\ + hence \z = y \ z pd\ y\ using a(2) by auto + thus \z = x \ z pd\ x\ proof(rule,simp add: \y pd\ x\) + assume \z pd\ y\ + show \z = x \ z pd\ x\ using \y pd\ x\ \z pd\ y\ pd_trans by blast + qed + qed + qed + then show \thesis\ using insert by blast + qed + qed + then show \thesis\ using insert by blast + qed + qed + hence ***: \y pd\ x\ \x\y\ by auto + have \\ z. z \ x \ z pd\ x \ z pd\ y\ proof (rule,rule) + fix z + assume a: \ z \ x \ z pd\ x\ + hence b: \z \ ?Q\ by auto + have \y pd\ z \ z pd\ y\ using pd_total ***(1) a by auto + thus \z pd\ y\ proof + assume c: \y pd\ z\ + hence \y = z\ using b ymax pdt_def pd_antisym by auto + thus \z pd\ y\ using c by simp + qed simp + qed + with *** have \y ipd\ x\ unfolding is_ipd_def by simp + thus \?thesis\ by blast +qed + +lemma ipd_unique: assumes yipd: \y ipd\ x\ and y'ipd: \y' ipd\ x\ shows \y = y'\ +proof - + have 1: \y pd\ y'\ and 2: \y' pd\ y\ using yipd y'ipd unfolding is_ipd_def by auto + show \?thesis\ using pd_antisym[OF 1 2] . +qed + +lemma ipd_is_ipd: assumes \x \ nodes\ and \x\return\ shows \ipd x ipd\ x\ proof - + from assms obtain y where \y ipd\ x\ using ipd_exists by auto + moreover + hence \\ z. z ipd\x \ z = y\ using ipd_unique by simp + ultimately show \?thesis\ unfolding ipd_def by (auto intro: theI2) +qed + +lemma is_ipd_in_pdt: \y ipd\ x \ (x,y) \ pdt\ unfolding is_ipd_def pdt_def by auto + +lemma ipd_in_pdt: \x \ nodes \ x\return \ (x,ipd x) \ pdt\ by (metis ipd_is_ipd is_ipd_in_pdt) + +lemma no_pd_path: assumes \x \ nodes\ and \\ y pd\ x\ +obtains \ n where \is_path \\ and \\ 0 = x\ and \\ n = return\ and \\ k \ n. \ k \ y\ +proof (rule ccontr) + assume \\ thesis\ + hence \\ \ n. is_path \ \ \ 0 = x \ \ n = return \ (\ k\n . \ k = y)\ using that by force + thus \False\ using assms unfolding is_pd_def by auto +qed + +lemma pd_pd_ipd: assumes \x \ nodes\ \x\return\ \y\x\ \y pd\ x\ shows \y pd\ ipd x\ +proof - + have \ipd x pd\ x\ by (metis assms(1,2) ipd_is_ipd is_ipd_def) + hence \y pd\ ipd x \ ipd x pd\ y\ by (metis assms(4) pd_total) + thus \?thesis\ proof + have 1: \ipd x ipd\ x\ by (metis assms(1,2) ipd_is_ipd) + moreover + assume \ipd x pd\ y\ + ultimately + show \y pd\ ipd x\ unfolding is_ipd_def using assms(3,4) by auto + qed auto +qed + +lemma pd_nodes: assumes \y pd\ x\ shows pd_node1: \y \ nodes\ and pd_node2: \x \ nodes\ +proof - + obtain \ k where \is_path \\ \\ k = y\ using assms unfolding is_pd_def using reaching_ret by force + thus \y \ nodes\ using path_nodes by auto + show \x \ nodes\ using assms unfolding is_pd_def by simp +qed + +lemma pd_ret_is_ret: \x pd\ return \ x = return\ by (metis pd_antisym pd_node1 return_pd) + +lemma ret_path_none_pd: assumes \x \ nodes\ \x\return\ +obtains \ n where \is_path \\ \\ 0 = x\ \\ n = return\ \\ i>0. \ x pd\ \ i\ +proof(rule ccontr) + assume \\thesis\ + hence *: \\ \ n. \is_path \; \ 0 = x; \ n = return\ \ \i>0. x pd\ \ i\ using that by blast + obtain \ n where **: \is_path \\ \\ 0 = x\ \\ n = return\ \\ i>0. \ i \ x\ using direct_path_return[OF assms] by metis + then obtain i where ***: \i>0\ \x pd\ \ i\ using * by blast + hence \\ i \ return\ using pd_ret_is_ret assms(2) by auto + hence \i < n\ using assms(2) term_path_stable ** by (metis linorder_neqE_nat less_imp_le) + hence \(\\i)(n-i) = return\ using **(3) by auto + moreover + have \(\\i) (0) = \ i\ by simp + moreover + have \is_path (\\i)\ using **(1) path_path_shift by metis + ultimately + obtain k where \(\\i) k = x\ using ***(2) unfolding is_pd_def by metis + hence \\ (i + k) = x\ by auto + thus \False\ using **(4) \i>0\ by auto +qed + +lemma path_pd_ipd0': assumes \is_path \\ and \\ n \ return\ \\ n \ \ 0\ and \\ n pd\ \ 0\ +obtains k where \k \ n\ and \\ k = ipd(\ 0)\ +proof(rule ccontr) + have *: \\ n pd\ ipd (\ 0)\ by (metis is_pd_def assms(3,4) pd_pd_ipd pd_ret_is_ret) + obtain \' n' where **: \is_path \'\ \\' 0 = \ n\ \\' n' = return\ \\ i>0. \ \ n pd\ \' i\ by (metis assms(2) assms(4) pd_node1 ret_path_none_pd) + hence \\ i>0. \' i \ ipd (\ 0)\ using * by metis + moreover + assume \\ thesis\ + hence \\ k\n. \ k \ ipd (\ 0)\ using that by blast + ultimately + have \\ i. (\@\<^bsup>n\<^esup> \') i \ ipd (\ 0)\ by (metis diff_is_0_eq neq0_conv path_append_def) + moreover + have \(\@\<^bsup>n\<^esup> \') (n + n') = return\ + by (metis \\' 0 = \ n\ \\' n' = return\ add_diff_cancel_left' assms(2) diff_is_0_eq path_append_def) + moreover + have \(\@\<^bsup>n\<^esup> \') 0 = \ 0\ by (metis le0 path_append_def) + moreover + have \is_path (\@\<^bsup>n\<^esup> \')\ by (metis \\' 0 = \ n\ \is_path \'\ assms(1) path_cons) + moreover + have \ipd (\ 0) pd\ \ 0\ by (metis **(2,3,4) assms(2) assms(4) ipd_is_ipd is_ipd_def neq0_conv pd_node2) + moreover + have \\ 0 \ nodes\ by (metis assms(1) path_nodes) + ultimately + show \False\ unfolding is_pd_def by blast +qed + +lemma path_pd_ipd0: assumes \is_path \\ and \\ 0 \ return\ \\ n \ \ 0\ and \\ n pd\ \ 0\ +obtains k where \k \ n\ and \\ k = ipd(\ 0)\ +proof cases + assume *: \\ n = return\ + have \ipd (\ 0) pd\ (\ 0)\ by (metis is_ipd_def is_pd_def assms(2,4) ipd_is_ipd) + with assms(1,2,3) * show \thesis\ unfolding is_pd_def by (metis that) +next + assume \\ n \ return\ + from path_pd_ipd0' [OF assms(1) this assms(3,4)] that show \thesis\ by auto +qed + +lemma path_pd_ipd: assumes \is_path \\ and \\ k \ return\ \\ n \ \ k\ and \\ n pd\ \ k\ and kn: \k < n\ +obtains l where \k < l\ and \l \ n\ and \\ l = ipd(\ k)\ +proof - + have \is_path (\ \ k)\ \(\ \ k) 0 \ return\ \(\ \ k) (n - k) \ (\ \ k) 0\ \(\ \ k) (n - k) pd\ (\ \ k) 0\ + using assms path_path_shift by auto + with path_pd_ipd0[of \\\k\ \n-k\] + obtain ka where \ka \ n - k\ \(\ \ k) ka = ipd ((\ \ k) 0)\ . + hence \k + ka \ n\ \\ (k + ka) = ipd (\ k)\ using kn by auto + moreover + hence \\ (k + ka) ipd\ \ k\ by (metis assms(1) assms(2) ipd_is_ipd path_nodes) + hence \k < k + ka\ unfolding is_ipd_def by (metis nat_neq_iff not_add_less1) + ultimately + show \thesis\ using that[of \k+ka\] by auto +qed + +lemma path_ret_ipd: assumes \is_path \\ and \\ k \ return\ \\ n = return\ +obtains l where \k < l\ and \l \ n\ and \\ l = ipd(\ k)\ +proof - + have \\ n \ \ k\ using assms by auto + moreover + have \k \ n\ apply (rule ccontr) using term_path_stable assms by auto + hence \k < n\ by (metis assms(2,3) dual_order.order_iff_strict) + moreover + have \\ n pd\ \ k\ by (metis assms(1,3) path_nodes return_pd) + ultimately + obtain l where \k < l\ \l \ n\ \\ l = ipd (\ k)\ using assms path_pd_ipd by blast + thus \thesis\ using that by auto +qed + +lemma pd_intro: assumes \l pd\ k\ \is_path \\ \\ 0 = k\ \\ n = return\ +obtains i where \i \ n\ \\ i = l\ using assms unfolding is_pd_def by metis + +lemma path_pd_pd0: assumes path: \is_path \\ and lpdn: \\ l pd\ n\ and npd0: \n pd\ \ 0\ +obtains k where \k \ l\ \\ k = n\ +proof (rule ccontr) + assume \\ thesis\ + hence notn: \\ k. k \ l \ \ k \ n\ using that by blast + have nret: \\ l \ return\ by (metis is_pd_def assms(1,3) notn) + + obtain \' n' where path': \is_path \'\ and \0': \\' 0 = \ l\ and \n': \\' n' = return\ and nonepd: \\ i>0. \ \ l pd\ \' i\ + using nret path path_nodes ret_path_none_pd by metis + + have \\ l \ n\ using notn by simp + hence \\ i. \' i \ n\ using nonepd \0' lpdn by (metis neq0_conv) + + hence notn': \\ i. (\@\<^bsup>l\<^esup> \') i \ n\ using notn \0' by auto + + have \is_path (\@\<^bsup>l\<^esup> \')\ using path path' by (metis \0' path_cons) + moreover + have \(\@\<^bsup>l\<^esup> \') 0 = \ 0\ by simp + moreover + have \(\@\<^bsup>l\<^esup> \') (n' + l) = return\ using \0' \n' by auto + ultimately + show \False\ using notn' npd0 unfolding is_pd_def by blast +qed + + +subsection \Facts about Control Dependencies\ + +lemma icd_imp_cd: \n icd\<^bsup>\\<^esup>\ k \ n cd\<^bsup>\\<^esup>\ k\ by (metis is_icdi_def) + +lemma ipd_impl_not_cd: assumes \j \ {k..i}\ and \\ j = ipd (\ k)\ shows \\ i cd\<^bsup>\\<^esup>\ k\ + by (metis assms(1) assms(2) is_cdi_def) + +lemma cd_not_ret: assumes \i cd\<^bsup>\\<^esup>\ k \ shows \\ k \ return\ by (metis is_cdi_def assms nat_less_le term_path_stable) + +lemma cd_path_shift: assumes \j \ k\ \is_path \ \ shows \(i cd\<^bsup>\\<^esup>\ k) = (i - j cd\<^bsup>\\j\<^esup>\ k-j)\ proof + assume a: \i cd\<^bsup>\\<^esup>\ k\ + hence b: \k < i\ by (metis is_cdi_def) + hence \is_path (\ \ j)\ \k - j < i - j\ using assms apply (metis path_path_shift) + by (metis assms(1) b diff_less_mono) + moreover + have c: \\ j \ {k..i}. \ j \ ipd (\ k)\ by (metis a ipd_impl_not_cd) + hence \\ ja \ {k - j..i - j}. (\ \ j) ja \ ipd ((\ \ j) (k - j))\ using b assms by auto fastforce + moreover + have \j < i\ using assms(1) b by auto + hence \(\\j) (i - j) \ return\ using a unfolding is_cdi_def by auto + ultimately + show \i - j cd\<^bsup>\\j\<^esup>\ k-j\ unfolding is_cdi_def by simp +next + assume a: \i - j cd\<^bsup>\\j\<^esup>\ k-j\ + hence b: \k - j < i-j\ by (metis is_cdi_def) + moreover + have c: \\ ja \ {k - j..i - j}. (\ \ j) ja \ ipd ((\ \ j) (k - j))\ by (metis a ipd_impl_not_cd) + have \\ j \ {k..i}. \ j \ ipd (\ k)\ proof (rule,goal_cases) case (1 n) + hence \n-j \ {k-j..i-j}\ using assms by auto + hence \\ (j + (n-j)) \ ipd(\ (j + (k-j)))\ by (metis c path_shift_def) + thus \?case\ using 1 assms(1) by auto + qed + moreover + have \j < i\ using assms(1) b by auto + hence \\ i \ return\ using a unfolding is_cdi_def by auto + ultimately + show \i cd\<^bsup>\\<^esup>\k\ unfolding is_cdi_def by (metis assms(1) assms(2) diff_is_0_eq' le_diff_iff nat_le_linear nat_less_le) +qed + +lemma cd_path_shift0: assumes \is_path \\ shows \(i cd\<^bsup>\\<^esup>\ k) = (i-k cd\<^bsup>\\k\<^esup>\0)\ + using cd_path_shift[OF _ assms] by (metis diff_self_eq_0 le_refl) + +lemma icd_path_shift: assumes \l \ k\ \is_path \\ shows \(i icd\<^bsup>\\<^esup>\ k) = (i - l icd\<^bsup>\\l\<^esup>\ k - l)\ +proof - + have \is_path (\\l)\ using path_path_shift assms(2) by auto + moreover + have \(i cd\<^bsup>\\<^esup>\ k) = (i - l cd\<^bsup>\\l\<^esup>\ k - l)\ using assms cd_path_shift by auto + moreover + have \(\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m) = (\ m \ {k - l<.. i - l cd\<^bsup>\ \ l\<^esup>\ m)\ + proof - + {fix m assume *: \\ m \ {k - l<.. i - l cd\<^bsup>\ \ l\<^esup>\ m\ \m \ {k<.. + hence \m-l \ {k-l<.. using assms(1) by auto + hence \\ i - l cd\<^bsup>\\l\<^esup>\(m-l)\ using * by blast + moreover + have \l \ m\ using * assms by auto + ultimately have \\ i cd\<^bsup>\\<^esup>\m\ using assms(2) cd_path_shift by blast + } + moreover + {fix m assume *: \\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m\ \m-l \ {k-l<.. + hence \m \ {k<.. using assms(1) by auto + hence \\ i cd\<^bsup>\\<^esup>\m\ using * by blast + moreover + have \l \ m\ using * assms by auto + ultimately have \\ i - l cd\<^bsup>\\l\<^esup>\(m-l)\ using assms(2) cd_path_shift by blast + } + ultimately show \?thesis\ by auto (metis diff_add_inverse) + qed + ultimately + show \?thesis\ unfolding is_icdi_def using assms by blast +qed + +lemma icd_path_shift0: assumes \is_path \\ shows \(i icd\<^bsup>\\<^esup>\ k) = (i-k icd\<^bsup>\\k\<^esup>\0)\ + using icd_path_shift[OF _ assms] by (metis diff_self_eq_0 le_refl) + +lemma cdi_path_swap: assumes \is_path \'\ \j cd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>j\<^esub> \'\ shows \j cd\<^bsup>\'\<^esup>\k\ using assms unfolding eq_up_to_def is_cdi_def by auto + +lemma cdi_path_swap_le: assumes \is_path \'\ \j cd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>n\<^esub> \'\ \j \ n\ shows \j cd\<^bsup>\'\<^esup>\k\ by (metis assms cdi_path_swap eq_up_to_le) + +lemma not_cd_impl_ipd: assumes \is_path \\ and \k < i\ and \\ i cd\<^bsup>\\<^esup>\ k\ and \\ i \ return\ obtains j where \j \ {k..i}\ and \\ j = ipd (\ k)\ +by (metis assms(1) assms(2) assms(3) assms(4) is_cdi_def) + +lemma icd_is_the_icd: assumes \i icd\<^bsup>\\<^esup>\ k\ shows \k = (THE k. i icd\<^bsup>\\<^esup>\ k)\ using assms icd_uniq + by (metis the1_equality) + +lemma all_ipd_imp_ret: assumes \is_path \\ and \\ i. \ i \ return \ (\ j>i. \ j = ipd (\ i))\ shows \\j. \ j = return\ +proof - + { fix x assume *: \\ 0 = x\ + have \?thesis\ using wf_pdt_inv * assms + proof(induction \x\ arbitrary: \\\ rule: wf_induct_rule ) + case (less x \) show \?case\ proof (cases \x = return\) + case True thus \?thesis\ using less(2) by auto + next + assume not_ret: \x \ return\ + moreover + then obtain k where k_ipd: \\ k = ipd x\ using less(2,4) by auto + moreover + have \x \ nodes\ using less(2,3) by (metis path_nodes) + ultimately + have \(x, \ k) \ pdt\ by (metis ipd_in_pdt) + hence a: \(\ k, x) \ pdt_inv\ unfolding pdt_inv_def by simp + have b: \is_path (\ \ k)\ by (metis less.prems(2) path_path_shift) + have c: \\ i. (\\k) i \ return \ (\j>i. (\\k) j = ipd ((\\k) i))\ using less(4) apply auto + by (metis (full_types) ab_semigroup_add_class.add_ac(1) less_add_same_cancel1 less_imp_add_positive) + from less(1)[OF a _ b c] + have \\j. (\\k) j = return\ by auto + thus \\j. \ j = return\ by auto + qed + qed + } + thus \?thesis\ by simp +qed + +lemma loop_has_cd: assumes \is_path \\ \0 < i\ \\ i = \ 0\ \\ 0 \ return\ shows \\ k < i. i cd\<^bsup>\\<^esup>\ k\ proof (rule ccontr) + let \?\\ = \(\ n. \ (n mod i))\ + assume \\ (\k\\<^esup>\ k)\ + hence \\ k i cd\<^bsup>\\<^esup>\ k\ by blast + hence *: \\ kj \ {k..i}. \ j = ipd (\ k))\ using assms(1,3,4) not_cd_impl_ipd by metis + have \\ k. (\ j > k. ?\ j = ipd (?\ k))\ proof + fix k + have \k mod i < i\ using assms(2) by auto + with * obtain j where \j \ {(k mod i)..i}\ \\ j = ipd (\ (k mod i))\ by auto + then obtain j' where 1: \j' < i\ \\ j' = ipd (\ (k mod i))\ + by (cases \j = i\, auto ,metis assms(2) assms(3),metis le_neq_implies_less) + then obtain j'' where 2: \j'' > k\ \j'' mod i = j'\ by (metis mod_bound_instance) + hence \?\ j'' = ipd (?\ k)\ using 1 by auto + with 2(1) + show \\ j > k. ?\ j = ipd (?\ k)\ by auto + qed + moreover + have \is_path ?\\ by (metis assms(1) assms(2) assms(3) is_path_loop) + ultimately + obtain k where \?\ k = return\ by (metis (lifting) all_ipd_imp_ret) + moreover + have \k mod i < i\ by (simp add: assms(2)) + ultimately + have \\ i = return\ by (metis assms(1) term_path_stable less_imp_le) + thus \False\ by (metis assms(3) assms(4)) +qed + +lemma loop_has_cd': assumes \is_path \\ \j < i\ \\ i = \ j\ \\ j \ return\ shows \\ k \ {j..\\<^esup>\ k\ +proof - + have \\ k'< i-j. i-j cd\<^bsup>\\j\<^esup>\k'\ + apply(rule loop_has_cd) + apply (metis assms(1) path_path_shift) + apply (auto simp add: assms less_imp_le) + done + then obtain k where k: \k \i-j cd\<^bsup>\\j\<^esup>\k\ by auto + hence k': \(k+j) < i\ \i-j cd\<^bsup>\\j\<^esup>\ (k+j)-j\ by auto + note cd_path_shift[OF _ assms(1)] + hence \i cd\<^bsup>\\<^esup>\ k+j\ using k'(2) by (metis le_add1 add.commute) + with k'(1) show \?thesis\ by force +qed + +lemma claim'': assumes path\: \is_path \\ and path\': \is_path \'\ +and \i: \\ i = \' i'\ and \j: \\ j = \' j'\ +and not_cd: \\ k. \ j cd\<^bsup>\\<^esup>\ k\ \\ k. \ i' cd\<^bsup>\'\<^esup>\ k\ +and nret: \\ i \ return\ +and ilj: \i < j\ +shows \i' < j'\ proof (rule ccontr) + assume \\ i' < j'\ + hence jlei: \j' \ i'\ by auto + show \False\ proof (cases) + assume j'li': \j' < i'\ + define \'' where \\'' \ (\@\<^bsup>j\<^esup>(\'\j'))\i\ + note \''_def[simp] + have \\ j = (\' \ j') 0\ by (metis path_shift_def Nat.add_0_right \j) + hence \is_path \''\ using path\ path\' \''_def path_path_shift path_cons by presburger + moreover + have \\'' (j-i+(i'-j')) = \'' 0\ using ilj jlei \i \j + by (auto, metis add_diff_cancel_left' le_antisym le_diff_conv le_eq_less_or_eq) + moreover + have \\'' 0 \ return\ by (simp add: ilj less_or_eq_imp_le nret) + moreover + have \0 < j-i+(i'-j')\ by (metis add_is_0 ilj neq0_conv zero_less_diff) + ultimately obtain k where k: \k < j-i+(i'-j')\ \j-i+(i'-j') cd\<^bsup>\''\<^esup>\ k\ by (metis loop_has_cd) + hence *: \\ l \ {k..j-i+(i'-j')}. \'' l \ ipd (\'' k)\ by (metis is_cdi_def) + show \False\ proof (cases \k < j-i\) + assume a: \k < j - i\ + hence b: \\'' k = \ (i + k)\ by auto + have \\ l \ {i+k..j}. \ l \ ipd (\ (i+k))\ proof + fix l assume l: \l \ {i + k..j}\ + hence \\ l = \'' (l - i)\ by auto + moreover + from a l have \l-i \ {k .. j-i + (i'-j')}\ by force + ultimately show \\ l \ ipd (\ (i + k))\ using * b by auto + qed + moreover + have \i + k < j\ using a by simp + moreover + have \\ j \ return\ by (metis \i \j j'li' nret path\' term_path_stable less_imp_le) + ultimately + have \j cd\<^bsup>\\<^esup>\ i+k\ by (metis not_cd_impl_ipd path\) + thus \False\ by (metis not_cd(1)) + next + assume \\ k < j - i\ + hence a: \j - i \ k\ by simp + hence b: \\'' k = \' (j' + (i + k) - j)\ unfolding \''_def path_shift_def path_append_def using ilj + by(auto,metis \j add_diff_cancel_left' le_antisym le_diff_conv add.commute) + have \\ l \ {j' + (i+k) - j..i'}. \' l \ ipd (\' (j' + (i+k) - j))\ proof + fix l assume l: \l \ {j' + (i+k) - j..i'}\ + hence \\' l = \'' (j + l - i - j')\ unfolding \''_def path_shift_def path_append_def using ilj + by (auto, metis Nat.diff_add_assoc \j a add.commute add_diff_cancel_left' add_leD1 le_antisym le_diff_conv) + moreover + from a l have \j + l - i - j' \ {k .. j-i + (i'-j')}\ by force + ultimately show \\' l \ ipd (\' (j' + (i + k) - j))\ using * b by auto + qed + moreover + have \j' + (i+k) - j < i'\ using a j'li' ilj k(1) by linarith + moreover + have \\' i' \ return\ by (metis \i nret) + ultimately + have \i' cd\<^bsup>\'\<^esup>\ j' + (i+k) - j\ by (metis not_cd_impl_ipd path\') + thus \False\ by (metis not_cd(2)) + qed + next + assume \\ j' < i'\ + hence \j' = i'\ by (metis \\ i' < j'\ linorder_cases) + hence \\ i = \ j\ by (metis \i \j) + thus \False\ by (metis ilj loop_has_cd' not_cd(1) nret path\) +qed +qed + +lemma other_claim': assumes path: \is_path \\ and eq: \\ i = \ j\ and \\ i \ return\ +and icd: \\ k. \ i cd\<^bsup>\\<^esup>\ k\ and \\ k. \ j cd\<^bsup>\\<^esup>\ k\ shows \i = j\ +proof (rule ccontr,cases) + assume \i < j\ thus \False\ using assms claim'' by blast +next + assume \\ i < j\ \i \ j\ + hence \j < i\ by auto + thus \False\ using assms claim'' by (metis loop_has_cd') +qed + +lemma icd_no_cd_path_shift: assumes \i icd\<^bsup>\\<^esup>\ 0\ shows \(\ k. \ i - 1 cd\<^bsup>\\1\<^esup>\ k)\ +proof (rule,rule ccontr,goal_cases) + case (1 k) + hence *: \i - 1 cd\<^bsup>\ \ 1\<^esup>\ k\ by simp + have **: \1 \ k + 1\ by simp + have ***: \is_path \\ by (metis assms is_icdi_def) + hence \i cd\<^bsup>\\<^esup>\ k+1\ using cd_path_shift[OF ** ***] * by auto + moreover + hence \k+1 < i\ unfolding is_cdi_def by simp + moreover + have \0 < k + 1\ by simp + ultimately show \False\ using assms[unfolded is_icdi_def] by auto +qed + +lemma claim': assumes path\: \is_path \\ and path\': \is_path \'\ and + \i: \\ i = \' i'\ and \j: \\ j = \' j'\ and not_cd: + \i icd\<^bsup>\\<^esup>\ 0\ \j icd\<^bsup>\\<^esup>\ 0\ + \i' icd\<^bsup>\'\<^esup>\ 0\ \j' icd\<^bsup>\'\<^esup>\ 0\ + and ilj: \i < j\ + and nret: \\ i \ return\ + shows \i' < j'\ +proof - + have g0: \0 < i\ \0 < j\ \0 < i'\ \0 < j'\using not_cd[unfolded is_icdi_def is_cdi_def] by auto + have \(\ \ 1) (i - 1) = (\' \ 1) (i' - 1)\ \(\ \ 1) (j - 1) = (\' \ 1) (j' - 1)\ using \i \j g0 by auto + moreover + have \\ k. \ (j - 1) cd\<^bsup>\\1\<^esup>\ k\ \\ k. \ (i' - 1) cd\<^bsup>\'\1\<^esup>\ k\ + by (metis icd_no_cd_path_shift not_cd(2)) (metis icd_no_cd_path_shift not_cd(3)) + moreover + have \is_path (\\1)\ \is_path (\'\1)\ using path\ path\' path_path_shift by blast+ + moreover + have \(\\1) (i - 1) \ return\ using g0 nret by auto + moreover + have \i - 1 < j - 1\ using g0 ilj by auto + ultimately have \i' - 1 < j' - 1\ using claim'' by blast + thus \i' by auto +qed + +lemma other_claim: assumes path: \is_path \\ and eq: \\ i = \ j\ and \\ i \ return\ +and icd: \i icd\<^bsup>\\<^esup>\ 0\ and \j icd\<^bsup>\\<^esup>\ 0\ shows \i = j\ proof (rule ccontr,cases) + assume \i < j\ thus \False\ using assms claim' by blast +next + assume \\ i < j\ \i \ j\ + hence \j < i\ by auto + thus \False\ using assms claim' by (metis less_not_refl) +qed + +lemma cd_trans0: assumes \j cd\<^bsup>\\<^esup>\ 0\ and \k cd\<^bsup>\\<^esup>\j\ shows \k cd\<^bsup>\\<^esup>\ 0\ proof (rule ccontr) + have path: \is_path \\ and ij: \0 < j\ and jk: \j < k\ + and nret: \\ j \ return\ \\ k \ return\ + and noipdi: \\ l \ {0..j}. \ l \ ipd (\ 0)\ + and noipdj: \\ l \ {j..k}. \ l \ ipd (\ j)\ + using assms unfolding is_cdi_def by auto + assume \\ k cd\<^bsup>\\<^esup>\ 0\ + hence \\l \ {0..k}. \ l = ipd (\ 0)\ unfolding is_cdi_def using path ij jk nret by force + then obtain l where \l \ {0..k}\ and l: \\ l = ipd (\ 0)\ by auto + hence jl: \j and lk: \l\k\ using noipdi ij by auto + have pdj: \ipd (\ 0) pd\ \ j\ proof (rule ccontr) + have \\ j \ nodes\ using path by (metis path_nodes) + moreover + assume \\ ipd (\ 0) pd\ \ j\ + ultimately + obtain \' n where *: \is_path \'\ \\' 0 = \ j\ \\' n = return\ \\ k\n. \' k \ ipd(\ 0)\ using no_pd_path by metis + hence path': \is_path (\ @\<^bsup>j\<^esup> \')\ by (metis path path_cons) + moreover + have \\ k \ j + n. (\@\<^bsup>j\<^esup> \') k \ ipd (\ 0)\ using noipdi *(4) by auto + moreover + have \(\@\<^bsup>j\<^esup> \') 0 = \ 0\ by auto + moreover + have \(\@\<^bsup>j\<^esup> \') (j + n) = return\ using *(2,3) by auto + ultimately + have \\ ipd (\ 0) pd\ \ 0\ unfolding is_pd_def by metis + thus \False\ by (metis is_ipd_def ij ipd_is_ipd nret(1) path path_nodes term_path_stable less_imp_le) + qed + hence \(\\j) (l-j) pd\ (\\j) 0\ using jl l by auto + moreover + have \is_path (\\j)\ by (metis path path_path_shift) + moreover + have \\ l \ return\ by (metis lk nret(2) path term_path_stable) + hence \(\\j) (l-j) \ return\ using jl by auto + moreover + have \\ j \ ipd (\ 0)\ using noipdi by force + hence \(\\j) (l-j) \ (\\j) 0\ using jl l by auto + ultimately + obtain k' where \k' \ l-j\ and \(\\j) k' = ipd ((\\j) 0)\ using path_pd_ipd0' by blast + hence \j + k' \ {j..k}\ \\ (j+k') = ipd (\ j)\ using jl lk by auto + thus \False\ using noipdj by auto +qed + +lemma cd_trans: assumes \j cd\<^bsup>\\<^esup>\ i\ and \k cd\<^bsup>\\<^esup>\j\ shows \k cd\<^bsup>\\<^esup>\ i\ proof - + have path: \is_path \\ using assms is_cdi_def by auto + have ij: \i using assms is_cdi_def by auto + let \?\\ = \\\i\ + have \j-i cd\<^bsup>?\\<^esup>\ 0\ using assms(1) cd_path_shift0 path by auto + moreover + have \k-i cd\<^bsup>?\\<^esup>\j-i\ by (metis assms(2) cd_path_shift is_cdi_def ij less_imp_le_nat) + ultimately + have \k-i cd\<^bsup>?\\<^esup>\ 0\ using cd_trans0 by auto + thus \k cd\<^bsup>\\<^esup>\ i\ using path cd_path_shift0 by auto +qed + +lemma excd_impl_exicd: assumes \\ k. i cd\<^bsup>\\<^esup>\k\ shows \\ k. i icd\<^bsup>\\<^esup>\k\ +using assms proof(induction \i\ arbitrary: \\\ rule: less_induct) + case (less i) + then obtain k where k: \i cd\<^bsup>\\<^esup>\k\ by auto + hence ip: \is_path \\ unfolding is_cdi_def by auto + show \?case\ proof (cases) + assume *: \\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m\ + hence \i icd\<^bsup>\\<^esup>\k\ using k ip unfolding is_icdi_def by auto + thus \?case\ by auto + next + assume \\ (\ m \ {k<.. i cd\<^bsup>\\<^esup>\ m)\ + then obtain m where m: \m \ {k<.. \i cd\<^bsup>\\<^esup>\ m\ by blast + hence \i - m cd\<^bsup>\\m\<^esup>\ 0\ by (metis cd_path_shift0 is_cdi_def) + moreover + have \i - m < i\ using m by auto + ultimately + obtain k' where k': \i - m icd\<^bsup>\\m\<^esup>\ k'\ using less(1) by blast + hence \i icd\<^bsup>\\<^esup>\ k' + m\ using ip + by (metis add.commute add_diff_cancel_right' icd_path_shift le_add1) + thus \?case\ by auto + qed +qed + +lemma cd_split: assumes \i cd\<^bsup>\\<^esup>\ k\ and \\ i icd\<^bsup>\\<^esup>\ k\ obtains m where \i icd\<^bsup>\\<^esup>\ m\ and \m cd\<^bsup>\\<^esup>\ k\ +proof - + have ki: \k < i\ using assms is_cdi_def by auto + obtain m where m: \i icd\<^bsup>\\<^esup>\ m\ using assms(1) by (metis excd_impl_exicd) + hence \k \ m\ unfolding is_icdi_def using ki assms(1) by force + hence km: \k < m\using m assms(2) by (metis le_eq_less_or_eq) + moreover have \\ m \ return\ using m unfolding is_icdi_def is_cdi_def by (simp, metis term_path_stable less_imp_le) + moreover have \m using m unfolding is_cdi_def is_icdi_def by auto + ultimately + have \m cd\<^bsup>\\<^esup>\ k\ using assms(1) unfolding is_cdi_def by auto + with m that show \thesis\ by auto +qed + +lemma cd_induct[consumes 1, case_names base IS]: assumes prem: \i cd\<^bsup>\\<^esup>\ k\ and base: \\ i. i icd\<^bsup>\\<^esup>\k \ P i\ +and IH: \\ k' i'. k' cd\<^bsup>\\<^esup>\ k \ P k' \ i' icd\<^bsup>\\<^esup>\ k' \ P i'\ shows \P i\ +using prem IH proof (induction \i\ rule: less_induct,cases) + case (less i) + assume \i icd\<^bsup>\\<^esup>\ k\ + thus \P i\ using base by simp +next + case (less i') + assume \\ i' icd\<^bsup>\\<^esup>\ k\ + then obtain k' where k': \ i' icd\<^bsup>\\<^esup>\ k'\ \k' cd\<^bsup>\\<^esup>\ k\ using less cd_split by blast + hence icdk: \i' cd\<^bsup>\\<^esup>\ k'\ using is_icdi_def by auto + note ih=less(3)[OF k'(2) _ k'(1)] + have ki: \k' < i'\ using k' is_icdi_def is_cdi_def by auto + have \P k'\ using less(1)[OF ki k'(2) ] less(3) by auto + thus \P i'\ using ih by simp +qed + +lemma cdi_prefix: \n cd\<^bsup>\\<^esup>\ m \ m < n' \ n' \ n \ n' cd\<^bsup>\\<^esup>\ m\ unfolding is_cdi_def + by (simp, metis term_path_stable) + +lemma cr_wn': assumes 1: \n cd\<^bsup>\\<^esup>\ m\ and nc: \\ m' cd\<^bsup>\\<^esup>\ m\ and 3: \m < m'\ shows \n < m'\ +proof (rule ccontr) + assume \\ n < m'\ + hence \m' \ n\ by simp + hence \m' cd\<^bsup>\\<^esup>\ m\ by (metis 1 3 cdi_prefix) + thus \False\ using nc by simp +qed + +lemma cr_wn'': assumes \i cd\<^bsup>\\<^esup>\ m\ and \j cd\<^bsup>\\<^esup>\ n\ and \\ m cd\<^bsup>\\<^esup>\ n\ and \i \ j\ shows \m \ n\ proof (rule ccontr) + assume \\m\n\ + hence nm: \n < m\ by auto + moreover + have \m using assms(1) assms(4) unfolding is_cdi_def by auto + ultimately + have \m cd\<^bsup>\\<^esup>\ n\ using assms(2) cdi_prefix by auto + thus \False\ using assms(3) by auto +qed + +lemma ret_no_cd: assumes \\ n = return\ shows \\ n cd\<^bsup>\\<^esup>\ k\ by (metis assms is_cdi_def) + +lemma ipd_not_self: assumes \x \ nodes\ \x\ return\ shows \x \ ipd x\ by (metis is_ipd_def assms ipd_is_ipd) + +lemma icd_cs: assumes \l icd\<^bsup>\\<^esup>\k\ shows \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ [\ l]\ +proof - + from assms have \k = (THE k. l icd\<^bsup>\\<^esup>\ k)\ by (metis icd_is_the_icd) + with assms show \?thesis\ by auto +qed + +lemma cd_not_pd: assumes \l cd\<^bsup>\\<^esup>\ k\ \\ l \ \ k\ shows \\ \ l pd\ \ k\ proof + assume pd: \\ l pd\ \ k\ + have nret: \\ k \ return\ by (metis assms(1) pd pd_ret_is_ret ret_no_cd) + have kl: \k < l\ by (metis is_cdi_def assms(1)) + have path: \is_path \\ by (metis is_cdi_def assms(1)) + from path_pd_ipd[OF path nret assms(2) pd kl] + obtain n where \k < n\ \n \ l\ \\ n = ipd (\ k)\ . + thus \False\ using assms(1) unfolding is_cdi_def by auto +qed + +lemma cd_ipd_is_cd: assumes \k \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ and mcdj: \m cd\<^bsup>\\<^esup>\ j\ shows \k cd\<^bsup>\\<^esup>\ j\ proof cases + assume \j < k\ thus \k cd\<^bsup>\\<^esup>\ j\ by (metis mcdj assms(1) cdi_prefix less_imp_le_nat) +next + assume \\ j < k\ + hence kj: \k \ j\ by simp + have \k < j\ apply (rule ccontr) using kj assms mcdj by (auto, metis is_cdi_def is_ipd_def cd_not_pd ipd_is_ipd path_nodes term_path_stable less_imp_le) + moreover + have \j < m\ using mcdj is_cdi_def by auto + hence \\ n \ {k..j}. \ n \ ipd(\ k)\ using assms(3) by force + ultimately + have \j cd\<^bsup>\\<^esup>\ k\ by (metis mcdj is_cdi_def term_path_stable less_imp_le) + hence \m cd\<^bsup>\\<^esup>\ k\ by (metis mcdj cd_trans) + hence \False\ by (metis is_cdi_def is_ipd_def assms(2) cd_not_pd ipd_is_ipd path_nodes term_path_stable less_imp_le) + thus \?thesis\ by simp +qed + +lemma ipd_pd_cd0: assumes lcd: \n cd\<^bsup>\\<^esup>\ 0\ shows \ipd (\ 0) pd\ (\ n)\ +proof - + obtain k l where \0: \\ 0 = k\ and \n: \\ n = l\ and cdi: \n cd\<^bsup>\\<^esup>\ 0\ using lcd unfolding is_cdi_def by blast + have nret: \k \ return\ by (metis is_cdi_def \0 cdi term_path_stable less_imp_le) + have path: \is_path \\ and ipd: \\ i\n. \ i \ ipd k\ using cdi unfolding is_cdi_def \0 by auto + { + fix \' n' + assume path': \is_path \'\ + and \'0: \\' 0 = l\ + and ret: \\' n' = return\ + have \is_path (\ @\<^bsup>n\<^esup> \')\ using path path' \n \'0 by (metis path_cons) + moreover + have \(\ @\<^bsup>n\<^esup> \') (n+n') = return\ using ret \n \'0 by auto + moreover + have \(\ @\<^bsup>n\<^esup> \') 0 = k\ using \0 by auto + moreover + have \ipd k pd\ k\ by (metis is_ipd_def path \0 ipd_is_ipd nret path_nodes) + ultimately + obtain k' where k': \k' \ n+n'\ \(\ @\<^bsup>n\<^esup> \') k' = ipd k\ by (metis pd_intro) + have \\ k'\ n\ proof + assume \k' \ n\ + hence \(\ @\<^bsup>n\<^esup> \') k' = \ k'\ by auto + thus \False\ using k'(2) ipd by (metis \k' \ n\) + qed + hence \(\ @\<^bsup>n\<^esup> \') k' = \' (k' - n)\ by auto + moreover + have \(k' - n) \ n'\ using k' by simp + ultimately + have \\ k'\n'. \' k' = ipd k\ unfolding k' by auto + } + moreover + have \l \ nodes\ by (metis \n path path_nodes) + ultimately show \ipd (\ 0) pd\ (\ n)\ unfolding is_pd_def by (simp add: \0 \n) +qed + +lemma ipd_pd_cd: assumes lcd: \l cd\<^bsup>\\<^esup>\ k\ shows \ipd (\ k) pd\ (\ l)\ +proof - + have \l-k cd\<^bsup>\\k\<^esup>\0\ using lcd cd_path_shift0 is_cdi_def by blast + moreover + note ipd_pd_cd0[OF this] + moreover + have \(\ \ k) 0 = \ k\ by auto + moreover + have \k < l\ using lcd unfolding is_cdi_def by simp + then have \(\ \ k) (l - k) = \ l\ by simp + ultimately show \?thesis\ by simp +qed + +lemma cd_is_cd_ipd: assumes km: \k and ipd: \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ and cdj: \k cd\<^bsup>\\<^esup>\ j\ and nipdj: \ipd (\ j) \ \ m\ shows \m cd\<^bsup>\\<^esup>\ j\ proof - + have path: \is_path \\ + and jk: \j < k\ + and nretj: \\ k \ return\ + and nipd: \\ l \ {j..k}. \ l \ ipd (\ j)\ using cdj is_cdi_def by auto + have pd: \ipd (\ j) pd\ \ m\ by (metis atLeastAtMost_iff cdj ipd(1) ipd_pd_cd jk le_refl less_imp_le nipd nretj path path_nodes pd_pd_ipd) + have nretm: \\ m \ return\ by (metis nipdj pd pd_ret_is_ret) + have jm: \j < m\ using jk km by simp + show \m cd\<^bsup>\\<^esup>\ j\ proof (rule ccontr) + assume ncdj: \\ m cd\<^bsup>\\<^esup>\ j\ + hence \\ l \ {j..m}. \ l = ipd (\ j)\ unfolding is_cdi_def by (metis jm nretm path) + then obtain l + where jl: \j \ l\ and \l \ m\ + and lipd: \\ l = ipd (\ j)\ by force + hence lm: \l < m\ using nipdj by (metis le_eq_less_or_eq) + have npd: \\ ipd (\ k) pd\ \ l\ by (metis ipd(1) lipd nipdj pd pd_antisym) + have nd: \\ l \ nodes\ using path path_nodes by simp + from no_pd_path[OF nd npd] + obtain \' n where path': \is_path \'\ and \'0: \\' 0 = \ l\ and \'n: \\' n = return\ and nipd: \\ ka\n. \' ka \ ipd (\ k)\ . + let \?\\ = \(\@\<^bsup>l\<^esup> \') \ k\ + have path'': \is_path ?\\ by (metis \'0 path path' path_cons path_path_shift) + moreover + have kl: \k < l\ using lipd cdj jl unfolding is_cdi_def by fastforce + have \?\ 0 = \ k\ using kl by auto + moreover + have \?\ (l + n - k) = return\ using \'n \'0 kl by auto + moreover + have \ipd (\ k) pd\ \ k\ by (metis is_ipd_def ipd_is_ipd nretj path path_nodes) + ultimately + obtain l' where l': \l' \ (l + n - k)\ \?\ l' = ipd (\ k)\ unfolding is_pd_def by blast + show \False\ proof (cases ) + assume *: \k + l' \ l\ + hence \\ (k + l') = ipd (\ k)\ using l' by auto + moreover + have \k + l' < m\ by (metis "*" dual_order.strict_trans2 lm) + ultimately + show \False\ using ipd(2) by simp + next + assume \\ k + l' \ l\ + hence \\' (k + l' - l) = ipd (\ k)\ using l' by auto + moreover + have \k + l' - l \ n\ using l' kl by linarith + ultimately + show \False\ using nipd by auto + qed + qed +qed + +lemma ipd_icd_greatest_cd_not_ipd: assumes ipd: \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ +and km: \k < m\ and icdj: \m icd\<^bsup>\\<^esup>\ j\ shows \j = (GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m)\ +proof - + let \?j\ = \GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m\ + have kcdj: \k cd\<^bsup>\\<^esup>\ j\ using assms cd_ipd_is_cd is_icdi_def by blast + have nipd: \ipd (\ j) \ \ m\ using icdj unfolding is_icdi_def is_cdi_def by auto + have bound: \\ j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m \ j \ k\ unfolding is_cdi_def by simp + have exists: \k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m\ (is \?P j\) using kcdj nipd by auto + note GreatestI_nat[of \?P\ _ \k\, OF exists] Greatest_le_nat[of \?P\ \j\ \k\, OF exists] + hence kcdj': \k cd\<^bsup>\\<^esup>\ ?j\ and ipd': \ipd (\ ?j) \ \ m\ and jj: \j \ ?j\ using bound by auto + hence mcdj': \m cd\<^bsup>\\<^esup>\ ?j\ using ipd km cd_is_cd_ipd by auto + show \j = ?j\ proof (rule ccontr) + assume \j \ ?j\ + hence jlj: \j < ?j\ using jj by simp + moreover + have \?j < m\ using kcdj' km unfolding is_cdi_def by auto + ultimately + show \False\ using icdj mcdj' unfolding is_icdi_def by auto + qed +qed + +lemma cd_impl_icd_cd: assumes \i cd\<^bsup>\\<^esup>\ l\ and \i icd\<^bsup>\\<^esup>\ k\ and \\ i icd\<^bsup>\\<^esup>\ l\ shows \k cd\<^bsup>\\<^esup>\ l\ + using assms cd_split icd_uniq by metis + +lemma cdi_is_cd_icdi: assumes \k icd\<^bsup>\\<^esup>\ j\ shows \k cd\<^bsup>\\<^esup>\ i \ j cd\<^bsup>\\<^esup>\ i \ i = j\ + by (metis assms cd_impl_icd_cd cd_trans icd_imp_cd icd_uniq) + +lemma same_ipd_stable: assumes \k cd\<^bsup>\\<^esup>\ i\ \k cd\<^bsup>\\<^esup>\ j\ \i \ipd (\ i) = ipd (\ k)\ shows \ipd (\ j) = ipd (\ k)\ +proof - + have jcdi: \j cd\<^bsup>\\<^esup>\ i\ by (metis is_cdi_def assms(1,2,3) cr_wn' le_antisym less_imp_le_nat) + have 1: \ipd (\ j) pd\ \ k \ by (metis assms(2) ipd_pd_cd) + have 2: \ipd (\ k) pd\ \ j \ by (metis assms(4) ipd_pd_cd jcdi) + have 3: \ipd (\ k) pd\ (ipd (\ j))\ by (metis 2 IFC_def.is_cdi_def assms(1,2,4) atLeastAtMost_iff jcdi less_imp_le pd_node2 pd_pd_ipd) + have 4: \ipd (\ j) pd\ (ipd (\ k))\ by (metis 1 2 IFC_def.is_ipd_def assms(2) cd_not_pd ipd_is_ipd jcdi pd_node2 ret_no_cd) + show \?thesis\ using 3 4 pd_antisym by simp +qed + +lemma icd_pd_intermediate': assumes icd: \i icd\<^bsup>\\<^esup>\ k\ and j: \k < j\ \j < i\ shows \\ i pd\ (\ j)\ +using j proof (induction \i - j\ arbitrary: \j\ rule: less_induct) + case (less j) + have \\ i cd\<^bsup>\\<^esup>\ j\ using less.prems icd unfolding is_icdi_def by force + moreover + have \is_path \\ using icd by (metis is_icdi_def) + moreover + have \\ i \ return\ using icd by (metis is_icdi_def ret_no_cd) + ultimately + have \\ l. j \ l \ l \ i \ \ l = ipd (\ j)\ unfolding is_cdi_def using less.prems by auto + then obtain l where l: \j \ l\ \l \ i\ \\ l = ipd (\ j)\ by blast + hence lpd: \\ l pd\ (\ j)\ by (metis is_ipd_def \\ i \ return\ \is_path \\ ipd_is_ipd path_nodes term_path_stable) + show \?case\ proof (cases) + assume \l = i\ + thus \?case\ using lpd by auto + next + assume \l \ i\ + hence \l < i\ using l by simp + moreover + have \j \ l\ using l by (metis is_ipd_def \\ i \ return\ \is_path \\ ipd_is_ipd path_nodes term_path_stable) + hence \j < l\ using l by simp + moreover + hence \i - l < i - j\ by (metis diff_less_mono2 less.prems(2)) + moreover + have \k < l\ by (metis l(1) less.prems(1) linorder_neqE_nat not_le order.strict_trans) + ultimately + have \\ i pd\ (\ l)\ using less.hyps by auto + thus \?case\ using lpd by (metis pd_trans) + qed +qed + +lemma icd_pd_intermediate: assumes icd: \i icd\<^bsup>\\<^esup>\ k\ and j: \k < j\ \j \ i\ shows \\ i pd\ (\ j)\ +using assms icd_pd_intermediate'[OF assms(1,2)] apply (cases \j < i\,metis) by (metis is_icdi_def le_neq_trans path_nodes pd_refl) + +lemma no_icd_pd: assumes path: \is_path \\ and noicd: \\ l\n. \ k icd\<^bsup>\\<^esup>\ l\ and nk: \n \ k\ shows \\ k pd\ \ n\ +proof cases + assume \\ k = return\ thus \?thesis\ by (metis path path_nodes return_pd) +next + assume nret: \\ k \ return\ + have nocd: \\ l. n\l \ \ k cd\<^bsup>\\<^esup>\ l\ proof + fix l assume kcd: \k cd\<^bsup>\\<^esup>\ l\ and nl: \n \ l\ + hence \(k - n) cd\<^bsup>\\n\<^esup>\ (l - n)\ using cd_path_shift[OF nl path] by simp + hence \\ l. (k - n) icd\<^bsup>\\n\<^esup>\ l\ using excd_impl_exicd by blast + then guess l' .. + hence \k icd\<^bsup>\\<^esup>\ (l' + n)\ using icd_path_shift[of \n\ \l' + n\ \\\ \k\] path by auto + thus \False\ using noicd by auto + qed + hence \\l. n \ l \ l \ j \ {l..k}. \ j = ipd (\ l)\ using path nret unfolding is_cdi_def by auto + thus \?thesis\ using nk proof (induction \k - n\ arbitrary: \n\ rule: less_induct,cases) + case (less n) + assume \n = k\ + thus \?case\ using pd_refl path path_nodes by auto + next + case (less n) + assume \n \ k\ + hence nk: \n < k\ using less(3) by auto + with less(2) obtain j where jnk: \j \ {n..k}\ and ipdj: \\ j = ipd (\ n)\ by blast + have nretn: \\ n \ return\ using nk nret term_path_stable path by auto + with ipd_is_ipd path path_nodes is_ipd_def ipdj + have jpdn: \\ j pd\ \ n\ by auto + show \?case\ proof cases + assume \j = k\ thus \?case\ using jpdn by simp + next + assume \j \ k\ + hence jk: \j < k\ using jnk by auto + have \j \ n\ using ipdj by (metis ipd_not_self nretn path path_nodes) + hence nj: \n < j\ using jnk by auto + have *: \k - j < k - n\ using jk nj by auto + + with less(1)[OF *] less(2) jk nj + have \\ k pd\ \ j\ by auto + + thus \?thesis\ using jpdn pd_trans by metis + qed + qed +qed + + +lemma first_pd_no_cd: assumes path: \is_path \\ and pd: \\ n pd\ \ 0\ and first: \\ l < n. \ l \ \ n\ shows \\ l. \ n cd\<^bsup>\\<^esup>\ l\ +proof (rule ccontr, goal_cases) + case 1 + then obtain l where ncdl: \n cd\<^bsup>\\<^esup>\ l\ by blast + hence ln: \l < n\ using is_cdi_def by auto + have \\ \ n pd\ \ l\ using ncdl cd_not_pd by (metis ln first) + then obtain \' n' where path': \is_path \'\ and \0: \\' 0 = \ l\ and \n: \\' n' = return\ and not\n: \\ j\ n'. \' j \ \ n\ unfolding is_pd_def using path path_nodes by auto + let \?\\ = \\@\<^bsup>l\<^esup> \'\ + + have \is_path ?\\ by (metis \0 path path' path_cons) + moreover + have \?\ 0 = \ 0\ by auto + moreover + have \?\ (n' + l) = return\ using \0 \n by auto + ultimately + obtain j where j: \j \ n' + l\ and jn: \?\ j = \ n\ using pd unfolding is_pd_def by blast + show \False\ proof cases + assume \j \ l\ thus \False\ using jn first ln by auto + next + assume \\ j \ l\ thus \False\ using j jn not\n by auto + qed +qed + +lemma first_pd_no_icd: assumes path: \is_path \\ and pd: \\ n pd\ \ 0\ and first: \\ l < n. \ l \ \ n\ shows \\ l. \ n icd\<^bsup>\\<^esup>\ l\ + by (metis first first_pd_no_cd icd_imp_cd path pd) + +lemma path_nret_ex_nipd: assumes \is_path \\ \\ i. \ i \ return\ shows \\ i. (\ j\i. (\ k>j. \ k \ ipd (\ j)))\ proof(rule, rule ccontr) + fix i + assume \\ (\j\i. \ k>j. \ k \ ipd (\ j))\ + hence *: \\ j\i. (\k>j. \ k = ipd (\ j))\ by blast + have \\ j. (\k>j. (\\i) k = ipd ((\\i) j))\ proof + fix j + have \i + j \ i\ by auto + then obtain k where k: \k>i+j\ \\ k = ipd (\ (i+j))\ using * by blast + hence \(\\i) (k - i) = ipd ((\\i) j)\ by auto + moreover + have \k - i > j\ using k by auto + ultimately + show \\k>j. (\\i) k = ipd ((\\i) j)\ by auto + qed + moreover + have \is_path (\\i)\ using assms(1) path_path_shift by simp + ultimately + obtain k where \(\\i) k = return\ using all_ipd_imp_ret by blast + thus \False\ using assms(2) by auto +qed + +lemma path_nret_ex_all_cd: assumes \is_path \\ \\ i. \ i \ return\ shows \\ i. (\ j\i. (\ k>j. k cd\<^bsup>\\<^esup>\ j))\ +unfolding is_cdi_def using assms path_nret_ex_nipd[OF assms] by (metis atLeastAtMost_iff ipd_not_self linorder_neqE_nat not_le path_nodes) + + +lemma path_nret_inf_all_cd: assumes \is_path \\ \\ i. \ i \ return\ shows \\ finite {j. \ k>j. k cd\<^bsup>\\<^esup>\ j}\ +using unbounded_nat_set_infinite path_nret_ex_all_cd[OF assms] by auto + +lemma path_nret_inf_icd_seq: assumes path: \is_path \\ and nret: \\ i. \ i \ return\ +obtains f where \\ i. f (Suc i) icd\<^bsup>\\<^esup>\ f i\ \range f = {i. \ j>i. j cd\<^bsup>\\<^esup>\ i}\ \\ (\i. f 0 cd\<^bsup>\\<^esup>\ i)\ +proof - + note path_nret_inf_all_cd[OF assms] + then obtain f where ran: \range f = {j. \ k>j. k cd\<^bsup>\\<^esup>\ j}\ and asc: \\ i. f i < f (Suc i)\ using infinite_ascending by blast + have mono: \\ i j. i < j \ f i < f j\ using asc by (metis lift_Suc_mono_less) + { + fix i + have cd: \f (Suc i) cd\<^bsup>\\<^esup>\ f i\ using ran asc by auto + have \f (Suc i) icd\<^bsup>\\<^esup>\ f i\ proof (rule ccontr) + assume \\ f (Suc i) icd\<^bsup>\\<^esup>\ f i\ + then obtain m where im: \f i < m\ and mi: \ m < f (Suc i)\ and cdm: \f (Suc i) cd\<^bsup>\\<^esup>\ m\ unfolding is_icdi_def using assms(1) cd by auto + have \\ k>m. k cd\<^bsup>\\<^esup>\m\ proof (rule,rule,cases) + fix k assume \f (Suc i) < k\ + hence \k cd\<^bsup>\\<^esup>\ f (Suc i)\ using ran by auto + thus \k cd\<^bsup>\\<^esup>\ m\ using cdm cd_trans by metis + next + fix k assume mk: \m < k\ and \\ f (Suc i) < k\ + hence ik: \k \ f (Suc i)\ by simp + thus \k cd\<^bsup>\\<^esup>\ m\ using cdm by (metis cdi_prefix mk) + qed + hence \m \ range f\ using ran by blast + then obtain j where m: \m = f j\ by blast + show \False\ using im mi mono unfolding m by (metis Suc_lessI le_less not_le) + qed + } + moreover + { + fix m + assume cdm: \f 0 cd\<^bsup>\\<^esup>\ m\ + have \\ k>m. k cd\<^bsup>\\<^esup>\m\ proof (rule,rule,cases) + fix k assume \f 0 < k\ + hence \k cd\<^bsup>\\<^esup>\ f 0\ using ran by auto + thus \k cd\<^bsup>\\<^esup>\ m\ using cdm cd_trans by metis + next + fix k assume mk: \m < k\ and \\ f 0 < k\ + hence ik: \k \ f 0\ by simp + thus \k cd\<^bsup>\\<^esup>\ m\ using cdm by (metis cdi_prefix mk) + qed + hence \m \ range f\ using ran by blast + then obtain j where m: \m = f j\ by blast + hence fj0: \f j < f 0\ using cdm m is_cdi_def by auto + hence \0 < j\ by (metis less_irrefl neq0_conv) + hence \False\ using fj0 mono by fastforce + } + ultimately show \thesis\ using that ran by blast +qed + +lemma cdi_iff_no_strict_pd: \i cd\<^bsup>\\<^esup>\ k \ is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ (\ k, \ j) \ pdt)\ +proof + assume cd:\i cd\<^bsup>\\<^esup>\ k\ + have 1: \is_path \ \ k < i \ \ i \ return\ using cd unfolding is_cdi_def by auto + have 2: \\ j \ {k..i}. \ (\ k, \ j) \ pdt\ proof (rule ccontr) + assume \ \ (\j\{k..i}. (\ k, \ j) \ pdt)\ + then obtain j where \j \ {k..i}\ and \(\ k, \ j) \ pdt\ by auto + hence \\ j \ \ k\ and \\ j pd\ \ k\ unfolding pdt_def by auto + thus \False\ using path_pd_ipd by (metis \j \ {k..i}\ atLeastAtMost_iff cd cd_not_pd cdi_prefix le_eq_less_or_eq) + qed + show \is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ (\ k, \ j) \ pdt)\ using 1 2 by simp +next + assume \is_path \ \ k < i \ \ i \ return \ (\ j \ {k..i}. \ (\ k, \ j) \ pdt)\ + thus \i cd\<^bsup>\\<^esup>\ k\ by (metis ipd_in_pdt term_path_stable less_or_eq_imp_le not_cd_impl_ipd path_nodes) +qed + + +subsection \Facts about Control Slices\ + +lemma last_cs: \last (cs\<^bsup>\\<^esup> i) = \ i\ by auto + +lemma cs_not_nil: \cs\<^bsup>\\<^esup> n \ []\ by (auto) + +lemma cs_return: assumes \\ n = return\ shows \cs\<^bsup>\\<^esup> n = [\ n]\ by (metis assms cs.elims icd_imp_cd ret_no_cd) + +lemma cs_0[simp]: \cs\<^bsup>\\<^esup> 0 = [\ 0]\ using is_icdi_def is_cdi_def by auto + +lemma cs_inj: assumes \is_path \\ \\ n \ return\ \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> n'\ shows \n = n'\ +using assms proof (induction \cs\<^bsup>\\<^esup> n\ arbitrary: \\\ \n\ \n'\ rule:rev_induct) + case Nil hence \False\ using cs_not_nil by metis thus \?case\ by simp +next + case (snoc x xs \ n n') show \?case\ proof (cases \xs\) + case Nil + hence *: \\ (\ k. n icd\<^bsup>\\<^esup>\k)\ using snoc(2) cs_not_nil + by (auto,metis append1_eq_conv append_Nil cs_not_nil) + moreover + have \[x] = cs\<^bsup>\\<^esup> n'\ using Nil snoc by auto + hence **: \\ (\ k. n' icd\<^bsup>\\<^esup>\k)\ using cs_not_nil + by (auto,metis append1_eq_conv append_Nil cs_not_nil) + ultimately + have \\ k. \ n cd\<^bsup>\\<^esup>\ k\ \\ k. \ n' cd\<^bsup>\\<^esup>\ k\ using excd_impl_exicd by auto blast+ + moreover + hence \\ n = \ n'\ using snoc(5,2) by auto (metis * ** list.inject) + ultimately + show \n = n'\ using other_claim' snoc by blast +next + case (Cons y ys) + hence *: \\ k. n icd\<^bsup>\\<^esup>\k\ using snoc(2) by auto (metis append_is_Nil_conv list.distinct(1) list.inject) + then obtain k where k: \n icd\<^bsup>\\<^esup>\k\ by auto + have \k = (THE k . n icd\<^bsup>\\<^esup>\ k)\ using k by (metis icd_is_the_icd) + hence xsk: \xs = cs\<^bsup>\\<^esup> k\ using * k snoc(2) unfolding cs.simps[of \\\ \n\] by auto + have **: \\ k. n' icd\<^bsup>\\<^esup>\k\ using snoc(2)[unfolded snoc(5)] by auto (metis Cons append1_eq_conv append_Nil list.distinct(1)) + then obtain k' where k': \n' icd\<^bsup>\\<^esup>\ k'\ by auto + hence \k' = (THE k . n' icd\<^bsup>\\<^esup>\ k)\ using k' by (metis icd_is_the_icd) + hence xsk': \xs = cs\<^bsup>\\<^esup> k'\ using ** k' snoc(5,2) unfolding cs.simps[of \\\ \n'\] by auto + hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> k'\ using xsk by simp + moreover + have kn: \k < n\ using k by (metis is_icdi_def is_cdi_def) + hence \\ k \ return\ using snoc by (metis term_path_stable less_imp_le) + ultimately + have kk'[simp]: \k' = k\ using snoc(1) xsk snoc(3) by metis + have nk0: \n - k icd\<^bsup>\\k\<^esup>\ 0\ \n' - k icd\<^bsup>\\k\<^esup>\ 0\ using k k' icd_path_shift0 snoc(3) by auto + moreover + have nkr: \(\\k)(n-k) \ return\ using snoc(4) kn by auto + moreover + have \is_path (\\k)\ by (metis path_path_shift snoc.prems(1)) + moreover + have kn': \k < n'\ using k' kk' by (metis is_icdi_def is_cdi_def) + have \\ n = \ n'\ using snoc(5) * ** by auto + hence \(\\k)(n-k) = (\\k)(n'-k)\ using kn kn' by auto + ultimately + have \n - k = n' - k\ using other_claim by auto + thus \n = n'\ using kn kn' by auto +qed +qed + +lemma cs_cases: fixes \ i +obtains (base) \cs\<^bsup>\\<^esup> i = [\ i]\ and \\ k. \ i cd\<^bsup>\\<^esup>\ k\ | +(depend) k where \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ and \i icd\<^bsup>\\<^esup>\ k\ +proof cases + assume *: \\ k. i icd\<^bsup>\\<^esup>\ k\ + then obtain k where k: \i icd\<^bsup>\\<^esup>\ k\ .. + hence \k = (THE k. i icd\<^bsup>\\<^esup>\ k)\ by (metis icd_is_the_icd) + hence \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ using * by auto + with k that show \thesis\ by simp +next + assume *: \\ (\ k. i icd\<^bsup>\\<^esup>\ k)\ + hence \\ k. \ i cd\<^bsup>\\<^esup>\ k\ by (metis excd_impl_exicd) + moreover + have \cs\<^bsup>\\<^esup> i = [\ i]\ using * by auto + ultimately + show \thesis\ using that by simp +qed + +lemma cs_length_one: assumes \length (cs\<^bsup>\\<^esup> i) = 1\ shows \cs\<^bsup>\\<^esup> i = [\ i]\ and \\ k. \ i cd\<^bsup>\\<^esup>\ k\ + apply (cases \i\ \\\ rule: cs_cases) + using assms cs_not_nil + apply auto + apply (cases \i\ \\\ rule: cs_cases) + using assms cs_not_nil + by auto + +lemma cs_length_g_one: assumes \length (cs\<^bsup>\\<^esup> i) \ 1\ obtains k where \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ and \i icd\<^bsup>\\<^esup>\ k\ + apply (cases \i\ \\\ rule: cs_cases) + using assms cs_not_nil by auto + +lemma claim: assumes path: \is_path \\ \is_path \'\ and ii: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ and jj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ +and bl: \butlast (cs\<^bsup>\\<^esup> i) = butlast (cs\<^bsup>\\<^esup> j)\ and nret: \\ i \ return\ and ilj: \i < j\ +shows \i' < j'\ +proof (cases ) + assume *: \length (cs\<^bsup>\\<^esup> i) = 1\ + hence **: \length (cs\<^bsup>\\<^esup> i) = 1\ \length (cs\<^bsup>\\<^esup> j) = 1\ \length (cs\<^bsup>\'\<^esup> i') = 1\ \length (cs\<^bsup>\'\<^esup> j') = 1\ + apply metis + apply (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil) + apply (metis "*" ii) + by (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil jj) + then obtain \cs\<^bsup>\\<^esup> i = [\ i]\ \cs\<^bsup>\\<^esup> j = [\ j]\ \cs\<^bsup>\'\<^esup> j' = [\' j']\ \cs\<^bsup>\'\<^esup> i'= [\' i']\ + \\ k. \ j cd\<^bsup>\\<^esup>\ k\ \\ k. \ i' cd\<^bsup>\'\<^esup>\ k\ \\ k. \ j' cd\<^bsup>\'\<^esup>\ k\ + by (metis cs_length_one ** ) + moreover + hence \\ i = \' i'\ \\ j = \' j'\ using assms by auto + ultimately + show \i' < j'\ using nret ilj path claim'' by blast +next + assume *: \length (cs\<^bsup>\\<^esup> i) \ 1\ + hence **: \length (cs\<^bsup>\\<^esup> i) \ 1\ \length (cs\<^bsup>\\<^esup> j) \ 1\ \length (cs\<^bsup>\'\<^esup> i') \ 1\ \length (cs\<^bsup>\'\<^esup> j') \ 1\ + apply metis + apply (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil) + apply (metis "*" ii) + by (metis "*" bl butlast.simps(2) butlast_snoc cs_length_g_one cs_length_one(1) cs_not_nil jj) + obtain k l k' l' where ***: + \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> k)@[\ i]\ \cs\<^bsup>\\<^esup> j = (cs\<^bsup>\\<^esup> l)@[\ j]\ \cs\<^bsup>\'\<^esup> i' = (cs\<^bsup>\'\<^esup> k')@[\' i']\ \cs\<^bsup>\'\<^esup> j' = (cs\<^bsup>\'\<^esup> l')@[\' j']\ and + icds: \i icd\<^bsup>\\<^esup>\ k\ \j icd\<^bsup>\\<^esup>\ l\ \i' icd\<^bsup>\'\<^esup>\ k'\ \j' icd\<^bsup>\'\<^esup>\ l'\ + by (metis ** cs_length_g_one) + hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> l\ \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\'\<^esup> l'\ using assms by auto + moreover + have \\ k \ return\ \\' k' \ return\ using nret + apply (metis is_icdi_def icds(1) is_cdi_def term_path_stable less_imp_le) + by (metis is_cdi_def is_icdi_def icds(3) term_path_stable less_imp_le) + ultimately + have lk[simp]: \l = k\ \l' = k'\ using path cs_inj by auto + let \?\\ = \\ \ k\ + let \?\'\ = \\'\k'\ + have \i-k icd\<^bsup>?\\<^esup>\ 0\ \j-k icd\<^bsup>?\\<^esup>\ 0\ \i'-k' icd\<^bsup>?\'\<^esup>\ 0\ \j'-k' icd\<^bsup>?\'\<^esup>\ 0\ using icd_path_shift0 path icds by auto + moreover + have ki: \k < i\ using icds by (metis is_icdi_def is_cdi_def) + hence \i-k < j-k\ by (metis diff_is_0_eq diff_less_mono ilj nat_le_linear order.strict_trans) + moreover + have \i: \\ i = \' i'\ \\ j = \' j'\ using assms *** by auto + have \k' < i'\ \k' < j'\ using icds unfolding lk by (metis is_cdi_def is_icdi_def)+ + hence \?\ (i-k) = ?\' (i'-k')\ \?\ (j-k) = ?\' (j'-k')\ using \i ki ilj by auto + moreover + have \?\ (i-k) \ return\ using nret ki by auto + moreover + have \is_path ?\\ \is_path ?\'\ using path path_path_shift by auto + ultimately + have \i'-k' < j' - k'\ using claim' by blast + thus \i' < j'\ by (metis diff_is_0_eq diff_less_mono less_nat_zero_code linorder_neqE_nat nat_le_linear) +qed + +lemma cs_split': assumes \cs\<^bsup>\\<^esup> i = xs@[x,x']@ys\ shows \\ m. cs\<^bsup>\\<^esup> m = xs@[x] \ i cd\<^bsup>\\<^esup>\ m\ +using assms proof (induction \ys\ arbitrary: \i\ rule:rev_induct ) + case (snoc y ys) + hence \length (cs\<^bsup>\\<^esup> i) \ 1\ by auto + then obtain i' where \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> i') @ [\ i]\ and *: \i icd\<^bsup>\\<^esup>\ i'\ using cs_length_g_one[of \\\ \i\] by metis + hence \cs\<^bsup>\\<^esup> i' = xs@[x,x']@ys\ using snoc(2) by (metis append1_eq_conv append_assoc) + then obtain m where **: \cs\<^bsup>\\<^esup> m = xs @ [x]\ and \i' cd\<^bsup>\\<^esup>\ m\ using snoc(1) by blast + hence \i cd\<^bsup>\\<^esup>\ m\ using * cd_trans by (metis is_icdi_def) + with ** show \?case\ by blast +next + case Nil + hence \length (cs\<^bsup>\\<^esup> i) \ 1\ by auto + then obtain i' where a: \cs\<^bsup>\\<^esup> i = (cs\<^bsup>\\<^esup> i') @ [\ i]\ and *: \i icd\<^bsup>\\<^esup>\ i'\ using cs_length_g_one[of \\\ \i\] by metis + have \cs\<^bsup>\\<^esup> i = (xs@[x])@[x']\ using Nil by auto + hence \cs\<^bsup>\\<^esup> i' = xs@[x]\ using append1_eq_conv a by metis + thus \?case\ using * unfolding is_icdi_def by blast +qed + +lemma cs_split: assumes \cs\<^bsup>\\<^esup> i = xs@[x]@ys@[\ i]\ shows \\ m. cs\<^bsup>\\<^esup> m = xs@[x] \ i cd\<^bsup>\\<^esup>\ m\ proof - + obtain x' ys' where \ys@[\ i] = [x']@ys'\ by (metis append_Cons append_Nil neq_Nil_conv) + thus \?thesis\ using cs_split'[of \\\ \i\ \xs\ \x\ \x'\ \ys'\] assms by auto +qed + +lemma cs_less_split: assumes \xs \ ys\ obtains a as where \ys = xs@a#as\ + using assms unfolding cs_less.simps apply auto +by (metis Cons_nth_drop_Suc append_take_drop_id) + +lemma cs_select_is_cs: assumes \is_path \\ \xs \ Nil\ \xs \ cs\<^bsup>\\<^esup> k\ shows \cs\<^bsup>\\<^esup> (\\xs) = xs\ \k cd\<^bsup>\\<^esup>\ (\\xs)\proof - + obtain b bs where b: \cs\<^bsup>\\<^esup> k = xs@b#bs\ using assms cs_less_split by blast + obtain a as where a: \xs = as@[a]\ using assms by (metis rev_exhaust) + have \cs\<^bsup>\\<^esup> k = as@[a,b]@bs\ using a b by auto + then obtain k' where csk: \cs\<^bsup>\\<^esup> k' = xs\ and is_cd: \k cd\<^bsup>\\<^esup>\ k'\ using cs_split' a by blast + hence nret: \\ k' \ return\ by (metis is_cdi_def term_path_stable less_imp_le) + show a: \cs\<^bsup>\\<^esup> (\\xs) = xs\ unfolding cs_select_def using cs_inj[OF assms(1) nret] csk the_equality[of _ \k'\] + by (metis (mono_tags)) + show \k cd\<^bsup>\\<^esup>\ (\\xs)\ unfolding cs_select_def by (metis a assms(1) cs_inj cs_select_def csk is_cd nret) +qed + +lemma cd_in_cs: assumes \n cd\<^bsup>\\<^esup>\ m\ shows \\ ns. cs\<^bsup>\\<^esup> n = (cs\<^bsup>\\<^esup> m) @ ns @[\ n]\ +using assms proof (induction rule: cd_induct) + case (base n) thus \?case\ by (metis append_Nil cs.simps icd_is_the_icd) +next + case (IS k n) + hence \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> k @ [\ n]\ by (metis cs.simps icd_is_the_icd) + thus \?case\ using IS by force +qed + +lemma butlast_cs_not_cd: assumes \butlast (cs\<^bsup>\\<^esup> m) = butlast (cs\<^bsup>\\<^esup> n)\ shows \\ m cd\<^bsup>\\<^esup>\n\ +by (metis append_Cons append_Nil append_assoc assms cd_in_cs cs_not_nil list.distinct(1) self_append_conv snoc_eq_iff_butlast) + +lemma wn_cs_butlast: assumes \butlast (cs\<^bsup>\\<^esup> m) = butlast (cs\<^bsup>\\<^esup> n)\ \i cd\<^bsup>\\<^esup>\ m\ \j cd\<^bsup>\\<^esup>\ n\ \m shows \i +proof (rule ccontr) + assume \\ i < j\ + moreover + have \\ n cd\<^bsup>\\<^esup>\ m\ by (metis assms(1) butlast_cs_not_cd) + ultimately + have \n \ m\ using assms(2,3) cr_wn'' by auto + thus \False\ using assms(4) by auto +qed + + +text \This is the central theorem making the control slice suitable for matching indices between executions.\ + +theorem cs_order: assumes path: \is_path \\ \is_path \'\ and csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ +and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and nret: \\ i \ return\ and ilj: \i < j\ +shows \i' +proof - + have \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\\<^esup> j\ using cs_inj[OF path(1) nret] ilj by blast + moreover + have \cs\<^bsup>\\<^esup> i \ Nil\ \cs\<^bsup>\\<^esup> j \ Nil\ by (metis cs_not_nil)+ + ultimately show \?thesis\ proof (cases rule: list_neq_prefix_cases) + case (diverge xs x x' ys ys') + note csx = \cs\<^bsup>\\<^esup> i = xs @ [x] @ ys\ + note csx' = \cs\<^bsup>\\<^esup> j = xs @ [x'] @ ys'\ + note xx = \x \ x'\ + show \i' < j'\ proof (cases \ys\) + assume ys: \ys = Nil\ + show \?thesis\ proof (cases \ys'\) + assume ys': \ys' = Nil\ + have cs: \cs\<^bsup>\\<^esup> i = xs @ [x]\ \cs\<^bsup>\\<^esup> j = xs @ [x']\ by (metis append_Nil2 csx ys, metis append_Nil2 csx' ys') + hence bl: \butlast (cs\<^bsup>\\<^esup> i) = butlast (cs\<^bsup>\\<^esup> j)\ by auto + show \i' < j'\ using claim[OF path csi csj bl nret ilj] . + next + fix y' zs' + assume ys': \ys' = y'#zs'\ + have cs: \cs\<^bsup>\\<^esup> i = xs @ [x]\ \cs\<^bsup>\\<^esup> j = xs @ [x',y']@ zs'\ by (metis append_Nil2 csx ys, metis append_Cons append_Nil csx' ys') + obtain n where n: \cs\<^bsup>\\<^esup> n = xs@[x']\ and jn: \j cd\<^bsup>\\<^esup>\ n\ using cs cs_split' by blast + obtain n' where n': \cs\<^bsup>\'\<^esup> n' = xs@[x']\ and jn': \j' cd\<^bsup>\'\<^esup>\ n'\ using cs cs_split' unfolding csj by blast + have csn : \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and bl: \butlast (cs\<^bsup>\\<^esup> i) = butlast (cs\<^bsup>\\<^esup> n)\ using n n' cs by auto + hence bl': \butlast (cs\<^bsup>\'\<^esup> i') = butlast (cs\<^bsup>\'\<^esup> n')\ using csi by auto + have notcd: \\ i cd\<^bsup>\\<^esup>\ n\ by (metis butlast_cs_not_cd bl) + have nin: \i \ n\ using cs n xx by auto + have iln: \i < n\ apply (rule ccontr) using cr_wn'[OF jn notcd] nin ilj by auto + note claim[OF path csi csn bl nret iln] + hence \i' < n'\ . + thus \i' < j'\ using jn' unfolding is_cdi_def by auto + qed + next + fix y zs + assume ys: \ys = y#zs\ + show \?thesis\ proof (cases \ys'\) + assume ys' : \ys' = Nil\ + have cs: \cs\<^bsup>\\<^esup> i = xs @ [x,y]@zs\ \cs\<^bsup>\\<^esup> j = xs @ [x']\ by (metis append_Cons append_Nil csx ys, metis append_Nil2 csx' ys') + obtain n where n: \cs\<^bsup>\\<^esup> n = xs@[x]\ and jn: \i cd\<^bsup>\\<^esup>\ n\ using cs cs_split' by blast + obtain n' where n': \cs\<^bsup>\'\<^esup> n' = xs@[x]\ and jn': \i' cd\<^bsup>\'\<^esup>\ n'\ using cs cs_split' unfolding csi by blast + have csn : \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and bl: \butlast (cs\<^bsup>\\<^esup> n) = butlast (cs\<^bsup>\\<^esup> j)\ using n n' cs by auto + hence bl': \butlast (cs\<^bsup>\'\<^esup> j') = butlast (cs\<^bsup>\'\<^esup> n')\ using csj by auto + have notcd: \\ j' cd\<^bsup>\'\<^esup>\ n'\ by (metis butlast_cs_not_cd bl') + have nin: \n < i\ using jn unfolding is_cdi_def by auto + have nlj: \n < j\ using nin ilj by auto + note claim[OF path csn csj bl _ nlj] + hence nj': \n' < j'\ using term_path_stable[OF path(1) _] less_imp_le nin nret by auto + show \i' < j'\ apply(rule ccontr) using cdi_prefix[OF jn' nj'] notcd by auto + next + fix y' zs' + assume ys' : \ys' = y'#zs'\ + have cs: \cs\<^bsup>\\<^esup> i = xs@[x,y]@zs\ \cs\<^bsup>\\<^esup> j = xs@[x',y']@zs'\ by (metis append_Cons append_Nil csx ys,metis append_Cons append_Nil csx' ys') + have neq: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\\<^esup> j\ using cs_inj path nret ilj by blast + obtain m where m: \cs\<^bsup>\\<^esup> m = xs@[x]\ and im: \i cd\<^bsup>\\<^esup>\ m\ using cs cs_split' by blast + obtain n where n: \cs\<^bsup>\\<^esup> n = xs@[x']\ and jn: \j cd\<^bsup>\\<^esup>\ n\ using cs cs_split' by blast + obtain m' where m': \cs\<^bsup>\'\<^esup> m' = xs@[x]\ and im': \i' cd\<^bsup>\'\<^esup>\ m'\ using cs cs_split' unfolding csi by blast + obtain n' where n': \cs\<^bsup>\'\<^esup> n' = xs@[x']\ and jn': \j' cd\<^bsup>\'\<^esup>\ n'\ using cs cs_split' unfolding csj by blast + have \m \ n\ using ilj m n wn_cs_butlast[OF _ jn im] by force + moreover + have \m \ n\ using m n xx by (metis last_snoc) + ultimately + have mn: \m < n\ by auto + moreover + have \\ m \ return\ by (metis last_cs last_snoc m mn n path(1) term_path_stable xx less_imp_le) + moreover + have \butlast (cs\<^bsup>\\<^esup> m) = butlast (cs\<^bsup>\\<^esup> n)\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ using m n n' m' by auto + ultimately + have \m' < n'\ using claim path by blast + thus \i' < j'\ using m' n' im' jn' wn_cs_butlast by (metis butlast_snoc) + qed + qed + next + case (prefix1 xs) + note pfx = \cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> j @ xs\ + note xs = \xs \ []\ + obtain a as where \xs = a#as\ using xs by (metis list.exhaust) + moreover + obtain bs b where bj: \cs\<^bsup>\\<^esup> j = bs@[b]\ using cs_not_nil by (metis rev_exhaust) + ultimately + have \cs\<^bsup>\\<^esup> i = bs@[b,a]@as\ using pfx by auto + then obtain m where \cs\<^bsup>\\<^esup> m = bs@[b]\ and cdep: \i cd\<^bsup>\\<^esup>\ m\ using cs_split' by blast + hence mi: \m = j\ using bj cs_inj by (metis is_cdi_def term_path_stable less_imp_le) + hence \i cd\<^bsup>\\<^esup>\ j\ using cdep by auto + hence \False\ using ilj unfolding is_cdi_def by auto + thus \i' < j'\ .. + next + case (prefix2 xs) + have pfx : \cs\<^bsup>\'\<^esup> i' @ xs = cs\<^bsup>\'\<^esup> j'\ using prefix2 csi csj by auto + note xs = \xs \ []\ + obtain a as where \xs = a#as\ using xs by (metis list.exhaust) + moreover + obtain bs b where bj: \cs\<^bsup>\'\<^esup> i' = bs@[b]\ using cs_not_nil by (metis rev_exhaust) + ultimately + have \cs\<^bsup>\'\<^esup> j' = bs@[b,a]@as\ using pfx by auto + then obtain m where \cs\<^bsup>\'\<^esup> m = bs@[b]\ and cdep: \j' cd\<^bsup>\'\<^esup>\ m\ using cs_split' by blast + hence mi: \m = i'\ using bj cs_inj by (metis is_cdi_def term_path_stable less_imp_le) + hence \j' cd\<^bsup>\'\<^esup>\ i'\ using cdep by auto + thus \i' < j'\ unfolding is_cdi_def by auto + qed +qed + +lemma cs_order_le: assumes path: \is_path \\ \is_path \'\ and csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ +and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and nret: \\ i \ return\ and ilj: \i \ j\ +shows \i'\j'\ proof cases + assume \i < j\ with cs_order[OF assms(1,2,3,4,5)] show \?thesis\ by simp +next + assume \\ i < j\ + hence \i = j\ using ilj by simp + hence csij: \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\'\<^esup> j'\ using csi csj by simp + have nret': \\' i' \ return\ using nret last_cs csi by metis + show \?thesis\ using cs_inj[OF path(2) nret' csij] by simp +qed + +lemmas cs_induct[case_names cs] = cs.induct + +lemma icdi_path_swap: assumes \is_path \'\ \j icd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>j\<^esub> \'\ shows \j icd\<^bsup>\'\<^esup>\k\ using assms unfolding eq_up_to_def is_icdi_def is_cdi_def by auto + +lemma icdi_path_swap_le: assumes \is_path \'\ \j icd\<^bsup>\\<^esup>\k\ \\ =\<^bsub>n\<^esub> \'\ \j \ n\ shows \j icd\<^bsup>\'\<^esup>\k\ by (metis assms icdi_path_swap eq_up_to_le) + +lemma cs_path_swap: assumes \is_path \\ \is_path \'\ \\ =\<^bsub>k\<^esub> \'\ shows \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k\ using assms(1,3) proof (induction \\\ \k\ rule:cs_induct,cases) + case (cs \ k) + let \?l\ = \(THE l. k icd\<^bsup>\\<^esup>\ l)\ + assume *: \\l. k icd\<^bsup>\\<^esup>\ l\ + have kicd: \k icd\<^bsup>\\<^esup>\ ?l\ by (metis "*" icd_is_the_icd) + hence \?l < k\ unfolding is_cdi_def[of \k\ \\\ \?l\] is_icdi_def[of \k\ \\\ \?l\] by auto + hence \\ i\?l. \ i = \' i\ using cs(2,3) unfolding eq_up_to_def by auto + hence csl: \cs\<^bsup>\\<^esup> ?l = cs\<^bsup>\'\<^esup> ?l\ using cs(1,2) * unfolding eq_up_to_def by auto + have kicd: \k icd\<^bsup>\\<^esup>\ ?l\ by (metis "*" icd_is_the_icd) + hence csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> ?l @ [\ k]\ using kicd by auto + have kicd': \k icd\<^bsup>\'\<^esup>\ ?l\ using kicd icdi_path_swap[OF assms(2) _ cs(3)] by simp + hence \?l = (THE l. k icd\<^bsup>\'\<^esup>\ l)\ by (metis icd_is_the_icd) + hence csk': \cs\<^bsup>\'\<^esup> k = cs\<^bsup>\'\<^esup> ?l @ [\' k]\ using kicd' by auto + have \\' k = \ k\ using cs(3) unfolding eq_up_to_def by auto + with csl csk csk' + show \?case\ by auto +next + case (cs \ k) + assume *: \\ (\l. k icd\<^bsup>\\<^esup>\ l)\ + hence csk: \cs\<^bsup>\\<^esup> k = [\ k]\ by auto + have \\ (\l. k icd\<^bsup>\'\<^esup>\ l)\ apply (rule ccontr) using * icdi_path_swap_le[OF cs(2) _, of \k\ \\'\] cs(3) by (metis eq_up_to_sym le_refl) + hence csk': \cs\<^bsup>\'\<^esup> k = [\' k]\ by auto + with csk show \?case\ using cs(3) eq_up_to_apply by auto +qed + +lemma cs_path_swap_le: assumes \is_path \\ \is_path \'\ \\ =\<^bsub>n\<^esub> \'\ \k \ n\ shows \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k\ by (metis assms cs_path_swap eq_up_to_le) + +lemma cs_path_swap_cd: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and \n cd\<^bsup>\\<^esup>\ k\ +obtains k' where \n' cd\<^bsup>\'\<^esup>\ k'\ and \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ +proof - + from cd_in_cs[OF assms(4)] + obtain ns where *: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> k @ ns @ [\ n]\ by blast + obtain xs x where csk: \cs\<^bsup>\\<^esup> k = xs @ [x]\ by (metis cs_not_nil rev_exhaust) + have \\ n = \' n'\ using assms(3) last_cs by metis + hence **: \cs\<^bsup>\'\<^esup> n' = xs@[x]@ns@[\' n']\ using * assms(3) csk by auto + from cs_split[OF **] + obtain k' where \cs\<^bsup>\'\<^esup> k' = xs @ [x]\ \n' cd\<^bsup>\'\<^esup>\ k'\ by blast + thus \thesis\ using that csk by auto +qed + +lemma path_ipd_swap: assumes \is_path \\ \\ k \ return\ \k < n\ +obtains \' m where \is_path \'\ \\ =\<^bsub>n\<^esub> \'\ \k < m\ \\' m = ipd (\' k)\ \\ l \ {k..' l \ ipd (\' k)\ +proof - + obtain \' r where *: \\' 0 = \ n\ \is_path \'\ \\' r = return\ by (metis assms(1) path_nodes reaching_ret) + let \?\\ = \\@\<^bsup>n\<^esup> \'\ + have path: \is_path ?\\ and ret: \?\ (n + r) = return\ and equpto: \?\ =\<^bsub>n\<^esub> \\ using assms path_cons * path_append_eq_up_to by auto + have \k: \?\ k = \ k\ by (metis assms(3) less_imp_le_nat path_append_def) + obtain j where j: \k < j \ j \ (n + r) \ ?\ j = ipd (\ k)\ (is \?P j\ )by (metis \k assms(2) path path_ret_ipd ret) + define m where m: \m \ LEAST m . ?P m\ + have Pm: \?P m\ using LeastI[of \?P\ \j\] j m by auto + hence km: \k < m\ \m \ (n + r)\ \?\ m = ipd (\ k)\ by auto + have le: \\l. ?P l \ m \ l\ using Least_le[of \?P\] m by blast + have \knipd: \?\ k \ ipd (\ k)\ by (metis \k assms(1) assms(2) ipd_not_self path_nodes) + have nipd': \\l. k < l \ l < m \ ?\ l \ ipd (\ k)\ apply (rule ccontr) using le km(2) by force + have \\ l \ {k.. l \ ipd(\ k)\ using \knipd nipd' by(auto, metis le_eq_less_or_eq,metis le_eq_less_or_eq) + thus \thesis\ using that by (metis \k eq_up_to_sym km(1) km(3) path path_append_eq_up_to) +qed + +lemma cs_sorted_list_of_cd': \cs\<^bsup>\\<^esup> k = map \ (sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i}) @ [\ k]\ +proof (induction \\\ \k\ rule: cs.induct, cases) + case (1 \ k) + assume \\ j. k icd\<^bsup>\\<^esup>\ j\ + then guess j .. + note j = this + hence csj: \cs\<^bsup>\\<^esup> j = map \ (sorted_list_of_set {i. j cd\<^bsup>\\<^esup>\ i}) @ [\ j]\ by (metis "1.IH" icd_is_the_icd) + have \{i. k cd\<^bsup>\\<^esup>\ i} = insert j {i. j cd\<^bsup>\\<^esup>\ i}\ using cdi_is_cd_icdi[OF j] by auto + moreover + have f: \finite {i. j cd\<^bsup>\\<^esup>\ i}\ unfolding is_cdi_def by auto + moreover + have \j \ {i. j cd\<^bsup>\\<^esup>\ i}\ unfolding is_cdi_def by auto + ultimately + have \sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i} = insort j (sorted_list_of_set { i . j cd\<^bsup>\\<^esup>\ i})\ using sorted_list_of_set_insert by auto + moreover + have \\ x \ {i. j cd\<^bsup>\\<^esup>\ i}. x < j\ unfolding is_cdi_def by auto + hence \\ x \ set (sorted_list_of_set {i. j cd\<^bsup>\\<^esup>\ i}). x < j\ by (simp add: f) + ultimately + have \sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i} = (sorted_list_of_set { i . j cd\<^bsup>\\<^esup>\ i})@[j]\ using insort_greater by auto + hence \cs\<^bsup>\\<^esup> j = map \ (sorted_list_of_set { i . k cd\<^bsup>\\<^esup>\ i})\ using csj by auto + thus \?case\ by (metis icd_cs j) +next + case (1 \ k) + assume *: \\ (\ j. k icd\<^bsup>\\<^esup>\ j)\ + hence \cs\<^bsup>\\<^esup> k = [\ k]\ by (metis cs_cases) + moreover + have \{ i . k cd\<^bsup>\\<^esup>\ i} = {}\ by (auto, metis * excd_impl_exicd) + ultimately + show \?case\ by (metis append_Nil list.simps(8) sorted_list_of_set_empty) +qed + +lemma cs_sorted_list_of_cd: \cs\<^bsup>\\<^esup> k = map \ (sorted_list_of_set ({ i . k cd\<^bsup>\\<^esup>\ i} \ {k}))\ proof - + have le: \\ x \ {i. k cd\<^bsup>\\<^esup>\i}.\ y \ {k}. x < y\ unfolding is_cdi_def by auto + have fin: \finite {i. k cd\<^bsup>\\<^esup>\i}\ \finite {k}\ unfolding is_cdi_def by auto + show \?thesis\ unfolding cs_sorted_list_of_cd'[of \\\ \k\] sorted_list_of_set_append[OF fin le] by auto +qed + +lemma cs_not_ipd: assumes \k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd (\ k)\ (is \?Q j\) +shows \cs\<^bsup>\\<^esup> (GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd (\ k)) = [n\cs\<^bsup>\\<^esup> k . ipd n \ ipd (\ k)]\ +(is \cs\<^bsup>\\<^esup> ?j = filter ?P _\) +proof - + have csk: \cs\<^bsup>\\<^esup> k = map \ (sorted_list_of_set ({ i . k cd\<^bsup>\\<^esup>\ i } \ {k}))\ by (metis cs_sorted_list_of_cd) + have csj: \cs\<^bsup>\\<^esup> ?j = map \ (sorted_list_of_set ({i. ?j cd\<^bsup>\\<^esup>\ i } \ {?j}))\ by (metis cs_sorted_list_of_cd) + + have bound: \\ j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd(\ k) \ j \ k\ unfolding is_cdi_def by simp + + have kcdj: \k cd\<^bsup>\\<^esup>\ ?j\ and ipd': \ipd (\ ?j) \ ipd(\ k)\ using GreatestI_nat[of \?Q\ \j\ \k\, OF assms] bound by auto + + have greatest: \\ j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ ipd (\ k) \ j \ ?j\ using Greatest_le_nat[of \?Q\ _ \k\] bound by auto + have less_not_ipdk: \\ j. k cd\<^bsup>\\<^esup>\ j \ j < ?j \ ipd (\ j) \ ipd (\ k)\ by (metis (lifting) ipd' kcdj same_ipd_stable) + hence le_not_ipdk: \\ j. k cd\<^bsup>\\<^esup>\ j \ j \ ?j \ ipd (\ j) \ ipd (\ k)\ using kcdj ipd' by (case_tac \j = ?j\,auto) + have *: \{j \ {i. k cd\<^bsup>\\<^esup>\i} \ {k}. ?P (\ j)} = insert ?j { i . ?j cd\<^bsup>\\<^esup>\ i} \ + apply auto + apply (metis (lifting, no_types) greatest cr_wn'' kcdj le_antisym le_refl) + apply (metis kcdj) + apply (metis ipd') + apply (metis (full_types) cd_trans kcdj) + apply (subgoal_tac \k cd\<^bsup>\\<^esup>\ x\) + apply (metis (lifting, no_types) is_cdi_def less_not_ipdk) + by (metis (full_types) cd_trans kcdj) + have \finite ({i . k cd\<^bsup>\\<^esup>\ i} \ {k})\ unfolding is_cdi_def by auto + note filter_sorted_list_of_set[OF this, of \?P o \\] + hence \[n\cs\<^bsup>\\<^esup> k . ipd n \ ipd(\ k)] = map \ (sorted_list_of_set {j \ {i. k cd\<^bsup>\\<^esup>\i} \ {k}. ?P (\ j)})\ unfolding csk filter_map by auto + also + have \\ = map \ (sorted_list_of_set (insert ?j { i . ?j cd\<^bsup>\\<^esup>\ i}))\ unfolding * by auto + also + have \\ = cs\<^bsup>\\<^esup> ?j\ using csj by auto + finally + show \?thesis\ by metis +qed + +lemma cs_ipd: assumes ipd: \\ m = ipd (\ k)\ \\ n \ {k.. n \ ipd (\ k)\ +and km: \k < m\ shows \cs\<^bsup>\\<^esup> m = [n\cs\<^bsup>\\<^esup> k . ipd n \ \ m] @ [\ m]\ +proof cases + assume \\ j. m icd\<^bsup>\\<^esup>\ j\ + then obtain j where jicd: \m icd\<^bsup>\\<^esup>\ j\ by blast + hence *: \cs\<^bsup>\\<^esup> m = cs\<^bsup>\\<^esup> j @ [\ m]\ by (metis icd_cs) + have j: \j = (GREATEST j. k cd\<^bsup>\\<^esup>\ j \ ipd (\ j) \ \ m)\ using jicd assms ipd_icd_greatest_cd_not_ipd by blast + moreover + have \ipd (\ j) \ ipd (\ k)\ by (metis is_cdi_def is_icdi_def is_ipd_def cd_not_pd ipd(1) ipd_is_ipd jicd path_nodes less_imp_le term_path_stable) + moreover + have \k cd\<^bsup>\\<^esup>\ j\ unfolding j by (metis (lifting, no_types) assms(3) cd_ipd_is_cd icd_imp_cd ipd(1) ipd(2) j jicd) + ultimately + have \cs\<^bsup>\\<^esup> j = [n\cs\<^bsup>\\<^esup> k . ipd n \ \ m]\ using cs_not_ipd[of \k\ \\\ \j\] ipd(1) by metis + thus \?thesis\ using * by metis +next + assume noicd: \\ (\ j. m icd\<^bsup>\\<^esup>\ j)\ + hence csm: \cs\<^bsup>\\<^esup> m = [\ m]\ by auto + have \\j. k cd\<^bsup>\\<^esup>\j \ ipd(\ j) = \ m\ using cd_is_cd_ipd[OF km ipd] by (metis excd_impl_exicd noicd) + hence *: \{j \ {i. k cd\<^bsup>\\<^esup>\ i} \ {k}. ipd (\ j) \ \ m} = {}\ using ipd(1) by auto + have **: \((\n. ipd n \ \ m) o \) = (\n. ipd (\ n) \ \ m)\ by auto + have fin: \finite ({i. k cd\<^bsup>\\<^esup>\ i} \ {k})\ unfolding is_cdi_def by auto + note csk = cs_sorted_list_of_cd[of \\\ \k\] + hence \[n\cs\<^bsup>\\<^esup> k . ipd n \ \ m] = [n\ (map \ (sorted_list_of_set ({i. k cd\<^bsup>\\<^esup>\ i} \ {k}))) . ipd n \ \ m]\ by simp + also + have \\ = map \ [n <- sorted_list_of_set ({i. k cd\<^bsup>\\<^esup>\ i} \ {k}). ipd (\ n) \ \ m]\ by (auto simp add: filter_map **) + also + have \\ = []\ unfolding * filter_sorted_list_of_set[OF fin, of \\n. ipd (\ n) \ \ m\] by auto + finally + show \?thesis\ using csm by (metis append_Nil) +qed + +lemma converged_ipd_same_icd: assumes path: \is_path \\ \is_path \'\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ +and csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and icd: \l icd\<^bsup>\\<^esup>\ k\ and suc: \\ (Suc k) = \' (Suc k')\ +and ipd: \\' m' = ipd (\ k)\ \\ n \ {k'..' n \ ipd (\ k)\ +shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ +proof cases + assume l: \l = Suc k\ + hence \Suc k cd\<^bsup>\\<^esup>\ k\ using icd by (metis is_icdi_def) + hence \\ (Suc k) \ ipd (\ k)\ unfolding is_cdi_def by auto + hence \\' (Suc k') \ ipd (\' k')\ by (metis csk last_cs suc) + moreover + have \\' (Suc k') \ return\ by (metis \Suc k cd\<^bsup>\\<^esup>\ k\ ret_no_cd suc) + ultimately + have \Suc k' cd\<^bsup>\'\<^esup>\ k'\ unfolding is_cdi_def using path(2) apply auto + by (metis ipd_not_self le_Suc_eq le_antisym path_nodes term_path_stable) + hence \Suc k' icd\<^bsup>\'\<^esup>\ k'\ unfolding is_icdi_def using path(2) by fastforce + hence \cs\<^bsup>\'\<^esup> (Suc k') = cs\<^bsup>\'\<^esup> k' @[\' (Suc k')]\ using icd_cs by auto + moreover + have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ [\ l]\ using icd icd_cs by auto + ultimately + have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> (Suc k')\ by (metis csk l suc) + thus \?thesis\ by blast +next + assume nsuck: \l \ Suc k\ + have kk'[simp]: \\' k' = \ k\ by (metis csk last_cs) + have kl: \k < l\ using icd unfolding is_icdi_def is_cdi_def by auto + hence skl: \Suc k < l\ by (metis Suc_lessI nsuck) + hence lpd: \\ l pd\ \ (Suc k)\ using icd icd_pd_intermediate by auto + have km: \k < m\ by (metis converge(1) kl order.strict_trans) + have lcd: \l cd\<^bsup>\\<^esup>\ k\ using icd is_icdi_def by auto + hence ipdk_pdl: \ipd (\ k) pd\ (\ l)\ by (metis ipd_pd_cd) + have *: \ipd (\ k) \ nodes\ by (metis ipdk_pdl pd_node1) + have nretk: \\ k \ return\ by (metis kl lcd path(1) ret_no_cd term_path_stable less_imp_le) + have **: \\ (\ l) pd\ ipd (\ k)\ proof + assume a: \\ l pd\ ipd (\ k)\ + hence \\ l pd\ (\ k)\ by (metis is_ipd_def \k < l\ ipd_is_ipd ipdk_pdl path(1) path_nodes pd_antisym term_path_stable less_imp_le) + moreover + have \\ l \ (\ k)\ by (metis "*" a ipd_not_self ipdk_pdl lcd pd_antisym ret_no_cd) + ultimately + show \False\ using lcd cd_not_pd by auto + qed + + have km': \k' < m'\ using cs_order[OF path csk converge(2) nretk km] . + + obtain \'' n'' where path'': \is_path \''\ and \''0: \\'' 0 = ipd (\ k)\ and \''n: \\'' n'' = return\ and not\l: \\ i\n''. \'' i \ \ l\ using no_pd_path[OF * **] . + let \?\'\ = \(\' @\<^bsup>m'\<^esup> \'') \ Suc k'\ + have \is_path ?\'\ by (metis \''0 ipd(1) path'' path(2) path_cons path_path_shift) + moreover + have \?\' 0 = \ (Suc k)\ using km' suc by auto + moreover + have \?\' (m' - Suc k' + n'') = return\ using \''n km' \''0 ipd(1) by auto + ultimately + obtain l'' where l'': \l'' \ m' - Suc k' + n''\ \?\' l'' = \ l\ using lpd unfolding is_pd_def by blast + have l''m: \l'' \ m' - Suc k'\ apply (rule ccontr) using l'' not\l km' by (cases \Suc (k' + l'') \ m'\, auto) + let \?l'\ = \Suc ( k' + l'')\ + have lm': \?l' \ m'\ using l''m km' by auto + + \ \Now we have found our desired l'\ + have 1: \\' ?l' = \ l\ using l'' l''m lm' by auto + have 2: \k' < ?l'\ by simp + have 3: \?l' < m'\ apply (rule ccontr) using lm' by (simp, metis "**" 1 ipd(1) ipdk_pdl) + + \ \Need the least such l'\ + + let \?P\ = \\ l'. \' l' = \ l \ k' < l' \ l' < m'\ + + have *: \?P ?l'\ using 1 2 3 by blast + + define l' where l': \l' == LEAST l'. ?P l'\ + + have \l': \\' l' = \ l\ using l' 1 2 3 LeastI[of \?P\] by blast + have kl': \k' < l'\ using l' 1 2 3 LeastI[of \?P\] by blast + have lm': \l' < m'\ using l' 1 2 3 LeastI[of \?P\] by blast + + have nretl': \\' l' \ return\ by (metis \''n \l' le_refl not\l) + + have nipd': \\ j \ {k'..l'}. \' j \ ipd (\' k')\ using lm' kk' ipd(2) kl' by force + + have lcd': \l' cd\<^bsup>\'\<^esup>\ k'\ by (metis is_cdi_def kl' nipd' nretl' path(2)) + + have licd: \l' icd\<^bsup>\'\<^esup>\ k'\ proof - + have \\ m \ {k'<.. l' cd\<^bsup>\'\<^esup>\ m\ proof (rule ccontr) + assume \\ (\ m \ {k'<.. l' cd\<^bsup>\'\<^esup>\ m)\ + then obtain j' where kj': \k' < j'\ and jl': \j' < l'\ and lcdj': \l' cd\<^bsup>\'\<^esup>\ j'\ by force + have jm': \j' by (metis jl' lm' order.strict_trans) + have \\' j' \ \ l\ apply (rule ccontr) using l' kj' jm' jl' Least_le[of \?P\ \j'\] by auto + hence \\ \' l' pd\ \' j'\ using cd_not_pd lcdj' \l' by metis + moreover have \\' j' \ nodes\ using path(2) path_nodes by auto + ultimately + obtain \\<^sub>1 n\<^sub>1 where path\<^sub>1: \is_path \\<^sub>1\ and \0\<^sub>1: \\\<^sub>1 0 = \' j'\ and \n\<^sub>1: \\\<^sub>1 n\<^sub>1 = return\ and nl': \\ l \n\<^sub>1. \\<^sub>1 l \ \' l'\ unfolding is_pd_def by blast + let \?\''\ = \(\'@\<^bsup>j'\<^esup> \\<^sub>1) \ Suc k'\ + have \is_path ?\''\ by (metis \0\<^sub>1 path(2) path\<^sub>1 path_cons path_path_shift) + moreover + have \?\'' 0 = \ (Suc k)\ by (simp, metis kj' less_eq_Suc_le suc) + moreover + have kj': \Suc k' \ j'\ by (metis kj' less_eq_Suc_le) + hence \?\'' (j' - Suc k' + n\<^sub>1) = return\ by (simp, metis \0\<^sub>1 \n\<^sub>1) + ultimately + obtain l'' where *: \?\'' l'' = \ l\ and **: \l'' \j' - Suc k' + n\<^sub>1\ using lpd is_pd_def by blast + show \False\ proof (cases) + assume a: \l'' \ j' - Suc k'\ + hence \\' (l'' + Suc k') = \ l\ using * kj' by(simp, metis Nat.le_diff_conv2 add_Suc diff_add_inverse le_add1 le_add_diff_inverse2) + moreover + have \l'' + Suc k' < l'\ by (metis a jl' add_diff_cancel_right' kj' le_add_diff_inverse less_imp_diff_less ordered_cancel_comm_monoid_diff_class.le_diff_conv2) + moreover + have \l'' + Suc k' < m'\ by (metis Suc_lessD calculation(2) less_trans_Suc lm') + moreover + have \k' < l'' + Suc k'\ by simp + ultimately + show \False\ using Least_le[of \?P\ \l'' + Suc k'\] l' by auto + next + assume a: \\ l'' \ j' - Suc k'\ + hence \\ Suc (k' + l'') \ j'\ by simp + hence \\\<^sub>1 (Suc (k' + l'') - j') = \ l\ using * kj' by simp + moreover + have \Suc (k' + l'') - j' \ n\<^sub>1\ using ** kj' by simp + ultimately + show \False\ using nl' by (metis \l') + qed + qed + thus \?thesis\ unfolding is_icdi_def using lcd' path(2) by simp + qed + hence \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\'\<^esup> k' @ [\' l']\ by (metis icd_cs) + hence \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l\ by (metis \l' csk icd icd_cs) + thus \?thesis\ by metis +qed + +lemma converged_same_icd: assumes path: \is_path \\ \is_path \'\ and converge: \l < n\ \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ +and csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and icd: \l icd\<^bsup>\\<^esup>\ k\ and suc: \\ (Suc k) = \' (Suc k')\ +shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ proof - + + have nret: \\ k \ return\ using icd unfolding is_icdi_def is_cdi_def using term_path_stable less_imp_le by metis + have kl: \k < l\ using icd unfolding is_icdi_def is_cdi_def by auto + have kn: \k < n\ using converge kl by simp + from path_ipd_swap[OF path(1) nret kn] + obtain \ m where path\: \is_path \\ and \\: \\ =\<^bsub>n\<^esub> \\ and km: \k < m\ and ipd: \\ m = ipd (\ k)\ \\ l \ {k.. l \ ipd (\ k)\ . + have csk1: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> k\ using cs_path_swap_le path path\ \\ kn by auto + have suc\: \\ (Suc k) = \ (Suc k)\ by (metis \\ eq_up_to_def kn less_eq_Suc_le) + + have nret': \\' k' \ return\ by (metis csk last_cs nret) + have kn': \k' < n'\ using cs_order[OF path csk converge(2) nret kn] . + from path_ipd_swap[OF path(2) nret' kn'] + obtain \' m' where path\': \is_path \'\ and \\': \\' =\<^bsub>n'\<^esub> \'\ and km': \k' < m'\ and ipd': \\' m' = ipd (\' k')\ \\ l \ {k'..' l \ ipd (\' k')\ . + have csk1': \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\'\<^esup> k'\ using cs_path_swap_le path path\' \\' kn' by auto + have suc\': \\' (Suc k') = \' (Suc k')\ by (metis \\' eq_up_to_def kn' less_eq_Suc_le) + + have icd\: \l icd\<^bsup>\\<^esup>\ k\ using icdi_path_swap_le[OF path\ icd \\] converge by simp + + have lm: \l < m\ using ipd(1) icd\ km unfolding is_icdi_def is_cdi_def by auto + + have csk': \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ using csk1 csk1' csk by auto + + hence kk': \\' k' = \ k\ using last_cs by metis + + have suc': \\ (Suc k) = \' (Suc k')\ using suc suc\ suc\' by auto + + have mm': \\' m' = \ m\ using ipd(1) ipd'(1) kk' by auto + + from cs_ipd[OF ipd km] cs_ipd[OF ipd' km',unfolded mm', folded csk'] + have csm: \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ by metis + + from converged_ipd_same_icd[OF path\ path\' lm csm csk' icd\ suc' ipd'[unfolded kk']] + obtain l' where csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ by blast + + have csl\: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> l\ using \\ converge(1) cs_path_swap_le less_imp_le_nat path(1) path\ by blast + + have nretl: \\ l \ return\ by (metis icd\ icd_imp_cd ret_no_cd) + + have csn': \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ using converge(2) cs_path_swap path path\ path\' \\ \\' by auto + + have ln': \l' < n'\ using cs_order[OF path\ path\' csl csn' nretl converge(1)] . + + have csl\': \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\'\<^esup> l'\ using cs_path_swap_le[OF path(2) path\' \\'] ln' by auto + + have csl': \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using csl\ csl\' csl by auto + thus \?thesis\ by blast +qed + +lemma cd_is_cs_less: assumes \l cd\<^bsup>\\<^esup>\ k\ shows \cs\<^bsup>\\<^esup> k \ cs\<^bsup>\\<^esup> l\ proof - + obtain xs where csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ xs @[\ l]\ using cd_in_cs[OF assms] by blast + hence len: \length(cs\<^bsup>\\<^esup> k) < length (cs\<^bsup>\\<^esup> l)\ by auto + have take: \take (length (cs\<^bsup>\\<^esup> k)) (cs\<^bsup>\\<^esup> l) = cs\<^bsup>\\<^esup> k\ using csl by auto + show \?thesis\ using cs_less.intros[OF len take] . +qed + +lemma cs_select_id: assumes \is_path \\ \\ k \ return\ shows \\\cs\<^bsup>\\<^esup> k = k\ (is \?k = k\) proof - + have *: \\ i . cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> k \ i = k\ using cs_inj[OF assms] by metis + hence \cs\<^bsup>\\<^esup> ?k = cs\<^bsup>\\<^esup> k\ unfolding cs_select_def using theI[of \\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> k\ \k\] by auto + thus \?k = k\ using * by auto +qed + +lemma cs_single_nocd: assumes \cs\<^bsup>\\<^esup> i = [x]\ shows \\ k. \ i cd\<^bsup>\\<^esup>\ k\ proof - + have \\ (\ k. i icd\<^bsup>\\<^esup>\ k)\ apply (rule ccontr) using assms cs_not_nil by auto + hence \\ (\ k. i cd\<^bsup>\\<^esup>\ k)\ by (metis excd_impl_exicd) + thus \?thesis\ by blast +qed + +lemma cs_single_pd_intermed: assumes \is_path \\ \cs\<^bsup>\\<^esup> n = [\ n]\ \k \ n\ shows \\ n pd\ \ k\ proof - + have \\ l. \ n icd\<^bsup>\\<^esup>\ l\ by (metis assms(2) cs_single_nocd icd_imp_cd) + thus \?thesis\ by (metis assms(1) assms(3) no_icd_pd) +qed + + +lemma cs_first_pd: assumes path: \is_path \\ and pd: \\ n pd\ \ 0\ and first: \\ l < n. \ l \ \ n\ shows \cs\<^bsup>\\<^esup> n = [\ n]\ +by (metis cs_cases first first_pd_no_cd icd_imp_cd path pd) + +lemma converged_pd_cs_single: assumes path: \is_path \\ \is_path \'\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ +and \0: \\ 0 = \' 0\ and mpdl: \\ m pd\ \ l\ and csl: \cs\<^bsup>\\<^esup> l = [\ l]\ +shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ proof - + have *: \\ l pd\ \' 0\ using cs_single_pd_intermed[OF path(1) csl] \0[symmetric] by auto + have \m: \\ m = \' m'\ by (metis converge(2) last_cs) + hence **: \\' m' pd\ \ l\ using mpdl by metis + + obtain l' where lm': \l' \ m'\ and \l: \\' l' = \ l\ (is \?P l'\) using path_pd_pd0[OF path(2) ** *] . + + let \?l\ = \(LEAST l'. \' l' = \ l)\ + + have \l': \\' ?l = \ l\ using LeastI[of \?P\,OF \l] . + moreover + have \\ i ' i \ \ l\ using Least_le[of \?P\] by (metis not_less) + hence \\ i ' i \ \' ?l\ using \l' by metis + moreover + have \\' ?l pd\ \' 0\ using * \l' by metis + ultimately + have \cs\<^bsup>\'\<^esup> ?l = [\' ?l]\ using cs_first_pd[OF path(2)] by metis + thus \?thesis\ using csl \l' by metis +qed + +lemma converged_cs_single: assumes path: \is_path \\ \is_path \'\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ +and \0: \\ 0 = \' 0\ and csl: \cs\<^bsup>\\<^esup> l = [\ l]\ +shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ proof cases + assume *: \\ l = return\ + hence \\ m = return\ by (metis converge(1) path(1) term_path_stable less_imp_le) + hence \cs\<^bsup>\\<^esup> m = [return]\ using cs_return by auto + hence \cs\<^bsup>\'\<^esup> m' = [return]\ using converge by simp + moreover + have \cs\<^bsup>\\<^esup> l = [return]\ using * cs_return by auto + ultimately show \?thesis\ by metis +next + assume nret: \\ l \ return\ + have \m: \\ m = \' m'\ by (metis converge(2) last_cs) + + obtain \\<^sub>1 n where path1: \is_path \\<^sub>1\ and upto: \\ =\<^bsub>m\<^esub> \\<^sub>1\ and \n: \\\<^sub>1 n = return\ using path(1) path_swap_ret by blast + + obtain \\<^sub>1' n' where path1': \is_path \\<^sub>1'\ and upto': \\' =\<^bsub>m'\<^esub> \\<^sub>1'\ and \n': \\\<^sub>1' n' = return\ using path(2) path_swap_ret by blast + + have \1l: \\\<^sub>1 l = \ l\ using upto converge(1) by (metis eq_up_to_def nat_less_le) + + have cs1l: \cs\<^bsup>\\<^sub>1\<^esup> l = cs\<^bsup>\\<^esup> l\ using cs_path_swap_le upto path1 path(1) converge(1) by auto + + have csl1: \cs\<^bsup>\\<^sub>1\<^esup> l = [\\<^sub>1 l]\ by (metis \1l cs1l csl) + + have converge1: \cs\<^bsup>\\<^sub>1\<^esup> n = cs\<^bsup>\\<^sub>1'\<^esup> n'\ using \n \n' cs_return by auto + + have ln: \l < n\ using nret \n \1l term_path_stable[OF path1 \n] by (auto, metis linorder_neqE_nat less_imp_le) + + have \01: \\\<^sub>1 0 = \\<^sub>1' 0\ using \0 eq_up_to_apply[OF upto] eq_up_to_apply[OF upto'] by auto + + have pd: \\\<^sub>1 n pd\ \\<^sub>1 l\ using \n by (metis path1 path_nodes return_pd) + + obtain l' where csl: \cs\<^bsup>\\<^sub>1\<^esup> l = cs\<^bsup>\\<^sub>1'\<^esup> l'\ using converged_pd_cs_single[OF path1 path1' ln converge1 \01 pd csl1] by blast + + have cs1m: \cs\<^bsup>\\<^sub>1\<^esup> m = cs\<^bsup>\\<^esup> m\ using cs_path_swap upto path1 path(1) by auto + have cs1m': \cs\<^bsup>\\<^sub>1'\<^esup> m' = cs\<^bsup>\'\<^esup> m'\ using cs_path_swap upto' path1' path(2) by auto + hence converge1: \cs\<^bsup>\\<^sub>1\<^esup> m = cs\<^bsup>\\<^sub>1'\<^esup> m'\ using converge(2) cs1m by metis + + have nret1: \\\<^sub>1 l \ return\ using nret \1l by auto + + have lm': \l' < m'\ using cs_order[OF path1 path1' csl converge1 nret1 converge(1)] . + + have \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^sub>1'\<^esup> l'\ using cs_path_swap_le[OF path(2) path1' upto'] lm' by auto + moreover + have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^sub>1\<^esup> l\ using cs_path_swap_le[OF path(1) path1 upto] converge(1) by auto + ultimately + have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using csl by auto + thus \?thesis\ by blast +qed + +lemma converged_cd_same_suc: assumes path: \is_path \\ \is_path \'\ and init: \\ 0 = \' 0\ +and cd_suc: \\ k k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\\<^esup>\ k \ \ (Suc k) = \' (Suc k')\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ +shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ +using path init cd_suc converge proof (induction \\\ \l\ rule: cs_induct,cases) + case (cs \ l) + assume *: \\k. l icd\<^bsup>\\<^esup>\ k\ + let \?k\ = \THE k. l icd\<^bsup>\\<^esup>\ k\ + have icd: \l icd\<^bsup>\\<^esup>\ ?k\ by (metis "*" icd_is_the_icd) + hence lcdk: \l cd\<^bsup>\\<^esup>\ ?k\ by (metis is_icdi_def) + hence kl: \?k using is_cdi_def by metis + + have \\ j. ?k cd\<^bsup>\\<^esup>\ j \ l cd\<^bsup>\\<^esup>\ j\ using icd cd_trans is_icdi_def by fast + hence suc': \\ j j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j' \ ?k cd\<^bsup>\\<^esup>\ j \ \ (Suc j) = \' (Suc j')\ using cs.prems(4) by blast + + from cs.IH[OF * cs(2) path(2) cs(4) suc'] cs.prems kl + have \\k'. cs\<^bsup>\\<^esup> (THE k. l icd\<^bsup>\\<^esup>\ k) = cs\<^bsup>\'\<^esup> k'\ by (metis Suc_lessD less_trans_Suc) + then obtain k' where csk: \cs\<^bsup>\\<^esup> ?k = cs\<^bsup>\'\<^esup> k'\ by blast + + have suc2: \\ (Suc ?k) = \' (Suc k')\ using cs.prems(4) lcdk csk by auto + + have km: \?k < m\ using kl cs.prems(5) by simp + + from converged_same_icd[OF cs(2) path(2) cs.prems(5) cs.prems(6) csk icd suc2] + show \?case\ . +next + case (cs \ l) + assume \\ (\k. l icd\<^bsup>\\<^esup>\ k)\ + hence \cs\<^bsup>\\<^esup> l = [\ l]\ by auto + with cs converged_cs_single + show \?case\ by metis +qed + +lemma converged_cd_diverge: +assumes path: \is_path \\ \is_path \'\ and init: \\ 0 = \' 0\ and notin: \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ and converge: \l < m\ \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ +obtains k k' where \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ \l cd\<^bsup>\\<^esup>\ k\ \\ (Suc k) \ \' (Suc k')\ +using assms converged_cd_same_suc by blast + + + +lemma converged_cd_same_suc_return: assumes path: \is_path \\ \is_path \'\ and \0: \\ 0 = \' 0\ +and cd_suc: \\ k k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\\<^esup>\ k \ \ (Suc k) = \' (Suc k')\ and ret: \\' n' = return\ +shows \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\proof cases + assume \\ l = return\ + hence \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> n'\ using ret cs_return by presburger + thus \?thesis\ by blast +next + assume nretl: \\ l \ return\ + have \\ l \ nodes\ using path path_nodes by auto + then obtain \l n where ipl: \is_path \l\ and \l: \\ l = \l 0\ and retn: \\l n = return\ and notl: \\ i>0. \l i \ \ l\ by (metis direct_path_return nretl) + hence ip: \is_path (\@\<^bsup>l\<^esup> \l)\ and l: \(\@\<^bsup>l\<^esup> \l) l = \ l\ and retl: \(\@\<^bsup>l\<^esup> \l) (l + n) = return\ and nl: \\ i>l. (\@\<^bsup>l\<^esup> \l) i \ \ l\ using path_cons[OF path(1) ipl \l] by auto + + have \0': \(\@\<^bsup>l\<^esup> \l) 0 = \' 0\ unfolding cs_0 using \l \0 by auto + + have csn: \cs\<^bsup>\@\<^bsup>l\<^esup> \l\<^esup> (l+n) = cs\<^bsup>\'\<^esup> n'\ using ret retl cs_return by metis + + have eql: \(\@\<^bsup>l\<^esup> \l) =\<^bsub>l\<^esub> \\ by (metis path_append_eq_up_to) + + have csl': \cs\<^bsup>\@\<^bsup>l\<^esup> \l\<^esup> l = cs\<^bsup>\\<^esup> l\ using eql cs_path_swap ip path(1) by metis + + have \0 < n\ using nretl[unfolded \l] retn by (metis neq0_conv) + hence ln: \l < l + n\ by simp + + have *: \\ k k'. cs\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup>\ k \ (\ @\<^bsup>l\<^esup> \l) (Suc k) = \' (Suc k')\ proof (rule,rule,rule) + fix k k' assume *: \cs\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup>\ k\ + hence kl: \k < l\ using is_cdi_def by auto + hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k' \ l cd\<^bsup>\\<^esup>\ k\ using eql * cs_path_swap_le[OF ip path(1) eql,of \k\] cdi_path_swap_le[OF path(1) _ eql,of \l\ \k\] by auto + hence \\ (Suc k) = \' (Suc k')\ using cd_suc by blast + then show \(\ @\<^bsup>l\<^esup> \l) (Suc k) = \' (Suc k')\ using cs_path_swap_le[OF ip path(1) eql,of \Suc k\] kl by auto + qed + obtain l' where \cs\<^bsup>\ @\<^bsup>l\<^esup> \l\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using converged_cd_same_suc[OF ip path(2) \0' * ln csn] by blast + moreover + have \cs\<^bsup>\@\<^bsup>l\<^esup> \l\<^esup> l = cs\<^bsup>\\<^esup> l\ using eql by (metis cs_path_swap ip path(1)) + ultimately + show \?thesis\ by metis +qed + +lemma converged_cd_diverge_return: assumes path: \is_path \\ \is_path \'\ and init: \\ 0 = \' 0\ +and notin: \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ and ret: \\' m' = return\ +obtains k k' where \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ \l cd\<^bsup>\\<^esup>\ k\ \\ (Suc k) \ \' (Suc k')\ using converged_cd_same_suc_return[OF path init _ ret, of \l\] notin by blast + +lemma returned_missing_cd_or_loop: assumes path: \is_path \\ \is_path \'\ and \0: \\ 0 = \' 0\ +and notin': \\(\ k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ and nret: \\ n'. \' n' \ return\ and ret: \\ n = return\ +obtains i i' where \i \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ \\ (Suc i) \ \' (Suc i')\ \k cd\<^bsup>\\<^esup>\ i \ (\ j'> i'. j' cd\<^bsup>\'\<^esup>\ i')\ +proof - + obtain f where icdf: \\ i'. f (Suc i') icd\<^bsup>\'\<^esup>\ f i'\ and ran: \range f = {i'. \ j'>i'. j' cd\<^bsup>\'\<^esup>\ i'}\ and icdf0: \\ (\i'. f 0 cd\<^bsup>\'\<^esup>\ i')\ using path(2) path_nret_inf_icd_seq nret by blast + show \thesis\ proof cases + assume \\ j. \ (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j))\ + then obtain j where ni\: \\ (\ i. cs\<^bsup>\'\<^esup> (f j) = cs\<^bsup>\\<^esup> i)\ by metis + note converged_cd_diverge_return[OF path(2,1) \0[symmetric] ni\ ret] that + then obtain i k' where csk: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> k'\ and cdj: \f j cd\<^bsup>\'\<^esup>\ k'\ and div: \\ (Suc i) \ \' (Suc k')\ by metis + have \k' \ range f\ using cdj proof (induction \j\) + case 0 thus \?case\ using icdf0 by blast + next + case (Suc j) + have icdfj: \f (Suc j) icd\<^bsup>\'\<^esup>\ f j\ using icdf by auto + show \?case\ proof cases + assume \f (Suc j) icd\<^bsup>\'\<^esup>\ k'\ + hence \k' = f j\ using icdfj by (metis icd_uniq) + thus \?case\ by auto + next + assume \\ f (Suc j) icd\<^bsup>\'\<^esup>\ k'\ + hence \f j cd\<^bsup>\'\<^esup>\ k'\ using cd_impl_icd_cd[OF Suc.prems icdfj] by auto + thus \?case\ using Suc.IH by auto + qed + qed + hence alldep: \\ i'>k'. i' cd\<^bsup>\'\<^esup>\ k'\ using ran by auto + show \thesis\ proof cases + assume \i < k\ with alldep that[OF _ csk div] show \thesis\ by blast + next + assume \\ i < k\ + hence ki: \k\i\ by auto + have \k \ i\ using notin' csk by auto + hence ki': \k using ki by auto + obtain ka k' where \cs\<^bsup>\\<^esup> ka = cs\<^bsup>\'\<^esup> k'\ \k cd\<^bsup>\\<^esup>\ ka\ \\ (Suc ka) \ \' (Suc k')\ + using converged_cd_diverge[OF path \0 notin' ki' csk] by blast + moreover + hence \ka < k\ unfolding is_cdi_def by auto + ultimately + show \?thesis\ using that by blast + qed + next + assume \\(\ j. \ (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j)))\ + hence allin: \\ j. (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j))\ by blast + define f' where f': \f' \ \ j. (SOME i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> (f j))\ + have \\ i. f' i < f' (Suc i)\ proof + fix i + have csi: \cs\<^bsup>\'\<^esup> (f i) = cs\<^bsup>\\<^esup> (f' i)\ unfolding f' using allin by (metis (mono_tags) someI_ex) + have cssuci: \cs\<^bsup>\'\<^esup> (f (Suc i)) = cs\<^bsup>\\<^esup> (f' (Suc i))\ unfolding f' using allin by (metis (mono_tags) someI_ex) + have fi: \f i < f (Suc i)\ using icdf unfolding is_icdi_def is_cdi_def by auto + have \f (Suc i) cd\<^bsup>\'\<^esup>\ f i\ using icdf unfolding is_icdi_def by blast + hence nreti: \\' (f i) \ return\ by (metis cd_not_ret) + show \f' i < f' (Suc i)\ using cs_order[OF path(2,1) csi cssuci nreti fi] . + qed + hence kle: \k < f' (Suc k)\ using mono_ge_id[of \f'\ \Suc k\] by auto + have cssk: \cs\<^bsup>\\<^esup> (f' (Suc k)) = cs\<^bsup>\'\<^esup> (f (Suc k))\ unfolding f' using allin by (metis (mono_tags) someI_ex) + obtain ka k' where \cs\<^bsup>\\<^esup> ka = cs\<^bsup>\'\<^esup> k'\ \k cd\<^bsup>\\<^esup>\ ka\ \\ (Suc ka) \ \' (Suc k')\ + using converged_cd_diverge[OF path \0 notin' kle cssk] by blast + moreover + hence \ka < k\ unfolding is_cdi_def by auto + ultimately + show \?thesis\ using that by blast + qed +qed + +lemma missing_cd_or_loop: assumes path: \is_path \\ \is_path \'\ and \0: \\ 0 = \' 0\ and notin': \\(\ k'. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ +obtains i i' where \i < k\ \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ \\ (Suc i) \ \' (Suc i')\ \k cd\<^bsup>\\<^esup>\ i \ (\ j'> i'. j' cd\<^bsup>\'\<^esup>\ i')\ +proof cases + assume \\ n'. \' n' = return\ + then obtain n' where retn: \\' n' = return\ by blast + note converged_cd_diverge_return[OF path \0 notin' retn] + then obtain ka k' where \cs\<^bsup>\\<^esup> ka = cs\<^bsup>\'\<^esup> k'\ \k cd\<^bsup>\\<^esup>\ ka\ \\ (Suc ka) \ \' (Suc k')\ by blast + moreover + hence \ka < k\ unfolding is_cdi_def by auto + ultimately show \thesis\ using that by simp +next + assume \\ (\ n'. \' n' = return)\ + hence notret: \\ n'. \' n' \ return\ by auto + then obtain \l n where ipl: \is_path \l\ and \l: \\ k = \l 0\ and retn: \\l n = return\ using reaching_ret path(1) path_nodes by metis + hence ip: \is_path (\@\<^bsup>k\<^esup>\l)\ and l: \(\@\<^bsup>k\<^esup>\l) k = \ k\ and retl: \(\@\<^bsup>k\<^esup>\l) (k + n) = return\ using path_cons[OF path(1) ipl \l] by auto + + have \0': \(\@\<^bsup>k\<^esup>\l) 0 = \' 0\ unfolding cs_0 using \l \0 by auto + + have eql: \(\@\<^bsup>k\<^esup>\l) =\<^bsub>k\<^esub> \\ by (metis path_append_eq_up_to) + + have csl': \cs\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup> k = cs\<^bsup>\\<^esup> k\ using eql cs_path_swap ip path(1) by metis + + hence notin: \\(\ k'. cs\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup> k = cs\<^bsup>\'\<^esup> k')\ using notin' by auto + + obtain i i' where *: \i < k\ and csi: \cs\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup> i = cs\<^bsup>\'\<^esup> i'\ and suci: \(\ @\<^bsup>k\<^esup> \l) (Suc i) \ \' (Suc i')\ and cdloop: \k cd\<^bsup>\@\<^bsup>k\<^esup>\l\<^esup>\ i \ (\ j'>i'. j' cd\<^bsup>\'\<^esup>\ i')\ + using returned_missing_cd_or_loop[OF ip path(2) \0' notin notret retl] by blast + + have \i \ k\ using notin csi by auto + hence ik: \i < k\ using * by auto + hence \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using csi cs_path_swap_le[OF ip path(1) eql] by auto + moreover + have \\ (Suc i) \ \' (Suc i')\ using ik eq_up_to_apply[OF eql, of \Suc i\] suci by auto + moreover + have \k cd\<^bsup>\\<^esup>\ i \ (\ j'>i'. j' cd\<^bsup>\'\<^esup>\ i')\ using cdloop cdi_path_swap_le[OF path(1) _ eql, of \k\ \i\] by auto + ultimately + show \thesis\ using that[OF *] by blast +qed + + +lemma path_shift_set_cd: assumes \is_path \\ shows \{k + j| j . n cd\<^bsup>\\k\<^esup>\ j } = {i. (k+n) cd\<^bsup>\\<^esup>\ i \ k \ i }\ +proof - + { fix i + assume \i\{k+j | j . n cd\<^bsup>\\k\<^esup>\ j }\ + then obtain j where \i = k+j\ \n cd\<^bsup>\\k\<^esup>\ j\ by auto + hence \k+n cd\<^bsup>\\<^esup>\ i \ k \ i\ using cd_path_shift[OF _ assms, of \k\ \k+j\ \k+n\] by simp + hence \i\{ i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i }\ by blast + } + moreover + { fix i + assume \i\{ i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i }\ + hence *: \k+n cd\<^bsup>\\<^esup>\ i \ k \ i\ by blast + then obtain j where i: \i = k+j\ by (metis le_Suc_ex) + hence \k+n cd\<^bsup>\\<^esup>\ k+j\ using * by auto + hence \n cd\<^bsup>\\k\<^esup>\ j\ using cd_path_shift[OF _ assms, of \k\ \k+j\ \k+n\] by simp + hence \i\{k+j | j . n cd\<^bsup>\\k\<^esup>\ j }\ using i by simp + } + ultimately show \?thesis\ by blast +qed + +lemma cs_path_shift_set_cd: assumes path: \is_path \\ shows \cs\<^bsup>\\k\<^esup> n = map \ (sorted_list_of_set {i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i }) @ [\ (k+n)]\ +proof - + have mono:\\n m. n < m \ k + n < k + m\ by auto + have fin: \finite {i. n cd\<^bsup>\ \ k\<^esup>\ i}\ unfolding is_cdi_def by auto + have *: \(\ x. k+x)`{i. n cd\<^bsup>\\k\<^esup>\ i} = {k + i | i. n cd\<^bsup>\\k\<^esup>\ i}\ by auto + have \cs\<^bsup>\\k\<^esup> n = map (\\k) (sorted_list_of_set {i. n cd\<^bsup>\\k\<^esup>\ i}) @ [(\\k) n]\ using cs_sorted_list_of_cd' by blast + also + have \\ = map \ (map (\ x. k+x) (sorted_list_of_set{i. n cd\<^bsup>\\k\<^esup>\ i})) @ [\ (k+n)]\ by auto + also + have \\ = map \ (sorted_list_of_set ((\ x. k+x)`{i. n cd\<^bsup>\\k\<^esup>\ i})) @ [\ (k+n)]\ using sorted_list_of_set_map_mono[OF mono fin] by auto + also + have \\ = map \ (sorted_list_of_set ({k + i | i. n cd\<^bsup>\\k\<^esup>\ i})) @ [\ (k+n)]\ using * by auto + also + have \\ = map \ (sorted_list_of_set ({i. k+n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ (k+n)]\ using path_shift_set_cd[OF path] by auto + finally + show \?thesis\ . +qed + +lemma cs_split_shift_cd: assumes \n cd\<^bsup>\\<^esup>\ j\ and \j < k\ and \k < n\ and \\j'\\<^esup>\ j' \ j' \ j\ shows \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\\k\<^esup> (n-k)\ +proof - + have path: \is_path \\ using assms unfolding is_cdi_def by auto + have 1: \{i. n cd\<^bsup>\\<^esup>\ i} = {i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}\ by auto + have le: \\ i\ {i. n cd\<^bsup>\\<^esup>\ i \ i < k}. \ j\ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}. i < j\ by auto + + have 2: \{i. n cd\<^bsup>\\<^esup>\ i \ i < k} = {i . j cd\<^bsup>\\<^esup>\ i} \ {j}\ proof - + { fix i assume \i\{i. n cd\<^bsup>\\<^esup>\ i \ i < k}\ + hence cd: \n cd\<^bsup>\\<^esup>\ i\ and ik:\i < k\ by auto + have \i\{i . j cd\<^bsup>\\<^esup>\ i} \ {j}\ proof cases + assume \i < j\ hence \j cd\<^bsup>\\<^esup>\ i\ by (metis is_cdi_def assms(1) cd cdi_prefix nat_less_le) + thus \?thesis\ by simp + next + assume \\ i < j\ + moreover + have \i \ j\ using assms(4) ik cd by auto + ultimately + show \?thesis\ by auto + qed + } + moreover + { fix i assume \i\{i . j cd\<^bsup>\\<^esup>\ i} \ {j}\ + hence \j cd\<^bsup>\\<^esup>\ i \ i = j\ by auto + hence \i\{i. n cd\<^bsup>\\<^esup>\ i \ i < k}\ using assms(1,2) cd_trans[OF _ assms(1)] apply auto unfolding is_cdi_def + by (metis (poly_guards_query) diff_diff_cancel diff_is_0_eq le_refl le_trans nat_less_le) + } + ultimately show \?thesis\ by blast + qed + + have \cs\<^bsup>\\<^esup> n = map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i}) @ [\ n]\ using cs_sorted_list_of_cd' by simp + also + have \\ = map \ (sorted_list_of_set ({i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ using 1 by metis + also + have \\ = map \ ((sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ i < k}) @ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ + using sorted_list_of_set_append[OF _ _ le] is_cdi_def by auto + also + have \\ = (map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ i < k})) @ (map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ by auto + also + have \\ = cs\<^bsup>\\<^esup> j @ (map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ unfolding 2 using cs_sorted_list_of_cd by auto + also + have \\ = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\\k\<^esup> (n-k)\ using cs_path_shift_set_cd[OF path, of \k\ \n-k\] assms(3) by auto + finally + show \?thesis\ . +qed + +lemma cs_split_shift_nocd: assumes \is_path \\ and \k < n\ and \\j. n cd\<^bsup>\\<^esup>\ j \ k \ j\ shows \cs\<^bsup>\\<^esup> n = cs\<^bsup>\\k\<^esup> (n-k)\ +proof - + have path: \is_path \\ using assms by auto + have 1: \{i. n cd\<^bsup>\\<^esup>\ i} = {i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}\ by auto + have le: \\ i\ {i. n cd\<^bsup>\\<^esup>\ i \ i < k}. \ j\ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}. i < j\ by auto + have 2: \{i. n cd\<^bsup>\\<^esup>\ i \ i < k} = {}\ using assms by auto + + have \cs\<^bsup>\\<^esup> n = map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i}) @ [\ n]\ using cs_sorted_list_of_cd' by simp + also + have \\ = map \ (sorted_list_of_set ({i. n cd\<^bsup>\\<^esup>\ i \ i < k} \ {i. n cd\<^bsup>\\<^esup>\ i \ k \ i})) @ [\ n]\ using 1 by metis + also + have \\ = map \ (sorted_list_of_set {i. n cd\<^bsup>\\<^esup>\ i \ k \ i}) @ [\ n]\ + unfolding 2 by auto + also + have \\ = cs\<^bsup>\\k\<^esup> (n-k)\ using cs_path_shift_set_cd[OF path, of \k\ \n-k\] assms(2) by auto + finally show \?thesis\ . +qed + +lemma shifted_cs_eq_is_eq: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and \cs\<^bsup>\\k\<^esup> n = cs\<^bsup>\'\k'\<^esup> n'\ shows \cs\<^bsup>\\<^esup> (k+n) = cs\<^bsup>\'\<^esup> (k'+n')\ +proof (rule ccontr) + note path = assms(1,2) + note csk = assms(3) + note csn = assms(4) + assume ne: \cs\<^bsup>\\<^esup> (k+n) \ cs\<^bsup>\'\<^esup> (k'+n')\ + have nretkn:\\ (k+n) \ return\ proof + assume 1:\\ (k+n) = return\ + hence \(\\k) n = return\ by auto + hence \(\'\k') n' = return\ using last_cs assms(4) by metis + hence \\' (k' + n') = return\ by auto + thus \False\ using ne 1 cs_return by auto + qed + hence nretk: \\ k \ return\ using term_path_stable[OF assms(1), of \k\ \k +n\] by auto + have nretkn': \\' (k'+n') \ return\ proof + assume 1:\\' (k'+n') = return\ + hence \(\'\k') n' = return\ by auto + hence \(\\k) n = return\ using last_cs assms(4) by metis + hence \\ (k + n) = return\ by auto + thus \False\ using ne 1 cs_return by auto + qed + hence nretk': \\' k' \ return\ using term_path_stable[OF assms(2), of \k'\ \k' +n'\] by auto + have n0:\n > 0\ proof (rule ccontr) + assume *: \\ 0 < n\ + hence 1:\cs\<^bsup>\\k\<^esup> 0 = cs\<^bsup>\'\k'\<^esup> n'\ using assms(3,4) by auto + have \(\\k) 0 = (\'\k') 0\ using assms(3) last_cs path_shift_def by (metis monoid_add_class.add.right_neutral) + hence \cs\<^bsup>\'\k'\<^esup> 0 = cs\<^bsup>\'\k'\<^esup> n'\ using 1 cs_0 by metis + hence n0': \n' = 0\ using cs_inj[of \\'\k'\ \0\ \n'\ ] * assms(2) by (metis path_shift_def assms(4) last_cs nretkn path_path_shift) + thus \False\ using ne * assms(3) by fastforce + qed + have n0':\n' > 0\ proof (rule ccontr) + assume *: \\ 0 < n'\ + hence 1:\cs\<^bsup>\'\k'\<^esup> 0 = cs\<^bsup>\\k\<^esup> n\ using assms(3,4) by auto + have \(\'\k') 0 = (\\k) 0\ using assms(3) last_cs path_shift_def by (metis monoid_add_class.add.right_neutral) + hence \cs\<^bsup>\\k\<^esup> 0 = cs\<^bsup>\\k\<^esup> n\ using 1 cs_0 by metis + hence n0: \n = 0\ using cs_inj[of \\\k\ \0\ \n\ ] * assms(1) by (metis path_shift_def assms(4) last_cs nretkn path_path_shift) + thus \False\ using ne * assms(3) by fastforce + qed + have cdleswap': \\ j'\'\<^esup>\ j' \ (\j\\<^esup>\ j \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ proof (rule,rule,rule, rule ccontr) + fix j' assume jk': \j' and ncdj': \(k'+n') cd\<^bsup>\'\<^esup>\ j'\ and ne: \\ (\j\\<^esup>\ j \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ + hence kcdj': \k' cd\<^bsup>\'\<^esup>\ j'\ using cr_wn' by blast + + then obtain j where kcdj: \k cd\<^bsup>\\<^esup>\ j\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using csk cs_path_swap_cd path by metis + hence jk: \j < k\ unfolding is_cdi_def by auto + + have ncdn: \\ (k+n) cd\<^bsup>\\<^esup>\ j\ using ne csj jk by blast + + obtain l' where lnocd': \l' = n' \ n' cd\<^bsup>\'\k'\<^esup>\ l'\ and cslsing': \cs\<^bsup>\'\k'\<^esup> l' = [(\'\k') l']\ + proof cases + assume \cs\<^bsup>\'\k'\<^esup> n' = [(\'\k') n']\ thus \thesis\ using that[of \n'\] by auto + next + assume *: \cs\<^bsup>\'\k'\<^esup> n' \ [(\'\k') n']\ + then obtain x ys where \cs\<^bsup>\'\k'\<^esup> n' = [x]@ys@[(\'\k') n']\ by (metis append_Cons append_Nil cs_length_g_one cs_length_one(1) neq_Nil_conv) + then obtain l' where \cs\<^bsup>\'\k'\<^esup> l' = [x]\ and cdl': \n' cd\<^bsup>\'\k'\<^esup>\ l'\ using cs_split[of \\'\k'\ \n'\ \Nil\ \x\ \ys\] by auto + hence \cs\<^bsup>\'\k'\<^esup> l' = [(\'\k') l']\ using last_cs by (metis last.simps) + thus \thesis\ using that cdl' by auto + qed + hence ln': \l'\n'\ unfolding is_cdi_def by auto + hence lcdj': \k'+l' cd\<^bsup>\'\<^esup>\ j'\ using jk' ncdj' by (metis add_le_cancel_left cdi_prefix trans_less_add1) + + obtain l where lnocd: \l = n \ n cd\<^bsup>\\k\<^esup>\ l\ and csl: \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using lnocd' proof + assume \l' = n'\ thus \thesis\ using csn that[of \n\] by auto + next + assume \n' cd\<^bsup>\'\k'\<^esup>\ l'\ + then obtain l where \n cd\<^bsup>\\k\<^esup>\ l\ \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using cs_path_swap_cd path csn by (metis path_path_shift) + thus \thesis\ using that by auto + qed + + have cslsing: \cs\<^bsup>\\k\<^esup> l = [(\\k) l]\ using cslsing' last_cs csl last.simps by metis + + have ln: \l\n\ using lnocd unfolding is_cdi_def by auto + hence nretkl: \\ (k + l) \ return\ using term_path_stable[of \\\ \k+l\ \k+n\] nretkn path(1) by auto + + have *: \n cd\<^bsup>\\k\<^esup>\ l \ k+n cd\<^bsup>\\<^esup>\ k+l\ using cd_path_shift[of \k\ \k+l\ \\\ \k+n\] path(1) by auto + + have ncdl: \\ (k+l) cd\<^bsup>\\<^esup>\ j\ apply rule using lnocd apply rule using ncdn apply blast using cd_trans ncdn * by blast + + hence \\ i\ {j..k+l}. \ i = ipd (\ j)\ unfolding is_cdi_def using path(1) jk nretkl by auto + hence \\ i\ {k<..k+l}. \ i = ipd (\ j)\ using kcdj unfolding is_cdi_def by force + + then obtain i where ki: \k < i\ and il: \i \ k+l\ and ipdi: \\ i = ipd (\ j)\ by force + + hence \(\\k) (i-k) = ipd (\ j)\ \i-k \ l\ by auto + hence pd: \(\\k) l pd\ ipd (\ j)\ using cs_single_pd_intermed[OF _ cslsing] path(1) path_path_shift by metis + moreover + have \(\\k) l = \' (k' + l')\ using csl last_cs by (metis path_shift_def) + moreover + have \\ j = \' j'\ using csj last_cs by metis + ultimately + have \\' (k'+l') pd\ ipd (\' j')\ by simp + moreover + have \ipd (\' j') pd\ \' (k'+l')\ using ipd_pd_cd[OF lcdj'] . + ultimately + have \\' (k'+l') = ipd (\' j')\ using pd_antisym by auto + thus \False\ using lcdj' unfolding is_cdi_def by force + qed + + \ \Symmetric version of the above statement\ + have cdleswap: \\ j\\<^esup>\ j \ (\j'\'\<^esup>\ j' \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ proof (rule,rule,rule, rule ccontr) + fix j assume jk: \j and ncdj: \(k+n) cd\<^bsup>\\<^esup>\ j\ and ne: \\ (\j'\'\<^esup>\ j' \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ + hence kcdj: \k cd\<^bsup>\\<^esup>\ j\ using cr_wn' by blast + + then obtain j' where kcdj': \k' cd\<^bsup>\'\<^esup>\ j'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using csk cs_path_swap_cd path by metis + hence jk': \j' < k'\ unfolding is_cdi_def by auto + + have ncdn': \\ (k'+n') cd\<^bsup>\'\<^esup>\ j'\ using ne csj jk' by blast + + obtain l where lnocd: \l = n \ n cd\<^bsup>\\k\<^esup>\ l\ and cslsing: \cs\<^bsup>\\k\<^esup> l = [(\\k) l]\ + proof cases + assume \cs\<^bsup>\\k\<^esup> n = [(\\k) n]\ thus \thesis\ using that[of \n\] by auto + next + assume *: \cs\<^bsup>\\k\<^esup> n \ [(\\k) n]\ + then obtain x ys where \cs\<^bsup>\\k\<^esup> n = [x]@ys@[(\\k) n]\ by (metis append_Cons append_Nil cs_length_g_one cs_length_one(1) neq_Nil_conv) + then obtain l where \cs\<^bsup>\\k\<^esup> l = [x]\ and cdl: \n cd\<^bsup>\\k\<^esup>\ l\ using cs_split[of \\\k\ \n\ \Nil\ \x\ \ys\] by auto + hence \cs\<^bsup>\\k\<^esup> l = [(\\k) l]\ using last_cs by (metis last.simps) + thus \thesis\ using that cdl by auto + qed + hence ln: \l\n\ unfolding is_cdi_def by auto + hence lcdj: \k+l cd\<^bsup>\\<^esup>\ j\ using jk ncdj by (metis add_le_cancel_left cdi_prefix trans_less_add1) + + obtain l' where lnocd': \l' = n' \ n' cd\<^bsup>\'\k'\<^esup>\ l'\ and csl: \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using lnocd proof + assume \l = n\ thus \thesis\ using csn that[of \n'\] by auto + next + assume \n cd\<^bsup>\\k\<^esup>\ l\ + then obtain l' where \n' cd\<^bsup>\'\k'\<^esup>\ l'\ \cs\<^bsup>\\k\<^esup> l = cs\<^bsup>\'\k'\<^esup> l'\ using cs_path_swap_cd path csn by (metis path_path_shift) + thus \thesis\ using that by auto + qed + + have cslsing': \cs\<^bsup>\'\k'\<^esup> l' = [(\'\k') l']\ using cslsing last_cs csl last.simps by metis + + have ln': \l'\n'\ using lnocd' unfolding is_cdi_def by auto + hence nretkl': \\' (k' + l') \ return\ using term_path_stable[of \\'\ \k'+l'\ \k'+n'\] nretkn' path(2) by auto + + have *: \n' cd\<^bsup>\'\k'\<^esup>\ l' \ k'+n' cd\<^bsup>\'\<^esup>\ k'+l'\ using cd_path_shift[of \k'\ \k'+l'\ \\'\ \k'+n'\] path(2) by auto + + have ncdl': \\ (k'+l') cd\<^bsup>\'\<^esup>\ j'\ apply rule using lnocd' apply rule using ncdn' apply blast using cd_trans ncdn' * by blast + + hence \\ i'\ {j'..k'+l'}. \' i' = ipd (\' j')\ unfolding is_cdi_def using path(2) jk' nretkl' by auto + hence \\ i'\ {k'<..k'+l'}. \' i' = ipd (\' j')\ using kcdj' unfolding is_cdi_def by force + + then obtain i' where ki': \k' < i'\ and il': \i' \ k'+l'\ and ipdi: \\' i' = ipd (\' j')\ by force + + hence \(\'\k') (i'-k') = ipd (\' j')\ \i'-k' \ l'\ by auto + hence pd: \(\'\k') l' pd\ ipd (\' j')\ using cs_single_pd_intermed[OF _ cslsing'] path(2) path_path_shift by metis + moreover + have \(\'\k') l' = \ (k + l)\ using csl last_cs by (metis path_shift_def) + moreover + have \\' j' = \ j\ using csj last_cs by metis + ultimately + have \\ (k+l) pd\ ipd (\ j)\ by simp + moreover + have \ipd (\ j) pd\ \ (k+l)\ using ipd_pd_cd[OF lcdj] . + ultimately + have \\ (k+l) = ipd (\ j)\ using pd_antisym by auto + thus \False\ using lcdj unfolding is_cdi_def by force + qed + + have cdle: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ (is \\ j. ?P j\) proof (rule ccontr) + assume \\ (\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k)\ + hence allge: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ k \ j\ by auto + have allge': \\j'. (k'+n') cd\<^bsup>\'\<^esup>\ j' \ k' \ j'\ proof (rule, rule, rule ccontr) + fix j' + assume *: \k' + n' cd\<^bsup>\'\<^esup>\ j'\ and \\ k' \ j'\ + then obtain j where \j \(k+n) cd\<^bsup>\\<^esup>\ j\ using cdleswap' by (metis le_neq_implies_less nat_le_linear) + thus \False\ using allge by auto + qed + have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_nocd[OF assms(1) _ allge] n0 by auto + moreover + have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_nocd[OF assms(2) _ allge'] n0' by auto + ultimately + show \False\ using ne assms(4) by auto + qed + + define j where \j == GREATEST j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ + have cdj:\(k+n) cd\<^bsup>\\<^esup>\ j\ and jk: \j < k\ and jge:\\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ proof - + have bound: \\ y. ?P y \ y \ k\ by auto + show \(k+n) cd\<^bsup>\\<^esup>\ j\ using GreatestI_nat[of \?P\] j_def bound cdle by blast + show \j < k\ using GreatestI_nat[of \?P\] bound j_def cdle by blast + show \\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ using Greatest_le_nat[of \?P\] bound j_def by blast + qed + + obtain j' where cdj':\(k'+n') cd\<^bsup>\'\<^esup>\ j'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and jk': \j' < k'\ using cdleswap cdj jk by blast + have jge':\\ i'< k'. (k'+n') cd\<^bsup>\'\<^esup>\ i' \ i' \ j'\ proof(rule,rule,rule) + fix i' + assume ik': \i' < k'\ and cdi': \k' + n' cd\<^bsup>\'\<^esup>\ i'\ + then obtain i where cdi:\(k+n) cd\<^bsup>\\<^esup>\ i\ and csi: \ cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ and ik: \i using cdleswap' by force + have ij: \i \ j\ using jge cdi ik by auto + show \i' \ j'\ using cs_order_le[OF assms(1,2) csi[symmetric] csj _ ij] cd_not_ret[OF cdi] by simp + qed + have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_cd[OF cdj jk _ jge] n0 by auto + moreover + have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\'\<^esup> j' @ cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_cd[OF cdj' jk' _ jge'] n0' by auto + ultimately + have \cs\<^bsup>\\<^esup> (k+n) = cs\<^bsup>\'\<^esup> (k'+n')\ using csj assms(4) by auto + thus \False\ using ne by simp +qed + +lemma cs_eq_is_eq_shifted: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and \cs\<^bsup>\\<^esup> (k+n) = cs\<^bsup>\'\<^esup> (k'+n')\ shows \cs\<^bsup>\\k\<^esup> n = cs\<^bsup>\'\k'\<^esup> n'\ +proof (rule ccontr) + assume ne: \cs\<^bsup>\ \ k\<^esup> n \ cs\<^bsup>\' \ k'\<^esup> n'\ + have nretkn:\\ (k+n) \ return\ proof + assume 1:\\ (k+n) = return\ + hence 2:\\' (k'+n') = return\ using assms(4) last_cs by metis + hence \(\\k) n = return\ \(\'\k') n' = return\ using 1 by auto + hence \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_return by metis + thus \False\ using ne by simp + qed + hence nretk: \\ k \ return\ using term_path_stable[OF assms(1), of \k\ \k +n\] by auto + have nretkn': \\' (k'+n') \ return\ proof + assume 1:\\' (k'+n') = return\ + hence 2:\\ (k+n) = return\ using assms(4) last_cs by metis + hence \(\\k) n = return\ \(\'\k') n' = return\ using 1 by auto + hence \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_return by metis + thus \False\ using ne by simp + qed + hence nretk': \\' k' \ return\ using term_path_stable[OF assms(2), of \k'\ \k' +n'\] by auto + have n0:\n > 0\ proof (rule ccontr) + assume *: \\ 0 < n\ + hence \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\'\<^esup> (k'+ n')\ using assms(3,4) by auto + hence n0: \n = 0\ \n' = 0\ using cs_inj[OF assms(2) nretkn', of \k'\] * by auto + have \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ unfolding n0 cs_0 by (auto , metis last_cs assms(3)) + thus \False\ using ne by simp + qed + have n0':\n' > 0\ proof (rule ccontr) + assume *: \\ 0 < n'\ + hence \cs\<^bsup>\\<^esup> k = cs\<^bsup>\\<^esup> (k+ n)\ using assms(3,4) by auto + hence n0: \n = 0\ \n' = 0\ using cs_inj[OF assms(1) nretkn, of \k\] * by auto + have \cs\<^bsup>\ \ k\<^esup> n = cs\<^bsup>\' \ k'\<^esup> n'\ unfolding n0 cs_0 by (auto , metis last_cs assms(3)) + thus \False\ using ne by simp + qed + have cdle: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ (is \\ j. ?P j\) proof (rule ccontr) + assume \\ (\j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k)\ + hence allge: \\j. (k+n) cd\<^bsup>\\<^esup>\ j \ k \ j\ by auto + have allge': \\j'. (k'+n') cd\<^bsup>\'\<^esup>\ j' \ k' \ j'\ proof (rule, rule) + fix j' + assume *: \k' + n' cd\<^bsup>\'\<^esup>\ j'\ + obtain j where cdj: \k+n cd\<^bsup>\\<^esup>\ j\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using cs_path_swap_cd[OF assms(2,1) assms(4)[symmetric] *] by metis + hence kj:\k \ j\ using allge by auto + thus kj': \k' \ j'\ using cs_order_le[OF assms(1,2,3) csj nretk] by simp + qed + have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_nocd[OF assms(1) _ allge] n0 by auto + moreover + have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_nocd[OF assms(2) _ allge'] n0' by auto + ultimately + show \False\ using ne assms(4) by auto + qed + define j where \j == GREATEST j. (k+n) cd\<^bsup>\\<^esup>\ j \ j < k\ + have cdj:\(k+n) cd\<^bsup>\\<^esup>\ j\ and jk: \j < k\ and jge:\\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ proof - + have bound: \\ y. ?P y \ y \ k\ by auto + show \(k+n) cd\<^bsup>\\<^esup>\ j\ using GreatestI_nat[of \?P\] bound j_def cdle by blast + show \j < k\ using GreatestI_nat[of \?P\] bound j_def cdle by blast + show \\ j'< k. (k+n) cd\<^bsup>\\<^esup>\ j' \ j' \ j\ using Greatest_le_nat[of \?P\] bound j_def by blast + qed + obtain j' where cdj':\(k'+n') cd\<^bsup>\'\<^esup>\ j'\ and csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using cs_path_swap_cd assms cdj by blast + have jge':\\ i'< k'. (k'+n') cd\<^bsup>\'\<^esup>\ i' \ i' \ j'\ proof(rule,rule,rule) + fix i' + assume ik': \i' < k'\ and cdi': \k' + n' cd\<^bsup>\'\<^esup>\ i'\ + then obtain i where cdi:\(k+n) cd\<^bsup>\\<^esup>\ i\ and csi: \ cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ using cs_path_swap_cd[OF assms(2,1) assms(4)[symmetric]] by blast + have nreti': \\' i' \ return\ by (metis cd_not_ret cdi') + have ik: \i < k\ using cs_order[OF assms(2,1) csi _ nreti' ik'] assms(3) by auto + have ij: \i \ j\ using jge cdi ik by auto + show \i' \ j'\ using cs_order_le[OF assms(1,2) csi[symmetric] csj _ ij] cd_not_ret[OF cdi] by simp + qed + have jk': \j' < k'\ using cs_order[OF assms(1,2) csj assms(3) cd_not_ret[OF cdj] jk] . + have \cs\<^bsup>\\<^esup> (k + n) = cs\<^bsup>\\<^esup> j @ cs\<^bsup>\ \ k\<^esup> n\ using cs_split_shift_cd[OF cdj jk _ jge] n0 by auto + moreover + have \cs\<^bsup>\'\<^esup> (k' + n') = cs\<^bsup>\'\<^esup> j' @ cs\<^bsup>\' \ k'\<^esup> n'\ using cs_split_shift_cd[OF cdj' jk' _ jge'] n0' by auto + ultimately + have \cs\<^bsup>\\k\<^esup> n = cs\<^bsup>\'\k'\<^esup> n'\ using csj assms(4) by auto + thus \False\ using ne by simp +qed + +lemma converged_cd_diverge_cs: assumes \is_path \\ and \is_path \'\ and \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and \j and \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ and \l < m\ and \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ +obtains k k' where \j\k\ \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and \l cd\<^bsup>\\<^esup>\ k\ and \\ (Suc k) \ \' (Suc k')\ + proof - + have \is_path (\\j)\ \is_path (\'\j')\ using assms(1,2) path_path_shift by auto + moreover + have \(\\j) 0 = (\'\j') 0\ using assms(3) last_cs by (metis path_shift_def add.right_neutral) + moreover + have \\(\l'. cs\<^bsup>\\j\<^esup> (l-j) = cs\<^bsup>\'\j'\<^esup> l')\ proof + assume \\l'. cs\<^bsup>\ \ j\<^esup> (l - j) = cs\<^bsup>\' \ j'\<^esup> l'\ + then obtain l' where csl: \cs\<^bsup>\\j\<^esup> (l - j) = cs\<^bsup>\'\j'\<^esup> l'\ by blast + + have \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> (j' + l')\ using shifted_cs_eq_is_eq[OF assms(1,2,3) csl] assms(4) by auto + thus \False\ using assms(5) by blast + qed + moreover + have \l-j < m-j\ using assms by auto + moreover + have \\ j \ return\ using cs_return assms(1-5) term_path_stable by (metis nat_less_le) + hence \j' using cs_order[OF assms(1,2,3,7)] assms by auto + hence \cs\<^bsup>\\j\<^esup> (m-j) = cs\<^bsup>\'\j'\<^esup> (m'-j')\ using cs_eq_is_eq_shifted[OF assms(1,2,3),of \m-j\ \m'-j'\] assms(4,6,7) by auto + ultimately + obtain k k' where csk: \cs\<^bsup>\\j\<^esup> k = cs\<^bsup>\'\j'\<^esup> k'\ and lcdk: \l-j cd\<^bsup>\\j\<^esup>\ k\ and suc:\(\\j) (Suc k) \ (\'\j') (Suc k')\ using converged_cd_diverge by blast + + have \cs\<^bsup>\\<^esup> (j+k) = cs\<^bsup>\'\<^esup> (j'+k')\ using shifted_cs_eq_is_eq[OF assms(1-3) csk] . + moreover + have \l cd\<^bsup>\\<^esup>\ j+k\ using lcdk assms(1,2,4) by (metis add.commute add_diff_cancel_right' cd_path_shift le_add1) + moreover + have \\ (Suc (j+k)) \ \' (Suc (j'+ k'))\ using suc by auto + moreover + have \j \ j+k\ by auto + ultimately + show \thesis\ using that[of \j+k\ \j'+k'\] by auto +qed + + +lemma cs_ipd_conv: assumes csk: \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ and ipd: \\ l = ipd (\ k)\ \\' l' = ipd(\' k')\ + and nipd: \\n\{k.. n \ ipd (\ k)\ \\n'\{k'..' n' \ ipd (\' k')\ and kl: \k < l\ \k' < l'\ +shows \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ using cs_ipd[OF ipd(1) nipd(1) kl(1)] cs_ipd[OF ipd(2) nipd(2) kl(2)] csk ipd by (metis (no_types) last_cs) + +lemma cp_eq_cs: assumes \((\,k),(\',k'))\cp\ shows \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ + using assms + apply(induction rule: cp.induct) + apply blast+ + apply simp + done + +lemma cd_cs_swap: assumes \l cd\<^bsup>\\<^esup>\ k\ \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ shows \l' cd\<^bsup>\'\<^esup>\ k'\ proof - + have \\ i. l icd\<^bsup>\\<^esup>\ i\ using assms(1) excd_impl_exicd by blast + hence \cs\<^bsup>\\<^esup> l \ [\ l]\ by auto + hence \cs\<^bsup>\'\<^esup> l' \ [\' l']\ using assms last_cs by metis + hence \\ i'. l' icd\<^bsup>\'\<^esup>\ i'\ by (metis cs_cases) + hence path': \is_path \'\ unfolding is_icdi_def is_cdi_def by auto + from cd_in_cs[OF assms(1)] + obtain ys where csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\\<^esup> k @ ys @ [\ l]\ by blast + obtain xs where csk: \cs\<^bsup>\\<^esup> k = xs@[\ k]\ by (metis append_butlast_last_id cs_not_nil last_cs) + have \l: \\ l = \' l'\ using assms last_cs by metis + have csl': \cs\<^bsup>\'\<^esup> l' = xs@[\ k]@ys@[\' l']\ by (metis \l append_eq_appendI assms(2) csk csl) + from cs_split[of \\'\ \l'\ \xs\ \\ k\ \ys\] + obtain m where csm: \cs\<^bsup>\'\<^esup> m = xs @ [\ k]\ and lcdm: \l' cd\<^bsup>\'\<^esup>\ m\ using csl' by metis + have csm': \cs\<^bsup>\'\<^esup> m = cs\<^bsup>\'\<^esup> k'\ by (metis assms(3) csk csm) + have \\' m \ return\ using lcdm unfolding is_cdi_def using term_path_stable by (metis nat_less_le) + hence \m = k'\ using cs_inj path' csm' by auto + thus \?thesis\ using lcdm by auto +qed + + +subsection \Facts about Observations\ +lemma kth_obs_not_none: assumes \is_kth_obs (path \) k i\ obtains a where \obsp \ i = Some a\ using assms unfolding is_kth_obs_def obsp_def by auto + +lemma kth_obs_unique: \is_kth_obs \ k i \ is_kth_obs \ k j \ i = j\ proof (induction \i\ \j\ rule: nat_sym_cases) + case sym thus \?case\ by simp +next + case eq thus \?case\ by simp +next + case (less i j) + have \obs_ids \ \ {.. obs_ids \ \ {.. using less(1) by auto + moreover + have \i \ obs_ids \ \ {.. using less unfolding is_kth_obs_def obs_ids_def by auto + moreover + have \i \ obs_ids \ \ {.. by auto + moreover + have \card (obs_ids \ \ {.. \ {.. using less.prems unfolding is_kth_obs_def by auto + moreover + have \finite (obs_ids \ \ {.. \finite (obs_ids \ \ {.. by auto + ultimately + have \False\ by (metis card_subset_eq) + thus \?case\ .. +qed + +lemma obs_none_no_kth_obs: assumes \obs \ k = None\ shows \\ (\ i. is_kth_obs (path \) k i)\ + apply rule + using assms + unfolding obs_def obsp_def + apply (auto split: option.split_asm) + by (metis assms kth_obs_not_none kth_obs_unique obs_def option.distinct(2) the_equality) + +lemma obs_some_kth_obs : assumes \obs \ k \ None\ obtains i where \is_kth_obs (path \) k i\ by (metis obs_def assms) + +lemma not_none_is_obs: assumes \att(\ i) \ None\ shows \is_kth_obs \ (card (obs_ids \ \ {.. unfolding is_kth_obs_def using assms by auto + +lemma in_obs_ids_is_kth_obs: assumes \i \ obs_ids \\ obtains k where \is_kth_obs \ k i\ proof + have \att (\ i) \ None\ using assms obs_ids_def by auto + thus \is_kth_obs \ (card (obs_ids \ \ {.. using not_none_is_obs by auto +qed + +lemma kth_obs_stable: assumes \is_kth_obs \ l j\ \k < l\ shows \\ i. is_kth_obs \ k i\ using assms proof (induction \l\ arbitrary: \j\ rule: less_induct ) + case (less l j) + have cardl: \card (obs_ids \ \ {.. using less is_kth_obs_def by auto + then obtain i where ex: \i \ obs_ids \ \ {.. (is \?P i\) using less(3) by (metis card.empty empty_iff less_irrefl subsetI subset_antisym zero_diff zero_less_diff) + have bound: \\ i. i \ obs_ids \ \ {.. i \ j\ by auto + let \?i\ = \GREATEST i. i \ obs_ids \ \ {.. + have *: \?i < j\ \?i \ obs_ids \\ using GreatestI_nat[of \?P\ \i\ \j\] ex bound by auto + have **: \\ i. i \ obs_ids \ \ i i \ ?i\ using Greatest_le_nat[of \?P\ _ \j\] ex bound by auto + have \(obs_ids \ \ {.. {?i} = obs_ids \ \ {.. apply rule apply auto using *[simplified] apply simp+ using **[simplified] by auto + moreover + have \?i \ (obs_ids \ \ {.. by auto + ultimately + have \Suc (card (obs_ids \ \ {.. using cardl by (metis Un_empty_right Un_insert_right card_insert_disjoint finite_Int finite_lessThan) + hence \card (obs_ids \ \ {.. by auto + hence iko: \is_kth_obs \ (l - 1) ?i\ using *(2) unfolding is_kth_obs_def obs_ids_def by auto + have ll: \l - 1 < l\ by (metis One_nat_def diff_Suc_less less.prems(2) not_gr0 not_less0) + note IV=less(1)[OF ll iko] + show \?thesis\ proof cases + assume \k < l - 1\ thus \?thesis\ using IV by simp + next + assume \\ k < l - 1\ + hence \k = l - 1\ using less by auto + thus \?thesis\ using iko by blast + qed +qed + +lemma kth_obs_mono: assumes \is_kth_obs \ k i\ \is_kth_obs \ l j\ \k < l\ shows \i < j\ proof (rule ccontr) + assume \\ i < j\ + hence \{.. {.. by auto + hence \obs_ids \ \ {.. obs_ids \ \ {.. by auto + moreover + have \finite (obs_ids \ \ {.. by auto + ultimately + have \card (obs_ids \ \ {.. card (obs_ids \ \ {.. by (metis card_mono) + thus \False\ using assms unfolding is_kth_obs_def by auto +qed + +lemma kth_obs_le_iff: assumes \is_kth_obs \ k i\ \is_kth_obs \ l j\ shows \k < l \ i < j\ by (metis assms kth_obs_unique kth_obs_mono not_less_iff_gr_or_eq) + +lemma ret_obs_all_obs: assumes path: \is_path \\ and iki: \is_kth_obs \ k i\ and ret: \\ i = return\ and kl: \k < l\ obtains j where \is_kth_obs \ l j\ +proof- + show \thesis\ + using kl iki ret proof (induction \l - k\ arbitrary: \k\ \i\ rule: less_induct) + case (less k i) + note kl = \k < l\ + note iki = \is_kth_obs \ k i\ + note ret = \\ i = return\ + have card: \card (obs_ids \ \ {.. and att_ret: \att return \ None\using iki ret unfolding is_kth_obs_def by auto + have rets: \\ (Suc i) = return\ using path ret term_path_stable by auto + hence attsuc: \att (\ (Suc i)) \ None\ using att_ret by auto + hence *: \i \ obs_ids \\ using att_ret ret unfolding obs_ids_def by auto + have \{..< Suc i} = insert i {.. by auto + hence a: \obs_ids \ \ {..< Suc i} = insert i (obs_ids \ \ {.. using * by auto + have b: \i \ obs_ids \ \ {.. by auto + have \finite (obs_ids \ \ {.. by auto + hence \card (obs_ids \ \ {.. by (metis card card_insert_disjoint a b) + hence iksuc: \is_kth_obs \ (Suc k) (Suc i)\ using attsuc unfolding is_kth_obs_def by auto + have suckl: \Suc k \ l\ using kl by auto + note less + thus \thesis\ proof (cases \Suc k < l\) + assume skl: \Suc k < l\ + from less(1)[OF _ skl iksuc rets] skl + show \thesis\ by auto + next + assume \\ Suc k < l\ + hence \Suc k = l\ using suckl by auto + thus \thesis\ using iksuc that by auto + qed + qed +qed + +lemma no_kth_obs_missing_cs: assumes path: \is_path \\ \is_path \'\ and iki: \is_kth_obs \ k i\ and not_in_\': \\(\i'. is_kth_obs \' k i')\ obtains l j where \is_kth_obs \ l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ +proof (rule ccontr) + assume \\ thesis\ + hence all_in_\': \\ l j. is_kth_obs \ l j \ (\ j' . cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ using that by blast + then obtain i' where csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using assms by blast + hence \att(\' i') \ None\ using iki by (metis is_kth_obs_def last_cs) + then obtain k' where ik': \is_kth_obs \' k' i'\ by (metis not_none_is_obs) + hence kk': \k' < k\ using not_in_\' kth_obs_stable by (auto, metis not_less_iff_gr_or_eq) + show \False\ proof (cases \\ i = return\) + assume \\ i \ return\ + thus \False\ using kk' ik' csi iki proof (induction \k\ arbitrary: \i\ \i'\ \k'\ ) + case 0 thus \?case\ by simp + next + case (Suc k i i' k') + then obtain j where ikj: \is_kth_obs \ k j\ by (metis kth_obs_stable lessI) + then obtain j' where csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using all_in_\' by blast + hence \att(\' j') \ None\ using ikj by (metis is_kth_obs_def last_cs) + then obtain k2 where ik2: \is_kth_obs \' k2 j'\ by (metis not_none_is_obs) + have ji: \j < i\ using kth_obs_mono [OF ikj \is_kth_obs \ (Suc k) i\] by auto + hence nretj: \\ j \ return\ using Suc(2) term_path_stable less_imp_le path(1) by metis + have ji': \j' < i'\ using cs_order[OF path _ _ nretj, of \j'\ \i\ \i'\] csj \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ ji by auto + have \k2 \ k'\ using ik2 Suc(4) ji' kth_obs_unique[of \\'\ \k'\ \i'\ \j'\] by (metis less_irrefl) + hence k2k': \k2 < k'\ using kth_obs_mono[OF \is_kth_obs \' k' i'\ ik2] ji' by (metis not_less_iff_gr_or_eq) + hence k2k: \k2 < k\ using Suc by auto + from Suc.IH[OF nretj k2k ik2 csj ikj] show \False\ . + qed + next + assume \\ i = return\ + hence reti': \\' i' = return\ by (metis csi last_cs) + from ret_obs_all_obs[OF path(2) ik' reti' kk', of \False\] not_in_\' + show \False\ by blast + qed +qed + +lemma kth_obs_cs_missing_cs: assumes path: \is_path \\ \is_path \'\ and iki: \is_kth_obs \ k i\ and iki': \is_kth_obs \' k i'\ and csi: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ +obtains l j where \j \ i\ \is_kth_obs \ l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ | l' j' where \j' \ i'\ \is_kth_obs \' l' j'\ \\ (\ j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ +proof (rule ccontr) + assume nt: \\ thesis\ + show \False\ using iki iki' csi that proof (induction \k\ arbitrary: \i\ \i'\ rule: less_induct) + case (less k i i') + hence all_in_\': \\ l j. j\i \ is_kth_obs \ l j \ (\ j' . cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ + and all_in_\: \\ l' j'. j' \ i' \ is_kth_obs \' l' j' \ (\ j . cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ by (metis nt) (metis nt less(6)) + obtain j j' where csji: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> i'\ and csij: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> j'\ using all_in_\ all_in_\' less by blast + then obtain l l' where ilj: \is_kth_obs \ l j\ and ilj': \is_kth_obs \' l' j'\ by (metis is_kth_obs_def last_cs less.prems(1,2)) + have lnk: \l \ k\ using ilj csji less(2) less(4) kth_obs_unique by auto + have lnk': \l' \ k\ using ilj' csij less(3) less(4) kth_obs_unique by auto + have cseq: \\ l j j'. l < k \ is_kth_obs \ l j \ is_kth_obs \' l j' \ cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ proof - + { fix t p p' assume tk: \t < k\ and ikp: \is_kth_obs \ t p\ and ikp': \is_kth_obs \' t p'\ + hence pi: \p < i\ and pi': \p' < i'\ by (metis kth_obs_mono less.prems(1)) (metis kth_obs_mono less.prems(2) tk ikp') + have *: \\j l. j \ p \ is_kth_obs \ l j \ \j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using pi all_in_\' by auto + have **: \\j' l'. j' \ p' \ is_kth_obs \' l' j' \ \j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ using pi' all_in_\ by auto + have \cs\<^bsup>\\<^esup> p = cs\<^bsup>\'\<^esup> p'\ apply(rule ccontr) using less(1)[OF tk ikp ikp'] * ** by blast + } + thus \?thesis\ by blast + qed + have ii'nret: \\ i \ return \ \' i' \ return\ using less cs_return by auto + have a: \k < l \ k < l'\ proof (rule ccontr) + assume \\(k < l \ k < l')\ + hence *: \l < k\ \l' < k\ using lnk lnk' by auto + hence ji: \j < i\ and ji': \j' < i'\ using ilj ilj' less(2,3) kth_obs_mono by auto + show \False\ using ii'nret proof + assume nreti: \\ i \ return\ + hence nretj': \\' j' \ return\ using last_cs csij by metis + show \False\ using cs_order[OF path(2,1) csij[symmetric] csji[symmetric] nretj' ji'] ji by simp + next + assume nreti': \\' i' \ return\ + hence nretj': \\ j \ return\ using last_cs csji by metis + show \False\ using cs_order[OF path csji csij nretj' ji] ji' by simp + qed + qed + have \l < k \ l' < k\ proof (rule ccontr) + assume \\ (l< k \ l' < k)\ + hence \k < l\ \k < l'\ using lnk lnk' by auto + hence ji: \i < j\ and ji': \i' < j'\ using ilj ilj' less(2,3) kth_obs_mono by auto + show \False\ using ii'nret proof + assume nreti: \\ i \ return\ + show \False\ using cs_order[OF path csij csji nreti ji] ji' by simp + next + assume nreti': \\' i' \ return\ + show \False\ using cs_order[OF path(2,1) csji[symmetric] csij[symmetric] nreti' ji'] ji by simp + qed + qed + hence \k < l \ l' < k \ k < l' \ l < k\ using a by auto + thus \False\ proof + assume \k < l \ l' < k\ + hence kl: \k < l\ and lk': \l' < k\ by auto + hence ij: \i < j\ and ji': \j' < i'\ using less(2,3) ilj ilj' kth_obs_mono by auto + have nreti: \\ i \ return\ by (metis csji ii'nret ij last_cs path(1) term_path_stable less_imp_le) + obtain h where ilh: \is_kth_obs \ l' h\ using ji' all_in_\ ilj' no_kth_obs_missing_cs path(1) path(2) by (metis kl lk' ilj kth_obs_stable) + hence \cs\<^bsup>\\<^esup> h = cs\<^bsup>\'\<^esup> j'\ using cseq lk' ilj' by blast + hence \cs\<^bsup>\\<^esup> i = cs\<^bsup>\\<^esup> h\ using csij by auto + hence hi: \h = i\ using cs_inj nreti path(1) by metis + have \l' = k\ using less(2) ilh unfolding hi by (metis is_kth_obs_def) + thus \False\ using lk' by simp + next + assume \k < l' \ l < k\ + hence kl': \k < l'\ and lk: \l < k\ by auto + hence ij': \i' < j'\ and ji: \j < i\ using less(2,3) ilj ilj' kth_obs_mono by auto + have nreti': \\' i' \ return\ by (metis csij ii'nret ij' last_cs path(2) term_path_stable less_imp_le) + obtain h' where ilh': \is_kth_obs \' l h'\ using all_in_\' ilj no_kth_obs_missing_cs path(1) path(2) kl' lk ilj' kth_obs_stable by metis + hence \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> h'\ using cseq lk ilj by blast + hence \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\'\<^esup> h'\ using csji by auto + hence hi: \h' = i'\ using cs_inj nreti' path(2) by metis + have \l = k\ using less(3) ilh' unfolding hi by (metis is_kth_obs_def) + thus \False\ using lk by simp + qed + qed +qed + + +subsection \Facts about Data\ + +lemma reads_restrict1: \\ \ (reads n) = \' \ (reads n) \ \ x \ reads n. \ x = \' x\ by (metis restrict_def) + +lemma reads_restrict2: \\ x \ reads n. \ x = \' x \ \ \ (reads n) = \' \ (reads n)\ unfolding restrict_def by auto + +lemma reads_restrict: \(\ \ (reads n) = \' \ (reads n)) = (\ x \ reads n. \ x = \' x)\ using reads_restrict1 reads_restrict2 by metis + +lemma reads_restr_suc: \\ \ (reads n) = \' \ (reads n) \ suc n \ = suc n \'\ by (metis reads_restrict uses_suc) + +lemma reads_restr_sem: \\ \ (reads n) = \' \ (reads n) \ \ v \ writes n. sem n \ v = sem n \' v\ by (metis reads_restrict1 uses_writes) + +lemma reads_obsp: assumes \path \ k = path \' k'\ \\\<^bsup>k\<^esup> \ (reads (path \ k)) = \'\<^bsup>k'\<^esup> \ (reads (path \ k))\ shows \obsp \ k = obsp \' k'\ + using assms(2) uses_att + unfolding obsp_def assms(1) reads_restrict + apply (cases \att (path \' k')\) + by auto + +lemma no_writes_unchanged0: assumes \\ l writes(path \ l)\ shows \(\\<^bsup>k\<^esup>) v = \ v\ using assms +proof (induction \k\) + case 0 thus \?case\ by(auto simp add: kth_state_def) +next + case (Suc k) + hence \(\\<^bsup>k\<^esup>) v = \ v\ by auto + moreover + have \\\<^bsup>Suc k\<^esup> = snd ( step (path \ k,\\<^bsup>k\<^esup>))\ by (metis kth_state_suc) + hence \\\<^bsup>Suc k\<^esup> = sem (path \ k) (\\<^bsup>k\<^esup>)\ by (metis step_suc_sem snd_conv) + moreover + have \v \ writes (path \ k)\ using Suc.prems by blast + ultimately + show \?case\ using writes by metis +qed + +lemma written_read_dd: assumes \is_path \\ \v \ reads (\ k) \ \v \ writes (\ j)\ \j obtains l where \k dd\<^bsup>\,v\<^esup>\ l\ +proof - + let \?l\ = \GREATEST l. l < k \ v \ writes (\ l)\ + have \?l < k\ by (metis (no_types, lifting) GreatestI_ex_nat assms(3) assms(4) less_or_eq_imp_le) + moreover + have \v \ writes (\ ?l)\ by (metis (no_types, lifting) GreatestI_nat assms(3) assms(4) less_or_eq_imp_le) + hence \v \ reads (\ k) \ writes (\ ?l)\ using assms(2) by auto + moreover + note is_ddi_def + have \\ l \ {?l<.. writes (\ l)\ by (auto, metis (lifting, no_types) Greatest_le_nat le_antisym nat_less_le) + ultimately + have \k dd\<^bsup>\,v\<^esup>\ ?l\ using assms(1) unfolding is_ddi_def by blast + thus \thesis\ using that by simp +qed + +lemma no_writes_unchanged: assumes \k \ l\ \\ j \ {k.. writes(path \ j)\ shows \(\\<^bsup>l\<^esup>) v = (\\<^bsup>k\<^esup>) v\ using assms +proof (induction \l - k\ arbitrary: \l\) + case 0 thus \?case\ by(auto) +next + case (Suc lk l) + hence kl: \k < l\ by auto + then obtain l' where lsuc: \l = Suc l'\ using lessE by blast + hence \lk = l' - k\ using Suc by auto + moreover + have \\ j \ {k.. writes (path \ j)\ using Suc(4) lsuc by auto + ultimately + have \(\\<^bsup>l'\<^esup>) v = (\\<^bsup>k\<^esup>) v\ using Suc(1)[of \l'\] lsuc kl by fastforce + moreover + have \\\<^bsup>l\<^esup> = snd ( step (path \ l',\\<^bsup>l'\<^esup>))\ by (metis kth_state_suc lsuc) + hence \\\<^bsup>l\<^esup> = sem (path \ l') (\\<^bsup>l'\<^esup>)\ by (metis step_suc_sem snd_conv) + moreover + have \l' < l\ \k \ l'\ using kl lsuc by auto + hence \v \ writes (path \ l')\ using Suc.prems(2) by auto + ultimately + show \?case\ using writes by metis +qed + +lemma ddi_value: assumes \l dd\<^bsup>(path \),v\<^esup>\ k\ shows \(\\<^bsup>l\<^esup>) v = (\\<^bsup>Suc k\<^esup> ) v\ +using assms no_writes_unchanged[of \Suc k\ \l\ \v\ \\\] unfolding is_ddi_def by auto + +lemma written_value: assumes \path \ l = path \' l'\ \\\<^bsup>l\<^esup> \ reads (path \ l) = \'\<^bsup>l'\<^esup> \ reads (path \ l)\ \v \ writes (path \ l)\ +shows \(\\<^bsup>Suc l\<^esup> ) v = (\'\<^bsup>Suc l'\<^esup> ) v\ +by (metis assms reads_restr_sem snd_conv step_suc_sem kth_state_suc) + + +subsection \Facts about Contradicting Paths\ + +lemma obsp_contradict: assumes csk: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ and obs: \obsp \ k \ obsp \' k'\ shows \(\', k') \ (\, k)\ +proof - + have pk: \path \ k = path \' k'\ using assms last_cs by metis + hence \\\<^bsup>k\<^esup>\(reads (path \ k)) \ \'\<^bsup>k'\<^esup>\(reads (path \ k))\ using obs reads_obsp[OF pk] by auto + thus \(\',k') \ (\,k)\ using contradicts.intros(2)[OF csk[symmetric]] by auto +qed + +lemma missing_cs_contradicts: assumes notin: \\(\ k'. cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k')\ and converge: \k \cs\<^bsup>path \\<^esup> n = cs\<^bsup>path \'\<^esup> n'\ shows \\ j'. (\', j') \ (\, k)\ +proof - + let \?\\ = \path \\ + let \?\'\ = \path \'\ + have init: \?\ 0 = ?\' 0\ unfolding path_def by auto + have path: \is_path ?\\ \is_path ?\'\ using path_is_path by auto + obtain j j' where csj: \cs\<^bsup>?\\<^esup> j = cs\<^bsup>?\'\<^esup> j'\ and cd: \k cd\<^bsup>?\\<^esup>\j\ and suc: \?\ (Suc j) \ ?\' (Suc j')\ using converged_cd_diverge[OF path init notin converge] . + have less: \cs\<^bsup>?\\<^esup> j \ cs\<^bsup>?\\<^esup> k\ using cd cd_is_cs_less by auto + have nretj: \?\ j \ return\ by (metis cd is_cdi_def term_path_stable less_imp_le) + have cs: \?\ \ cs\<^bsup>?\'\<^esup> j' = j\ using csj cs_select_id nretj path_is_path by metis + have \(\',j') \ (\,k)\ using contradicts.intros(1)[of \?\'\ \j'\ \?\\ \k\ \\\ \\'\,unfolded cs] less suc csj by metis + thus \?thesis\ by blast +qed + +theorem obs_neq_contradicts_term: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ assumes ret: \\ n = return\ \\' n' = return\ and obsne: \obs \ \ obs \'\ +shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ +proof - + have path: \is_path \\ \is_path \'\ using \ \' path_is_path by auto + obtain k1 where neq: \obs \ k1 \ obs \' k1\ using obsne ext[of \obs \\ \obs \'\] by blast + hence \(\k i i'. is_kth_obs \ k i \ is_kth_obs \' k i' \ obsp \ i \ obsp \' i' \ cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i') + \ (\ k i. is_kth_obs \ k i \ \ (\ i'. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')) + \ (\ k i'. is_kth_obs \' k i' \ \ (\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))\ + proof(cases rule: option_neq_cases) + case (none2 x) + have notin\': \\ (\ l. is_kth_obs \' k1 l)\ using none2(2) \' obs_none_no_kth_obs by auto + obtain i where in\: \is_kth_obs \ k1 i\ using obs_some_kth_obs[of \\\ \k1\] none2(1) \ by auto + obtain l j where \is_kth_obs \ l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ using path in\ notin\' by (metis no_kth_obs_missing_cs) + thus \?thesis\ by blast + next + case (none1 x) + have notin\: \\ (\ l. is_kth_obs \ k1 l)\ using none1(1) \ obs_none_no_kth_obs by auto + obtain i' where in\': \is_kth_obs \' k1 i'\ using obs_some_kth_obs[of \\'\ \k1\] none1(2) \' by auto + obtain l j where \is_kth_obs \' l j\ \\ (\ j'. cs\<^bsup>\\<^esup> j' = cs\<^bsup>\'\<^esup> j)\ using path in\' notin\ by (metis no_kth_obs_missing_cs) + thus \?thesis\ by blast + next + case (some x y) + obtain i where in\: \is_kth_obs \ k1 i\ using obs_some_kth_obs[of \\\ \k1\] some \ by auto + obtain i' where in\': \is_kth_obs \' k1 i'\ using obs_some_kth_obs[of \\'\ \k1\] some \' by auto + show \?thesis\ proof (cases) + assume *: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ + have \obsp \ i = obs \ k1\ by (metis obs_def \ in\ kth_obs_unique the_equality) + moreover + have \obsp \' i' = obs \' k1\ by (metis obs_def \' in\' kth_obs_unique the_equality) + ultimately + have \obsp \ i \ obsp \' i'\ using neq by auto + thus \?thesis\ using * in\ in\' by blast + next + assume *: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ + note kth_obs_cs_missing_cs[OF path in\ in\' *] + thus \?thesis\ by metis + qed + qed + thus \?thesis\ proof (cases rule: three_cases) + case 1 + then obtain k i i' where iki: \is_kth_obs \ k i\ \is_kth_obs \' k i'\ and obsne: \obsp \ i \ obsp \' i'\ and csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ by auto + note obsp_contradict[OF csi[unfolded \ \'] obsne] + moreover + have \\ i \ dom att\ using iki unfolding is_kth_obs_def by auto + ultimately + show \?thesis\ by blast + next + case 2 + then obtain k i where iki: \is_kth_obs \ k i\ and notin\': \\ (\i'. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')\ by auto + let \?n\ = \Suc (max n i)\ + have nn: \n < ?n\ by auto + have iln: \i < ?n\ by auto + have retn: \\ ?n = return\ using ret term_path_stable path by auto + hence \cs\<^bsup>\\<^esup> ?n = cs\<^bsup>\'\<^esup> n'\ using ret(2) cs_return by auto + then obtain i' where \(\',i') \ (\,i)\ using missing_cs_contradicts[OF notin\'[unfolded \ \'] iln] \ \' by auto + moreover + have \\ i \ dom att\ using iki is_kth_obs_def by auto + ultimately + show \?thesis\ by blast + next + case 3 + then obtain k i' where iki: \is_kth_obs \' k i'\ and notin\': \\ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')\ by auto + let \?n\ = \Suc (max n' i')\ + have nn: \n' < ?n\ by auto + have iln: \i' < ?n\ by auto + have retn: \\' ?n = return\ using ret term_path_stable path by auto + hence \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> ?n\ using ret(1) cs_return by auto + then obtain i where \(\,i) \ (\',i')\ using missing_cs_contradicts notin\' iln \ \' by metis + moreover + have \\' i' \ dom att\ using iki is_kth_obs_def by auto + ultimately + show \?thesis\ by blast + qed +qed + +lemma obs_neq_some_contradicts': fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ +assumes obsnecs: \obsp \ i \ obsp \' i' \ cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ +and iki: \is_kth_obs \ k i\ and iki': \is_kth_obs \' k i'\ +shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom att) \ ((\, k) \ (\' ,k') \ \' k' \ dom att)\ +using obsnecs iki iki' proof (induction \k\ arbitrary: \i\ \i'\ rule: less_induct ) + case (less k i i') + note iki = \is_kth_obs \ k i\ + and iki' = \is_kth_obs \' k i'\ + have domi: \\ i \ dom att\ by (metis is_kth_obs_def domIff iki) + have domi': \\' i' \ dom att\ by (metis is_kth_obs_def domIff iki') + note obsnecs = \obsp \ i \ obsp \' i' \ cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ + show \?thesis\ proof cases + assume csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ + hence *: \obsp \ i \ obsp \' i'\ using obsnecs by auto + note obsp_contradict[OF _ *] csi domi \ \' + thus \?thesis\ by blast + next + assume ncsi: \cs\<^bsup>\\<^esup> i \ cs\<^bsup>\'\<^esup> i'\ + have path: \is_path \\ \is_path \'\ using \ \' path_is_path by auto + have \0: \\ 0 = \' 0\ unfolding \ \' path_def by auto + note kth_obs_cs_missing_cs[of \\\ \\'\ \k\ \i\ \i'\] \ \' path_is_path iki iki' ncsi + hence \(\ l j .j \ i \ is_kth_obs \ l j \ \ (\ j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')) \ (\ l' j'. j' \ i' \ is_kth_obs \' l' j' \ \ (\ j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'))\ by metis + thus \?thesis\ proof + assume \\l j. j \ i \ is_kth_obs \ l j \ \ (\j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ + then obtain l j where ji: \j\i\ and iobs: \is_kth_obs \ l j\ and notin: \\ (\j'. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ by blast + have dom: \\ j \ dom att\ using iobs is_kth_obs_def by auto + obtain n n' where nj: \n < j\ and csn: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and sucn: \\ (Suc n) \ \' (Suc n')\ and cdloop: \j cd\<^bsup>\\<^esup>\ n \ (\ j'> n'. j' cd\<^bsup>\'\<^esup>\ n')\ + using missing_cd_or_loop[OF path \0 notin] by blast + show \?thesis\ using cdloop proof + assume cdjn: \j cd\<^bsup>\\<^esup>\ n\ + hence csnj: \cs\<^bsup>\'\<^esup> n' \ cs\<^bsup>\\<^esup> j\ using csn by (metis cd_is_cs_less) + have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> n')) = \ (Suc n)\ using csn by (metis cdjn cd_not_ret cs_select_id path(1)) + have \(\',n') \ (\,j)\ using csnj apply(rule contradicts.intros(1)) using cssel \ \' sucn by auto + thus \?thesis\ using dom by auto + next + assume loop: \\ j'>n'. j' cd\<^bsup>\'\<^esup>\ n'\ + show \?thesis\ proof cases + assume in': \i' \ n'\ + have nreti': \\' i' \ return\ by( metis le_eq_less_or_eq lessI loop not_le path(2) ret_no_cd term_path_stable) + show \?thesis\ proof cases + assume \\ \. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> \\ + then obtain \ where cs\: \cs\<^bsup>\\<^esup> \ = cs\<^bsup>\'\<^esup> i'\ by metis + have \n: \\ \ n\ using cs_order_le[OF path(2,1) cs\[symmetric] csn[symmetric] nreti' in'] . + hence \i: \\ < i\ using nj ji by auto + have dom\: \\ \ \ dom att\ using domi' cs\ last_cs by metis + obtain \ where i\\: \is_kth_obs \ \ \\ using dom\ by (metis is_kth_obs_def domIff) + hence \k: \\ < k\ using \i iki by (metis kth_obs_le_iff) + obtain \' where i\\': \is_kth_obs \' \ \'\ using \k iki' by (metis kth_obs_stable) + have \\' < i'\ using \k iki' i\\' by (metis kth_obs_le_iff) + hence cs\': \cs\<^bsup>\\<^esup> \ \ cs\<^bsup>\'\<^esup> \'\ unfolding cs\ using cs_inj[OF path(2) nreti', of \\'\] by blast + thus \?thesis\ using less(1)[OF \k _ i\\ i\\'] by auto + next + assume notin'': \\(\ \. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> \)\ + obtain \ \' where \i': \\' < i'\ and cs\: \cs\<^bsup>\\<^esup> \ = cs\<^bsup>\'\<^esup> \'\ and suc\: \\ (Suc \) \ \' (Suc \')\ and cdloop': \i' cd\<^bsup>\'\<^esup>\ \' \ (\ j>\. j cd\<^bsup>\\<^esup>\ \)\ + using missing_cd_or_loop[OF path(2,1) \0[symmetric] notin''] by metis + show \?thesis\ using cdloop' proof + assume cdjn: \i' cd\<^bsup>\'\<^esup>\ \'\ + hence csnj: \cs\<^bsup>\\<^esup> \ \ cs\<^bsup>\'\<^esup> i'\ using cs\ by (metis cd_is_cs_less) + have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> \)) = \' (Suc \')\ using cs\ by (metis cdjn cd_not_ret cs_select_id path(2)) + have \(\,\) \ (\',i')\ using csnj apply(rule contradicts.intros(1)) using cssel \ \' suc\ by auto + thus \?thesis\ using domi' by auto + next + assume loop': \\ j>\. j cd\<^bsup>\\<^esup>\ \\ + have \n': \\' < n'\ using in' \i' by auto + have nret\': \\' \' \ return\ by (metis cs\ last_cs le_eq_less_or_eq lessI path(1) path(2) suc\ term_path_stable) + have \\ < n\ using cs_order[OF path(2,1) cs\[symmetric] csn[symmetric] nret\' \n'] . + hence \\ < i\ using nj ji by auto + hence cdi\: \i cd\<^bsup>\\<^esup>\ \\ using loop' by auto + hence cs\i: \cs\<^bsup>\'\<^esup> \' \ cs\<^bsup>\\<^esup> i\ using cs\ by (metis cd_is_cs_less) + have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> \')) = \ (Suc \)\ using cs\ by (metis cdi\ cd_not_ret cs_select_id path(1)) + have \(\',\') \ (\,i)\ using cs\i apply(rule contradicts.intros(1)) using cssel \ \' suc\ by auto + thus \?thesis\ using domi by auto + qed + qed + next + assume \\ i' \ n'\ + hence ni': \n'< i'\ by simp + hence cdin: \i' cd\<^bsup>\'\<^esup>\ n'\ using loop by auto + hence csni: \cs\<^bsup>\\<^esup> n \ cs\<^bsup>\'\<^esup> i'\ using csn by (metis cd_is_cs_less) + have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> n)) = \' (Suc n')\ using csn by (metis cdin cd_not_ret cs_select_id path(2)) + have \(\,n) \ (\',i')\ using csni apply(rule contradicts.intros(1)) using cssel \ \' sucn by auto + thus \?thesis\ using domi' by auto + qed + qed + next + \ \Symmetric case as above, indices might be messy.\ + assume \\l j. j \ i' \ is_kth_obs \' l j \ \ (\j'. cs\<^bsup>\\<^esup> j' = cs\<^bsup>\'\<^esup> j)\ + then obtain l j where ji': \j\i'\ and iobs: \is_kth_obs \' l j\ and notin: \\ (\j'. cs\<^bsup>\'\<^esup> j = cs\<^bsup>\\<^esup> j')\ by metis + have dom: \\' j \ dom att\ using iobs is_kth_obs_def by auto + obtain n n' where nj: \n < j\ and csn: \cs\<^bsup>\'\<^esup> n = cs\<^bsup>\\<^esup> n'\ and sucn: \\' (Suc n) \ \ (Suc n')\ and cdloop: \j cd\<^bsup>\'\<^esup>\ n \ (\ j'> n'. j' cd\<^bsup>\\<^esup>\ n')\ + using missing_cd_or_loop[OF path(2,1) \0[symmetric] ] notin by metis + show \?thesis\ using cdloop proof + assume cdjn: \j cd\<^bsup>\'\<^esup>\ n\ + hence csnj: \cs\<^bsup>\\<^esup> n' \ cs\<^bsup>\'\<^esup> j\ using csn by (metis cd_is_cs_less) + have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> n')) = \' (Suc n)\ using csn by (metis cdjn cd_not_ret cs_select_id path(2)) + have \(\,n') \ (\',j)\ using csnj apply(rule contradicts.intros(1)) using cssel \' \ sucn by auto + thus \?thesis\ using dom by auto + next + assume loop: \\ j'>n'. j' cd\<^bsup>\\<^esup>\ n'\ + show \?thesis\ proof cases + assume in': \i \ n'\ + have nreti: \\ i \ return\ by (metis le_eq_less_or_eq lessI loop not_le path(1) ret_no_cd term_path_stable) + show \?thesis\ proof cases + assume \\ \. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> \\ + then obtain \ where cs\: \cs\<^bsup>\'\<^esup> \ = cs\<^bsup>\\<^esup> i\ by metis + have \n: \\ \ n\ using cs_order_le[OF path cs\[symmetric] csn[symmetric] nreti in'] . + hence \i': \\ < i'\ using nj ji' by auto + have dom\: \\' \ \ dom att\ using domi cs\ last_cs by metis + obtain \ where i\\: \is_kth_obs \' \ \\ using dom\ by (metis is_kth_obs_def domIff) + hence \k: \\ < k\ using \i' iki' by (metis kth_obs_le_iff) + obtain \' where i\\': \is_kth_obs \ \ \'\ using \k iki by (metis kth_obs_stable) + have \\' < i\ using \k iki i\\' by (metis kth_obs_le_iff) + hence cs\': \cs\<^bsup>\'\<^esup> \ \ cs\<^bsup>\\<^esup> \'\ unfolding cs\ using cs_inj[OF path(1) nreti, of \\'\] by blast + thus \?thesis\ using less(1)[OF \k _ i\\' i\\] by auto + next + assume notin'': \\(\ \. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> \)\ + obtain \ \' where \i: \\' < i\ and cs\: \cs\<^bsup>\'\<^esup> \ = cs\<^bsup>\\<^esup> \'\ and suc\: \\' (Suc \) \ \ (Suc \')\ and cdloop': \i cd\<^bsup>\\<^esup>\ \' \ (\ j>\. j cd\<^bsup>\'\<^esup>\ \)\ + using missing_cd_or_loop[OF path \0 notin''] by metis + show \?thesis\ using cdloop' proof + assume cdjn: \i cd\<^bsup>\\<^esup>\ \'\ + hence csnj: \cs\<^bsup>\'\<^esup> \ \ cs\<^bsup>\\<^esup> i\ using cs\ by (metis cd_is_cs_less) + have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> \)) = \ (Suc \')\ using cs\ by (metis cdjn cd_not_ret cs_select_id path(1)) + have \(\',\) \ (\,i)\ using csnj apply(rule contradicts.intros(1)) using cssel \' \ suc\ by auto + thus \?thesis\ using domi by auto + next + assume loop': \\ j>\. j cd\<^bsup>\'\<^esup>\ \\ + have \n': \\' < n'\ using in' \i by auto + have nret\': \\ \' \ return\ by (metis cs\ last_cs le_eq_less_or_eq lessI path(1) path(2) suc\ term_path_stable) + have \\ < n\ using cs_order[OF path cs\[symmetric] csn[symmetric] nret\' \n'] . + hence \\ < i'\ using nj ji' by auto + hence cdi\: \i' cd\<^bsup>\'\<^esup>\ \\ using loop' by auto + hence cs\i': \cs\<^bsup>\\<^esup> \' \ cs\<^bsup>\'\<^esup> i'\ using cs\ by (metis cd_is_cs_less) + have cssel: \\' (Suc (\' \ cs\<^bsup>\\<^esup> \')) = \' (Suc \)\ using cs\ by (metis cdi\ cd_not_ret cs_select_id path(2)) + have \(\,\') \ (\',i')\ using cs\i' apply(rule contradicts.intros(1)) using cssel \' \ suc\ by auto + thus \?thesis\ using domi' by auto + qed + qed + next + assume \\ i \ n'\ + hence ni: \n'< i\ by simp + hence cdin: \i cd\<^bsup>\\<^esup>\ n'\ using loop by auto + hence csni': \cs\<^bsup>\'\<^esup> n \ cs\<^bsup>\\<^esup> i\ using csn by (metis cd_is_cs_less) + have cssel: \\ (Suc (\ \ cs\<^bsup>\'\<^esup> n)) = \ (Suc n')\ using csn by (metis cdin cd_not_ret cs_select_id path(1)) + have \(\',n) \ (\,i)\ using csni' apply(rule contradicts.intros(1)) using cssel \' \ sucn by auto + thus \?thesis\ using domi by auto + qed + qed + qed + qed +qed + +theorem obs_neq_some_contradicts: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ +assumes obsne: \obs \ k \ obs \' k\ and not_none: \obs \ k \ None\ \obs \' k \ None\ +shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom att) \ ((\, k) \ (\' ,k') \ \' k' \ dom att)\ +proof - + obtain i where iki: \is_kth_obs \ k i\ using not_none(1) by (metis \ obs_some_kth_obs) + obtain i' where iki': \is_kth_obs \' k i'\ using not_none(2) by (metis \' obs_some_kth_obs) + have \obsp \ i = obs \ k\ by (metis \ iki kth_obs_unique obs_def the_equality) + moreover + have \obsp \' i' = obs \' k\ by (metis \' iki' kth_obs_unique obs_def the_equality) + ultimately + have obspne: \obsp \ i \ obsp \' i'\ using obsne by auto + show \?thesis\ using obs_neq_some_contradicts'[OF _ iki[unfolded \] iki'[unfolded \']] using obspne \ \' by metis +qed + +theorem obs_neq_ret_contradicts: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ +assumes ret: \\ n = return\ and obsne: \obs \' i \ obs \ i\ and obs:\obs \' i \ None\ +shows \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ +proof (cases \\ j k'. is_kth_obs \' j k' \ (\ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\) + case True + obtain l k' where jk': \is_kth_obs \' l k'\ and unmatched: \(\ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ using True by blast + have \0: \\ 0 = \' 0\ using \ \' path0 by auto + obtain j j' where csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j'\ and cd: \k' cd\<^bsup>\'\<^esup>\j'\ and suc: \\ (Suc j) \ \' (Suc j')\ + using converged_cd_diverge_return[of \\'\ \\\ \k'\ \n\] ret unmatched path_is_path \ \' \0 by metis + hence *: \(\, j) \ (\' ,k')\ using contradicts.intros(1)[of \\\ \j\ \\'\ \k'\ \\'\ \\\, unfolded csj] \ \' + using cd_is_cs_less cd_not_ret cs_select_id by auto + have \\' k' \ dom(att)\ using jk' by (meson domIff is_kth_obs_def) + thus \?thesis\ using * by blast +next + case False + hence *: \\ j k'. is_kth_obs \' j k' \ \ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ by auto + obtain k' where k': \is_kth_obs \' i k'\ using obs \' obs_some_kth_obs by blast + obtain l where \is_kth_obs \ i l\ using * \ \' k' no_kth_obs_missing_cs path_is_path by metis + thus \?thesis\ using \ \' obs obs_neq_some_contradicts obs_none_no_kth_obs obsne by metis +qed + + +subsection \Facts about Critical Observable Paths\ + +lemma contradicting_in_cp: assumes leq:\\ =\<^sub>L \'\ and cseq: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ +and readv: \v\reads(path \ k)\ and vneq: \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ shows \((\,k),(\',k')) \ cp\ + using cseq readv vneq proof(induction \k+k'\ arbitrary: \k\ \k'\ \v\ rule: less_induct) + fix k k' v + assume csk: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ + assume vread: \v \ reads (path \ k)\ + assume vneq: \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ + assume IH: \\ka k'a v. ka + k'a < k + k' \ cs\<^bsup>path \\<^esup> ka = cs\<^bsup>path \'\<^esup> k'a \ v \ reads (path \ ka) \ (\\<^bsup>ka\<^esup>) v \ (\'\<^bsup>k'a\<^esup>) v \ ((\, ka), \', k'a) \ cp\ + + define \ where \\ \ path \\ + define \' where \\' \ path \'\ + have path: \\ = path \\ \\' = path \'\ using \_def \'_def path_is_path by auto + have ip: \is_path \\ \is_path \'\ using path path_is_path by auto + + have \0: \\' 0 = \ 0\ unfolding path path_def by auto + have vread': \v \ reads (path \' k')\ using csk vread by (metis last_cs) + have cseq: \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k\ using csk path by simp + + show \((\, k), \', k') \ cp\ proof cases + assume vnw: \\ l < k. v\writes (\ l)\ + hence \v: \(\\<^bsup>k\<^esup>) v = \ v\ by (metis no_writes_unchanged0 path(1)) + show \?thesis\ proof cases + assume vnw': \\ l < k'. v\writes (\' l)\ + hence \v': \(\'\<^bsup>k'\<^esup>) v = \' v\ by (metis no_writes_unchanged0 path(2)) + with \v vneq have \\ v \ \' v\ by auto + hence vhigh: \v \ hvars\ using leq unfolding loweq_def restrict_def by (auto,metis) + thus \?thesis\ using cp.intros(1)[OF leq csk vread vneq] vnw vnw' path by simp + next + assume \\(\ l < k'. v\writes (\' l))\ + then obtain l' where kddl': \k' dd\<^bsup>\',v\<^esup>\ l'\ using path(2) path_is_path written_read_dd vread' by blast + hence lv': \v \ writes (\' l')\ unfolding is_ddi_def by auto + have lk': \l' < k'\ by (metis is_ddi_def kddl') + have nret: \\' l' \ return\ using lv' writes_return by auto + + have notin\: \\ (\l. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l)\ proof + assume \\l. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l\ + then guess l .. + note csl = \cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> l\ + have lk: \l < k\ using lk' cseq ip cs_order[of \\'\ \\\ \l'\ \l\ \k'\ \k\] csl nret path by force + + have \v \ writes (\ l)\ using csl lv' last_cs by metis + thus \False\ using lk vnw by blast + qed + + from converged_cd_diverge[OF ip(2,1) \0 notin\ lk' cseq] + obtain i i' where csi: \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ and lcdi: \l' cd\<^bsup>\'\<^esup>\ i'\ and div: \\' (Suc i') \ \ (Suc i)\ . + + have 1: \\ (Suc i) = suc (\ i) (\\<^bsup>i\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) + have 2: \\' (Suc i') = suc (\' i') (\'\<^bsup>i'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) + have 3: \\' i' = \ i\ using csi last_cs by metis + have nreads: \\\<^bsup>i\<^esup> \ reads (\ i) \ \'\<^bsup>i'\<^esup> \ reads (\ i)\ by (metis 1 2 3 div reads_restr_suc) + then obtain v' where v'read: \v'\ reads(path \ i)\ \(\\<^bsup>i\<^esup>) v' \ (\'\<^bsup>i'\<^esup>) v'\ unfolding path by (metis reads_restrict) + + have nreti: \\' i' \ return\ by (metis csi div ip(1) ip(2) last_cs lessI term_path_stable less_imp_le) + have ik': \i' < k'\ using lcdi lk' is_cdi_def by auto + have ik: \i < k\ using cs_order[OF ip(2,1) csi cseq nreti ik'] . + + have cpi: \((\, i), (\', i')) \ cp\ using IH[of \i\ \i'\] v'read csi ik ik' path by auto + hence cpi': \((\', i'), (\, i)) \ cp\ using cp.intros(4) by blast + + have nwvi: \\j'\{LEAST i'. i < i' \ (\i. cs\<^bsup>path \'\<^esup> i = cs\<^bsup>path \\<^esup> i').. writes (path \ j')\ using vnw[unfolded path] + by (metis (poly_guards_query) atLeastLessThan_iff) + + from cp.intros(3)[OF cpi' kddl'[unfolded path] lcdi[unfolded path] csk[symmetric] div[unfolded path] vneq[symmetric] nwvi] + + show \?thesis\ using cp.intros(4) by simp + qed + next + assume wv: \\ (\ l writes (\ l))\ + then obtain l where kddl: \k dd\<^bsup>\,v\<^esup>\ l\ using path(1) path_is_path written_read_dd vread by blast + hence lv: \v \ writes (\ l)\ unfolding is_ddi_def by auto + have lk: \l < k\ by (metis is_ddi_def kddl) + have nret: \\ l \ return\ using lv writes_return by auto + have nwb: \\ i \ {Suc l..< k}. v\writes(\ i)\ using kddl unfolding is_ddi_def by auto + have \vk: \(\\<^bsup>k\<^esup>) v = (\\<^bsup>Suc l\<^esup> ) v\ using kddl ddi_value path(1) by auto + + show \?thesis\ proof cases + assume vnw': \\ l < k'. v\writes (\' l)\ + hence \v': \(\'\<^bsup>k'\<^esup>) v = \' v\ by (metis no_writes_unchanged0 path(2)) + + have notin\': \\ (\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l')\ proof + assume \\l'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ + then guess l' .. + note csl = \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ + have lk: \l' < k'\ using lk cseq ip cs_order[of \\\ \\'\ \l\ \l'\ \k\ \k'\] csl nret by metis + + have \v \ writes (\' l')\ using csl lv last_cs by metis + thus \False\ using lk vnw' by blast + qed + + from converged_cd_diverge[OF ip(1,2) \0[symmetric] notin\' lk cseq[symmetric]] + obtain i i' where csi: \cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i\ and lcdi: \l cd\<^bsup>\\<^esup>\ i\ and div: \\ (Suc i) \ \' (Suc i')\ by metis + + have 1: \\ (Suc i) = suc (\ i) (\\<^bsup>i\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) + have 2: \\' (Suc i') = suc (\' i') (\'\<^bsup>i'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) + have 3: \\' i' = \ i\ using csi last_cs by metis + have nreads: \\\<^bsup>i\<^esup> \ reads (\ i) \ \'\<^bsup>i'\<^esup> \ reads (\ i)\ by (metis 1 2 3 div reads_restr_suc) + have contri: \(\',i') \ (\,i)\ using contradicts.intros(2)[OF csi path nreads] . + + have nreti: \\ i \ return\ by (metis csi div ip(1) ip(2) last_cs lessI term_path_stable less_imp_le) + have ik: \i < k\ using lcdi lk is_cdi_def by auto + have ik': \i' < k'\ using cs_order[OF ip(1,2) csi[symmetric] cseq[symmetric] nreti ik] . + have nreads: \\\<^bsup>i\<^esup> \ reads (\ i) \ \'\<^bsup>i'\<^esup> \ reads (\ i)\ by (metis 1 2 3 div reads_restr_suc) + then obtain v' where v'read: \v'\ reads(path \ i)\ \(\\<^bsup>i\<^esup>) v' \ (\'\<^bsup>i'\<^esup>) v'\ unfolding path by (metis reads_restrict) + + + have cpi: \((\, i), (\', i')) \ cp\ using IH[of \i\ \i'\] v'read csi ik ik' path by auto + hence cpi': \((\', i'), (\, i)) \ cp\ using cp.intros(4) by blast + + have vnwi: \\j'\{LEAST i'a. i' < i'a \ (\i. cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i'a).. writes (path \' j')\ using vnw'[unfolded path] + by (metis (poly_guards_query) atLeastLessThan_iff) + + from cp.intros(3)[OF cpi kddl[unfolded path] lcdi[unfolded path] csk div[unfolded path] vneq vnwi] + + show \?thesis\ using cp.intros(4) by simp + next + assume \\ (\ l writes (\' l))\ + then obtain l' where kddl': \k' dd\<^bsup>\',v\<^esup>\ l'\ using path(2) path_is_path written_read_dd vread' by blast + hence lv': \v \ writes (\' l')\ unfolding is_ddi_def by auto + have lk': \l' < k'\ by (metis is_ddi_def kddl') + have nretl': \\' l' \ return\ using lv' writes_return by auto + have nwb': \\ i' \ {Suc l'..< k'}. v\writes(\' i')\ using kddl' unfolding is_ddi_def by auto + have \vk': \(\'\<^bsup>k'\<^esup>) v = (\'\<^bsup>Suc l'\<^esup> ) v\ using kddl' ddi_value path(2) by auto + + show \?thesis\ proof cases + assume csl: \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> l'\ + hence \l: \\ l = \' l'\ by (metis last_cs) + have \vls: \(\\<^bsup>Suc l\<^esup> ) v \ (\'\<^bsup>Suc l'\<^esup> ) v\ by (metis \vk \vk' vneq) + have r\: \\\<^bsup>l\<^esup> \ reads (\ l) \ \'\<^bsup>l'\<^esup> \ reads (\ l)\ using path \l \vls written_value lv by blast + then obtain v' where v'read: \v'\ reads(path \ l)\ \(\\<^bsup>l\<^esup>) v' \ (\'\<^bsup>l'\<^esup>) v'\ unfolding path by (metis reads_restrict) + + + have cpl: \((\, l), (\', l')) \ cp\ using IH[of \l\ \l'\] v'read csl lk lk' path by auto + show \((\, k), (\', k')) \ cp\ using cp.intros(2)[OF cpl kddl[unfolded path] kddl'[unfolded path] csk vneq] . + next + assume csl: \cs\<^bsup>\\<^esup> l \ cs\<^bsup>\'\<^esup> l'\ + show \?thesis\ proof cases + assume \\ i'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> i'\ + then obtain i' where csli': \cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> i'\ by blast + have ilne': \i' \ l'\ using csl csli' by auto + have ij': \i' < k'\ using cs_order[OF ip csli' cseq[symmetric] nret lk] . + have iv': \v \ writes(\' i')\ using lv csli' last_cs by metis + have il': \i' < l'\ using kddl' ilne' ij' iv' unfolding is_ddi_def by auto + have nreti': \\' i' \ return\ using csli' nret last_cs by metis + + have l'notin\: \\(\i. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> i )\ proof + assume \\i. cs\<^bsup>\'\<^esup> l' = cs\<^bsup>\\<^esup> i\ + then obtain i where csil: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l'\ by metis + have ik: \i < k\ using cs_order[OF ip(2,1) csil[symmetric] cseq nretl' lk'] . + have li: \l < i\ using cs_order[OF ip(2,1) csli'[symmetric] csil[symmetric] nreti' il'] . + have iv: \v \ writes(\ i)\ using lv' csil last_cs by metis + show \False\ using kddl ik li iv is_ddi_def by auto + qed + + obtain n n' where csn: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ and sucn: \\ (Suc n) \ \' (Suc n')\ and in': \i' \ n'\ + using converged_cd_diverge_cs [OF ip(2,1) csli'[symmetric] il' l'notin\ lk' cseq] by metis + + \ \Can apply the IH to n and n'\ + + have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) + have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) + have 3: \\' n' = \ n\ using csn last_cs by metis + have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 sucn reads_restr_suc) + then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) + moreover + have nl': \n' < l'\ using lcdn' is_cdi_def by auto + have nk': \n' < k'\ using nl' lk' by simp + have nretn': \\' n' \ return\ by (metis ip(2) nl' nretl' term_path_stable less_imp_le) + have nk: \n < k\ using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . + hence lenn: \n+n' < k+k'\ using nk' by auto + ultimately + have \((\, n), (\', n')) \ cp\ using IH csn path by auto + hence ncp: \((\', n'), (\, n)) \ cp\ using cp.intros(4) by auto + + have nles: \n < (LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i'))\ (is \_ < (LEAST i. ?P i)\) using nk cseq LeastI[of \?P\ \k\] by metis + moreover + have ln: \l \ n\ using cs_order_le[OF ip(2,1) csli'[symmetric] csn[symmetric] nreti' in'] . + ultimately + have lles: \Suc l \ (LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i'))\ by auto + + have nwcseq: \\j'\{LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i').. writes (\ j')\ proof + fix j' assume *: \j' \ {LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i').. + hence \(LEAST i'. n < i' \ (\i. cs\<^bsup>\'\<^esup> i = cs\<^bsup>\\<^esup> i')) \ j'\ by (metis (poly_guards_query) atLeastLessThan_iff) + hence \Suc l \ j'\ using lles by auto + moreover + have \j' < k\ using * by (metis (poly_guards_query) atLeastLessThan_iff) + ultimately have \j'\ {Suc l.. by (metis (poly_guards_query) atLeastLessThan_iff) + thus \v\writes (\ j')\ using nwb by auto + qed + + from cp.intros(3)[OF ncp,folded path,OF kddl' lcdn' cseq sucn[symmetric] vneq[symmetric] nwcseq] + have \((\', k'), \, k) \ cp\ . + thus \((\, k), (\', k')) \ cp\ using cp.intros(4) by auto + next + assume lnotin\': \\ (\i'. cs\<^bsup>\\<^esup> l = cs\<^bsup>\'\<^esup> i')\ + show \?thesis\ proof cases + assume \\ i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l'\ + then obtain i where csli: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l'\ by blast + have ilne: \i \ l\ using csl csli by auto + have ij: \i < k\ using cs_order[OF ip(2,1) csli[symmetric] cseq nretl' lk'] . + have iv: \v \ writes(\ i)\ using lv' csli last_cs by metis + have il: \i < l\ using kddl ilne ij iv unfolding is_ddi_def by auto + have nreti: \\ i \ return\ using csli nretl' last_cs by metis + + obtain n n' where csn: \cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn: \l cd\<^bsup>\\<^esup>\ n\ and sucn: \\ (Suc n) \ \' (Suc n')\ and ilen: \i \ n\ + using converged_cd_diverge_cs [OF ip csli il lnotin\' lk cseq[symmetric]] by metis + + \ \Can apply the IH to n and n'\ + + have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) + have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) + have 3: \\' n' = \ n\ using csn last_cs by metis + have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 sucn reads_restr_suc) + then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) + moreover + have nl: \n < l\ using lcdn is_cdi_def by auto + have nk: \n < k\ using nl lk by simp + have nretn: \\ n \ return\ by (metis ip(1) nl nret term_path_stable less_imp_le) + have nk': \n' < k'\ using cs_order[OF ip csn cseq[symmetric] nretn nk] . + hence lenn: \n+n' < k+k'\ using nk by auto + ultimately + have ncp: \((\, n), (\', n')) \ cp\ using IH csn path by auto + + have nles': \n' < (LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))\ (is \_ < (LEAST i. ?P i)\) using nk' cseq LeastI[of \?P\ \k'\] by metis + moreover + have ln': \l' \ n'\ using cs_order_le[OF ip csli csn nreti ilen] . + ultimately + have lles': \Suc l' \ (LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))\ by auto + + have nwcseq': \\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')).. writes (\' j')\ proof + fix j' assume *: \j' \ {(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')).. + hence \(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')) \ j'\ by (metis (poly_guards_query) atLeastLessThan_iff) + hence \Suc l' \ j'\ using lles' by auto + moreover + have \j' < k'\ using * by (metis (poly_guards_query) atLeastLessThan_iff) + ultimately have \j'\ {Suc l'.. by (metis (poly_guards_query) atLeastLessThan_iff) + thus \v\writes (\' j')\ using nwb' by auto + qed + + from cp.intros(3)[OF ncp,folded path, OF kddl lcdn cseq[symmetric] sucn vneq nwcseq'] + + show \((\, k), (\', k')) \ cp\ . + next + assume l'notin\: \\ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> l')\ + define m where \m \ 0::nat\ + define m' where \m' \ 0::nat\ + have csm: \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ unfolding m_def m'_def cs_0 by (metis \0) + have ml: \m m' using csm csl unfolding m_def m'_def by (metis neq0_conv) + have \\ n n'. cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n' \ \ (Suc n) \ \' (Suc n') \ + (l cd\<^bsup>\\<^esup>\ n \ (\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j')) + \ l' cd\<^bsup>\'\<^esup>\ n' \ (\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)))\ + using csm ml proof (induction \k+k'-(m+m')\ arbitrary: \m\ \m'\ rule: less_induct) + case (less m m') + note csm = \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m'\ + note lm = \m < l \ m' < l'\ + note IH = \\ n n'. + k + k' - (n + n') < k + k' - (m + m') \ + cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n' \ + n < l \ n' < l' \ ?thesis\ + show \?thesis\ using lm proof + assume ml: \m < l\ + obtain n n' where mn: \m \ n\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn: \l cd\<^bsup>\\<^esup>\ n\ and suc: \\ (Suc n) \ \' (Suc n')\ + using converged_cd_diverge_cs[OF ip csm ml lnotin\' lk cseq[symmetric]] . + have nl: \n < l\ using lcdn is_cdi_def by auto + hence nk: \n using lk by auto + have nretn: \\ n \ return\ using lcdn by (metis cd_not_ret) + have nk': \n' using cs_order[OF ip csn cseq[symmetric] nretn nk] . + show \?thesis\ proof cases + assume \\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j')\ + thus \?thesis\ using lcdn csn suc by blast + next + assume \\(\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j'))\ + then obtain j' where jin': \j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')).. and vwrite: \v\writes (\' j')\ by blast + define i' where \i' \ LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i')\ + have Pk': \n' < k' \ (\ k. cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k')\ (is \?P k'\) using nk' cseq[symmetric] by blast + have ni': \n' < i'\ using LeastI[of \?P\, OF Pk'] i'_def by auto + obtain i where csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using LeastI[of \?P\, OF Pk'] i'_def by blast + have ij': \i'\j'\ using jin'[folded i'_def] by auto + have jk': \j' using jin'[folded i'_def] by auto + have jl': \j' \ l'\ using kddl' jk' vwrite unfolding is_ddi_def by auto + have nretn': \\' n' \ return\ using nretn csn last_cs by metis + have iln: \n using cs_order[OF ip(2,1) csn[symmetric] csi[symmetric] nretn' ni'] . + hence mi: \m < i\ using mn by auto + have nretm: \\ m \ return\ by (metis ip(1) mn nretn term_path_stable) + have mi': \m' using cs_order[OF ip csm csi nretm mi] . + have ik': \i' < k'\ using ij' jk' by auto + have nreti': \\' i' \ return\ by (metis ij' jl' nretl' ip(2) term_path_stable) + have ik: \i < k\ using cs_order[OF ip(2,1) csi[symmetric] cseq nreti' ik'] . + show \?thesis\ proof cases + assume il:\i < l\ + have le: \k + k' - (i +i') < k+k' - (m+m')\ using mi mi' ik ik' by auto + show \?thesis\ using IH[OF le] using csi il by blast + next + assume \\ i < l\ + hence li: \l \ i\ by auto + have \i' \ l'\ using ij' jl' by auto + hence il': \i' < l'\ using csi l'notin\ by fastforce + obtain n n' where in': \i' \ n'\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ and suc: \\ (Suc n) \ \' (Suc n')\ + using converged_cd_diverge_cs[OF ip(2,1) csi[symmetric] il' _ lk' cseq] l'notin\ by metis + have nk': \n' < k'\ using lcdn' is_cdi_def lk' by auto + have nretn': \\' n' \ return\ by (metis cd_not_ret lcdn') + have nk: \n < k\ using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . + define j where \j \ LEAST j. n < j \ (\j'. cs\<^bsup>\'\<^esup> j' = cs\<^bsup>\\<^esup> j)\ + have Pk: \n < k \ (\j'. cs\<^bsup>\'\<^esup> j' = cs\<^bsup>\\<^esup> k)\ (is \?P k\) using nk cseq by blast + have nj: \n using LeastI[of \?P\, OF Pk] j_def by auto + have ilen: \i \ n\ using cs_order_le[OF ip(2,1) csi[symmetric] csn[symmetric] nreti' in'] . + hence lj: \l using li nj by simp + have \\l\{l<.. writes (\ l)\ using kddl unfolding is_ddi_def by simp + hence nw: \\l\{j.. writes (\ l)\ using lj by auto + show \?thesis\ using csn lcdn' suc nw[unfolded j_def] by blast + qed + qed + next + assume ml': \m' < l'\ + obtain n n' where mn': \m' \ n'\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ and suc: \\ (Suc n) \ \' (Suc n')\ + using converged_cd_diverge_cs[OF ip(2,1) csm[symmetric] ml' _ lk' cseq] l'notin\ by metis + have nl': \n' < l'\ using lcdn' is_cdi_def by auto + hence nk': \n' using lk' by auto + have nretn': \\' n' \ return\ using lcdn' by (metis cd_not_ret) + have nk: \n using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . + show \?thesis\ proof cases + assume \\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)\ + thus \?thesis\ using lcdn' csn suc by blast + next + assume \\(\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j))\ + then obtain j where jin: \j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i)).. and vwrite: \v\writes (\ j)\ by blast + define i where \i \ LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i)\ + have Pk: \n < k \ (\ k'. cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k)\ (is \?P k\) using nk cseq by blast + have ni: \n < i\ using LeastI[of \?P\, OF Pk] i_def by auto + obtain i' where csi: \cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'\ using LeastI[of \?P\, OF Pk] i_def by metis + have ij: \i\j\ using jin[folded i_def] by auto + have jk: \j using jin[folded i_def] by auto + have jl: \j \ l\ using kddl jk vwrite unfolding is_ddi_def by auto + have nretn: \\ n \ return\ using nretn' csn last_cs by metis + have iln': \n' using cs_order[OF ip csn csi nretn ni] . + hence mi': \m' < i'\ using mn' by auto + have nretm': \\' m' \ return\ by (metis ip(2) mn' nretn' term_path_stable) + have mi: \m using cs_order[OF ip(2,1) csm[symmetric] csi[symmetric] nretm' mi'] . + have ik: \i < k\ using ij jk by auto + have nreti: \\ i \ return\ by (metis ij ip(1) jl nret term_path_stable) + have ik': \i' < k'\ using cs_order[OF ip csi cseq[symmetric] nreti ik] . + show \?thesis\ proof cases + assume il':\i' < l'\ + have le: \k + k' - (i +i') < k+k' - (m+m')\ using mi mi' ik ik' by auto + show \?thesis\ using IH[OF le] using csi il' by blast + next + assume \\ i' < l'\ + hence li': \l' \ i'\ by auto + have \i \ l\ using ij jl by auto + hence il: \i < l\ using csi lnotin\' by fastforce + obtain n n' where ilen: \i \ n\ and csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and lcdn: \l cd\<^bsup>\\<^esup>\ n\ and suc: \\ (Suc n) \ \' (Suc n')\ + using converged_cd_diverge_cs[OF ip csi il _ lk cseq[symmetric]] lnotin\' by metis + have nk: \n < k\ using lcdn is_cdi_def lk by auto + have nretn: \\ n \ return\ by (metis cd_not_ret lcdn) + have nk': \n' < k'\ using cs_order[OF ip csn cseq[symmetric] nretn nk] . + define j' where \j' \ LEAST j'. n' < j' \ (\j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> j')\ + have Pk': \n' < k' \ (\j. cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> k')\ (is \?P k'\) using nk' cseq[symmetric] by blast + have nj': \n' using LeastI[of \?P\, OF Pk'] j'_def by auto + have in': \i' \ n'\ using cs_order_le[OF ip csi csn nreti ilen] . + hence lj': \l' using li' nj' by simp + have \\l\{l'<.. writes (\' l)\ using kddl' unfolding is_ddi_def by simp + hence nw': \\l\{j'.. writes (\' l)\ using lj' by auto + show \?thesis\ using csn lcdn suc nw'[unfolded j'_def] by blast + qed + qed + qed + qed + then obtain n n' where csn: \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n'\ and suc: \\ (Suc n) \ \' (Suc n')\ + and cdor: + \(l cd\<^bsup>\\<^esup>\ n \ (\j'\{(LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i'))..writes (\' j')) + \ l' cd\<^bsup>\'\<^esup>\ n' \ (\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)))\ + by blast + show \?thesis\ using cdor proof + assume *: \l cd\<^bsup>\\<^esup>\ n \ (\j'\{LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i').. local.writes (\' j'))\ + hence lcdn: \l cd\<^bsup>\\<^esup>\ n\ by blast + have nowrite: \\j'\{LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i').. local.writes (\' j')\ using * by blast + show \?thesis\ proof (rule cp.intros(3)[of \\\ \n\ \\'\ \n'\,folded path]) + show \l cd\<^bsup>\\<^esup>\ n\ using lcdn . + show \k dd\<^bsup>\,v\<^esup>\ l\ using kddl . + show \cs\<^bsup>\\<^esup> k = cs\<^bsup>\'\<^esup> k'\ using cseq by simp + show \\ (Suc n) \ \' (Suc n')\ using suc by simp + show \\j'\{LEAST i'. n' < i' \ (\i. cs\<^bsup>\\<^esup> i = cs\<^bsup>\'\<^esup> i').. local.writes (\' j')\ using nowrite . + show \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ using vneq . + have nk: \n < k\ using lcdn lk is_cdi_def by auto + have nretn: \\ n \ return\ using cd_not_ret lcdn by metis + have nk': \n' < k'\ using cs_order[OF ip csn cseq[symmetric] nretn nk] . + hence le: \n + n' < k + k'\ using nk by auto + moreover + have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) + have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) + have 3: \\' n' = \ n\ using csn last_cs by metis + have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 suc reads_restr_suc) + then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) + ultimately + show \((\, n), (\', n')) \ cp\ using IH csn path by auto + qed + next + assume *: \l' cd\<^bsup>\'\<^esup>\ n' \ (\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j))\ + hence lcdn': \l' cd\<^bsup>\'\<^esup>\ n'\ by blast + have nowrite: \\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)\ using * by blast + show \?thesis\ proof (rule cp.intros(4), rule cp.intros(3)[of \\'\ \n'\ \\\ \n\,folded path]) + show \l' cd\<^bsup>\'\<^esup>\ n'\ using lcdn' . + show \k' dd\<^bsup>\',v\<^esup>\ l'\ using kddl' . + show \cs\<^bsup>\'\<^esup> k' = cs\<^bsup>\\<^esup> k\ using cseq . + show \\' (Suc n') \ \ (Suc n)\ using suc by simp + show \\j\{(LEAST i. n < i \ (\i'. cs\<^bsup>\'\<^esup> i' = cs\<^bsup>\\<^esup> i))..writes (\ j)\ using nowrite . + show \(\'\<^bsup>k'\<^esup>) v \ (\\<^bsup>k\<^esup>) v\ using vneq by simp + have nk': \n' < k'\ using lcdn' lk' is_cdi_def by auto + have nretn': \\' n' \ return\ using cd_not_ret lcdn' by metis + have nk: \n < k\ using cs_order[OF ip(2,1) csn[symmetric] cseq nretn' nk'] . + hence le: \n + n' < k + k'\ using nk' by auto + moreover + have 1: \\ (Suc n) = suc (\ n) (\\<^bsup>n\<^esup>)\ by (metis step_suc_sem fst_conv path(1) path_suc) + have 2: \\' (Suc n') = suc (\' n') (\'\<^bsup>n'\<^esup>)\ by (metis step_suc_sem fst_conv path(2) path_suc) + have 3: \\' n' = \ n\ using csn last_cs by metis + have nreads: \\\<^bsup>n\<^esup> \ reads (\ n) \ \'\<^bsup>n'\<^esup> \ reads (\ n)\ by (metis 1 2 3 suc reads_restr_suc) + then obtain v' where v'read: \v'\reads (path \ n)\ \(\\<^bsup>n\<^esup>) v' \ (\'\<^bsup>n'\<^esup>) v'\ by (metis path(1) reads_restrict) + ultimately + have \((\, n), (\', n')) \ cp\ using IH csn path by auto + thus \((\', n'), \, n) \ cp\ using cp.intros(4) by simp + qed + qed + qed + qed + qed + qed + qed +qed + + +theorem contradicting_in_cop: assumes \\ =\<^sub>L \'\ and \(\',k') \ (\,k)\ and \path \ k \ dom att\ +shows \((\,k),\',k') \ cop\ using assms(2) proof(cases) + case (1 \' \) + define j where \j \ \ \ cs\<^bsup>\'\<^esup> k'\ + have csj: \cs\<^bsup>\\<^esup> j = cs\<^bsup>\'\<^esup> k'\ unfolding j_def using 1 by (metis cs_not_nil cs_select_is_cs(1) path_is_path) + have suc: \\ (Suc j) \ \' (Suc k')\ using 1 j_def by simp + have kcdj: \k cd\<^bsup>\\<^esup>\ j\ by (metis cs_not_nil cs_select_is_cs(2) 1(1,2) j_def path_is_path) + obtain v where readv: \v\reads(path \ j)\ and vneq: \(\\<^bsup>j\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ using suc csj unfolding 1 by (metis IFC_def.suc_def 1(2) 1(3) last_cs path_suc reads_restr_suc reads_restrict) + have \((\,j),\',k') \ cp\ apply (rule contradicting_in_cp[OF assms(1)]) using readv vneq csj 1 by auto + thus \((\,k),\',k') \ cop\ using kcdj suc assms(3) cop.intros(2) unfolding 1 by auto + next + case (2 \' \) + obtain v where readv: \v\reads(path \ k)\ and vneq: \(\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v\ using 2(2-4) by (metis reads_restrict) + have \((\,k),\',k') \ cp\ apply (rule contradicting_in_cp[OF assms(1)]) using readv vneq 2 by auto + thus \((\,k),\',k') \ cop\ using assms(3) cop.intros(1) unfolding 2 by auto +qed + + +theorem cop_correct_term: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ +assumes ret: \\ n = return\ \\' n' = return\ and obsne: \obs \ \ obs \'\ and leq: \\ =\<^sub>L \'\ +shows \\ k k'. ((\,k),\',k')\ cop \ ((\',k'),\,k)\ cop\ +proof - + have *: \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ using obs_neq_contradicts_term ret obsne \ \' by auto + have leq' :\\' =\<^sub>L \\ using leq unfolding loweq_def by auto + from * contradicting_in_cop[OF leq] contradicting_in_cop[OF leq'] show \?thesis\ unfolding \ \' by metis +qed + + +theorem cop_correct_ret: fixes \ \' defines \: \\ \ path \\ and \': \\' \ path \'\ +assumes ret: \\ n = return\ and obsne: \obs \ i \ obs \' i\ and obs: \obs \' i \ None\ and leq: \\ =\<^sub>L \'\ +shows \\ k k'. ((\,k),\',k')\ cop \ ((\',k'),\,k)\ cop\ +proof - + have *: \\ k k'. ((\', k') \ (\ ,k) \ \ k \ dom (att)) \ ((\, k) \ (\' ,k') \ \' k' \ dom (att))\ + by (metis (no_types, lifting) \ \' obs obs_neq_ret_contradicts obsne ret) + have leq' :\\' =\<^sub>L \\ using leq unfolding loweq_def by auto + from * contradicting_in_cop[OF leq] contradicting_in_cop[OF leq'] show \?thesis\ unfolding \ \' by metis +qed + + +theorem cop_correct_nterm: assumes obsne: \obs \ k \ obs \' k\ \obs \ k \ None\ \obs \' k \ None\ +and leq: \\ =\<^sub>L \'\ +shows \\ k k'. ((\,k),\',k')\ cop \ ((\',k'),\,k)\ cop\ +proof - + obtain k k' where \((\', k') \ (\ ,k) \ path \ k \ dom att) \ ((\, k) \ (\' ,k') \ path \' k' \ dom att)\ + using obs_neq_some_contradicts[OF obsne] by metis + thus \?thesis\ proof + assume *: \(\', k') \ (\ ,k) \ path \ k \ dom att\ + hence \((\,k),\',k') \ cop\ using leq by (metis contradicting_in_cop) + thus \?thesis\ using * by blast + next + assume *: \(\, k) \ (\' ,k') \ path \' k' \ dom att\ + hence \((\',k'),\,k) \ cop\ using leq by (metis contradicting_in_cop loweq_def) + thus \?thesis\ using * by blast + qed +qed + + +subsection \Correctness of the Characterisation\ +text_raw \\label{sec:cor-cp}\ + +text \The following is our main correctness result. If there exist no critical observable paths, +then the program is secure.\ + +theorem cop_correct: assumes \cop = empty\ shows \secure\ proof (rule ccontr) + assume \\ secure\ + then obtain \ \' where leq: \ \ =\<^sub>L \'\ + and **: \\ obs \ \ obs \' \ (terminates \ \ \ obs \' \ obs \)\ + unfolding secure_def by blast + show \False\ using ** proof + assume \\ obs \ \ obs \'\ + then obtain k where \obs \ k \ obs \' k \ obs \ k \ None \ obs \' k \ None\ + unfolding obs_comp_def obs_prefix_def + by (metis kth_obs_stable linorder_neqE_nat obs_none_no_kth_obs obs_some_kth_obs) + thus \False\ using cop_correct_nterm leq assms by auto + next + assume *: \terminates \ \ \ obs \' \ obs \\ + then obtain n where ret: \path \ n = return\ + unfolding terminates_def by auto + obtain k where \obs \ k \ obs \' k \ obs \' k \ None\ using * unfolding obs_prefix_def by metis + thus \False\ using cop_correct_ret ret leq assms by (metis empty_iff) + qed +qed + + +text \Our characterisation is not only correct, it is also precise in the way that \cp\ characterises +exactly the matching indices in executions for low equivalent input states where diverging data is read. +This follows easily as the inverse implication to lemma \contradicting_in_cp\ can be shown by simple induction.\ + +theorem cp_iff_reads_contradict: \((\,k),(\',k')) \ cp \ \ =\<^sub>L \' \ cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k' \ (\ v\reads(path \ k). (\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v)\ +proof + assume \\ =\<^sub>L \' \ cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k' \ (\v\reads (path \ k). (\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v)\ + thus \((\, k), \', k') \ cp\ using contradicting_in_cp by blast +next + assume \((\, k), \', k') \ cp\ + thus \\ =\<^sub>L \' \ cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k' \ (\v\reads (path \ k). (\\<^bsup>k\<^esup>) v \ (\'\<^bsup>k'\<^esup>) v)\ + proof (induction) + case (1 \ \' n n' h) + then show \?case\ by blast + next + case (2 \ k \' k' n v n') + have \v\reads (path \ n)\ using 2(2) unfolding is_ddi_def by auto + then show \?case\ using 2 by auto + next + case (3 \ k \' k' n v l n') + have \v\reads (path \ n)\ using 3(2) unfolding is_ddi_def by auto + then show \?case\ using 3(4,6,8) by auto + next + case (4 \ k \' k') + hence \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ by simp + hence \path \' k' = path \ k\ by (metis last_cs) + moreover have \\' =\<^sub>L \\ using 4(2) unfolding loweq_def by simp + ultimately show \?case\ using 4 by metis + qed +qed + + +text \In the same way the inverse implication to \contradicting_in_cop\ follows easily +such that we obtain the following characterisation of \cop\.\ + +theorem cop_iff_contradicting: \((\,k),(\',k')) \ cop \ \ =\<^sub>L \' \ (\',k') \ (\,k) \ path \ k \ dom att\ +proof + assume \\ =\<^sub>L \' \ (\', k') \ (\, k) \ path \ k \ dom att\ thus \((\,k),(\',k')) \ cop\ using contradicting_in_cop by simp +next + assume \((\,k),(\',k')) \ cop\ + thus \ \ =\<^sub>L \' \ (\',k') \ (\,k) \ path \ k \ dom att\ proof (cases rule: cop.cases) + case 1 + then show \?thesis\ using cp_iff_reads_contradict contradicts.simps by (metis (full_types) reads_restrict1) + next + case (2 k) + then show \?thesis\ using cp_iff_reads_contradict contradicts.simps + by (metis cd_is_cs_less cd_not_ret contradicts.intros(1) cs_select_id path_is_path) + qed +qed + + +subsection \Correctness of the Single Path Approximation\ +text_raw \\label{sec:cor-scp}\ + +theorem cp_in_scp: assumes \((\,k),(\',k'))\cp\ shows \(path \,k)\scp \ (path \',k')\scp\ +using assms proof(induction \\\ \k\ \\'\ \k'\ rule:cp.induct[case_names read_high dd dcd sym]) + case (read_high \ \' k k' h) + have \\ h = (\\<^bsup>k\<^esup>) h\ using read_high(5) by (simp add: no_writes_unchanged0) + moreover have \\' h = (\'\<^bsup>k'\<^esup>) h\ using read_high(6) by (simp add: no_writes_unchanged0) + ultimately have \\ h \ \' h\ using read_high(4) by simp + hence *: \h\hvars\ using read_high(1) unfolding loweq_def by (metis Compl_iff IFC_def.restrict_def) + have 1: \(path \,k)\scp\ using scp.intros(1) read_high(3,5) * by auto + have \path \ k = path \' k'\ using read_high(2) by (metis last_cs) + hence \(path \',k')\scp\ using scp.intros(1) read_high(3,6) * by auto + thus \?case\ using 1 by auto +next + case dd show \?case\ using scp.intros(3) dd by auto +next + case sym thus \?case\ by blast +next + case (dcd \ k \' k' n v l n') + note scp.intros(4) is_dcdi_via_def cd_cs_swap cs_ipd + have 1: \(path \, n)\scp\ using dcd.IH dcd.hyps(2) dcd.hyps(3) scp.intros(2) scp.intros(3) by blast + have csk: \cs\<^bsup>path \\<^esup> k = cs\<^bsup>path \'\<^esup> k'\ using cp_eq_cs[OF dcd(1)] . + have kn: \k and kl: \k and ln: \l using dcd(2,3) unfolding is_ddi_def is_cdi_def by auto + have nret: \path \ k \ return\ using cd_not_ret dcd.hyps(3) by auto + have \k' < n'\ using kn csk dcd(4) cs_order nret path_is_path last_cs by blast + have 2: \(path \', n')\scp\ proof cases + assume j'ex: \\j'\{k'.. writes (path \' j')\ + hence \\j'. j'\{k'.. v \ writes (path \' j')\ by auto + note * = GreatestI_ex_nat[OF this] + define j' where \j' == GREATEST j'. j'\{k'.. v \ writes (path \' j')\ + note ** = *[of \j'\,folded j'_def] + have \k' \ j'\ \j' and j'write: \v \ writes (path \' j')\ + using "*" atLeastLessThan_iff j'_def nat_less_le by auto + have nowrite: \\ i'\{j'<.. writes(path \' i')\ proof (rule, rule ccontr) + fix i' assume \i' \ {j'<.. \\ v \ local.writes (path \' i')\ + hence \i' \ {k'.. v \ local.writes (path \' i')\ using \k' \ j'\ by auto + hence \i' \ j'\ using Greatest_le_nat + by (metis (no_types, lifting) atLeastLessThan_iff j'_def nat_less_le) + thus \False\ using \i' \ {j'<.. by auto + qed + have \path \' n' = path \ n\ using dcd(4) last_cs by metis + hence \v\reads(path \' n')\ using dcd(2) unfolding is_ddi_def by auto + hence nddj': \n' dd\<^bsup>path \',v\<^esup>\ j'\ using dcd(2) unfolding is_ddi_def using nowrite \j' j'write by auto + show \?thesis\ proof cases + assume \j' cd\<^bsup>path \'\<^esup>\ k'\ + thus \(path \',n') \ scp\ using scp.intros(2) scp.intros(3) dcd.IH nddj' by fast + next + assume jcdk': \\ j' cd\<^bsup>path \'\<^esup>\ k'\ + show \?thesis\ proof cases + assume \j' = k'\ + thus \?thesis\ using scp.intros(3) dcd.IH nddj' by fastforce + next + assume \j' \ k'\ hence \k' < j'\ using \k' \ j'\ by auto + have \path \' j' \ return\ using j'write writes_return by auto + hence ipdex':\\j. j \{k'..j'} \ path \' j = ipd (path \' k') \ using path_is_path \k' < j'\ jcdk' is_cdi_def by blast + define i' where \i' == LEAST j. j\ {k'..j'} \ path \' j = ipd (path \' k')\ + have iipd': \i'\ {k'..j'}\ \path \' i' = ipd (path \' k')\ unfolding i'_def using LeastI_ex[OF ipdex'] by simp_all + have *:\\ i \ {k'..' i \ ipd (path \' k')\ proof (rule, rule ccontr) + fix i assume *: \i \ {k'.. \\ path \' i \ ipd (path \' k')\ + hence **: \i \{k'..j'} \ path \' i = ipd (path \' k')\ (is \?P i\) using iipd'(1) by auto + thus \False\ using Least_le[of \?P\ \i\] i'_def * by auto + qed + have \i' \ k'\ using iipd'(2) by (metis csk last_cs nret path_in_nodes ipd_not_self) + hence \k' using iipd'(1) by simp + hence csi': \cs\<^bsup>path \'\<^esup> i' = [n\cs\<^bsup>path \'\<^esup> k' . ipd n \ path \' i'] @ [path \' i']\using cs_ipd[OF iipd'(2) *] by fast + + have ncdk': \\ n' cd\<^bsup>path \'\<^esup>\ k'\ using \j' < n'\ \k' < j'\ cdi_prefix jcdk' less_imp_le_nat by blast + hence ncdk: \\ n cd\<^bsup>path \\<^esup>\ k\ using cd_cs_swap csk dcd(4) by blast + have ipdex: \\i. i\{k..n} \ path \ i = ipd (path \ k)\ (is \\i. ?P i\) proof cases + assume *:\path \ n = return\ + from path_ret_ipd[of \path \\ \k\ \n\,OF path_is_path nret *] + obtain i where \?P i\ by fastforce thus \?thesis\ by auto + next + assume *:\path \ n \ return\ + show \?thesis\ using not_cd_impl_ipd [of \path \\ \k\ \n\, OF path_is_path \k ncdk *] by auto + qed + + define i where \i == LEAST j. j\ {k..n} \ path \ j = ipd (path \ k)\ + have iipd: \i\ {k..n}\ \path \ i = ipd (path \ k)\ unfolding i_def using LeastI_ex[OF ipdex] by simp_all + have **:\\ i' \ {k.. i' \ ipd (path \ k)\ proof (rule, rule ccontr) + fix i' assume *: \i' \ {k.. \\ path \ i' \ ipd (path \ k)\ + hence **: \i' \{k..n} \ path \ i' = ipd (path \ k)\ (is \?P i'\) using iipd(1) by auto + thus \False\ using Least_le[of \?P\ \i'\] i_def * by auto + qed + have \i \ k\ using iipd(2) by (metis nret path_in_nodes ipd_not_self) + hence \k using iipd(1) by simp + hence \cs\<^bsup>path \\<^esup> i = [n\cs\<^bsup>path \\<^esup> k . ipd n \ path \ i] @ [path \ i]\using cs_ipd[OF iipd(2) **] by fast + hence csi: \cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i'\ using csi' csk unfolding iipd'(2) iipd(2) by (metis last_cs) + hence \(LEAST i'. k' < i' \ (\i. cs\<^bsup>path \\<^esup> i = cs\<^bsup>path \'\<^esup> i')) \ i'\ (is \(LEAST x. ?P x) \ _\) + using \k' < i'\ Least_le[of \?P\ \i'\] by blast + hence nw: \\j'\{i'.. writes (path \' j')\ using dcd(7) allB_atLeastLessThan_lower by blast + moreover have \v \ writes (path \' j')\ using nddj' unfolding is_ddi_def by auto + moreover have \i' \ j'\ using iipd'(1) by auto + ultimately have \False\ using \j' < n'\ by auto + thus \?thesis\ .. + qed + qed + next + assume \\ (\j'\{k'.. writes (path \' j'))\ + + hence \n' dcd\<^bsup>path \',v\<^esup>\ k' via (path \) k\ unfolding is_dcdi_via_def using dcd(2-4) csk \k' path_is_path by metis + thus \?thesis\ using dcd.IH scp.intros(4) by blast + qed + with 1 show \?case\ .. +qed + + +theorem cop_in_scop: assumes \((\,k),(\',k'))\cop\ shows \(path \,k)\scop \ (path \',k')\scp\ + using assms + apply (induct rule: cop.induct) + apply (simp add: cp_in_scp) + using cp_in_scp scop.intros scp.intros(2) + apply blast + using cp_in_scp scop.intros scp.intros(2) + apply blast + done + +text \The main correctness result for out single execution approximation follows directly.\ + +theorem scop_correct: assumes \scop = empty\ shows \secure\ + using cop_correct assms cop_in_scop by fast + +end + +end \ No newline at end of file diff --git a/thys/IFC_Tracking/PDG.thy b/thys/IFC_Tracking/PDG.thy new file mode 100644 --- /dev/null +++ b/thys/IFC_Tracking/PDG.thy @@ -0,0 +1,63 @@ +section \Example: Program Dependence Graphs\ +text_raw \\label{sec:pdg}\ + +text \Program dependence graph (PDG) based slicing provides a very crude but direct approximation of +our characterisation. As such we can easily derive a corresponding correctness result.\ + +theory PDG imports IFC +begin + +context IFC +begin + +text \We utilise our established dependencies on program paths to define the PDG. Note that PDGs +usually only contain immediate control dependencies instead of the transitive ones we use here. +However as slicing is considering reachability questions this does not affect the result.\ +inductive_set pdg where +\\i cd\<^bsup>\\<^esup>\ k\ \ (\ k, \ i) \ pdg\ | +\\i dd\<^bsup>\,v\<^esup>\ k\ \ (\ k, \ i) \ pdg\ + +text \The set of sources is the set of nodes reading high variables.\ +inductive_set sources where +\n \ nodes \ h \ hvars \ h \ reads n \ n\ sources\ + +text \The forward slice is the set of nodes reachable in the PDG from the set of sources. To ensure +security slicing aims to prove that no observable node is contained in the \ +inductive_set slice where +\n\ sources \ n \ slice\ | +\m \ slice \ (m,n)\pdg \ n \ slice\ + + +text \As the PDG does not contain data control dependencies themselves we have to decompose these.\ +lemma dcd_pdg: assumes \n dcd\<^bsup>\,v\<^esup>\ m via \' m'\ obtains l where \(\ m,l)\ pdg\ and \(l,\ n)\pdg\ +proof - + assume r: \(\l. (\ m, l) \ pdg \ (l, \ n) \ pdg \ thesis)\ + obtain l' n' where ln: \cs\<^bsup>\\<^esup> m = cs\<^bsup>\'\<^esup> m' \ cs\<^bsup>\\<^esup> n = cs\<^bsup>\'\<^esup> n' \ n' dd\<^bsup>\',v\<^esup>\ l' \ l' cd\<^bsup>\'\<^esup>\ m'\ using assms unfolding is_dcdi_via_def by metis + hence mn: \\' m' = \ m \ \' n' = \ n\ by (metis last_cs ln) + have 1: \(\ m, \' l') \ pdg\ by (metis ln mn pdg.intros(1)) + have 2: \(\' l', \ n) \ pdg\ by (metis ln mn pdg.intros(2)) + show thesis using 1 2 r by auto +qed + +text \By induction it directly follows that the slice is an approximation of the single critical paths.\ +lemma scp_slice: \(\, i)\ scp \ \ i \ slice\ + apply (induction rule: scp.induct) + apply (simp add: path_in_nodes slice.intros(1) sources.intros) + using pdg.intros(1) slice.intros(2) apply blast + using pdg.intros(2) slice.intros(2) apply blast + by (metis dcd_pdg slice.intros(2)) + +lemma scop_slice: \(\, i) \ scop \ \ i \ slice \ dom(att)\ by (metis IntI scop.cases scp_slice) + +text \The requirement targeted by slicing, that no observable node is contained in the slice, +is thereby a sound criteria for security.\ +lemma pdg_correct: assumes \slice \ dom(att) = {}\ shows \secure\ +proof (rule ccontr) + assume \\ secure\ + then obtain \ i where \(\, i) \ scop\ using scop_correct by force + thus \False\ using scop_slice assms by auto +qed + +end + +end \ No newline at end of file diff --git a/thys/IFC_Tracking/ROOT b/thys/IFC_Tracking/ROOT new file mode 100644 --- /dev/null +++ b/thys/IFC_Tracking/ROOT @@ -0,0 +1,10 @@ +chapter AFP + +session IFC_Tracking (AFP) = HOL + + options [timeout = 600] + theories + IFC + PDG + document_files + "root.tex" + "root.bib" diff --git a/thys/IFC_Tracking/document/root.bib b/thys/IFC_Tracking/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/IFC_Tracking/document/root.bib @@ -0,0 +1,17 @@ +@inproceedings{Bohannon:2009:RN:1653662.1653673, + author = {Bohannon, Aaron and Pierce, Benjamin C. and Sj\"{o}berg, Vilhelm and Weirich, Stephanie and Zdancewic, Steve}, + title = {Reactive Noninterference}, + booktitle = {Proceedings of the 16th ACM Conference on Computer and Communications Security}, + series = {CCS '09}, + year = {2009}, + isbn = {978-1-60558-894-0}, + location = {Chicago, Illinois, USA}, + pages = {79--90}, + numpages = {12}, + url = {http://doi.acm.org/10.1145/1653662.1653673}, + doi = {10.1145/1653662.1653673}, + acmid = {1653673}, + publisher = {ACM}, + address = {New York, NY, USA}, + keywords = {information flow, noninterference, reactive programming, web applications, web browsers}, +} \ No newline at end of file diff --git a/thys/IFC_Tracking/document/root.tex b/thys/IFC_Tracking/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/IFC_Tracking/document/root.tex @@ -0,0 +1,97 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} +\usepackage[left=2cm, right=2cm, top=3cm, bottom=3cm]{geometry} +\usepackage[T1]{fontenc} + +% further packages required for unusual symbols (see also +% isabellesym.sty), use only when needed + +\usepackage{amssymb} + %for \, \, \, \, \, \, + %\, \, \, \, \, + %\, \, \ + +%\usepackage{eurosym} + %for \ + +\newcommand{\flqq}{\guillemotleft} + +\usepackage[english]{babel} + + +%\usepackage[only,bigsqcap]{stmaryrd} + %for \ + +%\usepackage{eufrak} + %for \ ... \, \ ... \ (also included in amssymb) + +%\usepackage{textcomp} + %for \, \, \, \, \, + %\ + +% This should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +% for uniform font size +%\renewcommand{\isastyle}{\isastyleminor} + +%\renewcommand{\isamarkupparagraph}[1]{\paragraph{#1}\mbox{}\\} +%\renewcommand{\isamarkupsubparagraph}[1]{\subparagraph{#1}\mbox{}\\} + +%\addto\extrasenglish{\renewcommand{\subsectionautorefname}{section}} +%\addto\extrasenglish{\renewcommand{\subsubsectionautorefname}{section}} + + + +\begin{document} + +\title{Information Flow Control via Dependency Tracking} +\author{Benedikt Nordhoff} +%\date{} +\maketitle + +\begin{abstract} + We provide a characterisation of how information is propagated by program executions based on the + tracking data and control dependencies within executions themselves. The characterisation might + be used for deriving approximative safety properties to be targeted by static analyses or checked + at runtime. We utilise a simple yet versatile control flow graph model as a program + representation. As our model is not assumed to be finite it can be instantiated for a broad class + of programs. The targeted security property is indistinguishable security where executions + produce sequences of observations and only non-terminating executions are allowed to drop a tail + of those. + + A very crude approximation of our characterisation is slicing based on program dependence graphs, + which we use as a minimal example and derive a corresponding soundness result. + + For further details and applications refer to the authors upcoming dissertation. +\end{abstract} + +\newpage + +\tableofcontents + +\newpage + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +\nocite{Bohannon:2009:RN:1653662.1653673} + +% optional bibliography +\bibliographystyle{abbrv} +\bibliography{root} + + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,592 +1,593 @@ ADS_Functor AI_Planning_Languages_Semantics AODV AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BNF_CC BNF_Operations BTree Banach_Steinhaus Bell_Numbers_Spivey Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Blue_Eyes Bondy Boolean_Expression_Checkers Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRDT CYK CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean ClockSynchInst Closest_Pair_Points CofGroups Coinductive Coinductive_Languages Collections Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Consensus_Refined Constructive_Cryptography Constructive_Cryptography_CM Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties CSP_RefTK DFS_Framework DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos DOM_Components DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Example-Submission Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Harrison FOL_Seq_Calc1 Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Fourier Free-Boolean-Algebra Free-Groups FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL IEEE_Floating_Point +IFC_Tracking IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpreter_Optimizations Interval_Arithmetic_Word32 Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model IsaGeoCoq Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knuth_Bendix_Order Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Of_Medians_Selection Menger Mereology Mersenne_Primes MiniML Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_arithmetic_LLL_and_HNF_algorithms Modular_Assembly_Kit_Security Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Myhill-Nerode Name_Carrying_Type_Inference Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PCF PLM POPLmark-deBruijn PSemigroupsConvolution Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Program-Conflict-Analysis Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem QHLProver QR_Decomposition Quantales Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Regular-Sets Regular_Algebras Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rep_Fin_Groups Residuated_Lattices Resolution_FOL Rewriting_Z Ribbon_Proofs Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Routing Roy_Floyd_Warshall SATSolverVerification SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Saturation_Framework_Extensions Shadow_DOM Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sort_Encodings Source_Coding_Theorem Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God UPF UPF_Firewall UTP Universal_Turing_Machine UpDown_Scheme Valuation VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Well_Quasi_Orders Winding_Number_Eval Word_Lib WorkerWrapper XML ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL