diff --git a/thys/HOL-CSP/Assertions.thy b/thys/HOL-CSP/Assertions.thy --- a/thys/HOL-CSP/Assertions.thy +++ b/thys/HOL-CSP/Assertions.thy @@ -1,620 +1,706 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * * Author : Burkhart Wolff, Safouan Taha, Lina Ye. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) theory Assertions - imports CSP Process_Order + imports CSP begin +default_sort "type" section\CSP Assertions\ definition DF :: "'a set \ 'a process" where "DF A \ \ X. \ x \ A \ X" lemma DF_unfold : "DF A = (\ z \ A \ DF A)" by(simp add: DF_def, rule trans, rule fix_eq, simp) definition deadlock_free :: "'a process \ bool" where "deadlock_free P \ DF UNIV \\<^sub>F\<^sub>D P" definition DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P :: "'a set \ 'a process" where "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \ \ X. ((\ x \ A \ X) \ SKIP)" (* TO DO: rename this in deadlock_free\<^sub>S\<^sub>K\<^sub>I\<^sub>P *) definition deadlock_free_v2 :: "'a process \ bool" where "deadlock_free_v2 P \ DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>F P" section\Some deadlock freeness laws\ lemma DF_subset: "A \ {} \ A \ B \ DF B \\<^sub>F\<^sub>D DF A" apply(subst DF_def, rule fix_ind) apply(rule le_FD_adm, simp_all add: monofunI, subst DF_unfold) by (meson mono_Mndetprefix_FD mono_Mndetprefix_FD_set trans_FD) lemma DF_Univ_freeness: "A \ {} \ (DF A) \\<^sub>F\<^sub>D P \ deadlock_free P" by (meson deadlock_free_def DF_subset failure_divergence_refine_def order.trans subset_UNIV) lemma deadlock_free_Ndet: "deadlock_free P \ deadlock_free Q \ deadlock_free (P \ Q)" by (simp add: D_Ndet F_Ndet deadlock_free_def failure_divergence_refine_def le_ref_def) section \Preliminaries\ lemma DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold : "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A = ((\ z \ A \ DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) \ SKIP)" by(simp add: DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule trans, rule fix_eq, simp) section \Deadlock Free\ lemma div_free_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P: "\(DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = {}" proof(simp add:DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def fix_def) define Y where "Y \ \i. iterate i\(\ x. (\xa\A \ x) \ SKIP)\\" hence a:"chain Y" by simp have "\ (Y (Suc i)) = {d. d \ [] \ hd d \ (ev ` A) \ tl d \ \(Y i)}" for i by (cases "A={}", auto simp add:Y_def D_STOP D_SKIP D_Mndetprefix write0_def D_Mprefix D_Ndet) hence b:"d \ \(Y i) \ length d \ i" for d i by (induct i arbitrary:d, simp_all add: Nitpick.size_list_simp(2)) { fix x assume c:"\ i. x \ \ (Y i)" from b have "x \ \ (Y (Suc (length x)))" using Suc_n_not_le_n by blast with c have False by simp } with a show "\ (\i. Y i) = {}" using D_LUB[OF a] limproc_is_thelub[OF a] by auto qed lemma div_free_DF: "\(DF A) = {}" proof - have "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>D DF A" apply (simp add:DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def) by(rule fix_ind, simp_all add: monofunI, subst DF_unfold, simp) then show ?thesis using divergence_refine_def div_free_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P by blast qed lemma deadlock_free_implies_div_free: "deadlock_free P \ \ P = {}" by (simp add: deadlock_free_def div_free_DF failure_divergence_refine_def le_ref_def) section \Run\ definition RUN :: "'a set \ 'a process" where "RUN A \ \ X. \ x \ A \ X" definition CHAOS :: "'a set \ 'a process" where "CHAOS A \ \ X. (STOP \ (\ x \ A \ X))" definition lifelock_free :: "'a process \ bool" where "lifelock_free P \ CHAOS UNIV \\<^sub>F\<^sub>D P" section \Reference processes and their unfolding rules\ definition CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P :: "'a set \ 'a process" where "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \ \ X. (SKIP \ STOP \ (\ x \ A \ X))" lemma RUN_unfold : "RUN A = (\ z \ A \ RUN A)" by(simp add: RUN_def, rule trans, rule fix_eq, simp) lemma CHAOS_unfold : "CHAOS A = (STOP \ (\ z \ A \ CHAOS A))" by(simp add: CHAOS_def, rule trans, rule fix_eq, simp) lemma CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \ SKIP \ STOP \ (\ x \ A \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" unfolding CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def using fix_eq[of "\ X. (SKIP \ STOP \ (\ x \ A \ X))"] by simp section \Process events and reference processes events\ definition events_of where "events_of P \ (\t\\ P. {a. ev a \ set t})" lemma events_DF: "events_of (DF A) = A" proof(auto simp add:events_of_def) fix x t show "t \ \ (DF A) \ ev x \ set t \ x \ A" proof(induct t) case Nil then show ?case by simp next case (Cons a t) from Cons(2) have "a # t \ \ (\z\A \ DF A)" using DF_unfold by metis then obtain aa where "a = ev aa \ aa \ A \ t \ \ (DF A)" by (cases "A={}", auto simp add: T_Mndetprefix write0_def T_Mprefix T_STOP) with Cons show ?case by auto qed next fix x show "x \ A \ \xa\\ (DF A). ev x \ set xa" apply(subst DF_unfold, cases "A={}", auto simp add:T_Mndetprefix write0_def T_Mprefix) by (metis Nil_elem_T list.sel(1) list.sel(3) list.set_intros(1)) qed lemma events_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P: "events_of (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = A" proof(auto simp add:events_of_def) fix x t show "t \ \ (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) \ ev x \ set t \ x \ A" proof(induct t) case Nil then show ?case by simp next case (Cons a t) from Cons(2) have "a # t \ \ ((\z\A \ DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) \ SKIP)" using DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold by metis with Cons obtain aa where "a = ev aa \ aa \ A \ t \ \ (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" by (cases "A={}", auto simp add:T_Mndetprefix write0_def T_Mprefix T_STOP T_SKIP T_Ndet) with Cons show ?case by auto qed next fix x show "x \ A \ \xa\\ (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A). ev x \ set xa" apply(subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, cases "A={}") apply(auto simp add:T_Mndetprefix write0_def T_Mprefix T_SKIP T_Ndet) by (metis Nil_elem_T list.sel(1) list.sel(3) list.set_intros(1)) qed lemma events_RUN: "events_of (RUN A) = A" proof(auto simp add:events_of_def) fix x t show "t \ \ (RUN A) \ ev x \ set t \ x \ A" proof(induct t) case Nil then show ?case by simp next case (Cons a t) from Cons(2) have "a # t \ \ (\z\A \ RUN A)" using RUN_unfold by metis then obtain aa where "a = ev aa \ aa \ A \ t \ \ (RUN A)" by (auto simp add:T_Mprefix) with Cons show ?case by auto qed next fix x show "x \ A \ \xa \ \(RUN A). ev x \ set xa" apply(subst RUN_unfold, simp add: T_Mprefix) by (metis Nil_elem_T list.sel(1) list.sel(3) list.set_intros(1)) qed lemma events_CHAOS: "events_of (CHAOS A) = A" proof(auto simp add:events_of_def) fix x t show "t \ \ (CHAOS A) \ ev x \ set t \ x \ A" proof(induct t) case Nil then show ?case by simp next case (Cons a t) from Cons(2) have "a # t \ \ (STOP \ (\ z \ A \ CHAOS A))" using CHAOS_unfold by metis then obtain aa where "a = ev aa \ aa \ A \ t \ \ (CHAOS A)" by (auto simp add:T_Ndet T_Mprefix T_STOP) with Cons show ?case by auto qed next fix x show "x \ A \ \xa\\ (CHAOS A). ev x \ set xa" apply(subst CHAOS_unfold, simp add:T_Ndet T_Mprefix T_STOP) by (metis Nil_elem_T list.sel(1) list.sel(3) list.set_intros(1)) qed lemma events_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P: "events_of (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = A" proof(auto simp add:events_of_def) fix x t show "t \ \ (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) \ ev x \ set t \ x \ A" proof(induct t) case Nil then show ?case by simp next case (Cons a t) from Cons(2) have "a # t \ \ (SKIP \ STOP \ (\ z \ A \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A))" using CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold by metis with Cons obtain aa where "a = ev aa \ aa \ A \ t \ \ (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" by (auto simp add:T_Ndet T_Mprefix T_STOP T_SKIP) with Cons show ?case by auto qed next fix x show "x \ A \ \xa\\ (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A). ev x \ set xa" apply(subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, simp add:T_Ndet T_Mprefix T_STOP T_SKIP) by (metis Nil_elem_T list.sel(1) list.sel(3) list.set_intros(1)) qed lemma events_div: "\(P) \ {} \ events_of (P) = UNIV" proof(auto simp add:events_of_def) fix x xa assume 1: "x \ \ P" show "\x\\ P. ev xa \ set x" proof(cases "tickFree x") case True hence "x@[ev xa] \ \ P" using 1 NT_ND front_tickFree_single is_processT7 by blast then show ?thesis by force next case False hence "(butlast x)@[ev xa] \ \ P" by (metis "1" D_T D_imp_front_tickFree append_single_T_imp_tickFree butlast_snoc front_tickFree_single nonTickFree_n_frontTickFree process_charn) then show ?thesis by force qed qed lemma DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_subset_FD: "A \ {} \ A \ B \ DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P B \\<^sub>F\<^sub>D DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" apply(subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule fix_ind, rule le_FD_adm, simp_all add:monofunI, subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold) by (rule mono_Ndet_FD, simp_all) (meson mono_Mndetprefix_FD mono_Mndetprefix_FD_set trans_FD) lemma RUN_subset_DT: "A \ B \ RUN B \\<^sub>D\<^sub>T RUN A" apply(subst RUN_def, rule fix_ind, rule le_DT_adm, simp_all add:monofunI, subst RUN_unfold) by (meson mono_Mprefix_DT mono_Mprefix_DT_set trans_DT) lemma CHAOS_subset_FD: "A \ B \ CHAOS B \\<^sub>F\<^sub>D CHAOS A" apply(subst CHAOS_def, rule fix_ind, rule le_FD_adm, simp_all add:monofunI, subst CHAOS_unfold) by (auto simp add: failure_divergence_refine_def le_ref_def D_Mprefix D_Ndet F_STOP F_Mprefix F_Ndet) lemma CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_subset_FD: "A \ B \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P B \\<^sub>F\<^sub>D CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" apply(subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule fix_ind, rule le_FD_adm) apply(simp_all add:monofunI, subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold) by (auto simp add: failure_divergence_refine_def le_ref_def D_Mprefix D_Ndet F_STOP F_Mprefix F_Ndet) section \Relations between refinements on reference processes\ lemma CHAOS_has_all_tickFree_failures: "tickFree a \ {x. ev x \ set a} \ A \ (a,b) \ \ (CHAOS A)" proof(induct a) case Nil then show ?case by (subst CHAOS_unfold, simp add:F_Ndet F_STOP) next case (Cons a0 al) hence "tickFree al" by (metis append.left_neutral append_Cons front_tickFree_charn front_tickFree_mono) with Cons show ?case apply (subst CHAOS_unfold, simp add:F_Ndet F_STOP F_Mprefix events_of_def) using event_set by blast qed lemma CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures: assumes as:"(events_of P) \ A" shows "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F P" proof - have "front_tickFree a \ set a \ (\t\\ P. set t) \ (a,b) \ \ (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" for a b proof(induct a) case Nil then show ?case by (subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, simp add:F_Ndet F_STOP) next case (Cons a0 al) hence "front_tickFree al" by (metis append.left_neutral append_Cons front_tickFree_charn front_tickFree_mono) with Cons show ?case apply (subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, simp add:F_Ndet F_STOP F_SKIP F_Mprefix events_of_def as) apply(cases "a0=tick") apply (metis append.simps(2) front_tickFree_charn front_tickFree_mono list.distinct(1) tickFree_Cons) using event_set image_iff as[simplified events_of_def] by fastforce qed thus ?thesis by (simp add: F_T SUP_upper failure_refine_def process_charn subrelI) qed corollary CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures_ev: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P (events_of P) \\<^sub>F P" and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures_Un: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>F P" by (simp_all add: CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures) lemma DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF_refine_F: "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F DF A" by (simp add:DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule fix_ind, simp_all add: monofunI, subst DF_unfold, simp) lemma DF_RUN_refine_F: "DF A \\<^sub>F RUN A" apply (simp add:DF_def, rule fix_ind, simp_all add: monofunI, subst RUN_unfold) by (meson Mprefix_refines_Mndetprefix_F mono_Mndetprefix_F trans_F) lemma CHAOS_DF_refine_F: "CHAOS A \\<^sub>F DF A" apply (simp add:CHAOS_def DF_def, rule parallel_fix_ind, simp_all add: monofunI) apply (rule le_F_adm, simp_all add: monofun_snd) by (cases "A={}", auto simp add:adm_def failure_refine_def F_Mndetprefix F_Mprefix write0_def F_Ndet F_STOP) corollary CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_CHAOS_refine_F: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F CHAOS A" and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_refine_F: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" by (simp_all add: CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures events_CHAOS events_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P trans_F[OF CHAOS_DF_refine_F DF_RUN_refine_F]) lemma div_free_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P: "\ (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = {}" proof - have "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>D CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" proof (simp add:DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule fix_ind, simp_all add: monofunI, subst CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold) fix x assume 1:"x \\<^sub>D CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" have a:"((\xa\A \ x) \ SKIP) = (SKIP \ SKIP \ (\xa\A \ x))" by (simp add: Ndet_commute Ndet_id) from 1 have b:"(SKIP \ SKIP \ (\xa\A \ x)) \\<^sub>D (SKIP \ STOP \ (\xa\A \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A))" by (meson Mprefix_refines_Mndetprefix_D idem_D leD_STOP mono_Mprefix_D mono_Ndet_D trans_D) from a b show "((\xa\A \ x) |-| SKIP) \\<^sub>D (SKIP |-| STOP |-| Mprefix A (\x. CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A))" by simp qed then show ?thesis using divergence_refine_def div_free_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P by blast qed lemma div_free_CHAOS: "\(CHAOS A) = {}" proof - have "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>D CHAOS A" apply (simp add:CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def, rule fix_ind) by (simp_all add: monofunI, subst CHAOS_unfold, simp) then show ?thesis using divergence_refine_def div_free_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P by blast qed lemma div_free_RUN: "\(RUN A) = {}" proof - have "CHAOS A \\<^sub>D RUN A" by (simp add:CHAOS_def, rule fix_ind, simp_all add: monofunI, subst RUN_unfold, simp) then show ?thesis using divergence_refine_def div_free_CHAOS by blast qed corollary DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF_refine_FD: "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F\<^sub>D DF A" and DF_RUN_refine_FD: "DF A \\<^sub>F\<^sub>D RUN A" and CHAOS_DF_refine_FD: "CHAOS A \\<^sub>F\<^sub>D DF A" and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_CHAOS_refine_FD: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F\<^sub>D CHAOS A" and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_refine_FD: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A \\<^sub>F\<^sub>D DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A" using div_free_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P[of A] div_free_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P[of A] div_free_DF[of A] div_free_RUN[of A] div_free_CHAOS[of A] leF_leD_imp_leFD[OF DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF_refine_F, of A] leF_leD_imp_leFD[OF DF_RUN_refine_F, of A] leF_leD_imp_leFD[OF CHAOS_DF_refine_F, of A] leF_leD_imp_leFD[OF CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_CHAOS_refine_F, of A] leF_leD_imp_leFD[OF CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_refine_F, of A] by (auto simp add:divergence_refine_def) lemma traces_CHAOS_sub: "\(CHAOS A) \ {s. set s \ ev ` A}" proof(auto) fix s sa assume "s \ \ (CHAOS A)" and "sa \ set s" then show "sa \ ev ` A" apply (induct s, simp) by (subst (asm) (2) CHAOS_unfold, cases "A={}", auto simp add:T_Ndet T_STOP T_Mprefix) qed lemma traces_RUN_sub: "{s. set s \ ev ` A} \ \(RUN A)" proof(auto) fix s assume "set s \ ev ` A" then show "s \ \ (RUN A)" by (induct s, simp add: Nil_elem_T) (subst RUN_unfold, auto simp add:T_Mprefix) qed corollary RUN_all_tickfree_traces1: "\(RUN A) = {s. set s \ ev ` A}" and DF_all_tickfree_traces1: "\(DF A) = {s. set s \ ev ` A}" and CHAOS_all_tickfree_traces1: "\(CHAOS A) = {s. set s \ ev ` A}" using DF_RUN_refine_F[THEN leF_imp_leT, simplified trace_refine_def] CHAOS_DF_refine_F[THEN leF_imp_leT,simplified trace_refine_def] traces_CHAOS_sub traces_RUN_sub by blast+ corollary RUN_all_tickfree_traces2: "tickFree s \ s \ \(RUN UNIV)" and DF_all_tickfree_traces2: "tickFree s \ s \ \(DF UNIV)" and CHAOS_all_tickfree_trace2: "tickFree s \ s \ \(CHAOS UNIV)" apply(simp_all add:tickFree_def RUN_all_tickfree_traces1 DF_all_tickfree_traces1 CHAOS_all_tickfree_traces1) by (metis event_set insertE subsetI)+ lemma traces_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_sub: "\(CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) \ {s. front_tickFree s \ set s \ (ev ` A \ {tick})}" proof(auto simp add:is_processT2_TR) fix s sa assume "s \ \ (CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" and "sa \ set s" and "sa \ ev ` A" then show "sa = tick" apply (induct s, simp) by (subst (asm) (2) CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, cases "A={}", auto simp add:T_Ndet T_STOP T_SKIP T_Mprefix) qed lemma traces_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_sub: "{s. front_tickFree s \ set s \ (ev ` A \ {tick})} \ \(DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A::'a process)" proof(auto) fix s assume a:"front_tickFree s" and b:"set s \ insert tick (ev ` A)" have c:"front_tickFree ((tick::'a event) # s) \ s = []" for s by (metis butlast.simps(2) butlast_snoc front_tickFree_charn list.distinct(1) tickFree_Cons) with a b show "s \ \ (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A)" apply (induct s, simp add: Nil_elem_T, subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, cases "A={}") apply (subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, cases "A={}") apply(auto simp add:T_Mprefix T_Mndetprefix write0_def T_SKIP T_Ndet T_STOP) apply (metis append_Cons append_Nil front_tickFree_charn front_tickFree_mono) by (metis append_Cons append_Nil front_tickFree_mono) qed corollary DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces1: "\(DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = {s. front_tickFree s \ set s \ (ev ` A \ {tick})}" and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces1: "\(CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P A) = {s. front_tickFree s \ set s \ (ev ` A \ {tick})}" using CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_refine_F[THEN leF_imp_leT, simplified trace_refine_def] traces_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_sub traces_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_sub by blast+ corollary DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces2: "front_tickFree s \ s \ \(DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV)" and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces2: "front_tickFree s \ s \ \(CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV)" apply(simp_all add:tickFree_def DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces1 CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces1) by (metis event_set subsetI)+ corollary DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_traces: "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>T P" and CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_traces: "CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>T P" apply (simp add:trace_refine_def DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces2 is_processT2_TR subsetI) by (simp add:trace_refine_def CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_all_front_tickfree_traces2 is_processT2_TR subsetI) lemma deadlock_free_implies_non_terminating: "deadlock_free (P::'a process) \ \s\\ P. tickFree s" unfolding deadlock_free_def apply(drule leFD_imp_leF, drule leF_imp_leT) unfolding trace_refine_def using DF_all_tickfree_traces1[of "(UNIV::'a set)"] tickFree_def by fastforce lemma deadlock_free_v2_is_right: "deadlock_free_v2 (P::'a process) \ (\s\\ P. tickFree s \ (s, UNIV::'a event set) \ \ P)" proof assume a:"deadlock_free_v2 P" have "tickFree s \ (s, UNIV) \ \ (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV)" for s::"'a event list" proof(induct s) case Nil then show ?case by (subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, simp add:F_Mndetprefix write0_def F_Mprefix F_Ndet F_SKIP) next case (Cons a s) then show ?case by (subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, simp add:F_Mndetprefix write0_def F_Mprefix F_Ndet F_SKIP) qed with a show "\s\\ P. tickFree s \ (s, UNIV) \ \ P" using deadlock_free_v2_def failure_refine_def by blast next assume as1:"\s\\ P. tickFree s \ (s, UNIV) \ \ P" have as2:"front_tickFree s \ (\aa \ UNIV. ev aa \ b) \ (s, b) \ \ (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P (UNIV::'a set))" for s b proof(induct s) case Nil then show ?case by (subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold, auto simp add:F_Mndetprefix write0_def F_Mprefix F_Ndet F_SKIP) next case (Cons hda tla) then show ?case proof(simp add:DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def fix_def) define Y where "Y \ \i. iterate i\(\ x. (\xa\(UNIV::'a set) \ x) \ SKIP)\\" assume a:"front_tickFree (hda # tla)" and b:"front_tickFree tla \ (tla, b) \ \ (\i. Y i)" and c:"\aa. ev aa \ b" from Y_def have cc:"chain Y" by simp from b have d:"front_tickFree tla \ \aa\UNIV. ev aa \ b \(tla, b) \ \ (Y i)" for i using F_LUB[OF cc] limproc_is_thelub[OF cc] by simp from Y_def have e:"\(Mndetprefix UNIV (\x. Y i) \ SKIP) \ \ (Y (Suc i))" for i by(simp) from a have f:"tla \ [] \ hda \ tick" "front_tickFree tla" apply (metis butlast.simps(2) butlast_snoc front_tickFree_charn list.distinct(1) tickFree_Cons) by (metis a append_Cons append_Nil front_tickFree_Nil front_tickFree_mono) have g:"(hda#tla, b) \ \ (Y (Suc i))" for i using f c e[of i] d[of i] by (auto simp: F_Mndetprefix write0_def F_Mprefix Y_def F_Ndet F_SKIP) (metis event.exhaust)+ have h:"(hda#tla, b) \ \ (Y 0)" using NF_ND cc g po_class.chainE proc_ord2a by blast from a b c show "(hda#tla, b) \ \ (\i. Y i)" using F_LUB[OF cc] is_ub_thelub[OF cc] by (metis D_LUB_2 cc g limproc_is_thelub po_class.chainE proc_ord2a process_charn) qed qed show "deadlock_free_v2 P" proof(auto simp add:deadlock_free_v2_def failure_refine_def) fix s b assume as3:"(s, b) \ \ P" hence a1:"s \ \ P" "front_tickFree s" using F_T apply blast using as3 is_processT2 by blast show "(s, b) \ \ (DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV)" proof(cases "tickFree s") case FT_True:True hence a2:"(s, UNIV) \ \ P" using a1 as1 by blast then show ?thesis by (metis FT_True UNIV_I UNIV_eq_I a1(2) as2 as3 emptyE event.exhaust is_processT6_S1 tickFree_implies_front_tickFree_single) next case FT_False: False then show ?thesis by (meson T_F_spec UNIV_witness a1(2) append_single_T_imp_tickFree as2 emptyE is_processT5_S7) qed qed qed lemma deadlock_free_v2_implies_div_free: "deadlock_free_v2 P \ \ P = {}" by (metis F_T append_single_T_imp_tickFree deadlock_free_v2_is_right ex_in_conv nonTickFree_n_frontTickFree process_charn) corollary deadlock_free_v2_FD: "deadlock_free_v2 P = DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>F\<^sub>D P" unfolding deadlock_free_v2_def using deadlock_free_v2_implies_div_free leFD_imp_leF leF_leD_imp_leFD deadlock_free_v2_def divergence_refine_def by fastforce lemma all_events_refusal: "(s, {tick} \ ev ` (events_of P)) \ \ P \ (s, UNIV::'a event set) \ \ P" proof - assume a1:"(s, {tick} \ ev ` events_of P) \ \ P" { assume "(s, UNIV) \ \ P" then obtain c where "c \ {tick} \ ev ` events_of P \ s @ [c] \ \ P" using is_processT5_S1[of s "{tick} \ ev ` events_of P" P "UNIV - ({tick} \ ev ` events_of P)", simplified] F_T a1 by auto hence False by (simp add:events_of_def, cases c) fastforce+ } with a1 show "(s, UNIV) \ \ P" by blast qed corollary deadlock_free_v2_is_right_wrt_events: "deadlock_free_v2 (P::'a process) \ (\s\\ P. tickFree s \ (s, {tick} \ ev ` (events_of P)) \ \ P)" unfolding deadlock_free_v2_is_right using all_events_refusal apply auto using is_processT4 by blast lemma deadlock_free_is_deadlock_free_v2: "deadlock_free P \ deadlock_free_v2 P" using DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_DF_refine_FD deadlock_free_def deadlock_free_v2_FD trans_FD by blast section \\<^const>\deadlock_free\ and \<^const>\deadlock_free_v2\ with \<^const>\SKIP\ and \<^const>\STOP\\ lemma deadlock_free_v2_SKIP: "deadlock_free_v2 SKIP" unfolding deadlock_free_v2_FD by (subst DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_unfold) simp lemma non_deadlock_free_SKIP: \\ deadlock_free SKIP\ by (metis T_SKIP deadlock_free_implies_non_terminating insertCI non_tickFree_tick) lemma non_deadlock_free_v2_STOP: \\ deadlock_free_v2 STOP\ by (simp add: F_STOP Nil_elem_T deadlock_free_v2_is_right) lemma non_deadlock_free_STOP: \\ deadlock_free STOP\ using deadlock_free_is_deadlock_free_v2 non_deadlock_free_v2_STOP by blast +section \Non-terminating Runs\ + +definition non_terminating :: "'a process \ bool" + where "non_terminating P \ RUN UNIV \\<^sub>T P" + +corollary non_terminating_refine_DF: "non_terminating P = DF UNIV \\<^sub>T P" + and non_terminating_refine_CHAOS: "non_terminating P = CHAOS UNIV \\<^sub>T P" + by (simp_all add: DF_all_tickfree_traces1 RUN_all_tickfree_traces1 CHAOS_all_tickfree_traces1 + non_terminating_def trace_refine_def) + +lemma non_terminating_is_right: "non_terminating (P::'a process) \ (\s\\ P. tickFree s)" + apply (rule iffI) + by (auto simp add:non_terminating_def trace_refine_def tickFree_def RUN_all_tickfree_traces1)[1] + (auto simp add:non_terminating_def trace_refine_def RUN_all_tickfree_traces2) + +lemma nonterminating_implies_div_free: "non_terminating P \ \ P = {}" + unfolding non_terminating_is_right + by (metis NT_ND equals0I front_tickFree_charn process_charn tickFree_Cons tickFree_append) + +lemma non_terminating_implies_F: "non_terminating P \ CHAOS UNIV \\<^sub>F P" + apply(auto simp add: non_terminating_is_right failure_refine_def) + using CHAOS_has_all_tickFree_failures F_T by blast + +corollary non_terminating_F: "non_terminating P = CHAOS UNIV \\<^sub>F P" + by (auto simp add:non_terminating_implies_F non_terminating_refine_CHAOS leF_imp_leT) + +corollary non_terminating_FD: "non_terminating P = CHAOS UNIV \\<^sub>F\<^sub>D P" + unfolding non_terminating_F + using div_free_CHAOS nonterminating_implies_div_free leFD_imp_leF + leF_leD_imp_leFD divergence_refine_def non_terminating_F + by fastforce + + +section \Lifelock Freeness\ + +thm lifelock_free_def + +corollary lifelock_free_is_non_terminating: "lifelock_free P = non_terminating P" + unfolding lifelock_free_def non_terminating_FD by rule + +lemma div_free_divergence_refine: + "\ P = {} \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>D P" + "\ P = {} \ CHAOS UNIV \\<^sub>D P" + "\ P = {} \ RUN UNIV \\<^sub>D P" + "\ P = {} \ DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>D P" + "\ P = {} \ DF UNIV \\<^sub>D P" + by (simp_all add: div_free_CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P div_free_CHAOS div_free_RUN div_free_DF + div_free_DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P divergence_refine_def) + +definition lifelock_free_v2 :: "'a process \ bool" + where "lifelock_free_v2 P \ CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P UNIV \\<^sub>F\<^sub>D P" + +lemma div_free_is_lifelock_free_v2: "lifelock_free_v2 P \ \ P = {}" + using CHAOS\<^sub>S\<^sub>K\<^sub>I\<^sub>P_has_all_failures_Un leFD_imp_leD leF_leD_imp_leFD + div_free_divergence_refine(1) lifelock_free_v2_def + by blast + +lemma lifelock_free_is_lifelock_free_v2: "lifelock_free P \ lifelock_free_v2 P" + by (simp add: leFD_imp_leD div_free_divergence_refine(2) div_free_is_lifelock_free_v2 lifelock_free_def) + +corollary deadlock_free_v2_is_lifelock_free_v2: "deadlock_free_v2 P \ lifelock_free_v2 P" + by (simp add: deadlock_free_v2_implies_div_free div_free_is_lifelock_free_v2) + + +section \New laws\ + +lemma non_terminating_Seq: + "non_terminating P \ (P \<^bold>; Q) = P" + apply(auto simp add: non_terminating_is_right Process_eq_spec D_Seq F_Seq F_T is_processT7) + using process_charn apply blast + using process_charn apply blast + using F_T is_processT5_S2a by fastforce + +lemma non_terminating_Sync: + "non_terminating P \ lifelock_free_v2 Q \ non_terminating (P \A\ Q)" + apply(auto simp add: non_terminating_is_right div_free_is_lifelock_free_v2 T_Sync) + apply (metis equals0D ftf_Sync1 ftf_Sync21 insertI1 tickFree_def) + apply (meson NT_ND is_processT7_S tickFree_append) + by (metis D_T empty_iff ftf_Sync1 ftf_Sync21 insertI1 tickFree_def) + +lemmas non_terminating_Par = non_terminating_Sync[where A = \UNIV\] + and non_terminating_Inter = non_terminating_Sync[where A = \{}\] + + + end \ No newline at end of file diff --git a/thys/HOL-CSP/Bot.thy b/thys/HOL-CSP/Bot.thy --- a/thys/HOL-CSP/Bot.thy +++ b/thys/HOL-CSP/Bot.thy @@ -1,100 +1,100 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * - * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * Author : Burkhart Wolff, Safouan Taha. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) chapter\ The CSP Operators \ section\The Undefined Process\ theory Bot imports Process begin -definition BOT :: "'\ process" -where "BOT \ Abs_process ({(s,X). front_tickFree s}, {d. front_tickFree d})" - -lemma is_process_REP_Bot : - "is_process ({(s,X). front_tickFree s}, {d. front_tickFree d})" -by(auto simp: tickFree_implies_front_tickFree is_process_def - FAILURES_def DIVERGENCES_def - elim: Process.front_tickFree_dw_closed - elim: Process.front_tickFree_append) +lift_definition BOT :: \'\ process\ + is \({(s,X). front_tickFree s}, {d. front_tickFree d})\ + unfolding is_process_def FAILURES_def DIVERGENCES_def + by (auto simp: tickFree_implies_front_tickFree + elim: front_tickFree_dw_closed front_tickFree_append) -lemma Rep_Abs_Bot :"Rep_process (Abs_process ({(s,X). front_tickFree s},{d. front_tickFree d})) = - ({(s,X). front_tickFree s},{d. front_tickFree d})" -by(subst Abs_process_inverse, simp_all only: CollectI Rep_process is_process_REP_Bot) +lemma F_BOT: "\ BOT = {(s,X). front_tickFree s}" + by (simp add: BOT.rep_eq FAILURES_def Failures.rep_eq) -lemma F_Bot: "\ BOT = {(s,X). front_tickFree s}" -by(simp add: BOT_def FAILURES_def Failures_def Rep_Abs_Bot) +lemma D_BOT: "\ BOT = {d. front_tickFree d}" + by (simp add: BOT.rep_eq DIVERGENCES_def Divergences.rep_eq) -lemma D_Bot: "\ BOT = {d. front_tickFree d}" -by(simp add: BOT_def DIVERGENCES_def D_def Rep_Abs_Bot) +lemma T_BOT: "\ BOT = {s. front_tickFree s}" + by (simp add: Traces.rep_eq TRACES_def Failures.rep_eq[symmetric] F_BOT) -lemma T_Bot: "\ BOT = {s. front_tickFree s}" -by(simp add: BOT_def TRACES_def Traces_def FAILURES_def Rep_Abs_Bot) text\ This is the key result: @{term "\"} --- which we know to exist -from the process instantiation --- is equal Bot . +from the process instantiation --- is equal \<^const>\BOT\ . \ lemma BOT_is_UU[simp]: "BOT = \" apply(auto simp: Pcpo.eq_bottom_iff Process.le_approx_def Ra_def min_elems_Collect_ftF_is_Nil Process.Nil_elem_T - F_Bot D_Bot T_Bot + F_BOT D_BOT T_BOT elim: D_imp_front_tickFree) apply(metis Process.is_processT2) done lemma F_UU: "\ \ = {(s,X). front_tickFree s}" - using F_Bot by auto + using F_BOT by auto lemma D_UU: "\ \ = {d. front_tickFree d}" - using D_Bot by auto + using D_BOT by auto lemma T_UU: "\ \ = {s. front_tickFree s}" - using T_Bot by auto + using T_BOT by auto + + +lemma BOT_iff_D: \P = \ \ [] \ \ P\ + apply (intro iffI, simp add: D_UU) + apply (subst Process_eq_spec, safe) + by (simp_all add: F_UU D_UU is_processT2 D_imp_front_tickFree) + (metis append_Nil is_processT tickFree_Nil)+ end diff --git a/thys/HOL-CSP/CSP.thy b/thys/HOL-CSP/CSP.thy --- a/thys/HOL-CSP/CSP.thy +++ b/thys/HOL-CSP/CSP.thy @@ -1,336 +1,430 @@ (*<*) \\ ******************************************************************** * Project : CSP-RefTK - A Refinement Toolkit for HOL-CSP * Version : 1.0 * - * Author : Burkhart Wolff, Safouan Taha, Lina Ye, Benoît Ballenghien, Catherine Dubois + * Author : Burkhart Wolff, Safouan Taha, Lina Ye, Benoît Ballenghien * * This file : A Normalization Theory * * Copyright (c) 2022 Université Paris-Saclay, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) -chapter\ Generalisation of Normalisation of Non Deterministic CSP Processes \ +chapter\ The refinements laws \ theory CSP - imports CSP_Laws Process_Order (*Induction*) + imports Process_Order CSP_Laws CSP_Induct begin section \Monotonicity\ lemma mono_Det_D[simp]: \\P \\<^sub>D P'; S \\<^sub>D S'\ \ (P \ S) \\<^sub>D (P' \ S')\ by (metis D_Det Un_mono divergence_refine_def) lemma mono_Det_T[simp]: \\P \\<^sub>T P'; S \\<^sub>T S'\ \ (P \ S) \\<^sub>T (P' \ S')\ by (metis T_Det Un_mono trace_refine_def) corollary mono_Det_DT[simp]: \\P \\<^sub>D\<^sub>T P'; S \\<^sub>D\<^sub>T S'\ \ (P \ S) \\<^sub>D\<^sub>T (P' \ S')\ by (simp_all add: trace_divergence_refine_def) lemmas mono_Det_FD[simp]= mono_Det_FD[folded failure_divergence_refine_def] \\Deterministic choice monotony doesn't hold for \\\<^sub>F\\ lemma mono_Ndet_F_left[simp]: \P \\<^sub>F Q \ (P \ S) \\<^sub>F Q\ by (simp add: F_Ndet failure_refine_def order_trans) lemma mono_Ndet_D_left[simp]: \P \\<^sub>D Q \ (P \ S) \\<^sub>D Q\ by (simp add: D_Ndet divergence_refine_def order_trans) lemma mono_Ndet_T_left[simp]: \P \\<^sub>T Q \ (P \ S) \\<^sub>T Q\ by (simp add: T_Ndet trace_refine_def order_trans) section \Monotonicity\ lemma mono_Ndet_F[simp]: \\P \\<^sub>F P'; S \\<^sub>F S'\ \ (P \ S) \\<^sub>F (P' \ S')\ by (metis F_Ndet Un_mono failure_refine_def) lemma mono_Ndet_D[simp]: \\P \\<^sub>D P'; S \\<^sub>D S'\ \ (P \ S) \\<^sub>D (P' \ S')\ by (metis D_Ndet Un_mono divergence_refine_def) lemma mono_Ndet_T[simp]: \\P \\<^sub>T P'; S \\<^sub>T S'\ \ (P \ S) \\<^sub>T (P' \ S')\ by (metis T_Ndet Un_mono trace_refine_def) corollary mono_Ndet_DT[simp]: \\P \\<^sub>D\<^sub>T P'; S \\<^sub>D\<^sub>T S'\ \ (P \ S) \\<^sub>D\<^sub>T (P' \ S')\ by (auto simp add: trace_divergence_refine_def) lemmas mono_Ndet_FD[simp]= mono_Ndet_FD[folded failure_divergence_refine_def] corollary mono_Ndet_DT_left[simp]: \P \\<^sub>D\<^sub>T Q \ (P \ S) \\<^sub>D\<^sub>T Q\ and mono_Ndet_F_right[simp]: \P \\<^sub>F Q \ (S \ P) \\<^sub>F Q\ and mono_Ndet_D_right[simp]: \P \\<^sub>D Q \ (S \ P) \\<^sub>D Q\ and mono_Ndet_T_right[simp]: \P \\<^sub>T Q \ (S \ P) \\<^sub>T Q\ and mono_Ndet_DT_right[simp]: \P \\<^sub>D\<^sub>T Q \ (S \ P) \\<^sub>D\<^sub>T Q\ by (simp_all add: trace_divergence_refine_def Ndet_commute) lemmas mono_Ndet_FD_left[simp] = mono_Ndet_FD_left[folded failure_divergence_refine_def] and mono_Ndet_FD_right[simp] = mono_Ndet_FD_right[folded failure_divergence_refine_def] lemma mono_Ndet_Det_FD[simp]: \(P \ S) \\<^sub>F\<^sub>D (P \ S)\ by (metis Det_id failure_divergence_refine_def mono_Det_FD mono_Ndet_FD_left mono_Ndet_FD_right order_refl) corollary mono_Ndet_Det_F[simp]: \(P \ S) \\<^sub>F (P \ S)\ and mono_Ndet_Det_D[simp]: \(P \ S) \\<^sub>D (P \ S)\ and mono_Ndet_Det_T[simp]: \(P \ S) \\<^sub>T (P \ S)\ and mono_Ndet_Det_DT[simp]: \(P \ S) \\<^sub>D\<^sub>T (P \ S)\ by (simp_all add: leFD_imp_leF leFD_imp_leD leF_imp_leT leD_leT_imp_leDT) lemma mono_Seq_F_right[simp]: \S \\<^sub>F S' \ (P \<^bold>; S) \\<^sub>F (P \<^bold>; S')\ - apply (auto simp: failure_refine_def F_Seq append_single_T_imp_tickFree) - using NF_ND by fastforce+ + by (auto simp: failure_refine_def F_Seq append_single_T_imp_tickFree) lemma mono_Seq_D_right[simp]: \S \\<^sub>D S' \ (P \<^bold>; S) \\<^sub>D (P \<^bold>; S')\ by (auto simp: divergence_refine_def D_Seq) lemma mono_Seq_T_right[simp]: \S \\<^sub>T S' \ (P \<^bold>; S) \\<^sub>T (P \<^bold>; S')\ - apply (auto simp: trace_refine_def T_Seq append_single_T_imp_tickFree) - using D_T by fastforce+ + by (auto simp: trace_refine_def T_Seq append_single_T_imp_tickFree) corollary mono_Seq_DT_right[simp]: \S \\<^sub>D\<^sub>T S' \ (P \<^bold>; S) \\<^sub>D\<^sub>T (P \<^bold>; S')\ by (simp add: trace_divergence_refine_def) lemma mono_Seq_DT_left[simp]: \P \\<^sub>D\<^sub>T P' \ (P \<^bold>; S) \\<^sub>D\<^sub>T (P' \<^bold>; S)\ apply (auto simp: trace_divergence_refine_def divergence_refine_def trace_refine_def D_Seq T_Seq) by (metis (no_types, lifting) Nil_elem_T Process.F_T T_F append.right_neutral is_processT5_S3 subset_eq) \\left Sequence monotony doesn't hold for \\\<^sub>F\, \\\<^sub>D\ and \\\<^sub>T\\ corollary mono_Seq_DT[simp]: \P \\<^sub>D\<^sub>T P' \ S \\<^sub>D\<^sub>T S' \ (P \<^bold>; S) \\<^sub>D\<^sub>T (P' \<^bold>; S')\ using mono_Seq_DT_left mono_Seq_DT_right trans_DT by blast lemmas mono_Seq_FD[simp]= mono_Seq_FD[folded failure_divergence_refine_def] lemma mono_Mprefix_F[simp]: \\x \ A. P x \\<^sub>F Q x \ Mprefix A P \\<^sub>F Mprefix A Q\ by (auto simp: failure_refine_def F_Mprefix) blast+ lemma mono_Mprefix_D[simp]: \\x \ A. P x \\<^sub>D Q x \ Mprefix A P \\<^sub>D Mprefix A Q\ by (auto simp: divergence_refine_def D_Mprefix) blast+ lemma mono_Mprefix_T[simp]: \\x \ A. P x \\<^sub>T Q x \ Mprefix A P \\<^sub>T Mprefix A Q\ by (auto simp: trace_refine_def T_Mprefix) corollary mono_Mprefix_DT[simp]: \\x \ A. P x \\<^sub>D\<^sub>T Q x \ Mprefix A P \\<^sub>D\<^sub>T Mprefix A Q\ by (simp add: trace_divergence_refine_def) lemmas mono_Mprefix_FD[simp] = mono_Mprefix_FD[folded failure_divergence_refine_def] lemma mono_Mprefix_DT_set[simp]: \A \ B \ Mprefix B P \\<^sub>D\<^sub>T Mprefix A P\ (* rename this in mono_Mprefix_D_subset ? *) by (auto simp add:T_Mprefix D_Mprefix trace_divergence_refine_def trace_refine_def divergence_refine_def) corollary mono_Mprefix_D_set[simp]: \A \ B \ Mprefix B P \\<^sub>D Mprefix A P\ and mono_Mprefix_T_set[simp]: \A \ B \ Mprefix B P \\<^sub>T Mprefix A P\ by (simp_all add: leDT_imp_leD leDT_imp_leT) \\Mprefix set monotony doesn't hold for \\\<^sub>F\ and \\\<^sub>F\<^sub>D\ where it holds for deterministic choice\ lemma mono_Hiding_F[simp]: \P \\<^sub>F Q \ (P \ A) \\<^sub>F (Q \ A)\ apply(cases \A = {}\, simp_all add: Hiding_set_empty failure_refine_def F_Hiding, intro conjI, blast) proof(subst (2) Un_commute, rule subsetI, rule UnCI, auto, goal_cases) case (1 b t u) then obtain a where a:\a \ A\ by blast define f where A: \f = rec_nat t (\i t. t @ [ev a])\ hence AA: \f (Suc i) = (f i) @ [ev a]\ for i by simp hence B: \strict_mono f\ by (simp add:strict_mono_def lift_Suc_mono_less_iff) from A have C: \t \ range f\ by (metis (mono_tags, lifting) old.nat.simps(6) range_eqI) { fix i have \f i \ \ Q \ tickFree (f i) \ trace_hide (f i) (ev ` A) = (trace_hide t (ev ` A))\ proof(induct i) case 0 then show ?case by (simp add: 1(4) 1(7) A) next case (Suc i) then show ?case apply (simp add:AA a) using is_processT7[rule_format, of \f i\ Q \[ev a]\] front_tickFree_single by blast qed } with B C have \isInfHiddenRun f P A \ t \ range f\ by (metis 1(1) D_T F_subset_imp_T_subset subsetD) with 1 show ?case by metis next case (2 b u f x) then show ?case using F_subset_imp_T_subset by blast qed lemma mono_Hiding_T[simp]: \P \\<^sub>T Q \ (P \ A) \\<^sub>T (Q \ A)\ apply(cases \A={}\, simp add: Hiding_set_empty, simp add:trace_refine_def T_Hiding, intro conjI) proof(goal_cases) case 1 then show ?case proof(subst Un_commute, rule_tac subsetI, rule_tac UnCI, auto, goal_cases) case 11:(1 t a) hence tt: \t \ \ P\ by (meson Process.F_T subset_eq) with 11(1) inf_hidden[of A t P] obtain f where \isInfHiddenRun f P A \ t \ range f\ by auto with 11(4)[rule_format, of t \[]\] show ?case by (metis 11(1) tt append_Nil2 front_tickFree_Nil is_processT2_TR nonTickFree_n_frontTickFree tick_T_F) qed next case 2 then show ?case proof(subst Un_commute, auto, goal_cases) case 21:(1 t u a) define f where A: \f = rec_nat t (\i t. t @ [ev a])\ hence AA: \f (Suc i) = (f i) @ [ev a]\ for i by simp hence B: \strict_mono f\ by (simp add:strict_mono_def lift_Suc_mono_less_iff) from A have C: \t \ range f\ by (metis (mono_tags, lifting) old.nat.simps(6) range_eqI) { fix i have \f i \ \ Q \ tickFree (f i) \ trace_hide (f i) (ev ` A) = (trace_hide t (ev ` A))\ proof(induct i) case 0 then show ?case by (simp add: 21(4) 21(7) A) next case (Suc i) then show ?case apply (simp add:AA 21(6)) using is_processT7[rule_format, of \f i\ Q \[ev a]\] front_tickFree_single by blast qed } with B C have \isInfHiddenRun f P A \ t \ range f\ by (metis 21(1) D_T subsetD) with 21 show ?case by metis next case 22:(2 b u f x) then show ?case by blast qed qed lemma mono_Hiding_DT[simp]: \P \\<^sub>D\<^sub>T Q \ (P \ A) \\<^sub>D\<^sub>T (Q \ A)\ proof - assume as: \P \\<^sub>D\<^sub>T Q\ then have \(P \ A) \\<^sub>D (Q \ A)\ apply(auto simp:trace_divergence_refine_def divergence_refine_def trace_refine_def D_Hiding T_Hiding) by blast+ with leDT_imp_leT[THEN mono_Hiding_T, OF as] show ?thesis by (simp add: trace_divergence_refine_def) qed lemmas mono_Hiding_FD[simp] = mono_Hiding_FD[folded failure_divergence_refine_def] \\Obviously, Hide monotony doesn't hold for \\\<^sub>D\\ lemma mono_Sync_DT[simp]: \P \\<^sub>D\<^sub>T P' \ Q \\<^sub>D\<^sub>T Q' \ (P \ A \ Q) \\<^sub>D\<^sub>T (P' \ A \ Q')\ by (simp add:trace_divergence_refine_def divergence_refine_def trace_refine_def T_Sync D_Sync) blast lemmas mono_Sync_FD[simp] = mono_Sync_FD[folded failure_divergence_refine_def] \\Synchronization monotony doesn't hold for \\\<^sub>F\, \\\<^sub>D\ and \\\<^sub>T\\ -(* choose between mono_Mndetprefix_F and mono_mndet_F *) lemma mono_Mndetprefix_F[simp]: \\x\A. P x \\<^sub>F Q x \ Mndetprefix A P \\<^sub>F Mndetprefix A Q\ by (cases \A = {}\, auto simp: failure_refine_def F_Mndetprefix write0_def F_Mprefix) lemma mono_Mndetprefix_D[simp]: \\x\A. P x \\<^sub>D Q x \ Mndetprefix A P \\<^sub>D Mndetprefix A Q\ by (cases \A = {}\, auto simp: divergence_refine_def D_Mndetprefix write0_def D_Mprefix) lemma mono_Mndetprefix_T[simp]: \\x\A. P x \\<^sub>T Q x \ Mndetprefix A P \\<^sub>T Mndetprefix A Q\ by (cases \A = {}\, auto simp: trace_refine_def T_Mndetprefix write0_def T_Mprefix) corollary mono_Mndetprefix_DT[simp]: \\x\A. P x \\<^sub>D\<^sub>T Q x \ Mndetprefix A P \\<^sub>D\<^sub>T Mndetprefix A Q\ by (simp add: trace_divergence_refine_def) lemmas mono_Mndetprefix_FD[simp] = mono_Mndetprefix_FD[folded failure_divergence_refine_def] lemmas mono_Mndetprefix_FD_set[simp] = Mndetprefix_FD_subset[folded failure_divergence_refine_def] corollary mono_Mndetprefix_F_set[simp] : \A \ {} \ A \ B \ Mndetprefix B P \\<^sub>F Mndetprefix A P\ - (* removed nonempty hypothesis in the 3 following propositions *) and mono_Mndetprefix_D_set[simp] : \A \ B \ Mndetprefix B P \\<^sub>D Mndetprefix A P\ and mono_Mndetprefix_T_set[simp] : \A \ B \ Mndetprefix B P \\<^sub>T Mndetprefix A P\ and mono_Mndetprefix_DT_set[simp]: \A \ B \ Mndetprefix B P \\<^sub>D\<^sub>T Mndetprefix A P\ by (cases \A = {}\, simp_all add: leFD_imp_leF leFD_imp_leD leF_imp_leT leD_leT_imp_leDT)+ lemmas Mprefix_refines_Mndetprefix_FD[simp] = Mprefix_refines_Mndetprefix[folded failure_divergence_refine_def] corollary Mprefix_refines_Mndetprefix_F[simp] : \Mndetprefix A P \\<^sub>F Mprefix A P\ and Mprefix_refines_Mndetprefix_D[simp] : \Mndetprefix A P \\<^sub>D Mprefix A P\ and Mprefix_refines_Mndetprefix_T[simp] : \Mndetprefix A P \\<^sub>T Mprefix A P\ and Mprefix_refines_Mndetprefix_DT[simp]: \Mndetprefix A P \\<^sub>D\<^sub>T Mprefix A P\ by (simp_all add: leFD_imp_leF leFD_imp_leD leF_imp_leT leD_leT_imp_leDT) -(* see where we put this *) + lemma mono_read_FD[simp, elim]: "(\x. P x \\<^sub>F\<^sub>D Q x) \ (c\<^bold>?x \ (P x)) \\<^sub>F\<^sub>D (c\<^bold>?x \ (Q x))" by (simp add: read_def) lemma mono_write_FD[simp, elim]: "(P \\<^sub>F\<^sub>D Q) \ (c\<^bold>!x \ P) \\<^sub>F\<^sub>D (c\<^bold>!x \ Q)" by (simp add: write_def) lemma mono_write0_FD[simp, elim]: "P \\<^sub>F\<^sub>D Q \ (a \ P) \\<^sub>F\<^sub>D (a \ Q)" by (simp add: write0_def) (* make this with \\<^sub>D \\<^sub>T and \\<^sub>F ? *) lemma mono_read_DT[simp, elim]: "(\x. P x \\<^sub>D\<^sub>T Q x) \ (c\<^bold>?x \ (P x)) \\<^sub>D\<^sub>T (c\<^bold>?x \ (Q x))" by (simp add: read_def) lemma mono_write_DT[simp, elim]: "(P \\<^sub>D\<^sub>T Q) \ (c\<^bold>!x \ P) \\<^sub>D\<^sub>T (c\<^bold>!x \ Q)" by (simp add: write_def) lemma mono_write0_DT[simp, elim]: "P \\<^sub>D\<^sub>T Q \ (a \ P) \\<^sub>D\<^sub>T (a \ Q)" by (simp add: write0_def) -\ \The following result is very useful, but we are not sure about where is its place\ + +lemma mono_Renaming_D: \P \\<^sub>D Q \ Renaming P f \\<^sub>D Renaming Q f\ + unfolding divergence_refine_def D_Renaming by blast + + +lemma mono_Renaming_FD: \P \\<^sub>F\<^sub>D Q \ Renaming P f \\<^sub>F\<^sub>D Renaming Q f\ + unfolding failure_divergence_refine_def le_ref_def + apply (simp add: mono_Renaming_D[unfolded divergence_refine_def]) + unfolding F_Renaming D_Renaming by blast + + +lemma mono_Renaming_DT: \P \\<^sub>D\<^sub>T Q \ Renaming P f \\<^sub>D\<^sub>T Renaming Q f\ + unfolding trace_divergence_refine_def + apply (simp add: mono_Renaming_D) + unfolding trace_refine_def divergence_refine_def T_Renaming by blast + + + + +text \Some new and very useful results\ + +lemma Ndet_is_STOP_iff: \P \ Q = STOP \ P = STOP \ Q = STOP\ + using Nil_subset_T by (simp add: STOP_iff_T T_Ndet, blast) + +lemma Det_is_STOP_iff: \P \ Q = STOP \ P = STOP \ Q = STOP\ + using Nil_subset_T by (simp add: STOP_iff_T T_Det, blast) + + +lemma Det_is_BOT_iff: \P \ Q = \ \ P = \ \ Q = \\ + by (simp add: BOT_iff_D D_Det) + +lemma Ndet_is_BOT_iff: \P \ Q = \ \ P = \ \ Q = \\ + by (simp add: BOT_iff_D D_Ndet) + +lemma Sync_is_BOT_iff: \P \S\ Q = \ \ P = \ \ Q = \\ + by (rule HOL.sym, intro iffI, metis Sync_BOT Sync_commute) + (use empty_setinterleaving in \auto simp add: BOT_iff_D D_Sync\) + + +lemma STOP_neq_BOT: \STOP \ \\ + by (simp add: D_STOP BOT_iff_D) + +lemma SKIP_neq_BOT: \SKIP \ \\ + by (simp add: D_SKIP BOT_iff_D) + + +lemma Mprefix_neq_BOT: \Mprefix A P \ \\ + by (simp add: BOT_iff_D) + +lemma Mndetprefix_neq_BOT: \Mndetprefix A P \ \\ + by (cases \A = {}\) (simp_all add: D_STOP BOT_iff_D D_Mndetprefix write0_def) + + +lemma STOP_T_iff: \STOP \\<^sub>T P \ P = STOP\ + by auto (metis Nil_elem_T STOP_iff_T empty_iff subset_singletonD trace_refine_def) + +lemma STOP_F_iff: \STOP \\<^sub>F P \ P = STOP\ + by auto (metis Nil_elem_T STOP_iff_T empty_iff leF_imp_leT subset_singletonD trace_refine_def) + +lemma STOP_FD_iff: \STOP \\<^sub>F\<^sub>D P \ P = STOP\ + using STOP_F_iff idem_FD leFD_imp_leF by blast + +lemma SKIP_FD_iff: \SKIP \\<^sub>F\<^sub>D P \ P = SKIP\ + apply (subst Process_eq_spec, auto simp add: failure_divergence_refine_def le_ref_def F_SKIP D_SKIP subset_iff) + by (metis F_T insertI1 is_processT5_S6 is_processT6_S2) + (metis F_T append.left_neutral insertI1 is_processT5_S6 tick_T_F) + +lemma SKIP_F_iff: \SKIP \\<^sub>F P \ P = SKIP\ + apply (subst Process_eq_spec, auto simp add: failure_refine_def le_ref_def F_SKIP D_SKIP subset_iff) + apply (metis F_T insertI1 is_processT5_S6 is_processT6_S2) + apply (metis F_T append.left_neutral insertI1 is_processT5_S6 tick_T_F) + by (metis append_Nil insertI1 neq_Nil_conv process_charn) + + + +lemma Seq_is_SKIP_iff: \P \<^bold>; Q = SKIP \ P = SKIP \ Q = SKIP\ + apply (intro iffI, simp_all add: Seq_SKIP) + apply (simp add: SKIP_F_iff[symmetric]) + unfolding failure_refine_def +proof (auto simp add: F_Seq F_SKIP D_Seq D_SKIP subset_iff intro: T_F, goal_cases) + case (1 s X) + thus ?case + apply (cases \tickFree s\) + apply (erule_tac x = s in allE, + metis F_T append.right_neutral is_processT4_empty is_processT5_S4 process_charn) + by (erule_tac x = \butlast s\ in allE, + metis F_T append.right_neutral append_butlast_last_id append_single_T_imp_tickFree + last_snoc nonTickFree_n_frontTickFree non_tickFree_tick process_charn self_append_conv2) +next + case (2 s X) + thus ?case + by (metis F_T append_Nil2 append_self_conv2 butlast_snoc front_tickFree_butlast + insert_Diff insert_Diff_single nonTickFree_n_frontTickFree non_tickFree_tick process_charn) +qed (metis F_T append.left_neutral insertI1 insert_absorb2 is_processT5_S6 tickFree_Nil)+ + + + + +\ \The following result is very useful (especially for ProcOmata).\ lemma cont_process_rec: \P = (\ X. f X) \ cont f \ P = f P\ by (simp add: def_cont_fix_eq) end \ No newline at end of file diff --git a/thys/HOL-CSP/CSP_Induct.thy b/thys/HOL-CSP/CSP_Induct.thy new file mode 100644 --- /dev/null +++ b/thys/HOL-CSP/CSP_Induct.thy @@ -0,0 +1,145 @@ +(*<*) +\\ ******************************************************************** + * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL + * Version : 2.0 + * + * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * + * This file : More Fixpoint and k-Induction Schemata + * + * Copyright (c) 2020 Université Paris-Saclay, France + * + * All rights reserved. + * + * Redistribution and use in source and binary forms, with or without + * modification, are permitted provided that the following conditions are + * met: + * + * * Redistributions of source code must retain the above copyright + * notice, this list of conditions and the following disclaimer. + * + * * Redistributions in binary form must reproduce the above + * copyright notice, this list of conditions and the following + * disclaimer in the documentation and/or other materials provided + * with the distribution. + * + * * Neither the name of the copyright holders nor the names of its + * contributors may be used to endorse or promote products derived + * from this software without specific prior written permission. + * + * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS + * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT + * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR + * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT + * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, + * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT + * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, + * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY + * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT + * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE + * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. + ******************************************************************************\ +(*>*) + +chapter \ Advanced Induction Schemata \ + +theory CSP_Induct + +imports HOLCF + +begin + + +default_sort type + +section \k-fixpoint-induction\ + +lemma nat_k_induct [case_names base step]: + fixes k::nat + assumes "\in\<^sub>0. (\i0+i)) \ P (n\<^sub>0+k)" + shows "P (n::nat)" +proof (induct rule: nat_less_induct) + case (1 n) + then show ?case + apply(cases "n < k") + using assms(1) apply blast + using assms(2)[rule_format, of "n-k"] by auto +qed + +thm fix_ind fix_ind2 + +lemma fix_ind_k [case_names admissibility base_k_steps step]: + fixes k::nat + assumes adm: "adm P" + and base_k_steps: "\if\\)" + and step: "\x. (\if\x)) \ P (iterate k\f\x)" + shows "P (fix\f)" + unfolding fix_def2 apply (rule admD [OF adm chain_iterate]) + apply(rule nat_k_induct[of k], simp add: base_k_steps) + using step by (subst (1 2) add.commute, unfold iterate_iterate[symmetric]) blast + +lemma nat_k_skip_induct [case_names lower_bound base_k step]: + fixes k::nat + assumes "k \ 1" and "\in\<^sub>0. P (n\<^sub>0) \ P (n\<^sub>0+k)" + shows "P (n::nat)" +proof(induct rule:nat_less_induct) + case (1 n) + then show ?case + apply(cases "n < k") + using assms(2) apply blast + using assms(3)[rule_format, of "n-k"] assms(1) by auto +qed + +lemma fix_ind_k_skip [case_names lower_bound admissibility base_k_steps step]: + fixes k::nat + assumes k_1: "k \ 1" + and adm: "adm P" + and base_k_steps: "\if\\)" + and step: "\x. P x \ P (iterate k\f\x)" + shows "P (fix\f)" + unfolding fix_def2 apply (rule admD [OF adm chain_iterate]) + apply(rule nat_k_skip_induct[of k]) + using k_1 base_k_steps apply auto + using step by (subst add.commute, unfold iterate_iterate[symmetric]) blast + +thm parallel_fix_ind + +section\Parallel fixpoint-induction\ + +lemma parallel_fix_ind_inc[case_names admissibility base_fst base_snd step]: + assumes adm: "adm (\x. P (fst x) (snd x))" + and base_fst: "\y. P \ y" and base_snd: "\x. P x \" + and step: "\x y. P x y \ P (G\x) y \ P x (H\y) \ P (G\x) (H\y)" + shows "P (fix\G) (fix\H)" +proof - + from adm have adm': "adm (case_prod P)" + unfolding split_def . + have "P (iterate i\G\\) (iterate j\H\\)" for i j + proof(induct "i+j" arbitrary:i j rule:nat_less_induct) + case 1 + { fix i' j' + assume i:"i = Suc i'" and j:"j = Suc j'" + have "P (iterate i'\G\\) (iterate j'\H\\)" + and "P (iterate i'\G\\) (iterate j\H\\)" + and "P (iterate i\G\\) (iterate j'\H\\)" + using "1.hyps" add_strict_mono i j apply blast + using "1.hyps" i apply auto[1] + using "1.hyps" j by auto + hence ?case by (simp add: i j step) + } + thus ?case + apply(cases i, simp add:base_fst) + apply(cases j, simp add:base_snd) + by assumption + qed + then have "\i. case_prod P (iterate i\G\\, iterate i\H\\)" + by simp + then have "case_prod P (\i. (iterate i\G\\, iterate i\H\\))" + by - (rule admD [OF adm'], simp, assumption) + then have "P (\i. iterate i\G\\) (\i. iterate i\H\\)" + by (simp add: lub_Pair) + then show "P (fix\G) (fix\H)" + by (simp add: fix_def2) +qed + +end \ No newline at end of file diff --git a/thys/HOL-CSP/CSP_Laws.thy b/thys/HOL-CSP/CSP_Laws.thy --- a/thys/HOL-CSP/CSP_Laws.thy +++ b/thys/HOL-CSP/CSP_Laws.thy @@ -1,4605 +1,3918 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * * Author : Burkhart Wolff, Safouan Taha, Lina Ye, Benoît Ballenghien. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) chapter\ The "Laws" of CSP \ theory CSP_Laws - imports Bot Skip Stop Det Ndet Mprefix Mndetprefix Seq Hiding Sync + imports Bot Skip Stop Det Ndet Mprefix Mndetprefix Seq Hiding Sync Renaming "HOL-Eisbach.Eisbach" begin method exI for y::'a = (rule_tac exI[where x = y]) section\ General Laws\ lemma SKIP_Neq_STOP: "SKIP \ STOP" by (auto simp: Process_eq_spec F_SKIP F_STOP D_SKIP D_STOP Un_def) lemma BOT_less1[simp]: "\ \ (X::'a process)" by (simp add: le_approx_implies_le_ref) lemma BOT_less2[simp]: "BOT \ (X::'a process)" by simp section\ Deterministic Choice Operator Laws \ lemma mono_Det_FD_onside[simp]: " P \ P' \ (P \ S) \ (P' \ S)" unfolding le_ref_def F_Det D_Det using F_subset_imp_T_subset by blast lemma mono_Det_FD[simp]: " \P \ P'; S \ S'\ \ (P \ S) \ (P' \ S')" by (metis Det_commute dual_order.trans mono_Det_FD_onside) lemma mono_Det_ref: " \P \ P'; S \ S'\ \ (P \ S) \ (P' \ S')" using below_trans mono_Det mono_Det_sym by blast lemma Det_BOT : "(P \ \) = \" by (auto simp add:Process_eq_spec D_Det F_Det is_processT2 D_imp_front_tickFree F_UU D_UU) lemma Det_STOP: "(P \ STOP) = P" by (auto simp add: Process_eq_spec D_Det F_Det D_STOP F_STOP T_STOP Un_def Sigma_def is_processT8 is_processT6_S2) lemma Det_id: "(P \ P) = P" by (auto simp: Process_eq_spec D_Det F_Det Un_def Sigma_def is_processT8 is_processT6_S2) lemma Det_assoc: "((M \ P) \ Q) = (M \ (P \ Q))" by (auto simp add: Process_eq_spec D_Det F_Det Un_def Sigma_def T_Det) section\ NonDeterministic Choice Operator Laws \ lemma mono_Ndet_FD[simp]: " \P \ P'; S \ S'\ \ (P \ S) \ (P' \ S')" by (auto simp: le_ref_def F_Ndet D_Ndet) lemma mono_Ndet_FD_left[simp]: "P \ Q \ (P \ S) \ Q" by (metis D_Ndet F_Ndet dual_order.trans le_ref_def le_sup_iff order_refl) lemma mono_Ndet_FD_right[simp]: "P \ Q \ (S \ P) \ Q" by (metis D_Ndet F_Ndet dual_order.trans le_ref_def le_sup_iff order_refl) lemma mono_Ndet_ref: " \P \ P'; S \ S'\ \ (P \ S) \ (P' \ S')" using below_trans mono_Ndet mono_Ndet_sym by blast lemma Ndet_BOT: "(P \ \) = \" by (auto simp: Process_eq_spec D_Ndet F_Ndet is_processT2 D_imp_front_tickFree F_UU D_UU) lemma Ndet_id: "(P \ P) = P" by (simp_all add: F_Ndet D_Ndet Process_eq_spec) lemma Ndet_assoc: "((M \ P) \ Q) = (M \ (P \ Q))" by (simp_all add: F_Ndet D_Ndet Process_eq_spec Un_assoc) subsection\ Multi-Operators laws \ lemma Det_distrib: "(M \ (P \ Q)) = ((M \ P) \ (M \ Q))" by (auto simp add: Process_eq_spec F_Det D_Det F_Ndet D_Ndet Un_def T_Ndet) lemma Ndet_distrib: "(M \ (P \ Q)) = ((M \ P) \ (M \ Q))" by (auto simp add: Process_eq_spec F_Det D_Det F_Ndet D_Ndet Un_def T_Ndet is_processT8 is_processT6_S2) lemma Ndet_FD_Det: "(P \ Q) \ (P \ Q)" apply(auto simp add:le_ref_def D_Ndet D_Det F_Ndet F_Det T_Ndet T_Det Ra_def min_elems_def) using is_processT6_S2 NF_ND by blast+ section\ Sequence Operator Laws \ subsection\Preliminaries\ definition F_minus_D_Seq where "F_minus_D_Seq P Q \ {(t, X). (t, X \ {tick}) \ \ P \ tickFree t} \ {(t, X). \t1 t2. t = t1 @ t2 \ t1 @ [tick] \ \ P \ (t2, X) \ \ Q}" lemma F_minus_D_Seq_opt: "(a,b) \ \(P \<^bold>; Q) = (a \ \(P \<^bold>; Q) \ (a,b) \ F_minus_D_Seq P Q)" - by (auto simp add: F_Seq D_Seq F_minus_D_Seq_def) + using NF_ND by (auto simp add: F_Seq D_Seq F_minus_D_Seq_def) blast+ + lemma Process_eq_spec_optimized_Seq : "((P \<^bold>; Q) = (U \<^bold>; S)) = (\ (P \<^bold>; Q) = \ (U \<^bold>; S) \ F_minus_D_Seq P Q \ \ (U \<^bold>; S) \ F_minus_D_Seq U S \ \ (P \<^bold>; Q))" unfolding Process_eq_spec_optimized[of "(P \<^bold>; Q)" "(U \<^bold>; S)"] by (auto simp:F_minus_D_Seq_opt) subsection\Laws\ lemma mono_Seq_FD[simp]: " \P \ P'; S \ S'\ \ (P \<^bold>; S) \ (P' \<^bold>; S')" apply (auto simp: le_ref_def F_Seq D_Seq) by (metis F_subset_imp_T_subset subsetCE)+ lemma mono_Seq_ref: " \P \ P'; S \ S'\ \ (P \<^bold>; S) \ (P' \<^bold>; S')" using below_trans mono_Seq mono_Seq_sym by blast lemma BOT_Seq: "(\ \<^bold>; P) = \" - apply(auto simp add: Process_eq_spec D_Seq F_Seq front_tickFree_append F_UU D_UU T_UU) - using front_tickFree_append front_tickFree_implies_tickFree is_processT2 apply blast - using D_imp_front_tickFree front_tickFree_append front_tickFree_implies_tickFree apply blast - using front_tickFree_charn tickFree_Nil apply blast - using D_imp_front_tickFree front_tickFree_append front_tickFree_implies_tickFree apply blast - using front_tickFree_Nil tickFree_Nil by blast + by (simp add: BOT_iff_D D_Seq D_UU) + lemma Seq_SKIP: "(P \<^bold>; SKIP) = P" - apply (auto simp add: Process_eq_spec F_Seq D_Seq F_SKIP D_SKIP is_processT7 is_processT8_S - T_F_spec is_processT6_S1) - apply (meson insertI2 is_processT4 subsetI) - apply (meson append_T_imp_tickFree is_processT5_S7 non_tickFree_tick not_Cons_self2 tickFree_append) - using T_F_spec insert_absorb is_processT5_S2 apply fastforce - apply (metis F_T is_processT nonTickFree_n_frontTickFree) - by (metis append_Nil2 front_tickFree_mono is_processT nonTickFree_n_frontTickFree not_Cons_self) + apply (auto simp add: Process_eq_spec F_Seq F_SKIP D_Seq D_SKIP T_F_spec is_processT6_S1) + apply (use is_processT4 in blast) + apply (meson append_single_T_imp_tickFree is_processT5_S7 non_tickFree_tick tickFree_append) + apply (meson is_processT8) + apply (use T_F_spec insert_absorb is_processT5_S2 in fastforce) + by (metis F_T is_processT nonTickFree_n_frontTickFree) lemma SKIP_Seq: "(SKIP \<^bold>; P) = P" by (auto simp add: Process_eq_spec D_Seq T_SKIP F_Seq F_SKIP D_SKIP is_processT8_Pair) lemma STOP_Seq: "(STOP \<^bold>; P) = STOP" - by (auto simp: Process_eq_spec F_Seq D_Seq F_STOP D_STOP T_STOP Un_def) + by (simp add: STOP_iff_T T_Seq T_STOP F_STOP D_STOP Collect_conv_if) + subsection\ Multi-Operators laws \ lemma Seq_Ndet_distrR: "((P \ Q) \<^bold>; S) = ((P \<^bold>; S) \ (Q \<^bold>; S))" by (auto simp add: Process_eq_spec D_Seq D_Ndet T_Ndet Un_def F_Seq T_Seq F_Ndet) lemma Seq_Ndet_distrL: "(P \<^bold>; (Q \ S)) = ((P \<^bold>; Q) \ (P \<^bold>; S))" by (auto simp add: Process_eq_spec D_Seq D_Ndet T_Ndet Un_def F_Seq F_Ndet) lemma Seq_assoc_D: "\ (P \<^bold>; (Q \<^bold>; S)) = \ ((P \<^bold>; Q) \<^bold>; S)" proof(safe, goal_cases) - case (1 x) - then show ?case - apply(auto simp add:D_Seq T_Seq) - using front_tickFree_Nil apply blast - apply (metis append.assoc append_single_T_imp_tickFree tickFree_append) - by (metis append.assoc) + show \s \ \ (P \<^bold>; (Q \<^bold>; S)) \ s \ \ (P \<^bold>; Q \<^bold>; S)\ for s + by (simp add: D_Seq T_Seq) (metis append.assoc) next - case (2 x) - then show ?case + fix s + assume \s \ \ (P \<^bold>; Q \<^bold>; S)\ + thus \s \ \ (P \<^bold>; (Q \<^bold>; S))\ proof(auto simp add:D_Seq T_Seq, goal_cases) - case (1 t2 t1a t2a) - then show ?case using front_tickFree_append by fastforce + case (1 t1 t2 t1a t2a) + from "1"(1)[rule_format, OF "1"(5)] show ?case + apply (cases \tickFree t2a\) + apply (metis "1"(4) "1"(5) append_single_T_imp_tickFree non_tickFree_tick tickFree_append) + by (metis "1"(3) "1"(4) "1"(6) T_nonTickFree_imp_decomp append.assoc butlast_snoc) next - case (2 t1 t2 t1a t2a) - then obtain t2b where "t2a = t2b@[tick]" - by (metis T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append) - with 2 show ?case by auto - next - case (3 t1 t2 t1a t2a) - then show ?case by (metis front_tickFree_implies_tickFree process_charn) - next - case (4 t1 t2 t1a t2a) - then obtain t2b where "t2a = t2b@[tick]" - by (metis D_T T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append) - with 4 show ?case - by (metis D_imp_front_tickFree append.assoc butlast_snoc front_tickFree_implies_tickFree process_charn) - next - case (5 t1 t2 t1a t2a) - then show ?case by (metis front_tickFree_implies_tickFree process_charn) - next - case (6 t1 t2 t1a t2a) - then obtain t2b where "t2a = t2b@[tick]" - by (metis D_T T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append) - with 6 show ?case - by (metis D_imp_front_tickFree append.assoc butlast_snoc front_tickFree_implies_tickFree process_charn) + case (2 t1 t2) + then show ?case + by (meson D_imp_front_tickFree front_tickFree_implies_tickFree is_processT7 is_processT9_S_swap) qed qed lemma Seq_assoc: "(P \<^bold>; (Q \<^bold>; S)) = ((P \<^bold>; Q) \<^bold>; S)" proof (auto simp: Process_eq_spec_optimized_Seq Seq_assoc_D, goal_cases) case (1 a b) then show ?case - proof(auto simp add:F_minus_D_Seq_def Seq_assoc_D F_minus_D_Seq_opt append_single_T_imp_tickFree del:conjI, + proof(auto simp add:F_minus_D_Seq_def Seq_assoc_D F_minus_D_Seq_opt append_single_T_imp_tickFree + del:conjI, goal_cases 11 12) case (11 t1 t2) then show ?case by (metis (mono_tags, lifting) D_Seq Seq_assoc_D UnCI mem_Collect_eq) next case (12 t1 t1a t2a) hence "(t1 @ t1a) @ [tick] \ \ (P \<^bold>; Q)" by (auto simp:T_Seq) then show ?case using 12(2)[rule_format, of "t1@t1a" t2a] 12(4,5,6) by simp qed next case (2 a b) then show ?case - proof(auto simp add:F_minus_D_Seq_def Seq_assoc_D F_minus_D_Seq_opt append_single_T_imp_tickFree T_Seq del:conjI, - goal_cases 21 22 23 24 25 26) + proof(auto simp add:F_minus_D_Seq_def Seq_assoc_D F_minus_D_Seq_opt + append_single_T_imp_tickFree T_Seq del:conjI, + goal_cases 21 22 23) case 21 - then show ?case - by (metis (mono_tags, lifting) D_Seq UnCI append_Nil2 front_tickFree_Nil mem_Collect_eq) + then show ?case + using D_Seq by force next case (22 t1 t2 t1a t2a) then obtain t2b where "t2a = t2b@[tick]" by (metis T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append) with 22 show ?case using append.assoc butlast_snoc by auto next - case (23 t1 t2 t1a t2a) + case (23 t1 t2) hence "t1 \ \ (P \<^bold>; (Q \<^bold>; S))" - by (simp add:D_Seq) (metis append_Nil2 front_tickFree_implies_tickFree process_charn) + by (simp add: D_Seq) (meson is_processT9_S_swap) with 23 Seq_assoc_D show ?case by (metis front_tickFree_implies_tickFree process_charn) - next - case (24 t1 t2 t1a t2a) - then obtain t2b where "t2a = t2b@[tick]" - by (metis D_T T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append) - with 24 show ?case by (metis D_T append.assoc butlast_snoc) - next - case (25 t1 t2 t1a t2a) - hence "t1 \ \ (P \<^bold>; (Q \<^bold>; S))" - by (simp add:D_Seq) (metis append_Nil2 front_tickFree_implies_tickFree process_charn) - with 25 Seq_assoc_D show ?case by (metis front_tickFree_implies_tickFree process_charn) - next - case (26 t1 t2 t1a t2a) - then obtain t2b where "t2a = t2b@[tick]" - by (metis D_T T_nonTickFree_imp_decomp append_single_T_imp_tickFree non_tickFree_tick tickFree_append) - with 26 show ?case by (metis D_T append.assoc butlast_snoc) qed qed section\ The Multi-Prefix Operator Laws \ lemma mono_Mprefix_FD[simp]: "\x \ A. P x \ P' x \ Mprefix A P \ Mprefix A P'" by (auto simp: le_ref_def F_Mprefix D_Mprefix) blast+ lemmas mono_Mprefix_ref = mono_Mprefix0 lemma Mprefix_STOP: "(Mprefix {} P) = STOP" by (auto simp:Process_eq_spec F_Mprefix D_Mprefix D_STOP F_STOP) subsection\ Multi-Operators laws \ lemma Mprefix_Un_distrib: "(Mprefix (A \ B) P) = ((Mprefix A P) \ (Mprefix B P))" apply (unfold Process_eq_spec, rule conjI) apply (auto, (elim disjE conjE | simp_all add: F_Det F_Mprefix Un_def image_def)+, auto) by (auto simp add: D_Det D_Mprefix Un_def) lemma Mprefix_Seq: "((Mprefix A P) \<^bold>; Q) = (Mprefix A (\x. (P x) \<^bold>; Q))" proof (unfold Process_eq_spec, rule conjI, rule subset_antisym, goal_cases) case 1 then show ?case apply(unfold split_def F_Seq D_Seq F_Mprefix T_Mprefix D_Mprefix) apply(rule subsetI, simp_all, elim disjE conjE exE) apply(rule disjI1, simp, meson tickFree_tl) apply (rule disjI2, rule conjI, simp) apply auto[1] apply (auto simp add:hd_append)[1] - using tickFree_tl apply fastforce - by (auto simp add: hd_append)[1] + using tickFree_tl by fastforce next case 2 then show ?case apply(unfold split_def F_Seq D_Seq F_Mprefix T_Mprefix D_Mprefix) apply(rule subsetI, simp_all, elim disjE conjE exE) apply(rule disjI1, simp, blast) apply(rule disjI1, metis event.simps(3) list.exhaust_sel tickFree_Cons) proof(goal_cases) case (1 x a t1 t2) then show ?case apply(rule_tac disjI2, rule_tac disjI1, rule_tac x="(ev a)#t1" in exI) using hd_Cons_tl image_iff by fastforce next - case (2 x a t1 t2) - then show ?case - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI1, rule_tac x="(ev a)#t1" in exI) - using hd_Cons_tl image_iff by fastforce - next - case (3 x a t1 t2) - then show ?case - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac x="(ev a)#t1" in exI) - using hd_Cons_tl image_iff by fastforce + case (2 x a) + then show ?case + by (metis prod.collapse) qed next case 3 then show ?case apply (auto simp add: D_Mprefix D_Seq T_Mprefix) - using tickFree_tl apply blast - apply (metis event.distinct(1) hd_append image_iff list.sel(1)) - apply (metis event.distinct(1) hd_append list.sel(1) tl_append2) - apply (metis (no_types, opaque_lifting) append_Cons event.distinct(1) image_eqI list.distinct(1) - list.exhaust_sel list.sel(1) list.sel(3) tickFree_Cons) - by (metis append_Cons list.exhaust_sel list.sel(1) list.sel(3)) + apply (metis event.distinct(1) hd_append image_iff list.sel(1)) + apply (metis event.distinct(1) hd_append list.sel(1) tl_append2) + by (metis Cons_eq_appendI hd_Cons_tl list.sel(1) list.sel(3)) qed subsection\ Derivative Operators laws \ lemma Mprefix_singl: "(Mprefix {a} P) = (a \ (P a))" by (simp add:write0_def Mprefix_def, rule arg_cong[of _ _ "\x. Abs_process x"]) fastforce lemma mono_read_FD: "(\x. P x \ Q x) \ (c\<^bold>?x \ (P x)) \ (c\<^bold>?x \ (Q x))" by (simp add: read_def) lemma mono_write_FD: "(P \ Q) \ (c\<^bold>!x \ P) \ (c\<^bold>!x \ Q)" by (simp add: write_def) lemma mono_write0_FD: "P \ Q \ (a \ P) \ (a \ Q)" by (simp add: write0_def) lemma mono_read_ref: "(\x. P x \ Q x) \ (c\<^bold>?x \ (P x)) \ (c\<^bold>?x \ (Q x))" by (simp add: mono_Mprefix0 read_def) lemma mono_write_ref: "(P \ Q) \ (c\<^bold>!x \ P) \ (c\<^bold>!x \ Q)" by (simp add: mono_Mprefix0 write_def) lemma mono_write0_ref: "P \ Q \ (a \ P) \ (a \ Q)" by (simp add: mono_Mprefix0 write0_def) lemma write0_Ndet: "(a \ (P \ Q)) = ((a \ P) \ (a \ Q))" by (auto simp: Process_eq_spec write0_def D_Ndet F_Ndet F_Mprefix D_Mprefix Un_def) lemma write0_Det_Ndet: "((a \ P) \ (a \ Q)) = ((a \ P) \ (a \ Q))" by (auto simp: Process_eq_spec write0_def D_Ndet F_Ndet F_Det D_Det) (simp add: F_Mprefix)+ lemma Mprefix_Det: \(\e\A \ P e) \ (\e\A \ Q e) = \e\A \ (P e \ Q e)\ by (auto simp: Process_eq_spec F_Det D_Det) (auto simp: D_Ndet F_Ndet F_Mprefix D_Mprefix) section\ The Hiding Operator Laws \ subsection\ Preliminaries \ lemma elemDIselemHD: \t \ \ P \ trace_hide t (ev ` A) \ \ (P \ A)\ proof (cases "tickFree t") case True assume "t \ \ P" with True show ?thesis by (simp add:D_Hiding, rule_tac x=t in exI, rule_tac x="[]" in exI, simp) next case False assume a:"t \ \ P" with False obtain t' where "t = t'@[tick]" using D_imp_front_tickFree nonTickFree_n_frontTickFree by blast with a show ?thesis apply (auto simp add:D_Hiding, rule_tac x=t' in exI, rule_tac x="[tick]" in exI) by (meson front_tickFree_implies_tickFree front_tickFree_single is_processT) qed lemma length_strict_mono: "strict_mono (f::nat \ 'a list) \ length (f i) \ i + length (f 0)" apply(induct i, simp) by (metis dual_order.trans lessI less_length_mono not_less not_less_eq_eq plus_nat.simps(2) strict_mono_def) lemma mono_trace_hide: "a \ b \ trace_hide a (ev ` A) \ trace_hide b (ev ` A)" by (metis filter_append le_list_def) lemma mono_constant: assumes "mono (f::nat \ 'a event list)" and "\i. f i \ a" shows "\i. \j\i. f j = f i" proof - from assms(2) have "\i. length (f i) \ length a" by (simp add: le_length_mono) hence aa:"finite {length (f i)|i. True}" using finite_nat_set_iff_bounded_le by auto define lm where l2:"lm = Max {length (f i)|i. True}" with aa obtain im where "length (f im) = lm" using Max_in by fastforce with l2 assms(1) show ?thesis apply(rule_tac x=im in exI, intro impI allI) by (metis (mono_tags, lifting) Max_ge aa antisym le_length_mono le_neq_trans less_length_mono mem_Collect_eq mono_def order_less_irrefl) qed lemma elemTIselemHT: \t \ \ P \ trace_hide t (ev ` A) \ \ (P \ A)\ proof (cases "tickFree t") case True assume a:"t \ \ P" with True show ?thesis proof (cases "(\ta. trace_hide t (ev ` A) = trace_hide ta (ev ` A) \ (ta, ev ` A) \ \ P)") case True thus ?thesis by (simp add:T_Hiding) next case False with a False inf_hidden[of A t P] obtain f where "isInfHiddenRun f P A \ t \ range f" by auto with True show ?thesis by (simp add:T_Hiding, rule_tac disjI2, rule_tac x=t in exI, rule_tac x="[]" in exI, auto) qed next case False assume a:"t \ \ P" with False obtain t' where tt:"t = t'@[tick]" using T_nonTickFree_imp_decomp by auto with a show ?thesis proof (cases "(\ta. trace_hide t (ev ` A) = trace_hide ta (ev ` A) \ (ta, ev ` A) \ \ P)") case True thus ?thesis by (simp add:T_Hiding) next case False assume "t \ \ P" with False inf_hidden[of A t P] obtain f where "isInfHiddenRun f P A \ t \ range f" by auto then show ?thesis apply (simp add:T_Hiding, rule_tac disjI2, rule_tac x=t' in exI, rule_tac x="[tick]" in exI, auto) apply (metis append_T_imp_tickFree list.distinct(1) tt) using tt apply force by (metis False append_T_imp_tickFree is_processT5_S7 non_tickFree_tick not_Cons_self2 tickFree_append tt) qed qed lemma Hiding_Un_D1: \\ (P \ (A \ B)) \ \ ((P \ A) \ B)\ proof (simp add:conj_commute D_Hiding, intro conjI subset_antisym subsetI, simp_all, goal_cases) case (1 x) - then obtain t u where B1:"x = trace_hide t (ev ` (A \ B)) @ u" and B2:"tickFree t" and B3:"front_tickFree u" and - B4:"(t \ \ P \ (\(f:: nat \ 'a event list). isInfHiddenRun f P (A \ B) \ t \ range f))" by auto + then obtain t u where B1:"x = trace_hide t (ev ` (A \ B)) @ u" + and B2:"tickFree t" and B3:"front_tickFree u" + and B4:"(t \ \ P \ (\(f:: nat \ 'a event list). isInfHiddenRun f P (A \ B) + \ t \ range f))" by auto thus ?case apply(erule_tac disjE) apply(rule_tac x="trace_hide t (ev ` A)" in exI, rule_tac x=u in exI) apply(simp add:Hiding_tickFree tickFree_def) apply(rule disjI1, rule_tac x=t in exI, rule_tac x="[]" in exI, simp) proof(goal_cases) case 1 then obtain f n where fc:"isInfHiddenRun f P (A \ B) \ t = f n" by auto define ff where "ff = (\i. take (i + length (f 0)) (f i))" with fc have ffc:"isInfHiddenRun ff P (A \ B) \ t \ range ff" proof (auto, goal_cases) case 1 { fix x from length_strict_mono[of f "Suc x ", OF 1(2)] have a:"take (x + length (f 0)) (f (Suc x)) < take ((Suc x) + length (f 0)) (f (Suc x))" by (simp add: take_Suc_conv_app_nth) from 1(2)[THEN strict_monoD, of x "Suc x", simplified] obtain t where "f (Suc x) = (f x @ t)" by (metis le_list_def less_imp_le) with length_strict_mono[of f x, OF 1(2)] have "take (x + length (f 0)) (f x) = take (x + length (f 0)) (f (Suc x))" by simp with a have "take (x + length (f 0)) (f x) < take (Suc x + length (f 0)) (f (Suc x))" by simp } thus ?case by (meson lift_Suc_mono_less strict_mono_def) next case (2 i) have "take (i + length (f 0)) (f i) \ f i" using append_take_drop_id le_list_def by blast also have "\x y. x \ y \ y \ \ P \ x \ \ P" using is_processT3_ST_pref by blast ultimately show ?case using "2"(3) by blast next case (3 i) hence "(f i) \ (f 0)" using strict_mono_less_eq by blast hence "take (i + length (f 0)) (f i) \ (f 0)" by (metis add.commute append_eq_conv_conj le_list_def take_add) - hence a:"[x\take (i + length (f 0)) (f i) . x \ ev ` A \ x \ ev ` B] \ [x\f 0 . x \ ev ` A \ x \ ev ` B]" + hence a:" [x\take (i + length (f 0)) (f i) . x \ ev ` A \ x \ ev ` B] + \ [x\f 0 . x \ ev ` A \ x \ ev ` B]" by (metis (no_types, lifting) filter_append le_list_def) have "take (i + length (f 0)) (f i) \ f i" using append_take_drop_id le_list_def by blast - hence "[x\take (i + length (f 0)) (f i) . x \ ev ` A \ x \ ev ` B] \ [x\f i . x \ ev ` A \ x \ ev ` B]" + hence " [x\take (i + length (f 0)) (f i) . x \ ev ` A \ x \ ev ` B] + \ [x\f i . x \ ev ` A \ x \ ev ` B]" by (metis (no_types, lifting) filter_append le_list_def) with a 3(4) show ?case by (metis (no_types, lifting) dual_order.antisym) next case 4 have "f (length (f n) - length (f 0)) \ f n" by (simp add: "4"(2) add_le_imp_le_diff length_strict_mono strict_mono_less_eq) hence "f n = (\i. take (i + length (f 0)) (f i)) (length (f n) - length (f 0))" by (metis "4"(2) add.commute append_eq_conv_conj diff_is_0_eq' le_add_diff_inverse le_list_def le_zero_eq nat_le_linear strict_mono_less_eq) then show ?case by blast qed thus ?case proof(cases "\m. (\i>m. last (ff i) \ (ev ` A))") case True then obtain m where mc:"\i>m. last (ff i) \ (ev ` A)" by blast hence mc2:"tickFree (ff m)" by (metis (no_types, lifting) B2 event.distinct(1) ffc image_iff mem_Collect_eq set_filter tickFree_def) with mc mc2 1 ffc show ?thesis apply(rule_tac x="trace_hide (ff m) (ev ` A)" in exI, rule_tac x=u in exI, simp, intro conjI) apply (metis (no_types, lifting) mem_Collect_eq set_filter tickFree_def) apply (metis (no_types, lifting) rangeE) apply(rule disjI1, rule_tac x="ff m" in exI, rule_tac x="[]" in exI, intro conjI, simp_all) apply(rule disjI2, rule_tac x="\i. ff(i + m)" in exI, intro conjI) apply (metis (no_types, lifting) add.commute add.right_neutral rangeI) apply (simp add: strict_mono_def) apply blast proof(rule allI, goal_cases) case (1 i) from ffc ff_def True have "\t. (ff (i + m)) = (ff m) @ t \ set t \ (ev ` A)" proof(induct i) case 0 then show ?case by fastforce next case (Suc i) then obtain tt where tc:"(ff (i + m)) = (ff m) @ tt \ set tt \ (ev ` A)" by blast - from ffc ff_def length_strict_mono[of ff] have lc:"length (ff (Suc i + m)) = Suc (length (ff (i + m)))" + from ffc ff_def length_strict_mono[of ff] have lc:"length (ff (Suc i + m)) + = Suc (length (ff (i + m)))" by (metis (no_types, lifting) add_Suc fc length_strict_mono length_take min.absorb2) from True obtain l where lc2:"l = last (ff (Suc i + m)) \ l \ (ev ` A)" by (meson less_add_same_cancel2 mc zero_less_Suc) from ffc obtain r where rc:"ff (Suc i + m) = ff (i + m)@r" by (metis add.commute add_Suc_right le_list_def lessI less_imp_le strict_mono_less_eq) with lc have "length r = 1" by (metis Suc_eq_plus1 add_left_cancel length_append) with rc lc2 have "r = [l]" by (metis (no_types, lifting) Nil_is_append_conv Suc_eq_plus1 append_butlast_last_id append_eq_append_conv append_eq_append_conv2 length_Cons length_append) with Suc lc2 tc rc show ?case by (rule_tac x="tt@[l]" in exI, auto) qed then show ?case using filter_empty_conv by fastforce qed next case False { fix i assume as:"(i::nat) > 0" with ffc obtain tt where ttc:"ff i = ff 0 @ tt \ set tt \ ev ` (A \ B)" unfolding isInfHiddenRun_1 by blast with ffc as have "tt \ []" using strict_mono_eq by fastforce with ttc have "last (ff i) \ ev ` (A \ B)" by auto } hence as2:"\i. \j>i. last(ff j) \ ((ev ` B) - (ev ` A))" by (metis DiffI False UnE gr_implies_not_zero gr_zeroI image_Un) define ffb where "ffb = rec_nat t (\i t. (let j = SOME j. ff j > t \ last(ff j) \ ((ev ` B) - (ev ` A)) in ff j))" with as2 have ffbff:"\n. ffb n \ range ff" by (metis (no_types, lifting) ffc old.nat.exhaust old.nat.simps(6) old.nat.simps(7) rangeI) from 1 ffb_def show ?thesis apply(rule_tac x="trace_hide t (ev ` A)" in exI, rule_tac x=u in exI, simp, intro conjI) apply (meson filter_is_subset set_rev_mp tickFree_def) proof(rule disjI2, rule_tac x="\i. trace_hide (ffb i) (ev ` A)" in exI, intro conjI, goal_cases) case 1 then show ?case by (metis (no_types, lifting) old.nat.simps(6) rangeI) next case 2 { fix n have a0:"(ffb (Suc n)) = ff (SOME j. ff j > ffb n \ last(ff j) \ ((ev ` B) - (ev ` A)))" by (simp add: ffb_def) from ffbff obtain i where a1:"ffb n = ff i" by blast with as2 have "\j. ff j > ffb n \ last(ff j) \ ((ev ` B) - (ev ` A))" by (metis ffc strict_mono_def) with a0 a1 have a:"(ffb (Suc n)) > (ffb n) \ last (ffb (Suc n)) \ (ev ` A)" by (metis (no_types, lifting) Diff_iff tfl_some) then obtain r where "ffb (Suc n) = (ffb n)@r \ last r \ (ev ` A)" by (metis append_self_conv last_append le_list_def less_list_def) hence "trace_hide (ffb (Suc n)) (ev ` A) > trace_hide (ffb n) (ev ` A)" by (metis (no_types, lifting) a append_self_conv filter_append filter_empty_conv last_in_set le_list_def less_list_def) } then show ?case by (metis (mono_tags, lifting) lift_Suc_mono_less_iff strict_monoI) next case 3 then show ?case by (metis (mono_tags) elemTIselemHT ffbff ffc rangeE) next case 4 from ffbff ffc show ?case by (metis rangeE trace_hide_union) qed qed qed qed lemma Hiding_Un_D2: \finite A \ \ ((P \ A) \ B) \ \ (P \ (A \ B))\ proof (simp add:conj_commute D_Hiding, intro conjI subset_antisym subsetI, simp_all, goal_cases) case (1 x) then obtain t u where B1:"x = trace_hide t (ev ` B) @ u" and B2:"tickFree t" and B3:"front_tickFree u" - and B4:"(t \ \ (P \\ A) \ - (\(f:: nat \ 'a event list). isInfHiddenRun f (P \\ A) B \ t \ range f))" + and B4:\(t \ \ (P \ A) \ + (\(f:: nat \ 'a event list). isInfHiddenRun f (P \ A) B \ t \ range f))\ by (simp add:D_Hiding) blast thus ?case proof(erule_tac disjE, auto simp add:D_Hiding, goal_cases) case (1 ta ua) then show ?case by (rule_tac x=ta in exI, rule_tac x = "trace_hide ua (ev ` B) @ u" in exI, auto simp add: front_tickFree_append tickFree_def) next case (2 ua f xa) then show ?case apply(rule_tac x="f xa" in exI, rule_tac x="trace_hide ua (ev ` B) @ u" in exI, intro conjI disjI2) apply(auto simp add: front_tickFree_append tickFree_def) by (rule_tac x="f" in exI) (metis (no_types) filter_filter rangeI) next case (3 f xx) note "3a" = 3 then show ?case - proof(cases "\i. f i \ \ (P \\ A)") + proof(cases \\i. f i \ \ (P \ A)\) case True with 3 show ?thesis proof(auto simp add:D_Hiding, goal_cases) case (1 i ta ua) then show ?case apply (rule_tac x=ta in exI, rule_tac x="trace_hide ua (ev ` B) @ u" in exI, intro conjI) apply (metis (full_types) front_tickFree_append Hiding_tickFree tickFree_append) apply(subgoal_tac "trace_hide (f xx) (ev ` B) = trace_hide (f i) (ev ` B)", auto) apply (metis (full_types) filter_append filter_filter) by (metis (full_types) filter_append filter_filter) next case (2 i ua fa xa) hence "trace_hide (f xx) (ev ` B) = trace_hide (f i) (ev ` B)" by metis with 2 show ?case apply (rule_tac x="fa xa" in exI, rule_tac x="trace_hide ua (ev ` B) @ u" in exI, intro conjI) apply (metis (full_types) front_tickFree_append Hiding_tickFree tickFree_append) apply (simp_all) apply(rule_tac disjI2, rule_tac x=fa in exI, auto) by (metis (no_types) filter_filter) qed next case False - with 3 have Falsebis:"\i. (f i \ \ (P \\ A) \ f i \ \ (P \\ A))" by blast + with 3 have Falsebis:\\i. (f i \ \ (P \ A) \ f i \ \ (P \ A))\ by blast with T_Hiding[of P A] D_Hiding[of P A] have "\i. (f i \ {trace_hide t (ev ` A) |t. (t, ev ` A) \ \ P})" by (metis (no_types, lifting) UnE) hence ff0:"\i. (\t. f i = trace_hide t (ev ` A) \ t \ \ P)" using F_T by fastforce define ff where ff1:"ff = (\i. SOME t. f i = trace_hide t (ev ` A) \ t \ \ P)" hence "inj ff" unfolding inj_def by (metis (mono_tags, lifting) "3"(4) ff0 strict_mono_eq tfl_some) hence ff2:"infinite (range ff)" using finite_imageD by blast { fix tt i assume "tt \ range ff" then obtain k where "ff k = tt" using finite_nat_set_iff_bounded_le by blast hence kk0:"f k = trace_hide tt (ev ` A) \ tt \ \ P" using ff1 by (metis (mono_tags, lifting) ff0 someI_ex) hence "set (take i tt) \ set (f i) \ (ev ` A)" proof(cases "k \ i") case True hence "set (f k) \ set (f i)" by (metis "3"(4) le_list_def set_append strict_mono_less_eq sup.cobounded1) moreover from kk0 have "set (take i tt) \ set (f k) \ (ev ` A)" using in_set_takeD by fastforce ultimately show ?thesis by blast next case False have a:"length (f i) \ i" by (meson "3"(4) dual_order.trans le_add1 length_strict_mono) have b:"f i \ f k" using "3"(4) False nat_le_linear strict_mono_less_eq by blast with a have c:"take i (f k) \ (f i)" by (metis append_eq_conv_conj le_add_diff_inverse le_list_def take_add) from kk0[THEN conjunct1] have c1:"f k = (trace_hide (take i tt) (ev ` A)) @ (trace_hide (drop i tt) (ev ` A))" by (metis append_take_drop_id filter_append) have "length (trace_hide (take i tt) (ev ` A)) \ i" by (metis length_filter_le length_take min.absorb2 nat_le_linear order.trans take_all) with c1 have "take i (f k) \ (trace_hide (take i tt) (ev ` A))" by (simp add: le_list_def) with c obtain t where d:"f i = (trace_hide (take i tt) (ev ` A))@t" by (metis append.assoc le_list_def) then show ?thesis using in_set_takeD by fastforce qed } note ee=this { fix i have "finite {(take i t)|t. t \ range ff}" proof(induct i, simp) case (Suc i) have ff:"{take (Suc i) t|t. t \ range ff} \ {(take i t)|t. t \ range ff} \ (\e\(set (f (Suc i)) \ (ev ` A)). {(take i t)@[e]|t. t \ range ff})" (is "?AA \ ?BB") proof fix t assume "t \ ?AA" then obtain t' where hh:"t' \ range ff \ t = take (Suc i) t'" using finite_nat_set_iff_bounded_le by blast with ee[of t'] show "t \ ?BB" proof(cases "length t' > i") case True hence ii:"take (Suc i) t' = take i t' @ [t'!i]" by (simp add: take_Suc_conv_app_nth) with ee[of t' "Suc i"] have "t'!i \ set (f (Suc i)) \ (ev ` A)" by (simp add: hh) with ii hh show ?thesis by blast next case False hence "take (Suc i) t' = take i t'" by fastforce with hh show ?thesis by auto qed qed { fix e have "{x @ [e] |x. \t. x = take i t \ t \ range ff} = {take i t' @ [e] |t'. t' \ range ff}" by auto hence "finite({(take i t') @ [e] |t'. t'\range ff})" using finite_image_set[of _ "\t. t@[e]", OF Suc] by auto } note gg=this have "finite(set (f (Suc i)) \ (ev ` A))" using 1(1) by simp with ff gg Suc show ?case by (metis (no_types, lifting) finite_UN finite_Un finite_subset) qed } note ff=this hence "\i. {take i t |t. t \ range ff} = {t. \t'. t = take i (ff t')}" by auto with KoenigLemma[of "range ff"] ff ff2 obtain f' where gg:"strict_mono (f':: nat \ 'a event list) \ range f' \ {t. \t'\range ff. t \ t'}" by auto { fix i from gg obtain n where aa:"f' i \ ff n" by blast have "\t. f n = f 0 @ t" by (metis "3a"(4) le0 le_list_def strict_mono_less_eq) with mono_trace_hide[OF aa, of A] ff0 ff1 have "\t. trace_hide (f' i) (ev ` A) \ f 0 @ t" by (metis (mono_tags, lifting) someI_ex) } note zz=this { define ff' where "ff' = (\i. trace_hide (f' i) (ev ` A))" with gg have "mono ff'" by (rule_tac monoI, simp add: mono_trace_hide strict_mono_less_eq) assume aa:"\i. trace_hide (f' i) (ev ` A) \ f 0" with aa mono_constant have "\i. \j\i. ff' j = ff' i" using \mono ff'\ ff'_def by blast then obtain m where bb:"\j\m. ff' j = ff' m" by blast - have "ff' m \ \ (P \\ A)" - proof(simp add:D_Hiding, rule_tac x="f' m" in exI, rule_tac x="[]" in exI, intro conjI, simp, goal_cases) + have \ff' m \ \ (P \ A)\ + proof(simp add:D_Hiding, rule_tac x="f' m" in exI, rule_tac x="[]" in exI, + intro conjI, simp, goal_cases) case 1 from gg have "f' m < f' (Suc m)" by (meson lessI strict_monoD) moreover from gg obtain k where "f' (Suc m) \ ff k" by blast moreover have "ff k \ \ P" by (metis (mono_tags, lifting) ff0 ff1 someI_ex) ultimately show ?case by (metis NF_NT append_Nil2 front_tickFree_mono is_processT le_list_def less_list_def) next case 2 then show ?case unfolding ff'_def by simp next case 3 then show ?case proof(rule disjI2, rule_tac x="\i. f' (m + i)" in exI, simp_all, intro conjI allI, goal_cases) case 1 show ?case using gg[THEN conjunct1] by (rule_tac strict_monoI, simp add: strict_monoD) next case (2 i) from gg obtain k where "f' (m + i) \ ff k" by blast moreover from ff0 ff1 have "ff k \ \ P" by (metis (mono_tags, lifting) someI_ex) ultimately have "f' (m + i) \ \ P" using is_processT3_ST_pref by blast then show ?case by (simp add: add.commute) next case (3 i) then show ?case using bb[THEN spec, of "m+i", simplified] unfolding ff'_def by assumption next case 4 then show ?case by (metis Nat.add_0_right rangeI) qed qed with gg False have "False" by (metis (no_types, lifting) Falsebis aa append_Nil2 ff'_def front_tickFree_mono is_processT is_processT2_TR le_list_def) } with zz obtain m where hh:"trace_hide (f' m) (ev ` A) \ f 0" unfolding le_list_def by (metis append_eq_append_conv2) from ff0 ff1 gg show ?thesis proof(auto simp add:T_Hiding, rule_tac x="f' m" in exI, rule_tac x=u in exI, intro conjI, simp_all add:3(3), goal_cases) case 1 hence "f' m < f' (Suc m)" by (meson lessI strict_monoD) moreover from gg obtain k where "f' (Suc m) \ ff k" by blast moreover have "ff k \ \ P" by (metis (mono_tags, lifting) ff0 ff1 someI_ex) ultimately show ?case by (metis NF_NT append_Nil2 front_tickFree_mono is_processT le_list_def less_list_def) next case 2 from gg obtain k where "f' m \ ff k" by blast with ff0 ff1 mono_trace_hide[of "f' m"] have "trace_hide (f' m) (ev ` A) \ f k" by (metis (mono_tags, lifting) someI_ex) with mono_trace_hide[OF this, of B] mono_trace_hide[OF hh, of B] 3(6)[THEN spec, of k] 3(6) show ?case by (metis (full_types) dual_order.antisym filter_filter) next case 3 show ?case proof(rule disjI2, rule_tac x="\i. f' (m + i)" in exI, simp_all, intro conjI allI, goal_cases) case 1 then show ?case by (metis Nat.add_0_right rangeI) next case 2 with 3(4) show ?case by (rule_tac strict_monoI, simp add: strict_monoD) next case (3 i) from gg obtain k where "f' (m + i) \ ff k" by blast moreover from ff0 ff1 have "ff k \ \ P" by (metis (mono_tags, lifting) someI_ex) ultimately have "f' (m + i) \ \ P" using is_processT3_ST_pref by blast then show ?case by (simp add: add.commute) next case (4 i) from gg obtain k where "f' (m + i) \ ff k" by blast with ff0 ff1 mono_trace_hide[of "f' (m + i)"] have ll:"trace_hide (f' (m + i)) (ev ` A) \ f k" by (metis (mono_tags, lifting) someI_ex) { fix a b c assume "(a::'a event list) \ b" and "b \ c" and "c \ a" hence "b = c" by auto} note jj=this from jj[OF mono_trace_hide[OF hh, of B], OF mono_trace_hide[THEN mono_trace_hide, of "f' m" "f' (m + i)" B A, OF gg[THEN conjunct1, THEN strict_mono_mono, THEN monoD, of "m" "m+i", simplified]]] mono_trace_hide[OF ll, of B] show ?case unfolding "3a"(6) [THEN spec, of k] by simp qed qed qed qed qed lemma Hiding_Un_D: \finite A \ \ ((P \ A) \ B) = \ (P \ (A \ B))\ using Hiding_Un_D1 Hiding_Un_D2 by blast subsection\ Laws \ lemma mono_Hiding_FD[simp]: \P \ Q \ P \ A \ Q \ A\ apply (auto simp: le_ref_def F_Hiding D_Hiding) using F_subset_imp_T_subset by blast+ lemmas mono_Hiding_ref = mono_Hiding lemma Hiding_Un: \finite A \ P \ (A \ B) = (P \ A) \ B\ proof (simp add:Process_eq_spec, intro conjI, unfold F_Hiding, goal_cases) case 1 thus ?case (is "{(s, X). ?A s X} \ {(s, X). ?B s} = {(s, X). (\t. (?C1 s t \ (t, X \ ev ` B) \ ?C2 \ ?C3))} \ {(s, X). ?D s}") proof(unfold F_Hiding set_eq_subset Un_subset_iff, intro conjI, goal_cases) case 1 then show ?case by (auto, metis (no_types) filter_filter image_Un sup_commute sup_left_commute) next case 2 then show ?case by (rule_tac Un_mono[of "{}", simplified], insert Hiding_Un_D[of A P B], simp add: D_Hiding) next case 3 have "{(s, X). (\t. (?C1 s t \ (t, X \ ev ` B) \ ?C2))} \ {(s, X). ?A s X}" by (auto, metis (no_types) filter_filter image_Un sup_commute sup_left_commute) moreover have "{(s, X). (\t. (?C1 s t \ (t, X \ ev ` B) \ ?C3))} \ {(s, X). ?B s}" proof(auto,goal_cases) case (1 ta u) then show ?case using Hiding_fronttickFree by blast next case (2 u f x) then show ?case apply(rule_tac x="f x" in exI, rule_tac x="trace_hide u (ev ` B)" in exI, auto) using Hiding_fronttickFree apply blast apply(erule_tac x=f in allE) by (metis (no_types) filter_filter rangeI) qed moreover have "{(s, X). (\t. (?C1 s t \ (t, X \ ev ` B) \ ?C2 \ ?C3))} = {(s, X). (\t. (?C1 s t \ (t, X \ ev ` B) \ ?C2 ))} \ {(s, X). (\t. (?C1 s t \ (t, X \ ev ` B) \ ?C3))}" by blast ultimately show ?case by (metis (no_types, lifting) Un_mono) next case 4 then show ?case by (rule_tac Un_mono[of "{}", simplified], insert Hiding_Un_D[of A P B], simp add:D_Hiding) qed next case 2 then show ?case by (simp add: Hiding_Un_D) qed -lemma Hiding_set_BOT: "(\ \\ A) = \" +lemma Hiding_set_BOT: \(\ \ A) = \\ apply(auto simp add:Process_eq_spec D_Hiding F_Hiding F_UU D_UU) using Hiding_fronttickFree apply blast using front_tickFree_append Hiding_tickFree apply blast using front_tickFree_append Hiding_tickFree apply blast apply (metis (full_types) append_Nil filter.simps(1) tickFree_Nil tickFree_implies_front_tickFree) using front_tickFree_append Hiding_tickFree apply blast using front_tickFree_append Hiding_tickFree apply blast using tickFree_Nil by fastforce -lemma Hiding_set_STOP: "(STOP \\ A) = STOP" +lemma Hiding_set_STOP: \(STOP \ A) = STOP\ apply(auto simp add:Process_eq_spec D_Hiding F_Hiding F_STOP D_STOP T_STOP) by (metis (full_types) lessI less_irrefl strict_mono_eq) + -lemma Hiding_set_SKIP: "(SKIP \\ A) = SKIP" +lemma Hiding_set_SKIP: \(SKIP \ A) = SKIP\ apply(auto simp add:Process_eq_spec D_Hiding F_Hiding F_SKIP D_SKIP T_SKIP split:if_splits) apply (metis filter.simps(1) non_tickFree_tick) apply (metis (full_types) Hiding_tickFree n_not_Suc_n non_tickFree_tick strict_mono_eq) apply (metis (full_types) Hiding_tickFree n_not_Suc_n non_tickFree_tick strict_mono_eq) apply (metis event.distinct(1) filter.simps(1) imageE) apply (metis event.distinct(1) filter.simps(1) filter.simps(2) imageE) by (metis (full_types) Hiding_tickFree n_not_Suc_n non_tickFree_tick strict_mono_eq) -lemma Hiding_set_empty: "(P \\ {}) = P" +lemma Hiding_set_empty: \(P \ {}) = P\ apply(auto simp add:Process_eq_spec D_Hiding F_Hiding is_processT7 is_processT8_S strict_mono_eq) by (metis append_Nil2 front_tickFree_implies_tickFree front_tickFree_single is_processT nonTickFree_n_frontTickFree) subsection\ Multi-Operators laws \ -lemma Hiding_Ndet: "(P \ Q) \\ A = ((P \\ A) \ (Q \\ A))" +lemma Hiding_Ndet: \(P \ Q) \ A = ((P \ A) \ (Q \ A))\ proof(auto simp add:Process_eq_spec D_Hiding F_Hiding, simp_all add: F_Ndet T_Ndet D_Ndet D_Hiding F_Hiding, goal_cases) case (1 b t) then show ?case by blast next case (2 b t u) then show ?case by blast next case (3 b u f x) from 3(4) have A:"infinite ({i. f i \ \ P}) \ infinite ({i. f i \ \ Q})" using finite_nat_set_iff_bounded by auto hence "(\i. f i \ \ P) \ (\i. f i \ \ Q)" by (metis (no_types, lifting) 3(3) finite_nat_set_iff_bounded is_processT3_ST_pref mem_Collect_eq not_less strict_mono_less_eq) with A show ?case by (metis (no_types, lifting) 3(1,2,3,5) rangeI) next case (4 a b) then show ?case by blast next case (5 t u) then show ?case by blast next case (6 u f x) from 6(4) have A:"infinite ({i. f i \ \ P}) \ infinite ({i. f i \ \ Q})" using finite_nat_set_iff_bounded by auto hence "(\i. f i \ \ P) \ (\i. f i \ \ Q)" by (metis (no_types, lifting) 6(3) finite_nat_set_iff_bounded is_processT3_ST_pref mem_Collect_eq not_less strict_mono_less_eq) with A show ?case by (metis (no_types, lifting) 6(1,2,3,5) rangeI) next case (7 x) then show ?case by blast qed lemma Hiding_Mprefix_distr: - "(B \ A) = {} \ ((Mprefix A P) \\ B) = (Mprefix A (\x. ((P x) \\ B)))" + \(B \ A) = {} \ ((Mprefix A P) \ B) = (Mprefix A (\x. ((P x) \ B)))\ proof (auto simp add: Process_eq_spec, simp_all add: F_Mprefix T_Mprefix D_Mprefix D_Hiding F_Hiding, goal_cases) case (1 x b) then show ?case proof(elim exE disjE conjE, goal_cases) case (1 t) then show ?case by (simp add: inf_sup_distrib2) next case (2 t a) then show ?case by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) imageE list.sel(1) list.sel(3) list.collapse neq_Nil_conv) next case (3 t u a) hence "hd t \ ev ` B" by force with 3 have"x = hd t # trace_hide (tl t) (ev ` B) @ u" by (metis append_Cons filter.simps(2) list.exhaust_sel) with 3 show ?case by (metis list.distinct(1) list.sel(1) list.sel(3) tickFree_tl) next case (4 t u f) then obtain k where kk:"t = f k" by blast from 4 obtain a where "f 1 \ [] \ ev a = hd (f 1)" and aa:"a \ A" by (metis less_numeral_extra(1) nil_le not_less strict_mono_less_eq) with 4(1) 4(6)[THEN spec, of 0] 4(7)[THEN spec, of 1] have "f 0 \ [] \ hd (f 0) = ev a" apply auto apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.distinct(1)) apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.distinct(1)) by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.sel(1)) with 4(1, 7) aa have nf: "\i. f i \ [] \ hd (f i) = ev a" by (metis (mono_tags, opaque_lifting) "4"(5) append_Cons le_list_def le_simps(2) list.distinct(1) list.exhaust_sel list.sel(1) neq0_conv old.nat.distinct(1) strict_mono_less_eq) with 4(5) have sm:"strict_mono (tl \ f)" by (simp add: less_tail strict_mono_def) with 4 aa kk nf show ?case apply(rule_tac disjI2, intro conjI) apply (metis (no_types, lifting) Nil_is_append_conv disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.distinct(1)) apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl hd_append2 image_iff list.distinct(1) list.sel(1)) apply(rule_tac x=a in exI, intro conjI disjI2) apply (metis disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl hd_append2 image_iff list.distinct(1) list.sel(1)) apply(rule_tac x="tl t" in exI, rule_tac x="u" in exI, intro conjI, simp_all) apply (metis tickFree_tl) apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.distinct(1) list.sel(3) tl_append2) apply(subst disj_commute, rule_tac disjCI) apply(rule_tac x="tl \ f" in exI, intro conjI) apply auto apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl list.sel(3)) done qed next case (2 x b) then show ?case proof(elim exE disjE conjE, goal_cases) case 1 then show ?case apply(rule_tac disjI1, rule_tac x="[]" in exI) by (simp add: disjoint_iff_not_equal inf_sup_distrib2) next case (2 a t) then show ?case apply(rule_tac disjI1, rule_tac x="(ev a)#t" in exI) using list.collapse by fastforce next case (3 a t u) then show ?case apply(rule_tac disjI2, rule_tac x="(ev a)#t" in exI, rule_tac x=u in exI) using list.collapse by fastforce next case (4 a t u f) then show ?case apply(rule_tac disjI2) apply(rule_tac x="(ev a) # t" in exI, rule_tac x="u" in exI, intro conjI, simp) apply auto[1] using hd_Cons_tl apply fastforce apply(rule_tac disjI2, rule_tac x="\i. (ev a) # (f i)" in exI) by (auto simp add: less_cons strict_mono_def) qed next case (3 x) then show ?case proof(elim exE disjE conjE, goal_cases) case (1 t u a) hence aa: "hd (trace_hide t (ev ` B)) = ev a \ trace_hide t (ev ` B) \ []" by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.distinct(1) list.sel(1)) with 1 have "hd x = ev a \ x \ []" by simp with 1 show ?case apply(intro conjI, simp_all, rule_tac x=a in exI, simp) apply(rule_tac x="tl t" in exI, rule_tac x="u" in exI, intro conjI, simp_all) using tickFree_tl apply blast using aa by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.sel(3) tl_append2) next case (2 t u f) then obtain k where kk:"t = f k" by blast from 2 obtain a where "f 1 \ [] \ ev a = hd (f 1)" and aa:"a \ A" by (metis less_numeral_extra(1) nil_le not_less strict_mono_less_eq) with 2(1) 2(6)[THEN spec, of 0] 2(7)[THEN spec, of 1] have "f 0 \ [] \ hd (f 0) = ev a" apply auto apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.distinct(1)) apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.distinct(1)) by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.sel(1)) with 2(1, 7) aa have nf: "\i. f i \ [] \ hd (f i) = ev a" by (metis (mono_tags, opaque_lifting) 2(5) append_Cons le_list_def le_simps(2) list.distinct(1) list.exhaust_sel list.sel(1) neq0_conv old.nat.distinct(1) strict_mono_less_eq) with 2(5) have sm:"strict_mono (tl \ f)" by (simp add: less_tail strict_mono_def) from 2(1,4) nf aa kk have x1:"x \ []" by (metis Nil_is_append_conv disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.distinct(1)) with 2(1,4) nf aa kk have x2: "hd (trace_hide t (ev ` B)) = ev a \ trace_hide t (ev ` B) \ []" by (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.distinct(1) list.sel(1)) with 2(1,4) nf aa kk x1 have x3:"hd x = ev a" by simp with 2 aa kk nf sm x1 show ?case apply(intro conjI, simp_all) apply(rule_tac x="tl t" in exI, rule_tac x="u" in exI, intro conjI, simp_all) apply (metis tickFree_tl) apply (metis (no_types, lifting) disjoint_iff_not_equal event.inject filter.simps(2) hd_Cons_tl imageE list.distinct(1) list.sel(3) tl_append2) apply(subst disj_commute, rule_tac disjCI) apply(rule_tac x="tl \ f" in exI, intro conjI) apply auto apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl list.sel(3)) by (metis (no_types, lifting) filter.simps(2) hd_Cons_tl list.sel(3)) qed next case (4 x) then show ?case proof(elim exE disjE conjE, goal_cases) case (1 a t u) then show ?case apply(rule_tac x="(ev a)#t" in exI, rule_tac x=u in exI) using list.collapse by fastforce next case (2 a t u f) then show ?case apply(rule_tac x="(ev a) # t" in exI, rule_tac x="u" in exI, intro conjI, simp_all) apply auto[1] using hd_Cons_tl apply fastforce apply(rule_tac disjI2, rule_tac x="\i. (ev a) # (f i)" in exI) by (auto simp add: less_cons strict_mono_def) qed qed -lemma no_Hiding_read: "(\y. c y \ B) \ ((c\<^bold>?x \ (P x)) \\ B) = (c\<^bold>?x \ ((P x) \\ B))" +lemma no_Hiding_read: \(\y. c y \ B) \ ((c\<^bold>?x \ (P x)) \ B) = (c\<^bold>?x \ ((P x) \ B))\ by (simp add: read_def o_def, subst Hiding_Mprefix_distr, auto) -lemma no_Hiding_write0: "a \ B \ ((a \ P) \\ B) = (a \ (P \\ B))" +lemma no_Hiding_write0: \a \ B \ ((a \ P) \ B) = (a \ (P \ B))\ by (simp add: Hiding_Mprefix_distr write0_def) -lemma Hiding_write0: "a \ B \ ((a \ P) \\ B) = (P \\ B)" +lemma Hiding_write0: \a \ B \ ((a \ P) \ B) = (P \ B)\ proof (auto simp add: write0_def Process_eq_spec, simp_all add: F_Mprefix T_Mprefix D_Mprefix D_Hiding F_Hiding, goal_cases) case (1 x b) then show ?case apply(elim exE disjE conjE) apply (metis filter.simps(2) hd_Cons_tl image_eqI) apply (metis (no_types, lifting) filter.simps(2) image_eqI list.sel(1) list.sel(3) neq_Nil_conv tickFree_tl) proof(goal_cases) case (1 t u f) have fS: "strict_mono (f \ Suc)" by (metis "1"(5) Suc_mono comp_def strict_mono_def) from 1 have aa:"\i. f (Suc i) \ []" by (metis (full_types) less_Suc_eq_le less_irrefl nil_le strict_mono_less_eq) with fS have sm:"strict_mono (tl \ f \ Suc)" by (simp add: less_tail strict_mono_def) with 1 aa show ?case apply(subst disj_commute, rule_tac disjCI, simp) apply(rule_tac x="tl (f 1)" in exI, rule_tac x="u" in exI, intro conjI, simp_all) apply (metis Hiding_tickFree imageE tickFree_tl) apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl image_eqI rangeE) apply(subst disj_commute, rule_tac disjCI) apply(rule_tac x="tl \ f \ Suc" in exI, intro conjI) apply auto apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl list.sel(3)) done qed next case (2 aa b) then show ?case apply(elim exE disjE conjE) apply (metis (no_types, lifting) filter.simps(2) image_eqI list.distinct(1) list.sel(1) list.sel(3)) proof(goal_cases) case (1 t u) then show ?case by (rule_tac disjI2, rule_tac x="(ev a)#t" in exI, auto) next case (2 t u f) then show ?case apply(rule_tac disjI2) apply(rule_tac x="(ev a) # t" in exI, rule_tac x="u" in exI, intro conjI, simp_all) apply(rule_tac disjI2) apply(rule_tac x="\i. (ev a) # (f i)" in exI, intro conjI) by (auto simp add: less_cons strict_mono_def) qed next case (3 x) then show ?case apply(elim exE disjE conjE) apply(rule_tac x="tl t" in exI, rule_tac x="u" in exI, intro conjI, simp_all) using tickFree_tl apply blast apply (metis filter.simps(2) hd_Cons_tl image_eqI) proof(goal_cases) case (1 t u f) have fS: "strict_mono (f \ Suc)" by (metis "1"(5) Suc_mono comp_def strict_mono_def) from 1 have aa:"\i. f (Suc i) \ []" by (metis (full_types) less_Suc_eq_le less_irrefl nil_le strict_mono_less_eq) with fS have sm:"strict_mono (tl \ f \ Suc)" by (simp add: less_tail strict_mono_def) with 1 aa show ?case apply(rule_tac x="tl (f 1)" in exI, rule_tac x="u" in exI, intro conjI, simp_all) apply (metis Hiding_tickFree imageE tickFree_tl) apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl image_eqI rangeE) apply(subst disj_commute, rule_tac disjCI) apply(rule_tac x="tl \ f \ Suc" in exI, intro conjI) apply auto apply (metis (no_types, lifting) filter.simps(2) hd_Cons_tl list.sel(3)) done qed next case (4 x) then show ?case apply(elim exE disjE conjE) proof(goal_cases) case (1 t u) then show ?case by (rule_tac x="(ev a)#t" in exI, auto) next case (2 t u f) then show ?case apply(rule_tac x="(ev a) # t" in exI, rule_tac x="u" in exI, intro conjI, simp_all) apply(rule_tac disjI2) apply(rule_tac x="\i. (ev a) # (f i)" in exI, intro conjI) by (auto simp add: less_cons strict_mono_def) qed qed -lemma no_Hiding_write: "(\y. c y \ B) \ ((c\<^bold>!a \ P) \\ B) = (c\<^bold>!a \ (P \\ B))" +lemma no_Hiding_write: \(\y. c y \ B) \ ((c\<^bold>!a \ P) \ B) = (c\<^bold>!a \ (P \ B))\ by(simp add: write_def, subst Hiding_Mprefix_distr, auto) -lemma Hiding_write: "(c a) \ B ==> ((c\<^bold>!a \ P) \\ B) = (P \\ B)" +lemma Hiding_write: \(c a) \ B \ ((c\<^bold>!a \ P) \ B) = (P \ B)\ by (simp add: write_def Hiding_write0 Mprefix_singl) section\ The Sync Operator Laws \ subsection\ Preliminaries \ lemma SyncTlEmpty:"a setinterleaves (([], u), A) \ tl a setinterleaves (([], tl u), A)" by (case_tac u, simp, case_tac a, simp_all split:if_splits) lemma SyncHd_Tl: "a setinterleaves ((t, u), A) \ hd t \ A \ hd u \ A \ hd a = hd u \ tl a setinterleaves ((t, tl u), A)" by (case_tac u) (case_tac t, auto split:if_splits)+ lemma SyncHdAddEmpty: "(tl a) setinterleaves (([], u), A) \ hd a \ A \ a \ [] \ a setinterleaves (([], hd a # u), A)" using hd_Cons_tl by fastforce lemma SyncHdAdd: "(tl a) setinterleaves ((t, u), A) \ hd a \ A \ hd t \ A \ a \ [] \ a setinterleaves ((t, hd a # u), A)" by (case_tac a, simp, case_tac t, auto) lemmas SyncHdAdd1 = SyncHdAdd[of "a#r", simplified] for a r lemma SyncSameHdTl: "a setinterleaves ((t, u), A) \ hd t \ A \ hd u \ A \ hd t = hd u \ hd a = hd t \ (tl a) setinterleaves ((tl t, tl u), A)" by (case_tac u) (case_tac t, auto split:if_splits)+ lemma SyncSingleHeadAdd: "a setinterleaves ((t, u), A) \ h \ A \ (h#a) setinterleaves ((h#t, u), A)" by (case_tac u, auto split:if_splits) lemma TickLeftSync: "tick \ A \ front_tickFree t \ s setinterleaves (([tick], t), A) \ s = t \ last t = tick" proof(induct t arbitrary: s) case Nil then show ?case by (simp add: Nil.prems) next case (Cons a t) then show ?case apply (auto split:if_splits) using emptyLeftProperty apply blast apply (metis last_ConsR last_snoc nonTickFree_n_frontTickFree tickFree_Cons) by (metis append_Cons append_Nil front_tickFree_mono)+ qed lemma EmptyLeftSync:"s setinterleaves (([], t), A) \ s = t \ set t \ A = {}" by (meson Int_emptyI emptyLeftNonSync emptyLeftProperty) lemma tick_T_F:"t@[tick] \ \ P \ (t@[tick], X) \ \ P" using append_T_imp_tickFree is_processT5_S7 by force lemma event_set: "(e::'a event) \ insert tick (range ev)" by (metis event.exhaust insert_iff rangeI) lemma Mprefix_Sync_distr1_D: " A \ S \ B \ S \ \ (Mprefix A P \S\ Mprefix B Q) = \ (\ x \ A \ B \ (P x \S\ Q x))" apply(auto,simp_all add:D_Sync F_Sync F_Mprefix T_Mprefix D_Mprefix) apply(elim exE disjE conjE) apply (metis (no_types, lifting) Sync.sym empty_iff image_iff insertI2 list.exhaust_sel setinterleaving.simps(2) subsetCE) proof(goal_cases) case (1 x t u r v a aa) from 1(1,2,6,8,11,13) have aa1: "hd t\insert tick (ev ` S)\hd u\insert tick (ev ` S)" by blast from 1(5,6,7,8,11,13) aa1 have aa2: " hd x \ ev ` (A \ B)" - by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff SyncSameHdTl) + by (metis (no_types, lifting) IntI empty_setinterleaving event.inject + hd_append2 image_iff SyncSameHdTl) then show ?case using 1(3,4,5,6,7,9,10,13,14) by (metis (no_types, lifting) Nil_is_append_conv aa1 empty_setinterleaving event.inject hd_append2 SyncSameHdTl tickFree_tl tl_append2) next case (2 x t u r v a) then show ?case by (metis (no_types, lifting) Sync.si_empty3 equals0D imageE image_eqI insertI2 list.exhaust_sel subsetCE) next case (3 x t u r v a aa) from 3(1,2,6,8,11,13) have aa1: "hd t\insert tick (ev ` S)\hd u\insert tick (ev ` S)" by blast from 3(5,6,7,8,11,13) aa1 have aa2: " hd x \ ev ` (A \ B)" by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff SyncSameHdTl) then show ?case using 3(3,4,5,6,7,9,10,13,14) by (metis (no_types, lifting) Nil_is_append_conv aa1 empty_setinterleaving event.inject hd_append2 SyncSameHdTl tickFree_tl tl_append2) next case (4 x t u r v a) then show ?case by (metis (no_types, lifting) Sync.si_empty3 equals0D imageE image_eqI insertI2 list.exhaust_sel subsetCE) next case (5 x t u r v a aa) from 5(1,2,6,8,11,13) have aa1: "hd t\insert tick (ev ` S)\hd u\insert tick (ev ` S)" by blast from 5(5,6,7,8,11,13) aa1 have aa2: " hd x \ ev ` (A \ B)" by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff SyncSameHdTl) then show ?case using 5(3,4,5,6,7,9,10,13,14) by (metis aa1 append_Nil2 empty_setinterleaving event.inject SyncSameHdTl) next case (6 x t u r v a) then show ?case by (metis (no_types, lifting) Sync.si_empty3 equals0D imageE image_eqI insertI2 list.exhaust_sel subsetCE) next case (7 x t u r v a aa) from 7(1,2,6,8,11,13) have aa1: "hd t\insert tick (ev ` S)\hd u\insert tick (ev ` S)" by blast from 7(5,6,7,8,11,13) aa1 have aa2: " hd x \ ev ` (A \ B)" by (metis (no_types, lifting) IntI empty_setinterleaving event.inject hd_append2 image_iff SyncSameHdTl) then show ?case using 7(3,4,5,6,7,9,10,13,14) by (metis aa1 append_Nil2 empty_setinterleaving event.inject SyncSameHdTl) next case (8 x) then show ?case apply(elim exE disjE conjE) proof(goal_cases) case (1 a t u r v) obtain r1 t1 u1 where aa0: "r1=hd x#r\t1=hd x#t\u1=hd x#u" by auto from 1(3,4,5,7,8,9) have aa1: "tickFree r1\x = r1 @ v" by (metis Cons_eq_appendI aa0 event.distinct(1) list.exhaust_sel tickFree_Cons) from 1(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` S))\t1 \ []" using aa0 subsetCE by auto from 1(4,5,10,11) aa0 have aa3: "(tl t1) \ \ (P a) \ (tl u1) \ \ (Q a) \ ev a = hd t1 \ ev a = hd u1 \ hd t1 \ ev ` A\ hd u1 \ ev ` B" by auto then show ?case using "1"(6) aa1 aa2 by fastforce next case (2 a t u r v) obtain r1 t1 u1 where aa0: "r1=hd x#r\t1=hd x#t\u1=hd x#u" by auto from 2(3,4,5,7,8,9) have aa1: "tickFree r1\x = r1 @ v" by (metis Cons_eq_appendI aa0 event.distinct(1) list.exhaust_sel tickFree_Cons) from 2(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` S))\t1 \ []" using aa0 subsetCE by auto from 2(4,5,10,11) aa0 have aa3: "(tl t1) \ \ (Q a) \ (tl u1) \ \ (P a) \ ev a = hd t1 \ ev a = hd u1 \ hd t1 \ ev ` A\ hd u1 \ ev ` B" by auto then show ?case using "2"(6) aa1 aa2 by fastforce next case (3 a t u r v) obtain r1 t1 u1 where aa0: "r1=hd x#r\t1=hd x#t\u1=hd x#u" by auto from 3(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` S))\t1 \ []" using aa0 subsetCE by auto from 3(3,4,5,7,8,10,11) aa0 have aa3: "(tl t1) \ \(P a)\(tl u1) \ \ (Q a)\ ev a = hd t1\ev a = hd u1 \ hd t1 \ ev ` A\ hd u1 \ ev ` B\x = r1 @ v" by (metis (no_types, lifting) Int_lower1 Int_lower2 append_Nil2 imageE image_eqI list.exhaust_sel list.sel(1) list.sel(3) subsetCE) then show ?case by (metis (no_types, lifting) "3"(6) "3"(7) aa2 event.inject imageE) next case (4 a t u r v) obtain r1 t1 u1 where aa0: "r1=hd x#r\t1=hd x#t\u1=hd x#u" by auto from 4(2,4,9) have aa2: "r1 setinterleaves ((t1, u1), insert tick (ev ` S))\t1 \ []" using aa0 subsetCE by auto from 4(3,4,5,7,8,10,11) aa0 have aa3: "(tl t1) \ \ (Q a)\(tl u1) \ \ (P a) \ ev a = hd t1\ev a = hd u1\ hd t1 \ ev ` A\ hd u1 \ ev ` B\x = r1 @ v" by (metis (no_types, lifting) Int_lower1 Int_lower2 append_Nil2 imageE image_eqI list.exhaust_sel list.sel(1) list.sel(3) subsetCE) then show ?case by (metis (no_types, lifting) "4"(6) "4"(7) aa2 event.inject imageE) qed qed lemma Hiding_interleave: "A \ C = {} \ r setinterleaves ((t, u), C) \ (trace_hide r A) setinterleaves ((trace_hide t A, trace_hide u A), C)" proof(induct r arbitrary:t u) case Nil then show ?case using EmptyLeftSync empty_setinterleaving by fastforce next case (Cons a r) then show ?case apply(cases t) using EmptyLeftSync apply fastforce apply(cases u) apply fastforce proof(simp split:if_splits, goal_cases) case (1 a list lista) then show ?case by fastforce next case (2 a list aa lista) then show ?case apply(erule_tac disjE) using Cons(1)[of list u] apply (metis (no_types, lifting) filter.simps(2) SyncSingleHeadAdd) using Cons(1)[of t lista] by (metis (no_types, lifting) Sync.sym filter.simps(2) SyncSingleHeadAdd) next case (3 a list lista) then show ?case by fastforce qed qed lemma non_Sync_interleaving: "(set t \ set u) \ C = {} \ setinterleaving (t, C, u) \ {}" (* FINALLY NON-USED*) proof(induct t arbitrary:u) case Nil then show ?case by (metis Un_empty_left disjoint_iff_not_equal emptyLeftSelf empty_iff empty_set) next case (Cons a t) then show ?case proof(induct u) case Nil then show ?case using Sync.sym by auto next case (Cons a u) then show ?case by auto qed qed lemma interleave_Hiding: "A \ C = {} \ ra setinterleaves ((trace_hide t A, trace_hide u A), C) \ \r. ra = trace_hide r A \ r setinterleaves ((t, u), C)" proof(induct "length t + length u" arbitrary:ra t u rule:nat_less_induct) case Ind:1 then show ?case proof (cases t) case Nilt: Nil from Ind(2) have a:"set (trace_hide u A) \ C = {} \ set u \ C = {}" by auto hence b: "u setinterleaves ((t, u), C)" by (metis (no_types, lifting) Ind(3) EmptyLeftSync Nilt disjoint_iff_not_equal emptyLeftSelf filter.simps(1)) then show ?thesis apply(rule_tac x=u in exI) using EmptyLeftSync[of "ra" C "trace_hide u A"] a b Cons Nilt Ind(3) by auto next case Const: (Cons ta tlist) then show ?thesis proof(cases u) case Nilu: Nil from Ind(2) have a:"set (trace_hide t A) \ C = {} \ set t \ C = {}" by auto hence b: "t setinterleaves ((t, u), C)" by (metis Ind(3) Nilu EmptyLeftSync disjoint_iff_not_equal emptyLeftSelf filter.simps(1) Sync.sym) then show ?thesis apply(rule_tac x=t in exI) using EmptyLeftSync[of "ra" C "trace_hide t A"] a b Ind Nilu by (metis Sync.sym filter.simps(1)) next case Consu: (Cons ua ulist) with Const Ind(2,3) show ?thesis proof(auto split:if_splits simp del:setinterleaving.simps, goal_cases) case 1 then show ?case proof (auto, goal_cases) case (1 raa) with Ind(1)[THEN spec, of "length tlist + length u", simplified Const, simplified, THEN spec, THEN spec, of tlist u, simplified Ind, simplified, THEN spec, of raa] obtain r where "raa = trace_hide r A \ r setinterleaves ((tlist, u), C)" by auto then show ?case using "1"(4) Consu by force next case (2 raa) with Ind(1)[THEN spec, of "length t + length ulist", simplified Consu, simplified, THEN spec, THEN spec, of t ulist, simplified Ind, simplified, THEN spec, of raa] obtain r where "raa = trace_hide r A \ r setinterleaves ((t, ulist), C)" by auto then show ?case using "2"(2) "2"(4) by force next case (3 raa) with Ind(1)[THEN spec, of "length tlist + length ulist", simplified Const Consu, simplified, THEN spec, THEN spec, of tlist ulist, simplified Ind, simplified, THEN spec, of raa] obtain r where "raa = trace_hide r A \ r setinterleaves ((tlist, ulist), C)" by auto then show ?case using "3"(4) by force next case (4 raa) with Ind(1)[THEN spec, of "length t + length ulist", simplified Consu, simplified, THEN spec, THEN spec, of t ulist, simplified Ind, simplified, THEN spec, of raa] obtain r where "raa = trace_hide r A \ r setinterleaves ((t, ulist), C)" by auto then show ?case using "4"(5) Const by force next case (5 raa) with Ind(1)[THEN spec, of "length tlist + length u", simplified Const, simplified, THEN spec, THEN spec, of tlist u, simplified Ind, simplified, THEN spec, of raa] obtain r where "raa = trace_hide r A \ r setinterleaves ((tlist, u), C)" by auto then show ?case using "5"(4) Consu by force next case (6 raa) with Ind(1)[THEN spec, of "length t + length ulist", simplified Consu, simplified, THEN spec, THEN spec, of t ulist, simplified Ind, simplified, THEN spec, of raa] obtain r where "raa = trace_hide r A \ r setinterleaves ((t, ulist), C)" by auto then show ?case using "6"(5) Const by force next case (7 raa) with Ind(1)[THEN spec, of "length tlist + length u", simplified Const, simplified, THEN spec, THEN spec, of tlist u, simplified Ind, simplified, THEN spec, of raa] obtain r where "raa = trace_hide r A \ r setinterleaves ((tlist, u), C)" by auto then show ?case using "7"(4) Consu by force qed next case 2 with Ind(1)[THEN spec, of "length t + length ulist", simplified Consu, simplified, THEN spec, THEN spec, of t ulist, simplified Ind, simplified, THEN spec, of ra] obtain r where "ra = trace_hide r A \ r setinterleaves ((t, ulist), C)" by auto then show ?case apply(rule_tac x="ua#r" in exI) using "2"(5) Const Ind.prems(1) by auto next case 3 with Ind(1)[THEN spec, of "length tlist + length u", simplified Const, simplified, THEN spec, THEN spec, of tlist u, simplified Ind, simplified, THEN spec, of ra] obtain r where "ra = trace_hide r A \ r setinterleaves ((tlist, u), C)" by auto then show ?case apply(rule_tac x="ta#r" in exI) using "3"(4) Consu Ind.prems(1) by auto next case 4 with Ind(1)[THEN spec, of "length tlist + length ulist", simplified Const Consu, simplified, THEN spec, THEN spec, of tlist ulist, simplified Ind, simplified, THEN spec, of ra] obtain r where "ra = trace_hide r A \ r setinterleaves ((tlist, ulist), C)" by auto then show ?case apply(rule_tac x="ta#ua#r" in exI) using "4"(4,5) Consu Const Ind.prems(1) apply auto using SyncSingleHeadAdd apply blast by (metis Sync.sym SyncSingleHeadAdd) qed qed qed qed lemma interleave_size: "s setinterleaves ((t,u), C) \ length s = length t + length u - length(filter (\x. x\C) t)" proof(induct s arbitrary:t u) case Nil then show ?case using EmptyLeftSync empty_setinterleaving by fastforce next case (Cons a list) then show ?case apply(cases t) using emptyLeftProperty apply fastforce apply(cases u) apply (metis (no_types, lifting) Sync.sym add_diff_cancel_right' emptyLeftNonSync emptyLeftProperty filter_empty_conv) proof(auto split:if_splits, goal_cases 11 12 13 14 15) case (11 tlist ulist) then show ?case by (metis (no_types, lifting) Suc_diff_le le_add1 length_filter_le order_trans) next case (12 ta tlist ulist) with 12(3)[rule_format, of "ta#tlist" ulist] show ?case by simp (metis Suc_diff_le le_add1 length_filter_le order_trans) next case (13 tlist ua ulist) with 13(3)[rule_format, of tlist "ua#ulist"] show ?case by simp (metis Suc_diff_le le_less_trans length_filter_le less_SucI less_Suc_eq_le less_add_Suc1) next case (14 ta tlist ulist) with 14(3)[rule_format, of "ta#tlist" ulist] show ?case by simp (metis Suc_diff_le le_less_trans length_filter_le less_SucI less_Suc_eq_le less_add_Suc1) next case (15 tlist ua ulist) with 15(3)[rule_format, of tlist "ua#ulist"] show ?case by simp (metis Suc_diff_le le_less_trans length_filter_le less_SucI less_Suc_eq_le less_add_Suc1) qed qed lemma interleave_eq_size: "s setinterleaves ((t,u), C) \ s' setinterleaves ((t,u), C) \ length s = length s'" by (simp add: interleave_size) lemma interleave_set: "s setinterleaves ((t,u), C) \ set t \ set u \ set s" proof(induct s arbitrary: t u) case Nil then show ?case using EmptyLeftSync empty_setinterleaving by blast next case (Cons a s) then show ?case apply(cases t) using emptyLeftProperty apply fastforce apply(cases u) apply (metis Sync.sym Un_empty_right emptyLeftProperty empty_set eq_refl) by (auto simp add: subset_iff split:if_splits) qed lemma interleave_order: "s setinterleaves ((t1@t2,u), C) \ set(t2) \ set(drop (length t1) s)" proof(induct s arbitrary: t1 t2 u) case Nil then show ?case using empty_setinterleaving by auto next case (Cons a s) then show ?case apply(cases t1) using append_self_conv2 interleave_set apply fastforce apply(cases u) apply (metis EmptyLeftSync Sync.sym append_eq_conv_conj order_refl) proof (auto simp add: subset_iff split:if_splits, goal_cases) case (1 a list lista t) then show ?case using Cons(1)[of "a#list" "t2" "lista", simplified, OF 1(6)] by (meson Suc_n_not_le_n contra_subsetD nat_le_linear set_drop_subset_set_drop) next case (2 a list lista t) then show ?case using Cons(1)[of "a#list" "t2" "lista", simplified, OF 2(7)] by (meson Suc_n_not_le_n contra_subsetD nat_le_linear set_drop_subset_set_drop) qed qed lemma interleave_append: "s setinterleaves ((t1@t2,u), C) \ \ u1 u2 s1 s2. u = u1@u2 \ s = s1@s2 \ s1 setinterleaves ((t1,u1), C) \ s2 setinterleaves ((t2,u2), C)" proof(induct s arbitrary: t1 t2 u) case Nil then show ?case using empty_setinterleaving by fastforce next case (Cons a s) then show ?case apply(cases t1) using append_self_conv2 interleave_set apply fastforce apply(cases u) apply(exI "[]", exI "[]") apply auto[1] apply (metis (no_types, opaque_lifting) Nil_is_append_conv append_Cons) proof (auto simp add: subset_iff split:if_splits, goal_cases) case (1 list lista) with Cons(1)[of "list" "t2" "lista", simplified, OF 1(5)] show ?case proof(elim exE conjE, goal_cases) case (1 u1 u2 s1 s2) then show ?case by (exI "a#u1", exI "u2", simp) (metis append_Cons) qed next case (2 aa list lista) with Cons(1)[of "aa#list" "t2" "lista", simplified, OF 2(6)] show ?case proof(elim exE conjE, goal_cases) case (1 u1 u2 s1 s2) then show ?case by (exI "a#u1", exI "u2", simp) (metis append_Cons) qed next case (3 list aa lista) with Cons(1)[of "list" "t2" "aa#lista", simplified, OF 3(6)] show ?case proof(elim exE conjE, goal_cases) case (1 u1 u2 s1 s2) then show ?case apply (exI "u1", exI "u2", simp) by (metis append_Cons SyncSingleHeadAdd) qed next case (4 aa list lista) with Cons(1)[of "aa#list" "t2" "lista", simplified, OF 4(6)] show ?case proof(elim exE conjE, goal_cases) case (1 u1 u2 s1 s2) then show ?case by (exI "a#u1", exI "u2", simp) (metis append_Cons) qed next case (5 list aa lista) with Cons(1)[of "list" "t2" "aa#lista", simplified, OF 5(6)] show ?case proof(elim exE conjE, goal_cases) case (1 u1 u2 s1 s2) then show ?case apply (exI "u1", exI "u2", simp) by (metis append_Cons SyncSingleHeadAdd) qed qed qed lemma interleave_append_sym: "s setinterleaves ((t,u1@u2), C) \ \ t1 t2 s1 s2. t = t1@t2 \ s = s1@s2 \ s1 setinterleaves ((t1,u1), C) \ s2 setinterleaves ((t2,u2), C)" by (metis (no_types) Sync.sym interleave_append) lemma interleave_append_tail: "s setinterleaves ((t1,u), C) \ (set t2) \ C = {} \ (s@t2) setinterleaves ((t1@t2,u), C)" proof(induct s arbitrary: t1 t2 u) case Nil then show ?case by (metis Set.set_insert Sync.sym append_Nil disjoint_insert(2) emptyLeftSelf empty_setinterleaving) next case (Cons a s) then show ?case apply(cases u) using EmptyLeftSync Sync.sym apply fastforce apply(cases t1, cases t2) apply simp apply fastforce proof(auto, goal_cases) case (1 list lista) with 1(1)[OF 1(7) 1(2)] show ?case by simp next case (2 list aa lista) with 2(1)[OF 2(2)] show ?case by simp next case (3 list aa lista) with 3(1)[OF 3(9) 3(2)] show ?case by simp next case (4 list aa lista) with 4(1)[OF 4(9) 4(2)] show ?case by simp qed qed lemma interleave_append_tail_sym: "s setinterleaves ((t,u1), C) \ (set u2) \ C = {} \ (s@u2) setinterleaves ((t,u1@u2), C)" by (metis (no_types) Sync.sym interleave_append_tail) lemma interleave_assoc_1: "tu setinterleaves ((t, u), A) \ tuv setinterleaves ((tu, v), A) \ \uv. uv setinterleaves ((u, v), A) \ tuv setinterleaves ((t, uv), A)" proof(induct tuv arbitrary: t tu u v) case Nil then show ?case using EmptyLeftSync empty_setinterleaving by blast next case Cons_tuv:(Cons a tuv) then show ?case proof(cases t) case Nil_t:Nil with Cons_tuv(2) have *:"tu = u" using EmptyLeftSync by blast show ?thesis proof(cases u) case Nil_u:Nil with Nil_t Cons_tuv show ?thesis using EmptyLeftSync by fastforce next case Cons_u:(Cons ua ulist) with Nil_t Cons_tuv(2) have "ua \ A" by force show ?thesis proof(cases v) case Nil_v:Nil with * Nil_t Cons_tuv(2,3) Cons_u show ?thesis using Sync.sym by blast next case Cons_v:(Cons va vlist) with * Nil_t Cons_tuv(2,3) Cons_u show ?thesis apply(exI "a#tuv", auto split:if_splits) using Cons_tuv.hyps Cons_tuv.prems(1) emptyLeftProperty by blast+ qed qed next case Cons_t:(Cons ta tlist) show ?thesis proof(cases u) case Nil_u:Nil with Cons_t Cons_tuv show ?thesis by (metis Sync.sym emptyLeftProperty emptyLeftSelf empty_set equals0D ftf_Sync21) next case Cons_u:(Cons ua ulist) show ?thesis proof(cases v) case Nil_v:Nil with Cons_tuv(3) have "a # tuv = tu" by (simp add: Nil_v EmptyLeftSync Sync.sym) with Nil_v Cons_u Cons_t Cons_tuv show ?thesis apply(exI u, auto split:if_splits) apply (metis Cons_t Nil_v Sync.sym emptyLeftNonSync list.set_intros(1)) using Cons_tuv(1)[of tuv tlist u] apply(metis (no_types, lifting) Sync.sym emptyLeftNonSync emptyLeftSelf list.sel(3) SyncTlEmpty) by (metis Cons_t Sync.sym emptyLeftProperty) fastforce+ next case Cons_v:(Cons va vlist) with Cons_t Cons_u Cons_tuv(2,3) show ?thesis proof(auto split:if_splits, goal_cases) case (1 tulist) from Cons_tuv(1)[OF 1(5) 1(10)] obtain uvlist - where "uvlist setinterleaves ((ulist, vlist), A) \ tuv setinterleaves ((tlist, uvlist), A)" by blast + where " uvlist setinterleaves ((ulist, vlist), A) + \ tuv setinterleaves ((tlist, uvlist), A)" by blast with 1 show ?case by(exI "va#uvlist", simp) next case (2 u) then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym SyncSingleHeadAdd) next case (3 u) then show ?case by (metis Cons_tuv.hyps Sync.sym SyncSingleHeadAdd) next case (4 u) then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym SyncSingleHeadAdd) next case (5 u) then show ?case by (metis Cons_tuv.hyps Sync.sym SyncSingleHeadAdd) next case (6 u) then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym SyncSingleHeadAdd) next case (7 u) then show ?case by (metis Cons_tuv.hyps Sync.sym SyncSingleHeadAdd) next case (8 tulist) from Cons_tuv(1)[OF 8(5) 8(9)] obtain uvlist where "uvlist setinterleaves ((va#ulist, va#vlist), A) \ tuv setinterleaves ((tlist, uvlist), A)" by blast with 8 show ?case using SyncSingleHeadAdd by (exI "uvlist", simp) blast next case (9 u) then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym SyncSingleHeadAdd) next case (10 tulist) from Cons_tuv(1)[OF 10(6) 10(10)] obtain uvlist where "uvlist setinterleaves ((ua#ulist, va#vlist), A) \ tuv setinterleaves ((tlist, uvlist), A)" by blast with 10 show ?case using SyncSingleHeadAdd by (exI "uvlist", simp) blast next case (11 u) then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym SyncSingleHeadAdd) next case (12 u) then show ?case using Cons_tuv.hyps by fastforce next case (13 u) then show ?case by (metis Cons_tuv.hyps Sync.sym SyncSingleHeadAdd) next case (14 u) then show ?case by (metis Cons_tuv.hyps Sync.sym SyncSingleHeadAdd) next case (15 u) then show ?case by (metis Cons_tuv.hyps Sync.sym SyncSingleHeadAdd) next case (16 u) then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym SyncSingleHeadAdd) next case (17 u) then show ?case by (metis Cons_tuv.hyps Sync.sym SyncSingleHeadAdd) next case (18 u) then show ?case using Cons_tuv.hyps by fastforce next case (19 u) then show ?case using Cons_tuv.hyps by fastforce next case (20 u) then show ?case by (metis Cons_tuv.hyps Cons_tuv.prems(1) Sync.sym SyncSingleHeadAdd) next case (21 u) then show ?case using Cons_tuv.hyps by fastforce qed qed qed qed qed lemma interleave_assoc_2: assumes *:"uv setinterleaves ((u, v), A)" and **:"tuv setinterleaves ((t, uv), A)" shows "\tu. tu setinterleaves ((t, u), A) \ tuv setinterleaves ((tu, v), A)" using "*" "**" Sync.sym interleave_assoc_1 by blast subsection\ Laws \ lemma Sync_commute: "(P \ S \ Q) = (Q \ S \ P)" by (simp add: Process_eq_spec mono_D_Sync mono_F_Sync) lemma mono_Sync_FD_oneside[simp]: "P \ P' \ (P \ S \ Q) \ (P' \ S \ Q)" apply(auto simp:le_ref_def F_Sync D_Sync) using F_subset_imp_T_subset front_tickFree_Nil by blast+ lemma mono_Sync_FD[simp]: "\P \ P'; Q \ Q'\ \ (P \ S \ Q) \ (P' \ S \ Q')" using mono_Sync_FD_oneside[of P P' S Q] mono_Sync_FD_oneside[of Q Q' S P] by (simp add: order_trans Sync_commute) lemma mono_Sync_ref: "\P \ P'; Q \ Q'\ \ (P \ S \ Q) \ (P' \ S \ Q')" using mono_Sync_sym[of Q Q' P S] mono_Sync[of P P' S Q'] below_trans by blast lemma Sync_BOT: "(P \ S \ \) = \" apply(auto simp add:Process_eq_spec, simp_all add:D_Sync F_Sync F_UU D_UU) apply (metis D_imp_front_tickFree append_Nil2 front_tickFree_append ftf_Sync is_processT is_processT2_TR) apply (metis Nil_elem_T Sync.si_empty1 append_Nil front_tickFree_Nil insertI1 non_tickFree_implies_nonMt) apply (metis D_imp_front_tickFree append_Nil2 front_tickFree_append ftf_Sync is_processT2_TR) by (metis Nil_elem_T Sync.si_empty1 append_Nil front_tickFree_Nil insertI1 non_tickFree_implies_nonMt) lemma Sync_SKIP_SKIP: "(SKIP \ S \ SKIP) = SKIP" apply(auto simp add:Process_eq_spec, simp_all add:D_Sync F_Sync F_SKIP D_SKIP) apply(elim exE conjE disjE, auto) apply (metis Sync.si_empty1 inf.idem insertI1 sup_commute sup_inf_absorb) apply(exI "[tick]", exI "[tick]", simp) by blast lemma Sync_SKIP_STOP: "(SKIP \S\ STOP) = STOP" proof(auto simp add:Process_eq_spec, simp_all add:D_Sync F_Sync F_SKIP D_SKIP F_STOP D_STOP, goal_cases) case (1 a b) then show ?case by (elim exE conjE disjE, auto) next case (2 a b) then show ?case apply(rule_tac x="[]" in exI, simp, rule_tac x="b-{tick}" in exI, simp, rule_tac x="b" in exI) by blast qed subsection\ Multi-Operators laws \ -lemma Mprefix_Sync_distr: - "\A \ S = {}; A' \ S; B \ S = {}; B' \ S\ \ - (\ x \ A \ A' \ P x) \ S \ (\ y \ B \ B' \ Q y) = (\ x \ A \ ( P x \ S \ (\ y \ B \ B' \ Q y))) - \ (\ y \ B \ ((\ x \ A \ A' \ P x) \ S \ Q y)) - \ (\ x \ A' \ B' \ ( P x \ S \ Q x))" -proof(auto simp add:Process_eq_spec, simp_all add:D_Sync F_Sync D_Mprefix F_Mprefix T_Mprefix D_Det F_Det T_Det, - elim exE disjE conjE, goal_cases) - case (1 a b t u X Y) - then show ?case by auto -next - case (2 a b t u X Y aa) - have 21:\a = u\ - using 2(5,7) EmptyLeftSync by blast - hence 22:\hd a \ ev ` B\ - using emptyLeftNonSync[rule_format, of a \insert tick (ev ` S)\ u \hd u\, simplified] - 2(10,4,5,7,9) by auto - with 2 show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, simp add:21) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, simp add:21) - apply(intro disjI2 conjI, simp_all add:21 22) - apply(exI aa, simp add:21) - apply(rule_tac disjI1, exI t, exI \tl u\, exI X, rule_tac conjI, simp, exI Y, simp) - by (meson SyncTlEmpty) -next - case (3 a b t u X Y aa) - have 31:\a = t\ - using "3"(6) "3"(8) EmptyLeftSync Sync.sym by blast - hence 32:\hd a \ ev ` A\ - using emptyLeftNonSync[rule_format, of a \insert tick (ev ` S)\ t \hd t\] - 3(10,2,5,6,8) Sync.sym by auto blast - with 3 show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, simp add:31) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, simp add:31) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, simp_all add:31 32) - apply(exI aa, simp) - apply(rule_tac disjI1, exI \tl t\, exI u, exI X, rule_tac conjI, simp, exI Y, simp) - by (simp add: Sync.sym SyncTlEmpty) -next - case (4 a b t u X Y aa ab) - hence \ (hd a = ev aa \ hd a \ ev ` A \ tl a setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd a = ev ab \ hd a \ ev ` B \ tl a setinterleaves ((t, tl u), insert tick (ev ` S))) - \ (hd a = ev ab \ aa = ab \ hd a \ ev ` (A' \ B') \ tl a setinterleaves ((tl t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp del:si_neq split:if_splits) - apply (metis (no_types, lifting) IntI UnE disjoint_iff event.simps(1) event.simps(3) imageE imageI list.sel(1) list.sel(3)) - by(metis (no_types, opaque_lifting) UnE image_Un list.sel(1) list.sel(3) subset_eq subset_image_iff)+ - with 4 show ?case - proof(elim disjE conjE, goal_cases) - case 41:1 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis empty_setinterleaving, simp_all) - apply(exI aa, simp) - apply(rule_tac disjI1, exI \tl t\, exI u, exI X, rule_tac conjI, simp, exI Y, simp) - by blast - next - case 42:2 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis empty_setinterleaving, simp_all) - apply(exI ab, simp) - apply(rule_tac disjI1, exI t, exI \tl u\, exI X, rule_tac conjI) - by blast+ - next - case 43:3 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis empty_setinterleaving, simp) - by(exI aa, simp, rule_tac disjI1, exI \tl t\, exI \tl u\, exI X, rule_tac conjI, simp, exI Y, simp) - qed -next - case (5 a b t u r v aa) - have 51:\a = t@v\ - using 5(7,8,11) EmptyLeftSync Sync.sym by blast - hence 52:\hd a \ ev ` A\ - using emptyLeftNonSync[rule_format, of r \insert tick (ev ` S)\ t \hd t\] - 5(10,11,2,8,9) Sync.sym by auto blast - with 5 show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, simp add:51) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, simp_all add:51) - apply(rule_tac disjI1, exI aa, simp) - apply(rule_tac disjI2, exI \tl t\, exI u, exI \tl r\, exI v, simp) - by (simp add: Sync.sym SyncTlEmpty tickFree_tl) -next - case (6 a b t u r v aa ab) - hence \(hd r = ev aa \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev ab \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply(simp del:si_neq split:if_splits, blast, metis empty_iff list.collapse) - apply(metis (no_types, lifting) event.simps(3) image_eqI insert_iff SyncHd_Tl) - defer - apply blast - - proof - - assume a1: \hd t \ ev ` S\ - and a2: \setinterleaving (t, insert tick (ev ` S), ev ab # tl u) = - {hd t # ua |ua. ua setinterleaves ((tl t, ev ab # tl u), insert tick (ev ` S))} - \ {ev ab # ua |ua. ua setinterleaves ((t, tl u), insert tick (ev ` S))}\ - have f3: "ev ab \ ev ` B" - using "6"(13) by blast - have f4: "hd t \ ev ` A" - using a1 "6"(10) "6"(2) by blast - show \hd r = hd t \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S)) \ - hd r = ev ab \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S))\ - apply auto - using "6"(8) a2 "6"(14) "6"(15) hd_Cons_tl apply fastforce - using "6"(8) a2 "6"(14) "6"(15) f4 hd_Cons_tl apply fastforce - using "6"(8) a2 "6"(14) "6"(15) hd_Cons_tl apply fastforce - using "6"(8) a2 "6"(14) "6"(15) f3 hd_Cons_tl apply fastforce - using "6"(8) a2 "6"(14) "6"(15) f3 hd_Cons_tl apply fastforce - using "6"(8) a2 "6"(14) "6"(15) f4 f3 hd_Cons_tl apply fastforce - using "6"(13) "6"(8) a2 "6"(14) "6"(15) hd_Cons_tl apply fastforce - using "6"(8) a2 "6"(14) "6"(15) f4 hd_Cons_tl apply fastforce - using "6"(8) a2 "6"(14) "6"(15) hd_Cons_tl by fastforce - qed - with 6 show ?case - proof(elim disjE conjE, goal_cases) - case 61:1 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving, simp_all) - apply(rule_tac conjI, metis empty_setinterleaving hd_append2) - apply(exI aa, rule_tac conjI, metis empty_setinterleaving hd_append2) - apply(rule_tac disjI2, exI \tl t\, exI u, exI \tl r\, exI v) - by (metis empty_setinterleaving tickFree_tl tl_append2) - next - case 62:2 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac conjI, metis empty_setinterleaving hd_append2) - apply(exI ab, rule_tac conjI, metis empty_setinterleaving hd_append2) - apply(rule_tac disjI2, exI t, exI \tl u\, exI \tl r\, exI v) - by (metis empty_setinterleaving tickFree_tl tl_append2) - qed -next - case (7 a b t u r v aa ab) - hence \ (hd r = ev aa \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev ab \ aa = ab \ hd r \ ev ` (A' \ B') \ tl r setinterleaves ((tl t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp del:si_neq split:if_splits) - apply (metis (no_types, opaque_lifting) IntI UnE empty_iff event.inject imageE imageI insert_iff SyncSameHdTl) - apply (metis empty_iff hd_Cons_tl, blast, blast) - proof - assume a1: "ev ab \ ev ` S" - assume a2: "hd t \ ev ` S" - assume a3: "hd t \ tick" - assume a4: "hd u = ev ab" - assume a5: "hd t \ ev ` (A \ A')" - assume a7: "A' \ S" - show \hd r = hd t \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))\ - apply safe - apply (metis "7"(8) Sync.sym a1 a2 a3 a4 insertCI insertE SyncHd_Tl) - apply (metis "7"(8) Sync.sym Un_iff a1 a2 a4 a5 a7 image_Un insert_iff sup.absorb_iff2 SyncHd_Tl SyncSameHdTl) - by (metis "7"(8) Sync.sym a1 a2 a3 a4 insertCI insertE SyncHd_Tl) - qed - with 7 show ?case - proof(elim disjE conjE, goal_cases) - case 71:1 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving, simp_all) - apply(rule_tac conjI, metis empty_setinterleaving hd_append2) - apply(exI aa, rule_tac conjI, metis empty_setinterleaving hd_append2) - apply(rule_tac disjI2, exI \tl t\, exI u, exI \tl r\, exI v) - by (metis empty_setinterleaving tickFree_tl tl_append2) +lemma Mprefix_Sync_distr: + \(\ x \ A \ A' \ P x) \ S \ (\ y \ B \ B' \ Q y) = + (\ x \ A \ ( P x \ S \ (\ y \ B \ B' \ Q y))) \ + (\ y \ B \ ((\ x \ A \ A' \ P x) \ S \ Q y)) \ + (\ x \ A' \ B' \ ( P x \ S \ Q x))\ + (is \?lhs A A' P B B' Q = ?rhs A A' P B B' Q\) + if sets_assms: \A \ S = {}\ \A' \ S\ \B \ S = {}\ \B' \ S\ + (* new proof by Benoît Ballenghien *) +proof (subst Process_eq_spec_optimized, safe) + fix s X + assume \(s, X) \ \ (?lhs A A' P B B' Q)\ + and same_div : \\ (?lhs A A' P B B' Q) = \ (?rhs A A' P B B' Q)\ + from this(1) consider + \\s_P s_Q X_P X_Q. (s_P, X_P) \ \ (Mprefix (A \ A') P) \ + (s_Q, X_Q) \ \ (Mprefix (B \ B') Q) \ + s setinterleaves ((s_P, s_Q), insert tick (ev ` S)) \ + X = (X_P \ X_Q) \ insert tick (ev ` S) \ X_P \ X_Q\ + | \s \ \ (?lhs A A' P B B' Q)\ + by (simp add: F_Sync D_Sync) blast + thus \(s, X) \ \ (?rhs A A' P B B' Q)\ + proof cases + case 1 + then obtain s_P s_Q X_P X_Q + where * : \(s_P, X_P) \ \ (Mprefix (A \ A') P)\ \(s_Q, X_Q) \ \ (Mprefix (B \ B') Q)\ + \s setinterleaves ((s_P, s_Q), insert tick (ev ` S))\ + \X = (X_P \ X_Q) \ insert tick (ev ` S) \ X_P \ X_Q\ by blast + thus \(s, X) \ \ (?rhs A A' P B B' Q)\ + proof (cases \s = [] \ hd s = tick\) + case True + with *(1, 2, 3) have ** : \s = [] \ s_P = [] \ s_Q = []\ + by (cases s_P; cases s_Q; force simp add: F_Mprefix subset_iff split: if_split_asm) + with *(1, 2, 4) sets_assms(1, 3) show \(s, X) \ \ (?rhs A A' P B B' Q)\ + by (auto simp add: subset_iff F_Det D_Det T_Det F_Mprefix image_Un)[1] + next + case False + then obtain e where *** : \s \ []\ \hd s = ev e\ by (meson event.exhaust) + from *(1, 2, 3) *** sets_assms consider + \e \ A \ s_P \ [] \ hd s_P = ev e \ (tl s_P, X_P) \ \ (P e) \ + tl s setinterleaves ((tl s_P, s_Q), insert tick (ev ` S))\ | + \e \ B \ s_Q \ [] \ hd s_Q = ev e \ (tl s_Q, X_Q) \ \ (Q e) \ + tl s setinterleaves ((s_P, tl s_Q), insert tick (ev ` S))\ | + \e \ A' \ e \ B' \ s_P \ [] \ s_Q \ [] \ hd s_P = ev e \ hd s_Q = ev e \ (tl s_P, X_P) \ \ (P e) \ + (tl s_Q, X_Q) \ \ (Q e) \ tl s setinterleaves ((tl s_P, tl s_Q), insert tick (ev ` S))\ + by (cases s_P; cases s_Q; auto simp add: F_Mprefix split: if_split_asm) + thus \(s, X) \ \ (?rhs A A' P B B' Q)\ + apply cases + apply (simp_all add: F_Det ***(1)) + apply (rule disjI1, simp add: F_Mprefix ***, simp add: F_Sync, rule disjI1) + apply (rule_tac x = \tl s_P\ in exI, rule_tac x = s_Q in exI, rule_tac x = X_P in exI, + simp, rule_tac x = X_Q in exI, simp add: *(2, 4)) + apply (rule disjI2, rule disjI1, simp add: F_Mprefix ***, simp add: F_Sync, rule disjI1) + apply (rule_tac x = s_P in exI, rule_tac x = \tl s_Q\ in exI, rule_tac x = X_P in exI, + simp add: *(1), rule_tac x = X_Q in exI, simp add: *(4)) + apply (rule disjI2, rule disjI2, simp add: F_Mprefix *** F_Sync, rule disjI1) + by (rule_tac x = \tl s_P\ in exI, rule_tac x = \tl s_Q\ in exI, rule_tac x = X_P in exI, + simp, rule_tac x = X_Q in exI, simp add: *(4)) + qed next - case 72:2 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac conjI, metis empty_setinterleaving hd_append2) - apply (exI aa, rule_tac conjI, metis empty_setinterleaving hd_append2) - apply (rule_tac disjI2, exI \tl t\, exI \tl u\, exI \tl r\, exI v) - by (metis empty_setinterleaving tickFree_tl tl_append2) - qed -next - case (8 a b u t r v aa) - have 81:\r = u\ - using 8 EmptyLeftSync Sync.sym by blast - hence 82:\hd r \ ev ` B\ - using emptyLeftNonSync[rule_format, of r \insert tick (ev ` S)\ u \hd u\] - 8 Sync.sym by auto blast - with 8 show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, simp add:81) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(intro disjI2 conjI, simp_all add:81 82) - apply(exI aa, simp add:81) - apply(rule_tac disjI2, exI \tl u\, exI t, exI \tl r\, exI v) - by (simp add: "81" Sync.sym SyncTlEmpty tickFree_tl) -next - case (9 a b u t r v ab aa) - hence \(hd r = ev aa \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev ab \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply(simp del:si_neq split:if_splits, metis disjoint_iff event.inject image_iff, blast, blast) - apply (simp add: Sync.sym image_Un list.exhaust_sel rev_image_eqI sup.absorb_iff2) - defer - apply (metis (no_types, lifting) Sync.sym event.simps(3) imageI insertE insertI2 SyncHd_Tl) - proof - - assume a1: "hd u \ ev ` B \ hd u \ ev ` B'" - assume a2: "r setinterleaves ((t, u), insert tick (ev ` S))" - assume a3: "setinterleaving (u, insert tick (ev ` S), ev aa # tl t) = - {ev aa # ua |ua. ua setinterleaves ((u, tl t), insert tick (ev ` S))} \ - {hd u # ua |ua. ua setinterleaves ((tl u, ev aa # tl t), insert tick (ev ` S))}" - assume a4: "hd t = ev aa" - assume a5: "t \ []" - assume a6: "aa \ A" - assume a7: "hd u \ ev ` S" - assume a8: "B' \ S = S" - obtain ees :: "'a event list \ 'a event list" where - f9: "((\es. r = hd u # es \ es setinterleaves ((tl u, ev aa # tl t), insert tick (ev ` S))) \ - r = hd u # ees r \ ees r setinterleaves ((tl u, ev aa # tl t), insert tick (ev ` S))) \ - ((\es. r = hd u # es \ es setinterleaves ((tl u, ev aa # tl t), insert tick (ev ` S))) \ - (\es. r \ hd u # es \ es \ setinterleaving (tl u, insert tick (ev ` S), ev aa # tl t)))" - by fastforce - have f10: "ev aa # tl t = t" - using a5 a4 hd_Cons_tl by fastforce - then have f11: "r \ {ev aa # es |es. es setinterleaves ((u, tl t), insert tick (ev ` S))} \ - {hd u # es |es. es setinterleaves ((tl u, ev aa # tl t), insert tick (ev ` S))}" - using a3 a2 by (simp add: Sync.sym) - have f12: "\e f a A. (e::'a event) \ f (a::'a) \ a \ A \ e \ f ` A" - by (meson image_eqI) - moreover - { assume "r \ hd u # ees r \ ees r \ setinterleaving (tl u, insert tick (ev ` S), ev aa # tl t)" - then obtain eesa :: "'a event list \ 'a event list" where - "r = ev aa # eesa r \ eesa r setinterleaves ((u, tl t), insert tick (ev ` S))" - using f11 f9 by fast - then have "hd r = ev aa \ hd r \ ev ` A \ tl r setinterleaves ((u, tl t), insert tick (ev ` S)) \ hd r = hd u \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S))" - using f12 a6 by (metis list.sel(1) list.sel(3)) } - ultimately show "hd r = ev aa \ hd r \ ev ` A \ tl r setinterleaves ((u, tl t), insert tick (ev ` S)) \ hd r = hd u \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S))" - using f10 a8 a7 a1 by (metis Sync.sym UnI1 image_Un list.sel(1) list.sel(3)) - qed - with 9 show ?case - proof(elim disjE conjE, goal_cases) - case 91:1 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving, simp) - apply(rule_tac conjI, metis empty_setinterleaving hd_append image_eqI) - apply(exI aa, rule_tac conjI, metis empty_setinterleaving hd_append2) - apply(rule_tac disjI2, exI u, exI \tl t\, exI \tl r\, exI v, simp) - by (metis Sync.sym empty_setinterleaving tickFree_tl tl_append2) - next - case 92:2 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving, simp) - apply(rule_tac conjI, metis empty_setinterleaving hd_append) - apply(exI ab, rule_tac conjI, metis empty_setinterleaving hd_append2) - apply(rule_tac disjI2, exI \tl u\, exI \t\, exI \tl r\, exI v, simp) - by (metis Sync.sym empty_setinterleaving tickFree_tl tl_append2) - qed -next - case (10 a b u t r v ab aa) - hence \ (hd r = ev ab \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S))) - \ (hd r = ev ab \ aa = ab \ hd r \ ev ` (A' \ B') \ tl r setinterleaves ((tl t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp del:si_neq split:if_splits) - apply (metis (no_types, opaque_lifting) IntI Sync.sym UnE empty_iff event.inject image_iff insert_iff SyncSameHdTl) - apply (metis insert_iff SyncSameHdTl) - defer - apply (metis imageI subsetD) - apply (metis imageI subsetD) - proof - assume a1: "hd u \ tick \ hd u \ ev ` S" - assume a2: "ev aa \ ev ` S" - assume a3: "hd t = ev aa" - assume a4: "r setinterleaves ((u, t), insert tick (ev ` S))" - assume a5: "B' \ S" - show \hd r = hd u \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S))\ - apply safe - apply (metis "10"(8) Sync.sym a1 a2 a3 insertCI insertE SyncHd_Tl) - apply (metis (no_types, opaque_lifting) "10"(10) "10"(8) Sync.sym UnE a1 a2 a3 a5 image_Un insert_iff subset_eq subset_image_iff SyncHd_Tl) - by (metis Sync.sym a1 a2 a3 a4 insertCI insertE SyncHd_Tl) - qed - with 10 show ?case - proof(elim disjE conjE, goal_cases) - case 101:1 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving, simp) - apply(rule_tac conjI, metis empty_setinterleaving hd_append) - apply(exI ab, rule_tac conjI, metis empty_setinterleaving hd_append2) - apply(rule_tac disjI2, exI \tl u\, exI \t\, exI \tl r\, exI v, simp) - by (metis Sync.sym empty_setinterleaving tickFree_tl tl_append2) - next - case 102:2 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac conjI, metis empty_setinterleaving hd_append2) - apply (exI aa, rule_tac conjI, metis empty_setinterleaving hd_append2) - apply (rule_tac disjI2, exI \tl u\, exI \tl t\, exI \tl r\, exI v, simp) - by (metis Sync.sym empty_setinterleaving tickFree_tl tl_append2) - qed -next - case (11 a b t u r v aa) - have 111:\a = t@v\ - using 11(7,8,11) EmptyLeftSync Sync.sym by blast - hence 112:\hd a \ ev ` A\ - using emptyLeftNonSync[rule_format, of r \insert tick (ev ` S)\ t \hd t\] - 11(10,11,2,8,9) Sync.sym by auto blast - with 11 show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, simp add:111) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, simp_all add:111) - apply(rule_tac disjI1, exI aa, simp) - apply(rule_tac disjI2, exI \tl t\, exI u, exI \tl r\, exI v, simp) - by (simp add: Sync.sym SyncTlEmpty tickFree_tl) -next - case (12 a b t u r v aa ab) - hence \(hd r = ev aa \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev ab \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply(simp del:si_neq split:if_splits, metis disjoint_iff event.inject image_iff, blast) - apply (metis (no_types, lifting) event.simps(3) imageI insert_iff SyncHd_Tl) - defer apply blast - proof - - assume a1: "ab \ B" - assume a2: "r setinterleaves ((t, u), insert tick (ev ` S))" - assume a3: "setinterleaving (t, insert tick (ev ` S), ev ab # tl u) = - {hd t # ua |ua. ua setinterleaves ((tl t, ev ab # tl u), insert tick (ev ` S))} \ - {ev ab # ua |ua. ua setinterleaves ((t, tl u), insert tick (ev ` S))}" - assume a4: "hd u = ev ab" - assume a5: "u \ []" - assume a6: "hd t \ ev ` (A \ A')" - assume a7: "A' \ S" - assume a8: "hd t \ ev ` S" - have f9: "hd t \ ev ` A" - using a8 a7 a6 by blast - show "hd r = hd t \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S)) \ - hd r = ev ab \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S))" - proof - - obtain ees :: "'a event list \ 'a event list" where - f1: "((\es. r = hd t # es \ es setinterleaves ((tl t, ev ab # tl u), insert tick (ev ` S))) \ - r = hd t # ees r \ ees r setinterleaves ((tl t, ev ab # tl u), insert tick (ev ` S))) \ - ((\es. r = hd t # es \ es setinterleaves ((tl t, ev ab # tl u), insert tick (ev ` S))) \ - (\es. r \ hd t # es \ es \ setinterleaving (tl t, insert tick (ev ` S), ev ab # tl u)))" - by meson - have f2: "ev ab # tl u = u" - using a5 by (subst a4[symmetric], rule list.collapse) - moreover - { assume "r \ hd t # ees r \ ees r \ setinterleaving (tl t, insert tick (ev ` S), ev ab # tl u)" - then obtain eesa :: "'a event list \ 'a event list" where - "r = ev ab # eesa r \ eesa r setinterleaves ((t, tl u), insert tick (ev ` S))" - using f2 f1 a2 a3 by fastforce - then have ?thesis - by (metis a1 imageI list.sel(1) list.sel(3)) } - ultimately show ?thesis - by (metis (no_types) f9 list.sel(1) list.sel(3)) - qed - qed - with 12 show ?case - proof(elim disjE conjE, goal_cases) - case 121:1 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving, simp) - apply(exI aa, rule_tac conjI, simp) - by (rule_tac disjI2, exI \tl t\, exI u, exI \tl r\, exI v, simp) - next - case 122:2 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving, simp) - by (rule_tac disjI2, exI \t\, exI \tl u\, exI \tl r\, exI v, simp) blast + show \s \ \ (?lhs A A' P B B' Q) \ (s, X) \ \ (?rhs A A' P B B' Q)\ + using same_div NF_ND by blast qed next - case (13 a b t u r v aa ab) - hence \ (hd r = ev aa \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev ab \ aa = ab \ hd r \ ev ` (A' \ B') \ tl r setinterleaves ((tl t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp del:si_neq split:if_splits) - apply (metis (no_types, opaque_lifting) IntI UnE empty_iff event.inject image_iff insert_iff SyncSameHdTl) - apply (metis empty_iff hd_Cons_tl, blast, blast) - proof - assume a1: \r setinterleaves ((t, u), insert tick (ev ` S))\ - assume a2: \hd t \ ev ` (A \ A')\ - assume a3: \hd t \ tick\ - assume a4: \hd t \ ev ` S\ - assume a5: \ev ab \ ev ` S\ - assume a6: \setinterleaving (t, insert tick (ev ` S), ev ab # tl u) = {hd t # ua |ua. ua setinterleaves ((tl t, ev ab # tl u), insert tick (ev ` S))}\ - show \hd r = hd t \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))\ - apply safe - using "13"(14) "13"(15) "13"(8) a6 list.collapse apply fastforce - using "13"(14) "13"(15) "13"(2) a1 a2 a4 a6 list.collapse apply fastforce - by (metis "13"(15) Sync.sym a5 a3 a4 a1 insertCI insertE SyncHd_Tl) - qed - with 13 show ?case - proof(elim disjE conjE, goal_cases) - case 131:1 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving, simp_all) - apply(exI aa, rule_tac conjI, simp) - apply(rule_tac disjI2, exI \tl t\, exI u, exI \tl r\, exI v) - using 131(5) by blast - next - case 132:2 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac conjI, metis empty_setinterleaving hd_append2) - apply (exI aa, rule_tac conjI, metis empty_setinterleaving hd_append2) - apply (rule_tac disjI2, exI \tl t\, exI \tl u\, exI \tl r\, exI v) - by (metis empty_setinterleaving tl_append2) - qed -next - case (14 a b u t r v aa) - have 141:\r = u\ - using 14 EmptyLeftSync Sync.sym by blast - hence 142:\hd r \ ev ` B\ - using emptyLeftNonSync[rule_format, of r \insert tick (ev ` S)\ u \hd u\] - 14 Sync.sym by auto blast - with 14 show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, simp add:141) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(intro disjI2 conjI, simp_all add:141 142) - apply(exI aa, simp add:141) - apply(rule_tac disjI2, exI \tl u\, exI t, exI \tl r\, exI v) - by (simp add: "141" Sync.sym SyncTlEmpty tickFree_tl) -next - case (15 a b u t r v ab aa) - hence \(hd r = ev aa \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev ab \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply(simp del:si_neq split:if_splits, metis disjoint_iff event.inject image_iff, blast, blast) - defer - apply (metis (no_types, lifting) Sync.sym event.simps(3) imageI insertE insertI2 SyncHd_Tl) - proof - - assume a1: "hd u \ ev ` (B \ B')" - assume a2: "B' \ S" - assume a3: "hd u \ ev ` S" - assume a4: "r setinterleaves ((u, t), insert tick (ev ` S))" - assume a5: "setinterleaving (ev aa # tl t, insert tick (ev ` S), u) = - {ev aa # ua |ua. ua setinterleaves ((tl t, u), insert tick (ev ` S))} \ - {hd u # ua |ua. ua setinterleaves ((ev aa # tl t, tl u), insert tick (ev ` S))}" - assume a6: "hd t = ev aa" - assume a7: "t \ []" - assume a8: "aa \ A" - have f9: "hd u \ ev ` B" - using a3 a2 a1 by fast - have f10: "ev aa # tl t = t" - using a7 a6 hd_Cons_tl by fastforce - then have "r \ {ev aa # es |es. es setinterleaves ((tl t, u), insert tick (ev ` S))} \ - {hd u # es |es. es setinterleaves ((ev aa # tl t, tl u), insert tick (ev ` S))}" - using a5 a4 by (simp add: Sync.sym) - then show "hd r = ev aa \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S)) \ - hd r = hd u \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S))" - using f10 f9 a8 by fastforce - qed - with 15 show ?case - proof(elim disjE conjE, goal_cases) - case 151:1 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving, simp) - apply(rule_tac disjI2, exI u, exI \tl t\, exI \tl r\, exI v, simp) - using Sync.sym by blast - next - case 152:2 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving, simp) - apply(exI ab, rule_tac conjI, simp) - apply(rule_tac disjI2, exI \tl u\, exI \t\, exI \tl r\, exI v, simp) - by (metis Sync.sym) - qed -next - case (16 a b u t r v ab aa) - hence \ (hd r = ev ab \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S))) - \ (hd r = ev ab \ aa = ab \ hd r \ ev ` (A' \ B') \ tl r setinterleaves ((tl t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp del:si_neq split:if_splits) - apply (metis (no_types, opaque_lifting) "16"(11) "16"(3) Int_iff Sync.sym UnE empty_iff event.inject imageE imageI insertCI SyncSameHdTl) - apply (metis insert_iff SyncSameHdTl) - defer - apply (metis imageI subsetD) - apply (metis imageI subsetD) - proof - - assume a1: "hd u \ tick \ hd u \ ev ` S" - assume a2: "B' \ S" - assume a3: "hd u \ ev ` (B \ B')" - assume a4: "r setinterleaves ((u, t), insert tick (ev ` S))" - assume a5: "setinterleaving (ev aa # tl t, insert tick (ev ` S), u) = - {hd u # ua |ua. ua setinterleaves ((ev aa # tl t, tl u), insert tick (ev ` S))}" - assume a6: "hd t = ev aa" - assume a7: "t \ []" - have f8: "B' \ S = S" - using a2 by blast - have f9: "ev aa # tl t = t" - using a7 a6 hd_Cons_tl by fastforce - then have "\es. r = hd u # es \ es setinterleaves ((ev aa # tl t, tl u), insert tick (ev ` S))" - using a5 a4 by (simp add: Sync.sym) - then show "hd r = hd u \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S)) \ - hd r = hd u \ aa = ab \ hd r \ ev ` (A' \ B') \ tl r setinterleaves ((tl t, tl u), insert tick (ev ` S))" - using f9 f8 a3 a1 by (metis Un_iff image_Un list.sel(1) list.sel(3)) - qed - with 16 show ?case - proof(elim disjE conjE, goal_cases) - case 161:1 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving, simp) - apply(exI ab, rule_tac conjI, simp) - apply(rule_tac disjI2, exI \tl u\, exI \t\, exI \tl r\, exI v, simp) - by (metis Sync.sym) - next - case 162:2 - then show ?case - apply(rule_tac disjI2, rule_tac disjI1, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac disjI2, rule_tac conjI, metis append_is_Nil_conv empty_setinterleaving) - apply(rule_tac conjI, metis empty_setinterleaving hd_append2) - apply (exI aa, rule_tac conjI, metis empty_setinterleaving hd_append2) - apply (rule_tac disjI2, exI \tl u\, exI \tl t\, exI \tl r\, exI v, simp) - by (metis Sync.sym) - qed -next - case (17 aA bB) - have tact:\\P Q. \t u r v. P t u r v \ Q \ (\ t u r v. P t u r v) \ Q\ - by fastforce - from 17 show ?case - proof (safe, goal_cases) - case 1 - obtain X where 11: "(X::'a event set)= bB - ev ` A'" by auto - obtain Y where 12: "(Y::'a event set)= bB - ev ` B'" by auto - from 1 11 12 have 13: "bB = (X \ Y) \ insert tick (ev ` S) \ X \ Y" - by auto - with 11 12 1 show ?case - apply(exI \[]\, exI \[]\, exI X, rule_tac conjI) - apply safe - apply (metis disjoint_iff_not_equal imageI) - apply (metis image_iff) - apply (metis disjoint_iff imageI) - apply (metis IntI image_eqI) - apply (metis image_eqI) - apply (metis rev_image_eqI) - by (exI Y) (simp add: disjoint_iff image_Un) - next - case (2 x a t u X Y) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI1, exI \ev x#t\, exI \[]\, exI X, rule conjI, simp) - by (exI Y, simp, metis disjoint_iff event.inject hd_Cons_tl imageE) - next - case (3 x a t u X Y xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI1, exI \ev x#t\, exI \u\, exI X, rule conjI, simp) - apply (exI Y, simp) - using hd_Cons_tl SyncSingleHeadAdd by fastforce+ - next - case (4 x a t u X Y xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI1, exI \ev x#t\, exI \u\, exI X, rule conjI, simp) - apply (exI Y, simp) - using hd_Cons_tl SyncSingleHeadAdd by fastforce+ - next - case (5 x a t u r v) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#t\, exI \[]\, exI \ev x#r\, exI v, simp) - by (metis disjoint_iff event.inject hd_Cons_tl imageE) - next - case (6 x a t u r v aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#t\, exI \u\, exI \ev x#r\, exI v, simp) - using hd_Cons_tl SyncSingleHeadAdd by fastforce+ - next - case (7 x a t u r v aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#t\, exI \u\, exI \ev x#r\, exI v, simp) - using hd_Cons_tl SyncSingleHeadAdd by fastforce+ - next - case (8 x a u t r v xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \u\, exI \ev x#t\, exI \ev x#r\, exI \v\, simp) - using si_neq[of \ev x\ \t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast) - by (metis Sync.sym SyncSingleHeadAdd event.distinct(1) insert_iff list.collapse) blast - next - case (9 x a u t r v xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \u\, exI \ev x#t\, exI \ev x#r\, exI \v\, simp) - using si_neq[of \ev x\ \t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast, blast) - by (metis SyncHdAdd1 event.distinct(1) insert_iff list.collapse) - next - case (10 x a t u r v) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#t\, exI \[]\, exI \ev x#r\, exI \[]\, simp) - by (metis disjoint_iff event.inject hd_Cons_tl imageE) - next - case (11 x a t u r v aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#t\, exI \u\, exI \ev x#r\, exI \[]\, simp) - using SyncSingleHeadAdd list.collapse by fastforce - next - case (12 x a t u r v aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#t\, exI \u\, exI \ev x#r\, exI \[]\, simp) - using SyncSingleHeadAdd list.collapse by fastforce - next - case (13 x a u t r v xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \u\, exI \ev x#t\, exI \ev x#r\, exI \[]\, simp) - using si_neq[of \ev x\ \t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast) - by (metis Sync.sym SyncSingleHeadAdd event.distinct(1) insert_iff list.collapse) blast - next - case (14 x a u t r v xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \u\, exI \ev x#t\, exI \ev x#r\, exI \[]\, simp) - using si_neq[of \ev x\ \t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast, blast) - by (metis SyncHdAdd1 event.distinct(1) insert_iff list.collapse) - next - case (15 x a t u X Y) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI1, exI \[]\, exI \ev x#u\, exI X, rule conjI, simp) - by (exI Y, simp, metis disjoint_iff event.inject hd_Cons_tl imageE) - next - case (16 x a t u X Y xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI1, exI \t\, exI \ev x#u\, exI X, rule conjI, simp, exI Y, simp) - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \ev x\ \tl u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast, blast) - by (metis Sync.sym SyncSingleHeadAdd event.distinct(1) insert_iff list.collapse) blast - next - case (17 x a t u X Y xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI1, exI \t\, exI \ev x#u\, exI X, rule conjI, simp) - apply (exI Y, simp) - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \ev x\ \tl u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast) sledgehammer - using SyncHdAdd by fastforce blast+ - next - case (18 x a t u r v xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \t\, exI \ev x#u\, exI \ev x#r\, exI \v\, simp) - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \ev x\ \u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast, blast) - by (metis Sync.sym SyncSingleHeadAdd event.distinct(1) insert_iff list.collapse) blast - next - case (19 x a t u r v xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \t\, exI \ev x#u\, exI \ev x#r\, exI \v\, simp) - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \ev x\ \u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast) - by (metis SyncHdAdd1 event.distinct(1) insert_iff list.collapse) blast+ - next - case (20 x a u t r v) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#u\, exI \[]\, exI \ev x#r\, exI \v\, simp) - by (metis disjoint_iff event.inject hd_Cons_tl imageE) - next - case (21 x a u t r v aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#u\, exI \t\, exI \ev x#r\, exI \v\, simp) - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \ev x\ \u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast, blast) - by (metis list.collapse) blast - next - case (22 x a u t r v aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#u\, exI \t\, exI \ev x#r\, exI \v\, simp) - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \ev x\ \u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast) - by (metis list.exhaust_sel) blast+ - next - case (23 x a t u r v xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \t\, exI \ev x#u\, exI \ev x#r\, exI \[]\, simp) - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \ev x\ \u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast, blast) - by (metis Sync.sym SyncSingleHeadAdd event.distinct(1) insert_iff list.collapse) blast - next - case (24 x a t u r v xa aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \t\, exI \ev x#u\, exI \ev x#r\, exI \[]\, simp) - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \ev x\ \u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast) sledgehammer - using SyncHdAdd1 list.collapse by fastforce blast+ - next - case (25 x a t u r v) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#t\, exI \[]\, exI \ev x#r\, exI \[]\, simp) - by (metis disjoint_iff event.inject hd_Cons_tl imageE) - next - case (26 x a u t r v aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#u\, exI \t\, exI \ev x#r\, exI \[]\, simp) - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \ev x\ \u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast, blast) - by (metis list.collapse) blast - next - case (27 x a u t r v aa) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#u\, exI \t\, exI \ev x#r\, exI \[]\, simp) - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \ev x\ \u\] - apply (simp add: SyncSingleHeadAdd split: if_splits, blast, blast) sledgehammer - by (metis list.collapse) blast+ - next - case (28 x a t u X Y) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI1, exI \ev x#t\, exI \ev x#u\, exI X, rule conjI, simp) - by (exI Y, simp, metis hd_Cons_tl image_eqI in_mono) - next - case (29 x a t u r v) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#t\, exI \ev x#u\, exI \ev x#r\, exI \v\, simp) - by (metis image_eqI list.exhaust_sel subsetD) - next - case (30 x a t u r v) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#t\, exI \ev x#u\, exI \ev x#r\, exI \v\, simp) - by (metis image_eqI list.exhaust_sel subsetD) - next - case (31 x a t u r v) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#t\, exI \ev x#u\, exI \ev x#r\, exI \[]\, simp) - by (metis image_eqI list.exhaust_sel subsetD) - next - case (32 x a t u r v) - then show ?case - apply(erule_tac tact) - apply(rule_tac disjI2, exI \ev x#t\, exI \ev x#u\, exI \ev x#r\, exI \[]\, simp) - by (metis image_eqI list.exhaust_sel subsetD) - qed + + fix s X + { fix P Q A A' B B' + assume assms : \s \ []\ \(s, X) \ \ (\x \ A \ (P x \S\ Mprefix (B \ B') Q))\ + \A \ S = {}\ \A' \ S\ \B \ S = {}\ \B' \ S\ + and same_div : \\ (?lhs A A' P B B' Q) = \ (?rhs A A' P B B' Q)\ + from assms(1, 2) obtain e + where * : \e \ A\ \hd s = ev e\ \(tl s, X) \ \ (P e \S\ Mprefix (B \ B') Q)\ + by (auto simp add: F_Mprefix) + from "*" assms(1) consider + \\s_P s_Q X_P X_Q. (s_P, X_P) \ \ (P e) \ (s_Q, X_Q) \ \ (Mprefix (B \ B') Q) \ + tl s setinterleaves ((s_P, s_Q), ev ` S \ {tick}) \ X = (X_P \ X_Q) \ (ev ` S \ {tick}) \ X_P \ X_Q\ + | \s \ \ (\x \ A \ (P x \S\ Mprefix (B \ B') Q))\ + by (force simp add: F_Sync D_Mprefix D_Sync) + hence \(s, X) \ \ (?lhs A A' P B B' Q)\ + proof cases + case 1 + then obtain s_P s_Q X_P X_Q + where ** : \(s_P, X_P) \ \ (P e)\ \(s_Q, X_Q) \ \ (Mprefix (B \ B') Q)\ + \tl s setinterleaves ((s_P, s_Q), ev ` S \ {tick})\ + \X = (X_P \ X_Q) \ (ev ` S \ {tick}) \ X_P \ X_Q\ by blast + show \(s, X) \ \ (?lhs A A' P B B' Q)\ + apply (simp add: F_Sync, rule disjI1) + apply (rule_tac x = \ev e # s_P\ in exI, rule_tac x = s_Q in exI, + rule_tac x = X_P in exI, intro conjI) + apply (simp add: F_Mprefix image_iff "*"(1) "**"(1)) + apply (rule_tac x = X_Q in exI, simp add: "**"(2, 4)) + apply (subst list.collapse[OF assms(1), symmetric]) + using "*"(1, 2) "**"(3) assms(3) by (cases s_Q) auto + next + assume \s \ \ (\x\A \ (P x \S\ Mprefix (B \ B') Q))\ + hence \s \ \ (?lhs A A' P B B' Q)\ by (simp add: same_div D_Det) + thus \(s, X) \ \ (?lhs A A' P B B' Q)\ using NF_ND by blast + qed + } note * = this + + { assume assms : \s \ []\ \(s, X) \ \ (\x\A' \ B' \ (P x \S\ Q x))\ + and same_div : \\ (?lhs A A' P B B' Q) = \ (?rhs A A' P B B' Q)\ + then obtain e where * : \e \ A'\ \e \ B'\ \hd s = ev e\ \(tl s, X) \ \ (P e \S\ Q e)\ + by (auto simp add: F_Mprefix) + have inside: \e \ S\ using "*"(2) that(4) by blast + from "*" assms(1) consider + \\s_P s_Q X_P X_Q. (s_P, X_P) \ \ (P e) \ (s_Q, X_Q) \ \ (Q e) \ + tl s setinterleaves ((s_P, s_Q), ev ` S \ {tick}) \ + X = (X_P \ X_Q) \ (ev ` S \ {tick}) \ X_P \ X_Q\ + | \s \ \ (\x\A' \ B' \ (P x \S\ Q x))\ + by (force simp add: F_Sync D_Mprefix D_Sync) + hence \(s, X) \ \ (?lhs A A' P B B' Q)\ + proof cases + case 1 + then obtain s_P s_Q X_P X_Q + where ** : \(s_P, X_P) \ \ (P e)\ \(s_Q, X_Q) \ \ (Q e)\ + \tl s setinterleaves ((s_P, s_Q), ev ` S \ {tick})\ + \X = (X_P \ X_Q) \ (ev ` S \ {tick}) \ X_P \ X_Q\ by blast + show \(s, X) \ \ (?lhs A A' P B B' Q)\ + apply (subst list.collapse[OF assms(1), symmetric], simp add: F_Sync "*"(3), rule disjI1) + apply (rule_tac x = \ev e # s_P\ in exI, rule_tac x = \ev e # s_Q\ in exI) + apply (rule_tac x = X_P in exI, intro conjI, simp add: F_Mprefix "*"(1) "**"(1)) + using "**"(3) by (rule_tac x = X_Q in exI) (simp add: "*"(2) "**"(2, 4) inside F_Mprefix) + next + assume \s \ \ (\x\A' \ B' \ (P x \S\ Q x))\ + hence \s \ \ (?lhs A A' P B B' Q)\ by (simp add: same_div D_Det) + thus \(s, X) \ \ (?lhs A A' P B B' Q)\ using NF_ND by blast + qed + } note ** = this + + have *** : \X \ ev ` (A' \ B') = {} \ X \ ev ` A = {} \ X \ ev ` B = {} \ + ([], X) \ \ (?lhs A A' P B B' Q)\ + apply (simp add: F_Sync, rule disjI1) + apply (rule_tac x = \[]\ in exI, rule_tac x = \[]\ in exI, simp add: F_Mprefix) + apply (rule_tac x = \X - (ev ` (A' - B'))\ in exI, intro conjI, fastforce) + apply (rule_tac x = \X - (ev ` (B' - A'))\ in exI, intro conjI, fastforce) + using sets_assms(2, 4) by auto + + assume same_div : \\ (?lhs A A' P B B' Q) = \ (?rhs A A' P B B' Q)\ + hence same_div_sym : \\ (?lhs B B' Q A A' P) = \ (?rhs B B' Q A A' P)\ + by (subst Sync_commute, subst Int_commute, subst Det_commute, simp add: Sync_commute) + show \(s, X) \ \ (?rhs A A' P B B' Q) \ (s, X) \ \ (?lhs A A' P B B' Q)\ + apply (auto simp add: F_Det) + apply (rule "***"; simp add: F_Mprefix) + apply (rule "*"; simp add: sets_assms same_div) + apply (subst (asm) Sync_commute, subst Sync_commute) + apply (rule "*"; simp add: sets_assms same_div_sym) + apply (erule "**", assumption, rule same_div) + apply (simp add: D_Mprefix D_Det)[1] + by (simp add: T_Mprefix T_Det) next - case (18 aA) - then show ?case - proof(elim exE disjE conjE, goal_cases) - case (1 t u r v a) - have 11:\r = t\ - using 1 EmptyLeftSync Sync.sym by blast - hence 12:\hd aA \ ev ` A\ - using emptyLeftNonSync[rule_format, of r \insert tick (ev ` S)\ t \hd t\] - 1 Sync.sym by auto blast - with 1 11 show ?case - apply(rule_tac disjI1, simp) - by (metis Sync.sym SyncTlEmpty tickFree_tl) - next - case (2 t u r v a aa) - hence \(hd r = ev a \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev aa \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply(simp del:si_neq split:if_splits, blast , metis empty_iff hd_Cons_tl) - apply (metis (no_types, lifting) event.simps(3) imageI insert_iff SyncHd_Tl) - defer - apply blast - proof(safe, simp_all add: in_mono, goal_cases) - case (1 x) - show ?case - using "1"(18) "1"(21) "2"(14) "2"(15) "2"(8) list.collapse by fastforce - next - case (2 x) - show ?case - using "2"(13) "2"(14) "2"(18) "2"(20) "2"(21) "2"(8) list.collapse by fastforce - next - case (3 x) - show ?case - using "2"(14) "2"(15) "2"(8) "3"(18) "3"(21) list.collapse by fastforce - next - case (4 x) - show ?case - using "2"(13) "2"(14) "2"(15) "2"(8) "4"(18) "4"(21) list.collapse by fastforce - next - case (5 x) - show ?case - using "2"(14) "2"(15) "2"(8) "5"(18) "5"(21) list.collapse by fastforce - next - case (6 x) - show ?case - using "6"(13) "6"(14) "6"(18) "6"(8) hd_Cons_tl "2"(13) "6"(20) "6"(21) by fastforce - next - case (7 x) - show ?case - using "2"(13) "2"(14) "2"(15) "2"(8) "7"(18) "7"(21) list.collapse by fastforce - next - case (8 x) - show ?case - using "8"(20) "8"(21) "8"(13) "8"(14) "8"(18) "8"(8) hd_Cons_tl by fastforce - next - case (9 x) - show ?case - using "9"(21) "9"(13) "9"(14) "9"(13) "9"(14) "9"(18) "9"(8) hd_Cons_tl by fastforce - qed - with 2 show ?case - by (simp add: hd_append) (metis empty_setinterleaving tickFree_tl) - next - case (3 t u r v a aa) - hence \ (hd r = ev a \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev aa \ aa = a \ hd r \ ev ` (A' \ B') \ tl r setinterleaves ((tl t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp del:si_neq split:if_splits) - apply (simp add: UnE disjoint_iff imageE insertI2 rev_image_eqI) - apply (metis (no_types, lifting) Int_iff Un_iff event.inject imageE image_eqI insertCI SyncSameHdTl) - apply (metis insert_iff SyncSameHdTl) - apply (metis imageI subsetD) - apply (metis imageI subsetD) - by (metis (no_types, opaque_lifting) "3"(10) "3"(2) Sync.sym Un_iff image_Un subset_insertI sup.orderE SyncHd_Tl SyncSameHdTl) - with 3 show ?case - by (simp add: hd_append) (metis Sync.sym empty_setinterleaving tickFree_tl) - next - case (4 u t r v a) - have 41:\r = u\ - using 4 EmptyLeftSync Sync.sym by blast - hence 42:\hd aA \ ev ` B\ - using emptyLeftNonSync[rule_format, of r \insert tick (ev ` S)\ u \hd u\] - 4 Sync.sym by auto blast - with 4 41 show ?case - apply(rule_tac disjI2, rule_tac disjI1, simp) - by (metis Sync.sym SyncTlEmpty tickFree_tl) - next - case (5 u t r v aa a) - hence \(hd r = ev a \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev aa \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd u\ \tl u\ \insert tick (ev ` S)\ \hd t\ \tl t\ ] - apply(simp del:si_neq split:if_splits, blast, blast) - apply (metis (no_types, lifting) Sync.sym event.simps(3) imageI insert_iff SyncHd_Tl) - defer - apply blast - proof(safe, simp_all add: in_mono, goal_cases) - case (1 x) - show ?case - using "1"(13) "1"(14) "1"(18) "1"(21) "5"(8) list.collapse by fastforce + + { fix s t u r v P Q A A' B B' + assume assms : \front_tickFree v\ \tickFree r \ v = []\ \s = r @ v\ + \r setinterleaves ((t, u), insert tick (ev ` S))\ + \t \ \ (Mprefix (A \ A') P)\ \u \ \ (Mprefix (B \ B') Q)\ + \A \ S = {}\ \A' \ S\ \B \ S = {}\ \B' \ S\ + from assms(5) obtain e where * : \t \ []\ \hd t = ev e\ \tl t \ \ (P e)\ \e \ A \ e \ A'\ + by (force simp add: D_Mprefix) + have nonNil: \r \ [] \ s \ []\ + using *(1) assms(3, 4) empty_setinterleaving by blast + have \s \ \ (?rhs A A' P B B' Q)\ + proof (cases \u = []\) + case True + hence \hd s = ev e\ by (metis "*"(2) EmptyLeftSync Sync.sym assms(3, 4) hd_append nonNil) + also from "*"(1, 2, 4) Sync.sym assms(4, 8)[simplified True] have \e \ A\ + using emptyLeftNonSync hd_in_set by fastforce + ultimately show \s \ \ (?rhs A A' P B B' Q)\ + apply (simp add: D_Det) + apply (rule disjI1, simp add: D_Mprefix nonNil D_Sync) + apply (rule_tac x = \tl t\ in exI, rule_tac x = \[]\ in exI, + rule_tac x = \tl r\ in exI, rule_tac x = v in exI) + apply (auto simp add: assms(1, 3, 6)[simplified True] * nonNil)[1] + apply (metis assms(2) tickFree_tl) + using Sync.sym SyncTlEmpty True assms(4) by blast next - case (2 x) - show ?case - using "5"(14) "5"(15) "2"(18) "5"(8) "2"(21) "5"(13) list.collapse by fastforce - next - case (3 x) - show ?case - using "3"(18) "3"(21) "5"(14) "5"(15) "5"(8) Sync.sym hd_Cons_tl by fastforce - next - case (4 x) - show ?case - using "4"(18) "4"(20) "4"(21) "5"(14) "5"(15) "5"(8) list.collapse by fastforce - next - case (5 x) - show ?case - using "5"(13) "5"(14) "5"(18) "5"(21) "5"(8) Sync.sym hd_Cons_tl by fastforce - next - case (6 x) - show ?case - using "6"(20) "6"(21) "6"(13) "6"(14) "6"(18) "6"(8) "6"(12) Sync.sym hd_Cons_tl by fastforce - next - case (7 x) - show ?case - using "7"(18) "7"(13) "7"(14) "7"(8) "7"(20) "7"(21) Sync.sym hd_Cons_tl by fastforce - next - case (8 x) - show ?case - using "8"(13) "8"(14) hd_Cons_tl Sync.sym "8"(18) "8"(8) "8"(21) "8"(12) by fastforce - next - case (9 x) - show ?case - using "9"(14) "9"(13) list.collapse "9"(21) Sync.sym "9"(18) "9"(8) by fastforce - qed - with 5 show ?case - by (simp add: hd_append) (metis Sync.sym empty_setinterleaving tickFree_tl) - next - case (6 u t r v aa a) - hence \ (hd r = ev aa \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S))) - \ (hd r = ev aa \ aa = a \ hd r \ ev ` (A' \ B') \ tl r setinterleaves ((tl t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp del:si_neq split:if_splits) - apply (simp add: UnE disjoint_iff imageE insertI2 rev_image_eqI) - apply (metis (no_types, lifting) Int_iff Sync.sym Un_iff event.inject imageE image_eqI insertCI SyncSameHdTl) - apply (metis insert_iff SyncSameHdTl) - apply (metis (no_types, opaque_lifting) Sync.sym UnE image_Un insert_iff subset_eq subset_image_iff SyncHd_Tl SyncSameHdTl) - apply (metis imageI subsetD) - by (metis image_eqI subsetD) - with 6 show ?case - by (simp add: hd_append) (metis Sync.sym empty_setinterleaving tickFree_tl) - next - case (7 t u r v a) - have 71:\r = t\ - using 7 EmptyLeftSync Sync.sym by blast - hence 72:\hd aA \ ev ` A\ - using emptyLeftNonSync[rule_format, of r \insert tick (ev ` S)\ t \hd t\] - 7 Sync.sym by auto blast - with 7 71 show ?case - apply(rule_tac disjI1, simp) - by (metis Sync.sym append_self_conv front_tickFree_Nil SyncTlEmpty) - next - case (8 t u r v a aa) - hence \(hd r = ev a \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev aa \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply(simp del:si_neq split:if_splits, blast , metis empty_iff hd_Cons_tl) - apply (metis (no_types, lifting) event.simps(3) imageI insert_iff SyncHd_Tl) - defer - apply blast - proof(safe, simp_all add: in_mono, goal_cases) - case (1 x) - show ?case - using "1"(17) "1"(20) "8"(14) "8"(15) "8"(8) list.collapse by fastforce - next - case (2 x) - show ?case - using "2"(12) "2"(13) hd_Cons_tl "2"(11) "8"(8) "2"(17) "2"(20) "2"(19) by fastforce - next - case (3 x) - show ?case - using "3"(12) "3"(13) "3"(17) "3"(20) "3"(7) "3"(8) Sync.sym hd_Cons_tl by fastforce - next - case (4 x) - show ?case - using "4"(11) "4"(17) "4"(20) "8"(14) "8"(15) "8"(8) list.collapse by fastforce - next - case (5 x) - show ?case - using "5"(17) "5"(12) "5"(13) "5"(7) "5"(8) "5"(20) list.collapse Sync.sym hd_Cons_tl by fastforce - next - case (6 x) - show ?case - using "6"(11) "6"(12) "6"(13) "6"(17) "6"(19) "6"(20) "6"(7) list.collapse by fastforce - next - case (7 x) - show ?case - using "7"(17) "7"(20) "8"(13) "8"(14) "8"(15) "8"(8) list.collapse by fastforce - next - case (8 x) - show ?case - using "8"(12) "8"(13) "8"(17) "8"(19) "8"(20) "8"(7) list.collapse by fastforce - next - case (9 x) - show ?case - using "9"(20) Sync.sym "9"(12) "9"(13) list.exhaust_sel "9"(12) "9"(13) "9"(17) "9"(7) Sync.sym hd_Cons_tl by fastforce + case False + with assms(6) obtain e' where ** : \hd u = ev e'\ \tl u \ \ (Q e')\ \e' \ B \ e' \ B'\ + by (auto simp add: T_Mprefix) + consider \e \ A \ hd s = ev e \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))\ | + \e' \ B \ hd s = ev e' \ tl r setinterleaves ((t, tl u), insert tick (ev ` S))\ | + \e = e' \ e \ A' \ e \ B' \ hd s = ev e \ + tl r setinterleaves ((tl t, tl u), insert tick (ev ` S))\ + using assms(4) + apply (subst (asm) list.collapse[OF nonNil[THEN conjunct1], symmetric], + subst (asm) list.collapse[OF *(1), symmetric, simplified *(2)], + subst (asm) list.collapse[OF False, symmetric, simplified **(1)]) + apply (simp add: assms(3) nonNil image_iff split: if_split_asm) + apply (metis (no_types, opaque_lifting) *(4) **(3) Int_iff + assms(7, 9) empty_iff list.sel(1, 3)) + apply (metis (no_types, opaque_lifting) *(1, 2) **(3) Int_iff + assms(10) inf.order_iff list.exhaust_sel list.sel(1, 3)) + apply (metis (no_types, opaque_lifting) *(1, 2, 4) **(1, 3) Int_iff + False assms(8, 10) inf.order_iff list.exhaust_sel list.sel(1, 3)) + by (metis (no_types, opaque_lifting) *(4) **(1) False assms(8) + list.collapse list.sel(1, 3) subsetD) + thus \s \ \ (?rhs A A' P B B' Q)\ + apply cases + apply (simp_all add: D_Det) + apply (rule disjI1, simp add: D_Mprefix nonNil D_Sync) + apply (rule_tac x = \tl t\ in exI, rule_tac x = u in exI, + rule_tac x = \tl r\ in exI, rule_tac x = v in exI) + apply (simp add: assms(1, 3, 6) *(3) nonNil, use assms(2) nonNil tickFree_tl in blast) + apply (rule disjI2, rule disjI1, simp add: D_Mprefix nonNil D_Sync) + apply (rule_tac x = t in exI, rule_tac x = \tl u\ in exI, + rule_tac x = \tl r\ in exI, rule_tac x = v in exI) + apply (simp add: assms(1, 3, 5) **(2) nonNil, + metis "*" UnCI assms(2) image_eqI tickFree_tl) + apply (rule disjI2, rule disjI2, auto simp add: D_Mprefix nonNil image_iff) + apply (simp add: D_Sync) + apply (rule_tac x = \tl t\ in exI, rule_tac x = \tl u\ in exI, + rule_tac x = \tl r\ in exI, rule_tac x = v in exI) + by (use *(3) **(2) assms(1, 2, 3) nonNil tickFree_tl in \auto simp add: nonNil\) qed - with 8 show ?case - by (metis empty_setinterleaving self_append_conv) - next - case (9 t u r v a aa) - hence \ (hd r = ev a \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev aa \ aa = a \ hd r \ ev ` (A' \ B') \ tl r setinterleaves ((tl t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp del:si_neq split:if_splits) - apply (simp add: disjoint_iff) - apply (metis (no_types, lifting) Int_iff Un_iff event.inject imageE image_eqI insertCI SyncSameHdTl) - apply (metis insert_iff SyncSameHdTl) - apply (metis rev_image_eqI subsetD) - apply (metis imageI subsetD) - by (metis (no_types, opaque_lifting) Sync.sym Un_iff image_Un insert_iff sup.absorb_iff2 SyncHd_Tl) - with 9 show ?case - by (metis empty_setinterleaving self_append_conv) - next - case (10 u t r v a) - have 101:\r = u\ - using 10 EmptyLeftSync Sync.sym by blast - hence 102:\hd aA \ ev ` B\ - using emptyLeftNonSync[rule_format, of r \insert tick (ev ` S)\ u \hd u\] - 10 Sync.sym by auto blast - with 10 101 show ?case - apply(rule_tac disjI2, rule_tac disjI1, simp) - by (metis Sync.sym append.right_neutral front_tickFree_Nil SyncTlEmpty) - next - case (11 u t r v aa a) - hence \(hd r = ev a \ hd r \ ev ` A \ tl r setinterleaves ((tl t, u), insert tick (ev ` S))) - \ (hd r = ev aa \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd u\ \tl u\ \insert tick (ev ` S)\ \hd t\ \tl t\ ] - apply(simp del:si_neq split:if_splits, blast, blast) - apply (metis (no_types, lifting) Sync.sym event.simps(3) imageI insert_iff SyncHd_Tl) - defer - apply blast - proof(safe, simp_all add: in_mono, goal_cases) - case (1 x) - show ?case - using "1"(17) "1"(20) "11"(14) "11"(15) "11"(8) list.collapse by fastforce - next - case (2 x) - show ?case - using "2"(12) "2"(13) hd_Cons_tl "2"(11) "11"(8) "2"(17) "2"(20) by fastforce - next - case (3 x) - show ?case - using "3"(12) "3"(13) "3"(17) "3"(20) "3"(7) "3"(8) Sync.sym hd_Cons_tl by fastforce - next - case (4 x) - show ?case - using "11"(14) "11"(15) "11"(8) "4"(17) "4"(19) "4"(20) list.collapse by fastforce - next - case (5 x) - show ?case - using "5"(17) "5"(12) "5"(13) "5"(7) "5"(8) "5"(20) list.collapse Sync.sym hd_Cons_tl by fastforce - next - case (6 x) - show ?case - using "6"(11) "6"(12) "6"(13) "6"(17) "6"(19) "6"(20) "6"(7) list.collapse by fastforce - next - case (7 x) - show ?case - using "7"(12) "7"(13) "7"(17) "7"(19) "7"(20) "7"(7) Sync.sym list.collapse by fastforce - next - case (8 x) - show ?case - using "8"(20) "8"(12) "8"(11) "8"(13) list.collapse Sync.sym "8"(12) "8"(13) "8"(17) "8"(7) hd_Cons_tl by fastforce - next - case (9 x) - show ?case - using "9"(20) Sync.sym "9"(12) "9"(13) list.exhaust_sel "9"(12) "9"(13) "9"(17) "9"(7) Sync.sym hd_Cons_tl by fastforce - qed - with 11 show ?case - by (metis Sync.sym append.right_neutral empty_setinterleaving) - next - case (12 u t r v aa a) - hence \ (hd r = ev aa \ hd r \ ev ` B \ tl r setinterleaves ((t, tl u), insert tick (ev ` S))) - \ (hd r = ev aa \ aa = a \ hd r \ ev ` (A' \ B') \ tl r setinterleaves ((tl t, tl u), insert tick (ev ` S)))\ - using si_neq[of \hd t\ \tl t\ \insert tick (ev ` S)\ \hd u\ \tl u\] - apply (simp del:si_neq split:if_splits) - apply (simp add: UnE disjoint_iff imageE insertI2 rev_image_eqI) - apply (metis (no_types, lifting) Int_iff Sync.sym Un_iff event.inject imageE image_eqI insertCI SyncSameHdTl) - apply (metis insert_iff SyncSameHdTl) - apply (metis (no_types, opaque_lifting) Sync.sym UnE image_Un insert_iff subset_eq subset_image_iff SyncHd_Tl SyncSameHdTl) - apply (metis imageI subsetD) - by (metis image_eqI subsetD) - with 12 show ?case - by (metis (no_types, lifting) Sync.sym append_Nil2 empty_setinterleaving) - qed + } note * = this + + fix s + assume \s \ \ (?lhs A A' P B B' Q)\ + then obtain t u r v + where ** : \front_tickFree v\ \tickFree r \ v = []\ \s = r @ v\ + \r setinterleaves ((t, u), insert tick (ev ` S))\ + \t \ \ (Mprefix (A \ A') P) \ u \ \ (Mprefix (B \ B') Q) \ + t \ \ (Mprefix (B \ B') Q) \ u \ \ (Mprefix (A \ A') P)\ + by (simp add: D_Sync) blast + have same_div : \\ (?rhs A A' P B B' Q) = \ (?rhs B B' Q A A' P)\ + by (subst Det_commute, subst Int_commute, simp add: Sync_commute) + from **(5) show \s \ \ (?rhs A A' P B B' Q)\ + apply (rule disjE) + by (rule *[OF **(1, 2, 3, 4)]; simp add: sets_assms) + (subst same_div, rule *[OF **(1, 2, 3, 4)]; simp add: sets_assms) next - case (19 x) - then show ?case - proof(elim exE disjE conjE, goal_cases) - case (1 a t u r v) - then show ?case - apply(exI \ev a#t\, exI u, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff SyncSingleHeadAdd) - by (metis list.exhaust_sel) - next - case (2 a t u r v aa) - then show ?case - apply(exI \ev a#t\, exI u, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff SyncSingleHeadAdd) - by (metis list.exhaust_sel) - next - case (3 a t u r v aa) - then show ?case - apply(exI \ev a#t\, exI u, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff SyncSingleHeadAdd) - by (metis list.exhaust_sel) - next - case (4 a t u r v aa) - then show ?case - apply(exI \t\, exI \ev a#u\, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff in_mono SyncHdAdd1 SyncSingleHeadAdd) - apply (metis list.exhaust_sel) - apply (metis (no_types, opaque_lifting) Sync.sym event.inject event.simps(3) imageE insertE SyncSingleHeadAdd) - by (metis list.exhaust_sel) - next - case (5 a t u r v) - then show ?case - apply(exI \ev a#t\, exI u, exI \ev a#r\, exI v, simp) - apply (safe, simp_all) - apply blast - by (metis list.exhaust_sel) - next - case (6 a t u r v aa) - then show ?case - apply(exI \ev a#t\, exI u, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff_not_equal image_iff SyncSingleHeadAdd) - by (metis list.exhaust_sel) - next - case (7 a t u r v aa) - then show ?case - apply(exI \ev a#t\, exI u, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff_not_equal image_iff SyncSingleHeadAdd) - by (metis list.exhaust_sel) - next - case (8 a t u r v aa) - then show ?case - apply(exI \t\, exI \ev a#u\, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff in_mono SyncHdAdd1) - apply (metis list.exhaust_sel) - apply (metis (no_types, lifting) Sync.sym event.distinct(1) event.inject imageE insertE SyncSingleHeadAdd) - by (metis list.collapse) - next - case (9 a t u r v aa) - then show ?case - apply(exI \t\, exI \ev a#u\, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff in_mono SyncHdAdd1) - apply (metis list.exhaust_sel) - apply (metis (no_types, lifting) Sync.sym event.distinct(1) event.inject imageE insertE SyncSingleHeadAdd) - by (metis list.collapse) - next - case (10 a t u r v) - then show ?case - apply(exI \ev a#t\, exI u, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff SyncSingleHeadAdd) - by (metis list.exhaust_sel) - next - case (11 a t u r v aa) - then show ?case - apply(exI \ev a#t\, exI \u\, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff SyncSingleHeadAdd) - by (metis list.exhaust_sel) - next - case (12 a t u r v aa) - then show ?case - apply(exI \ev a#t\, exI \u\, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff SyncSingleHeadAdd) - by (metis list.exhaust_sel) - next - case (13 a t u r v aa) - then show ?case - apply(exI \t\, exI \ev a#u\, exI \ev a#r\, exI v, simp) - apply (safe, auto simp add: disjoint_iff image_iff subset_eq SyncHdAdd1) - apply (metis list.exhaust_sel) - apply (metis (no_types, lifting) Sync.sym event.inject event.simps(3) imageE insertE SyncSingleHeadAdd) - by (metis hd_Cons_tl) - next - case (14 a t u r v) - then show ?case - apply(exI \ev a#t\, exI u, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff SyncSingleHeadAdd) - by (metis list.collapse) - next - case (15 a t u r v aa) - then show ?case - apply(exI \ev a#t\, exI \u\, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff SyncSingleHeadAdd) - by (metis list.exhaust_sel) - next - case (16 a t u r v aa) - then show ?case - apply(exI \ev a#t\, exI \u\, exI \ev a#r\, exI v, simp) - apply (safe, simp_all add: disjoint_iff image_iff SyncSingleHeadAdd) - by (metis list.exhaust_sel) - next - case (17 a t u r v) - then show ?case - apply(exI \ev a#t\, exI \ev a#u\, exI \ev a#r\, exI v, simp) - apply (safe, simp_all) - apply (metis list.exhaust_sel) - by blast+ - next - case (18 a t u r v) - then show ?case - apply(exI \ev a#t\, exI \ev a#u\, exI \ev a#r\, exI v, simp) - apply (safe, simp_all) - apply (metis list.exhaust_sel) - by blast+ - next - case (19 a t u r v) - then show ?case - apply(exI \ev a#t\, exI \ev a#u\, exI \ev a#r\, exI v, simp) - apply (safe, simp_all) - apply (metis list.exhaust_sel) - by blast+ - next - case (20 a t u r v) - then show ?case - apply(exI \ev a#t\, exI \ev a#u\, exI \ev a#r\, exI v, simp) - apply (safe, simp_all) - apply (metis list.exhaust_sel) - by blast+ + + { fix A A' B B' P Q s + assume set_assm : \A \ S = {}\ + have \s \ \ (\x\A \ (P x \S\ Mprefix (B \ B') Q)) \ s \ \ (?lhs A A' P B B' Q)\ + proof (auto simp add: D_Mprefix) + fix e + assume assms: \s \ []\ \hd s = ev e\ \e \ A\ \tl s \ \ (P e \S\ Mprefix (B \ B') Q)\ + from assms(4) obtain t u r v + where * : \front_tickFree v\ \tickFree r \ v = []\ \tl s = r @ v\ + \r setinterleaves ((t, u), insert tick (ev ` S))\ + \t \ \ (P e) \ u \ \ (Mprefix (B \ B') Q) \ + t \ \ (Mprefix (B \ B') Q) \ u \ \ (P e)\ by (simp add: D_Sync) blast + have notin: \e \ S\ using assms(3) set_assm by blast + show \s \ \ (?lhs A A' P B B' Q)\ + apply (subst list.collapse[OF assms(1), symmetric], simp add: D_Sync assms(2)) + using *(5) apply (rule disjE) + + apply (rule_tac x = \ev e # t\ in exI, rule_tac x = u in exI, + rule_tac x = \ev e # r\ in exI, rule_tac x = v in exI) + apply (simp add: *(1, 2, 3, 4)) + apply (cases u; simp add: notin image_iff T_Mprefix D_Mprefix) + using *(4) assms(3) apply (fast+)[2] + + apply (rule_tac x = t in exI, rule_tac x = \ev e # u\ in exI, + rule_tac x = \ev e # r\ in exI, rule_tac x = v in exI) + apply (simp add: *(1, 2, 3, 4)) + apply (cases t; simp add: notin image_iff T_Mprefix D_Mprefix) + using *(4) assms(3) by blast + qed + } note * = this + + show \s \ \ (?rhs A A' P B B' Q) \ s \ \ (?lhs A A' P B B' Q)\ for s + apply (simp add: D_Det) + apply (erule disjE, rule *, simp_all add: sets_assms(1)) + apply (erule disjE, subst Sync_commute, rule *, simp_all add: sets_assms(3) Sync_commute) + proof - + assume \s \ \ (\x \ A' \ B' \ (P x \S\ Q x))\ + then obtain e t u r v + where * : \e \ A'\ \e \ B'\ \s \ []\ \hd s = ev e\ \front_tickFree v\ \tickFree r \ v = []\ + \tl s = r @ v\ \r setinterleaves ((t, u), insert tick (ev ` S))\ + \t \ \ (P e) \ u \ \ (Q e) \ t \ \ (Q e) \ u \ \ (P e)\ + by (simp add: D_Mprefix D_Sync) force + have inside: \e \ S\ using *(1) sets_assms(2) by blast + show \s \ \ (?lhs A A' P B B' Q)\ + apply (subst list.collapse[OF *(3), symmetric], simp add: D_Sync) + apply (rule_tac x = \ev e # t\ in exI, rule_tac x = \ev e # u\ in exI, + rule_tac x = \ev e # r\ in exI, rule_tac x = v in exI) + by (simp add: * inside D_Mprefix T_Mprefix) qed qed (* useful version for later *) lemma Mprefix_Sync_distr_bis: \(Mprefix A P) \S\ (Mprefix B Q) = - (\x\A - S \ (P x \S\ Mprefix B Q)) \ (\y\B - S \ (Mprefix A P \S\ Q y)) \ (\x\A \ B \ S \ (P x \S\ Q x))\ + (\x\A - S \ (P x \S\ Mprefix B Q)) + \ (\y\B - S \ (Mprefix A P \S\ Q y)) + \ (\x\A \ B \ S \ (P x \S\ Q x))\ by (subst Mprefix_Sync_distr[of \A - S\ S \A \ S\ \B - S\ \B \ S\, simplified Un_Diff_Int]) (simp_all add: Int_commute inf_left_commute) -lemmas Mprefix_Sync_distr_subset = Mprefix_Sync_distr[where A = \{}\ and B = \{}\, simplified, simplified Mprefix_STOP Det_STOP trans[OF Det_commute Det_STOP]] - and Mprefix_Sync_distr_indep = Mprefix_Sync_distr[where A'= \{}\ and B'= \{}\, simplified, simplified Mprefix_STOP Det_STOP] - and Mprefix_Sync_distr_right = Mprefix_Sync_distr[where A = \{}\ and B'= \{}\, simplified, simplified Mprefix_STOP Det_STOP trans[OF Det_commute Det_STOP], rotated] - and Mprefix_Sync_distr_left = Mprefix_Sync_distr[where A'= \{}\ and B = \{}\, simplified, simplified Mprefix_STOP Det_STOP trans[OF Det_commute Det_STOP]] +lemmas Mprefix_Sync_distr_subset = Mprefix_Sync_distr[where A = \{}\ + and B = \{}\, simplified, simplified Mprefix_STOP Det_STOP trans[OF Det_commute Det_STOP]] + and Mprefix_Sync_distr_indep = Mprefix_Sync_distr[where A'= \{}\ + and B'= \{}\, simplified, simplified Mprefix_STOP Det_STOP] + and Mprefix_Sync_distr_right = Mprefix_Sync_distr[where A = \{}\ + and B'= \{}\, simplified, simplified Mprefix_STOP Det_STOP trans[OF Det_commute Det_STOP], rotated] + and Mprefix_Sync_distr_left = Mprefix_Sync_distr[where A'= \{}\ + and B = \{}\, simplified, simplified Mprefix_STOP Det_STOP trans[OF Det_commute Det_STOP]] lemma Mprefix_Sync_SKIP: "(Mprefix B P \ S \ SKIP) = (\ x \ B - S \ (P x \ S \ SKIP))" proof (auto simp add:Process_eq_spec_optimized, goal_cases) case (1 x) then show ?case proof(simp_all add:D_Sync D_Mprefix D_SKIP T_SKIP, elim exE conjE, goal_cases) case (1 t u r v a) then show ?case apply(intro conjI) using empty_setinterleaving apply blast apply(elim disjE) - apply (metis (no_types, lifting) DiffI Sync.sym emptyLeftProperty empty_iff hd_append2 image_diff_subset insertI2 list.exhaust_sel setinterleaving.simps(2) subsetCE) - apply (metis D_imp_front_tickFree Sync.sym TickLeftSync empty_iff empty_setinterleaving event.distinct(1) ftf_Sync21 insertI1 list.expand list.sel(1) list.set_intros(1) SyncHd_Tl SyncSameHdTl tickFree_def) + apply (metis (no_types, lifting) DiffI Sync.sym emptyLeftProperty + empty_iff hd_append2 image_diff_subset insertI2 list.exhaust_sel + setinterleaving.simps(2) subsetCE) + apply (metis D_imp_front_tickFree Sync.sym TickLeftSync empty_iff + empty_setinterleaving event.distinct(1) ftf_Sync21 insertI1 SyncHd_Tl + list.expand list.sel(1) list.set_intros(1) SyncSameHdTl tickFree_def) using Sync.sym emptyLeftNonSync emptyLeftProperty hd_in_set apply fastforce - apply (metis (no_types, lifting) DiffI Sync.sym append_Nil2 event.distinct(1) image_diff_subset insertI1 insertI2 list.sel(1) subsetCE SyncHd_Tl SyncSameHdTl) + apply (metis (no_types, lifting) DiffI Sync.sym append_Nil2 event.distinct(1) + image_diff_subset insertI1 insertI2 list.sel(1) subsetCE SyncHd_Tl SyncSameHdTl) apply(rule_tac x="a" in exI, intro conjI) - apply (metis Sync.sym emptyLeftProperty empty_setinterleaving hd_append2 insertI1 list.sel(1) SyncHd_Tl SyncSameHdTl) - apply(rule_tac x="tl t" in exI, rule_tac x="u" in exI, rule_tac x="tl r" in exI, rule_tac x="v" in exI) - by (metis (no_types, lifting) Sync.sym empty_setinterleaving event.distinct(1) insertI1 list.sel(1) SyncHd_Tl SyncSameHdTl SyncTlEmpty tickFree_tl tl_append2) + apply (metis Sync.sym emptyLeftProperty empty_setinterleaving hd_append2 insertI1 + list.sel(1) SyncHd_Tl SyncSameHdTl) + apply(rule_tac x="tl t" in exI, rule_tac x="u" in exI, + rule_tac x="tl r" in exI, rule_tac x="v" in exI) + by (metis (no_types, lifting) Sync.sym empty_setinterleaving + event.distinct(1) insertI1 list.sel(1) SyncHd_Tl SyncSameHdTl + SyncTlEmpty tickFree_tl tl_append2) qed next case (2 x) then show ?case proof(simp_all add:D_Sync D_Mprefix D_SKIP T_SKIP, elim exE conjE, goal_cases) case (1 a t u r v) then show ?case apply(rule_tac x="hd x#t" in exI, rule_tac x="u" in exI, rule_tac x="hd x#r" in exI, rule_tac x="v" in exI, auto) using hd_Cons_tl list.exhaust_sel by force+ qed next case (3 a b) then show ?case proof(simp add:F_Sync F_Mprefix F_SKIP, elim exE disjE, simp_all, goal_cases) case (1 t u X) { fix X Y A B assume "b = (X \ Y) \ insert tick (ev ` B) \ X \ Y" and "X \ ev ` A = {}" and "tick \ Y" hence "b \ ev ` (A - B) = {}" by (rule_tac equals0I, auto) } note l1 = this from 1 show ?case apply(elim exE conjE disjE) using l1 apply simp apply(rule disjI2, simp) apply(intro disjI2 conjI) using empty_setinterleaving apply blast apply (metis (no_types, opaque_lifting) DiffI Sync.sym UnI2 contra_subsetD emptyLeftNonSync emptyLeftProperty hd_in_set image_diff_subset insert_is_Un) apply (metis Sync.sym Un_commute emptyLeftProperty list.sel(1) list.sel(3) neq_Nil_conv SyncTlEmpty) apply(intro disjI2 conjI) using empty_setinterleaving apply blast apply (metis (no_types, opaque_lifting) DiffI Sync.sym event.distinct(1) event.inject imageE image_eqI insertI1 insertI2 list.sel(1) SyncHd_Tl SyncSameHdTl) apply (metis Sync.sym event.distinct(1) insertI1 list.sel(1) SyncHd_Tl SyncSameHdTl) done next case (2 t u r v) from 2(2) have "a \ \ (Mprefix B P \S\ SKIP)" apply(simp add:D_Sync) using T_SKIP by blast with 2(3) have "a \ \ (\ x \ B - S \ (P x \ S \ SKIP))" by simp with 2(2) show ?case by (simp add:D_Sync D_Mprefix, rule_tac disjI2, metis) qed next case (4 a b) then show ?case proof(simp add: F_Sync F_Mprefix F_SKIP, elim exE disjE, simp_all, goal_cases) case 1 then show ?case apply(rule_tac disjI1, rule_tac x="[]" in exI, simp, rule_tac x="[]" in exI, simp) apply(rule_tac x="b - (ev ` B)" in exI, rule_tac conjI) by (blast, rule_tac x="b - {tick}" in exI, blast) next case 2 then show ?case proof (elim conjE exE, elim disjE exE, simp_all, goal_cases) case (1 aa t u X) then show ?case apply(erule_tac conjE, erule_tac exE, simp_all) apply(rule_tac disjI1, rule_tac x="(ev aa)#t" in exI, rule_tac x=u in exI, rule_tac x=X in exI, intro conjI, simp_all) apply auto[1] by (metis (no_types, lifting) DiffE event.distinct(1) event.inject imageE insertE list.exhaust_sel SyncSingleHeadAdd) next case (2 aa t u r v) from 2(2) have "a \ \ (\ x \ B - S \ (P x \ S \ SKIP))" apply(simp add:D_Mprefix D_Sync) by (metis "2"(1) "2"(3) "2"(4)) with 2(5) have "a \ \ (Mprefix B P \S\ SKIP)" by simp with 2(2) show ?case by (simp add:D_Sync D_Mprefix) qed qed qed lemma Sync_Ndet_left_distrib: "((P \ Q) \ S \ M) = ((P \ S \ M) \ (Q \ S \ M))" by (auto simp: Process_eq_spec, simp_all add: D_Sync F_Sync D_Ndet F_Ndet T_Ndet) blast+ lemma Sync_Ndet_right_distrib: "(M \ S \ (P \ Q)) = ((M \ S \ P) \ (M \ S \ Q))" by (metis Sync_commute Sync_Ndet_left_distrib) lemma prefix_Sync_SKIP1: "a \ S \ (a \ P \S\ SKIP) = (a \ (P \S\ SKIP))" by (metis (no_types, lifting) write0_def empty_Diff insert_Diff_if Mprefix_Sync_SKIP Mprefix_singl) lemma prefix_Sync_SKIP2: "a \ S \ (a \ P \S\ SKIP) = STOP" by (simp add: Mprefix_STOP Mprefix_Sync_SKIP write0_def) lemma prefix_Sync_SKIP: "(a \ P \S\ SKIP) = (if a \ S then STOP else (a \ (P \S\ SKIP)))" by (auto simp add: prefix_Sync_SKIP2 prefix_Sync_SKIP1) lemma prefix_Sync1: "\a \ S; b \ S; a \ b\ \ (a \ P \S\ (b \ Q)) = STOP" using Mprefix_Sync_distr_subset[of "{a}" S "{b}" "\x. P" "\x. Q", simplified] by (auto, simp add: Mprefix_STOP Mprefix_singl) lemma prefix_Sync2: "a \ S \ ((a \ P) \S\ (a \ Q)) = (a \ (P \S\ Q))" by (simp add:write0_def Mprefix_Sync_distr_subset[of "{a}" S "{a}" "\x. P" "\x. Q", simplified]) lemma prefix_Sync3: "\a \ S; b \ S\ \ (a \ P \S\ (b \ Q)) = (b \ ((a \ P) \S\ Q))" using Mprefix_Sync_distr_right[of "{b}" S "{a}" "\x. P" "\x. Q", simplified] by (auto simp add:write0_def Sync_commute) lemma Hiding_Sync_D1: \finite A \ A \ S = {} \ \ ((P \S\ Q) \ A) \ \ ((P \ A) \S\ (Q \ A))\ proof (simp only:D_Sync T_Sync D_Hiding T_Hiding, intro subsetI CollectI, elim CollectE exE conjE, elim exE CollectE disjE, goal_cases) case (1 x t u ta ua r v) then show ?case (is "\t ua ra va. ?P t ua ra va") apply (subgoal_tac "?P (trace_hide ta (ev ` A)) (trace_hide ua (ev ` A)) (trace_hide r (ev ` A)) ((trace_hide v (ev ` A))@u)") apply (metis (no_types, lifting)) apply(simp_all add:front_tickFree_append Hiding_tickFree tickFree_def disjoint_iff_not_equal Hiding_interleave[of "ev ` A" "insert tick (ev ` S)" r ta ua]) apply(elim conjE, erule_tac disjE, rule_tac disjI1, rule_tac conjI, case_tac "tickFree ta") apply(meson front_tickFree_charn self_append_conv tickFree_def) apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_Sync21 ftf_Sync32 in_set_conv_decomp_last insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def) apply(subst disj_commute, rule_tac disjCI, simp) subgoal proof(goal_cases) case 1 hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7) from 1(11) 1(10) inf_hidden[of A ua Q] obtain ff where "isInfHiddenRun ff Q A \ ua \ range ff" by auto with 1(2,3,4,7) a show ?case apply(rule_tac x=ua in exI, rule_tac x="[]" in exI) using front_tickFree_Nil tickFree_def by blast qed apply(rule_tac disjI2, rule_tac conjI, case_tac "tickFree ta") apply(meson front_tickFree_charn self_append_conv tickFree_def) apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_Sync21 ftf_Sync32 in_set_conv_decomp_last insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def) apply(subst disj_commute, rule_tac disjCI, simp) proof(goal_cases) case 1 hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7) from 1(10) 1(11) inf_hidden[of A ua P] obtain ff where "isInfHiddenRun ff P A \ ua \ range ff" by auto with 1(2,3,4,7) a show ?case apply(rule_tac x=ua in exI, rule_tac x="[]" in exI) using front_tickFree_Nil tickFree_def by blast qed next case (2 x t u f) then show ?case (is "\t ua ra va. ?P t ua ra va") proof(elim UnE exE conjE rangeE CollectE, goal_cases 21) case (21 xa) note aa = 21(7)[rule_format, of xa] with 21 show ?case proof(elim UnE exE conjE rangeE CollectE, goal_cases 211 212) case (211 taa uaa) then show ?case (is "\t ua ra va. ?P t ua ra va") proof (cases "\i. f i \ {s. \t u. t \ \ P \ u \ \ Q \ s setinterleaves ((t, u), ev ` S \ {tick})}") case True define ftu where "ftu \ \i. (SOME (t,u). t \ \ P \ u \ \ Q \ (f i) setinterleaves ((t, u), ev ` S \ {tick}))" define ft fu where "ft \ fst \ ftu" and "fu \ snd \ ftu" have "inj ftu" proof fix x y assume ftuxy: "ftu x = ftu y" obtain t u where tu:"(t, u) = ftu x" by (metis surj_pair) have "\t u. t \ \ P \ u \ \ Q \ f x setinterleaves ((t, u), ev ` S \ {tick})" using True[rule_format, of x] by simp with tu have a:"(f x) setinterleaves ((t, u), ev ` S \ {tick})" unfolding ftu_def by (metis (mono_tags, lifting) exE_some old.prod.case tu) have "\t u. t \ \ P \ u \ \ Q \ f y setinterleaves ((t, u), ev ` S \ {tick})" using True[rule_format, of y] by simp with ftuxy tu have b:"(f y) setinterleaves ((t, u), ev ` S \ {tick})" unfolding ftu_def by (metis (mono_tags, lifting) exE_some old.prod.case tu) from interleave_eq_size[OF a b] have "length (f x) = length (f y)" by assumption with 211(6) show "x = y" by (metis add.left_neutral less_length_mono linorder_neqE_nat not_add_less2 strict_mono_def) qed hence inf_ftu: "infinite (range ftu)" using finite_imageD by blast have "range ftu \ range ft \ range fu" by (clarify, metis comp_apply fst_conv ft_def fu_def rangeI snd_conv) with inf_ftu have inf_ft_fu: "infinite (range ft) \ infinite (range fu)" by (meson finite_SigmaI infinite_super) have a:"isInfHiddenRun f (P \S\ Q) A" using "2"(6) T_Sync by blast { fix i from a obtain t where "f i = f 0 @ t \ set t \ (ev ` A)" unfolding isInfHiddenRun_1 by blast hence "set (f i) \ set (f 0) \ (ev ` A)" by auto } note b = this { fix x obtain t u where tu:"(t, u) = ftu x" by (metis surj_pair) have "\t u. t \ \ P \ u \ \ Q \ f x setinterleaves ((t, u), ev ` S \ {tick})" using True[rule_format, of x] by simp with tu have a:"t \ \ P \ u \ \ Q \ (f x) setinterleaves ((t, u), ev ` S \ {tick})" unfolding ftu_def by (metis (mono_tags, lifting) exE_some old.prod.case tu) hence "ft x \ \ P \ fu x \ \ Q \ (f x) setinterleaves ((ft x, fu x), ev ` S \ {tick})" by (metis comp_apply fst_conv ft_def fu_def snd_conv tu) hence "ft x \ \ P \ fu x \ \ Q \ set (ft x) \ set (fu x) \ set (f x) \ (f x) setinterleaves ((ft x, fu x), ev ` S \ {tick})" using interleave_set by blast } note ft_fu_f = this with b have c:"ft i \ \ P \ fu i \ \ Q \ set (ft i) \ set (fu i) \ set (f 0) \ ev ` A" for i by blast with inf_ft_fu show ?thesis proof(elim disjE, goal_cases 2111 2112) case 2111 have "\t'\range ft. t = take i t' \ (set t \ set (f 0) \ ev ` A \ length t \ i)" for i by simp (metis "211"(9) b) - hence "{t. \t'\range ft. t = take i t'} \ {t. set t \ set (f 0) \ ev ` A \ length t \ i} " for i + hence " {t. \t'\range ft. t = take i t'} + \ {t. set t \ set (f 0) \ ev ` A \ length t \ i} " for i by auto (meson UnE c in_set_takeD subsetCE sup.boundedE) - with 2(1) finite_lists_length_le[of "set (f 0) \ ev ` A"] have "finite {t. \t'\range ft. t = take i t'}" for i + with 2(1) finite_lists_length_le[of "set (f 0) \ ev ` A"] + have "finite {t. \t'\range ft. t = take i t'}" for i by (meson List.finite_set finite_Un finite_imageI finite_subset) - from KoenigLemma[of "range ft", OF 2111(2), rule_format, OF this] obtain ftf::"nat \ 'a event list" where + from KoenigLemma[of "range ft", OF 2111(2), rule_format, OF this] + obtain ftf::"nat \ 'a event list" where d:"strict_mono ftf \ range ftf \ {t. \t'\range ft. t \ t'}" by blast with c d[THEN conjunct2] have e:"range ftf \ \ P" using is_processT3_ST_pref by blast define ftfs where f:"ftfs = ftf \ (\n. n + length (f 0))" from e d[THEN conjunct1] strict_mono_def have f1:"range ftfs \ \ P" and f2:"strict_mono ftfs" by (auto simp add: strict_mono_def f) { fix i have t0:"length (ftfs 0) \ length (f 0)" by (metis d[THEN conjunct1] add_leE comp_def f length_strict_mono) obtain tt where t1:"ftfs i = (ftfs 0) @ tt" by (metis f2 le0 le_list_def strict_mono_less_eq) from d[THEN conjunct2] f obtain j where t2:"ftfs i \ ft j" by simp blast obtain tt1 where t3: "ft j = (ftfs 0) @ tt @ tt1" by (metis append.assoc le_list_def t1 t2) with t0 interleave_order[of "f j" "ftfs 0" "tt@tt1" "ev ` S \ {tick}" "fu j"] ft_fu_f have t4:"set tt \ set (drop (length (f 0)) (f j))" by (metis Un_subset_iff set_append set_drop_subset_set_drop subset_Un_eq) with 21 isInfHiddenRun_1[of f _ A] have t5: "set (drop (length (f 0)) (f j)) \ ev ` A" by (metis (full_types) a append_eq_conv_conj) with t4 have "set tt \ ev ` A" by simp with t1 have "trace_hide (ftfs i) (ev ` A) = trace_hide (ftfs 0) (ev ` A)" by (simp add: subset_eq) } note f3 = this from d[THEN conjunct2] f obtain i where f4:"ftfs 0 \ ft i" by simp blast show ?case - apply(rule_tac x="trace_hide (ft i) (ev ` A)" in exI, rule_tac x="trace_hide (fu i) (ev ` A)" in exI) + apply(rule_tac x="trace_hide (ft i) (ev ` A)" in exI, + rule_tac x="trace_hide (fu i) (ev ` A)" in exI) apply(rule_tac x="trace_hide (f i) (ev ` A)" in exI, rule_tac x="u" in exI) proof (intro conjI, goal_cases 21111 21112 21113 21114 21115) case 21111 then show ?case using 2 by (simp add: "211"(3)) next case 21112 then show ?case by (metis "211"(4) "211"(9) a Hiding_tickFree) next case 21113 then show ?case by (metis "211"(5) "211"(8) "211"(9)) next case 21114 then show ?case apply(rule Hiding_interleave) using "211"(2) apply blast using ft_fu_f by simp next case 21115 from f4 obtain u where f5:"ft i = (ftfs 0) @ u" by (metis le_list_def) have "tickFree (f i)" by (metis "211"(4) "211"(8) "211"(9) Hiding_tickFree) with ft_fu_f have f6:"tickFree (ft i)" by (meson subsetCE sup.boundedE tickFree_def) show ?case apply (rule disjI1, rule conjI, simp) apply(rule_tac x="ftfs 0" in exI, rule_tac x="trace_hide u (ev ` A)" in exI, intro conjI) - apply (metis f5 front_tickFree_Nil front_tickFree_mono ft_fu_f Hiding_fronttickFree is_processT2_TR) + apply (metis f5 front_tickFree_Nil front_tickFree_mono ft_fu_f + Hiding_fronttickFree is_processT2_TR) apply (metis f5 f6 tickFree_append) apply (simp add: f5, rule disjI2, rule_tac x=ftfs in exI) using f1 f2 f3 apply blast using elemTIselemHT[of "fu i" Q A, simplified T_Hiding] ft_fu_f by blast qed next case 2112 have "\t'\range fu. t = take i t' \ (set t \ set (f 0) \ ev ` A \ length t \ i)" for i by simp (metis "211"(9) b) - hence "{t. \t'\range fu. t = take i t'} \ {t. set t \ set (f 0) \ ev ` A \ length t \ i} " for i + hence " {t. \t'\range fu. t = take i t'} + \ {t. set t \ set (f 0) \ ev ` A \ length t \ i} " for i by auto (meson UnE c in_set_takeD subsetCE sup.boundedE) - with 2(1) finite_lists_length_le[of "set (f 0) \ ev ` A"] have "finite {t. \t'\range fu. t = take i t'}" for i + with 2(1) finite_lists_length_le[of "set (f 0) \ ev ` A"] + have "finite {t. \t'\range fu. t = take i t'}" for i by (meson List.finite_set finite_Un finite_imageI finite_subset) - from KoenigLemma[of "range fu", OF 2112(2), rule_format, OF this] obtain fuf::"nat \ 'a event list" where + from KoenigLemma[of "range fu", OF 2112(2), rule_format, OF this] + obtain fuf::"nat \ 'a event list" where d:"strict_mono fuf \ range fuf \ {t. \t'\range fu. t \ t'}" by blast with c d[THEN conjunct2] have e:"range fuf \ \ Q" using is_processT3_ST_pref by blast define fufs where f:"fufs = fuf \ (\n. n + length (f 0))" from e d[THEN conjunct1] strict_mono_def have f1:"range fufs \ \ Q" and f2:"strict_mono fufs" by (auto simp add: strict_mono_def f) { fix i have t0:"length (fufs 0) \ length (f 0)" by (metis d[THEN conjunct1] add_leE comp_def f length_strict_mono) obtain tt where t1:"fufs i = (fufs 0) @ tt" by (metis f2 le0 le_list_def strict_mono_less_eq) from d[THEN conjunct2] f obtain j where t2:"fufs i \ fu j" by simp blast obtain tt1 where t3: "fu j = (fufs 0) @ tt @ tt1" by (metis append.assoc le_list_def t1 t2) with t0 interleave_order[of "f j" "fufs 0" "tt@tt1" "ev ` S \ {tick}" "ft j"] ft_fu_f have t4:"set tt \ set (drop (length (f 0)) (f j))" by (metis (no_types, opaque_lifting) Sync.sym Un_subset_iff order_trans set_append set_drop_subset_set_drop) with 21 isInfHiddenRun_1[of f _ A] have t5: "set (drop (length (f 0)) (f j)) \ ev ` A" by (metis (full_types) a append_eq_conv_conj) with t4 have "set tt \ ev ` A" by simp with t1 have "trace_hide (fufs i) (ev ` A) = trace_hide (fufs 0) (ev ` A)" by (simp add: subset_eq) } note f3 = this from d[THEN conjunct2] f obtain i where f4:"fufs 0 \ fu i" by simp blast show ?case apply(rule_tac x="trace_hide (fu i) (ev ` A)" in exI, rule_tac x="trace_hide (ft i) (ev ` A)" in exI) apply(rule_tac x="trace_hide (f i) (ev ` A)" in exI, rule_tac x="u" in exI) proof (intro conjI, goal_cases 21111 21112 21113 21114 21115) case 21111 then show ?case using 2 by (simp add: "211"(3)) next case 21112 then show ?case by (metis "211"(4) "211"(9) a Hiding_tickFree) next case 21113 then show ?case by (metis "211"(5) "211"(8) "211"(9)) next case 21114 then show ?case apply(rule Hiding_interleave) using "211"(2) apply blast using ft_fu_f using Sync.sym by blast next case 21115 from f4 obtain u where f5:"fu i = (fufs 0) @ u" by (metis le_list_def) have "tickFree (f i)" by (metis "211"(4) "211"(8) "211"(9) Hiding_tickFree) with ft_fu_f have f6:"tickFree (fu i)" by (meson subsetCE sup.boundedE tickFree_def) show ?case apply (rule disjI2, rule conjI, simp) apply(rule_tac x="fufs 0" in exI, rule_tac x="trace_hide u (ev ` A)" in exI, intro conjI) - apply (metis f5 front_tickFree_Nil front_tickFree_mono ft_fu_f Hiding_fronttickFree is_processT2_TR) + apply (metis f5 front_tickFree_Nil front_tickFree_mono ft_fu_f + Hiding_fronttickFree is_processT2_TR) apply (metis f5 f6 tickFree_append) apply (simp add: f5, rule disjI2, rule_tac x=fufs in exI) using f1 f2 f3 apply blast using elemTIselemHT[of "ft i" P A, simplified T_Hiding] ft_fu_f by blast qed qed next case False then obtain xaa where "f xaa \ {s. \t u. t \ \ P \ u \ \ Q \ s setinterleaves ((t, u), ev ` S \ {tick})}" by blast with 211(7)[rule_format, of xaa] obtain ta ua ra va where bb:"front_tickFree va \ (tickFree ra \ va = []) \ f xaa = ra @ va \ ra setinterleaves ((ta, ua), ev ` S \ {tick}) \ (ta \ \ P \ ua \ \ Q \ ta \ \ Q \ ua \ \ P)" by blast from 211 have cc:"x = trace_hide (f xaa) (ev ` A) @ u" by (metis (no_types, lifting)) from bb 211(9) "211"(8) "21"(4) have "tickFree ra \ tickFree va" by (metis Hiding_tickFree tickFree_append) with cc bb 211 show ?thesis apply(subgoal_tac "?P (trace_hide ta (ev ` A)) (trace_hide ua (ev ` A)) (trace_hide ra (ev ` A)) ((trace_hide va (ev ` A))@u)") apply (metis (no_types, lifting)) apply(simp_all add:front_tickFree_append Hiding_tickFree tickFree_def disjoint_iff_not_equal Hiding_interleave[of "ev ` A" "insert tick (ev ` S)" ra ta ua]) apply(elim conjE disjE, rule_tac disjI1, rule_tac conjI, case_tac "tickFree ta") apply(meson front_tickFree_charn self_append_conv tickFree_def) apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_Sync21 ftf_Sync32 in_set_conv_decomp_last insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def) apply(subst disj_commute, rule_tac disjCI, simp) subgoal proof(goal_cases 2111) case 2111 hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7) from 2111(20) 2111(21) inf_hidden[of A ua Q] obtain ff where "isInfHiddenRun ff Q A \ ua \ range ff" by auto with 2111(2,3,4,7) a show ?case apply(rule_tac x=ua in exI, rule_tac x="[]" in exI) using front_tickFree_Nil tickFree_def by blast qed apply(rule_tac disjI2, rule_tac conjI, case_tac "tickFree ta") apply(meson front_tickFree_charn self_append_conv tickFree_def) apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_Sync21 ftf_Sync32 in_set_conv_decomp_last insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def) apply(subst disj_commute, rule_tac disjCI, simp) proof(goal_cases 2111) case 2111 hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7) from 2111(20) 2111(21) inf_hidden[of A ua P] obtain ff where "isInfHiddenRun ff P A \ ua \ range ff" by auto with 2111(2,3,4,7) a show ?case apply(rule_tac x=ua in exI, rule_tac x="[]" in exI) using front_tickFree_Nil tickFree_def by blast qed qed next case (212 ta ua r v) then show ?case (is "\t ua ra va. ?P t ua ra va") apply (subgoal_tac "?P (trace_hide ta (ev ` A)) (trace_hide ua (ev ` A)) (trace_hide r (ev ` A)) ((trace_hide v (ev ` A))@u)") apply (metis (no_types, lifting)) apply(simp_all add:front_tickFree_append Hiding_tickFree tickFree_def disjoint_iff_not_equal Hiding_interleave[of "ev ` A" "insert tick (ev ` S)" r ta ua]) apply(erule_tac disjE, rule_tac disjI1, rule_tac conjI, case_tac "tickFree ta") apply(meson front_tickFree_charn self_append_conv tickFree_def) apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_Sync21 ftf_Sync32 in_set_conv_decomp_last insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def) apply(subst disj_commute, rule_tac disjCI, simp) subgoal proof(goal_cases 2121) case 2121 hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7) from 2121(14) 2121(13) inf_hidden[of A ua Q] obtain ff where "isInfHiddenRun ff Q A \ ua \ range ff" by auto with 2121(2,3,4,7) a show ?case apply(rule_tac x=ua in exI, rule_tac x="[]" in exI) using front_tickFree_Nil tickFree_def by blast qed apply(rule_tac disjI2, rule_tac conjI, case_tac "tickFree ta") apply(meson front_tickFree_charn self_append_conv tickFree_def) apply (metis D_imp_front_tickFree emptyLeftNonSync ftf_Sync21 ftf_Sync32 in_set_conv_decomp_last insertI1 is_processT2_TR nonTickFree_n_frontTickFree tickFree_Nil tickFree_def) apply(subst disj_commute, rule_tac disjCI, simp) proof(goal_cases 2121) case 2121 hence a:"tickFree ua" by (metis front_tickFree_implies_tickFree is_processT2_TR is_processT5_S7) from 2121(14) 2121(13) inf_hidden[of A ua P] obtain ff where "isInfHiddenRun ff P A \ ua \ range ff" by auto with 2121(2,3,4,7) a show ?case apply(rule_tac x=ua in exI, rule_tac x="[]" in exI) using front_tickFree_Nil tickFree_def by blast qed qed qed qed lemma Hiding_Sync_D2: assumes *:"A \ S = {}" shows \\ ((P \ A) \S\ (Q \ A)) \ \ ((P \S\ Q) \ A)\ proof - { fix P Q - have "{s. \t u r v. front_tickFree v \ (tickFree r \ v = []) \ + have \{s. \t u r v. front_tickFree v \ (tickFree r \ v = []) \ s = r @ v \ r setinterleaves ((t, u), insert tick (ev ` S)) \ - (t \ \ (P \\ A) \ u \ \ (Q \\ A))} \ \ ((P \S\ Q) \\ A)" + (t \ \ (P \ A) \ u \ \ (Q \ A))} \ \ ((P \S\ Q) \ A)\ proof(simp add:D_Hiding D_Sync, intro subsetI CollectI, elim exE conjE CollectE, goal_cases a) case (a x t u r v ta1 ta2) from a(1,3-9) show ?case proof(erule_tac disjE, goal_cases) case 1 with interleave_append[of r "trace_hide ta1 (ev ` A)" "ta2" "insert tick (ev ` S)" u] obtain u1 u2 r1 r2 where a1:"u = u1 @ u2" and a2:"r = r1 @ r2" and a3:"r1 setinterleaves ((trace_hide ta1 (ev ` A), u1), insert tick (ev ` S))" and a4:"r2 setinterleaves ((ta2, u2), insert tick (ev ` S))" by blast with 1 show ?case proof(auto simp only:T_Hiding, goal_cases 11 12 13) case (11 ua) from trace_hide_append[OF 11(12)] obtain ua1 where a5:"u1 = trace_hide ua1 (ev ` A) \ ua1 \ ua \ ua1 \ \ Q" using "11"(13) F_T is_processT3_ST_pref le_list_def by blast from interleave_Hiding[OF _ 11(10)[simplified a5]] obtain ra1 where a6:"r1 = trace_hide ra1 (ev ` A) \ ra1 setinterleaves ((ta1, ua1), insert tick (ev ` S))" using * by blast from a(2) 11 a5 a6 show ?case apply(exI ra1, exI "r2@v", intro conjI, simp_all (asm_lr)) - apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_Sync is_processT2_TR) + apply (metis a(5) append_Nil2 front_tickFree_append + front_tickFree_mono ftf_Sync is_processT2_TR) apply (metis equals0D ftf_Sync1 ftf_Sync21 insertI1 tickFree_def) using front_tickFree_Nil by blast next case (12 ua1 ua2) then show ?case proof(cases "u1 \ trace_hide ua1 (ev ` A)") case True then obtain uu1 where "u1@uu1 = trace_hide ua1 (ev ` A)" using le_list_def by blast with trace_hide_append[OF this] obtain uaa1 where a5:"u1 = trace_hide uaa1 (ev ` A) \ uaa1 \ ua1 \ uaa1 \ \ Q" using "12"(15) NT_ND is_processT3_ST_pref le_list_def by blast from interleave_Hiding[OF _ 12(10)[simplified a5]] obtain rr1 where a6:"r1 = trace_hide rr1 (ev ` A) \ rr1 setinterleaves ((ta1, uaa1), insert tick (ev ` S))" using * by blast from a(2) 12 a5 a6 show ?thesis apply(exI rr1, exI "r2@v", intro conjI, simp_all (asm_lr)) apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_Sync is_processT2_TR) apply (metis equals0D ftf_Sync1 ftf_Sync21 insertI1 tickFree_def) using front_tickFree_Nil by blast next case False with 12(14) obtain uaa1 where "u1 = trace_hide ua1 (ev ` A)@uaa1" by (metis append_eq_append_conv2 le_list_def) with a3 interleave_append_sym[of r1 "trace_hide ta1 (ev ` A)" "insert tick (ev ` S)" "trace_hide ua1 (ev ` A)" uaa1] obtain taa1 taa2 ra1 ra2 where aa1:"trace_hide ta1 (ev ` A) = taa1 @ taa2" and aa2:"r1 = ra1 @ ra2" and aa3:"ra1 setinterleaves ((taa1, trace_hide ua1 (ev ` A)), insert tick (ev ` S))" and aa4:"ra2 setinterleaves ((taa2,uaa1), insert tick (ev ` S))" by blast with trace_hide_append[OF aa1[symmetric]] obtain taaa1 where a5:"taa1 = trace_hide taaa1 (ev ` A) \ taaa1 \ ta1 \ taaa1 \ \ P" using "12"(7) D_T is_processT3_ST_pref le_list_def by blast with interleave_Hiding[OF _ aa3[simplified a5]] obtain rr1 where a6:"ra1 = trace_hide rr1 (ev ` A) \ rr1 setinterleaves ((taaa1, ua1), insert tick (ev ` S))" using * by blast from a(2) 12 a5 show ?thesis apply(exI rr1, exI "ra2@r2@v", intro conjI, simp_all (asm_lr)) apply (metis aa2 front_tickFree_append front_tickFree_mono ftf_Sync Hiding_tickFree self_append_conv tickFree_append) apply (metis a6 ftf_Sync1 le_list_def tickFree_append tickFree_def) using a6 aa2 apply blast using Sync.sym a6 front_tickFree_Nil by blast qed next case (13 ua fa xa) with isInfHiddenRun_1[of fa Q] have a0:"\i. \t. fa i = fa 0 @ t \ set t \ (ev ` A)" by simp define tlf where "tlf \ \i. drop (length (fa 0)) (fa i)" have tlf_def2:"(fa x) = (fa 0) @ (tlf x)" for x by (metis a0 append_eq_conv_conj tlf_def) { fix x y assume *:"(x::nat) < y" hence "fa x = fa 0 @ tlf x \ fa y = fa 0 @ tlf y" by (metis tlf_def2) hence "tlf x < tlf y" using strict_monoD[OF 13(16), of x y] by (simp add: A "*" le_list_def) } note tlf_strict_mono = this have tlf_range:"set (tlf i) \ (ev ` A)" for i by (metis a0 append_eq_conv_conj tlf_def) from 13 show ?case proof(cases "u1 \ trace_hide (fa xa) (ev ` A)") case True then obtain uu1 where "u1@uu1 = trace_hide (fa xa) (ev ` A)" using le_list_def by blast with trace_hide_append[OF this] obtain uaa1 where a5:"u1 = trace_hide uaa1 (ev ` A) \ uaa1 \ (fa xa) \ uaa1 \ \ Q" by (metis "13"(17) is_processT3_ST le_list_def) from interleave_Hiding[OF _ 13(10)[simplified a5]] obtain rr1 - where a6:"r1 = trace_hide rr1 (ev ` A) \ rr1 setinterleaves ((ta1, uaa1), insert tick (ev ` S))" + where a6:" r1 = trace_hide rr1 (ev ` A) + \ rr1 setinterleaves ((ta1, uaa1), insert tick (ev ` S))" using * by blast from a(2) 13 a5 a6 show ?thesis apply(exI rr1, exI "r2@v", intro conjI, simp_all (asm_lr)) apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_Sync is_processT2_TR) apply (metis equals0D ftf_Sync1 ftf_Sync21 insertI1 tickFree_def) using front_tickFree_Nil by blast next case False with 13(14) obtain uaa1 where "u1 = trace_hide (fa 0) (ev ` A)@uaa1" by (metis "13"(18) append_eq_append_conv2 le_list_def) with a3 interleave_append_sym[of r1 "trace_hide ta1 (ev ` A)" "insert tick (ev ` S)" "trace_hide (fa 0) (ev ` A)" uaa1] obtain taa1 taa2 ra1 ra2 where aa1:"trace_hide ta1 (ev ` A) = taa1 @ taa2" and aa2:"r1 = ra1 @ ra2" and aa3:"ra1 setinterleaves ((taa1, trace_hide (fa 0) (ev ` A)), insert tick (ev ` S))" and aa4:"ra2 setinterleaves ((taa2,uaa1), insert tick (ev ` S))" by blast with trace_hide_append[OF aa1[symmetric]] obtain taaa1 where a5:"taa1 = trace_hide taaa1 (ev ` A) \ taaa1 \ ta1 \ taaa1 \ \ P" using "13"(7) D_T is_processT3_ST_pref le_list_def by blast with interleave_Hiding[OF _ aa3[simplified a5]] obtain rr1 where a6:"ra1 = trace_hide rr1 (ev ` A) \ rr1 setinterleaves ((taaa1, (fa 0)), insert tick (ev ` S))" using * by blast from a(2) 13 a0 a5 a6 aa1 aa2 aa3 show ?thesis apply(exI rr1, exI "ra2@r2@v", intro conjI, simp_all (asm_lr)) apply (metis front_tickFree_append front_tickFree_mono ftf_Sync Hiding_tickFree self_append_conv tickFree_append) apply (metis a6 ftf_Sync1 le_list_def tickFree_append tickFree_def) proof (rule_tac disjI2, exI "\i. rr1@(tlf i)", intro conjI, goal_cases 131 132 133 134) case 131 then show ?case apply(rule_tac strict_monoI, drule_tac tlf_strict_mono,rule_tac less_append) by assumption next case 132 have "(rr1@tlf i) setinterleaves ((taaa1, fa i),insert tick (ev ` S))" for i apply(subst tlf_def2, rule interleave_append_tail_sym[OF a6[THEN conjunct2], of "tlf i"]) using * tlf_range by blast then show ?case unfolding T_Sync using a5 13(17) by auto next case 133 from tlf_range have "trace_hide (y @ tlf i) (ev ` A) = trace_hide y (ev ` A)" for i y apply(induct y) using filter_empty_conv apply fastforce by simp then show ?case by simp next case 134 have "tlf 0 = []" by (simp add: tlf_def) then show ?case by simp qed qed qed next case 2 from 2(8) obtain nta::nat and ff::"nat \ 'a event list" where "strict_mono ff" and ff:"(\i. ff i \ \ P) \ ta1 = ff nta \ (\i. trace_hide (ff i) (ev ` A) = trace_hide (ff 0) (ev ` A))" by blast define f where "f \ (\i. ff (i + nta))" with ff have f1:"strict_mono f" and f2:"(\i. f i \ \ P)" and f3:"ta1 = f 0" and f4:"(\i. trace_hide (f i) (ev ` A) = trace_hide ta1 (ev ` A))" apply auto apply(rule strict_monoI, rule \strict_mono ff\[THEN strict_monoD]) by simp metis with isInfHiddenRun_1[of f P] have a0:"\i. \t. f i = f 0 @ t \ set t \ (ev ` A)" by simp define tlf where "tlf \ \i. drop (length (f 0)) (f i)" have tlf_def2:"(f x) = ta1 @ (tlf x)" for x by (metis a0 f3 append_eq_conv_conj tlf_def) { fix x y assume *:"(x::nat) < y" hence "f x = f 0 @ tlf x \ f y = f 0 @ tlf y" by (metis f3 tlf_def2) hence "tlf x < tlf y" using strict_monoD[OF f1, of x y] by (simp add: A "*" le_list_def) } note tlf_strict_mono = this have tlf_range:"set (tlf i) \ (ev ` A)" for i by (metis a0 append_eq_conv_conj tlf_def) with 2 interleave_append[of r "trace_hide ta1 (ev ` A)" "ta2" "insert tick (ev ` S)" u] obtain u1 u2 r1 r2 where a1:"u = u1 @ u2" and a2:"r = r1 @ r2" and a3:"r1 setinterleaves ((trace_hide ta1 (ev ` A), u1), insert tick (ev ` S))" and a4:"r2 setinterleaves ((ta2, u2), insert tick (ev ` S))" by blast with 2(1-6) f1 f2 f3 f4 show ?case proof(auto simp only:T_Hiding, goal_cases 21 22 23) case (21 ua) from trace_hide_append[OF 21(14)] obtain ua1 where a5:"u1 = trace_hide ua1 (ev ` A) \ ua1 \ ua \ ua1 \ \ Q" using "21"(15) F_T is_processT3_ST_pref le_list_def by blast from interleave_Hiding[OF _ a3[simplified a5]] obtain ra1 where a6:"r1 = trace_hide ra1 (ev ` A) \ ra1 setinterleaves ((ta1, ua1), insert tick (ev ` S))" using * by blast from a(2) 21 a5 a6 show ?case apply(exI ra1, exI "r2@v", intro conjI, simp_all (asm_lr)) apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_Sync is_processT2_TR) apply (metis equals0D ftf_Sync1 ftf_Sync21 insertI1 tickFree_def) proof (rule_tac disjI2, exI "\i. ra1@(tlf i)", intro conjI, goal_cases 211 212 213 214) case 211 then show ?case by(rule_tac strict_monoI, drule_tac tlf_strict_mono, rule_tac less_append)(assumption) next case 212 have "(ra1@tlf i) setinterleaves ((f i, ua1),insert tick (ev ` S))" for i apply(subst tlf_def2, rule interleave_append_tail[OF a6[THEN conjunct2], of "tlf i"]) using * tlf_range by blast then show ?case unfolding T_Sync using a5 21(15) f2 by auto next case 213 from tlf_range have "trace_hide (y @ tlf i) (ev ` A) = trace_hide y (ev ` A)" for i y apply(induct y) using filter_empty_conv apply fastforce by simp then show ?case by simp next case 214 have "tlf 0 = []" by (simp add: tlf_def) then show ?case by simp qed next case (22 ua1 ua2) then show ?case proof(cases "u1 \ trace_hide ua1 (ev ` A)") case True then obtain uu1 where "u1@uu1 = trace_hide ua1 (ev ` A)" using le_list_def by blast from trace_hide_append[OF this] obtain uaa1 where a5:"u1 = trace_hide uaa1 (ev ` A) \ uaa1 \ ua1 \ uaa1 \ \ Q" using "22"(17) D_T is_processT3_ST le_list_def by blast from interleave_Hiding[OF _ a3[simplified a5]] obtain ra1 - where a6:"r1 = trace_hide ra1 (ev ` A) \ ra1 setinterleaves ((ta1, uaa1), insert tick (ev ` S))" + where a6:" r1 = trace_hide ra1 (ev ` A) + \ ra1 setinterleaves ((ta1, uaa1), insert tick (ev ` S))" using * by blast from a(2) 22 a5 a6 show ?thesis apply(exI ra1, exI "r2@v", intro conjI, simp_all (asm_lr)) - apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_Sync is_processT2_TR) + apply (metis a(5) append_Nil2 front_tickFree_append + front_tickFree_mono ftf_Sync is_processT2_TR) apply (metis equals0D ftf_Sync1 ftf_Sync21 insertI1 tickFree_def) proof (rule_tac disjI2, exI "\i. ra1@(tlf i)", intro conjI, goal_cases 211 212 213 214) case 211 then show ?case apply(rule_tac strict_monoI, drule_tac tlf_strict_mono, rule_tac less_append) by assumption next case 212 have "(ra1@tlf i) setinterleaves ((f i, uaa1),insert tick (ev ` S))" for i apply(subst tlf_def2, rule interleave_append_tail[OF a6[THEN conjunct2], of "tlf i"]) using * tlf_range by blast then show ?case unfolding T_Sync using a5 22(15) f2 by auto next case 213 from tlf_range have "trace_hide (y @ tlf i) (ev ` A) = trace_hide y (ev ` A)" for i y apply(induct y) using filter_empty_conv apply fastforce by simp then show ?case by simp next case 214 have "tlf 0 = []" by (simp add: tlf_def) then show ?case by simp qed next case False with 22(16) obtain uaa1 where "u1 = trace_hide ua1 (ev ` A)@uaa1" by (metis append_eq_append_conv2 le_list_def) with a3 interleave_append_sym[of r1 "trace_hide ta1 (ev ` A)" "insert tick (ev ` S)" " trace_hide ua1 (ev ` A)" uaa1] obtain taa1 taa2 ra1 ra2 where aa1:"trace_hide ta1 (ev ` A) = taa1 @ taa2" and aa2:"r1 = ra1 @ ra2" and aa3:"ra1 setinterleaves ((taa1, trace_hide ua1 (ev ` A)), insert tick (ev ` S))" and aa4:"ra2 setinterleaves ((taa2,uaa1), insert tick (ev ` S))" by blast with trace_hide_append[OF aa1[symmetric]] obtain taaa1 where a5:"taa1 = trace_hide taaa1 (ev ` A) \ taaa1 \ ta1 \ taaa1 \ \ P" using f2 f3 is_processT3_ST_pref le_list_def by metis with interleave_Hiding[OF _ aa3[simplified a5]] obtain rr1 where a6:"ra1 = trace_hide rr1 (ev ` A) \ rr1 setinterleaves ((taaa1, ua1), insert tick (ev ` S))" using * by blast from a(2) 22 a5 show ?thesis apply(exI rr1, exI "ra2@r2@v", intro conjI, simp_all (asm_lr)) - apply (metis a(5) a(8) aa2 front_tickFree_append front_tickFree_mono ftf_Sync Hiding_tickFree - is_processT2_TR self_append_conv tickFree_append) + apply (metis a(5) a(8) aa2 front_tickFree_append front_tickFree_mono is_processT2_TR + ftf_Sync Hiding_tickFree self_append_conv tickFree_append) apply (metis a6 ftf_Sync1 le_list_def tickFree_append tickFree_def) using a6 aa2 apply blast using Sync.sym a6[THEN conjunct2] front_tickFree_Nil by (metis append_Nil2) qed next case (23 ua fa xa) with isInfHiddenRun_1[of fa Q] have a0:"\i. \t. fa i = fa 0 @ t \ set t \ (ev ` A)" by simp define tlfa where "tlfa \ \i. drop (length (fa 0)) (fa i)" have tlfa_def2:"(fa x) = (fa 0) @ (tlfa x)" for x by (metis a0 append_eq_conv_conj tlfa_def) { fix x y assume *:"(x::nat) < y" hence "fa x = fa 0 @ tlfa x \ fa y = fa 0 @ tlfa y" by (metis tlfa_def2) hence "tlfa x < tlfa y" using strict_monoD[OF 23(18), of x y] by (simp add: A "*" le_list_def) } note tlfa_strict_mono = this have tlfa_range:"set (tlfa i) \ (ev ` A)" for i by (metis a0 append_eq_conv_conj tlfa_def) from 23 show ?case proof(cases "u1 \ trace_hide (fa xa) (ev ` A)") case True then obtain uu1 where "u1@uu1 = trace_hide (fa xa) (ev ` A)" using le_list_def by blast with trace_hide_append[OF this] obtain uaa1 where a5:"u1 = trace_hide uaa1 (ev ` A) \ uaa1 \ (fa xa) \ uaa1 \ \ Q" by (metis "23"(19) is_processT3_ST le_list_def) from interleave_Hiding[OF _ a3[simplified a5]] obtain rr1 where a6:"r1 = trace_hide rr1 (ev ` A) \ rr1 setinterleaves ((ta1, uaa1), insert tick (ev ` S))" using * by blast from a(2) 23 a5 a6 show ?thesis apply(exI rr1, exI "r2@v", intro conjI, simp_all (asm_lr)) - apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono ftf_Sync is_processT2_TR) - apply (metis equals0D ftf_Sync1 ftf_Sync21 insertI1 tickFree_def) + apply (metis a(5) append_Nil2 front_tickFree_append front_tickFree_mono + ftf_Sync is_processT2_TR) + apply (metis equals0D ftf_Sync1 ftf_Sync21 insertI1 tickFree_def) proof (rule_tac disjI2, exI "\i. rr1@(tlf i)", intro conjI, goal_cases 211 212 213 214) case 211 - then show ?case apply(rule_tac strict_monoI, drule_tac tlf_strict_mono, rule_tac less_append) by assumption + then show ?case + apply(rule_tac strict_monoI, drule_tac tlf_strict_mono, rule_tac less_append) by assumption next case 212 have "(rr1@tlf i) setinterleaves ((f i, uaa1),insert tick (ev ` S))" for i apply(subst tlf_def2, rule interleave_append_tail[OF a6[THEN conjunct2], of "tlf i"]) using * tlf_range by blast then show ?case unfolding T_Sync using a5 23(17) f2 by auto next case 213 from tlf_range have "trace_hide (y @ tlf i) (ev ` A) = trace_hide y (ev ` A)" for i y apply(induct y) using filter_empty_conv apply fastforce by simp then show ?case by simp next case 214 have "tlf 0 = []" by (simp add: tlf_def) then show ?case by simp qed next case False with 23(16) obtain uaa1 where "u1 = trace_hide (fa 0) (ev ` A)@uaa1" by (metis "23"(20) append_eq_append_conv2 le_list_def) - with a3 interleave_append_sym[of r1 "trace_hide ta1 (ev ` A)" "insert tick (ev ` S)" "trace_hide (fa 0) (ev ` A)" uaa1] + with a3 interleave_append_sym[of r1 "trace_hide ta1 (ev ` A)" + "insert tick (ev ` S)" + "trace_hide (fa 0) (ev ` A)" uaa1] obtain taa1 taa2 ra1 ra2 where aa1:"trace_hide ta1 (ev ` A) = taa1 @ taa2" and aa2:"r1 = ra1 @ ra2" and aa3:"ra1 setinterleaves ((taa1, trace_hide (fa 0) (ev ` A)), insert tick (ev ` S))" and aa4:"ra2 setinterleaves ((taa2,uaa1), insert tick (ev ` S))" by blast with trace_hide_append[OF aa1[symmetric]] obtain taaa1 where a5:"taa1 = trace_hide taaa1 (ev ` A) \ taaa1 \ ta1 \ taaa1 \ \ P" by (metis f2 f3 is_processT3_ST le_list_def) with interleave_Hiding[OF _ aa3[simplified a5]] obtain rr1 where a6:"ra1 = trace_hide rr1 (ev ` A) \ rr1 setinterleaves ((taaa1, (fa 0)), insert tick (ev ` S))" using * by blast from a(2) 23 a0 a5 a6 aa1 aa2 aa3 show ?thesis apply(exI rr1, exI "ra2@r2@v", intro conjI, simp_all (asm_lr)) apply (metis a(5) a(8) front_tickFree_append front_tickFree_mono ftf_Sync Hiding_tickFree is_processT2_TR self_append_conv) apply (metis a6 ftf_Sync1 le_list_def tickFree_append tickFree_def) proof (rule_tac disjI2, exI "\i. rr1@(tlfa i)", intro conjI, goal_cases 131 132 133 134) case 131 then show ?case by(rule_tac strict_monoI, drule_tac tlfa_strict_mono, rule_tac less_append) assumption next case 132 have "(rr1@tlfa i) setinterleaves ((taaa1, fa i),insert tick (ev ` S))" for i apply(subst tlfa_def2, rule interleave_append_tail_sym[OF a6[THEN conjunct2], of "tlfa i"]) using * tlfa_range by blast then show ?case unfolding T_Sync using a5 23(19) by auto next case 133 from tlfa_range have "trace_hide (y @ tlfa i) (ev ` A) = trace_hide y (ev ` A)" for i y apply(induct y) using filter_empty_conv apply fastforce by simp then show ?case by simp next case 134 have "tlfa 0 = []" by (simp add: tlfa_def) then show ?case by simp qed qed qed qed qed } from this[of P Q] this[of Q P] Sync_commute[of P S Q] show ?thesis by (simp add:D_Sync T_Sync) blast qed lemma Hiding_Sync: \finite A \ A \ S = {} \ ((P \S\ Q) \ A) = ((P \ A) \S\ (Q \ A))\ proof(auto simp add: Process_eq_spec_optimized subset_antisym[OF Hiding_Sync_D1 Hiding_Sync_D2], simp_all add: F_Sync F_Hiding, goal_cases) case (1 a b) then show ?case proof(elim disjE, goal_cases 11 12) case 11 then show ?case proof(elim exE conjE, elim disjE, goal_cases 111 112) case (111 c) from 111(7) obtain t u X Y where a: "(t, X) \ \ P" and b: "(u, Y) \ \ Q" and c: "c setinterleaves ((t, u), insert tick (ev ` S))" and d: "b \ ev ` A = (X \ Y) \ insert tick (ev ` S) \ X \ Y" by auto from 1(2) d have e:"ev ` A \ X \ Y" by blast from a b c d e 111(2,3) show ?case apply(rule_tac disjI1) apply(rule_tac x="trace_hide t (ev ` A)" in exI, rule_tac x="trace_hide u (ev ` A)" in exI) apply(rule_tac x="X - ((ev ` A) - b)" in exI, intro conjI disjI1, rule_tac x="t" in exI) apply (metis Int_subset_iff Un_Diff_Int Un_least is_processT4 sup_ge1) apply(rule_tac x="Y - ((ev ` A) - b)" in exI, intro conjI disjI1, rule_tac x="u" in exI) apply (metis Int_subset_iff Un_Diff_Int Un_least is_processT4 sup_ge1) using Hiding_interleave[of "ev ` A" "insert tick (ev ` S)" c t u] by blast+ next case (112 t) from 112(7) have a:"front_tickFree t" by (metis (mono_tags, opaque_lifting) D_imp_front_tickFree append_Nil2 front_tickFree_append ftf_Sync is_processT2_TR) from 112 have "t \ \ (P \S\ Q)" by (simp add:D_Sync) - with 112(1,2,3) have "a \ \ ((P \S\ Q) \\ A)" + with 112(1,2,3) have \a \ \ ((P \S\ Q) \ A)\ apply (simp add:D_Hiding) proof(case_tac "tickFree t", goal_cases 1121 1122) case 1121 then show ?case by (rule_tac x=t in exI, rule_tac x="[]" in exI, simp) next case 1122 with a obtain t' where "t = t'@[tick]" using nonTickFree_n_frontTickFree by blast with a show ?case apply (rule_tac x="t'" in exI, rule_tac x="[tick]" in exI, auto) using front_tickFree_mono apply blast using 1122(4) is_processT9 by blast qed then show ?case apply(rule_tac disjI2) - using Hiding_Sync_D1[of A S P Q, OF 112(1,2)] D_Sync[of "P \\ A" S "Q \\ A"] by auto + using Hiding_Sync_D1[of A S P Q, OF 112(1,2)] D_Sync[of \P \ A\ S \Q \ A\] by auto qed next case 12 - hence "a \ \ ((P \S\ Q) \\ A)" by (simp add:D_Hiding) + hence \a \ \ ((P \S\ Q) \ A)\ by (simp add:D_Hiding) then show ?case apply(rule_tac disjI2) - using Hiding_Sync_D1[of A S P Q, OF 12(1,2)] D_Sync[of "P \\ A" S "Q \\ A"] by auto + using Hiding_Sync_D1[of A S P Q, OF 12(1,2)] D_Sync[of \P \ A\ S \Q \ A\] by auto qed next case (2 a b) then show ?case proof(elim disjE, goal_cases 21 22) case 21 then show ?case proof(elim exE conjE, elim disjE, goal_cases 211 212 213 214) case (211 t u X Y) then obtain ta ua where t211: "t = trace_hide ta (ev ` A) \ (ta, X \ ev ` A) \ \ P" and u211: "u = trace_hide ua (ev ` A) \ (ua, Y \ ev ` A) \ \ Q" by blast from interleave_Hiding[of "(ev ` A)" "insert tick (ev ` S)" a ta ua] "211"(1,2,3) t211 u211 obtain aa where a211: "a = trace_hide aa (ev ` A) \ aa setinterleaves ((ta, ua), insert tick (ev ` S))" by blast with 211(1,2,3,4) t211 u211 show ?case apply(rule_tac disjI1, rule_tac x="aa" in exI, simp) apply(rule_tac disjI1, rule_tac x="ta" in exI, rule_tac x="ua" in exI) apply(rule_tac x="X \ ev ` A" in exI, rule conjI) using is_processT4 apply blast apply(rule_tac x="Y \ ev ` A" in exI, rule conjI) using is_processT4 by blast+ next case (212 t u X Y) - hence "u \ \ (Q \\ A)" and "t \ \ (P \\ A)" + hence \u \ \ (Q \ A)\ and \t \ \ (P \ A)\ apply (simp add:D_Hiding) using "212"(8) F_T elemTIselemHT by blast - with 212(3,8) have "a \ \ (((P \\ A) \S\ (Q \\ A)))" + with 212(3,8) have \a \ \ (((P \ A) \S\ (Q \ A)))\ apply (simp add:D_Sync) using Sync.sym front_tickFree_charn by blast then show ?case apply (rule_tac disjI2) using Hiding_Sync_D2[of A S P Q, OF 2(2)] D_Hiding[of "(P \S\ Q)", simplified] by auto next case (213 t u X Y) - hence "u \ \ (Q \\ A)" and "t \ \ (P \\ A)" + hence \u \ \ (Q \ A)\ and \t \ \ (P \ A)\ by (simp_all add:D_Hiding, metis F_T elemTIselemHT) - with 213(3,8) have "a \ \ (((P \\ A) \S\ (Q \\ A)))" + with 213(3,8) have \a \ \ (((P \ A) \S\ (Q \ A)))\ apply (simp add:D_Sync) using Sync.sym front_tickFree_charn by blast then show ?case apply (rule_tac disjI2) using Hiding_Sync_D2[of A S P Q, OF 2(2)] D_Hiding[of "(P \S\ Q)", simplified] by auto next case (214 t u X Y) - hence "u \ \ (Q \\ A)" and "t \ \ (P \\ A)" + hence \u \ \ (Q \ A)\ and \t \ \ (P \ A)\ by (simp_all add:D_Hiding T_Hiding) - with 214(3,8) have "a \ \ (((P \\ A) \S\ (Q \\ A)))" + with 214(3,8) have \a \ \ (((P \ A) \S\ (Q \ A)))\ apply (simp add:D_Sync) using Sync.sym front_tickFree_charn by blast then show ?case apply (rule_tac disjI2) using Hiding_Sync_D2[of A S P Q, OF 2(2)] D_Hiding[of "(P \S\ Q)", simplified] by auto qed next case 22 hence \a \ \ (((P \ A) \S\ (Q \ A)))\ by (simp add:D_Sync) then show ?case apply (rule_tac disjI2) using Hiding_Sync_D2[of A S P Q, OF 2(2)] D_Hiding[of "(P \S\ Q)", simplified] by auto qed qed lemma Sync_assoc_oneside_D: "\ (P \S\ (Q \S\ T)) \ \ (P \S\ Q \S\ T)" proof(auto simp add:D_Sync T_Sync del:disjE, goal_cases a) case (a tt uu rr vv) from a(3,4) show ?case proof(auto, goal_cases) case (1 t u) with interleave_assoc_2[OF 1(5) 1(1)] obtain tu where *:"tu setinterleaves ((tt, t), insert tick (ev ` S))" and **:"rr setinterleaves ((tu, u), insert tick (ev ` S))" by blast with a(1,2) 1 show ?case by (metis append.right_neutral front_tickFree_charn) next case (2 t u u1 u2) with interleave_append_sym [OF 2(1)] obtain t1 t2 r1 r2 where a1:"tt = t1 @ t2" and a2:"rr = r1 @ r2" and a3:"r1 setinterleaves ((t1, u1), insert tick (ev ` S))" and a4:"r2 setinterleaves ((t2, u2), insert tick (ev ` S))" by blast with interleave_assoc_2[OF 2(5) a3] obtain tu where *:"tu setinterleaves ((t1, t), insert tick (ev ` S))" and **:"r1 setinterleaves ((tu, u), insert tick (ev ` S))" by blast with a1 a2 a3 a(1,2) 2 show ?case apply(exI tu, exI u, exI r1, exI "r2@vv", intro conjI, simp_all) apply (metis D_imp_front_tickFree append_Nil2 front_tickFree_append front_tickFree_mono ftf_Sync) using D_imp_front_tickFree front_tickFree_append front_tickFree_mono ftf_Sync apply blast by (metis NT_ND Sync.sym append_Nil2 front_tickFree_Nil is_processT3_ST) next case (3 v u u1 u2) with interleave_append_sym [OF 3(1)] obtain t1 t2 r1 r2 where a1:"tt = t1 @ t2" and a2:"rr = r1 @ r2" and a3:"r1 setinterleaves ((t1, u1), insert tick (ev ` S))" and a4:"r2 setinterleaves ((t2, u2), insert tick (ev ` S))" by blast with interleave_assoc_2[OF _ a3, of u v] obtain tu where *:"tu setinterleaves ((t1, u), insert tick (ev ` S))" and **:"r1 setinterleaves ((tu, v), insert tick (ev ` S))" using Sync.sym 3(5) by blast with a1 a2 a3 a(1,2) 3 show ?case apply(exI v, exI tu, exI r1, exI "r2@vv", intro conjI, simp_all) apply (metis D_imp_front_tickFree append_Nil2 front_tickFree_append front_tickFree_mono ftf_Sync) using D_imp_front_tickFree front_tickFree_append front_tickFree_mono ftf_Sync apply blast using Sync.sym apply blast by (meson D_T is_processT3_ST) next case (4 t u) with interleave_assoc_2[OF 4(3) 4(1)] obtain tu where *:"tu setinterleaves ((tt, t), insert tick (ev ` S))" and **:"rr setinterleaves ((tu, u), insert tick (ev ` S))" by blast with a(1,2) 4 show ?case apply(exI tu, exI u, exI rr, exI vv, simp) using NT_ND front_tickFree_Nil by blast next case (5 t u) with interleave_assoc_2[OF _ 5(1), of u t] obtain tu where *:"tu setinterleaves ((tt, u), insert tick (ev ` S))" and **:"rr setinterleaves ((tu, t), insert tick (ev ` S))" using Sync.sym by blast with a(1,2) 5 show ?case apply(exI tu, exI t, exI rr, exI vv, simp) using NT_ND front_tickFree_Nil by blast next case (6 t u u1 u2) with interleave_append [OF 6(1)] obtain t1 t2 r1 r2 where a1:"uu = t1 @ t2" and a2:"rr = r1 @ r2" and a3:"r1 setinterleaves ((t1, u1), insert tick (ev ` S))" and a4:"r2 setinterleaves ((t2, u2), insert tick (ev ` S))" using Sync.sym by blast with interleave_assoc_2[OF 6(5) a3] obtain tu where *:"tu setinterleaves ((t1, t), insert tick (ev ` S))" and **:"r1 setinterleaves ((tu, u), insert tick (ev ` S))" by blast with a1 a2 a3 a(1,2) 6 show ?case apply(exI tu, exI u, exI r1, exI "r2@vv", intro conjI, simp_all) apply (metis append_Nil2 front_tickFree_append front_tickFree_mono ftf_Sync is_processT2_TR) apply (meson front_tickFree_append front_tickFree_mono ftf_Sync is_processT2_TR) by (metis Sync.sym append_Nil2 front_tickFree_Nil is_processT3_ST) next case (7 v u t1 t2) with interleave_append [OF 7(1)] obtain u1 u2 r1 r2 where a1:"uu = u1 @ u2" and a2:"rr = r1 @ r2" and a3:"r1 setinterleaves ((t1, u1), insert tick (ev ` S))" and a4:"r2 setinterleaves ((t2, u2), insert tick (ev ` S))" by blast with interleave_assoc_2[of t1 u _ v r1 u1] obtain tu where *:"tu setinterleaves ((u1, u), insert tick (ev ` S))" and **:"r1 setinterleaves ((tu, v), insert tick (ev ` S))" using 7(5) Sync.sym by blast with a1 a2 a3 a(1,2) 7 show ?case apply(exI v, exI tu, exI r1, exI "r2@vv", intro conjI, simp_all) apply (metis append_Nil2 front_tickFree_append front_tickFree_mono ftf_Sync is_processT2_TR) apply (meson front_tickFree_append front_tickFree_mono ftf_Sync is_processT2_TR) using Sync.sym apply blast by (meson D_T is_processT3_ST) next case (8 t u) with interleave_assoc_2[OF 8(3), of rr uu] obtain tu where *:"tu setinterleaves ((uu, t), insert tick (ev ` S))" and **:"rr setinterleaves ((tu, u), insert tick (ev ` S))" using Sync.sym by blast with a(1,2) 8 show ?case apply(exI tu, exI u, exI rr, exI vv, simp) using Sync.sym front_tickFree_Nil by blast next case (9 t u) with interleave_assoc_1[OF 9(3), of rr uu] obtain uv where *:"uv setinterleaves ((u, uu), insert tick (ev ` S))" and **:"rr setinterleaves ((t, uv), insert tick (ev ` S))" by blast with a(1,2) 9 show ?case apply(exI t, exI uv, exI rr, exI vv, simp) using Sync.sym front_tickFree_Nil by blast qed qed lemma Sync_assoc_oneside: "((P \S\ Q) \S\ T) \ (P \S\ (Q \S\ T))" proof(unfold le_ref_def, rule conjI, goal_cases D F) case D then show ?case using Sync_assoc_oneside_D by assumption next case F then show ?case unfolding F_Sync proof(rule Un_least, goal_cases FF DD) case (FF) then show ?case proof(rule subsetI, simp add:D_Sync T_Sync split_paired_all, elim exE conjE disjE, goal_cases) case (1 r Xtuv t uv Xt Xuv u v Xu Xv) with interleave_assoc_2[OF 1(6), OF 1(2)] obtain tu where *:"tu setinterleaves ((t, u), insert tick (ev ` S))" and **:"r setinterleaves ((tu, v), insert tick (ev ` S))" by blast with 1(1,3,4,5,7) show ?case apply(rule_tac disjI1, exI tu, exI v, exI "(Xt \ Xu) \ insert tick (ev ` S) \ Xt \ Xu", intro conjI disjI1, simp_all) by blast+ next case (2 r Xtuv t uv Xt Xuv u v uv1 uv2) with interleave_append_sym [of r t _ uv1 uv2] obtain t1 t2 r1 r2 where a1:"t = t1 @ t2" and a2:"r = r1 @ r2" and a3:"r1 setinterleaves ((t1, uv1), insert tick (ev ` S))" and a4:"r2 setinterleaves ((t2, uv2), insert tick (ev ` S))" by metis with interleave_assoc_2[OF 2(6), OF a3] obtain tu where *:"tu setinterleaves ((t1, u), insert tick (ev ` S))" and **:"r1 setinterleaves ((tu, v), insert tick (ev ` S))" by blast from 2(1) a1 obtain Xt1 where ***:"(t1, Xt1) \ \ P" using is_processT3_SR by blast from 2 a1 a2 a3 a4 * ** *** show ?case apply(rule_tac disjI2, exI tu, exI v, exI r1, exI r2, intro conjI disjI1, simp_all) apply (metis front_tickFree_charn front_tickFree_mono ftf_Sync is_processT2) apply (metis equals0D ftf_Sync1 ftf_Sync21 insertI1 tickFree_def) using F_T Sync.sym front_tickFree_Nil by blast next case (3 r Xtuv t uv Xt Xuv v u uv1 uv2) with interleave_append_sym [of r t _ uv1 uv2] obtain t1 t2 r1 r2 where a1:"t = t1 @ t2" and a2:"r = r1 @ r2" and a3:"r1 setinterleaves ((t1, uv1), insert tick (ev ` S))" and a4:"r2 setinterleaves ((t2, uv2), insert tick (ev ` S))" by metis with interleave_assoc_2[OF _ a3, of u v] obtain tu where *:"tu setinterleaves ((t1, u), insert tick (ev ` S))" and **:"r1 setinterleaves ((tu, v), insert tick (ev ` S))" using "3"(6) Sync.sym by blast from 3(1) a1 obtain Xt1 where ***:"(t1, Xt1) \ \ P" using is_processT3_SR by blast from 3 a1 a2 a3 a4 * ** *** show ?case apply(rule_tac disjI2, exI v, exI tu, exI r1, exI r2, intro conjI, simp_all) apply (metis front_tickFree_charn front_tickFree_mono ftf_Sync is_processT2) apply (metis equals0D ftf_Sync1 ftf_Sync21 insertI1 tickFree_def) using Sync.sym F_T by blast+ next case (4 r Xtuv t uv Xt Xuv u v uv1 uv2) with interleave_assoc_2[OF 4(6), of r t] obtain tu where *:"tu setinterleaves ((t, u), insert tick (ev ` S))" and **:"r setinterleaves ((tu, v), insert tick (ev ` S))" by auto from 4 * ** show ?case apply(rule_tac disjI2, exI tu, exI v, exI r, exI "[]", intro conjI, simp_all) by (metis "4"(4) F_T Sync.sym append.right_neutral) next case (5 r Xtuv t uv Xt Xuv v u uv1 uv2) from 5(2,7,5,6) interleave_assoc_2[of uv u _ v r t] obtain tu where *:"tu setinterleaves ((t, u), insert tick (ev ` S))" and **:"r setinterleaves ((tu, v), insert tick (ev ` S))" by (metis Sync.sym append_Nil2) from 5 * ** show ?case apply(rule_tac disjI2, exI v, exI tu, exI r, exI "[]", intro conjI, simp_all) using Sync.sym F_T by blast+ qed next case (DD) then show ?case using Sync_assoc_oneside_D[of P S Q T] by (simp add:D_Sync Collect_mono_iff sup.coboundedI2) qed qed lemma Sync_assoc: "((P \S\ Q) \S\ T) = (P \S\ (Q \S\ T))" by (metis antisym_conv Sync_assoc_oneside Sync_commute) subsection\ Derivative Operators laws \ (* Trivial results have been removed *) lemmas Par_BOT = Sync_BOT[where S = \UNIV\] and Par_SKIP_SKIP = Sync_SKIP_SKIP[where S = \UNIV\] and Par_SKIP_STOP = Sync_SKIP_STOP[where S = \UNIV\] and Par_commute = Sync_commute[where S = \UNIV\] and Par_assoc = Sync_assoc[where S = \UNIV\] and Par_Ndet_right_distrib = Sync_Ndet_right_distrib[where S = \UNIV\] and Par_Ndet_left_distrib = Sync_Ndet_left_distrib[where S = \UNIV\] and Mprefix_Par_SKIP = Mprefix_Sync_SKIP[where S = \UNIV\] and prefix_Par_SKIP = prefix_Sync_SKIP[where S = \UNIV\, simplified] and mono_Par_FD = mono_Sync_FD[where S = \UNIV\] and mono_Par_ref = mono_Sync_ref[where S = \UNIV\] and prefix_Par1 = prefix_Sync1[where S = \UNIV\, simplified] and prefix_Par2 = prefix_Sync2[where S = \UNIV\, simplified] and Mprefix_Par_distr = Mprefix_Sync_distr_subset[where S = \UNIV\, simplified] and Inter_BOT = Sync_BOT[where S = \{}\] and Inter_SKIP_SKIP = Sync_SKIP_SKIP[where S = \{}\] and Inter_SKIP_STOP = Sync_SKIP_STOP[where S = \{}\] and Inter_commute = Sync_commute[where S = \{}\] and Inter_assoc = Sync_assoc[where S = \{}\] and Inter_Ndet_right_distrib = Sync_Ndet_right_distrib[where S = \{}\] and Inter_Ndet_left_distrib = Sync_Ndet_left_distrib[where S = \{}\] and Mprefix_Inter_SKIP = Mprefix_Sync_SKIP[where S = \{}\, simplified, simplified Mprefix_STOP] and prefix_Inter_SKIP = prefix_Sync_SKIP[where S = \{}\, simplified] and mono_Inter_FD = mono_Sync_FD[where S = \{}\] and mono_Inter_ref = mono_Sync_ref[where S = \{}\] and Hiding_Inter = Hiding_Sync[where S = \{}\, simplified] and Mprefix_Inter_distr = Mprefix_Sync_distr_indep[where S = \{}\, simplified] and prefix_Inter = Mprefix_Sync_distr_indep[of \{a}\ \{}\ \{b}\ \\x. P\ \\x. Q\, simplified, folded write0_def] for a P b Q lemma Inter_SKIP: "(P ||| SKIP) = P" proof(auto simp:Process_eq_spec D_Sync F_Sync F_SKIP D_SKIP T_SKIP, goal_cases) case (1 a t X Y) then have aa:"(X \ Y) \ {tick} \ X \ Y \ X" by (simp add: Int_Un_distrib2) have "front_tickFree t" using "1"(1) is_processT2 by blast with 1 EmptyLeftSync[of a "{tick}" t] have bb:"a=t" using Sync.sym by blast from 1(1) aa bb show ?case using is_processT4 by blast next case (2 a t X Y) have "front_tickFree t" using 2(1) is_processT2 by blast with 2 TickLeftSync[of "{tick}" t a] have bb:"a=t \ last t=tick" using Sync.sym by blast with 2 have "a \ \ P \ last a = tick" using F_T by blast with tick_T_F[of "butlast a" P "(X \ Y) \ {tick} \ X \ Y"] show ?case by (metis "2"(2) EmptyLeftSync bb self_append_conv2 snoc_eq_iff_butlast) next case (3 b t r v) with 3 EmptyLeftSync[of r "{tick}" t] have bb:"r=t" using Sync.sym by blast with 3 show ?case using is_processT by blast next case (4 b t r v) with TickLeftSync[of "{tick}" t r, simplified] have bb:"r=t" using D_imp_front_tickFree Sync.sym by blast with 4 show ?case using is_processT by blast next case (5 b t r) with EmptyLeftSync[of r "{tick}" t] have bb:"r=t" using Sync.sym by blast with 5 show ?case using is_processT by blast next case (6 b t r) with TickLeftSync[of "{tick}" t r, simplified] have bb:"r=t" using D_imp_front_tickFree Sync.sym by blast with 6 show ?case using is_processT by blast next case (7 a b) then show ?case apply(cases "tickFree a") apply(rule_tac x=a in exI, rule_tac x="[]" in exI, rule_tac x=b in exI, simp) apply(rule_tac x="b - {tick}" in exI, simp, intro conjI) using emptyLeftSelf[of a "{tick}"] Sync.sym tickFree_def apply fastforce apply blast apply(rule_tac x=a in exI, rule_tac x="[tick]" in exI, rule_tac x=b in exI, simp, rule conjI) subgoal proof - assume a1: "\ tickFree a" assume a2: "(a, b) \ \ P" then obtain ees :: "'a event list \ 'a event list" where f3: "a = ees a @ [tick]" using a1 by (meson is_processT2 nonTickFree_n_frontTickFree) then have "ees a setinterleaves (([], ees a), {tick})" using a2 by (metis (no_types) DiffD2 Diff_insert_absorb emptyLeftSelf front_tickFree_implies_tickFree is_processT2 tickFree_def) then show "a setinterleaves ((a, [tick]), {tick})" using f3 by (metis (no_types) Sync.sym addSyncEmpty insertI1) qed apply(rule_tac x="b - {tick}" in exI) by blast next case (8 t r v) with EmptyLeftSync[of r "{tick}" t] have bb:"r=t" using Sync.sym by blast with 8 show ?case using is_processT by blast next case (9 t r v) with TickLeftSync[of "{tick}" t r, simplified] have bb:"r=t" using D_imp_front_tickFree Sync.sym by blast with 9 show ?case using is_processT by blast next case (10 t r) with EmptyLeftSync[of r "{tick}" t] have bb:"r=t" using Sync.sym by blast with 10 show ?case by simp next case (11 t r) with TickLeftSync[of "{tick}" t r, simplified] have bb:"r=t" using D_imp_front_tickFree Sync.sym by blast with 11 show ?case by simp next case (12 x) then show ?case proof(cases "tickFree x") case True with 12 show ?thesis apply(rule_tac x=x in exI, rule_tac x="[]" in exI, rule_tac x=x in exI, rule_tac x="[]" in exI, simp) by (metis Sync.sym emptyLeftSelf singletonD tickFree_def) next case False with 12 show ?thesis apply(rule_tac x="butlast x" in exI, rule_tac x="[]" in exI, rule_tac x="butlast x" in exI, rule_tac x="[tick]" in exI, simp, intro conjI) apply (metis D_imp_front_tickFree append_butlast_last_id front_tickFree_mono list.distinct(1) tickFree_Nil) using NT_ND T_nonTickFree_imp_decomp apply fastforce apply (metis NT_ND Sync.sym append_T_imp_tickFree append_butlast_last_id emptyLeftSelf list.distinct(1) singletonD tickFree_Nil tickFree_def) using D_imp_front_tickFree is_processT9 nonTickFree_n_frontTickFree by fastforce qed qed lemma Inter_STOP_Seq_STOP: "(P ||| STOP) = (P \<^bold>; STOP)" proof (auto simp add:Process_eq_spec D_Sync F_Sync F_STOP D_STOP T_STOP F_Seq D_Seq, goal_cases) case (1 a t X Y) have "insert tick ((X \ Y) \ {tick} \ X \ Y) \ X \ {tick}" by blast with 1 show ?case by (metis (mono_tags, lifting) EmptyLeftSync Sync.sym is_processT is_processT5_S2 no_Trace_implies_no_Failure) next case (2 a t X) then show ?case using EmptyLeftSync Sync.sym tickFree_def by fastforce next case (3 b t r v) - then show ?case using EmptyLeftSync Sync.sym tickFree_def by blast + then show ?case using EmptyLeftSync Sync.sym is_processT7 by blast next case (4 t r v) - then show ?case using EmptyLeftSync Sync.sym by blast + then show ?case using EmptyLeftSync Sync.sym is_processT7 by blast next case (5 b t r) - then show ?case by (metis EmptyLeftSync Sync.sym is_processT) + then show ?case using EmptyLeftSync Sync.sym is_processT7 by blast next case (6 t r) - then show ?case using EmptyLeftSync Sync.sym tickFree_def by fastforce + then show ?case using EmptyLeftSync Sync.sym is_processT7 by blast next case (7 a b) then show ?case apply(rule_tac x=a in exI, rule_tac x="b" in exI, intro conjI) apply (meson insert_iff is_processT subsetI) using emptyLeftSelf[of a "{tick}"] Sync.sym tickFree_def apply fastforce by blast next case (8 a b) then show ?case apply(rule_tac x=a in exI, rule_tac x="b-{tick}" in exI, intro conjI) apply (meson NF_NT is_processT6) using emptyLeftSelf[of a "{tick}"] Sync.sym tickFree_def apply (metis append_T_imp_tickFree list.distinct(1) singletonD) by blast next - case (9 b t1 t2) - hence "t1 setinterleaves ((t1, []), {tick})" - using emptyLeftSelf[of t1 "{tick}"] Sync.sym tickFree_def by fastforce - with 9(1)[THEN spec, THEN spec, of t1 t1, simplified] have False - using "9"(2) "9"(3) "9"(4) by blast - then show ?case .. + case (9 t X) + then show ?case + apply (cases \tickFree t\) + by (metis Sync.sym append.right_neutral emptyLeftSelf front_tickFree_Nil singletonD tickFree_def) + (metis D_T D_imp_front_tickFree Sync.sym append_single_T_imp_tickFree emptyLeftSelf + front_tickFree_single is_processT9_S_swap nonTickFree_n_frontTickFree singletonD tickFree_def) next case (10 t r v) - then show ?case - apply(rule_tac x=r in exI, rule_tac x="v" in exI, auto) - using EmptyLeftSync Sync.sym by blast + then show ?case using EmptyLeftSync Sync.sym is_processT7 by blast next case (11 t r) - then show ?case - apply(rule_tac x=r in exI, rule_tac x="[]" in exI, simp) - using EmptyLeftSync Sync.sym tickFree_def by fastforce + then show ?case using Sync.sym emptyLeftProperty by blast next - case (12 t1 t2) + case (12 t) then show ?case - apply(rule_tac x=t1 in exI, rule_tac x=t1 in exI, rule_tac x=t2 in exI, simp) - by (metis Sync.sym emptyLeftSelf singletonD tickFree_def) + apply (cases \tickFree t\) + apply (rule_tac x = t in exI, rule_tac x = t in exI, simp, + metis Sync.sym emptyLeftSelf singletonD tickFree_def) + by (rule_tac x = \butlast t\ in exI, rule_tac x = \butlast t\ in exI, + rule_tac x = \[tick]\ in exI, simp, + metis D_imp_front_tickFree Sync.sym emptyLeftSelf front_tickFree_implies_tickFree tickFree_def + is_processT9_S_swap nonTickFree_n_frontTickFree singletonD snoc_eq_iff_butlast) qed lemma Par_STOP: "P \ \ \ (P || STOP) = STOP" proof(auto simp add:Process_eq_spec, simp_all add:D_Sync F_Sync F_STOP D_STOP T_STOP is_processT2 D_imp_front_tickFree F_UU D_UU, goal_cases) case (1 a b aa ba) then show ?case proof(elim exE conjE disjE, goal_cases) case (1 t X Y) then show ?case proof(cases t) case Nil then show ?thesis using "1"(4) EmptyLeftSync by blast next case (Cons a list) with 1 show ?thesis by (simp split:if_splits, cases a, simp_all) qed next case (2 t r v) hence "[] \ \ P" proof(cases t) case Nil with 2 show ?thesis by simp next case (Cons a list) with 2 show ?thesis by (simp split:if_splits, cases a, simp_all) qed with 2(1,2) show ?case using NF_ND is_processT7 by fastforce next case (3 t r v) hence "[] \ \ P" proof(cases t) case Nil with 3 show ?thesis by simp next case (Cons a list) with 3 show ?thesis by (simp split:if_splits, cases a, simp_all) qed with 3(1,2) show ?case using NF_ND is_processT7 by fastforce qed next case (2 a b aa ba) then show ?case apply(rule_tac disjI1, rule_tac x="[]" in exI, rule_tac x="{}" in exI, simp, rule_tac conjI) using is_processT apply blast apply(rule_tac x=ba in exI, rule_tac set_eqI, rule_tac iffI) using event_set apply blast by fast next case (3 a b x) then show ?case proof(elim exE conjE disjE, goal_cases) case (1 t r v) hence "[] \ \ P" proof(cases t) case Nil with 1 show ?thesis by simp next case (Cons a list) with 1 show ?thesis by (simp split:if_splits, cases a, simp_all) qed with 3(1,2) show ?case using NF_ND is_processT7 by fastforce next case (2 t r v) hence "[] \ \ P" proof(cases t) case Nil with 2 show ?thesis by simp next case (Cons a list) with 2 show ?thesis by (simp split:if_splits, cases a, simp_all) qed with 3(1,2) show ?case using NF_ND is_processT7 by fastforce qed next case (4 x a b) then show ?case proof(elim exE conjE disjE, goal_cases) case (1 t X Y) then show ?case proof(cases t) case Nil then show ?thesis using "1"(4) EmptyLeftSync by blast next case (Cons a list) with 1 show ?thesis by (simp split:if_splits, cases a, simp_all) qed next case (2 t r v) hence "[] \ \ P" proof(cases t) case Nil with 2 show ?thesis by simp next case (Cons a list) with 2 show ?thesis by (simp split:if_splits, cases a, simp_all) qed with 2(1,2) show ?case using NF_ND is_processT7 by fastforce next case (3 t r v) hence "[] \ \ P" proof(cases t) case Nil with 3 show ?thesis by simp next case (Cons a list) with 3 show ?thesis by (simp split:if_splits, cases a, simp_all) qed with 3(1,2) show ?case using NF_ND is_processT7 by fastforce qed next case (5 x a b) then show ?case apply(rule_tac disjI1, rule_tac x="[]" in exI, rule_tac x="{}" in exI, simp, rule_tac conjI) using is_processT apply blast apply(rule_tac x=b in exI, rule_tac set_eqI, rule_tac iffI) using event_set apply blast by fast next case (6 x xa) then show ?case proof(elim exE conjE disjE, goal_cases) case (1 t r v) hence "[] \ \ P" proof(cases t) case Nil with 1 show ?thesis by simp next case (Cons a list) with 1 show ?thesis by (simp split:if_splits, cases a, simp_all) qed with 6(1,2) show ?case using NF_ND is_processT7 by fastforce next case (2 t r v) hence "[] \ \ P" proof(cases t) case Nil with 2 show ?thesis by simp next case (Cons a list) with 2 show ?thesis by (simp split:if_splits, cases a, simp_all) qed with 6(1,2) show ?case using NF_ND is_processT7 by fastforce qed qed section\Multiple Non Deterministic Operator Laws\ lemma Mprefix_refines_Mndetprefix: "(\ x \ A \ P x) \ (\ x \ A \ P x)" proof(cases "A = {}") case True then show ?thesis by (simp add: Mprefix_STOP) next case False then show ?thesis by (auto simp add:le_ref_def D_Mprefix write0_def F_Mprefix F_Mndetprefix D_Mndetprefix) qed lemma Mndetprefix_refine_FD: "a \ A \ (\ x \ A \ (P x)) \ (a \ (P a))" by(subgoal_tac "A \ {}", auto simp:le_ref_def D_Mndetprefix F_Mndetprefix) lemma Mndetprefix_subset_FD:"A \ {} \ (\xa\ A \ B \ P) \ (\xa\ A \ P)" by (simp add: D_Mndetprefix F_Mndetprefix le_ref_def) lemma mono_Mndetprefix_FD[simp]: "\x \ A. P x \ P' x \ (\ x \ A \ (P x)) \ (\ x \ A \ (P' x))" apply(cases "A \ {}", auto simp:le_ref_def D_Mndetprefix F_Mndetprefix) by(meson contra_subsetD le_ref_def mono_write0_FD)+ lemma Mndetprefix_FD_subset : "A \ {} \ A \ B \ ((\ x \ B \ P x) \ (\ x \ A \ P x))" by (metis bot.extremum_uniqueI Mndetprefix_Un_distrib mono_Ndet_FD_right order_refl sup.absorb_iff1) section \ Infra-structure for Communication Primitives \ lemma read_read_Sync: assumes contained: "(\y. c y \ S)" shows "((c\<^bold>?x \ P x) \S\ (c\<^bold>?x \ Q x)) = (c\<^bold>?x \ ((P x) \S\ (Q x)))" proof - have A: "range c \ S" by (insert contained, auto) show ?thesis by (auto simp: read_def o_def Set.Int_absorb2 Mprefix_Sync_distr_subset A) qed lemma read_read_non_Sync: "\\y. c y \ S; \y. d y \ S\ \ (c\<^bold>?x \ (P x)) \S\ (d\<^bold>?x \(Q x)) = c\<^bold>?x \ ((P x) \S\ (d\<^bold>?x \(Q x)))" apply (auto simp add: read_def o_def) using Mprefix_Sync_distr_right[of "range c" S "range d" "\x. Q (inv d x)" "\x. P (inv c x)"] apply (subst (1 2) Sync_commute) by auto lemma write_read_Sync: assumes contained: "\y. c y \ S" and is_construct: "inj c" shows "(c\<^bold>!a \ P) \S\ (c\<^bold>?x \ Q x) = (c\<^bold>!a \ (P \S\ Q a))" proof - have A : "range c \ S" by (insert contained, auto) have B : " {c a} \ range c \ S = {c a}" by (insert contained, auto) show ?thesis apply (simp add: read_def write_def o_def, subst Mprefix_Sync_distr_subset) by (simp_all add: A B contained Mprefix_singl is_construct) qed lemma write_read_non_Sync: "\d a \ S; \y. c y \ S \ \ (d\<^bold>!a \ P) \S\ (c\<^bold>?x \ Q x) = d\<^bold>!a \ (P \S\ (c\<^bold>?x \ Q x))" apply (simp add: read_def o_def write_def) using Mprefix_Sync_distr_right[of "{d a}" S "range c" "\x. (Q (inv c x))" "\x. P"] apply (subst (1 2) Sync_commute) by auto lemma write0_read_non_Sync: "\d \ S; \y. c y \ S\ \ (d \ P) \S\ (c\<^bold>?x \ Q x) = c\<^bold>?x \ ((d \ P) \S\ Q x)" apply (simp add: read_def o_def) apply (subst (1 2) Mprefix_singl[symmetric]) apply (subst Mprefix_Sync_distr_right) by auto lemma write0_write_non_Sync: "\d a \ C; c \ C\ \ (c \ Q) \C\ (d\<^bold>!a \ P) = d\<^bold>!a \ ((c \ Q) \C\ P)" apply(simp add: write_def ) apply(subst (1 2) Mprefix_singl[symmetric]) by (auto simp: Mprefix_Sync_distr_right) + + +section \Renaming operator laws\ + +lemma Renaming_BOT: \Renaming \ f = \\ + by (fastforce simp add: F_UU D_UU F_Renaming D_Renaming EvExt_tF EvExt_ftF + Process_eq_spec front_tickFree_append intro: tickFree_Nil)+ +lemma Renaming_STOP: \Renaming STOP f = STOP\ + by (subst Process_eq_spec) (simp add: EvExt_ftF F_STOP D_STOP F_Renaming D_Renaming) + + +lemma Renaming_SKIP: \Renaming SKIP f = SKIP\ + by (subst Process_eq_spec) (force simp add: EvExt_def F_SKIP D_SKIP F_Renaming D_Renaming) + + + +lemma Renaming_Ndet: \Renaming (P \ Q) f = Renaming P f \ Renaming Q f\ + by (subst Process_eq_spec) (auto simp add: F_Renaming D_Renaming F_Ndet D_Ndet) + + +lemma Renaming_Det: \Renaming (P \ Q) f = Renaming P f \ Renaming Q f\ +proof (subst Process_eq_spec, intro iffI conjI subset_antisym) + show \\ (Renaming (P \ Q) f) \ \ (Renaming P f \ Renaming Q f)\ + apply (simp add: F_Renaming D_Renaming T_Renaming F_Det D_Det, safe, + simp_all add: is_processT6_S2) + by blast+ (metis EvExt_eq_tick, meson map_EvExt_tick)+ +next + show \\ (Renaming P f \ Renaming Q f) \ \ (Renaming (P \ Q) f)\ + apply (simp add: F_Renaming D_Renaming T_Renaming F_Det D_Det, safe, simp_all) + by blast+ (metis EvExt_eq_tick, metis Cons_eq_append_conv EvExt_tF + list.map_disc_iff tickFree_Cons)+ +next + show \\ (Renaming (P \ Q) f) \ \ (Renaming P f \ Renaming Q f)\ + by (auto simp add: D_Renaming D_Det) +next + show \\ (Renaming P f \ Renaming Q f) \ \ (Renaming (P \ Q) f)\ + by (auto simp add: D_Renaming D_Det) +qed + + + + +(* lemma THE_inj_trivial: \inj f \ (THE x. f x = f a) = a\ by (simp add: inj_eq) + +lemma THE_inj_on_trivial: \inj_on f A \ a \ A \ (THE x. x \ A \ f x = f a) = a\ + by (simp add: inj_onD the_equality) + *) + +(* TODO: Move this in Mprefix and Mndetprefix theories *) + +lemma mono_Mprefix_eq: \\a \ A. P a = Q a \ Mprefix A P = Mprefix A Q\ + by (subst Process_eq_spec, auto simp add: F_Mprefix D_Mprefix) + +lemma mono_Mndetprefix_eq: \\a \ A. P a = Q a \ Mndetprefix A P = Mndetprefix A Q\ + by (cases \A = {}\, simp) (subst Process_eq_spec, auto simp add: F_Mndetprefix D_Mndetprefix) + + + + +lemma Renaming_Mprefix: \Renaming (\ a \ A \ P (f a)) f = \ b \ f ` A \ Renaming (P b) f\ +proof (subst Process_eq_spec, intro conjI subset_antisym) + show \\ (Renaming (\ a \ A \ P (f a)) f) \ \ (\ b \ f ` A \ Renaming (P b) f)\ + apply (simp add: F_Renaming F_Mprefix D_Mprefix, safe; + auto simp add: hd_map_EvExt map_tl tickFree_tl) + by ((meson disjoint_iff ev_elem_anteced1 image_eqI)+)[3] +next + show \\ (\ b \ f ` A \ Renaming (P b) f) \ \ (Renaming (\ a \ A \ P (f a)) f)\ + apply (simp add: F_Renaming F_Mprefix D_Mprefix, safe, simp_all) + apply (auto simp add: vimage_def image_def, insert hd_map_EvExt, fastforce)[1] + apply (rule_tac x = \ev xa # s1\ in exI, simp, + metis hd_map hd_map_EvExt list.collapse list.discI list.sel(1)) + proof goal_cases + case (1 a b xa aa s1 s2) + have \s1 \ \ (P (f xa))\ + apply (rule "1"(1)[rule_format, of \ev xa # s1\ s2, simplified]) + by (simp_all add: "1"(4, 6, 7)) + (metis "1"(2, 3, 8) hd_map_EvExt list.collapse list.distinct(1) list.sel(1) list.simps(9)) + with \s1 \ \ (P (f xa))\ show ?case by blast + qed +next + show \\ (Renaming (\ a \ A \ P (f a)) f) \ \ (\ b \ f ` A \ Renaming (P b) f)\ + apply (auto simp add: D_Mprefix D_Renaming hd_map_EvExt) + by (rule_tac x = \tl s1\ in exI, simp add: map_tl tickFree_tl) +next + show \\ (\ b \ f ` A \ Renaming (P b) f) \ \ (Renaming (\ a \ A \ P (f a)) f)\ + apply (auto simp add: D_Mprefix D_Renaming) + by (rule_tac x = \ev xb # s1\ in exI, simp, rule_tac x = s2 in exI, simp) + (metis EvExt_def event.case(1) list.collapse o_apply) +qed + + + +lemma Renaming_Mprefix_inj_on: + \Renaming (Mprefix A P) f = \ b \ f ` A \ Renaming (P (THE a. a \ A \ f a = b)) f\ + if inj_on_f: \inj_on f A\ + apply (subst Renaming_Mprefix[symmetric]) + apply (intro arg_cong[where f = \\Q. Renaming Q f\] mono_Mprefix_eq) + by (metis that the_inv_into_def the_inv_into_f_f) + + +corollary Renaming_Mprefix_inj: + \Renaming (Mprefix A P) f = \ b \ f ` A \ Renaming (P (THE a. f a = b)) f\ if inj_f: \inj f\ + apply (subst Renaming_Mprefix_inj_on, metis inj_eq inj_onI that) + apply (rule mono_Mprefix_eq[rule_format]) + by (metis imageE inj_eq[OF inj_f]) + + + +text \A smart application (as \<^term>\f\ is of course injective on the singleton \<^term>\{a}\)\ + +corollary Renaming_prefix: \Renaming (a \ P) f = (f a \ Renaming P f)\ + unfolding write0_def by (simp add: Renaming_Mprefix_inj_on) + +(* TODO: write corollaries on read and write *) + + + +lemma Renaming_Mndetprefix: \Renaming (\ a \ A \ P (f a)) f = \ b \ f ` A \ Renaming (P b) f\ + apply (cases \A = {}\, simp add: Renaming_STOP) + by (subst Process_eq_spec) + (auto simp add: F_Renaming F_Mndetprefix D_Renaming D_Mndetprefix Renaming_prefix[symmetric]) + + +corollary Renaming_Mndetprefix_inj_on: + \Renaming (Mndetprefix A P) f = \ b \ f ` A \ Renaming (P (THE a. a \ A \ f a = b)) f\ + if inj_on_f: \inj_on f A\ + apply (subst Renaming_Mndetprefix[symmetric]) + apply (intro arg_cong[where f = \\Q. Renaming Q f\] mono_Mndetprefix_eq) + by (metis that the_inv_into_def the_inv_into_f_f) + + + +corollary Renaming_Mndetprefix_inj: + \Renaming (Mndetprefix A P) f = \ b \ f ` A \ Renaming (P (THE a. f a = b)) f\ + if inj_f: \inj f\ + apply (subst Renaming_Mndetprefix_inj_on, metis inj_eq inj_onI that) + apply (rule mono_Mndetprefix_eq[rule_format]) + by (metis imageE inj_eq[OF inj_f]) + + + + + + +lemma Renaming_Seq: \Renaming (P \<^bold>; Q) f = Renaming P f \<^bold>; Renaming Q f\ +proof (subst Process_eq_spec, intro conjI impI subset_antisym) + show \\ (Renaming (P \<^bold>; Q) f) \ \ (Renaming P f \<^bold>; Renaming Q f)\ + apply (simp add: F_Seq D_Seq F_Renaming D_Renaming T_Renaming, safe; + auto simp add: EvExt_tF append_single_T_imp_tickFree) + apply (rule_tac x = s1 in exI) + apply (metis Diff_insert_absorb Diff_single_insert anteced_EvExt_diff_tick + insert_absorb insert_absorb2 is_processT4 subset_insertI) + apply (rule_tac x = \t1 @ t2\ in exI, metis map_EvExt_tick map_append) + apply (metis map_EvExt_tick map_append) + apply (metis NF_ND) + apply (metis EvExt_ftF front_tickFree_mono map_append nonTickFree_n_frontTickFree + non_tickFree_tick process_charn tickFree_Nil) + apply (rule_tac x = \t1 @ t2\ in exI) + by (metis map_EvExt_tick map_append)+ +next + show \\ (Renaming P f \<^bold>; Renaming Q f) \ \ (Renaming (P \<^bold>; Q) f)\ + proof ((simp add: F_Seq D_Seq F_Renaming D_Renaming T_Renaming subset_iff, safe; auto), + goal_cases) + case (1 Y s1) + then show ?case + apply (rule_tac x = \s1\ in exI, safe; simp add: EvExt_tF) + by (metis (no_types, lifting) Diff_insert_absorb anteced_EvExt_diff_tick insert_Diff + insert_absorb insert_iff tick_elem_anteced1) + next + case (2 Y t1 t1a s1) + then show ?case + apply (rule_tac x = \butlast t1a @ s1\ in exI) + by (metis (no_types, opaque_lifting) EvExt_tF T_nonTickFree_imp_decomp butlast_snoc + map_append map_butlast non_tickFree_tick tickFree_append) + next + case (3 Y t1 t1a s1 s2) + have \butlast t1a @ s1 \ \ P \ + (\t1. t1 @ [tick] \ \ P \ (\t2. butlast t1a @ s1 = t1 @ t2 \ t2 \ \ Q))\ + apply (rule "3"(1)[rule_format, of \butlast t1a @ s1\, rotated 2, OF "3"(5)]) + using "3"(2) "3"(4) front_tickFree_butlast is_processT2_TR tickFree_append apply blast + by (metis "3"(3) append.assoc butlast_snoc map_append map_butlast) + then show ?case + by (metis "3"(2, 3, 6) EvExt_tF T_nonTickFree_imp_decomp snoc_eq_iff_butlast tickFree_butlast) + next + case (4 Y t1 t1a t2a s1) + have \(if t2a = [] then butlast t1a else t1a) \ \ P \ + (\t1. t1 @ [tick] \ \ P \ + (\t2. (if t2a = [] then butlast t1a else t1a) = t1 @ t2 \ t2 \ \ Q))\ + apply (rule "4"(1)[rule_format, of _ \butlast t2a @ map (EvExt f) s1\]) + using "4"(2) front_tickFree_butlast tickFree_implies_front_tickFree apply fastforce + apply (auto, metis "4"(4) append.right_neutral butlast_snoc map_butlast, + metis "4"(4) butlast_append butlast_snoc) + by (metis "4"(3, 6) EvExt_ftF front_tickFree_append front_tickFree_butlast is_processT2) + then show ?case + by (metis "4"(2, 4, 5) EvExt_tF append.right_neutral tickFree_Cons tickFree_append) + next + case (5 Y t1 t1a t2a s1 s2) + have \(if t2a = [] then butlast t1a else t1a) \ \ P \ + (\t1. t1 @ [tick] \ \ P \ + (\t2. (if t2a = [] then butlast t1a else t1a) = t1 @ t2 \ t2 \ \ Q))\ + apply (rule "5"(1)[rule_format, of _ \butlast t2a @ map (EvExt f) s1 @ s2\]) + using "5"(2) tickFree_butlast apply fastforce + apply (auto, metis "5"(4) append.right_neutral butlast_snoc map_butlast, + metis "5"(4) butlast_append butlast_snoc) + by (metis "5"(3, 6, 7) EvExt_tF front_tickFree_append front_tickFree_butlast) + then show ?case + by (metis "5"(2, 4, 5) EvExt_tF append_Nil2 non_tickFree_tick tickFree_append) + qed +next + show \\ (Renaming (P \<^bold>; Q) f) \ \ (Renaming P f \<^bold>; Renaming Q f)\ + apply (auto simp add: D_Seq D_Renaming T_Renaming) + by (metis map_EvExt_tick map_append) +next + show \\ (Renaming P f \<^bold>; Renaming Q f) \ \ (Renaming (P \<^bold>; Q) f)\ + apply (auto simp add: D_Seq D_Renaming T_Renaming) + apply (rule_tac x = \butlast t1a @ s1\ in exI, auto, + metis append_butlast_last_id append_single_T_imp_tickFree butlast.simps(1) tickFree_Nil) + apply (rule_tac x = s2 in exI, auto, metis butlast_snoc map_butlast, + metis EvExt_tF T_nonTickFree_imp_decomp snoc_eq_iff_butlast tickFree_butlast) + apply (rule_tac x = \t1a\ in exI, simp) + apply (rule_tac x = \butlast t2a @ map (EvExt f) s1 @ s2\ in exI, auto) + apply (meson EvExt_tF front_tickFree_append front_tickFree_butlast) + by (metis EvExt_tF butlast_append butlast_snoc non_tickFree_tick tickFree_append) +qed + + +(* TODO: Renaming of Hiding and Renaming of Sync *) + + + + + + + + + + + (* ************************************************************************* *) (* *) (* Setup for rewriting *) (* *) (* ************************************************************************* *) lemmas Sync_rules = prefix_Sync2 read_read_Sync read_read_non_Sync write_read_Sync write_read_non_Sync write0_read_non_Sync write0_write_non_Sync lemmas Hiding_rules = no_Hiding_read no_Hiding_write Hiding_write Hiding_write0 lemmas mono_rules = mono_read_ref mono_write_ref mono_write0_ref end diff --git a/thys/HOL-CSP/Conclusion.thy b/thys/HOL-CSP/Conclusion.thy --- a/thys/HOL-CSP/Conclusion.thy +++ b/thys/HOL-CSP/Conclusion.thy @@ -1,114 +1,117 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * * Author : Burkhart Wolff, Safouan Taha. * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) chapter\Conclusion\ (*<*) theory Conclusion imports Assertions begin (*>*) section\Related Work\ text\As mentioned earlier, this work has its very ancient roots in a first formalization of A. Camilieri in the early 90s in HOL. This work was reformulated and substantially extended in HOL-CSP 1.0 published in 1997. In 2005, Roggenbach and Isobe published CSP-Prover, a formal theory of a (fragment of) the Failures model of CSP. This work led to a couple of publications culminating in @{cite "IsobeRoggenbach2010"}; emphasis was put on actually completing the CSP theory up to the point where it is sufficiently tactically supported to serve as a kind of tool. This theory is still maintained and last releases (the latest one was released on 18 February 2019) can be found under \<^url>\https://staff.aist.go.jp/y-isobe/CSP-Prover/CSP-Prover.html\. This theory represents the first half of Roscoes theory of a Failures/Divergence model, i.e. the Failures part. More recently, Pasquale Noce @{cite "Noninterference_CSP-AFP" and "Noninterference_Sequential_Composition-AFP" and "Noninterference_Concurrent_Composition-AFP"} developed a theory of non-interference notions based on an abstract denotational model fragment of the Failure/Divergence Model of CSP (without continuity and algebraic laws); this theory could probably be rebuilt on top of our work. The present work could be another, more "classic" foundation of test-generation techniques of this kind paving the way to an interaction with FDR and its possibility to generate labelled transition systems as output that could drive specialized tactics in HOL-CSP 2.0. \ section\Lessons learned\ text\We have ported a first formalization in Isabelle/HOL on the Failure/Divergence model of CSP, done with Isabelle93-7 in 1997, to a modern Isabelle version. Particularly, we use the modern declarative proof style available in Isabelle/Isar instead of imperative proof one, the latter being used in the old version. On the one hand, it is worth noting that some of the old theories still have a surprisingly high value: Actually it took time to develop the right granularity of abstraction in lemmas, which is thus still quite helpful and valuable to reconstruct the theory in the new version. If a substantially large body of lemmas is available, the degree of automation tends to increase. On the other hand, redevelopment from scratch is unavoidable in parts where basic libraries change. For example, this was a necessary consequence of our decision to base HOL-CSP 2.0 on HOLCF instead of continuing the development of an older fixed-point theory; nearly all continuity proofs had to be redeveloped. Moreover, a fresh look on old proof-obligations may incite unexpected generalizations and some newly proved lemmas that cannot be constructed in the old version even with several attempts. The influence of the chosen strategy (from scratch or refactoring) on the proof length is inconclusive. Note that our data does not allow to make a prediction on the length of a porting project --- the effort was distributed over a too long period and performed by a team with initially very different knowledge about CSP and interactive theorem proving. \ section\A Summary on New Results\ text\Compared to the original version of HOL-CSP 1.0, the present theory is complete relative to Roscoe's Book@{cite "roscoe:csp:1998" }. It contains a number of new theorems and some interesting (and unexpected) generalizations: \<^enum> @{thm mono_Hiding} is now also valid for the infinite case (arbitrary hide-set A). \<^enum> @{term \P \ (A \ B) = (P \ A) \ B\} is true for @{term "finite A"} (see @{thm Hiding_Un}); this was not even proven in HOL-CSP 1.0 for the singleton case! It can be considered as the most complex theorem of this theory. -\<^enum> distribution laws of hiding over synchronisation @{thm Hiding_Sync}; +\<^enum> distribution laws of \<^const>\Hiding\ over \<^const>\Sync\ @{thm Hiding_Sync}; however, this works only in the finite case. A true monster proof. +\<^enum> distribution of \<^const>\Mprefix\ over \<^const>\Sync\ @{thm Mprefix_Sync_distr} in the most + generalized case. Also a true monster proof, but reworked using symmetries and + abstractions to be more reasonable (and faster) \<^enum> the synchronization operator is associative @{thm "Sync_assoc"}. (In HOL-CSP 1.0, this had only be shown for special cases like @{thm "Par_assoc"}). \<^enum> the generalized non-deterministic choice operator --- relevant for proofs of deadlock-freeness --- has been added to the theory @{thm "Mndetprefix_def"}; it is proven monotone in the general case and continuous for the special case @{thm Mndetprefix_cont} relevant for the definition of the deadlock reference processes @{thm DF_def} and @{thm "DF\<^sub>S\<^sub>K\<^sub>I\<^sub>P_def"}.\ end diff --git a/thys/HOL-CSP/CopyBuffer.thy b/thys/HOL-CSP/CopyBuffer.thy --- a/thys/HOL-CSP/CopyBuffer.thy +++ b/thys/HOL-CSP/CopyBuffer.thy @@ -1,268 +1,269 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * * Author : Burkhart Wolff, Safouan Taha, Lina Ye. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) chapter\ Annex: Refinement Example with Buffer over infinite Alphabet\ -theory CopyBuffer - imports "Assertions" +theory CopyBuffer + imports CSP "Assertions" begin section\ Defining the Copy-Buffer Example \ + datatype 'a channel = left 'a | right 'a | mid 'a | ack definition SYN :: "('a channel) set" where "SYN \ (range mid) \ {ack}" definition COPY :: "('a channel) process" where "COPY \ (\ COPY. left\<^bold>?x \ (right\<^bold>!x \ COPY))" definition SEND :: "('a channel) process" where "SEND \ (\ SEND. left\<^bold>?x \ (mid\<^bold>!x \( ack \ SEND)))" definition REC :: "('a channel) process" where "REC \ (\ REC. mid\<^bold>?x \ (right\<^bold>!x \ (ack \ REC)))" definition SYSTEM :: "('a channel) process" where \SYSTEM \ (SEND \ SYN \ REC) \ SYN\ thm SYSTEM_def section\ The Standard Proof \ subsection\ Channels and Synchronization Sets\ text\ First part: abstract properties for these events to SYN. This kind of stuff could be automated easily by some extra-syntax for channels and SYN-sets. \ lemma [simp]: "left x \ SYN" by(auto simp: SYN_def) lemma [simp]: "right x \ SYN" by(auto simp: SYN_def) lemma [simp]: "ack \ SYN" by(auto simp: SYN_def) lemma [simp]: "mid x \ SYN" by(auto simp: SYN_def) lemma [simp]: "inj mid" by(auto simp: inj_on_def) lemma "finite (SYN:: 'a channel set) \ finite {(t::'a). True}" by (metis (no_types) SYN_def UNIV_def channel.inject(3) finite_Un finite_imageD inj_on_def) subsection\ Definitions by Recursors \ text\ Second part: Derive recursive process equations, which are easier to handle in proofs. This part IS actually automated if we could reuse the fixrec-syntax below. \ lemma COPY_rec: "COPY = left\<^bold>?x \ right\<^bold>!x \ COPY" by(simp add: COPY_def,rule trans, rule fix_eq, simp) lemma SEND_rec: "SEND = left\<^bold>?x \ mid\<^bold>!x \ (ack \ SEND)" by(simp add: SEND_def,rule trans, rule fix_eq, simp) lemma REC_rec: "REC = mid\<^bold>?x \ right\<^bold>!x \ (ack \ REC)" by(simp add: REC_def,rule trans, rule fix_eq, simp) subsection\ A Refinement Proof \ text\ Third part: No comes the proof by fixpoint induction. Not too bad in automation considering what is inferred, but wouldn't scale for large examples. \ lemma impl_refines_spec : "COPY \\<^sub>F\<^sub>D SYSTEM" apply(simp add: SYSTEM_def COPY_def) apply(rule fix_ind, simp_all) apply (intro le_FD_adm, simp_all add: cont_fun monofunI) apply (subst SEND_rec, subst REC_rec) by (simp add: Sync_rules Hiding_rules) lemma spec_refines_impl : assumes fin: "finite (SYN:: 'a channel set)" shows "SYSTEM \\<^sub>F\<^sub>D (COPY :: 'a channel process)" apply(simp add: SYSTEM_def SEND_def) apply(rule fix_ind, simp_all) apply (intro le_FD_adm) apply (simp add: fin) apply (simp add: cont2mono) apply (simp add: Hiding_set_BOT Sync_BOT Sync_commute) apply (subst COPY_rec, subst REC_rec) by (simp add: Sync_rules Hiding_rules) text\ Note that this was actually proven for the Process ordering, not the refinement ordering. But the former implies the latter. And due to anti-symmetry, equality follows for the case of finite alphabets \ldots \ lemma spec_equal_impl : assumes fin: "finite (SYN::('a channel)set)" shows "SYSTEM = (COPY::'a channel process)" by (simp add: FD_antisym fin impl_refines_spec spec_refines_impl) subsection\Deadlock Freeness Proof \ text\HOL-CSP can be used to prove deadlock-freeness of processes with infinite alphabet. In the case of the @{term COPY} - process, this can be formulated as the following refinement problem:\ -lemma "(DF (range left \ range right)) \\<^sub>F\<^sub>D COPY" +lemma DF_COPY : "(DF (range left \ range right)) \\<^sub>F\<^sub>D COPY" apply(simp add:DF_def,rule fix_ind2) proof - show "adm (\a. a \\<^sub>F\<^sub>D COPY)" by(rule le_FD_adm, simp_all add: monofunI) next show "\ \\<^sub>F\<^sub>D COPY" by fastforce next have 1: "(\xa\ range left \ range right \ \) \\<^sub>F\<^sub>D (\xa\ range left \ \)" by (rule mono_Mndetprefix_FD_set, simp, blast) have 2: "(\xa\ range left \ \) \\<^sub>F\<^sub>D (left\<^bold>?x \ \)" unfolding read_def by (meson Mprefix_refines_Mndetprefix_FD BOT_leFD mono_Mndetprefix_FD trans_FD) show "(\ x. \xa\range left \ range right \ x)\\ \\<^sub>F\<^sub>D COPY" by simp (metis (mono_tags, lifting) 1 2 COPY_rec mono_read_FD BOT_leFD trans_FD) next fix P::"'a channel process" assume *: "P \\<^sub>F\<^sub>D COPY" and ** : "(\ x. \xa\range left \ range right \ x)\P \\<^sub>F\<^sub>D COPY" have 1:"(\xa\ range left \ range right \ P) \\<^sub>F\<^sub>D (\xa\ range right \ P)" by (rule mono_Mndetprefix_FD_set, simp, blast) have 2:"(\xa\ range right \ P) \\<^sub>F\<^sub>D (right\<^bold>!x \ P)" for x apply (unfold write_def, rule trans_FD[OF mono_Mndetprefix_FD_set[of \{right x}\ \range right\]]) by simp_all from 1 2 have ab:"(\xa\ range left \ range right \ P) \\<^sub>F\<^sub>D (right\<^bold>!x \ P)" for x using trans_FD by blast hence 3:"(left\<^bold>?x \ (\xa\ range left \ range right \ P)) \\<^sub>F\<^sub>D (left\<^bold>?x \(right\<^bold>!x \ P))" by simp have 4:"\X. (\xa\ range left \ range right \ X) \\<^sub>F\<^sub>D (\xa\ range left \ X)" by (rule mono_Mndetprefix_FD_set, simp, blast) have 5:"\X. (\xa\ range left \ X) \\<^sub>F\<^sub>D (left\<^bold>?x \ X)" by (unfold read_def, subst K_record_comp, fact Mprefix_refines_Mndetprefix_FD) from 3 4[of "(\xa\ range left \ range right \ P)"] 5 [of "(\xa\ range left \ range right \ P)"] have 6:"(\xa\ range left \ range right \ (\xa\ range left \ range right \ P)) \\<^sub>F\<^sub>D (left\<^bold>?x \ (right\<^bold>!x \ P))" using trans_FD by blast from * ** have 7:"left\<^bold>?x \ right\<^bold>!x \ P \\<^sub>F\<^sub>D left\<^bold>?x \ right\<^bold>!x \ COPY" by simp show "(\ x. \xa\range left \ range right \ x)\ ((\ x. \xa\range left \ range right \ x)\P) \\<^sub>F\<^sub>D COPY" by simp (metis (mono_tags, lifting) "6" "7" COPY_rec trans_FD) qed section\ An Alternative Approach: Using the fixrec-Package \ subsection\ Channels and Synchronisation Sets \ text\ As before. \ subsection\ Process Definitions via fixrec-Package \ fixrec COPY' :: "('a channel) process" and SEND' :: "('a channel) process" and REC' :: "('a channel) process" where COPY'_rec[simp del]: "COPY' = left\<^bold>?x \ right\<^bold>!x \ COPY'" | SEND'_rec[simp del]: "SEND' = left\<^bold>?x \ mid\<^bold>!x \ (ack \ SEND')" | REC'_rec[simp del] : "REC' = mid\<^bold>?x \ right\<^bold>!x \ (ack \ REC')" thm COPY'_rec definition SYSTEM' :: "('a channel) process" where \SYSTEM' \ ((SEND' \ SYN \ REC') \ SYN)\ subsection\ Another Refinement Proof on fixrec-infrastructure \ text\ Third part: No comes the proof by fixpoint induction. Not too bad in automation considering what is inferred, but wouldn't scale for large examples. \ thm COPY'_SEND'_REC'.induct lemma impl_refines_spec' : "(COPY'::'a channel process) \\<^sub>F\<^sub>D SYSTEM'" apply (unfold SYSTEM'_def) apply (rule_tac P=\\ a b c. a \\<^sub>F\<^sub>D ((SEND' \SYN\ REC') \ SYN)\ in COPY'_SEND'_REC'.induct) apply (subst case_prod_beta')+ apply (intro le_FD_adm, simp_all add: monofunI) apply (subst SEND'_rec, subst REC'_rec) by (simp add: Sync_rules Hiding_rules) lemma spec_refines_impl' : assumes fin: "finite (SYN::('a channel)set)" shows "SYSTEM' \\<^sub>F\<^sub>D (COPY'::'a channel process)" proof(unfold SYSTEM'_def, rule_tac P=\\ a b c. ((b \SYN\ REC') \ SYN) \\<^sub>F\<^sub>D COPY'\ in COPY'_SEND'_REC'.induct, goal_cases) case 1 have aa:\adm (\(a::'a channel process). ((a \SYN\ REC') \ SYN) \\<^sub>F\<^sub>D COPY')\ apply (intro le_FD_adm) by (simp_all add: fin cont2mono) thus ?case using adm_subst[of "\(a,b,c). b", simplified, OF aa] by (simp add: split_def) next case 2 then show ?case by (simp add: Hiding_set_BOT Sync_BOT Sync_commute) next case (3 a aa b) then show ?case apply (subst COPY'_rec, subst REC'_rec) by (simp add: Sync_rules Hiding_rules) qed lemma spec_equal_impl' : assumes fin: "finite (SYN::('a channel)set)" shows "SYSTEM' = (COPY'::'a channel process)" apply (rule FD_antisym) apply (rule spec_refines_impl'[OF fin]) apply (rule impl_refines_spec') done end diff --git a/thys/HOL-CSP/Det.thy b/thys/HOL-CSP/Det.thy --- a/thys/HOL-CSP/Det.thy +++ b/thys/HOL-CSP/Det.thy @@ -1,298 +1,279 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * - * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * Author : Burkhart Wolff, Safouan Taha. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) section\ Deterministic Choice Operator Definition \ theory Det imports Process begin -subsection\Definition\ -definition +subsection\The Det Operator Definition\ +lift_definition Det :: "['\ process,'\ process] \ '\ process" (infixl "[+]" 79) -where "P [+] Q \ Abs_process( {(s,X). s = [] \ (s,X) \ \ P \ \ Q} - \ {(s,X). s \ [] \ (s,X) \ \ P \ \ Q} - \ {(s,X). s = [] \ s \ \ P \ \ Q} - \ {(s,X). s = [] \ tick \ X \ [tick] \ \ P \ \ Q}, - \ P \ \ Q)" + is "\P Q. ( {(s,X). s = [] \ (s,X) \ \ P \ \ Q} + \ {(s,X). s \ [] \ (s,X) \ \ P \ \ Q} + \ {(s,X). s = [] \ s \ \ P \ \ Q} + \ {(s,X). s = [] \ tick \ X \ [tick] \ \ P \ \ Q}, + \ P \ \ Q)" +proof - + show \is_process ( {(s,X). s = [] \ (s,X) \ \ (P ::'\ process) \ \ Q} + \ {(s,X). s \ [] \ (s,X) \ \ P \ \ Q} + \ {(s,X). s = [] \ s \ \ P \ \ Q} + \ {(s,X). s = [] \ tick \ X \ [tick] \ \ P \ \ Q}, + \ P \ \ Q)\ (is \is_process (?f, ?d)\) for P Q + proof (simp only: fst_conv snd_conv Rep_process is_process_def + DIVERGENCES_def FAILURES_def, intro conjI) + show "([], {}) \ ?f" + by(simp add: is_processT1) + next + show "\s X. (s, X) \ ?f \ front_tickFree s" + by(auto simp: is_processT2) + next + show "\s t. (s @ t, {}) \ ?f \ (s, {}) \ ?f" + by(auto simp: is_processT1 dest!: is_processT3[rule_format]) + next + show "\s X Y. (s, Y) \ ?f \ X \ Y \ (s, X) \ ?f" + by(auto dest: is_processT4[rule_format,OF conj_commute[THEN iffD1],OF conjI]) + next + show "\sa X Y. (sa, X) \ ?f \ (\c. c \ Y \ (sa @ [c], {}) \ ?f) \ (sa, X \ Y) \ ?f" + by(auto simp: is_processT5 T_F) + next + show " \s X. (s @ [tick], {}) \ ?f \ (s, X - {tick}) \ ?f" + apply((rule allI)+, rule impI, rename_tac s X) + apply(case_tac "s=[]", simp_all add: is_processT6[rule_format] T_F_spec) + by(auto simp: is_processT6[rule_format] T_F_spec) + next + show "\s t. s \ ?d \ tickFree s \ front_tickFree t \ s @ t \ ?d" + by(auto simp: is_processT7) + next + show "\s X. s \ ?d \ (s, X) \ ?f" + by(auto simp: is_processT8[rule_format]) + next + show "\s. s @ [tick] \ ?d \ s \ ?d" + by(auto intro!:is_processT9[rule_format]) + qed +qed notation Det (infixl "\" 79) term \(A \ B) \ D' = C\ -lemma is_process_REP_Det: -"is_process ({(s, X). s = [] \ (s, X) \ \ P \ \ Q} \ - {(s, X). s \ [] \ (s, X) \ \ P \ \ Q} \ - {(s, X). s = [] \ s \ \ P \ \ Q} \ - {(s, X). s = [] \ tick \ X \ [tick] \ \ P \ \ Q}, - \ P \ \ Q)" -(is "is_process(?T,?D)") -proof (simp only: fst_conv snd_conv Rep_process is_process_def - DIVERGENCES_def FAILURES_def, intro conjI) - show "([], {}) \ ?T" - by(simp add: is_processT1) -next - show "\s X. (s, X) \ ?T \ front_tickFree s" - by(auto simp: is_processT2) -next - show "\s t. (s @ t, {}) \ ?T \ (s, {}) \ ?T" - by(auto simp: is_processT1 dest!: is_processT3[rule_format]) -next - show "\s X Y. (s, Y) \ ?T \ X \ Y \ (s, X) \ ?T" - by(auto dest: is_processT4[rule_format,OF conj_commute[THEN iffD1],OF conjI]) -next - show "\sa X Y. (sa, X) \ ?T \ (\c. c \ Y \ (sa @ [c], {}) \ ?T) \ (sa, X \ Y) \ ?T" - by(auto simp: is_processT5 T_F) -next - show " \s X. (s @ [tick], {}) \ ?T \ (s, X - {tick}) \ ?T" - apply((rule allI)+, rule impI, rename_tac s X) - apply(case_tac "s=[]", simp_all add: is_processT6[rule_format] T_F_spec) - by(auto simp: is_processT6[rule_format] T_F_spec) -next - show "\s t. s \ ?D \ tickFree s \ front_tickFree t \ s @ t \ ?D" - by(auto simp: is_processT7) -next - show "\s X. s \ ?D \ (s, X) \ ?T" - by(auto simp: is_processT8[rule_format]) -next - show "\s. s @ [tick] \ ?D \ s \ ?D" - by(auto intro!:is_processT9[rule_format]) -qed - - -lemma Rep_Abs_Det: -"Rep_process - (Abs_process - ({(s, X). s = [] \ (s, X) \ \ P \ \ Q} \ - {(s, X). s \ [] \ (s, X) \ \ P \ \ Q} \ - {(s, X). s = [] \ s \ \ P \ \ Q} \ - {(s, X). s = [] \ tick \ X \ [tick] \ \ P \ \ Q}, - \ P \ \ Q)) = - ({(s, X). s = [] \ (s, X) \ \ P \ \ Q} \ - {(s, X). s \ [] \ (s, X) \ \ P \ \ Q} \ - {(s, X). s = [] \ s \ \ P \ \ Q} \ - {(s, X). s = [] \ tick \ X \ [tick] \ \ P \ \ Q}, - \ P \ \ Q)" - by(simp only: CollectI Rep_process Abs_process_inverse is_process_REP_Det) - subsection\The Projections\ lemma F_Det : "\ (P \ Q) = {(s,X). s = [] \ (s,X) \ \ P \ \ Q} \ {(s,X). s \ [] \ (s,X) \ \ P \ \ Q} \ {(s,X). s = [] \ s \ \ P \ \ Q} \ {(s,X). s = [] \ tick \ X \ [tick] \ \ P \ \ Q}" - by(subst Failures_def, simp only: Det_def Rep_Abs_Det FAILURES_def fst_conv) + by (subst Failures.rep_eq, simp add: Det.rep_eq FAILURES_def) +lemma D_Det: "\ (P \ Q) = \ P \ \ Q" + by (subst Divergences.rep_eq, simp add: Det.rep_eq DIVERGENCES_def) + +lemma T_Det: "\ (P \ Q) = \ P \ \ Q" + apply (subst Traces.rep_eq, simp add: TRACES_def Failures.rep_eq[symmetric] F_Det) + apply(auto simp: T_F F_T is_processT1 Nil_elem_T) + by(rule_tac x="{}" in exI, simp add: T_F F_T is_processT1 Nil_elem_T)+ + subsection\Basic Laws\ text\The following theorem of Commutativity helps to simplify the subsequent continuity proof by symmetry breaking. It is therefore already developed here:\ -lemma D_Det: "\ (P \ Q) = \ P \ \ Q" - by(subst D_def, simp only: Det_def Rep_Abs_Det DIVERGENCES_def snd_conv) - - -lemma T_Det: "\ (P \ Q) = \ P \ \ Q" - apply(subst Traces_def, simp only: Det_def Rep_Abs_Det TRACES_def FAILURES_def - fst_def snd_conv) - apply(auto simp: T_F F_T is_processT1 Nil_elem_T) - apply(rule_tac x="{}" in exI, simp add: T_F F_T is_processT1 Nil_elem_T)+ - done lemma Det_commute:"(P \ Q) = (Q \ P)" by(auto simp: Process_eq_spec F_Det D_Det) subsection\The Continuity-Rule\ lemma mono_Det1: "P \ Q \ \ (Q \ S) \ \ (P \ S)" apply (drule le_approx1) by (auto simp: Process_eq_spec F_Det D_Det) lemma mono_Det2: assumes ordered: "P \ Q" shows "(\ s. s \ \ (P \ S) \ Ra (P \ S) s = Ra (Q \ S) s)" proof - have A:"\s t. [] \ \ P \ [] \ \ S \ s = [] \ ([], t) \ \ P \ ([], t) \ \ S \ ([], t) \ \ Q" by (insert ordered,frule_tac X="t" and s="[]" in proc_ord2a, simp_all) have B:"\s t. s \ \ P \ s \ \ S \ (s \ [] \ ((s, t) \ \ P \ (s, t) \ \ S) \ (s, t) \ \ Q \ (s, t) \ \ S) \ (s = [] \ tick \ t \ ([tick] \ \ P \ [tick] \ \ S) \ ([], t) \ \ Q \ ([], t) \ \ S \ [] \ \ Q \ [tick] \ \ Q \ [tick] \ \ S)" apply(intro conjI impI, elim conjE disjE, rule disjI1) apply(simp_all add: proc_ord2a[OF ordered,symmetric]) apply(elim conjE disjE,subst le_approx2T[OF ordered]) apply(frule is_processT9_S_swap, simp_all) done have C: "\s. s \ \ P \ s \ \ S \ {X. s = [] \ (s, X) \ \ Q \ (s, X) \ \ S \ s \ [] \ ((s, X) \ \ Q \ (s, X) \ \ S) \ s = [] \ s \ \ Q \ s = [] \ tick \ X \ ([tick] \ \ Q \ [tick] \ \ S)} \ {X. s = [] \ (s, X) \ \ P \ (s, X) \ \ S \ s \ [] \ ((s, X) \ \ P \ (s, X) \ \ S) \ s = [] \ tick \ X \ ([tick] \ \ P \ [tick] \ \ S)}" apply(intro subsetI, frule is_processT9_S_swap, simp) apply(elim conjE disjE, simp_all add: proc_ord2a[OF ordered,symmetric] is_processT8_S) apply(insert le_approx1[OF ordered] le_approx_lemma_T[OF ordered]) by (auto simp: proc_ord2a[OF ordered,symmetric]) show ?thesis apply(simp add: Process_eq_spec F_Det D_Det Ra_def min_elems_def) apply(intro allI impI equalityI C, simp_all) apply(intro allI impI subsetI, simp) apply(metis A B) done qed lemma mono_Det3: "P \ Q \ min_elems (\ (P \ S)) \ \ (Q \ S)" apply (frule le_approx3) apply (simp add: Process_eq_spec F_Det T_Det D_Det Ra_def min_elems_def subset_iff) apply (auto intro:D_T) done lemma mono_Det[simp] : "P \ Q \ (P \ S) \ (Q \ S)" by (auto simp: le_approx_def mono_Det1 mono_Det2 mono_Det3) lemma mono_Det_sym[simp] : "P \ Q \ (S \ P) \ (S \ Q)" by (simp add: Det_commute) lemma cont_Det0: assumes C : "chain Y" shows "(lim_proc (range Y) \ S) = lim_proc (range (\i. Y i \ S))" proof - have C':"chain (\i. Y i \ S)" by(auto intro!:chainI simp:chainE[OF C]) show ?thesis apply (subst Process_eq_spec) apply (simp add: D_Det F_Det) apply(simp add: F_LUB[OF C] D_LUB[OF C] T_LUB[OF C] F_LUB[OF C'] D_LUB[OF C'] T_LUB[OF C']) apply(safe) apply(auto simp: D_Det F_Det T_Det) using NF_ND apply blast using is_processT6_S2 is_processT8 apply blast apply (metis D_T append_Nil front_tickFree_single process_charn tickFree_Nil) using NF_ND is_processT6_S2 apply blast using NF_ND is_processT6_S2 by blast qed lemma cont_Det: assumes C: "chain Y" shows " ((\ i. Y i) \ S) = (\ i. (Y i \ S))" apply(insert C) apply(subst limproc_is_thelub, simp) apply(subst limproc_is_thelub) apply(rule chainI) apply(frule_tac i="i" in chainE) apply(simp) apply(erule cont_Det0) done lemma cont_Det': assumes chain:"chain Y" shows "((\ i. Y i) \ S) = (\ i. (Y i \ S))" proof - have chain':"chain (\i. Y i \ S)" by(auto intro!:chainI simp: chainE[OF chain]) have B:"\ (lim_proc (range Y) \ S) \ \ (lim_proc (range (\i. Y i \ S)))" apply(simp add: D_Det F_Det F_LUB T_LUB D_LUB chain chain') apply(intro conjI subsetI, simp_all) by(auto split:prod.split prod.split_asm) have C:"\ (lim_proc (range (\i. Y i \ S))) \ \(lim_proc (range Y) \ S)" proof - have C1 : "\ba ab ac. \a. ([], ba) \ \ (Y a) \ ([], ba) \ \ S \ [] \ \ (Y a) \ [] \ \ (Y ab) \ [] \ \ S \ tick \ ba \ ([], ba) \ \ (Y ac) " using is_processT8 by auto have C2 : "\ba ab ac ad D. \a. ([], ba) \ \ (Y a) \ ([], ba) \ \ S \ [] \ \ (Y a) \ tick \ ba \ [tick] \ \ (Y a) \ []\D(Y ab) \ []\D S \ ([],ba)\ \(Y ac) \ [tick] \ \ S \ [tick]\\(Y ad)" using NF_ND is_processT6_S2 by blast have C3: "\ba ab ac. \a. [] \ \ (Y a) \ tick \ ba \ [tick] \ \ (Y a) \ [] \ \ (Y ab) \ [] \ \ S \ ([], ba) \ \ S \ [tick] \ \ S \ [tick] \ \ (Y ac)" by (metis D_T append_Nil front_tickFree_single process_charn tickFree_Nil) show ?thesis apply(simp add: D_Det F_Det F_LUB T_LUB D_LUB chain chain') apply(rule subsetI, simp) apply(simp split:prod.split prod.split_asm) apply(intro allI impI,simp) by (metis C3 is_processT6_S2 is_processT8_S) qed have D:"\ (lim_proc (range Y) \ S) = \ (lim_proc (range (\i. Y i \ S)))" by(simp add: D_Det F_Det F_LUB T_LUB D_LUB chain chain') show ?thesis by(simp add: chain chain' limproc_is_thelub Process_eq_spec equalityI B C D) qed lemma Det_cont[simp]: assumes f:"cont f" and g:"cont g" shows "cont (\x. f x \ g x)" proof - have A : "\x. cont f \ cont (\y. y \ f x)" apply (rule contI2,rule monofunI) apply (auto) apply (subst Det_commute,subst cont_Det) apply (auto simp: Det_commute) done have B : "\y. cont f \ cont (\x. y \ f x)" apply (rule_tac c="(\ x. y \ x)" in cont_compose) apply (rule contI2,rule monofunI) apply (auto) apply (subst Det_commute,subst cont_Det) by (simp_all add: Det_commute) show ?thesis using f g apply(rule_tac f="(\x. (\ g. f x \ g))" in cont_apply) apply(auto intro:contI2 monofunI simp:Det_commute A B) done qed end \ No newline at end of file diff --git a/thys/HOL-CSP/Hiding.thy b/thys/HOL-CSP/Hiding.thy --- a/thys/HOL-CSP/Hiding.thy +++ b/thys/HOL-CSP/Hiding.thy @@ -1,660 +1,653 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * - * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * Author : Burkhart Wolff, Safouan Taha. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) chapter\Concurrent CSP Operators\ section\The Hiding Operator\ theory Hiding imports Process begin subsection\Preliminaries : primitives and lemmas\ abbreviation "trace_hide t A \ filter (\x. x \ A) t" lemma Hiding_tickFree : "tickFree s \ tickFree (trace_hide s (ev`A))" by (auto simp add: tickFree_def) lemma Hiding_fronttickFree : "front_tickFree s \ front_tickFree (trace_hide s (ev`A))" apply(auto simp add: front_tickFree_charn tickFree_def split:if_splits) by (metis Hiding_tickFree list_nonMt_append tickFree_append tickFree_def) lemma trace_hide_union[simp] : "trace_hide t (ev ` (A \ B)) = trace_hide (trace_hide t (ev ` A)) (ev ` B)" by (subgoal_tac "ev ` (A \ B) = (ev ` A) \ (ev ` B)") auto abbreviation "isInfHiddenRun f P A \ strict_mono f \ (\ i. f i \ \ P) \ (\ i. trace_hide (f i) (ev ` A) = trace_hide (f (0::nat)) (ev ` A))" lemma isInfHiddenRun_1: "isInfHiddenRun f P A \ strict_mono f \ (\ i. f i \ \ P) \ (\i. \t. f i = f 0 @ t \ set t \ (ev ` A))" (is "?A \ ?B") proof assume A: ?A { fix i from A have "f 0 \ f i" using strict_mono_less_eq by blast then obtain t where B:"f i = f 0 @ t" by (metis le_list_def) hence "trace_hide (f i) (ev ` A) = (trace_hide (f 0) (ev ` A)) @ (trace_hide t (ev ` A))" by simp with A have "trace_hide t (ev ` A) = []" by (metis append_self_conv) with B have "\t. f i = f 0 @ t \ set t \ (ev ` A)" using filter_empty_conv[of "\x. x \ (ev ` A)"] by auto } with A show ?B by simp next assume B: ?B { fix i from B obtain t where B:"f i = f 0 @ t" and "set t \ (ev ` A)" by blast hence "trace_hide (f i) (ev ` A) = (trace_hide (f 0) (ev ` A))" by (simp add: subset_iff) } with B show ?A by simp qed -subsection\The Hiding Operator Definition\ abbreviation "div_hide P A \ {s. \ t u. front_tickFree u \ tickFree t \ s = trace_hide t (ev ` A) @ u \ (t \ \ P \ (\ f. isInfHiddenRun f P A \ t \ range f))}" -definition Hiding :: "['\ process ,'\ set] => '\ process" (\_ \ _\ [57,56] 57) where - \P \ A \ Abs_process ({(s,X). \ t. s = trace_hide t (ev`A) \ (t,X \ (ev`A)) \ \ P} \ - {(s,X). s \ div_hide P A}, div_hide P A)\ - lemma inf_hidden: assumes as1:"\t. trace_hide t (ev ` A) = trace_hide s (ev ` A) \ (t, ev ` A) \ \ P" and as2:"s \ \ P" shows "\f. isInfHiddenRun f P A \ s \ range f" proof define f where A:"f = rec_nat s (\i t. (let e = SOME e. e \ ev`A \ t @ [e] \ \ P in t @ [e]))" hence B:"strict_mono f" by (simp add:strict_mono_def lift_Suc_mono_less_iff) from A have C:"s \ range f" by (metis (mono_tags, lifting) old.nat.simps(6) range_eqI) { fix i have "f i \ \ P \ trace_hide (f i) (ev ` A) = (trace_hide s (ev ` A))" proof(induct i, simp add: A Nil_elem_T as2) case (Suc i) with as1[THEN spec, of "f i"] have a:"\e. e \ ev`A \ f i @ [e] \ \ P" using is_processT5_S7 by force from A have b:"f (Suc i) = (let e = SOME e. e \ ev`A \ f i @ [e] \ \ P in f i @ [e])" by simp with Suc a[THEN someI_ex] show ?case by simp qed } with B C show "isInfHiddenRun f P A \ s \ range f" by simp qed -subsection\Consequences\ + lemma trace_hide_append: "s @ t = trace_hide ta (ev ` A) \ \ss tt. ta = ss@tt \ s = trace_hide ss (ev ` A) \ t = trace_hide tt (ev ` A)" proof(induct "ta" arbitrary:s t) case Nil thus ?case by simp next case (Cons a ta) thus ?case proof(cases "a \ (ev ` A)") case True hence "s @ t = trace_hide ta (ev ` A)" by (simp add: Cons) with Cons obtain ss tt where "ta = ss @ tt \ s = trace_hide ss (ev ` A) \ t = trace_hide tt (ev ` A)" by blast with True Cons show ?thesis by (rule_tac x="a#ss" in exI, rule_tac x=tt in exI, auto) next case False thus ?thesis proof(cases s) case Nil thus ?thesis using Cons by fastforce next case Cons2:(Cons aa tls) with False have A:"a = aa \ tls @ t = trace_hide ta (ev ` A)" using Cons by auto with Cons obtain ss tt where "ta = ss @ tt \ tls = trace_hide ss (ev ` A) \ t = trace_hide tt (ev ` A)" by blast with Cons2 A False show ?thesis by (rule_tac x="a#ss" in exI, rule_tac x=tt in exI, auto) qed qed qed -lemma Hiding_maintains_is_process: - "is_process ({(s,X). \ t. s = trace_hide t (ev`A) \ (t,X \ (ev`A)) \ \ P} \ - {(s,X). s \ div_hide P A}, div_hide P A)" (is "is_process(?f, ?d)") -proof (simp only: fst_conv snd_conv Failures_def is_process_def FAILURES_def DIVERGENCES_def, - fold FAILURES_def DIVERGENCES_def,fold Failures_def,intro conjI, goal_cases) - case 1 thus ?case - proof (auto, rule not_not[THEN iffD1], rule notI, simp, goal_cases) - case 1 - from inf_hidden[of A "[]" P] obtain f where A:"isInfHiddenRun f P A \ [] \ range f" - using "1"(2) Nil_elem_T by auto - from A 1(1)[THEN spec, of "[]"] filter.simps(1) tickFree_Nil show ?case by auto - qed -next - case 2 thus ?case - using is_processT2 Hiding_fronttickFree front_tickFree_append Hiding_tickFree by blast+ -next - case 3 thus ?case - proof(auto del:disjE, goal_cases) - case (1 s t ta) show ?case - proof(cases "tickFree s") - case True - from 1(2) obtain ss tt where A:"ta = ss@tt \ s = trace_hide ss (ev ` A) \ ss \ \ P" - using trace_hide_append[of s t A ta, OF 1(1)] by (metis F_T is_processT3_SR) - with True have B:"tickFree ss" - by (metis event.distinct(1) filter_set imageE member_filter tickFree_def) - show ?thesis - using 1(3) A B inf_hidden[of A ss P] by (metis append_Nil2 front_tickFree_Nil) - next - case False - with 1(1,2) obtain ss tt where A:"s = ss@[tick] \ ta = tt@[tick]" - by (metis append_Nil2 contra_subsetD filter_is_subset front_tickFree_mono - Hiding_fronttickFree is_processT nonTickFree_n_frontTickFree tickFree_def) - with 1(1,2) have "ss = trace_hide tt (ev ` A)" - by (metis (no_types, lifting) butlast_append butlast_snoc contra_subsetD - filter.simps(2) filter_append filter_is_subset front_tickFree_implies_tickFree - front_tickFree_single is_processT nonTickFree_n_frontTickFree non_tickFree_tick - self_append_conv2 tickFree_append tickFree_def) - with False 1(1,2) A show ?thesis - by (metis append_Nil2 front_tickFree_mono Hiding_fronttickFree is_processT) +subsection\The Hiding Operator Definition\ + +lift_definition Hiding :: "['\ process ,'\ set] => '\ process" (\_ \ _\ [57,56] 57) is + \\P A. ({(s,X). \ t. s = trace_hide t (ev`A) \ (t,X \ (ev`A)) \ \ P} \ + {(s,X). s \ div_hide P A}, div_hide P A)\ +proof - + show "is_process ({(s,X). \ t. s = trace_hide t (ev`(A :: '\ set)) \ (t,X \ (ev`A)) \ \ P} \ + {(s,X). s \ div_hide P A}, div_hide P A)" (is "is_process(?f, ?d)") for P A + proof (simp only: is_process_def FAILURES_def DIVERGENCES_def fst_conv snd_conv, + intro conjI, goal_cases) + case 1 thus ?case + proof (auto, rule not_not[THEN iffD1], rule notI, simp, goal_cases) + case 1 + from inf_hidden[of A "[]" P] obtain f where A:"isInfHiddenRun f P A \ [] \ range f" + using "1"(2) Nil_elem_T by auto + from A 1(1)[THEN spec, of "[]"] filter.simps(1) tickFree_Nil show ?case by auto qed next - case (2 s t ta u) show ?case - proof(cases "length s \ length (trace_hide ta (ev ` A))") - case True - with append_eq_append_conv2[THEN iffD1, OF 2(3)] - obtain tt where "s@tt = trace_hide ta (ev ` A)" by auto - with 2(4) obtain ss ttt where A:"ta = ss@ttt \ s = trace_hide ss (ev ` A) \ ss \ \ P" - using trace_hide_append[of s tt A ta] by (metis D_T imageE is_processT3_ST) - with 2(2) have B:"tickFree ss" using tickFree_append by blast - show ?thesis - using 2(4,5) A B inf_hidden[of A ss P] - by (metis (no_types, lifting) append_Nil2 is_processT) + case 2 thus ?case + using is_processT2 Hiding_fronttickFree front_tickFree_append Hiding_tickFree by blast+ + next + case 3 thus ?case + proof(auto del:disjE, goal_cases) + case (1 s t ta) show ?case + proof(cases "tickFree s") + case True + from 1(2) obtain ss tt where A:"ta = ss@tt \ s = trace_hide ss (ev ` A) \ ss \ \ P" + using trace_hide_append[of s t A ta, OF 1(1)] by (metis F_T is_processT3_SR) + with True have B:"tickFree ss" + by (metis event.distinct(1) filter_set imageE member_filter tickFree_def) + show ?thesis + using 1(3) A B inf_hidden[of A ss P] by (metis append_Nil2 front_tickFree_Nil) + next + case False + with 1(1,2) obtain ss tt where A:"s = ss@[tick] \ ta = tt@[tick]" + by (metis append_Nil2 contra_subsetD filter_is_subset front_tickFree_mono + Hiding_fronttickFree is_processT nonTickFree_n_frontTickFree tickFree_def) + with 1(1,2) have "ss = trace_hide tt (ev ` A)" + by (metis (no_types, lifting) butlast_append butlast_snoc contra_subsetD + filter.simps(2) filter_append filter_is_subset front_tickFree_implies_tickFree + front_tickFree_single is_processT nonTickFree_n_frontTickFree non_tickFree_tick + self_append_conv2 tickFree_append tickFree_def) + with False 1(1,2) A show ?thesis + by (metis append_Nil2 front_tickFree_mono Hiding_fronttickFree is_processT) + qed next - case False - with 2(3) obtain uu uuu where A:"s = trace_hide ta (ev ` A) @ uu \ u = uu @ uuu" - by (auto simp add: append_eq_append_conv2, metis le_length_mono le_list_def) - with 2(1,2,4) 2(5)[rule_format, of ta uu] show ?thesis - using front_tickFree_dw_closed by blast + case (2 s t ta u) show ?case + proof(cases "length s \ length (trace_hide ta (ev ` A))") + case True + with append_eq_append_conv2[THEN iffD1, OF 2(3)] + obtain tt where "s@tt = trace_hide ta (ev ` A)" by auto + with 2(4) obtain ss ttt where A:"ta = ss@ttt \ s = trace_hide ss (ev ` A) \ ss \ \ P" + using trace_hide_append[of s tt A ta] by (metis D_T imageE is_processT3_ST) + with 2(2) have B:"tickFree ss" using tickFree_append by blast + show ?thesis + using 2(4,5) A B inf_hidden[of A ss P] + by (metis (no_types, lifting) append_Nil2 is_processT) + next + case False + with 2(3) obtain uu uuu where A:"s = trace_hide ta (ev ` A) @ uu \ u = uu @ uuu" + by (auto simp add: append_eq_append_conv2, metis le_length_mono le_list_def) + with 2(1,2,4) 2(5)[rule_format, of ta uu] show ?thesis + using front_tickFree_dw_closed by blast + qed qed - qed -next - case 4 thus ?case - by (auto, metis (mono_tags) Un_subset_iff is_processT4 sup.cobounded2 sup.coboundedI1) -next - case 5 thus ?case - proof(intro impI allI, auto, rule not_not[THEN iffD1], rule notI, simp, goal_cases) - case (1 X Y t) - from 1(3) 1(4)[THEN spec, of t, simplified] obtain c where A1:"c \ Y" and A2:"c \ (ev`A)" - and A3:"t@[c] \ \ P" - using is_processT5_S7'[of t "X \ (ev`A)" P Y] by (metis UnCI sup_commute sup_left_commute) - hence "trace_hide t (ev ` A) @ [c] = trace_hide (t@[c]) (ev ` A)" by simp - thus ?case using 1(1)[rule_format, OF A1] inf_hidden[of A "t@[c]", rotated, OF A3] - by (metis (no_types, lifting) append.right_neutral append_T_imp_tickFree - front_tickFree_Nil is_processT5_S7 not_Cons_self2 rangeE) - qed -next - case 6 thus ?case - proof (auto, goal_cases) - case (1 s X t) - hence "front_tickFree t" by (simp add:is_processT2) - with 1(1) obtain t' where A:"t = t'@[tick]" - by (metis filter_is_subset nonTickFree_n_frontTickFree non_tickFree_tick - subset_iff tickFree_append tickFree_def) - with 1(1,2) have B:"s = trace_hide t' (ev ` A)" - by (auto simp add:tickFree_def split:if_splits) - with A 1(1,2) is_processT6[of P, THEN spec, THEN spec, of t' "X \ ev ` A"] - is_processT4_empty[of t "ev ` A" P] show ?case - by (auto simp add: Un_Diff split:if_splits) next - case (2 s X t u) - then obtain u' where A:"u = u'@[tick]" - by (metis filter_is_subset nonTickFree_n_frontTickFree non_tickFree_tick - subset_iff tickFree_append tickFree_def) - with 2(3) have B:"s = trace_hide t (ev ` A) @ u'" - by (auto simp add:tickFree_def split:if_splits) - with 2(1,2,5) 2(4)[THEN spec, THEN spec, of t u'] show ?case - using front_tickFree_dw_closed A by blast + case 4 thus ?case + by (auto, metis (mono_tags) Un_subset_iff is_processT4 sup.cobounded2 sup.coboundedI1) next - case (3 s X u f x) - then obtain u' where A:"u = u'@[tick]" - by (metis filter_is_subset nonTickFree_n_frontTickFree non_tickFree_tick - subset_iff tickFree_append tickFree_def) - with 3(3) have B:"s = trace_hide (f x) (ev ` A) @ u'" - by (auto simp add:tickFree_def split:if_splits) - with 3(1,2,3,5,6,7) 3(4)[THEN spec, THEN spec, of "f x" u'] show ?case - using front_tickFree_dw_closed[of u' "[tick]"] by auto - qed -next - case 7 thus ?case using front_tickFree_append by (auto, blast +) -next - case 8 thus ?case by simp -next - case 9 thus ?case - proof (intro allI impI, simp, elim exE, goal_cases) - case (1 s t u) - then obtain u' where "u = u'@[tick]" - by (metis Hiding_tickFree nonTickFree_n_frontTickFree non_tickFree_tick tickFree_append) - with 1 show ?case - apply(rule_tac x=t in exI, rule_tac x=u' in exI) - using front_tickFree_dw_closed by auto + case 5 thus ?case + proof(intro impI allI, auto, rule not_not[THEN iffD1], rule notI, simp, goal_cases) + case (1 X Y t) + from 1(3) 1(4)[THEN spec, of t, simplified] obtain c where A1:"c \ Y" and A2:"c \ (ev`A)" + and A3:"t@[c] \ \ P" + using is_processT5_S7'[of t "X \ (ev`A)" P Y] by (metis UnCI sup_commute sup_left_commute) + hence "trace_hide t (ev ` A) @ [c] = trace_hide (t@[c]) (ev ` A)" by simp + thus ?case using 1(1)[rule_format, OF A1] inf_hidden[of A "t@[c]", rotated, OF A3] + by (metis (no_types, lifting) append.right_neutral append_T_imp_tickFree + front_tickFree_Nil is_processT5_S7 not_Cons_self2 rangeE) + qed + next + case 6 thus ?case + proof (auto, goal_cases) + case (1 s X t) + hence "front_tickFree t" by (simp add:is_processT2) + with 1(1) obtain t' where A:"t = t'@[tick]" + by (metis filter_is_subset nonTickFree_n_frontTickFree non_tickFree_tick + subset_iff tickFree_append tickFree_def) + with 1(1,2) have B:"s = trace_hide t' (ev ` A)" + by (auto simp add:tickFree_def split:if_splits) + with A 1(1,2) is_processT6[of P, THEN spec, THEN spec, of t' "X \ ev ` A"] + is_processT4_empty[of t "ev ` A" P] show ?case + by (auto simp add: Un_Diff split:if_splits) + next + case (2 s X t u) + then obtain u' where A:"u = u'@[tick]" + by (metis filter_is_subset nonTickFree_n_frontTickFree non_tickFree_tick + subset_iff tickFree_append tickFree_def) + with 2(3) have B:"s = trace_hide t (ev ` A) @ u'" + by (auto simp add:tickFree_def split:if_splits) + with 2(1,2,5) 2(4)[THEN spec, THEN spec, of t u'] show ?case + using front_tickFree_dw_closed A by blast + next + case (3 s X u f x) + then obtain u' where A:"u = u'@[tick]" + by (metis filter_is_subset nonTickFree_n_frontTickFree non_tickFree_tick + subset_iff tickFree_append tickFree_def) + with 3(3) have B:"s = trace_hide (f x) (ev ` A) @ u'" + by (auto simp add:tickFree_def split:if_splits) + with 3(1,2,3,5,6,7) 3(4)[THEN spec, THEN spec, of "f x" u'] show ?case + using front_tickFree_dw_closed[of u' "[tick]"] by auto + qed + next + case 7 thus ?case using front_tickFree_append by (auto, blast +) + next + case 8 thus ?case by simp + next + case 9 thus ?case + proof (intro allI impI, simp, elim exE, goal_cases) + case (1 s t u) + then obtain u' where "u = u'@[tick]" + by (metis Hiding_tickFree nonTickFree_n_frontTickFree non_tickFree_tick tickFree_append) + with 1 show ?case + apply(rule_tac x=t in exI, rule_tac x=u' in exI) + using front_tickFree_dw_closed by auto + qed qed qed -lemma Rep_Abs_Hiding: "Rep_process (Abs_process - ({(s,X). \ t. s = trace_hide t (ev`A) \ (t,X \ (ev`A)) \ \ P} \ - {(s,X). s \ div_hide P A}, div_hide P A)) - = ({(s,X). \ t. s = trace_hide t (ev`A) \ (t,X \ (ev`A)) \ \ P} \ - {(s,X). s \ div_hide P A}, div_hide P A)" - by (simp only:CollectI Rep_process Abs_process_inverse Hiding_maintains_is_process) subsection\Projections\ lemma F_Hiding: \\ (P \ A) = {(s,X). \ t. s = trace_hide t (ev`A) \ (t,X \ (ev`A)) \ \ P} \ {(s,X). s \ div_hide P A}\ - by (subst Failures_def, simp only: Hiding_def Rep_Abs_Hiding FAILURES_def fst_conv) + by (simp add: Failures.rep_eq Hiding.rep_eq FAILURES_def) lemma D_Hiding: \\ (P \ A) = div_hide P A\ - by (subst D_def, simp only: Hiding_def Rep_Abs_Hiding DIVERGENCES_def snd_conv) + by (simp add: Divergences.rep_eq Hiding.rep_eq DIVERGENCES_def) lemma T_Hiding: \\ (P \ A) = {s. \t. s = trace_hide t (ev`A) \ (t, ev`A) \ \ P} \ div_hide P A\ - apply (unfold Traces_def, simp only:Rep_Abs_Hiding Hiding_def TRACES_def FAILURES_def fst_conv, auto) - apply (metis is_processT sup.cobounded2) - apply (metis FAILURES_def Failures_def NF_NT range_eqI) - apply (metis sup.idem) - by (metis FAILURES_def F_T Failures_def range_eqI) + by (auto simp add: Traces.rep_eq TRACES_def Failures.rep_eq[symmetric] F_Hiding) + (metis is_processT4 sup_ge2, metis sup_bot_left) + subsection\Continuity Rule\ lemma mono_Hiding[simp]: \(P::'a process) \ Q \ (P \ A) \ (Q \ A)\ proof (auto simp only:le_approx_def D_Hiding Ra_def F_Hiding T_Hiding, goal_cases) case (1 t u) thus ?case by blast next case (2 u f xa) thus ?case apply(rule_tac x="f xa" in exI, rule_tac x=u in exI) by (metis D_T Ra_def le_approx2T le_approx_def rangeI) next case (3 x t) hence A:"front_tickFree t" by (meson is_processT2) show ?case proof(cases "tickFree t") case True thus ?thesis by (metis "3"(2) "3"(4) "3"(5) "3"(6) front_tickFree_charn mem_Collect_eq self_append_conv) next case False with A obtain t' where "t = t'@[tick]" using nonTickFree_n_frontTickFree by blast with 3 show ?thesis by (metis (no_types, lifting) filter.simps(1) filter.simps(2) front_tickFree_mono front_tickFree_Nil filter_append is_processT9 list.distinct(1) local.A mem_Collect_eq) qed next case (4 x t) thus ?case using NF_ND by blast next case (5 x t u) thus ?case by blast next case (6 x u f xa) thus ?case by (metis D_T Ra_def le_approx2T le_approx_def rangeI) next case (7 x) from 7(4) have A:"x \ min_elems (div_hide P A)" by simp from elem_min_elems[OF 7(4), simplified] obtain t u where B1:"x = trace_hide t (ev ` A) @ u" and B2:"tickFree t" and B3:"front_tickFree u" and B4:"(t \ \ P \ (\(f:: nat \ 'a event list). strict_mono f \ (\i. f i \ \ P) \ (\i. trace_hide (f i) (ev ` A) = trace_hide (f 0) (ev ` A)) \ t \ range f))" by blast show ?case proof(cases "t \ \ P") case True then obtain t' t'' where C1:"t'@t''=t" and C2:"t' \ min_elems (\ P)" by (metis min_elems_charn) hence C3:"trace_hide t' (ev ` A) \ div_hide P A" apply (simp, rule_tac x=t' in exI, rule_tac x="[]" in exI, simp) using B2 elem_min_elems tickFree_append by blast from C1 B1 have D1:"trace_hide t' (ev ` A) @ trace_hide t'' (ev ` A) = trace_hide t (ev ` A)" by fastforce from B1 C1 D1 min_elems_no[OF A C3] have E1:"x=trace_hide t' (ev ` A)" by (metis (no_types, lifting) append.assoc le_list_def) with B1 B2 B3 C1 D1 7(5)[simplified, rule_format, of "t'" "[]"] have E2:"(\(f:: nat \ 'a event list). strict_mono f \ (\i. f i \ \ Q) \ (\i. trace_hide (f i) (ev ` A) \ trace_hide (f 0) (ev ` A)) \ t' \ range f)" apply simp using front_tickFree_append Hiding_tickFree tickFree_append by blast with E1 7(3) C2 inf_hidden[of A t' Q] show ?thesis by (metis (no_types, lifting) contra_subsetD) next case False from B1 B2 B3 B4 have C:"trace_hide t (ev ` A) \ div_hide P A" by (simp, rule_tac x=t in exI, rule_tac x="[]" in exI, simp) from B1 min_elems_no[OF A C] have E1:"x=trace_hide t (ev ` A)" using le_list_def by auto from B4 False 7(2)[rule_format, of t, OF False] have "t \ \ Q" using F_T T_F by blast with E1 7(5)[simplified, rule_format, of t "[]", simplified, OF E1 B2] inf_hidden[of A t Q] show ?thesis by metis qed qed lemma cont_Hiding1 :\chain Y \ chain (\ i. (Y i \ A))\ by (simp add: po_class.chain_def) lemma KoenigLemma: assumes *:"infinite (Tr::'a list set)" and **:"\i. finite{t. \t'\Tr. t = take i t'}" shows "\(f::nat \ 'a list). strict_mono f \ range f \ {t. \t'\Tr. t \ t'}" proof - define Tr' where "Tr' = {t. \t'\Tr. t \ t'}" (* prefix closure *) have a:"infinite Tr'" by (metis (mono_tags, lifting) * Tr'_def infinite_super mem_Collect_eq order_refl subsetI) { fix i have "{t \ Tr'. length t = i} \ {t. \t'\Tr. t = take i t'}" by (auto simp add:Tr'_def, metis append_eq_conv_conj le_list_def) with ** have "finite({t\Tr'. length t = i})" using infinite_super by blast } note b=this { fix t define E where "E = {e |e. t@[e]\ Tr'}" have aa:"finite E" proof - have "inj_on (\e. t @ [e]) E" by (simp add: inj_on_def) with b[of "Suc (length t)"] inj_on_finite[of "\e. t@[e]" E "{t' \ Tr'. length t' = Suc (length t)}"] show ?thesis by (simp add: E_def image_Collect_subsetI) qed hence bb:"finite {t@[e] |e. e\E}" by simp have "{t'\Tr'. t < t'} = {t@[e] |e. e\E} \ (\e\E. {t'\Tr'. t@[e] < t'})" proof (auto simp add:Let_def E_def Tr'_def, goal_cases) case (1 x xa) then obtain u e where "x = t @ [e] @ u" by (metis A append_Cons append_Nil append_Nil2 le_list_def list.exhaust) with 1 1(4)[rule_format, of e] show ?case by (metis append_assoc le_list_def less_list_def) next case (2 x xa xb xc) thus ?case by (meson less_self less_trans) qed with aa bb have "infinite {t'\Tr'. t < t'} \ \e. infinite {t'\Tr'. t@[e] < t'}" by auto } note c=this define ff where d:"ff =rec_nat [] (\i t. (let e = SOME e. infinite {t'\Tr'. t@[e] < t'} in t @ [e]))" hence dd:"\n. ff (Suc n) > ff n" by simp (* funny *) hence e:"strict_mono ff" by (simp add: lift_Suc_mono_less strict_monoI) { fix n have "ff n \ Tr' \ infinite {t' \ Tr'. ff n < t'}" proof(induct n) case 0 from a Tr'_def have "Tr' = {t' \ Tr'. [] < t'} \ {[]}" by (auto simp add: le_neq_trans) with a have "infinite {t' \ Tr'. [] < t'}" by (metis (no_types, lifting) finite.emptyI finite.insertI finite_UnI) with d Tr'_def show ?case by auto next case (Suc n) from d have "ff (Suc n) = (let e = SOME e. infinite {t'\Tr'. ff n@[e] < t'} in ff n @ [e])" by simp with c[rule_format, of "ff n"] obtain e where a1:"ff (Suc n) = (ff n) @ [e] \ infinite {t' \ Tr'. ff n @ [e] < t'}" by (metis (no_types, lifting) Suc.hyps someI_ex) then obtain i where "i \ Tr' \ ff (Suc n) < i" using not_finite_existsD by auto with Tr'_def have "ff (Suc n) \ Tr'" using dual_order.trans less_imp_le by fastforce with a1 show ?case by simp qed } note g=this hence h:"range ff \ Tr'" by auto show ?thesis using Tr'_def e h by blast qed lemma div_Hiding_lub : \finite (A::'a set) \ chain Y \ \ (\ i. (Y i \ A)) \ \ ((\ i. Y i) \ A)\ proof (auto simp add:limproc_is_thelub cont_Hiding1 D_Hiding T_Hiding D_LUB T_LUB, goal_cases) case (1 x) { fix xa t u f assume a:"front_tickFree u \ tickFree t \ x = trace_hide t (ev ` A) @ u \ isInfHiddenRun f (Y xa) A \ (\ i. f i \ \ (Y xa)) \ t \ range f" hence "(\i n. f i \ \ (Y n))" using 1(2) NT_ND chain_lemma le_approx2T by blast with a have ?case by blast } note aa=this { fix xa t u f j assume a:"front_tickFree u \ tickFree t \ x = trace_hide t (ev ` A) @ u \ isInfHiddenRun f (Y xa) A \ (f j \ \ (Y xa)) \ t \ range f" hence "\t u. front_tickFree u \ tickFree t \ x = trace_hide t (ev ` A) @ u \ t \ \ (Y xa)" apply(rule_tac x="f j" in exI, rule_tac x=u in exI) using Hiding_tickFree[of "f j" A] Hiding_tickFree[of t A] by (metis imageE) } note bb=this have cc: "\xa. \t u. front_tickFree u \ tickFree t \ x = trace_hide t (ev ` A) @ u \ t \ \(Y xa) \ ?case" (is "\xa. \t. ?S t xa \ ?case") proof - assume dd:"\xa. \t u. front_tickFree u \ tickFree t \ x = trace_hide t (ev ` A) @ u \ t \ \(Y xa)" (is "\xa. \t. ?S t xa") define f where "f = (\n. SOME t. ?S t n)" thus ?case proof (cases "finite(range f)") case True obtain t where gg:"infinite (f -` {t})" using f_def True inf_img_fin_dom by blast then obtain k where "f k = t" using finite_nat_set_iff_bounded_le by blast then obtain u where uu:"front_tickFree u \ x = trace_hide t (ev ` A) @ u \ tickFree t" using f_def dd[rule_format, of k] some_eq_ex[of "\t. ?S t k"] by blast { fix m from gg obtain n where gg:"n \ m \ n \ (f -` {t})" by (meson finite_nat_set_iff_bounded_le nat_le_linear) hence "t \ \ (Y n)" using f_def dd[rule_format, of n] some_eq_ex[of "\t. ?S t n"] by auto with gg 1(2) have "t \ \ (Y m)" by (meson contra_subsetD le_approx_def po_class.chain_mono) } with gg uu show ?thesis by blast next case False { fix t assume "t \ range f" then obtain k where "f k = t" using finite_nat_set_iff_bounded_le by blast then obtain u where uu:"front_tickFree u \ x = trace_hide t (ev ` A) @ u \ tickFree t" using f_def dd[rule_format, of k] some_eq_ex[of "\t. ?S t k"] by blast hence "set t \ set x \ (ev ` A)" by auto } note ee=this { fix i have "finite {(take i t)|t. t \ range f}" proof(induct i, simp) case (Suc i) have ff:"{take (Suc i) t|t. t \ range f} \ {(take i t)|t. t \ range f} \ (\e\(set x \ (ev ` A)). {(take i t)@[e]|t. t \ range f})" (is "?AA \ ?BB") proof fix t assume "t \ ?AA" then obtain t' where hh:"t' \ range f \ t = take (Suc i) t'" using finite_nat_set_iff_bounded_le by blast with ee[of t'] show "t \ ?BB" proof(cases "length t' > i") case True hence ii:"take (Suc i) t' = take i t' @ [t'!i]" by (simp add: take_Suc_conv_app_nth) with ee[of t'] have "t'!i \ set x \ (ev ` A)" by (meson True contra_subsetD hh nth_mem) with ii hh show ?thesis by blast next case False hence "take (Suc i) t' = take i t'" by fastforce with hh show ?thesis by auto qed qed { fix e have "{x @ [e] |x. \t. x = take i t \ t \ range f} = {take i t' @ [e] |t'. t' \ range f}" by auto hence "finite({(take i t') @ [e] |t'. t'\range f})" using finite_image_set[of _ "\t. t@[e]", OF Suc] by auto } note gg=this have "finite(set x \ (ev ` A))" using 1(1) by simp with ff gg Suc show ?case by (metis (no_types, lifting) finite_UN finite_Un finite_subset) qed } note ff=this hence "\i. {take i t |t. t \ range f} = {t. \t'. t = take i (f t')}" by auto with KoenigLemma[of "range f", OF False] ff obtain f' where gg:"strict_mono (f':: nat \ 'a event list) \ range f' \ {t. \t'\range f. t \ t'}" by auto { fix n define M where "M = {m. f' n \ f m }" assume "finite M" hence l1:"finite {length (f m)|m. m \ M}" by simp obtain lm where l2:"lm = Max {length (f m)|m. m \ M}" by blast { fix k have "length (f' k) \ k" by(induct k, simp, metis (full_types) gg lessI less_length_mono linorder_not_le not_less_eq_eq order_trans strict_mono_def) } with gg obtain m where r1:"length (f' m) > lm" by (meson lessI less_le_trans) from gg obtain r where r2:"f' (max m n) \ f r" by blast with gg have r3: "r \ M" by (metis (no_types, lifting) M_def max.bounded_iff mem_Collect_eq order_refl order_trans strict_mono_less_eq) with l1 l2 have f1:"length (f r) \ lm" using Max_ge by blast from r1 r2 have f2:"length (f r) > lm" by (meson dual_order.strict_trans1 gg le_length_mono max.bounded_iff order_refl strict_mono_less_eq) from f1 f2 have False by simp } note ii=this { fix i n from ii obtain m where jj:"m \ n \ f m \ f' i" by (metis finite_nat_set_iff_bounded_le mem_Collect_eq nat_le_linear) have kk: "(f m) \ \ (Y m)" using f_def dd[rule_format, of m] some_eq_ex[of "\t. ?S t m"] by auto with jj gg have "(f' i) \ \ (Y m)" by (meson D_T is_processT3_ST_pref) with jj 1(2) have "(f' i) \ \ (Y n)" using D_T le_approx2T po_class.chain_mono by blast } note jj=this from gg have kk:"mono (\n. trace_hide (f' n) (ev ` A))" unfolding strict_mono_def mono_def by (metis (no_types, lifting) filter_append gg le_list_def mono_def strict_mono_mono) { fix n from gg obtain k r where "f k = f' n @ r" by (metis ii le_list_def not_finite_existsD) hence "trace_hide (f' n) (ev ` A) \ x" using f_def dd[rule_format, of k] some_eq_ex[of "\t. ?S t k"] le_list_def by auto blast } note ll=this { assume llll:"\m. \n. trace_hide (f' n) (ev ` A) > trace_hide (f' m) (ev ` A)" hence lll:"\m. \n. length (trace_hide (f' n) (ev ` A)) > length (trace_hide (f' m) (ev ` A))" using less_length_mono by blast define ff where lll':"ff = rec_nat (length (trace_hide (f' 0) (ev ` A))) (\i t. (let n = SOME n. (length (trace_hide (f' n) (ev ` A))) > t in length (trace_hide (f' n) (ev ` A))))" { fix n from lll' lll[rule_format, of n] have "ff (Suc n) > ff n" apply simp apply (cases n) apply (metis (no_types, lifting) old.nat.simps(6) someI_ex) by (metis (no_types, lifting) llll less_length_mono old.nat.simps(7) someI_ex) } note lll''=this with lll'' have "strict_mono ff" by (simp add: lll'' lift_Suc_mono_less strict_monoI) hence lll''':"infinite(range ff)" using finite_imageD strict_mono_imp_inj_on by auto from lll lll' have "range ff \ range (\n. length (trace_hide (f' n) (ev ` A)))" by (auto, metis (mono_tags, lifting) old.nat.exhaust old.nat.simps(6) old.nat.simps(7) range_eqI) with lll''' have "infinite (range (\n. length (trace_hide (f' n) (ev ` A))))" using finite_subset by auto hence "\m. length (trace_hide (f' m) (ev ` A)) > length x" using finite_nat_set_iff_bounded_le by (metis (no_types, lifting) not_less rangeE) with ll have False using le_length_mono not_less by blast } then obtain m where mm:"\n. trace_hide (f' n) (ev ` A) \ trace_hide (f' m) (ev ` A)" by (metis (mono_tags, lifting) A kk le_cases mono_def) with gg obtain k where nn:"f k \ f' m" by blast then obtain u where oo:"front_tickFree u \ x = trace_hide (f' m) (ev ` A) @ u \ tickFree (f' m)" using f_def dd[rule_format, of k] some_eq_ex[of "\t. ?S t k"] by (auto, metis (no_types, lifting) contra_subsetD filter_is_subset front_tickFree_append front_tickFree_mono le_list_def ll tickFree_Nil tickFree_append tickFree_def tickFree_implies_front_tickFree) show ?thesis apply(rule_tac x="f' m" in exI, rule_tac x=u in exI) apply(simp add:oo, rule disjI2, rule_tac x="\n. f' (n + m)" in exI) using gg jj kk mm apply (auto simp add: strict_mono_def dual_order.antisym mono_def) by (metis plus_nat.add_0 rangeI) qed qed show ?case proof (cases "\ xa t u f. front_tickFree u \ tickFree t \ (\ i. f i \ \ (Y xa)) \ t \ range f \ x = trace_hide t (ev ` A) @ u \ isInfHiddenRun f (Y xa) A") case True then show ?thesis using aa by blast next case False have dd:"\xa. \t u. front_tickFree u \ tickFree t \ x = trace_hide t (ev ` A) @ u \ (t \ \ (Y xa) \ (\f. isInfHiddenRun f (Y xa) A \ (\i. f i \ \ (Y xa)) \ t \ range f))" (is "\xa. ?dd xa") proof (rule_tac allI) fix xa from 1(3) obtain t u where "front_tickFree u \ tickFree t \ x = trace_hide t (ev ` A) @ u \ (t \ \ (Y xa) \ (\f. isInfHiddenRun f (Y xa) A \ t \ range f))" by blast thus "?dd xa" apply(rule_tac x=t in exI, rule_tac x=u in exI, intro conjI, simp_all, elim conjE disjE, simp_all) using 1(1) False NT_ND chain_lemma le_approx2T by blast qed hence ee:"\xa. \t u. front_tickFree u \ tickFree t \ x = trace_hide t (ev ` A) @ u \ t \ \(Y xa)" using bb by blast with cc show ?thesis by simp qed qed lemma cont_Hiding2 : \finite A \ chain Y \ ((\ i. Y i) \ A) = (\ i. (Y i \ A))\ proof(auto simp add:limproc_is_thelub cont_Hiding1 Process_eq_spec D_Hiding Ra_def F_Hiding T_Hiding F_LUB D_LUB T_LUB, goal_cases) case (1 b x t u) thus ?case by blast next case (2 b x u f xa) thus ?case by blast next case (3 s X) hence \s \ \ ((\ i. Y i) \ A)\ by (simp add:limproc_is_thelub cont_Hiding1 F_LUB D_LUB T_LUB D_Hiding) with 3(1,2) obtain n where a: \s \ \ (Y n \ A)\ by (metis (no_types, lifting) D_LUB_2 div_Hiding_lub subsetCE limproc_is_thelub cont_Hiding1) with 3(3) obtain t where b:"s = trace_hide t (ev ` A) \ (t, X \ ev ` A) \ \ (Y n)" unfolding D_Hiding by blast hence c:"front_tickFree t" using is_processT2 by blast have d:"t \ \(Y n)" proof(cases "tickFree t") case True with a b show ?thesis using front_tickFree_Nil by (simp add: D_Hiding) next case False with c obtain t' where "t = t'@[tick]" using nonTickFree_n_frontTickFree by blast with a b show ?thesis apply(simp add: D_Hiding, erule_tac x=t' in allE, erule_tac x="[tick]" in allE, simp) by (metis event.distinct(1) filter.simps(1) front_tickFree_implies_tickFree imageE is_processT) qed with b show ?case using 3(2) NF_ND chain_lemma proc_ord2a by blast next case (4 xa t u) thus ?case by blast next case (5 xa u f xb) thus ?case by blast next case (6 s) hence \s \ \ (\ i. (Y i \ A))\ by (simp add:limproc_is_thelub cont_Hiding1 6(1) D_Hiding D_LUB) with div_Hiding_lub[OF 6(1,2)] have \s \ \ ((\i. Y i) \ A)\ by blast thus ?case by (simp add:limproc_is_thelub 6(2) D_Hiding D_LUB T_LUB) qed lemma cont_Hiding_base[simp] : \finite A \ cont (\x. x \ A)\ by (simp add: cont_def cont_Hiding1 cont_Hiding2 cpo_lubI) lemma Hiding_cont[simp]: \finite A \ cont f \ cont (\x. f x \ A)\ by (rule_tac f=f in cont_compose, simp_all) end \ No newline at end of file diff --git a/thys/HOL-CSP/Introduction.thy b/thys/HOL-CSP/Introduction.thy --- a/thys/HOL-CSP/Introduction.thy +++ b/thys/HOL-CSP/Introduction.thy @@ -1,314 +1,316 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * * Author : Burkhart Wolff, Safouan Taha, Lina Ye. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * Copyright (c) 2023 Université Paris-Saclay, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) chapter\Context\ (*<*) theory Introduction imports HOLCF begin (*>*) section\Preface\ text\ This is a formalization in Isabelle/HOL of the work of Hoare and Roscoe on the denotational semantics of the Failure/Divergence Model of CSP. It follows essentially the presentation of CSP in Roscoe's Book "Theory and Practice of Concurrency" @{cite "roscoe:csp:1998" } and the semantic details in a joint Paper of Roscoe and Brooks "An improved failures model for communicating processes" @{cite "brookes-roscoe85"}. The original version of this formalization, called HOL-CSP 1.0 @{cite "tej.ea:corrected:1997"}, revealed minor, but omnipresent foundational errors in key concepts like the process invariant. A correction was proposed slightly before the apparition of Roscoe's book (all three authors were in e-mail contact at that time). \ text\ In recent years, a team of authors undertook the task to port HOL-CSP 1.0 to modern Isabelle/HOL and structured proof techniques. This results in the present version are called HOL-CSP 2.0. The effort is motivated by the following assets of CSP: \<^item> the theory is interesting in itself, and reworking its formal structure might help to make it more widely accessible, given that it is a particularly advanced example of the shallow embedding technique using the denotational semantics of a language, \<^item> it is interesting to \<^emph>\compare\ the ancient, imperative, ML-heavy proof style to the more recent declarative one in Isabelle/Isar; this comparison (not presented here) gives a source of empirical evidence that such proofs are more stable wrt. the constant changes in the Isabelle itself, \<^item> the \<^emph>\semantic\ presentation of CSP lends itself to a semantically clean and well-understood \<^emph>\combination\ of specification languages, which represents a major step to our longterm goal of heterogenuous, yet semantically clean system specifications consisting of different formalisms describing components or system aspects separately, \<^item> the resulting HOL-CSP environment could one day be used as a tool that certifies traces of other CSP model-checkers such as FDR4 or PAT. \ text\ In contrast to HOL-CSP 1.0, which came with an own fixpoint theory partly inspired by previous work of Alberto Camilieri under HOL4 @{cite"Camilleri91"}, it is the goal of the redesign of HOL-CSP 2.0 to reuse the HOLCF theory that emmerged from Franz Regensburgers work and was substantially extended by Brian Huffman. Thus, the footprint of the HOL-CSP 2.0 theory should be reduced drastically. Moreover, all proofs have been heavily revised or re-constructed to reflect the drastically improved state of the art of interactive theory development with Isabelle. \ section\Introduction\ text\DRAWN FROM THE PAPER @{cite "tej.ea:corrected:1997"}\ text\ In his invited lecture at FME'96, C.A.R. Hoare presented his view on the status quo of formal methods in industry. With respect to formal proof methods, he ruled that they "are now sufficiently advanced that a [...] formal methodologist could occasionally detect [...] obscure latent errors before they occur in practice" and asked for their publication as a possible "milestone in the acceptance of formal methods" in industry. In this paper, we report of a larger verification effort as part of the UniForM project @{cite "KriegBrueckner95"}. It revealed an obscure latent error that was not detected within a decade. It cannot be said that the object of interest is a "large software system" whose failure may "cost millions", but it is a well-known subject in the center of academic interest considered foundational for several formal methods tools: the theory of the failure- divergence model of CSP (@{cite "Hoare:1985:CSP:3921"}, @{cite "brookes-roscoe85"}). And indeed we hope that this work may further encourage the use of formal proof methods at least in the academic community working on formal methods. Implementations of proof support for a formal method can roughly be divided into two categories. In direct tools like FDR @{cite "FDRTutorial2000"}, the logical rules of a method (possibly integrated into complex proof techniques) are hard-wired into the code of their implementation. Such tools tend to be difficult to modify and to formally reason about, but can possess enviable automatic proof power in specific problem domains and comfortable user interfaces. The other category can be labelled as logical embeddings. Formal methods such as CSP or Z can be logically embedded into an LCF-style tactical theorem prover such as HOL @{cite "gordon.ea:hol:1993"} or Isabelle@{cite "nipkow.ea:isabelle:2002"}. Coming with an open system design going back to Milner, these provers allow for user-programmed extensions in a logically sound way. Their strength is flexibility, generality and expressiveness that makes them to symbolic programming environments. In this paper we present a tool of the latter category (as a step towards a future combination with the former). After a brief introduction into the failure divergence semantics in the traditional CSP-literature, we will discuss the revealed problems and present a correction. Although the error is not "mathematically deep", it stings since its correction affects many definitions. It is shown that the corrected CSP still fulfils the desired algebraic laws. The addition of fixpoint-theory and specialised tactics extends the embedding in Isabelle/HOL to a formally proven consistent proof environment for CSP. Its use is demonstrated in a final example. \ section\An Outline of Failure-Divergence Semantics\ text\ A very first approach to give denotational semantics to CSP is to view it as a kind of a regular expression. This way, it can be understood as an automata and the denotations are just the language of the automaton; this way, synchronization and concurrency can be basically understood as the construction of a product automaton with potential interleaving. The semantics becomes compositional, and internal communication between sub-components of a component can be modeled by the concealment operator. Hoares work @{cite "Hoare:1985:CSP:3921"} was strongly inspired by this initial idea. However, it became quickly clear that the simplistic automata vision is not a satisfying paradigm for all aspects of concurrency. Particularly regarding the nature of communication, where one "sends" actively information and the other "receives" it, the bi-directional product construction seems to be misleading. Furthermore, it is an obvious difference if a group of processes remains in a passive deadlock because all possible communications contradict each other, or if a group of processes is too busy with internal chatter and never reaches the point where this component is again ready for communication. Hoare solved these apparent problems by presenting a multi-layer approach, in which the denotational models were refined more and more allowing to distinguish the above critical situations. An ingenious concept in the overall scheme is to distinguish \<^emph>\non-deterministic\ choice from \<^emph>\deterministic\ one \<^footnote>\which in itself produces problems with recursion which had to be overcome by some restrictions on its use.\ in order to solve the sender/receiver problem. Hoare proposed 3 denotational semantics for CSP: \<^item> the \<^emph>\trace\ model, which is basically the above naive automata model not allowing to distinguish non-deterministic choice from deterministic one, neither to distinguish deadlock from infinite internal chatter, \<^item> the \<^emph>\failure\ model is able to distinguish non-deterministic choice from deterministic one by different maximum refusal sets, which is however cannot differentiate deadlock from infinite internal chatter, \<^item> the \<^emph>\failure-divergence\ model overcomes additionally the unresolved problem of failure model. In the sequel, we explain these two problems in more detail, giving some motivation for the daunting complexity of the latter model. It is this complexity which finally raises general interest in a formal verification. \ subsection\Non-Determinism\ (*<*) consts dummyPrefix :: "'a::cpo \ 'b::cpo \ 'b::cpo" consts dummyDet :: "'a::cpo \ 'b::cpo \ 'b::cpo" consts dummyNDet :: "'a::cpo \ 'b::cpo \ 'b::cpo" consts dummyHide :: "'b::cpo \ 'a \ 'b::cpo" notation dummyPrefix (infixr "\" 50) notation dummyDet (infixr "\" 50) notation dummyNDet (infixr "\" 50) notation dummyHide (infixr "\" 50) (*>*) text\ Let a and b be any two events in some set of events @{term "\"}. The two processes \<^descr> @{term "(a \ Stop) \ (b \ Stop)"} \hspace{7cm} (1) and \<^descr> @{term "(a \ Stop) \ (b \ Stop)"} \hspace{7cm} (2) \ text\ cannot be distinguished under the trace semantics, in which both processes are capable of performing the same sequences of events, i.e. both have the same set of traces @{term "{{},{a},{b}}"}. This is because both processes can either engage in @{term "a"} and then @{term "Stop"}, or engage in @{term "b"} and then @{term "Stop"}. We would, however, like to distinguish between @{term "a"} \<^emph>\deterministic\ choice of @{term "a"} or @{term "b"} (1) and @{term "a"} \<^emph>\non-deterministic\ choice of @{term "a"} or @{term "b"} (2). This can be done by considering the events that a process can refuse to engage in when these events are offered by the environment; it cannot refuse either, so we say its maximal refusal set is the set containing all elements of @{term "\"} other than @{term "a"} and @{term "b"}, written @{term "\-{a,b}"}, i.e. it can refuse all elements in @{term "\"} other than @{term "a"} or @{term "b"}. In the case of the non-deterministic process (2), however, we wish to express that if the environment offers the event @{term "a"} say, the process non-deterministically chooses either to engage in @{term "a"}, to refuse it and engage in @{term "b"} (likewise for @{term "b"}). We say therefore, that process (2) has two maximal refusal sets, @{term "\-{a}"} and @{term "\-{b}"}, because it can refuse to engage in either @{term "a"} or @{term "b"}, but not both. The notion of refusal sets is in this way used to distinguish non-determinism from determinism in choices. \ subsection\Infinite Chatter\ text\ Consider the infinite process @{term "\ X. a \ X"} which performs an infinite stream of @{term "a"}'s. If one now conceals the event a in this process by writing \<^descr> @{term "(\ X. a \ X) \ a"} \hspace{7.8cm} (3) \ text\ it no longer becomes possible to distinguish the behaviour of this process from that of the deadlock process Stop. We would like to be able to make such a distinction, since the former process has clearly not stopped but is engaging in an unbounded sequence of internal actions invisible to the environment. We say the process has diverged, and introduce the notion of a divergence set to denote all sequences events that can cause a process to diverge. Hence, the process Stop is assigned the divergence set @{term "{}"}, since it can not diverge, whereas the process (3) above diverges on any sequence of events since the process begins to diverge immediately, i.e. its divergence set is @{term "\\<^sup>*"} , where @{term "\\<^sup>*"} denotes the set of all sequences with elements in @{term "\"}. Divergence is undesirable and so it is essential to be able to express it to ensure that it is avoided. \ subsection\The Global Architecture of HOL-CSP 2.0\ text\ The global architecture of HOL-CSP 2.0 has been substantially simplified compared to HOL-CSP 1.0: the fixpoint reasoning is now entirely based on HOLCF (which meant that the continuity proofs for CSP operators had basically been re-done).\ text\ The theory \<^verbatim>\Process\ establishes the basic common notions for events, traces, ticks and tickfree-ness, the type definitions for failures and divergences as well as the global constraints on them (called the "axioms" in Hoare's Book.) captured in a predicate \<^verbatim>\is_process\. On this basis, the set of failures and divergences satisfying \<^verbatim>\is_process\ is turned into the type \<^verbatim>\'a process\ via a type-definition (making \<^verbatim>\is_process\ as the central data invariant). In the sequel, it is shown that \<^verbatim>\'a process\ belongs to the type-class @{class "cpo"} stemming from @{theory HOLCF} which makes the concepts of complete partial order, continuity, fixpoint-induction and general recursion available to all expressions of type \<^verbatim>\'a process\. \ text\ The theory \<^verbatim>\Process\ also establishes the two partial orderings @{term "P \ P'"} for refinements and @{term "P \ P'"} for the approximation on processes used to give semantics to recursion. The latter is well-known to be logically weaker than the former. Note that, unfortunately, the use of these two symbols in HOL-CSP 2.0, where the latter is already used in the @{theory HOLCF}-theory, is just the other way round as in the literature. \ text\ Each CSP operator is described in an own theory which comprises: \<^item> The denotational core definition in terms of a pair of Failures and Divergences \<^item> The establishment of \<^verbatim>\is_process\ for the Failures and Divergences in the - range of the given operator (thus, the invariance of \<^verbatim>\is_process\ for this operator) + range of the given operator (thus, the invariance of \<^verbatim>\is_process\ for this operator). + In this new version, the proof is required immediatly after the definition of the + operator since we use lift\_definition instead of definition \<^item> The proof of the projections \<^verbatim>\T\ and \<^verbatim>\F\ and \<^verbatim>\D\ for this operator \<^item> The proof of continuity of the operator (which is always possible except for - hide if applied to infinite hide-sets).\ + Hiding if applied to infinite hide-sets).\ text\ Finally, the theory CSP contains the "Laws" of CSP, i.e. the derived rules allowing abstractly to reason over CSP processes. The overall dependency graph is shown in \autoref{fig:fig1}.\ text\ \begin{figure}[ht] \centering \includegraphics[width=0.9\textwidth]{session_graph} \caption{The HOL-CSP 2.0 Theory Graph} \label{fig:fig1} \end{figure}\ (*<*) no_notation dummyPrefix (infixr "\" 50) no_notation dummyDet (infixr "\" 50) no_notation dummyNDet (infixr "\" 50) no_notation dummyHide (infixr "\" 50) end (*>*) diff --git a/thys/HOL-CSP/Mndetprefix.thy b/thys/HOL-CSP/Mndetprefix.thy --- a/thys/HOL-CSP/Mndetprefix.thy +++ b/thys/HOL-CSP/Mndetprefix.thy @@ -1,252 +1,235 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * - * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * Author : Burkhart Wolff, Safouan Taha. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) theory Mndetprefix imports Process Stop Mprefix Ndet begin section\Multiple non deterministic operator\ -definition - Mndetprefix :: "['\ set, '\ \ '\ process] \ '\ process" - where "Mndetprefix A P \ if A = {} - then STOP - else Abs_process(\ x\A. \(x \ P x), - \ x\A. \(x \ P x))" -syntax - "_Mndetprefix" :: "pttrn \ 'a set \ 'a process \ 'a process" - ("(3\_\_ \ / _)" [0, 0, 70] 70) - -translations - "\x\A\ P" \ "CONST Mndetprefix A (\x. P)" - -lemma mt_Mndetprefix[simp] : "Mndetprefix {} P = STOP" - unfolding Mndetprefix_def by simp - -lemma Mndetprefix_is_process : "A \ {} \ is_process (\ x\A. \(x \ P x), \ x\A. \(x \ P x))" +lift_definition Mndetprefix :: "['\ set, '\ \ '\ process] \ '\ process" + is "\A P. if A = {} then Rep_process STOP + else (\ x\A. \(x \ P x), \ x\A. \(x \ P x))" + apply auto + using Rep_process apply blast unfolding is_process_def FAILURES_def DIVERGENCES_def apply auto using is_processT1 apply auto[1] using is_processT2 apply blast using is_processT3_SR apply blast using is_processT4 apply blast using is_processT5_S1 apply blast using is_processT6 apply blast using is_processT7 apply blast using NF_ND apply auto[1] using is_processT9 by blast -lemma T_Mndetprefix1 : "\ (Mndetprefix {} P) = {[]}" - unfolding Mndetprefix_def by(simp add: T_STOP) +syntax + "_Mndetprefix" :: "pttrn \ 'a set \ 'a process \ 'a process" + ("(3\_\_ \ / _)" [0, 0, 70] 70) -lemma rep_abs_Mndetprefix[simp]: "A \ {} \ +translations + "\x\A\ P" \ "CONST Mndetprefix A (\x. P)" + +lemma mt_Mndetprefix[simp] : "Mndetprefix {} P = STOP" + by (simp add: Mndetprefix_def Rep_process_inverse) + +lemma rep_abs_Mndetprefix: "A \ {} \ (Rep_process (Abs_process(\x\A. \(x \ P x),\x\A. \ (x \ P x)))) = (\x\A. \(x \ P x), \x\A. \ (x \ P x))" - apply(subst Process.process.Abs_process_inverse) - by(auto intro: Mndetprefix_is_process[simplified]) + by (metis (mono_tags, lifting) Abs_process_inverse Mndetprefix.rep_eq Rep_process) -lemma T_Mndetprefix: "A\{} \ \ (Mndetprefix A P) = (\ x\A. \ (x \ P x))" - unfolding Mndetprefix_def - apply(simp, subst Traces_def, simp add: TRACES_def FAILURES_def) - apply(auto intro: Mndetprefix_is_process[simplified]) - unfolding TRACES_def FAILURES_def apply(cases "A = {}") - apply(auto intro: F_T D_T simp: Nil_elem_T) - using NF_NT by blast - -lemma F_Mndetprefix1 : "\ (Mndetprefix {} P) = {(s, X). s = []}" - unfolding Mndetprefix_def by(simp add: F_STOP) - -lemma F_Mndetprefix: "A\{} \ \ (Mndetprefix A P) = (\ x\A. \(x \ P x))" - unfolding Mndetprefix_def - by(simp, subst Failures_def, auto simp : FAILURES_def F_Mndetprefix1) -lemma D_Mndetprefix1 : "\ (Mndetprefix {} P) = {}" - unfolding Mndetprefix_def by(simp add: D_STOP) +lemma F_Mndetprefix: + "\ (Mndetprefix A P) = (if A = {} then {(s, X). s = []} else \ x\A. \ (x \ P x))" + by (simp add: Failures.rep_eq FAILURES_def STOP.rep_eq Mndetprefix.rep_eq) -lemma D_Mndetprefix : "A\{} \ \(Mndetprefix A P) = (\ x\A. \ (x \ P x))" - unfolding Mndetprefix_def - apply(simp, subst D_def, subst Process.process.Abs_process_inverse) - by(auto intro: Mndetprefix_is_process[simplified] simp: DIVERGENCES_def) +lemma D_Mndetprefix : "\ (Mndetprefix A P) = (if A = {} then {} else \ x\A. \ (x \ P x))" + by (simp add: Divergences.rep_eq DIVERGENCES_def STOP.rep_eq Mndetprefix.rep_eq) + +lemma T_Mndetprefix: "\ (Mndetprefix A P) = (if A = {} then {[]} else \ x\A. \ (x \ P x))" + by (auto simp add: Traces.rep_eq TRACES_def Failures.rep_eq[symmetric] F_Mndetprefix) + text\ Thus we know now, that Mndetprefix yields processes. Direct consequences are the following distributivities: \ lemma Mndetprefix_unit : "(\ x \ {a} \ P x) = (a \ P a)" by(auto simp : Process.Process_eq_spec F_Mndetprefix D_Mndetprefix) lemma Mndetprefix_Un_distrib : "A \{} \ B \{} \ (\ x \ A\B \ P x) = ((\ x \ A \ P x) \ (\ x \ B \ P x))" by(auto simp : Process.Process_eq_spec F_Ndet D_Ndet F_Mndetprefix D_Mndetprefix) text\ The two lemmas @{thm Mndetprefix_unit} and @{thm Mndetprefix_Un_distrib} together give us that @{const Mndetprefix} can be represented by a fold in the finite case. \ lemma Mndetprefix_distrib_unit : "A-{a} \ {} \ (\ x \ insert a A \ P x) = ((a \ P a) \ (\ x \ A-{a} \ P x))" by (metis Un_Diff_cancel insert_is_Un insert_not_empty Mndetprefix_Un_distrib Mndetprefix_unit) subsection\Finite case Continuity\ text\ This also implies that Mndetprefix is continuous for the finite @{term A} and an arbitrary body @{term f}: \ lemma Mndetprefix_cont_finite[simp]: assumes "finite A" and "\x. cont (f x)" shows "cont (\y. \ z \ A \ f z y)" proof(rule Finite_Set.finite_induct[OF `finite A`]) show "cont (\y. \z\{} \ f z y)" by auto next fix A fix a assume "cont (\y. \z\A \ f z y)" and "a \ A" show "cont (\y. \z\insert a A \ f z y)" proof(cases "A={}") case True then show ?thesis by(simp add: Mndetprefix_unit True `\x. cont (f x)`) next case False have * : "A-{a} \ {}" by (simp add: False \a \ A\) have ** : "A-{a} = A" by (simp add: \a \ A\) show ?thesis apply(simp only: Mndetprefix_distrib_unit[OF *], simp only: **) by (simp add: \cont (\y. \z\A \ f z y)\ assms(2)) qed qed subsection\General case Continuity\ lemma mono_Mndetprefix[simp] : "monofun (Mndetprefix (A::'a set))" proof(cases "A={}") case True then show ?thesis by(auto simp: monofun_def) next case False then show ?thesis apply(simp add: monofun_def, intro allI impI) unfolding le_approx_def proof(simp add:T_Mndetprefix F_Mndetprefix D_Mndetprefix, intro conjI) fix x::"'a \ 'a process" and y::"'a \ 'a process" assume "A \ {}" and "x \ y" show "(\x\A. \ (x\ y x)) \ (\xa\A. \ (xa \ x xa))" by (metis (mono_tags, lifting) SUP_mono \x \ y\ fun_below_iff le_approx1 mono_Mprefix0 write0_def) next fix x::"'a \ 'a process" and y::"'a \ 'a process" assume *:"A \ {}" and **:"x \ y" have *** : "\z. z \ A \ x z \ y z " by (simp add: \x \ y\ fun_belowD) with * show "\s. (\xa\A. s \ \ (xa \ x xa)) \ Ra (Mndetprefix A x) s = Ra (Mndetprefix A y) s" unfolding Ra_def by (auto simp:proc_ord2a mono_Mprefix0 write0_def F_Mndetprefix) (meson le_approx2 mono_Mprefix0 write0_def)+ next fix x::"'a \ 'a process" and y::"'a \ 'a process" assume *:"A \ {}" and "x \ y" have *** : "\z. z \ A \ (z \ x z) \ (z \ y z) " by (metis \x \ y\ fun_below_iff mono_Mprefix0 write0_def) with * show "min_elems (\xa\A. \ (xa \ x xa)) \ (\x\A. \ (x \ y x))" unfolding min_elems_def apply auto by (metis Set.set_insert elem_min_elems insert_subset le_approx3 le_less min_elems5) qed qed lemma Mndetprefix_chainpreserving: "chain Y \ chain (\i. (\ z \ A \ Y i z))" apply(rule chainI, rename_tac i) apply(frule_tac i="i" in chainE) by (simp add: below_fun_def mono_Mprefix0 write0_def monofunE) lemma contlub_Mndetprefix : "contlub (Mndetprefix A)" proof(cases "A={}") case True then show ?thesis by(auto simp: contlub_def) next case False show ?thesis proof(rule contlubI, rule proc_ord_proc_eq_spec) fix Y :: "nat \ 'a \ 'a process" assume a:"chain Y" show "(\x\A \ (\i. Y i) x) \ (\i. \x\A \ Y i x)" proof(simp add:le_approx_def, intro conjI allI impI) show "\ (\i. \x\A \ Y i x) \ \ (\x\A \ Lub Y x)" using a False D_LUB[OF Mndetprefix_chainpreserving[OF a], of A] limproc_is_thelub[OF Mndetprefix_chainpreserving[OF a], of A] apply (auto simp add:write0_def D_Mprefix D_LUB[OF ch2ch_fun[OF a]] limproc_is_thelub_fun[OF a] D_Mndetprefix) by (metis (mono_tags, opaque_lifting) event.inject) next fix s :: "'a event list" assume "s \ \ (\x\A \ Lub Y x)" show "Ra (\x\A \ Lub Y x) s = Ra (\i. \x\A \ Y i x) s" unfolding Ra_def using a False F_LUB[OF Mndetprefix_chainpreserving[OF a], of A] limproc_is_thelub[OF Mndetprefix_chainpreserving[OF a], of A] apply (auto simp add:write0_def F_Mprefix F_LUB[OF ch2ch_fun[OF a]] limproc_is_thelub_fun[OF a] F_Mndetprefix) by (metis (mono_tags, opaque_lifting) event.inject) next show "min_elems (\ (\x\A \ Lub Y x)) \ \ (\i. \x\A \ Y i x)" unfolding min_elems_def using a False limproc_is_thelub[OF Mndetprefix_chainpreserving[OF a], of A] D_LUB[OF Mndetprefix_chainpreserving[OF a], of A] F_LUB[OF Mndetprefix_chainpreserving[OF a], of A] by (auto simp add:D_T write0_def D_Mprefix F_Mprefix D_Mndetprefix F_Mndetprefix D_LUB[OF ch2ch_fun[OF a]] F_LUB[OF ch2ch_fun[OF a]] limproc_is_thelub_fun[OF a]) qed next fix Y :: "nat \ 'a \ 'a process" assume a:"chain Y" show "\ (\x\A \ (\i. Y i) x) \ \ (\i. \x\A \ Y i x)" using a False D_LUB[OF Mndetprefix_chainpreserving[OF a], of A] limproc_is_thelub[OF Mndetprefix_chainpreserving[OF a], of A] by (auto simp add:write0_def D_Mprefix D_Mndetprefix D_LUB[OF ch2ch_fun[OF a]] limproc_is_thelub_fun[OF a]) qed qed lemma Mndetprefix_cont[simp]: "(\x. cont (f x)) \ cont (\y. (\ z \ A \ (f z y)))" apply(rule_tac f = "\z y. (f y z)" in Cont.cont_compose, rule monocontlub2cont) by (auto intro: mono_Mndetprefix contlub_Mndetprefix Fun_Cpo.cont2cont_lambda) end \ No newline at end of file diff --git a/thys/HOL-CSP/Mprefix.thy b/thys/HOL-CSP/Mprefix.thy --- a/thys/HOL-CSP/Mprefix.thy +++ b/thys/HOL-CSP/Mprefix.thy @@ -1,439 +1,434 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * - * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * Author : Burkhart Wolff. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) section\ The Multi-Prefix Operator Definition \ theory Mprefix imports Process begin subsection\ The Definition and some Consequences \ -definition Mprefix :: "['a set,'a => 'a process] => 'a process" where - "Mprefix A P \ Abs_process( - {(tr,ref). tr = [] \ ref Int (ev ` A) = {}} \ - {(tr,ref). tr \ [] \ hd tr \ (ev ` A) \ - (\ a. ev a = (hd tr)\(tl tr,ref)\\(P a))}, - {d. d \ [] \ hd d \ (ev ` A) \ - (\ a. ev a = hd d \ tl d \ \ (P a))})" - -syntax - "_Mprefix" :: "[pttrn,'a set,'a process] \ 'a process" ("(3\(_/\_)/ \ (_))"[0,0,64]64) - -term "Ball A (\x. P)" - -translations - "\x\A \ P" \ "CONST Mprefix A (\x. P)" - -text\Syntax Check:\ -term \\x\A \ \y\A \ \z\A \ P z x y = Q\ - -subsection\ Well-foundedness of Mprefix \ - -lemma is_process_REP_Mprefix : -"is_process ({(tr,ref). tr=[] \ ref \ (ev ` A) = {}} \ - {(tr,ref). tr \ [] \ hd tr \ (ev ` A) \ - (\ a. ev a = (hd tr) \ (tl tr,ref) \ \(P a))}, - {d. d \ [] \ hd d \ (ev ` A) \ - (\ a. ev a = hd d \ tl d \ \(P a))})" - (is "is_process(?f, ?d)") -proof (simp only:is_process_def FAILURES_def DIVERGENCES_def - Product_Type.fst_conv Product_Type.snd_conv, +lift_definition Mprefix :: "['a set,'a => 'a process] => 'a process" + is "\A P. ({(tr,ref). tr = [] \ ref Int (ev ` A) = {}} \ + {(tr,ref). tr \ [] \ hd tr \ (ev ` A) \ + (\ a. ev a = (hd tr)\(tl tr,ref)\\(P a))}, + {d. d \ [] \ hd d \ (ev ` A) \ (\ a. ev a = hd d \ tl d \ \ (P a))})" +proof - + show \is_process ({(tr, ref). tr = [] \ ref \ ev ` (A :: 'a set) = {}} \ + {(tr, ref). tr \ [] \ hd tr \ ev ` A \ + (\a. ev a = hd tr \ (tl tr, ref) \ \ (P a))}, + {d. d \ [] \ hd d \ ev ` A \ (\a. ev a = hd d \ tl d \ \ (P a))}) \ + (is \is_process(?f, ?d)\) for A P +proof (simp only: is_process_def FAILURES_def DIVERGENCES_def + Product_Type.fst_conv Product_Type.snd_conv, intro conjI allI impI) show "([],{}) \ ?f" by(simp) next fix s:: "'a event list" fix X::"'a event set" assume H : "(s, X) \ ?f" show "front_tickFree s" apply(insert H, auto simp: front_tickFree_def tickFree_def dest!:list_nonMt_append) apply(rename_tac aa ta x, case_tac "ta", auto simp: front_tickFree_charn dest! : is_processT2[rule_format]) apply(simp add: tickFree_def) done next fix s t :: "'a event list" assume "(s @ t, {}) \ ?f" then show "(s, {}) \ ?f" by(auto elim: is_processT3[rule_format]) next fix s:: "'a event list" fix X Y::"'a event set" assume "(s, Y) \ ?f \ X \ Y" then show "(s, X) \ ?f" by(auto intro: is_processT4[rule_format]) next fix s:: "'a event list" fix X Y::"'a event set" assume "(s, X) \ ?f \ (\ c. c\Y \ (s @ [c], {}) \ ?f)" then show "(s, X \ Y) \ ?f " by(auto intro!: is_processT1 is_processT5[rule_format]) next fix s:: "'a event list" fix X::"'a event set" assume "(s @ [tick], {}) \ ?f" then show "(s, X - {tick}) \ ?f" by(cases s, auto dest!: is_processT6[rule_format]) next fix s t:: "'a event list" fix X::"'a event set" assume "s \ ?d \ tickFree s \ front_tickFree t" then show "s @ t \ ?d" by(auto intro!: is_processT7_S, cases s, simp_all) next fix s:: "'a event list" fix X::"'a event set" assume "s \ ?d" then show "(s, X) \ ?f" by(auto simp: is_processT8_S) next fix s:: "'a event list" assume "s @ [tick] \ ?d" then show "s \ ?d" apply(auto) apply(cases s, simp_all) apply(cases s, auto intro: is_processT9[rule_format]) done -qed + qed + qed + + +syntax + "_Mprefix" :: "[pttrn,'a set,'a process] \ 'a process" ("(3\(_/\_)/ \ (_))"[0,0,64]64) + +term "Ball A (\x. P)" + +translations + "\x\A \ P" \ "CONST Mprefix A (\x. P)" + +text\Syntax Check:\ +term \\x\A \ \y\A \ \z\A \ P z x y = Q\ (* bizarre exercise in style proposed by Makarius... *) lemma is_process_REP_Mprefix' : "is_process ({(tr,ref). tr=[] \ ref \ (ev ` A) = {}} \ {(tr,ref). tr \ [] \ hd tr \ (ev ` A) \ (\ a. ev a = (hd tr) \ (tl tr,ref) \ \(P a))}, {d. d \ [] \ hd d \ (ev ` A) \ (\ a. ev a = hd d \ tl d \ \(P a))})" (is "is_process(?f, ?d)") proof (simp only:is_process_def FAILURES_def DIVERGENCES_def Product_Type.fst_conv Product_Type.snd_conv, intro conjI allI impI,goal_cases) case 1 show "([],{}) \ ?f" by(simp) next case (2 s X) assume H : "(s, X) \ ?f" have "front_tickFree s" apply(insert H, auto simp: front_tickFree_def tickFree_def dest!:list_nonMt_append) apply(rename_tac aa ta x, case_tac "ta", auto simp: front_tickFree_charn dest! : is_processT2[rule_format]) apply(simp add: tickFree_def) done then show "front_tickFree s" by(auto) next case (3 s t) assume H : "(s @ t, {}) \ ?f" show "(s, {}) \ ?f" using H by(auto elim: is_processT3[rule_format]) next case (4 s X Y) assume H1: "(s, Y) \ ?f \ X \ Y" then show "(s, X) \ ?f" by(auto intro: is_processT4[rule_format]) next case (5 s X Y) assume "(s, X) \ ?f \ (\ c. c\Y \ (s @ [c], {}) \ ?f)" then show "(s, X \ Y) \ ?f" by(auto intro!: is_processT1 is_processT5[rule_format]) next case (6 s X) assume "(s @ [tick], {}) \ ?f" then show "(s, X - {tick}) \ ?f" by(cases s, auto dest!: is_processT6[rule_format]) next case (7 s t) assume H1 : "s \ ?d \ tickFree s \ front_tickFree t" have 7: "s @ t \ ?d" using H1 by(auto intro!: is_processT7_S, cases s, simp_all) then show "s @ t \ ?d" using H1 by(auto) next case (8 s X) assume H : "s \ ?d" then show "(s, X) \ ?f" using H by(auto simp: is_processT8_S) next case (9 s) assume H: "s @ [tick] \ ?d" then have 9: "s \ ?d" apply(auto) apply(cases s, simp_all) apply(cases s, auto intro: is_processT9[rule_format]) done then show "s \ ?d" by(auto) qed -lemmas Rep_Abs_Mprefix[simp] = Abs_process_inverse[simplified, OF is_process_REP_Mprefix] -thm Rep_Abs_Mprefix lemma Rep_Abs_Mprefix'' : assumes H1 : "f = {(tr, ref). tr = [] \ ref \ ev ` A = {}} \ {(tr, ref). tr \ [] \ hd tr \ ev ` A \ (\a. ev a = hd tr \ (tl tr, ref) \ \ (P a))}" and H2 : "d = {d. d \ [] \ hd d \ (ev ` A) \ (\ a. ev a = hd d \ tl d \ \(P a))}" -shows "Rep_process (Abs_process (f,d)) = (f,d)" -by(subst Abs_process_inverse, - simp_all only: H1 H2 CollectI Rep_process is_process_REP_Mprefix) + shows "Rep_process (Abs_process (f,d)) = (f,d)" + using Abs_process_inverse H1 H2 is_process_REP_Mprefix' by blast + subsection \ Projections in Prefix \ lemma F_Mprefix : "\ (\ x \ A \ P x) = {(tr,ref). tr=[] \ ref \ (ev ` A) = {}} \ {(tr,ref). tr \ [] \ hd tr \ (ev ` A) \ (\ a. ev a = (hd tr) \ (tl tr,ref) \ \(P a))}" - by(simp add:Mprefix_def Failures_def Rep_Abs_Mprefix'' FAILURES_def) - + by (subst Failures.rep_eq, simp add: Mprefix.rep_eq FAILURES_def) lemma D_Mprefix: "\ (\ x \ A \ P x) = {d. d \ [] \ hd d \ (ev ` A) \ - (\ a. ev a = hd d \ tl d \ \(P a))}" - by(simp add:Mprefix_def D_def Rep_Abs_Mprefix'' DIVERGENCES_def) + (\ a. ev a = hd d \ tl d \ \(P a))}" + by (subst Divergences.rep_eq, simp add: Mprefix.rep_eq DIVERGENCES_def) lemma T_Mprefix: "\ (\ x \ A \ P x)={s. s=[] \ (\ a. a\A \ s\[] \ hd s = ev a \ tl s \ \(P a))}" - by(auto simp: T_F_spec[symmetric] F_Mprefix) + by (subst Traces.rep_eq, auto simp add: TRACES_def Failures.rep_eq[symmetric] F_Mprefix) + (use F_T T_F in \blast+\) + subsection \ Basic Properties \ lemma tick_T_Mprefix [simp]: "[tick] \ \ (\ x \ A \ P x)" by(simp add:T_Mprefix) lemma Nil_Nin_D_Mprefix [simp]: "[] \ \ (\ x \ A \ P x)" by(simp add: D_Mprefix) subsection\ Proof of Continuity Rule \ subsubsection\ Backpatch Isabelle 2009-1\ \\re-introduced from Isabelle/HOLCF 2009-1; clearly a derived concept, but still a useful shortcut\ definition contlub :: "('a::cpo \ 'b::cpo) \ bool" where "contlub f = (\Y. chain Y \ f (\ i. Y i) = (\ i. f (Y i)))" lemma contlubE: "\contlub f; chain Y\ \ f (\ i. Y i) = (\ i. f (Y i))" by (simp add: contlub_def) lemma monocontlub2cont: "\monofun f; contlub f\ \ cont f" apply (rule contI) apply (rule thelubE) apply (erule (1) ch2ch_monofun) apply (erule (1) contlubE [symmetric]) done lemma contlubI: "(\Y. chain Y \ f (\ i. Y i) = (\ i. f (Y i))) \ contlub f" by (simp add: contlub_def) lemma cont2contlub: "cont f \ contlub f" apply (rule contlubI) apply (rule Porder.po_class.lub_eqI [symmetric]) apply (erule (1) contE) done subsubsection\ Core of Proof \ lemma mono_Mprefix1: "\a\A. P a \ Q a \ \ (Mprefix A Q) \ \ (Mprefix A P)" apply(auto simp: D_Mprefix) using le_approx1 by blast lemma mono_Mprefix2: "\x\A. P x \ Q x \ \s. s \ \ (Mprefix A P) \ Ra (Mprefix A P) s = Ra (Mprefix A Q) s" apply (auto simp: Ra_def D_Mprefix F_Mprefix) using proc_ord2a by blast+ lemma mono_Mprefix3 : assumes H:"\x\A. P x \ Q x" shows " min_elems (\ (Mprefix A P)) \ \ (Mprefix A Q)" proof(auto simp: min_elems_def D_Mprefix T_Mprefix image_def, goal_cases) case (1 x a) with H[rule_format, of a, OF 1(2)] show ?case apply(auto dest!: le_approx3 simp: min_elems_def) apply(subgoal_tac "\t. t \ \ (P a) \ \ t < tl x", auto) apply(rename_tac t, erule_tac x="(ev a)#t" in allE, auto) using less_cons hd_Cons_tl by metis qed lemma mono_Mprefix0: "\x\A. P x \ Q x \ Mprefix A P \ Mprefix A Q" apply(simp add: le_approx_def mono_Mprefix1 mono_Mprefix3) apply(rule mono_Mprefix2) apply(auto simp: le_approx_def) done lemma mono_Mprefix[simp]: "monofun(Mprefix A)" by(auto simp: Fun_Cpo.below_fun_def monofun_def mono_Mprefix0) lemma proc_ord2_set: "P \ Q \ {(s, X). s \ \ P \ (s, X) \ \ P} = {(s, X). s \ \ P \ (s, X) \ \ Q}" by(auto simp: le_approx2) lemma proc_ord_proc_eq_spec: "P \ Q \ \ P \ \ Q \ P = Q" by (metis (mono_tags, lifting) below_antisym below_refl le_approx_def subset_antisym) lemma Mprefix_chainpreserving: "chain Y \ chain (\i. Mprefix A (Y i))" apply(rule chainI, rename_tac i) apply(frule_tac i="i" in chainE) by(simp add: mono_Mprefix0 fun_belowD) lemma limproc_is_thelub_fun: assumes "chain S" shows "(Lub S c) = lim_proc (range (\x. (S x c)))" proof - have "\xa. chain (\x. S x xa)" using `chain S` by(auto intro!: chainI simp: chain_def fun_belowD ) then show ?thesis by (metis contlub_lambda limproc_is_thelub) qed lemma contlub_Mprefix : "contlub(%P. Mprefix A P)" proof(rule contlubI, rule proc_ord_proc_eq_spec) fix Y :: "nat \ 'a \ 'a process" assume C : "chain Y" have C': "\xa. chain (\x. Y x xa)" apply(insert C,rule chainI) by(auto simp: chain_def fun_belowD) show "Mprefix A (\ i. Y i) \ (\ i. Mprefix A (Y i))" by(auto simp: Process.le_approx_def F_Mprefix D_Mprefix T_Mprefix C C' Mprefix_chainpreserving limproc_is_thelub limproc_is_thelub_fun D_T D_LUB D_LUB_2 F_LUB T_LUB_2 Ra_def min_elems_def) next fix Y :: "nat \ 'a \ 'a process" assume C : "chain Y" show "\ (Mprefix A (\ i. Y i)) \ \ (\ i. Mprefix A (Y i))" apply(auto simp: Process.le_approx_def F_Mprefix D_Mprefix T_Mprefix C Mprefix_chainpreserving limproc_is_thelub D_LUB_2) by (meson C fun_below_iff in_mono is_ub_thelub le_approx1) qed lemma Mprefix_cont [simp]: "(\x. cont (f x)) \ cont (\y. \ z \ A \ f z y)" apply(rule_tac f = "\z y. (f y z)" in Cont.cont_compose) apply(rule monocontlub2cont) apply(auto intro: mono_Mprefix contlub_Mprefix Fun_Cpo.cont2cont_lambda) done subsection\ High-level Syntax for Read and Write \ text\ The following syntax introduces the usual channel notation for CSP. Slightly deviating from conventional wisdom, we view a channel not as a tag in a pair, rather than as a function of type @{typ "'a\'b"}. This paves the way for \emph{typed} channels over a common universe of events. \ definition read :: "['a \ 'b,'a set, 'a \ 'b process] \ 'b process" where "read c A P \ Mprefix(c ` A) (P o (inv_into A c))" definition "write" :: "['a \ 'b, 'a, 'b process] \ 'b process" where "write c a P \ Mprefix {c a} (\ x. P)" definition write0 :: "['a, 'a process] \ 'a process" where "write0 a P \ Mprefix {a} (\ x. P)" syntax "_read" :: "[id, pttrn, 'a process] => 'a process" ("(3(_\<^bold>?_) /\ _)" [0,0,78] 78) "_readX" :: "[id, pttrn, bool,'a process] => 'a process" ("(3(_\<^bold>?_)\<^bold>|_ /\ _)" [0,0,78] 78) "_readS" :: "[id, pttrn, 'b set,'a process] => 'a process" ("(3(_\<^bold>?_)\_ /\ _)" [0,0,78] 78) "_write" :: "[id, 'b, 'a process] => 'a process" ("(3(_\<^bold>!_) /\ _)" [0,0,78] 78) "_writeS" :: "['a, 'a process] => 'a process" ("(3_ /\ _)" [0,78]78) subsection\CSP$_M$-Style Syntax for Communication Primitives\ translations "_read c p P" \ "CONST read c CONST UNIV (\p. P)" "_write c p P" \ "CONST write c p P" "_readX c p b P" => "CONST read c {p. b} (\p. P)" "_writeS a P" \ "CONST write0 a P" "_readS c p A P" => "CONST read c A (\p. P)" text\Syntax Check:\ term\ d\<^bold>?y \ c\<^bold>!x \ P = Q\ lemma read_cont[simp]: "cont P \ cont (\y. read c A (P y))" unfolding read_def o_def by (rule Mprefix_cont) (rule cont2cont_fun) lemma read_cont'[simp]: "(\x. cont (f x)) \ cont (\y. c\<^bold>?x \ f x y)" by simp lemma write_cont[simp]: "cont (P::('b::cpo \ 'a process)) \ cont(\x. (c\<^bold>!a \ P x))" by(simp add:write_def) corollary write0_cont_lub : "contlub(Mprefix {a})" using contlub_Mprefix by blast lemma write0_contlub : "contlub(write0 a)" unfolding write0_def contlub_def proof (auto) fix Y :: "nat \ 'a process" assume "chain Y" have * : "chain (\i. (\_. Y i)::'b \ 'a process)" by (meson \chain Y\ fun_below_iff po_class.chain_def) have **: "(\x. Lub Y) = Lub (\i. (\_. Y i))" by(rule ext,metis * ch2ch_fun limproc_is_thelub limproc_is_thelub_fun lub_eq) show "Mprefix {a} (\x. Lub Y) = (\i. Mprefix {a} (\x. Y i))" apply(subst **, subst contlubE[OF contlub_Mprefix]) by (simp_all add: *) qed lemma write0_cont[simp]: "cont (P::('b::cpo \ 'a process)) \ cont(\x. (a \ P x))" by(simp add:write0_def) end diff --git a/thys/HOL-CSP/Ndet.thy b/thys/HOL-CSP/Ndet.thy --- a/thys/HOL-CSP/Ndet.thy +++ b/thys/HOL-CSP/Ndet.thy @@ -1,180 +1,170 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * - * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * Author : Burkhart Wolff, Safouan Taha. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) section\ Nondeterministic Choice Operator Definition \ theory Ndet imports Process begin -subsection\ Definition and Consequences \ -definition +subsection\The Ndet Operator Definition\ +lift_definition Ndet :: "['\ process,'\ process] \ '\ process" (infixl "|-|" 80) -where "P |-| Q \ Abs_process(\ P \ \ Q , \ P \ \ Q)" - -notation - Ndet (infixl "\" 80) - - -lemma is_process_REP_Ndet: -"is_process (\ P \ \ Q , \ P \ \ Q)" +is "\P Q. (\ P \ \ Q , \ P \ \ Q)" proof(simp only: fst_conv snd_conv Rep_process is_process_def DIVERGENCES_def FAILURES_def, intro conjI) - show "([], {}) \ (\ P \ \ Q)" + show "\P Q. ([], {}) \ (\ P \ \ Q)" by(simp add: is_processT1) next - show "\s X. (s, X) \ (\ P \ \ Q) \ front_tickFree s" + show "\P Q. \s X. (s, X) \ (\ P \ \ Q) \ front_tickFree s" by(auto simp: is_processT2) next - show "\s t. (s @ t, {}) \(\ P \ \ Q) \ (s, {}) \ (\ P \ \ Q)" + show "\P Q. \s t. (s @ t, {}) \(\ P \ \ Q) \ (s, {}) \ (\ P \ \ Q)" by (auto simp: is_processT1 dest!: is_processT3[rule_format]) next - show "\s X Y. (s, Y) \ (\ P \ \ Q) \ X \ Y \ (s, X) \ (\ P \ \ Q)" + show "\P Q. \s X Y. (s, Y) \ (\ P \ \ Q) \ X \ Y \ (s, X) \ (\ P \ \ Q)" by(auto dest: is_processT4[rule_format,OF conj_commute[THEN iffD1], OF conjI]) next - show "\sa X Y. (sa, X) \ (\ P \ \ Q) \ (\c. c \ Y \ (sa @ [c], {}) \ (\ P \ \ Q)) + show "\P Q. \sa X Y. (sa, X) \ (\ P \ \ Q) \ (\c. c \ Y \ (sa @ [c], {}) \ (\ P \ \ Q)) \ (sa, X \ Y) \ (\ P \ \ Q)" by(auto simp: is_processT5 T_F) next - show "\s X. (s @ [tick], {}) \ (\ P \ \ Q) \ (s, X - {tick}) \ (\ P \ \ Q)" + show "\P Q. \s X. (s @ [tick], {}) \ (\ P \ \ Q) \ (s, X - {tick}) \ (\ P \ \ Q)" apply((rule allI)+, rule impI) apply(rename_tac s X, case_tac "s=[]", simp_all add: is_processT6[rule_format] T_F_spec) apply(erule disjE,simp_all add: is_processT6[rule_format] T_F_spec) apply(erule disjE,simp_all add: is_processT6[rule_format] T_F_spec) done next - show "\s t. s \ \ P \ \ Q \ tickFree s \ front_tickFree t \ s @ t \ \ P \ \ Q" + show "\P Q. \s t. s \ \ P \ \ Q \ tickFree s \ front_tickFree t \ s @ t \ \ P \ \ Q" by(auto simp: is_processT7) next - show "\s X. s \ \ P \ \ Q \ (s, X) \ (\ P \ \ Q)" + show "\P Q. \s X. s \ \ P \ \ Q \ (s, X) \ (\ P \ \ Q)" by(auto simp: is_processT8[rule_format]) next - show "\s. s @ [tick] \ \ P \ \ Q \ s \ \ P \ \ Q" + show "\P Q. \s. s @ [tick] \ \ P \ \ Q \ s \ \ P \ \ Q" by(auto intro!:is_processT9[rule_format]) -qed + qed + +notation + Ndet (infixl "\" 80) -lemma Rep_Abs_Ndet: -"Rep_process(Abs_process(\ P \ \ Q , \ P \ \ Q)) = (\ P \ \ Q , \ P \ \ Q)" - by(simp only:CollectI Rep_process Abs_process_inverse is_process_REP_Ndet) +subsection \The Projections\ + +lemma F_Ndet : "\ (P \ Q) = \ P \ \ Q" + by (simp add: FAILURES_def Failures.rep_eq Ndet.rep_eq) + +lemma D_Ndet : "\ (P \ Q) = \ P \ \ Q" + by (simp add: DIVERGENCES_def Divergences.rep_eq Ndet.rep_eq) + +lemma T_Ndet : "\ (P \ Q) = \ P \ \ Q" + apply (subst Traces.rep_eq, simp add: TRACES_def Failures.rep_eq[symmetric] F_Ndet) + apply(auto simp: T_F F_T is_processT1 Nil_elem_T) + by(rule_tac x="{}" in exI, simp add: T_F F_T is_processT1 Nil_elem_T)+ -subsection\Some Laws\ +subsection\Basic Laws\ text \The commutativity of the operator helps to simplify the subsequent continuity proof and is therefore developed here: \ -lemma F_Ndet : "\ (P \ Q) = \ P \ \ Q" - apply (subst Failures_def) - apply (simp only: Ndet_def) - by (simp add: FAILURES_def Rep_Abs_Ndet) - - -lemma D_Ndet : "\(P \ Q) = \ P \ \ Q" - by(subst D_def, simp only: Ndet_def Rep_Abs_Ndet DIVERGENCES_def snd_conv) - - -lemma T_Ndet : "\ (P \ Q) = \ P \ \ Q" -apply(subst Traces_def, simp only: Ndet_def Rep_Abs_Ndet TRACES_def FAILURES_def - fst_def snd_conv) -apply(auto simp: T_F F_T is_processT1 Nil_elem_T) -apply(rule_tac x="{}" in exI, simp add: T_F F_T is_processT1 Nil_elem_T)+ -done lemma Ndet_commute: "(P \ Q) = (Q \ P)" -by(auto simp: Process_eq_spec F_Ndet D_Ndet) + by(auto simp: Process_eq_spec F_Ndet D_Ndet) + subsection\The Continuity Rule\ lemma mono_Ndet1: "P \ Q \ \ (Q \ S) \ \ (P \ S)" apply(drule le_approx1) by (auto simp: Process_eq_spec F_Ndet D_Ndet) lemma mono_Ndet2: "P \ Q \ (\ s. s \ \ (P \ S) \ Ra (P \ S) s = Ra (Q \ S) s)" apply(subst (asm) le_approx_def) by (auto simp: Process_eq_spec F_Ndet D_Ndet Ra_def) lemma mono_Ndet3: "P \ Q \ min_elems (\ (P \ S)) \ \ (Q \ S)" apply(auto dest!:le_approx3 simp: min_elems_def subset_iff F_Ndet D_Ndet T_Ndet) apply(erule_tac x="t" in allE, auto) by (erule_tac x="[]" in allE, auto simp: less_list_def Nil_elem_T D_T) lemma mono_Ndet[simp] : "P \ Q \ (P \ S) \ (Q \ S)" by(auto simp:le_approx_def mono_Ndet1 mono_Ndet2 mono_Ndet3) lemma mono_Ndet_sym[simp] : "P \ Q \ (S \ P) \ (S \ Q)" by (auto simp: Ndet_commute) lemma cont_Ndet1: assumes chain:"chain Y" shows "((\ i. Y i) \ S) = (\ i. (Y i \ S))" proof - have A : "chain (\i. Y i \ S)" apply(insert chain,rule chainI) apply(frule_tac i="i" in chainE) by(simp) show ?thesis using chain by(auto simp add: limproc_is_thelub Process_eq_spec D_Ndet F_Ndet F_LUB D_LUB A) qed lemma Ndet_cont[simp]: assumes f: "cont f" and g: "cont g" shows "cont (\x. f x \ g x)" proof - have A:"\x. cont f \ cont g \ cont (\X. (f x)\ X)" apply(rule contI2, rule monofunI) apply(auto) apply(subst Ndet_commute, subst cont_Ndet1) by (auto simp:Ndet_commute) have B:"\y. cont f \ cont g \ cont (\x. f x \ y)" apply(rule_tac c="(\ g. g \ y)" in cont_compose) apply(rule contI2,rule monofunI) by (simp_all add: cont_Ndet1) show ?thesis using f g by (rule_tac f="(\ x. (\ g. f x \ g))" in cont_apply, auto simp: A B) qed end diff --git a/thys/HOL-CSP/Process.thy b/thys/HOL-CSP/Process.thy --- a/thys/HOL-CSP/Process.thy +++ b/thys/HOL-CSP/Process.thy @@ -1,1293 +1,1306 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * - * Author : Benoit Ballenghien, Safouan Taha, Burkhart Wolff & Lina Ye + * Author : Benoit Ballenghien, Safouan Taha, Burkhart Wolff * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * Copyright (c) 2023 Université Paris-Saclay, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) chapter\The Notion of Processes\ text\As mentioned earlier, we base the theory of CSP on HOLCF, a Isabelle/HOL library providing a theory of continuous functions, fixpoint induction and recursion.\ theory Process imports HOLCF begin text\Since HOLCF sets the default type class to @{class "cpo"}, while our Process theory establishes links between standard types and @{class "pcpo"} types. Consequently, we reset the default type class to the default in HOL.\ default_sort type section\Pre-Requisite: Basic Traces and tick-Freeness\ text\The denotational semantics of CSP assumes a distinguishable special event, called \verb+tick+ and written $\checkmark$, that is required to occur only in the end of traces in order to signalize successful termination of a process. (In the original text of Hoare, this treatment was more liberal and lead to foundational problems: the process invariant could not be established for the sequential composition operator of CSP; see @{cite "tej.ea:corrected:1997"} for details.)\ datatype '\ event = ev '\ | tick type_synonym '\ trace = "('\ event) list" text\We chose as standard ordering on traces the prefix ordering.\ instantiation list :: (type) order begin definition le_list_def : "(s::'a list) \ t \ (\ r. s @ r = t)" definition less_list_def: "(s::'a list) < t \ s \ t \ s \ t" lemma A : "((x::'a list) < y) = (x \ y \ \ y \ x)" by(auto simp: le_list_def less_list_def) instance proof fix x y z ::"'a list" show "(x < y) = (x \ y \ \ y \ x)" by(auto simp: le_list_def less_list_def) show "x \ x" by(simp add: le_list_def) assume A:"x \ y" and B:"y \ z" thus "x \ z" apply(insert A B, simp add: le_list_def, safe) apply(rename_tac r ra, rule_tac x="r@ra" in exI, simp) done assume A:"x \ y" and B:"y \ x" thus "x = y" by(insert A B, auto simp: le_list_def) qed end text\Some facts on the prefix ordering.\ lemma nil_le[simp]: "[] \ s" by(induct "s", simp_all, auto simp: le_list_def) lemma nil_le2[simp]: "s \ [] = (s = [])" by(induct "s", auto simp:le_list_def) lemma nil_less[simp]: "\ t < []" by(simp add: less_list_def) lemma nil_less2[simp]: "[] < t @ [a]" by(simp add: less_list_def) lemma less_self[simp]: "t < t@[a]" by(simp add:less_list_def le_list_def) lemma le_length_mono: "s \ t \ length s \ length t" by(auto simp: le_list_def) lemma less_length_mono: "s < t \ length s < length t" by(auto simp: less_list_def le_list_def) lemma less_cons: "s < t \ a # s < a # t" by (simp add: le_list_def less_list_def) lemma less_append: "s < t \ a @ s < a @ t" by (simp add: le_list_def less_list_def) lemma less_tail: "s \ [] \ s < t \ tl s < tl t" by (auto simp add: le_list_def less_list_def) \\should be in the List library\ lemma list_nonMt_append: "s \ [] \ \ a t. s = t @ [a]" by(erule rev_mp,induct "s",simp_all,case_tac "s = []",auto) lemma append_eq_first_pref_spec[rule_format]: "s @ t = r @ [x] \ t \ [] \ s \ r" apply(rule_tac x=s in spec) apply(induct r,auto) apply(erule rev_mp) apply(rename_tac xa, rule_tac list=xa in list.induct, simp_all) apply(simp add: le_list_def) apply(drule list_nonMt_append, auto) done lemma prefixes_fin: "let prefixes = {t. \t2. x = t @ t2} in finite prefixes \ card prefixes = length x + 1" proof(induct x, simp) case (Cons a x) hence A:"finite {t. (\t2. x = t @ t2)}" using not_add_less2 by fastforce have B:"inj_on Cons UNIV" by (metis injI list.inject) from Cons A B inj_on_iff_eq_card have C:"card ((\x. a#x)`{t. (\t2. x = t @ t2)}) = length x + 1" by fastforce have D:"{t. \t2. a # x = t @ t2} = {[]} \ (\x. a#x)`{t. (\t2. x = t @ t2)}" by(intro set_eqI iffI, auto simp add:Cons_eq_append_conv) from C D card_insert_if[of "(\x. a#x)`{t. (\t2. x = t @ t2)}"] show ?case by (metis (no_types, lifting) One_nat_def Suc_eq_plus1 Un_insert_left finite.insertI finite_imageI image_iff list.distinct(1) list.size(4) local.A sup_bot.left_neutral) qed lemma sublists_fin: "finite {t. \t1 t2. x = t1 @ t @ t2}" proof(induct x, simp) case (Cons a x) have "{t. \t1 t2. a # x = t1 @ t @ t2} \ {t. \t1 t2. x = t1 @ t @ t2} \ {t. \t2. a#x = t @ t2}" apply(auto) by (metis Cons_eq_append_conv) with Cons prefixes_fin[of "a#x"] show ?case by (meson finite_UnI finite_subset) qed lemma suffixes_fin: "finite {t. \t1. x = t1 @ t}" apply(subgoal_tac "{t. \t1. x = t1 @ t} \ {t. \t1 t2. x = t1 @ t @ t2}") using infinite_super sublists_fin apply blast by blast text\For the process invariant, it is a key element to reduce the notion of traces to traces that may only contain one tick event at the very end. This is captured by the definition of the predicate \verb+front_tickFree+ and its stronger version \verb+tickFree+. Here is the theory of this concept.\ definition tickFree :: "'\ trace \ bool" where "tickFree s = (tick \ set s)" definition front_tickFree :: "'\ trace \ bool" where "front_tickFree s = (s =[] \ tickFree(tl(rev s)))" lemma tickFree_Nil [simp]: "tickFree []" by(simp add: tickFree_def) lemma tickFree_Cons [simp]: "tickFree (a # t) = (a \ tick \ tickFree t)" by(auto simp add: tickFree_def) -lemma tickFree_tl : "[|s ~= [] ; tickFree s|] ==> tickFree(tl s)" +lemma tickFree_tl : "tickFree s \ tickFree(tl s)" by(case_tac s, simp_all) lemma tickFree_append[simp]: "tickFree(s@t) = (tickFree s \ tickFree t)" by(simp add: tickFree_def member_def) lemma non_tickFree_tick [simp]: "\ tickFree [tick]" by(simp add: tickFree_def) lemma non_tickFree_implies_nonMt: "\ tickFree s \ s \ []" by(simp add:tickFree_def,erule rev_mp, induct s, simp_all) lemma tickFree_rev : "tickFree(rev t) = (tickFree t)" by(simp add: tickFree_def member_def) lemma front_tickFree_Nil[simp]: "front_tickFree []" by(simp add: front_tickFree_def) lemma front_tickFree_single[simp]: "front_tickFree [a]" by(simp add: front_tickFree_def) lemma tickFree_implies_front_tickFree: "tickFree s \ front_tickFree s" apply(simp add: tickFree_def front_tickFree_def member_def,safe) apply(erule contrapos_np, simp,(erule rev_mp)+) apply(rule_tac xs=s in List.rev_induct,simp_all) done lemma front_tickFree_charn: "front_tickFree s = (s = [] \ (\a t. s = t @ [a] \ tickFree t))" apply(simp add: front_tickFree_def) apply(cases "s=[]", simp_all) apply(drule list_nonMt_append, auto simp: tickFree_rev) done lemma front_tickFree_implies_tickFree: "front_tickFree (t @ [a]) \ tickFree t" by(simp add: tickFree_def front_tickFree_def member_def) lemma tickFree_implies_front_tickFree_single: "tickFree t \ front_tickFree (t @ [a])" by(simp add:front_tickFree_charn) lemma nonTickFree_n_frontTickFree: "\\ tickFree s; front_tickFree s \ \ \t. s = t @ [tick]" apply(frule non_tickFree_implies_nonMt) apply(drule front_tickFree_charn[THEN iffD1], auto) done lemma front_tickFree_dw_closed : "front_tickFree (s @ t) \ front_tickFree s" apply(erule rev_mp, rule_tac x= s in spec) apply(rule_tac xs=t in List.rev_induct, simp, safe) apply(rename_tac x xs xa) apply(simp only: append_assoc[symmetric]) apply(rename_tac x xs xa, erule_tac x="xa @ xs" in all_dupE) apply(drule front_tickFree_implies_tickFree) apply(erule_tac x="xa" in allE, auto) apply(auto dest!:tickFree_implies_front_tickFree) done lemma front_tickFree_append: "\ tickFree s; front_tickFree t\ \ front_tickFree (s @ t)" apply(drule front_tickFree_charn[THEN iffD1], auto) apply(erule tickFree_implies_front_tickFree) apply(subst append_assoc[symmetric]) apply(rule tickFree_implies_front_tickFree_single) apply(auto intro: tickFree_append) done lemma front_tickFree_mono: "front_tickFree (t @ r) \ r \ [] \ tickFree t \ front_tickFree r" by(metis append_assoc append_butlast_last_id front_tickFree_charn front_tickFree_implies_tickFree tickFree_append) +lemma tickFree_butlast: \tickFree s \ tickFree (butlast s) \ (s \ [] \ last s \ tick)\ + by (induct s, simp_all) + +lemma front_tickFree_butlast: \front_tickFree s \ tickFree (butlast s)\ + by (induct s, auto simp add: front_tickFree_def) + + section\ Basic Types, Traces, Failures and Divergences \ type_synonym '\ refusal = "('\ event) set" type_synonym '\ failure = "'\ trace \ '\ refusal" type_synonym '\ divergence = "'\ trace set" type_synonym '\ process\<^sub>0 = "'\ failure set \ '\ divergence" definition FAILURES :: "'\ process\<^sub>0 \ ('\ failure set)" where "FAILURES P = fst P" definition TRACES :: "'\ process\<^sub>0 \ ('\ trace set)" where "TRACES P = {tr. \ a. a \ FAILURES P \ tr = fst a}" definition DIVERGENCES :: "'\ process\<^sub>0 \ '\ divergence" where "DIVERGENCES P = snd P" definition REFUSALS :: "'\ process\<^sub>0 \ ('\ refusal set)" where "REFUSALS P = {ref. \ F. F \ FAILURES P \ F = ([],ref)}" section\ The Process Type Invariant \ definition is_process :: "'\ process\<^sub>0 \ bool" where "is_process P = (([],{}) \ FAILURES P \ (\ s X. (s,X) \ FAILURES P \ front_tickFree s) \ (\ s t . (s@t,{}) \ FAILURES P \ (s,{}) \ FAILURES P) \ (\ s X Y. (s,Y) \ FAILURES P & X <= Y \ (s,X) \ FAILURES P) \ (\ s X Y. (s,X) \ FAILURES P \ (\ c. c \ Y \ ((s@[c],{})\FAILURES P)) \ (s,X \ Y)\FAILURES P) \ (\ s X. (s@[tick],{}) : FAILURES P \ (s,X-{tick}) \ FAILURES P) \ (\ s t. s \ DIVERGENCES P \ tickFree s \ front_tickFree t \ s@t \ DIVERGENCES P) \ (\ s X. s \ DIVERGENCES P \ (s,X) \ FAILURES P) \ (\ s. s @ [tick] : DIVERGENCES P \ s \ DIVERGENCES P))" lemma is_process_spec: "is_process P = (([],{}) \ FAILURES P \ (\ s X. (s,X) \ FAILURES P \ front_tickFree s) \ (\ s t . (s @ t,{}) \ FAILURES P \ (s,{}) \ FAILURES P) \ (\ s X Y. (s,Y) \ FAILURES P \ \(X\Y) | (s,X) \ FAILURES P) \ (\ s X Y.(s,X) \ FAILURES P \ (\ c. c \ Y \ ((s@[c],{}) \ FAILURES P)) \(s,X \ Y) \ FAILURES P) \ (\ s X. (s@[tick],{}) \ FAILURES P \ (s,X - {tick}) \ FAILURES P) \ (\ s t. s \ DIVERGENCES P \ \tickFree s \ \front_tickFree t \ s @ t \ DIVERGENCES P) \ (\ s X. s \ DIVERGENCES P \ (s,X) \ FAILURES P) \ (\ s. s @ [tick] \ DIVERGENCES P \ s \ DIVERGENCES P))" by(simp only: is_process_def HOL.nnf_simps(1) HOL.nnf_simps(3) [symmetric] HOL.imp_conjL[symmetric]) lemma Process_eqI : assumes A: "FAILURES P = FAILURES Q " assumes B: "DIVERGENCES P = DIVERGENCES Q" shows "(P::'\ process\<^sub>0) = Q" apply(insert A B, unfold FAILURES_def DIVERGENCES_def) apply(rule_tac t=P in surjective_pairing[symmetric,THEN subst]) apply(rule_tac t=Q in surjective_pairing[symmetric,THEN subst]) apply(simp) done lemma process_eq_spec: "((P::'a process\<^sub>0) = Q) = (FAILURES P = FAILURES Q \ DIVERGENCES P = DIVERGENCES Q)" apply(auto simp: FAILURES_def DIVERGENCES_def) apply(rule_tac t=P in surjective_pairing[symmetric,THEN subst]) apply(rule_tac t=Q in surjective_pairing[symmetric,THEN subst]) apply(simp) done lemma process_surj_pair: "(FAILURES P,DIVERGENCES P) = P" by(auto simp:FAILURES_def DIVERGENCES_def) lemma Fa_eq_imp_Tr_eq: "FAILURES P = FAILURES Q \ TRACES P = TRACES Q" by(auto simp:FAILURES_def DIVERGENCES_def TRACES_def) lemma is_process1: "is_process P \ ([],{})\ FAILURES P " by(auto simp: is_process_def) lemma is_process2: "is_process P \ \ s X. (s,X) \ FAILURES P \ front_tickFree s " by(simp only: is_process_spec, metis) lemma is_process3: "is_process P \ \ s t. (s @ t,{}) \ FAILURES P \ (s, {}) \ FAILURES P" by(simp only: is_process_spec, metis) lemma is_process3_S_pref: "\is_process P; (t, {}) \ FAILURES P; s \ t\ \ (s, {}) \ FAILURES P" by(auto simp: le_list_def intro: is_process3 [rule_format]) lemma is_process4: "is_process P \ \s X Y. (s, Y) \ FAILURES P \ \ X \ Y \ (s, X) \ FAILURES P" by(simp only: is_process_spec, simp) lemma is_process4_S: "\is_process P; (s, Y) \ FAILURES P; X \ Y\ \ (s, X) \ FAILURES P" by(drule is_process4, auto) lemma is_process4_S1: "\is_process P; x \ FAILURES P; X \ snd x\ \ (fst x, X) \ FAILURES P" by(drule is_process4_S, auto) lemma is_process5: "is_process P \ \sa X Y. (sa, X) \ FAILURES P \ (\c. c \ Y \ (sa @ [c], {}) \ FAILURES P) \ (sa, X \ Y) \ FAILURES P" by(drule is_process_spec[THEN iffD1],metis) lemma is_process5_S: "\is_process P; (sa, X) \ FAILURES P; \c. c \ Y \ (sa @ [c], {}) \ FAILURES P\ \ (sa, X \ Y) \ FAILURES P" by(drule is_process5, metis) lemma is_process5_S1: "\is_process P; (sa, X) \ FAILURES P; (sa, X \ Y) \ FAILURES P\ \ \c. c \ Y \ (sa @ [c], {}) \ FAILURES P" by(erule contrapos_np, drule is_process5_S, simp_all) lemma is_process6: "is_process P \ \ s X. (s@[tick],{}) \ FAILURES P \ (s,X-{tick}) \ FAILURES P" by(drule is_process_spec[THEN iffD1], metis) lemma is_process6_S: "\is_process P ;(s@[tick],{}) \ FAILURES P\ \ (s,X-{tick}) \ FAILURES P" by(drule is_process6, metis) lemma is_process7: "is_process P \ \ s t. s \ DIVERGENCES P \ \ tickFree s \ \ front_tickFree t \ s @ t \ DIVERGENCES P" by(drule is_process_spec[THEN iffD1], metis) lemma is_process7_S: "\ is_process P;s : DIVERGENCES P;tickFree s;front_tickFree t\ \ s @ t \ DIVERGENCES P" by(drule is_process7, metis) lemma is_process8: "is_process P \ \ s X. s \ DIVERGENCES P \ (s,X) \ FAILURES P" by(drule is_process_spec[THEN iffD1], metis) lemma is_process8_S: "\ is_process P; s \ DIVERGENCES P \ \ (s,X) \ FAILURES P" by(drule is_process8, metis) lemma is_process9: "is_process P \ \ s. s@[tick] \ DIVERGENCES P \ s \ DIVERGENCES P" by(drule is_process_spec[THEN iffD1], metis) lemma is_process9_S: "\ is_process P;s@[tick] \ DIVERGENCES P \ \ s \ DIVERGENCES P" by(drule is_process9, metis) lemma Failures_implies_Traces: " \is_process P; (s, X) \ FAILURES P\ \ s \ TRACES P" by(simp add: TRACES_def, metis) lemma is_process5_sing: "\ is_process P ; (s,{x}) \ FAILURES P;(s,{}) \ FAILURES P\ \ (s @ [x],{}) \ FAILURES P" by(drule_tac X="{}" in is_process5_S1, auto) lemma is_process5_singT: "\ is_process P ; (s,{x}) \ FAILURES P;(s,{}) \ FAILURES P\ \ s @ [x] \ TRACES P" apply(drule is_process5_sing, auto) by(simp add: TRACES_def, auto) lemma front_trace_is_tickfree: assumes A: "is_process P" and B: "(t @ [tick],X) \ FAILURES P" shows "tickFree t" proof - have C: "front_tickFree(t @ [tick])" by(insert A B, drule is_process2, metis) show ?thesis by(rule front_tickFree_implies_tickFree[OF C]) qed lemma trace_with_Tick_implies_tickFree_front : "\is_process P; t @ [tick] \ TRACES P\ \ tickFree t" by(auto simp: TRACES_def intro: front_trace_is_tickfree) section \ The Abstraction to the process-Type \ typedef '\ process = "{p :: '\ process\<^sub>0 . is_process p}" proof - - have "({(s, X). s = []},{}) \ {p::'\ process\<^sub>0. is_process p}" - by(simp add: is_process_def front_tickFree_def - FAILURES_def TRACES_def DIVERGENCES_def ) - thus ?thesis by auto + have "({(s, X). s = []},{}) \ {p::'\ process\<^sub>0. is_process p}" + by(simp add: is_process_def front_tickFree_def + FAILURES_def TRACES_def DIVERGENCES_def ) + thus ?thesis by auto qed -definition Failures :: "'\ process \ ('\ failure set)" ("\") - where "\ P = FAILURES (Rep_process P)" - -definition Traces :: "'\ process \ ('\ trace set)" ("\") - where "\ P = TRACES (Rep_process P)" +setup_lifting type_definition_process -definition Divergences :: "'\ process \ '\ divergence" ("\") - where "\ P = DIVERGENCES (Rep_process P)" +text \This is where we differ from previous versions: we lift definitions + using Isabelle's machinery instead of doing it by hand.\ -lemmas D_def = Divergences_def +lift_definition Failures :: "'\ process \ ('\ failure set)" ("\") + is "\P. FAILURES P" . -definition R :: "'\ process \ ('\ refusal set)" ("\") - where "\ P = REFUSALS (Rep_process P)" + +lift_definition Traces :: "'\ process \ ('\ trace set)" ("\") + is "\P. TRACES P" . + + +lift_definition Divergences :: "'\ process \ '\ divergence" ("\") + is "\P. DIVERGENCES P" . + + +lift_definition Refusals :: "'\ process \ ('\ refusal set)" ("\") + is "\P. REFUSALS P" . + lemma is_process_Rep : "is_process (Rep_process P)" -by(rule_tac P=is_process in CollectD, rule Rep_process) + using Rep_process by blast + lemma Process_spec: "Abs_process (\ P, \ P) = P" -by(simp add: Failures_def FAILURES_def D_def - DIVERGENCES_def Rep_process_inverse) + by (simp add: Divergences.rep_eq Failures.rep_eq Rep_process_inverse process_surj_pair) + lemma Process_eq_spec: "(P = Q) = (\ P = \ Q \ \ P = \ Q)" - apply(rule iffI,simp) - apply(rule_tac t=P in Process_spec[THEN subst]) - apply(rule_tac t=Q in Process_spec[THEN subst]) - by simp + by (metis Process_spec) lemma Process_eq_spec_optimized: "(P = Q) = (\ P = \ Q \ (\ P = \ Q \ \ P = \ Q))" using Process_eq_spec by auto lemma is_processT: "([],{}) \ \ P \ (\ s X. (s,X) \ \ P \ front_tickFree s) \ (\ s t .(s@t,{}) \ \ P \ (s,{}) \ \ P) \ (\ s X Y.(s,Y) \ \ P \ (X\Y) \ (s,X) \ \ P) \ (\ s X Y.(s,X) \ \ P \ (\c. c \ Y \((s@[c],{})\ \ P)) \ (s,X \ Y) \ \ P) \ (\ s X. (s@[tick],{}) \ \ P \ (s, X-{tick}) \ \ P) \ (\ s t. s \ \ P \ tickFree s \ front_tickFree t \ s @ t \ \ P) \ (\ s X. s \ \ P \ (s,X) \ \ P) \ - (\ s. s@[tick] \ \ P \ s \ \ P)" -apply(simp only: Failures_def D_def Traces_def) -apply(rule is_process_def[THEN iffD1]) -apply(rule is_process_Rep) -done + (\ s. s@[tick] \ \ P \ s \ \ P)" + apply (simp add: Divergences.rep_eq Failures.rep_eq) + by (meson is_process_spec[of \Rep_process P\, simplified Rep_process[simplified]]) + + lemma process_charn: "([], {}) \ \ P \ (\s X. (s, X) \ \ P \ front_tickFree s) \ (\s t. (s @ t, {}) \ \ P \ (s, {}) \ \ P) \ (\s X Y. (s, Y) \ \ P \ \ X \ Y \ (s, X) \ \ P) \ (\s X Y. (s, X) \ \ P \ (\c. c \ Y \ (s @ [c], {}) \ \ P) \ (s, X \ Y) \ \ P) \ (\s X. (s @ [tick], {}) \ \ P \ (s, X - {tick}) \ \ P) \ (\s t. s \ \ P \ \ tickFree s \ \ front_tickFree t \ s @ t \ \ P) \ (\s X. s \ \ P \ (s, X) \ \ P) \ (\s. s @ [tick] \ \ P \ s \ \ P)" + (* by (meson is_processT) *) proof - { have A: "(\s t. (s @ t, {}) \ \ P \ (s, {}) \ \ P) = (\s t. (s @ t, {}) \ \ P \ (s, {}) \ \ P)" by metis have B : "(\s X Y. (s, Y) \ \ P \ \ X \ Y \ (s, X) \ \ P) = (\s X Y. (s, Y) \ \ P \ X \ Y \ (s, X) \ \ P) " by metis have C : "(\s t. s \ \ P \ \ tickFree s \ \ front_tickFree t \ s @ t \ \ P) = (\s t. s \ \ P \ tickFree s \ front_tickFree t \ s @ t \ \ P) " by metis have D:"(\s X. s \ \ P \ (s, X) \ \ P) = (\s X. s \ \ P \ (s, X) \ \ P)" by metis have E:"(\s. s @ [tick] \ \ P \ s \ \ P) = (\s. s @ [tick] \ \ P \ s \ \ P)" by metis note A B C D E } note a = this then show ?thesis apply(simp only: a) apply(rule is_processT) done qed text\ split of \verb+is_processT+: \ lemma is_processT1: "([], {}) \ \ P" by(simp add:process_charn) lemma is_processT2: "\s X. (s, X) \ \ P \ front_tickFree s" by(simp add:process_charn) lemma is_processT2_TR : "\s. s \ \ P \ front_tickFree s" -apply(simp add: Failures_def [symmetric] Traces_def TRACES_def, safe) -apply (drule is_processT2[rule_format], assumption) -done + apply (simp add: Traces.rep_eq Traces_def TRACES_def Failures.rep_eq[symmetric]) + using is_processT2 by blast + lemma is_proT2: assumes A : " (s, X) \ \ P" and B : " s \ []" shows "tick \ set (tl (rev s))" proof - have C: "front_tickFree s" by(insert A B, simp add: is_processT2) show ?thesis by(insert C,simp add: B tickFree_def front_tickFree_def) qed lemma is_processT3 : "\s t. (s @ t, {}) \ \ P \ (s, {}) \ \ P" by(simp only: process_charn HOL.nnf_simps(3), simp) lemma is_processT3_S_pref : "\(t, {}) \ \ P; s \ t\ \ (s, {}) \ \ P" apply(simp only: le_list_def, safe) apply(erule is_processT3[rule_format]) done lemma is_processT4 : "\s X Y. (s, Y) \ \ P \ X \ Y \ (s, X) \ \ P" by(insert process_charn [of P], metis) lemma is_processT4_S1 : "\x \ \ P; X \ snd x\ \ (fst x, X) \ \ P" apply(rule_tac Y = "snd x" in is_processT4[rule_format]) apply simp done lemma is_processT5: "\s X Y. (s, X) \ \ P \ (\c. c \ Y \ (s @ [c], {}) \ \ P) \ (s, X \ Y) \ \ P" by(simp add: process_charn) lemma is_processT5_S1: "(s, X) \ \ P \ (s, X \ Y) \ \ P \ \c. c \ Y \ (s @ [c], {}) \ \ P" by(erule contrapos_np, simp add: is_processT5[rule_format]) lemma is_processT5_S2: "(s, X) \ \ P \ (s @ [c], {}) \ \ P \ (s, X \ {c}) \ \ P" by(rule is_processT5[rule_format,OF conjI], metis, safe) lemma is_processT5_S2a: "(s, X) \ \ P \ (s, X \ {c}) \ \ P \ (s @ [c], {}) \ \ P" apply(erule contrapos_np) apply(rule is_processT5_S2) apply(simp_all) done lemma is_processT5_S3: assumes A: "(s, {}) \ \ P" and B: "(s @ [c], {}) \ \ P" shows "(s, {c}) \ \ P" proof - have C : " {c} = ({} Un {c})" by simp show ?thesis by(subst C, rule is_processT5_S2, simp_all add: A B) qed lemma is_processT5_S4: "(s, {}) \ \ P \ (s, {c}) \ \ P \ (s @ [c], {}) \ \ P" by(erule contrapos_np, simp add: is_processT5_S3) lemma is_processT5_S5: "(s, X) \ \ P \ \c. c \ Y \ (s, X \ {c}) \ \ P \ \c. c \ Y \ (s @ [c], {}) \ \ P" by(erule_tac Q = "\c. c \ Y \ (s, X \ {c}) \ \ P" in contrapos_pp, metis is_processT5_S2) lemma is_processT5_S6: "([], {c}) \ \ P \ ([c], {}) \ \ P" apply(rule_tac t="[c]" and s="[]@[c]" in subst, simp) apply(rule is_processT5_S4, simp_all add: is_processT1) done lemma is_processT6: "\s X. (s @ [tick], {}) \ \ P \ (s, X - {tick}) \ \ P" by(simp add: process_charn) lemma is_processT7: "\s t. s \ \ P \ tickFree s \ front_tickFree t \ s @ t \ \ P" by(insert process_charn[of P], metis) lemmas is_processT7_S = is_processT7[rule_format,OF conjI[THEN conjI, THEN conj_commute[THEN iffD1]]] lemma is_processT8: "\s X. s \ \ P \ (s, X) \ \ P" by(insert process_charn[of P], metis) lemmas is_processT8_S = is_processT8[rule_format] lemma is_processT8_Pair: "fst s \ \ P \ s \ \ P" apply(subst surjective_pairing) apply(rule is_processT8_S, simp) done lemma is_processT9: "\s. s @ [tick] \ \ P \ s \ \ P" by(insert process_charn[of P], metis) lemma is_processT9_S_swap: "s \ \ P \ s @ [tick] \ \ P" by(erule contrapos_nn,simp add: is_processT9[rule_format]) section\ Some Consequences of the Process Characterization\ lemma no_Trace_implies_no_Failure: "s \ \ P \ (s, {}) \ \ P" by(simp add: Traces_def TRACES_def Failures_def) lemmas NT_NF = no_Trace_implies_no_Failure lemma T_def_spec: "\ P = {tr. \a. a \ \ P \ tr = fst a}" by(simp add: Traces_def TRACES_def Failures_def) lemma F_T: "(s, X) \ \ P \ s \ \ P" by(simp add: T_def_spec split_def, metis) lemma F_T1: "a \ \ P \ fst a \ \ P" by(rule_tac X="snd a" in F_T,simp) lemma T_F: "s \ \ P \ (s, {}) \ \ P" apply(auto simp: T_def_spec) apply(drule is_processT4_S1, simp_all) done lemmas is_processT4_empty [elim!]= F_T [THEN T_F] lemma NF_NT: "(s, {}) \ \ P \ s \ \ P" by(erule contrapos_nn, simp only: T_F) lemma is_processT6_S1: "tick \ X \ (s @ [tick], {}) \ \ P \ (s, X) \ \ P" by(subst Diff_triv[of X "{tick}", symmetric], simp, erule is_processT6[rule_format]) lemmas is_processT3_ST = T_F [THEN is_processT3[rule_format,THEN F_T]] lemmas is_processT3_ST_pref = T_F [THEN is_processT3_S_pref [THEN F_T]] lemmas is_processT3_SR = F_T [THEN T_F [THEN is_processT3[rule_format]]] lemmas D_T = is_processT8_S [THEN F_T] lemma D_T_subset : "\ P \ \ P" by(auto intro!:D_T) lemma NF_ND : "(s, X) \ \ P \ s \ \ P" by(erule contrapos_nn, simp add: is_processT8_S) lemmas NT_ND = D_T_subset[THEN Set.contra_subsetD] lemma T_F_spec : "((t, {}) \ \ P) = (t \ \ P)" by(auto simp:T_F F_T) lemma is_processT5_S7: "t \ \ P \ (t, A) \ \ P \ \x. x \ A \ t @ [x] \ \ P" apply(erule contrapos_np, simp) apply(rule is_processT5[rule_format, OF conjI,of _ "{}", simplified]) apply(auto simp: T_F_spec) done lemma is_processT5_S7': "(t, X) \ \ P \ (t, X \ A) \ \ P \ \x. x \ A \ x \ X \ t @ [x] \ \ P" apply(erule contrapos_np, simp, subst Un_Diff_cancel[of X A, symmetric]) apply(rule is_processT5[rule_format]) apply(auto simp: T_F_spec) done lemma Nil_subset_T: "{[]} \ \ P" by(auto simp: T_F_spec[symmetric] is_processT1) lemma Nil_elem_T: "[] \ \ P" by(simp add: Nil_subset_T[THEN subsetD]) lemmas D_imp_front_tickFree = is_processT8_S[THEN is_processT2[rule_format]] lemma D_front_tickFree_subset : "\ P \ Collect front_tickFree" by(auto simp: D_imp_front_tickFree) lemma F_D_part : "\ P = {(s, x). s \ \ P} \ {(s, x). s \ \ P \ (s, x) \ \ P}" by(auto intro:is_processT8_Pair) lemma D_F : "{(s, x). s \ \ P} \ \ P" by(auto intro:is_processT8_Pair) lemma append_T_imp_tickFree: "t @ s \ \ P \ s \ [] \ tickFree t" by(frule is_processT2_TR[rule_format], simp add: front_tickFree_def tickFree_rev) corollary append_single_T_imp_tickFree : "t @ [a] \ \ P \ tickFree t" by (simp add: append_T_imp_tickFree) lemma F_subset_imp_T_subset: "\ P \ \ Q \ \ P \ \ Q" by(auto simp: subsetD T_F_spec[symmetric]) lemma is_processT6_S2: "tick \ X \ [tick] \ \ P \ ([], X) \ \ P" by(erule is_processT6_S1, simp add: T_F_spec) lemma is_processT9_tick: "[tick] \ \ P \ front_tickFree s \ s \ \ P" apply(rule append.simps(1) [THEN subst, of _ s]) apply(rule is_processT7_S, simp_all) apply(rule is_processT9 [rule_format], simp) done lemma T_nonTickFree_imp_decomp: "t \ \ P \ \ tickFree t \ \s. t = s @ [tick]" by(auto elim: is_processT2_TR[rule_format] nonTickFree_n_frontTickFree) section\ Process Approximation is a Partial Ordering, a Cpo, and a Pcpo \ text\The Failure/Divergence Model of CSP Semantics provides two orderings: The \emph{approximation ordering} (also called \emph{process ordering}) will be used for giving semantics to recursion (fixpoints) over processes, the \emph{refinement ordering} captures our intuition that a more concrete process is more deterministic and more defined than an abstract one. We start with the key-concepts of the approximation ordering, namely the predicates $min\_elems$ and $Ra$ (abbreviating \emph{refusals after}). The former provides just a set of minimal elements from a given set of elements of type-class $ord$ \ldots \ definition min_elems :: "('s::ord) set \ 's set" where "min_elems X = {s \ X. \t. t \ X \ \ (t < s)}" lemma Nil_min_elems : "[] \ A \ [] \ min_elems A" by(simp add: min_elems_def) lemma min_elems_le_self[simp] : "(min_elems A) \ A" by(auto simp: min_elems_def) lemmas elem_min_elems = Set.set_mp[OF min_elems_le_self] lemma min_elems_Collect_ftF_is_Nil : "min_elems (Collect front_tickFree) = {[]}" apply(auto simp: min_elems_def le_list_def) apply(drule front_tickFree_charn[THEN iffD1]) apply(auto dest!: tickFree_implies_front_tickFree) done lemma min_elems5 : assumes A: "(x::'a list) \ A" shows "\s\x. s \ min_elems A" proof - have * : "!! (x::'a list) (A::'a list set) (n::nat). x \ A \ length x \ n \ (\s\x. s \ min_elems A)" apply(rule_tac x=x in spec) apply(rule_tac n=n in nat_induct) (* quirk in induct *) apply(auto simp: Nil_min_elems) apply(case_tac "\ y. y \ A \ y < x",auto) apply(rename_tac A na x y, erule_tac x=y in allE, simp) apply(erule impE,drule less_length_mono, arith) apply(safe, rename_tac s, rule_tac x=s in exI,simp) apply(rule_tac x=x in exI, simp add:min_elems_def) done show ?thesis by(rule_tac n="length x" in *[rule_format],simp add:A) qed lemma min_elems4: "A \ {} \ \s. (s :: 'a trace) \ min_elems A" by(auto dest: min_elems5) lemma min_elems_charn: "t \ A \ \ t' r. t = (t' @ r) \ t' \ min_elems A" by(drule min_elems5[simplified le_list_def], auto) lemmas min_elems_ex = min_elems_charn (* Legacy *) lemma min_elems_no: "(x::'a list) \ min_elems A \ t \ A \ t \ x \ x = t" by (metis (no_types, lifting) less_list_def mem_Collect_eq min_elems_def) text\ \ldots while the second returns the set of possible refusal sets after a given trace $s$ and a given process $P$: \ definition Ra :: "['\ process, '\ trace] \ ('\ refusal set)" ("\\<^sub>a") where "\\<^sub>a P s = {X. (s, X) \ \ P}" text\ In the following, we link the process theory to the underlying fixpoint/domain theory of HOLCF by identifying the approximation ordering with HOLCF's pcpo's. \ instantiation process :: ("type") below begin text\ declares approximation ordering $\_ \sqsubseteq \_$ also written \verb+_ << _+. \ definition le_approx_def : "P \ Q \ \ Q \ \ P \ (\s. s \ \ P \ \\<^sub>a P s = \\<^sub>a Q s) \ min_elems (\ P) \ \ Q" text\ The approximation ordering captures the fact that more concrete processes should be more defined by ordering the divergence sets appropriately. For defined positions in a process, the failure sets must coincide pointwise; moreover, the minimal elements (wrt.~prefix ordering on traces, i.e.~lists) must be contained in the trace set of the more concrete process.\ instance .. end lemma le_approx1: "P\Q \ \ Q \ \ P" by(simp add: le_approx_def) lemma le_approx2: "\ P\Q; s \ \ P\ \ (s,X) \ \ Q = ((s,X) \ \ P)" by(auto simp: Ra_def le_approx_def) lemma le_approx3: "P \ Q \ min_elems(\ P) \ \ Q" by(simp add: le_approx_def) lemma le_approx2T: "\ P\Q; s \ \ P\ \ s \ \ Q = (s \ \ P)" by(auto simp: le_approx2 T_F_spec[symmetric]) lemma le_approx_lemma_F : "P\Q \ \ Q \ \ P" apply(subst F_D_part[of Q], subst F_D_part[of P]) apply(auto simp:le_approx_def Ra_def min_elems_def) done lemmas order_lemma = le_approx_lemma_F lemma le_approx_lemma_T: "P\Q \ \ Q \ \ P" by(auto dest!:le_approx_lemma_F simp: T_F_spec[symmetric]) lemma proc_ord2a : " P \ Q \ s \ \ P \ ((s, X) \ \ P) = ((s, X) \ \ Q)" by(auto simp: le_approx_def Ra_def) instance process :: ("type") po proof fix P::"'\ process" show "P \ P" by(auto simp: le_approx_def min_elems_def elim: Process.D_T) next fix P Q ::"'\ process" assume A:"P \ Q" and B:"Q \ P" thus "P = Q" apply(insert A[THEN le_approx1] B[THEN le_approx1]) apply(insert A[THEN le_approx_lemma_F] B[THEN le_approx_lemma_F]) by(auto simp: Process_eq_spec) next fix P Q R ::"'\ process" assume A: "P \ Q" and B: "Q \ R" thus "P \ R" proof - have C : "\ R \ \ P" by(insert A[THEN le_approx1] B[THEN le_approx1], auto) have D : "\s. s \ \ P \ {X. (s, X) \ \ P} = {X. (s, X) \ \ R}" apply(rule allI, rule impI, rule set_eqI, simp) apply(frule A[THEN le_approx1, THEN Set.contra_subsetD]) apply(frule B[THEN le_approx1, THEN Set.contra_subsetD]) apply(drule A[THEN le_approx2], drule B[THEN le_approx2]) apply auto done have E : "min_elems (\ P) \ \ R" apply(insert B[THEN le_approx3] A[THEN le_approx3] ) apply(insert B[THEN le_approx_lemma_T] A[THEN le_approx1] ) apply(rule subsetI, simp add: min_elems_def, auto) apply(rename_tac x, case_tac "x \ \ Q") apply(drule_tac B = "\ R" and t=x in subset_iff[THEN iffD1,rule_format], auto) apply(subst B [THEN le_approx2T],simp) apply(rename_tac x, drule_tac B = "\ Q" and t=x in subset_iff[THEN iffD1,rule_format],auto) done show ?thesis by(insert C D E, simp add: le_approx_def Ra_def) qed qed text\ At this point, we inherit quite a number of facts from the underlying HOLCF theory, which comprises a library of facts such as \verb+chain+, \verb+directed+(sets), upper bounds and least upper bounds, etc. \ text\ Some facts from the theory of complete partial orders: \begin{itemize} \item \verb+Porder.chainE+ : @{thm "Porder.chainE"} \item \verb+Porder.chain_mono+ : @{thm "Porder.chain_mono"} \item \verb+Porder.is_ubD+ : @{thm "Porder.is_ubD"} \item \verb+Porder.ub_rangeI+ : \\ @{thm "Porder.ub_rangeI"} \item \verb+Porder.ub_imageD+ : @{thm "Porder.ub_imageD"} \item \verb+Porder.is_ub_upward+ : @{thm "Porder.is_ub_upward"} \item \verb+Porder.is_lubD1+ : @{thm "Porder.is_lubD1"} \item \verb+Porder.is_lubI+ : @{thm "Porder.is_lubI"} \item \verb+Porder.is_lub_maximal+ : @{thm "Porder.is_lub_maximal"} \item \verb+Porder.is_lub_lub+ : @{thm "Porder.is_lub_lub"} \item \verb+Porder.is_lub_range_shift+: \\ @{thm "Porder.is_lub_range_shift"} \item \verb+Porder.is_lub_rangeD1+: @{thm "Porder.is_lub_rangeD1"} \item \verb+Porder.lub_eqI+: @{thm "Porder.lub_eqI"} \item \verb+Porder.is_lub_unique+:@{thm "Porder.is_lub_unique"} \end{itemize} \ definition lim_proc :: "('\ process) set \ '\ process" where "lim_proc (X) = Abs_process (\ (\ ` X), \ (\ ` X))" lemma min_elems3: " s @ [c] \ \ P \ s @ [c] \ min_elems (\ P) \ s \ \ P" apply(auto simp: min_elems_def le_list_def less_list_def) apply(rename_tac t r) apply(subgoal_tac "t \ s") apply(subgoal_tac "r \ []") apply(simp add: le_list_def) apply(auto intro!: is_processT7_S append_eq_first_pref_spec) apply(auto dest!: D_T) apply(simp_all only: append_assoc[symmetric], drule append_T_imp_tickFree, simp_all add: tickFree_implies_front_tickFree)+ done lemma min_elems1: "s \ \ P \ s @ [c] \ \ P \ s @ [c] \ min_elems (\ P)" by(erule contrapos_np, auto elim!: min_elems3) lemma min_elems2: "s \ \ P \ s @ [c] \ \ P \ P \ S \ Q \ S \ (s @ [c], {}) \ \ Q" apply(frule_tac P=Q and Q=S in le_approx_lemma_T) apply(simp add: T_F_spec) apply(rule_tac A="\ S" in subsetD, assumption) apply(rule_tac A="min_elems(\ P)" in subsetD) apply(simp_all add: le_approx_def min_elems1) done lemma min_elems6: "s \ \ P \ s @ [c] \ \ P \ P \ S \ (s @ [c], {}) \ \ S" by(auto intro!: min_elems2) lemma ND_F_dir2: "s \ \ P \ (s, {}) \ \ P \ P \ S \ Q \ S \ (s, {}) \ \ Q" apply(frule_tac P=Q and Q=S in le_approx_lemma_T) apply(simp add: le_approx_def Ra_def T_F_spec, safe) apply((erule_tac x=s in allE)+,simp) apply(drule_tac x="{}" in eqset_imp_iff, auto simp: T_F_spec) done \\orig version\ lemma ND_F_dir2': "s \ \ P \ s \ \ P \ P \ S \ Q \ S \ s \ \ Q" apply(frule_tac P=Q and Q=S in le_approx_lemma_T) apply(simp add: le_approx_def Ra_def T_F_spec, safe) apply((erule_tac x=s in allE)+,simp) apply(drule_tac x="{}" in eqset_imp_iff, auto simp: T_F_spec) done lemma chain_lemma: "\chain S\ \ S i \ S k \ S k \ S i" by(case_tac "i \ k", auto intro:chain_mono chain_mono_less) lemma is_process_REP_LUB: assumes chain: "chain S" shows "is_process (\ (\ ` range S), \ (\ ` range S))" proof (auto simp: is_process_def) show "([], {}) \ FAILURES (\ a::nat. \ (S a), \ a::nat. \ (S a))" by(auto simp: FAILURES_def is_processT) next fix s::"'a trace" fix X::"'a event set" assume A : "(s, X) \ (FAILURES (\ a :: nat. \ (S a), \ a :: nat. \ (S a)))" thus "front_tickFree s" by(auto simp: DIVERGENCES_def FAILURES_def intro!: is_processT2[rule_format]) next fix s t::"'a trace" assume " (s @ t, {}) \ FAILURES (\ a::nat. \ (S a), \ a::nat. \ (S a)) " thus "(s, {}) \ FAILURES (\ a::nat. \ (S a), \ a::nat. \ (S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro: is_processT3[rule_format]) next fix s::"'a trace" fix X Y ::"'a event set" assume "(s, Y) \ FAILURES (\ a::nat. \ (S a), \ a::nat. \(S a))" and "X \ Y" thus "(s, X) \ FAILURES (\ a::nat. \ (S a), \ a::nat. \(S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro: is_processT4[rule_format]) next fix s::"'a trace" fix X Y ::" 'a event set" assume A:"(s, X) \ FAILURES (\ a::nat. \ (S a), \ a::nat. \ (S a))" assume B:"\ c. c\Y \ (s@[c],{})\FAILURES(\ a::nat. \(S a),\ a::nat. \(S a))" thus "(s, X \ Y) \ FAILURES (\ a::nat. \ (S a), \ a::nat. \ (S a))" txt\ What does this mean: All trace prolongations $c$ in all $Y$, which are blocking in the limit, will also occur in the refusal set of the limit. \ using A B chain proof (auto simp: DIVERGENCES_def FAILURES_def, case_tac "\ x. x \ (range S) \ (s, X \ Y) \ \ x", simp_all add:DIVERGENCES_def FAILURES_def,rename_tac a, case_tac "s \ \ (S a)",simp_all add: is_processT8) fix a::nat assume X: "\a. (s, X) \ \ (S a)" have X_ref_at_a: "(s, X) \ \ (S a)" using X by auto assume Y: "\c. c \ Y \ (\a. (s @ [c], {}) \ \ (S a))" assume defined: "s \ \ (S a)" show "(s::'a trace, X \ Y) \ \ (S a)" proof(auto simp:X_ref_at_a intro!: is_processT5[rule_format], frule Y[THEN spec, THEN mp], erule exE, erule_tac Q="(s @ [c], {}) \ \ (S a)" in contrapos_pp) fix c::"'a event" fix a' :: nat assume s_c_trace_not_trace_somewhere: "(s @ [c], {}) \ \ (S a')" show "(s @ [c], {}) \ \ (S a)" proof(insert chain_lemma[OF chain, of "a" "a'"],erule disjE) assume before: "S a \ S a'" show "(s @ [c], {}) \ \ (S a)" using s_c_trace_not_trace_somewhere before apply(case_tac "s @ [c] \ \ (S a)", simp_all add: T_F_spec before[THEN le_approx2T,symmetric]) apply(erule contrapos_nn) apply(simp only: T_F_spec[symmetric]) apply(auto dest!:min_elems6[OF defined]) done next assume after:"S a' \ S a" show "(s @ [c], {}) \ \ (S a)" using s_c_trace_not_trace_somewhere by(simp add:T_F_spec after[THEN le_approx2T] s_c_trace_not_trace_somewhere[THEN NF_ND]) qed qed qed next fix s::"'a trace" fix X::"'a event set" assume "(s @ [tick], {}) \ FAILURES (\ a::nat. \ (S a), \ a::nat. \ (S a))" thus "(s, X - {tick}) \ FAILURES (\ a::nat. \ (S a), \ a::nat. \ (S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro! : is_processT6[rule_format]) next fix s t ::"'a trace" assume "s : DIVERGENCES (\ a::nat. \ (S a), \ a::nat. \ (S a))" and "tickFree s" and " front_tickFree t" thus "s @ t \ DIVERGENCES (\ a::nat. \ (S a), \ a::nat. \ (S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro: is_processT7[rule_format]) next fix s::"'a trace" fix X::"'a event set" assume "s \ DIVERGENCES (\ a::nat. \ (S a), \ a::nat. \ (S a)) " thus "(s, X) \ FAILURES (\ a::nat. \ (S a), \ a::nat. \ (S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro: is_processT8[rule_format]) next fix s::"'a trace" assume "s @ [tick] \ DIVERGENCES (\ a::nat. \ (S a), \ a::nat. \ (S a)) " thus "s \ DIVERGENCES (\ a::nat. \ (S a), \ a::nat. \ (S a))" by(auto simp: DIVERGENCES_def FAILURES_def intro: is_processT9[rule_format]) qed lemmas Rep_Abs_LUB = Abs_process_inverse[simplified Rep_process, simplified, OF is_process_REP_LUB, simplified] lemma F_LUB: "chain S \ \(lim_proc(range S)) = \ (\ ` range S)" by(simp add: lim_proc_def , subst Failures_def, auto simp: FAILURES_def Rep_Abs_LUB) lemma D_LUB: "chain S \ \(lim_proc(range S)) = \ (\ ` range S)" -by(simp add: lim_proc_def , subst D_def, auto simp: DIVERGENCES_def Rep_Abs_LUB) +by(simp add: lim_proc_def , subst Divergences_def, auto simp: DIVERGENCES_def Rep_Abs_LUB) lemma T_LUB: "chain S \ \ (lim_proc(range S)) = \ (\ ` range S)" apply(simp add: lim_proc_def , subst Traces_def) apply(simp add: TRACES_def FAILURES_def Rep_Abs_LUB) apply(auto intro: F_T, rule_tac x="{}" in exI, auto intro: T_F) done schematic_goal D_LUB_2: "chain S \ t \ \(lim_proc(range S)) = ?X" apply(subst D_LUB, simp) apply(rule trans, simp) apply(simp) done schematic_goal T_LUB_2: "chain S \ (t \ \ (lim_proc (range S))) = ?X" apply(subst T_LUB, simp) apply(rule trans, simp) apply(simp) done section\ Process Refinement is a Partial Ordering\ text\ The following type instantiation declares the refinement order $\_ \le \_ $ written \verb+_ <= _+. It captures the intuition that more concrete processes should be more deterministic and more defined.\ instantiation process :: (type) ord begin definition le_ref_def : "P \ Q \ \ Q \ \ P \ \ Q \ \ P" definition less_ref_def : "(P::'a process) < Q \ P \ Q \ P \ Q" instance .. end text\Note that this just another syntax to our standard process refinement order defined in the theory Process. \ lemma le_approx_implies_le_ref: "(P::'\ process) \ Q \ P \ Q" by(simp add: le_ref_def le_approx1 le_approx_lemma_F) lemma le_ref1: "P \ Q \ \ Q \ \ P" by(simp add: le_ref_def) lemma le_ref2: "P\Q \ \ Q \ \ P" by(simp add: le_ref_def) lemma le_ref2T : "P\Q \ \ Q \ \ P" by (rule subsetI) (simp add: T_F_spec[symmetric] le_ref2[THEN subsetD]) instance process :: (type) order proof fix P Q R :: "'\ process" { show "(P < Q) = (P \ Q \ \ Q \ P)" by(auto simp: le_ref_def less_ref_def Process_eq_spec) next show "P \ P" by(simp add: le_ref_def) next assume "P \ Q" and "Q \ R" then show "P \ R" by (simp add: le_ref_def, auto) next assume "P \ Q" and "Q \ P" then show "P = Q" by(auto simp: le_ref_def Process_eq_spec) } qed lemma lim_proc_is_ub:"chain S \ range S <| lim_proc (range S)" apply(auto simp: is_ub_def le_approx_def F_LUB D_LUB T_LUB Ra_def) using chain_lemma is_processT8 le_approx2 apply blast using D_T chain_lemma le_approx2T le_approx_def by blast lemma lim_proc_is_lub1: "chain S \ \ u . (range S <| u \ \ u \ \ (lim_proc (range S)))" by(auto simp: D_LUB, frule_tac i=a in Porder.ub_rangeD, auto dest: le_approx1) lemma lim_proc_is_lub2: "chain S \ \ u . range S <| u \ (\ s. s \ \ (lim_proc (range S)) \ Ra (lim_proc (range S)) s = Ra u s)" apply(auto simp: is_ub_def D_LUB F_LUB Ra_def) using proc_ord2a apply blast using is_processT8_S proc_ord2a by blast lemma lim_proc_is_lub3a: "front_tickFree s \ s \ \ P \ (\t. t \ \ P \ \ t < s @ [c])" apply(rule impI, erule contrapos_np, auto) -apply(auto simp: le_list_def less_list_def) -by (metis D_def butlast_append butlast_snoc - front_tickFree_mono is_process7 is_process_Rep self_append_conv) + apply(auto simp: le_list_def less_list_def) + by (metis butlast_append butlast_snoc front_tickFree_mono is_processT self_append_conv) + lemma lim_proc_is_lub3b: assumes 1 : "\x. x \ X \ (\xa. xa \ \ x \ (\t. t \ \ x \ \ t < xa) \ xa \ \ u)" and 2 : "xa \ X" and 3 : "\xa. xa \ X \ x \ \ xa" and 4 : "\t. (\x. x \ X \ t \ \ x) \ \ t < x" shows "x \ \ u" proof (cases "\t. t \ \ xa \ \ t < x") case True from \\t. t \ \ xa \ \ t < x\ show ?thesis using 1 2 3 by simp next case False from \\(\t. t \ \ xa \ \ t < x)\ and 3 4 have A: "\y r c. y \ X \ r \ \ y \ r \ \ xa \ r @ [c] = x" by (metis D_imp_front_tickFree front_tickFree_charn less_self lim_proc_is_lub3a nil_less tickFree_implies_front_tickFree) show ?thesis apply(insert A) apply(erule exE)+ using "1" "3" D_imp_front_tickFree lim_proc_is_lub3a by blast qed lemma lim_proc_is_lub3c: assumes *:"chain S" and **:"X = range S" \\protection for range - otherwise auto unfolds and gets lost\ shows "\ u. X <| u \ min_elems(\ (lim_proc X)) \ \ u" proof - have B : "\ (lim_proc X) = \ (\ ` X)" by(simp add: * ** D_LUB) show ?thesis apply(auto simp: is_ub_def * **) apply(auto simp: B min_elems_def le_approx_def HOL.imp_conjR HOL.all_conj_distrib Ball_def) apply(simp add: subset_iff imp_conjL[symmetric]) apply(rule lim_proc_is_lub3b[of "range S", simplified]) using "**" B by auto qed lemma limproc_is_lub: "chain S \ range S <<| lim_proc (range S)" apply (auto simp: is_lub_def lim_proc_is_ub) apply (simp add: le_approx_def is_lub_def lim_proc_is_ub) by (simp add: lim_proc_is_lub1 lim_proc_is_lub2 lim_proc_is_lub3c) lemma limproc_is_thelub: "chain S \ Lub S = lim_proc (range S)" by (frule limproc_is_lub,frule Porder.po_class.lub_eqI, simp) instance process :: (type) cpo proof fix S ::"nat \ '\ process" assume C:"chain S" thus "\ x. range S <<| x" using limproc_is_lub by blast qed instance process :: (type) pcpo proof - show "\ x::'a process. \ y::'a process. x \ y" + show "\ x::'a process. \ y::'a process. x \ y" proof - have is_process_witness : "is_process({(s,X). front_tickFree s},{d. front_tickFree d})" apply(auto simp:is_process_def FAILURES_def DIVERGENCES_def) apply(auto elim!: tickFree_implies_front_tickFree front_tickFree_dw_closed front_tickFree_append) done have bot_inverse : "Rep_process(Abs_process({(s, X). front_tickFree s},Collect front_tickFree))= ({(s, X). front_tickFree s}, Collect front_tickFree)" by(subst Abs_process_inverse, simp_all add: Rep_process is_process_witness) have divergences_frontTickFree: "\y x. x \ snd (Rep_process y) \ front_tickFree x" - by(rule D_imp_front_tickFree, simp add: D_def DIVERGENCES_def) + by(rule D_imp_front_tickFree, simp add: Divergences_def DIVERGENCES_def) have failures_frontTickFree: "\y s x. (s, x) \ fst (Rep_process y) \ front_tickFree s" by(rule is_processT2[rule_format], simp add: Failures_def FAILURES_def) have minelems_contains_mt: "\y x. x \ min_elems (Collect front_tickFree) \ x = []" by(simp add: min_elems_def front_tickFree_charn,safe, auto simp: Nil_elem_T) show ?thesis apply(rule_tac x="Abs_process ({(s,X). front_tickFree s},{d. front_tickFree d})" in exI) apply(auto simp: le_approx_def bot_inverse Ra_def - Failures_def D_def FAILURES_def DIVERGENCES_def + Failures_def Divergences_def FAILURES_def DIVERGENCES_def divergences_frontTickFree failures_frontTickFree) apply (metis minelems_contains_mt Nil_elem_T ) done qed qed section\ Process Refinement is Admissible \ lemma le_adm[simp]: "cont (u::('a::cpo) \ 'b process) \ monofun v \ adm(\x. u x \ v x)" proof(auto simp add:le_ref_def cont2contlubE adm_def, goal_cases) case (1 Y x) hence "v (Y i) \ v (\i. Y i)" for i by (simp add: is_ub_thelub monofunE) hence "\ (v (\i. Y i)) \ \ (u (Y i))" for i using "1"(4) le_approx1 by blast then show ?case using D_LUB[OF ch2ch_cont[OF 1(1) 1(3)]] limproc_is_thelub[OF ch2ch_cont[OF 1(1) 1(3)]] "1"(5) by force next case (2 Y a b) hence "v (Y i) \ v (\i. Y i)" for i by (simp add: is_ub_thelub monofunE) hence "\ (v (\i. Y i)) \ \ (u (Y i))" for i using "2"(4) le_approx_lemma_F by blast then show ?case using F_LUB[OF ch2ch_cont[OF 2(1) 2(3)]] limproc_is_thelub[OF ch2ch_cont[OF 2(1) 2(3)]] "2"(5) by force qed lemmas le_adm_cont[simp] = le_adm[OF _ cont2mono] section\ The Conditional Statement is Continuous \ text\The conditional operator of CSP is obtained by a direct shallow embedding. Here we prove it continuous\ lemma if_then_else_cont[simp]: assumes *:"(\x. P x \ cont ((f::'c \ ('a::cpo) \ 'b process) x))" and **:"(\x. \ P x \ cont (g x))" shows "\x. cont(\y. if P x then f x y else g x y)" using * ** by (auto simp:cont_def) end diff --git a/thys/HOL-CSP/Process_Order.thy b/thys/HOL-CSP/Process_Order.thy --- a/thys/HOL-CSP/Process_Order.thy +++ b/thys/HOL-CSP/Process_Order.thy @@ -1,200 +1,200 @@ (*<*) \\ ******************************************************************** * Project : CSP-RefTK - A Refinement Toolkit for HOL-CSP * Version : 1.0 * - * Author : Burkhart Wolff, Safouan Taha, Lina Ye, Benoît Ballenghien, Catherine Dubois + * Author : Burkhart Wolff, Safouan Taha, Lina Ye, Benoît Ballenghien * * This file : A Normalization Theory * * Copyright (c) 2022 Université Paris-Saclay, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) chapter \ Refinements \ (* find a better name *) theory Process_Order imports Process Stop begin definition trace_refine :: \'a process \ 'a process \ bool\ (infix \\\<^sub>T\ 60) where \P \\<^sub>T Q \ \ Q \ \ P\ definition failure_refine :: \'a process \ 'a process \ bool\ (infix \\\<^sub>F\ 60) where \P \\<^sub>F Q \ \ Q \ \ P\ definition divergence_refine :: \'a process \ 'a process \ bool\ (infix \\\<^sub>D\ 53) where \P \\<^sub>D Q \ \ Q \ \ P\ definition failure_divergence_refine :: \'a process \ 'a process \ bool\ (infix \\\<^sub>F\<^sub>D\ 60) where \P \\<^sub>F\<^sub>D Q \ P \ Q\ definition trace_divergence_refine :: \'a process \ 'a process \ bool\ (infix \\\<^sub>D\<^sub>T\ 53) (* an experimental order considered also in our theories*) where \P \\<^sub>D\<^sub>T Q \ P \\<^sub>T Q \ P \\<^sub>D Q\ section \Idempotency\ lemma idem_F[simp]: \P \\<^sub>F P\ and idem_D[simp]: \P \\<^sub>D P\ and idem_T[simp]: \P \\<^sub>T P\ and idem_FD[simp]: \P \\<^sub>F\<^sub>D P\ and idem_DT[simp]: \P \\<^sub>D\<^sub>T P\ by (simp_all add: failure_refine_def divergence_refine_def trace_refine_def failure_divergence_refine_def trace_divergence_refine_def) section \Some obvious refinements\ lemma BOT_leF[simp]: \\ \\<^sub>F Q\ and BOT_leD[simp]: \\ \\<^sub>D Q\ and BOT_leT[simp]: \\ \\<^sub>T Q\ by (simp_all add: failure_refine_def le_approx_lemma_F trace_refine_def le_approx1 divergence_refine_def le_approx_lemma_T) section \Antisymmetry\ lemma FD_antisym: \P \\<^sub>F\<^sub>D Q \ Q \\<^sub>F\<^sub>D P \ P = Q\ by (simp add: failure_divergence_refine_def) lemma DT_antisym: \P \\<^sub>D\<^sub>T Q \ Q \\<^sub>D\<^sub>T P \ P = Q\ apply (simp add: trace_divergence_refine_def) oops section \Transitivity\ lemma trans_F: \P \\<^sub>F Q \ Q \\<^sub>F S \ P \\<^sub>F S\ and trans_D: \P \\<^sub>D Q \ Q \\<^sub>D S \ P \\<^sub>D S\ and trans_T: \P \\<^sub>T Q \ Q \\<^sub>T S \ P \\<^sub>T S\ and trans_FD: \P \\<^sub>F\<^sub>D Q \ Q \\<^sub>F\<^sub>D S \ P \\<^sub>F\<^sub>D S\ and trans_DT: \P \\<^sub>D\<^sub>T Q \ Q \\<^sub>D\<^sub>T S \ P \\<^sub>D\<^sub>T S\ by (meson failure_refine_def order_trans, meson divergence_refine_def order_trans, meson trace_refine_def order_trans, meson failure_divergence_refine_def order_trans, meson divergence_refine_def order_trans trace_divergence_refine_def trace_refine_def) section \Relations between refinements\ lemma leF_imp_leT: \P \\<^sub>F Q \ P \\<^sub>T Q\ and leFD_imp_leF: \P \\<^sub>F\<^sub>D Q \ P \\<^sub>F Q\ and leFD_imp_leD: \P \\<^sub>F\<^sub>D Q \ P \\<^sub>D Q\ and leDT_imp_leD: \P \\<^sub>D\<^sub>T Q \ P \\<^sub>D Q\ and leDT_imp_leT: \P \\<^sub>D\<^sub>T Q \ P \\<^sub>T Q\ and leF_leD_imp_leFD: \P \\<^sub>F Q \ P \\<^sub>D Q \ P \\<^sub>F\<^sub>D Q\ and leD_leT_imp_leDT: \P \\<^sub>D Q \ P \\<^sub>T Q \ P \\<^sub>D\<^sub>T Q\ by (simp_all add: F_subset_imp_T_subset failure_refine_def trace_refine_def divergence_refine_def trace_divergence_refine_def failure_divergence_refine_def le_ref_def) section \More obvious refinements\ lemma BOT_leFD[simp]: \\ \\<^sub>F\<^sub>D Q\ and BOT_leDT[simp]: \\ \\<^sub>D\<^sub>T Q\ by (simp_all add: leF_leD_imp_leFD leD_leT_imp_leDT) lemma leDT_STOP[simp]: \P \\<^sub>D\<^sub>T STOP\ by (simp add: D_STOP leD_leT_imp_leDT Nil_elem_T T_STOP divergence_refine_def trace_refine_def) lemma leD_STOP[simp]: \P \\<^sub>D STOP\ and leT_STOP[simp]: \P \\<^sub>T STOP\ by (simp_all add: leDT_imp_leD leDT_imp_leT) section \Admissibility\ lemma le_F_adm[simp]: \cont (u::('a::cpo) \ 'b process) \ monofun v \ adm(\x. u x \\<^sub>F v x)\ proof(auto simp add:cont2contlubE adm_def failure_refine_def) fix Y a b assume 1: \cont u\ and 2: \monofun v\ and 3: \chain Y\ and 4: \\i. \ (v (Y i)) \ \ (u (Y i))\ and 5: \(a, b) \ \ (v (\x. Y x))\ hence \v (Y i) \ v (\i. Y i)\ for i by (simp add: is_ub_thelub monofunE) hence \\ (v (\i. Y i)) \ \ (u (Y i))\ for i using 4 le_approx_lemma_F by blast then show \(a, b) \ \ (\i. u (Y i))\ using F_LUB[OF ch2ch_cont[OF 1 3]] limproc_is_thelub[OF ch2ch_cont[OF 1 3]] 5 by force qed lemma le_T_adm[simp]: \cont (u::('a::cpo) \ 'b process) \ monofun v \ adm(\x. u x \\<^sub>T v x)\ proof(auto simp add:cont2contlubE adm_def trace_refine_def) fix Y x assume 1: \cont u\ and 2: \monofun v\ and 3: \chain Y\ and 4: \\i. \ (v (Y i)) \ \ (u (Y i))\ and 5: \x \ \ (v (\i. Y i))\ hence \v (Y i) \ v (\i. Y i)\ for i by (simp add: is_ub_thelub monofunE) hence \\ (v (\i. Y i)) \ \ (u (Y i))\ for i using 4 le_approx_lemma_T by blast then show \x \ \ (\i. u (Y i))\ using T_LUB[OF ch2ch_cont[OF 1 3]] limproc_is_thelub[OF ch2ch_cont[OF 1 3]] 5 by force qed lemma le_D_adm[simp]: \cont (u::('a::cpo) \ 'b process) \ monofun v \ adm(\x. u x \\<^sub>D v x)\ proof(auto simp add:cont2contlubE adm_def divergence_refine_def) fix Y x assume 1: \cont u\ and 2: \monofun v\ and 3: \chain Y\ and 4: \\i. \ (v (Y i)) \ \ (u (Y i))\ and 5: \x \ \ (v (\i. Y i))\ hence \v (Y i) \ v (\i. Y i)\ for i by (simp add: is_ub_thelub monofunE) hence \\ (v (\i. Y i)) \ \ (u (Y i))\ for i using 4 le_approx1 by blast then show \x \ \ (\i. u (Y i))\ using D_LUB[OF ch2ch_cont[OF 1 3]] limproc_is_thelub[OF ch2ch_cont[OF 1 3]] 5 by force qed lemmas le_FD_adm[simp] = le_adm[folded failure_divergence_refine_def] lemma le_DT_adm[simp]: \cont (u::('a::cpo) \ 'b process) \ monofun v \ adm(\x. u x \\<^sub>D\<^sub>T v x)\ using adm_conj[OF le_T_adm[of u v] le_D_adm[of u v]] by (simp add: trace_divergence_refine_def) end \ No newline at end of file diff --git a/thys/HOL-CSP/ROOT b/thys/HOL-CSP/ROOT --- a/thys/HOL-CSP/ROOT +++ b/thys/HOL-CSP/ROOT @@ -1,68 +1,67 @@ chapter AFP (***************************************************************************** * * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * * Author : Benoit Ballenghien, Safouan Taha, Burkhart Wolff & Lina Ye * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * Copyright (c) 2023 Université Paris-Saclay, France * * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************) (* $Id:$ *) -session "HOL-CSP" = HOLCF + +session "HOL-CSP" (AFP) = HOLCF + description \HOL-CSP based on the original Roscoe-Book.\ - options [timeout = 1200] + options [document = pdf, document_output = "output", timeout = 1200] sessions "HOL-Eisbach" theories "Introduction" "Process" - Bot Skip Stop Det Ndet Seq Hiding Sync Mndetprefix Mprefix - "CSP" "Assertions" + Bot Skip Stop Det Ndet Seq Renaming Hiding Sync Mndetprefix Mprefix + "CSP_Induct" "CSP" "Assertions" "Conclusion" "CopyBuffer" - Induction_ext document_files "root.tex" "adb-long.bib" "root.bib" diff --git a/thys/HOL-CSP/Renaming.thy b/thys/HOL-CSP/Renaming.thy new file mode 100644 --- /dev/null +++ b/thys/HOL-CSP/Renaming.thy @@ -0,0 +1,580 @@ +(*<*) +\\ ******************************************************************** + * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL + * Version : 2.0 + * + * Author : Benoît Ballenghien. + * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) + * + * This file : A Combined CSP Theory + * + * Copyright (c) 2009 Université Paris-Sud, France + * + * All rights reserved. + * + * Redistribution and use in source and binary forms, with or without + * modification, are permitted provided that the following conditions are + * met: + * + * * Redistributions of source code must retain the above copyright + * notice, this list of conditions and the following disclaimer. + * + * * Redistributions in binary form must reproduce the above + * copyright notice, this list of conditions and the following + * disclaimer in the documentation and/or other materials provided + * with the distribution. + * + * * Neither the name of the copyright holders nor the names of its + * contributors may be used to endorse or promote products derived + * from this software without specific prior written permission. + * + * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS + * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT + * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR + * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT + * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, + * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT + * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, + * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY + * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT + * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE + * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. + ******************************************************************************\ +(*>*) + +section\ The Renaming Operator \ + +theory Renaming + imports Process +begin + +subsection\Some preliminaries\ + +definition EvExt where \EvExt f x \ case_event (ev o f) tick x\ + + +term \f -` B\ (* instead of defining our own antecedent function *) + +definition finitary :: \('a \ 'b) \ bool\ + where \finitary f \ \x. finite(f -` {x})\ + + + +text \We start with some simple results.\ + +lemma \f -` {} = {}\ by simp + +lemma \X \ Y \ f -` X \ f -` Y\ by (rule vimage_mono) + +lemma \f -`(X \ Y) = f -` X \ f -` Y\ by (rule vimage_Un) + + +lemma EvExt_id: \EvExt id = id\ + unfolding id_def EvExt_def by (metis comp_apply event.exhaust event.simps(4, 5)) + +lemma EvExt_eq_tick: \EvExt f a = tick \ a = tick\ + by (metis EvExt_def comp_apply event.exhaust event.simps(3, 4, 5)) + +lemma tick_eq_EvExt: \tick = EvExt f a \ a = tick\ + by (metis EvExt_eq_tick) + +lemma EvExt_ev1: \EvExt f b = ev a \ (\c. b = ev c \ EvExt f (ev c) = ev a)\ + by (metis event.exhaust event.simps(3) tick_eq_EvExt) + +lemma EvExt_tF: \tickFree (map (EvExt f) s) \ tickFree s\ + unfolding tickFree_def by (auto simp add: tick_eq_EvExt image_iff) + +lemma inj_EvExt: \inj EvExt\ + unfolding inj_on_def EvExt_def + by auto (metis comp_eq_dest_lhs event.simps(4, 5) ext) + + + + +lemma EvExt_ftF: \front_tickFree (map (EvExt f) s) \ front_tickFree s\ + unfolding front_tickFree_def by safe (metis (no_types, opaque_lifting) EvExt_tF map_tl rev_map)+ + + +lemma map_EvExt_tick: \[tick] = map (EvExt f) t \ t = [tick]\ + using tick_eq_EvExt by fastforce + + +lemma anteced_EvExt_diff_tick: \EvExt f -` (X - {tick}) = EvExt f -` X - {tick}\ + by (rule subset_antisym) (simp_all add: EvExt_eq_tick subset_Diff_insert subset_vimage_iff) + + + +lemma ev_elem_anteced1: \ev a \ EvExt f -` A \ ev (f a) \ A\ + and tick_elem_anteced1: \tick \ EvExt f -` A \ tick \ A\ + unfolding EvExt_def by simp_all + + +lemma hd_map_EvExt: \t \ [] \ hd t = ev a \ hd (map (EvExt f) t) = ev (f a)\ + and tl_map_EvExt: \t \ [] \ tl (map (EvExt f) t) = map (EvExt f) (tl t)\ + unfolding EvExt_def by (simp_all add: list.map_sel(1) map_tl) + + + +subsection\The Renaming Operator Definition\ + +lift_definition Renaming :: \['a process, 'a \ 'b] \ 'b process\ + is \\P f. ({(s, R). \ s1. (s1, (EvExt f) -` R) \ \ P \ s = map (EvExt f) s1} \ + {(s ,R). \ s1 s2. tickFree s1 \ front_tickFree s2 \ s = (map (EvExt f) s1) @ s2 \ s1 \ \ P}, + {t. \ s1 s2. tickFree s1 \ front_tickFree s2 \ t = (map (EvExt f) s1) @ s2 \ s1 \ \ P})\ +proof - + show "is_process ({(s, R). \ s1. (s1, (EvExt f) -` R) \ \ P \ s = map (EvExt f) s1} \ + {(s ,R). \ s1 s2. tickFree s1 \ front_tickFree s2 + \ s = (map (EvExt (f :: 'a \ 'b)) s1) @ s2 \ s1 \ \ P}, + {t. \ s1 s2. tickFree s1 \ front_tickFree s2 + \ t = (map (EvExt f) s1) @ s2 \ s1 \ \ P})" + (is "is_process(?f, ?d)") for P f +unfolding is_process_def FAILURES_def DIVERGENCES_def + proof (simp only: fst_conv snd_conv, intro conjI allI impI) + show \([], {}) \ ?f\ + by (simp add: process_charn) + next + show \(s, X) \ ?f \ front_tickFree s\ for s X + by (auto simp add: EvExt_ftF is_processT2 EvExt_tF front_tickFree_append) + next + show \(s @ t, {}) \ ?f \ (s, {}) \ ?f\ for s t + proof (induct t rule: rev_induct, simp, auto, + metis (no_types, lifting) is_processT3_SR map_eq_append_conv, goal_cases) + case (1 x t s1 s2) + show ?case + proof (rule "1"(1), auto) + assume a1: \\s1. tickFree s1 \ + (\s2. s @ t = map (EvExt f) s1 @ s2 \ front_tickFree s2 \ s1 \ \ P)\ + have \(if s2 = [] then butlast s1 else s1) \ \ P\ + apply (rule a1[rule_format, of _ \butlast s2\]) + using "1"(3, 4, 5) apply (auto simp add: tickFree_def front_tickFree_def + intro: in_set_butlastD) + apply ((metis append_assoc butlast_snoc map_butlast)+)[2] + apply (metis append_assoc butlast_append butlast_snoc) + by (metis butlast_rev list.set_sel(2) rev_is_Nil_conv rev_rev_ident) + with \s1 \ \ P\ have \s2 = []\ by presburger + with "1"(5, 6) show \\s1. (s1, {}) \ \ P \ s @ t = map (EvExt f) s1\ + apply (rule_tac x = \butlast s1\ in exI, intro conjI) + by (metis append_butlast_last_id butlast.simps(1) process_charn) + (metis butlast_append map_butlast snoc_eq_iff_butlast) + qed + qed + next + show \(s, Y) \ ?f \ X \ Y \ (s, X) \ ?f\ for s X Y + apply (induct s rule: rev_induct, simp_all) + by (meson is_processT4 vimage_mono) + (metis process_charn vimage_mono) + next + show \(s, X) \ ?f \ (\c. c \ Y \ (s @ [c], {}) \ ?f) \ (s, X \ Y) \ ?f\ for s X Y + by auto (metis (no_types, lifting) is_processT5 list.simps(8, 9) map_append vimageE) + next + show \(s @ [tick], {}) \ ?f \ (s, X - {tick}) \ ?f\ for s X + apply auto + by (metis (no_types, lifting) anteced_EvExt_diff_tick is_processT6 map_EvExt_tick + map_eq_append_conv) + (metis EvExt_tF append.assoc butlast_snoc front_tickFree_charn non_tickFree_tick + tickFree_Nil tickFree_append tickFree_implies_front_tickFree) + next + show \s \ ?d \ tickFree s \ front_tickFree t \ s @ t \ ?d\ for s t + using front_tickFree_append by safe auto + next + show \s \ ?d \ (s, X) \ ?f\ for s X by blast + + next + show \s @ [tick] \ ?d \ s \ ?d\ for s + apply (induct s, auto) + by (metis EvExt_tF Nil_is_map_conv hd_append2 list.exhaust_sel list.sel(1) tickFree_Cons) + (metis Cons_eq_appendI EvExt_tF append.assoc butlast_snoc front_tickFree_charn + non_tickFree_tick tickFree_Nil tickFree_append tickFree_implies_front_tickFree) + qed +qed + + +text \Some syntaxic sugar\ + +syntax + "_Renaming" :: \'a process \ updbinds \ 'a process\ (\_\_\\ [100, 100]) (*see the values we need, at least 51*) +translations + "_Renaming P updates" \ "CONST Renaming P (_Update (CONST id) updates)" + + +text \Now we can write \<^term>\P\a := b\\. But like in \<^theory>\HOL.Fun\, we can write this kind of + expression with as many updates we want: \<^term>\P\a := b, c := d, e := f, g := h\\. + By construction we also inherit all the results about \<^const>\fun_upd\, for example + @{thm fun_upd_other fun_upd_upd fun_upd_twist}.\ + +lemma \P\x := y, x := z\ = P\x := z\\ by simp + +lemma \a \ c \ P\a := b, c := d\ = P\c := d, a := b\\ by (simp add: fun_upd_twist) + + + + +subsection\The Projections\ + +lemma F_Renaming: \\ (Renaming P f) = + {(s, R). \s1. (s1, EvExt f -` R) \ \ P \ s = map (EvExt f) s1} \ + {(s, R). \s1 s2. tickFree s1 \ front_tickFree s2 \ s = map (EvExt f) s1 @ s2 \ s1 \ \ P}\ + by (simp add: Failures.rep_eq FAILURES_def Renaming.rep_eq) + +lemma D_Renaming: \\ (Renaming P f) = {t. \s1 s2. tickFree s1 \ front_tickFree s2 \ + t = map (EvExt f) s1 @ s2 \ s1 \ \ P}\ + by (simp add: Divergences.rep_eq DIVERGENCES_def Renaming.rep_eq) + +lemma T_Renaming: \\ (Renaming P f) = + {t. \ t1. t1 \ \ P \ t = map (EvExt f) t1} \ + {t.\ t1 t2. tickFree t1 \ front_tickFree t2 \ t = map (EvExt f) t1 @ t2 \ t1 \ \ P}\ + apply (subst Traces.rep_eq, auto simp add: TRACES_def Failures.rep_eq[symmetric] F_Renaming) + by (use F_T in blast) (metis T_F vimage_empty) + + + +subsection\Continuity Rule\ + +subsubsection \Monotonicity of \<^const>\Renaming\.\ + +lemma mono_Renaming[simp] : \P \ Q \ (Renaming P f) \ (Renaming Q f)\ +proof (subst le_approx_def, intro conjI) + show \P \ Q \ \ (Renaming Q f) \ \ (Renaming P f)\ + unfolding D_Renaming using le_approx1 by blast +next + show \P \ Q \ \s. s \ \ (Renaming P f) \ \\<^sub>a (Renaming P f) s = \\<^sub>a (Renaming Q f) s\ + apply (auto simp add: Ra_def D_Renaming F_Renaming) + apply (metis EvExt_ftF append.right_neutral front_tickFree_charn front_tickFree_single + le_approx2 map_append process_charn tickFree_Cons tickFree_Nil tickFree_append) + using le_approx_lemma_F by blast (metis le_approx_def subset_eq) +next + show \P \ Q \ min_elems (\ (Renaming P f)) \ \ (Renaming Q f)\ + proof (unfold min_elems_def, auto simp add: D_Renaming T_Renaming, goal_cases) + case (1 s1 s2) + obtain t2 where f1: \map (EvExt f) t2 = s2\ + by (metis "1"(3, 4, 5, 6) front_tickFree_charn less_append map_is_Nil_conv nil_less2) + with "1" show ?case + apply (rule_tac x = \s1 @ t2\ in exI, simp) + apply (rule le_approx3[simplified subset_iff, OF "1"(1), rule_format]) + apply (induct s2 rule: rev_induct, induct s1 rule: rev_induct, auto) + apply (simp add: Nil_min_elems) + apply (metis "1"(5) append.right_neutral f1 less_self list.simps(8) min_elems3) + by (metis append.assoc front_tickFree_dw_closed less_self) + qed +qed + + + + +lemma \{ev c |c. f c = a} = ev`{c . f c = a}\ by (rule setcompr_eq_image) + + +lemma vimage_EvExt_subset_vimage: \EvExt f -` (ev ` A) \ insert tick (ev ` (f -` A))\ + and vimage_subset_vimage_EvExt: \(ev ` (f -` A)) \ (EvExt f -` (ev ` A)) - {tick}\ + unfolding EvExt_def by auto (metis comp_apply event.exhaust event.simps(4) image_iff vimage_eq) + + + +subsubsection \Some useful results about \<^const>\finitary\, and preliminaries lemmas for continuity.\ + +lemma inj_imp_finitary[simp] : \inj f \ finitary f\ by (simp add: finitary_def finite_vimageI) + +lemma finitary_comp_iff : \finitary g \ finitary (g o f) \ finitary f\ +proof (unfold finitary_def, intro iffI allI; + (subst vimage_comp[symmetric] | subst (asm) vimage_comp[symmetric])) + have f1: \f -` g -` {x} = (\y \ g -` {x}. f -` {y})\ for x by blast + show \finite (f -` g -` {x})\ if \\x. finite (f -` {x})\ and \\x. finite (g -` {x})\ for x + apply (subst f1, subst finite_UN) + by (rule that(2)[unfolded finitary_def, rule_format]) + (intro ballI that(1)[rule_format]) +qed (metis finite_Un insert_absorb vimage_insert vimage_singleton_eq) + + + +lemma finitary_updated_iff[simp] : \finitary (f (a := b)) \ finitary f\ +proof (unfold fun_upd_def finitary_def vimage_def, intro iffI allI) + show \finite {x. (if x = a then b else f x) \ {y}}\ if \\y. finite {x. f x \ {y}}\ for y + apply (rule finite_subset[of _ \{x. x = a} \ {x. f x \ {y}}\], (auto)[1]) + apply (rule finite_UnI) + by simp (fact that[rule_format]) +next + show \finite {x. f x \ {y}}\ if \\y. finite {x. (if x = a then b else f x) \ {y}}\ for y + apply (rule finite_subset[of _ \{x. x = a \ f x \ {y}} \ {x. x \ a \ f x \ {y}}\], (auto)[1]) + apply (rule finite_UnI) + using that[rule_format, of y] by (simp_all add: Collect_mono rev_finite_subset) +qed + +text \By declaring this simp, we automatically obtain this kind of result.\ + +lemma \finitary f \ finitary (f (a := b, c := d, e:= g, h := i))\ by simp + + + +lemma Cont_RenH2: \finitary (EvExt f) \ finitary f\ +proof - + have inj_ev: \inj ev\ by (simp add: inj_def) + show \finitary (EvExt f) \ finitary f\ + apply (subst finitary_comp_iff[of ev f, symmetric], simp add: inj_ev) + proof (intro iffI; subst finitary_def, intro allI) + have inj_ev: \inj ev\ by (simp add: inj_def) + show \finite ((ev \ f) -` {x})\ if \finitary (EvExt f)\ for x + apply (fold vimage_comp) + proof (cases \x = tick\, (insert finite.simps)[1], blast) + assume \x \ tick\ + with subset_singletonD[OF inj_vimage_singleton[OF inj_ev, of x]] obtain y + where f1: \ev -` {x} = {y}\ by auto (metis empty_iff event.exhaust vimage_singleton_eq) + have f2: \(f -` {y}) = ev -` ev ` f -` {y}\ by (simp add: inj_ev inj_vimage_image_eq) + show \finite (f -` ev -` {x})\ + apply (subst f1, subst f2) + apply (rule finite_vimageI[OF _ inj_ev]) + apply (rule finite_subset[OF vimage_subset_vimage_EvExt]) + apply (rule finite_Diff) + using finitary_def that by fastforce + qed + next + show \finite (EvExt f -` {x})\ if \finitary (ev \ f)\ for x + proof (cases \x = tick\, + metis Diff_cancel anteced_EvExt_diff_tick finite.emptyI infinite_remove vimage_empty) + assume \x \ tick\ + with subset_singletonD[OF inj_vimage_singleton[OF inj_ev, of x]] obtain y + where f1: \{x} = ev ` {y}\ using event.exhaust by auto + show \finite (EvExt f -` {x})\ + apply (subst f1) + apply (rule finite_subset[OF vimage_EvExt_subset_vimage]) + apply (subst finite_insert) + apply (rule finite_imageI) + by (fact finitary_comp_iff[OF inj_imp_finitary[OF inj_ev], of f, + simplified that, simplified, unfolded finitary_def, rule_format]) + qed + qed +qed + + + +lemma Cont_RenH3: + \{s1 @ t1 |s1 t1. (\ b. s1 = [b] & f b = a) \ list = map f t1} = + (\ b \ f -`{a}. (\t. b # t) ` {t. list = map f t})\ + by auto (metis append_Cons append_Nil) + + +lemma Cont_RenH4: \finitary f \ (\s. finite {t. s = map f t})\ +proof (unfold finitary_def, intro iffI allI) + show \finite {t. s = map f t}\ if \\x. finite (f -` {x})\for s + proof (induct s, simp) + case (Cons a s) + have \{t. a # s = map f t} = (\b \ f -` {a}. {b # t |t. s = map f t})\ + by (subst Cons_eq_map_conv) blast + thus ?case by (simp add: finite_UN[OF that[rule_format]] local.Cons) + qed +next + have map1: \[a] = map f s = (\b. s = [b] \ f b = a)\ for a s by fastforce + show \finite (f -` {x}) \ if \\s. finite {t. s = map f t}\ for x + by (simp add: vimage_def) + (fact finite_vimageI[OF that[rule_format, of \[x]\, simplified map1], + of \\x. [x]\, unfolded inj_on_def, simplified]) +qed + + + +lemma Cont_RenH5: \finitary f \ finite (\t \ {t. t \ s}. {s. t=map (EvExt f) s})\ + apply (rule finite_UN_I) + unfolding le_list_def + apply (induct s rule: rev_induct, simp) + apply (subgoal_tac \{t. \r. t @ r = xs @ [x]} = insert (xs @ [x]) {t. \r. t @ r = xs}\, auto) + apply (metis butlast_append butlast_snoc self_append_conv) + using Cont_RenH2 Cont_RenH4 by blast + + + +lemma Cont_RenH6: \{t. \ t2. tickFree t \ front_tickFree t2 \ x = map (EvExt f) t @ t2} + \ {y. \xa. xa \ {t. t <= x} \ y \ {s. xa = map (EvExt f) s}}\ + by (auto simp add: le_list_def) + + + +lemma Cont_RenH7: \finite {t. \t2. tickFree t \ front_tickFree t2 \ s = map (EvExt f) t @ t2}\ + if \finitary f\ +proof - + have f1: \{y. map (EvExt f) y \ s} = (\z \ {z. z \ s}. {y. z = map (EvExt f) y})\ by fast + show ?thesis + apply (rule finite_subset[OF Cont_RenH6]) + apply (simp add: f1) + apply (rule finite_UN_I) + apply (induct s rule: rev_induct, simp) + apply (subgoal_tac \{z. z \ xs @ [x]} = insert (xs @ [x]) {z. z \ xs}\, auto) + apply (metis append_eq_first_pref_spec append_self_conv le_list_def) + apply (meson le_list_def order_trans) + using Cont_RenH2 Cont_RenH4 that by blast +qed + + +lemma chain_Renaming: \chain Y \ chain (\i. Renaming (Y i) f)\ + by (simp add: chain_def) + + + +text \A key lemma for the continuity.\ + +lemma Inter_nonempty_finite_chained_sets: + \\i. S i \ {} \ finite (S 0) \ \i. S (Suc i) \ S i \ (\i. S i) \ {}\ +proof (induct \card (S 0)\ arbitrary: S rule: nat_less_induct) + case 1 + show ?case + proof (cases \\i. S i = S 0\) + case True + thus ?thesis by (metis "1.prems"(1) INT_iff ex_in_conv) + next + case False + have f1: \i \ j \ S j \ S i\ for i j by (simp add: "1.prems"(3) lift_Suc_antimono_le) + with False obtain j m where f2: \m < card (S 0)\ and f3: \m = card (S j)\ + by (metis "1.prems"(2) psubsetI psubset_card_mono zero_le) + define T where \T i \ S (i + j)\ for i + have f4: \m = card (T 0)\ unfolding T_def by (simp add: f3) + from f1 have f5: \(\i. S i) = (\i. T i)\ unfolding T_def by (auto intro: le_add1) + show ?thesis + apply (subst f5) + apply (rule "1.hyps"[rule_format, OF f2, of T, OF f4], unfold T_def) + by (simp_all add: "1.prems"(1, 3) lift_Suc_antimono_le) + (metis "1.prems"(2) add_0 f1 finite_subset le_add1) + qed +qed + + + +subsubsection \Finally, continuity !\ + +lemma Cont_Renaming_prem: + \(\ i. Renaming (Y i) f) = (Renaming (\ i. Y i) f)\ if finitary: \finitary f\ + and chain: \chain Y\ +proof (subst Process_eq_spec, safe, goal_cases) + show \s \ \ (\i. Renaming (Y i) f) \ s \ \ (Renaming (Lub Y) f)\ for s + proof (simp add: limproc_is_thelub chain chain_Renaming D_Renaming D_LUB) + assume a1: \\i. \s1. tickFree s1 + \ (\s2. front_tickFree s2 + \ s = map (EvExt f) s1 @ s2 \ s1 \ \ (Y i))\ + define S where \S i \ {s1. \t2. tickFree s1 \ front_tickFree t2 + \ s = map (EvExt f) s1 @ t2 \ s1 \ \ (Y i)}\ for i + + have \(\i. S i) \ {}\ + apply (rule Inter_nonempty_finite_chained_sets, unfold S_def) + apply (insert a1, blast)[1] + apply (rule finite_subset[OF _ Cont_RenH7[OF finitary, of s]], blast) + using chain unfolding chain_def le_approx_def by blast + then obtain s1 where f5: \s1 \ (\i. S i)\ by blast + + thus \\s1. tickFree s1 \ (\s2. front_tickFree s2 + \ s = map (EvExt f) s1 @ s2 + \ (\i. s1 \ \ (Y i)))\ + by (rule_tac x = s1 in exI, unfold S_def) auto + qed +next + show \s \ \ (Renaming (Lub Y) f) \ s \ \ (\i. Renaming (Y i) f)\ for s + by (auto simp add: limproc_is_thelub chain chain_Renaming D_Renaming D_LUB) +next + show \(s, X) \ \ (\i. Renaming (Y i) f) \ (s, X) \ \ (Renaming (Lub Y) f)\ for s X + proof (auto simp add: limproc_is_thelub chain chain_Renaming F_Renaming D_LUB F_LUB) + assume a1: \\i. (\s1. (s1, EvExt f -` X) \ \ (Y i) \ s = map (EvExt f) s1) \ + (\s1. tickFree s1 \ + (\s2. front_tickFree s2 + \ s = map (EvExt f) s1 @ s2 + \ s1 \ \ (Y i)))\ + and a2: \\s1. tickFree s1 \ + (\s2. s = map (EvExt f) s1 @ s2 + \ front_tickFree s2 + \ (\i. s1 \ \ (Y i)))\ + define S where \S i \ {s1. (s1, EvExt f -` X) \ \ (Y i) \ s = map (EvExt f) s1} \ + {s1. \t2. tickFree s1 + \ front_tickFree t2 + \ s = map (EvExt f) s1 @ t2 + \ s1 \ \ (Y i)}\ for i + + have \s1 \ (\i. S i) \ (\i. (s1, EvExt f -` X) \ \ (Y i) \ s = map (EvExt f) s1) + \ tickFree s1 \ (\s2. front_tickFree s2 + \ s = map (EvExt f) s1 @ s2 + \ (\i. s1 \ \ (Y i)))\ for s1 + unfolding S_def by auto (meson NF_ND) + hence f1: \s1 \ (\i. S i) \ (\i. (s1, EvExt f -` X) \ \(Y i) \ s = map(EvExt f) s1)\ for s1 + by (meson a2) + have \(\i. S i) \ {}\ + apply (rule Inter_nonempty_finite_chained_sets, unfold S_def) + apply (insert a1, blast)[1] + apply (subst finite_Un, intro conjI) + apply (rule finite_subset[of _ \{s1. s = map (EvExt f) s1}\], blast) + apply (insert Cont_RenH2 Cont_RenH4 finitary, blast)[1] + apply (rule finite_subset[OF _ Cont_RenH7[OF finitary, of s]], blast) + using NF_ND chain po_class.chainE proc_ord2a chain le_approx_def by safe blast+ + + then obtain s1 where \s1 \ (\i. S i)\ by blast + thus \\s1. (\x. (s1, EvExt f -` X) \ \ (Y x)) \ s = map (EvExt f) s1\ + using f1 by (rule_tac x = s1 in exI) blast + qed +next + show \(s, X) \ \ (Renaming (Lub Y) f) \ (s, X) \ \ (\i. Renaming (Y i) f)\ for s X + by (auto simp add: limproc_is_thelub chain chain_Renaming F_Renaming D_LUB F_LUB) +qed + + +lemma Renaming_cont[simp] : \cont P \ finitary f \ cont (\x. (Renaming (P x) f))\ + by (rule contI2) + (simp add: cont2monofunE monofunI, simp add: Cont_Renaming_prem ch2ch_cont cont2contlubE) + + +lemma RenamingF_cont : \cont P \ cont (\x. (P x)\a := b\)\ + (* by simp *) + by (simp only: Renaming_cont finitary_updated_iff inj_imp_finitary inj_on_id) + + +lemma \cont P \ cont (\x. (P x)\a := b, c := d, e := f, g := a\)\ + (* by simp *) + apply (erule Renaming_cont) + apply (simp only: finitary_updated_iff) + apply (rule inj_imp_finitary) + by (rule inj_on_id) + + + + + +subsection \Some nice properties\ + + +lemma EvExt_comp: \EvExt (g \ f) = EvExt g \ EvExt f\ + unfolding EvExt_def by (rule ext, auto) (metis comp_apply event.exhaust event.simps(4, 5)) + + + +lemma Renaming_comp : \Renaming P (g o f) = Renaming (Renaming P f) g\ +proof (subst Process_eq_spec, intro conjI subset_antisym) + show \\ (Renaming P (g \ f)) \ \ (Renaming (Renaming P f) g)\ + apply (simp add: F_Renaming D_Renaming subset_iff, safe) + by (metis (no_types, opaque_lifting) EvExt_comp list.map_comp set.compositionality) + (metis (no_types, opaque_lifting) EvExt_comp EvExt_tF append.right_neutral + front_tickFree_Nil list.map_comp) +next + show \\ (Renaming (Renaming P f) g) \ \ (Renaming P (g \ f))\ + apply (auto simp add: F_Renaming D_Renaming EvExt_comp EvExt_ftF EvExt_tF front_tickFree_append) + by (metis (no_types, opaque_lifting) EvExt_comp list.map_comp set.compositionality) +next + show \\ (Renaming P (g \ f)) \ \ (Renaming (Renaming P f) g)\ + apply (simp add: D_Renaming subset_iff, safe) + by (metis (no_types, opaque_lifting) EvExt_comp EvExt_tF append.right_neutral + front_tickFree_Nil list.map_comp) +next + show \\ (Renaming (Renaming P f) g) \ \ (Renaming P (g \ f))\ + apply (auto simp add: D_Renaming subset_iff) + by (metis EvExt_comp EvExt_tF front_tickFree_append) +qed + + + + +lemma Renaming_id: \Renaming P id = P\ + by (subst Process_eq_spec, auto simp add: F_Renaming D_Renaming EvExt_id + is_processT7_S is_processT8_S) + (metis append.right_neutral front_tickFree_mono list.distinct(1) + nonTickFree_n_frontTickFree process_charn) + + +lemma Renaming_inv: \Renaming (Renaming P f) (inv f) = P\ if \inj f\ + apply (subst Renaming_comp[symmetric]) + apply (subst inv_o_cancel[OF that]) + by (fact Renaming_id) + + +end + diff --git a/thys/HOL-CSP/Seq.thy b/thys/HOL-CSP/Seq.thy --- a/thys/HOL-CSP/Seq.thy +++ b/thys/HOL-CSP/Seq.thy @@ -1,442 +1,402 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * - * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * Author : Burkhart Wolff, Safouan Taha. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) section\The Sequence Operator\ theory Seq imports Process begin subsection\Definition\ abbreviation "div_seq P Q \ {t1 @ t2 |t1 t2. t1 \ \ P \ tickFree t1 \ front_tickFree t2} \ {t1 @ t2 |t1 t2. t1 @ [tick] \ \ P \ t2 \ \ Q}" -definition Seq :: "['a process,'a process] \ 'a process" (infixl "\<^bold>;" 74) -where "P \<^bold>; Q \ Abs_process - ({(t, X). (t, X \ {tick}) \ \ P \ tickFree t} \ - {(t, X). \t1 t2. t = t1 @ t2 \ t1 @ [tick] \ \ P \ (t2, X) \ \ Q} \ - {(t, X). t \ div_seq P Q}, - div_seq P Q)" - -lemma Seq_maintains_is_process: - "is_process ({(t, X). (t, X \ {tick}) \ \ P \ tickFree t} \ - {(t, X). \t1 t2. t = t1 @ t2 \ t1 @ [tick] \ \ P \ (t2, X) \ \ Q} \ - {(t, X). t \ div_seq P Q}, - div_seq P Q)" - (is "is_process(?f, ?d)") -proof(simp only: fst_conv snd_conv Failures_def is_process_def FAILURES_def DIVERGENCES_def, - fold FAILURES_def DIVERGENCES_def,fold Failures_def,intro conjI) - show "([], {}) \ ?f" - apply(cases "[tick] \ \ P", simp_all add: is_processT1) - using F_T is_processT5_S6 by blast -next - show " \s X. (s, X) \ ?f \ front_tickFree s" - by (auto simp:is_processT2 append_T_imp_tickFree front_tickFree_append D_imp_front_tickFree) -next - show "\s t. (s @ t, {}) \ ?f \ (s, {}) \ ?f" - apply auto - apply (metis F_T append_Nil2 is_processT is_processT3_SR is_processT5_S3) - apply(simp add:append_eq_append_conv2) - apply (metis T_F_spec append_Nil2 is_processT is_processT5_S4) - apply (metis append_self_conv front_tickFree_append front_tickFree_mono is_processT2_TR - no_Trace_implies_no_Failure non_tickFree_implies_nonMt non_tickFree_tick) - apply (metis (mono_tags, lifting) F_T append_Nil2 is_processT5_S3 process_charn) - apply (metis front_tickFree_append front_tickFree_mono self_append_conv) - apply(simp add:append_eq_append_conv2) - apply (metis T_F_spec append_Nil2 is_processT is_processT5_S4) - by (metis D_imp_front_tickFree append_T_imp_tickFree front_tickFree_append - front_tickFree_mono not_Cons_self2 self_append_conv) -next - show "\s X Y. (s, Y) \ ?f \ X \ Y \ (s, X) \ ?f" - apply auto - apply (metis insert_mono is_processT4_S1 prod.sel(1) prod.sel(2)) - apply (metis is_processT4) - apply (simp add: append_T_imp_tickFree) - by (metis process_charn) -next - { fix sa X Y - have "(sa, X \ {tick}) \ \ P \ - tickFree sa \ - \c. c \ Y \ c \ tick \ (sa @ [c], {}) \ \ P \ - (sa, X \ Y \ {tick}) \ \ P \ tickFree sa" - apply(rule_tac t="X \ Y \ {tick}" and s="X \ {tick} \ (Y-{tick})" in subst,auto) - by (metis DiffE Un_insert_left is_processT5 singletonI) - } note is_processT5_SEQH3 = this - have is_processT5_SEQH4 : - "\ sa X Y. (sa, X \ {tick}) \ \ P - \ tickFree sa - \ \c. c \ Y \ (sa@[c],{tick}) \ \ P \ \ tickFree(sa@[c]) - \ \c. c \ Y \ (\t1 t2. sa@[c]\t1@t2 \ t1@[tick]\\ P \ (t2,{})\\ Q) - \ (sa, X \ Y \ {tick}) \ \ P \ tickFree sa" - by (metis append_Nil2 is_processT1 is_processT5_S3 is_processT5_SEQH3 - no_Trace_implies_no_Failure tickFree_Cons tickFree_append) - let ?f1 = "{(t, X). (t, X \ {tick}) \ \ P \ tickFree t} \ - {(t, X). \t1 t2. t = t1 @ t2 \ t1 @ [tick] \ \ P \ (t2, X) \ \ Q}" - have is_processT5_SEQ2: "\ sa X Y. (sa, X) \ ?f1 - \ (\c. c \ Y \ (sa@[c], {})\?f) - \ (sa, X\Y) \ ?f1" - apply (clarsimp,rule is_processT5_SEQH4[simplified]) - by (auto simp: is_processT5) - show "\s X Y. (s, X) \ ?f \ (\c. c \ Y \ (s @ [c], {}) \ ?f) \ (s, X \ Y) \ ?f" - apply(intro allI impI, elim conjE UnE) - apply(rule rev_subsetD) - apply(rule is_processT5_SEQ2) - apply auto - using is_processT5_S1 apply force - apply (simp add: append_T_imp_tickFree) - using is_processT5[rule_format, OF conjI] by force -next - show "\s X. (s @ [tick], {}) \ ?f \ (s, X - {tick}) \ ?f" - apply auto - apply (metis (no_types, lifting) append_T_imp_tickFree butlast_append - butlast_snoc is_processT2 is_processT6 nonTickFree_n_frontTickFree - non_tickFree_tick tickFree_append) - apply (metis append_T_imp_tickFree front_tickFree_append front_tickFree_mono - is_processT2 non_tickFree_implies_nonMt non_tickFree_tick) - apply (metis process_charn) - apply (metis front_tickFree_append front_tickFree_implies_tickFree) - apply (metis D_T T_nonTickFree_imp_decomp append_T_imp_tickFree append_assoc - append_same_eq non_tickFree_implies_nonMt non_tickFree_tick process_charn - tickFree_append) - by (metis D_imp_front_tickFree append_T_imp_tickFree front_tickFree_append - front_tickFree_mono non_tickFree_implies_nonMt non_tickFree_tick) -next - show "\s t. s \ ?d \ tickFree s \ front_tickFree t \ s @ t \ ?d" - apply auto - using front_tickFree_append apply blast - by (metis process_charn) -next - show "\s X. s \ ?d \ (s, X) \ ?f" - by blast -next - show "\s. s @ [tick] \ ?d \ s \ ?d" - apply auto - apply (metis append_Nil2 front_tickFree_implies_tickFree process_charn) - by (metis append1_eq_conv append_assoc front_tickFree_implies_tickFree is_processT2_TR - nonTickFree_n_frontTickFree non_tickFree_tick process_charn tickFree_append) +lift_definition Seq :: "['a process,'a process] \ 'a process" (infixl "\<^bold>;" 74) +is "\P Q. ({(t, X). (t, X \ {tick}) \ \ P \ tickFree t} \ + {(t, X). \t1 t2. t = t1 @ t2 \ t1 @ [tick] \ \ P \ (t2, X) \ \ Q} \ + {(t, X). t \ div_seq P Q}, + div_seq P Q)" +proof - + show "is_process ({(t, X). (t, X \ {tick}) \ \ (P :: 'a process) \ tickFree t} \ + {(t, X). \t1 t2. t = t1 @ t2 \ t1 @ [tick] \ \ P \ (t2, X) \ \ Q} \ + {(t, X). t \ div_seq P Q}, + div_seq P Q)" + (is "is_process(?f, ?d)") for P Q + unfolding is_process_def FAILURES_def DIVERGENCES_def + proof (simp only: fst_conv snd_conv, intro conjI) + show "([], {}) \ ?f" + apply(cases "[tick] \ \ P", simp_all add: is_processT1) + using F_T is_processT5_S6 by blast + next + show " \s X. (s, X) \ ?f \ front_tickFree s" + by (auto simp:is_processT2 append_T_imp_tickFree front_tickFree_append D_imp_front_tickFree) + next + show "\s t. (s @ t, {}) \ ?f \ (s, {}) \ ?f" + apply auto + apply (metis F_T append_Nil2 is_processT is_processT3_SR is_processT5_S3) + apply(simp add:append_eq_append_conv2) + apply (metis T_F_spec append_Nil2 is_processT is_processT5_S4) + apply (metis append_self_conv front_tickFree_append front_tickFree_mono is_processT2_TR + no_Trace_implies_no_Failure non_tickFree_implies_nonMt non_tickFree_tick) + apply (metis (mono_tags, lifting) F_T append_Nil2 is_processT5_S3 process_charn) + apply (metis front_tickFree_append front_tickFree_mono self_append_conv) + apply(simp add:append_eq_append_conv2) + apply (metis T_F_spec append_Nil2 is_processT is_processT5_S4) + by (metis D_imp_front_tickFree append_T_imp_tickFree front_tickFree_append + front_tickFree_mono not_Cons_self2 self_append_conv) + next + show "\s X Y. (s, Y) \ ?f \ X \ Y \ (s, X) \ ?f" + apply auto + apply (metis insert_mono is_processT4_S1 prod.sel(1) prod.sel(2)) + apply (metis is_processT4) + apply (simp add: append_T_imp_tickFree) + by (metis process_charn) + next + { fix sa X Y + have "(sa, X \ {tick}) \ \ P \ + tickFree sa \ + \c. c \ Y \ c \ tick \ (sa @ [c], {}) \ \ P \ + (sa, X \ Y \ {tick}) \ \ P \ tickFree sa" + apply(rule_tac t="X \ Y \ {tick}" and s="X \ {tick} \ (Y-{tick})" in subst,auto) + by (metis DiffE Un_insert_left is_processT5 singletonI) + } note is_processT5_SEQH3 = this + have is_processT5_SEQH4 : + "\ sa X Y. (sa, X \ {tick}) \ \ P + \ tickFree sa + \ \c. c \ Y \ (sa@[c],{tick}) \ \ P \ \ tickFree(sa@[c]) + \ \c. c \ Y \ (\t1 t2. sa@[c]\t1@t2 \ t1@[tick]\\ P \ (t2,{})\\ Q) + \ (sa, X \ Y \ {tick}) \ \ P \ tickFree sa" + by (metis append_Nil2 is_processT1 is_processT5_S3 is_processT5_SEQH3 + no_Trace_implies_no_Failure tickFree_Cons tickFree_append) + let ?f1 = "{(t, X). (t, X \ {tick}) \ \ P \ tickFree t} \ + {(t, X). \t1 t2. t = t1 @ t2 \ t1 @ [tick] \ \ P \ (t2, X) \ \ Q}" + have is_processT5_SEQ2: "\ sa X Y. (sa, X) \ ?f1 + \ (\c. c \ Y \ (sa@[c], {})\?f) + \ (sa, X\Y) \ ?f1" + apply (clarsimp,rule is_processT5_SEQH4[simplified]) + by (auto simp: is_processT5) + show "\s X Y. (s, X) \ ?f \ (\c. c \ Y \ (s @ [c], {}) \ ?f) \ (s, X \ Y) \ ?f" + apply(intro allI impI, elim conjE UnE) + apply(rule rev_subsetD) + apply(rule is_processT5_SEQ2) + apply auto + using is_processT5_S1 apply force + apply (simp add: append_T_imp_tickFree) + using is_processT5[rule_format, OF conjI] by force + next + show "\s X. (s @ [tick], {}) \ ?f \ (s, X - {tick}) \ ?f" + apply auto + apply (metis (no_types, lifting) append_T_imp_tickFree butlast_append + butlast_snoc is_processT2 is_processT6 nonTickFree_n_frontTickFree + non_tickFree_tick tickFree_append) + apply (metis append_T_imp_tickFree front_tickFree_append front_tickFree_mono + is_processT2 non_tickFree_implies_nonMt non_tickFree_tick) + apply (metis process_charn) + apply (metis front_tickFree_append front_tickFree_implies_tickFree) + apply (metis D_T T_nonTickFree_imp_decomp append_T_imp_tickFree append_assoc + append_same_eq non_tickFree_implies_nonMt non_tickFree_tick process_charn + tickFree_append) + by (metis D_imp_front_tickFree append_T_imp_tickFree front_tickFree_append + front_tickFree_mono non_tickFree_implies_nonMt non_tickFree_tick) + next + show "\s t. s \ ?d \ tickFree s \ front_tickFree t \ s @ t \ ?d" + apply auto + using front_tickFree_append apply blast + by (metis process_charn) + next + show "\s X. s \ ?d \ (s, X) \ ?f" + by blast + next + show "\s. s @ [tick] \ ?d \ s \ ?d" + apply auto + apply (metis append_Nil2 front_tickFree_implies_tickFree process_charn) + by (metis append1_eq_conv append_assoc front_tickFree_implies_tickFree is_processT2_TR + nonTickFree_n_frontTickFree non_tickFree_tick process_charn tickFree_append) + qed qed subsection\The Projections\ - -lemmas Rep_Abs_Seq[simp] = Abs_process_inverse[simplified, OF Seq_maintains_is_process] - -lemma - F_Seq : "\ (P \<^bold>; Q) = {(t, X). (t, X \ {tick}) \ \ P \ tickFree t} \ - {(t, X). \t1 t2. t = t1 @ t2 \ t1 @ [tick] \ \ P \ (t2, X) \ \ Q} \ - {(t, X). \t1 t2. t = t1 @ t2 \ t1 \ \ P \ tickFree t1 \ front_tickFree t2} \ - {(t, X). \t1 t2. t = t1 @ t2 \ t1 @ [tick] \ \ P \ t2 \ \ Q}" -unfolding Seq_def by(subst Failures_def, simp only:Rep_Abs_Seq, auto simp: FAILURES_def) +lemma F_Seq: \\ (P \<^bold>; Q) = {(t, X). (t, X \ {tick}) \ \ P \ tickFree t} \ + {(t1 @ t2, X) |t1 t2 X. t1 @ [tick] \ \ P \ (t2, X) \ \ Q} \ + {(t1, X) |t1 X. t1 \ \ P}\ + apply (subst Failures.rep_eq, auto simp add: Seq.rep_eq FAILURES_def) + using is_processT7 is_processT apply (blast+)[5] + by (meson is_processT) + (metis front_tickFree_implies_tickFree front_tickFree_single + is_processT nonTickFree_n_frontTickFree) -lemma - D_Seq : "\ (P \<^bold>; Q) = {t1 @ t2 |t1 t2. t1 \ \ P \ tickFree t1 \ front_tickFree t2} \ - {t1 @ t2 |t1 t2. t1 @ [tick] \ \ P \ t2 \ \ Q}" -unfolding Seq_def by(subst D_def,simp only:Rep_Abs_Seq, simp add:DIVERGENCES_def) + +lemma D_Seq: \\ (P \<^bold>; Q) = \ P \ {t1 @ t2 |t1 t2. t1 @ [tick] \ \ P \ t2 \ \ Q}\ + apply (subst Divergences.rep_eq, auto simp add: Seq.rep_eq DIVERGENCES_def) + by (simp add: is_processT7_S) + (metis elem_min_elems front_tickFree_mono is_processT min_elems_charn + nonTickFree_n_frontTickFree non_tickFree_tick tickFree_Nil tickFree_append) -lemma - T_Seq : "\ (P \<^bold>; Q) = {t. \ X. (t, X \ {tick}) \ \ P \ tickFree t} \ - {t. \ t1 t2. t = t1 @ t2 \ t1 @ [tick] \ \ P \ t2 \ \ Q} \ - {t1 @ t2 |t1 t2. t1 \ \ P \ tickFree t1 \ front_tickFree t2} \ - {t1 @ t2 |t1 t2. t1 @ [tick] \ \ P \ t2 \ \ Q} \ - {t1 @ t2 |t1 t2. t1 \ \ P \ tickFree t1 \ front_tickFree t2} \ - {t1 @ t2 |t1 t2. t1 @ [tick] \ \ P \ t2 \ \ Q}" - unfolding Seq_def - apply(subst Traces_def, simp only: Rep_Abs_Seq, auto simp: TRACES_def FAILURES_def) - using F_T apply fastforce - apply (simp add: append_T_imp_tickFree) - using F_T apply fastforce - using T_F by blast + +lemma T_Seq: \\ (P \<^bold>; Q) = {t. \X. (t, X \ {tick}) \ \ P \ tickFree t} \ + {t1 @ t2 |t1 t2. t1 @ [tick] \ \ P \ t2 \ \ Q} \ + \ P\ + by (auto simp add: Traces.rep_eq TRACES_def Failures.rep_eq[symmetric] F_Seq) + subsection\ Continuity Rule \ lemma mono_Seq_D11: "P \ Q \ \ (Q \<^bold>; S) \ \ (P \<^bold>; S)" apply(auto simp: D_Seq) using le_approx1 apply blast using le_approx_lemma_T by blast lemma mono_Seq_D12: assumes ordered: "P \ Q" shows "(\ s. s \ \ (P \<^bold>; S) \ Ra (P \<^bold>; S) s = Ra (Q \<^bold>; S) s)" proof - have mono_SEQI2a:"P \ Q \ \s. s \ \ (P \<^bold>; S) \ Ra (Q \<^bold>; S) s \ Ra (P \<^bold>; S) s" apply(simp add: Ra_def D_Seq F_Seq) apply(insert le_approx_lemma_F[of P Q] le_approx_lemma_T[of P Q], auto) using le_approx1 by blast+ have mono_SEQI2b:"P \ Q \ \s. s \ \ (P \<^bold>; S) \ Ra (P \<^bold>; S) s \ Ra (Q \<^bold>; S) s" apply(simp add: Ra_def D_Seq F_Seq) apply(insert le_approx_lemma_F[of P Q] le_approx_lemma_T[of P Q] le_approx1[of P Q] le_approx2T[of P Q], auto) using le_approx2 apply fastforce apply (metis front_tickFree_implies_tickFree is_processT2_TR process_charn) apply (simp add: append_T_imp_tickFree) by (metis front_tickFree_implies_tickFree is_processT2_TR process_charn) show ?thesis using ordered mono_SEQI2a mono_SEQI2b by(blast) qed lemma minSeqInclu: "min_elems(\ (P \<^bold>; S)) \ min_elems(\ P) \ {t1@t2|t1 t2. t1@[tick]\\ P\t1\\ P\t2\min_elems(\ S)}" apply(auto simp: D_Seq min_elems_def) - apply (meson process_charn) - apply (metis append_Nil2 front_tickFree_Nil front_tickFree_append front_tickFree_mono - le_list_def less_list_def) - apply (metis (no_types, lifting) D_imp_front_tickFree Nil_is_append_conv append_T_imp_tickFree - append_butlast_last_id butlast_append process_charn less_self lim_proc_is_lub3a - list.distinct(1) nil_less) - by (metis D_imp_front_tickFree append_Nil2 front_tickFree_Nil front_tickFree_mono process_charn - list.distinct(1) nonTickFree_n_frontTickFree) + by (metis append_single_T_imp_tickFree less_append process_charn) + lemma mono_Seq_D13 : assumes ordered: "P \ Q" shows "min_elems (\ (P \<^bold>; S)) \ \ (Q \<^bold>; S)" apply (insert ordered, frule le_approx3, rule subset_trans [OF minSeqInclu]) apply (auto simp: F_Seq T_Seq min_elems_def append_T_imp_tickFree) apply(rule_tac x="{}" in exI, rule is_processT5_S3) apply (metis (no_types, lifting) T_F elem_min_elems le_approx3 less_list_def min_elems5 subset_eq) using Nil_elem_T no_Trace_implies_no_Failure apply fastforce - apply (metis (no_types, lifting) less_self nonTickFree_n_frontTickFree process_charn) - apply(rule_tac x="{}" in exI, metis (no_types, lifting) le_approx2T process_charn) + apply (metis (no_types, lifting) less_self nonTickFree_n_frontTickFree process_charn) + apply (metis (no_types, opaque_lifting) D_T le_approx2T process_charn) by (metis (no_types, lifting) less_self nonTickFree_n_frontTickFree process_charn) lemma mono_Seq[simp] : "P \ Q \ (P \<^bold>; S) \ (Q \<^bold>; S)" by (auto simp: le_approx_def mono_Seq_D11 mono_Seq_D12 mono_Seq_D13) lemma mono_Seq_D21: "P \ Q \ \ (S \<^bold>; Q) \ \ (S \<^bold>; P)" apply(auto simp: D_Seq) using le_approx1 by blast lemma mono_Seq_D22: assumes ordered: "P \ Q" shows "(\ s. s \ \ (S \<^bold>; P) \ Ra (S \<^bold>; P) s = Ra (S \<^bold>; Q) s)" proof - have mono_SEQI2a:"P \ Q \ \s. s \ \ (S \<^bold>; P) \ Ra (S \<^bold>; Q) s \ Ra (S \<^bold>; P) s" apply(simp add: Ra_def D_Seq F_Seq) - apply(insert le_approx_lemma_F[of P Q] le_approx_lemma_T[of P Q], auto) - using le_approx1 by fastforce+ + by(use le_approx_lemma_F[of P Q] le_approx_lemma_T[of P Q] in auto) have mono_SEQI2b:"P \ Q \ \s. s \ \ (S \<^bold>; P) \ Ra (S \<^bold>; P) s \ Ra (S \<^bold>; Q) s" apply(simp add: Ra_def D_Seq F_Seq) apply(insert le_approx_lemma_F[of P Q] le_approx_lemma_T[of P Q] le_approx1[of P Q] le_approx2T[of P Q], auto) using le_approx2 by fastforce+ show ?thesis using ordered mono_SEQI2a mono_SEQI2b by(blast) qed lemma mono_Seq_D23 : assumes ordered: "P \ Q" shows "min_elems (\ (S \<^bold>; P)) \ \ (S \<^bold>; Q)" apply (insert ordered, frule le_approx3, auto simp: D_Seq T_Seq min_elems_def) apply (metis (no_types, lifting) D_imp_front_tickFree Nil_elem_T append.assoc below_refl front_tickFree_charn less_self min_elems2 no_Trace_implies_no_Failure) - apply (simp add: append_T_imp_tickFree) - by (metis (no_types, lifting) D_def D_imp_front_tickFree append_butlast_last_id append_is_Nil_conv - butlast_append butlast_snoc is_process9 is_process_Rep less_self nonTickFree_n_frontTickFree) + apply (simp add: append_T_imp_tickFree) + by (metis (no_types, lifting) append.assoc is_processT less_self nonTickFree_n_frontTickFree) lemma mono_Seq_sym[simp] : "P \ Q \ (S \<^bold>; P) \ (S \<^bold>; Q)" by (auto simp: le_approx_def mono_Seq_D21 mono_Seq_D22 mono_Seq_D23) lemma chain_Seq1: "chain Y \ chain (\i. Y i \<^bold>; S)" by(simp add: chain_def) lemma chain_Seq2: "chain Y \ chain (\i. S \<^bold>; Y i)" by(simp add: chain_def) lemma limproc_Seq_D1: "chain Y \ \ (lim_proc (range Y) \<^bold>; S) = \ (lim_proc (range (\i. Y i \<^bold>; S)))" apply(auto simp:Process_eq_spec D_Seq F_Seq F_LUB D_LUB T_LUB chain_Seq1) - apply(blast) - proof - - fix x - assume A:"\xa. (\t1 t2. x = t1 @ t2 \ t1 \ \ (Y xa) \ tickFree t1 \ front_tickFree t2) \ - (\t1 t2. x = t1 @ t2 \ t1 @ [tick] \ \ (Y xa) \ t2 \ \ S)" - and B: "\t1 t2. x = t1 @ t2 \ (\x. t1 @ [tick] \ \ (Y x)) \ t2 \ \ S" - and C: "chain Y" - thus "\t1 t2. x = t1 @ t2 \ (\x. t1 \ \ (Y x)) \ tickFree t1 \ front_tickFree t2" - proof (cases "\n. \t1 \ x. t1 \ \ (Y n)") - case True - then obtain n where "\t1 \ x. t1 \ \ (Y n)" by blast - with A B C show ?thesis - apply(erule_tac x=n in allE, elim exE disjE, auto simp add:le_list_def) - by (metis D_T chain_lemma is_processT le_approx2T) - next - case False - from False obtain t1 where D:"t1 \ x \ (\n. \t \ x. t \ \ (Y n) \ t \ t1)" by blast - with False have E:"\n. t1 \ \ (Y n)" - by (metis append_Nil2 append_T_imp_tickFree front_tickFree_append front_tickFree_mono - is_processT le_list_def local.A not_Cons_self2) - from A B C D E show ?thesis - by (metis D_imp_front_tickFree Traces_def append_Nil2 front_tickFree_append - front_tickFree_implies_tickFree front_tickFree_mono is_processT - is_process_Rep le_list_def nonTickFree_n_frontTickFree - trace_with_Tick_implies_tickFree_front) - qed - qed + by (metis ND_F_dir2' append_single_T_imp_tickFree is_processT is_ub_thelub) + lemma limproc_Seq_F1: "chain Y \ \ (lim_proc (range Y) \<^bold>; S) = \ (lim_proc (range (\i. Y i \<^bold>; S)))" apply(auto simp add:Process_eq_spec D_Seq F_Seq F_LUB D_LUB T_LUB chain_Seq1) proof (auto, goal_cases) - case (1 a b x) + case (1 a b x y) then show ?case - apply(erule_tac x=x in allE, elim disjE exE, auto simp add: is_processT7 is_processT8_S) - apply(rename_tac t1 t2, erule_tac x=t1 in allE, erule_tac x=t1 in allE, erule_tac x=t1 in allE) - apply (metis D_T append_T_imp_tickFree chain_lemma is_processT le_approx2T not_Cons_self2) - by (metis D_T append_T_imp_tickFree chain_lemma is_processT le_approx2T not_Cons_self2) + apply (erule_tac x=x in allE, elim disjE exE, auto simp add: is_processT7 is_processT8_S) + apply (erule_tac x=t1 in allE, erule_tac x=t2 in allE) + by (metis ND_F_dir2' append_single_T_imp_tickFree is_processT is_ub_thelub) next - case (2 a b) - assume A1:"\t1 t2. a = t1 @ t2 \ (\x. t1 @ [tick] \ \ (Y x)) \ t2 \ \ S" - and A2:"\t1. tickFree t1 \ (\t2. a = t1 @ t2 \ (\x. t1 \ \ (Y x)) \ \ front_tickFree t2)" + case (2 a b x) + assume A1:"a \ \ (Y x)" and A3:"\t1 t2. a = t1 @ t2 \ (\x. t1 @ [tick] \ \ (Y x)) \ (t2, b) \ \ S" - and B: "\x. (a, insert tick b) \ \ (Y x) \ tickFree a \ - (\t1 t2. a = t1 @ t2 \ t1 @ [tick] \ \ (Y x) \ (t2, b) \ \ S) \ - (\t1 t2. a = t1 @ t2 \ t1 \ \ (Y x) \ tickFree t1 \ front_tickFree t2) \ - (\t1 t2. a = t1 @ t2 \ t1 @ [tick] \ \ (Y x) \ t2 \ \ S)" + and B: "\x. (a, insert tick b) \ \ (Y x) \ tickFree a \ + (\t1 t2. a = t1 @ t2 \ t1 @ [tick] \ \ (Y x) \ (t2, b) \ \ S) \ + a \ \ (Y x)" and C:"chain Y" have E:"\ tickFree a \ False" proof - assume F:"\ tickFree a" from A obtain f where D:"f = (\t2. {n. \t1. a = t1 @ t2 \ t1 @ [tick] \ \ (Y n) \ (t2, b) \ \ S} \ {n. \t1. a = t1 @ t2 \ t1 \ \ (Y n) \ tickFree t1 \ front_tickFree t2})" by simp with B F have "\n. n \ (\x\{t. \t1. a = t1 @ t}. f x)" - (is "\n. n \ ?S f") using NF_ND by fastforce + (is "\n. n \ ?S f") using NF_ND + by (smt (verit, best) A1 A3 C ND_F_dir2' append_single_T_imp_tickFree is_processT is_ub_thelub) hence "infinite (?S f)" by (simp add: Sup_set_def) then obtain t2 where E:"t2\{t. \t1. a = t1 @ t} \ infinite (f t2)" using suffixes_fin by blast { assume E1:"infinite{n. \t1. a = t1 @ t2 \ t1 @ [tick] \ \ (Y n) \ (t2, b) \ \ S}" (is "infinite ?E1") with E obtain t1 where F:"a = t1 @ t2 \ (t2, b) \ \ S" using D not_finite_existsD by blast with A3 obtain m where G:"t1@ [tick] \ \ (Y m)" by blast with E1 obtain n where "n \ m \ n \ ?E1" by (meson finite_nat_set_iff_bounded_le nat_le_linear) with D have "n \ m \ t1@ [tick] \ \ (Y n)" by (simp add: F) with G C have False using le_approx_lemma_T chain_mono by blast } note E1 = this { assume E2:"infinite{n. \t1. a = t1 @ t2 \ t1 \ \ (Y n) \ tickFree t1 \ front_tickFree t2}" (is "infinite ?E2") with E obtain t1 where F:"a = t1 @ t2 \ tickFree t1 \ front_tickFree t2" using D not_finite_existsD by blast - with A2 obtain m where G:"t1 \ \ (Y m)" by blast + with A1 obtain m where G:"t1 \ \ (Y m)" using is_processT7_S by blast with E2 obtain n where "n \ m \ n \ ?E2" by (meson finite_nat_set_iff_bounded_le nat_le_linear) with D have "n \ m \ t1 \ \ (Y n)" by (simp add: F) with G C have False using le_approx1 chain_mono by blast } note E2 = this from D E E1 E2 show False by blast qed from E show "tickFree a" by blast qed lemma cont_left_D_Seq : "chain Y \ ((\ i. Y i) \<^bold>; S) = (\ i. (Y i \<^bold>; S))" by (simp add: Process_eq_spec chain_Seq1 limproc_Seq_D1 limproc_Seq_F1 limproc_is_thelub) lemma limproc_Seq_D2: "chain Y \ \ (S \<^bold>; lim_proc (range Y)) = \ (lim_proc (range (\i. S \<^bold>; Y i )))" apply(auto simp add:Process_eq_spec D_Seq F_Seq F_LUB D_LUB T_LUB chain_Seq2) - apply(blast) - proof - - fix x - assume A:"\t1. t1 @ [tick] \ \ S \ (\t2. x = t1 @ t2 \ (\x. t2 \ \ (Y x)))" - and B: "\n. \t1 t2. x = t1 @ t2 \ t1 @ [tick] \ \ S \ t2 \ \ (Y n)" - and C: "chain Y" - from A obtain f where D:"f = (\t2. {n. \t1. x = t1 @ t2 \ t1 @ [tick] \ \ S \ t2 \ \ (Y n)})" - by simp - with B have "\n. n \ (\x\{t. \t1. x = t1 @ t}. f x)" (is "\n. n \ ?S f") by fastforce - hence "infinite (?S f)" by (simp add: Sup_set_def) - then obtain t2 where E:"t2\{t. \t1. x = t1 @ t} \ infinite (f t2)" using suffixes_fin by blast - then obtain t1 where F:"x = t1 @ t2 \ t1 @ [tick] \ \ S" using D not_finite_existsD by blast - from A F obtain m where G:"t2 \ \ (Y m)" by blast +proof - + fix x + assume B: "\n. \t1 t2. x = t1 @ t2 \ t1 @ [tick] \ \ S \ t2 \ \ (Y n)" + and C: "chain Y" + from A obtain f where D:"f = (\t2. {n. \t1. x = t1 @ t2 \ t1 @ [tick] \ \ S \ t2 \ \ (Y n)})" + by simp + with B have "\n. n \ (\x\{t. \t1. x = t1 @ t}. f x)" (is "\n. n \ ?S f") by fastforce + hence "infinite (?S f)" by (simp add: Sup_set_def) + then obtain t2 where E:"t2\{t. \t1. x = t1 @ t} \ infinite (f t2)" using suffixes_fin by blast + then obtain t1 where F:"x = t1 @ t2 \ t1 @ [tick] \ \ S" using D not_finite_existsD by blast + thus "\t1 t2. x = t1 @ t2 \ t1 @ [tick] \ \ S \ (\x. t2 \ \ (Y x))" + apply (rule_tac x = t1 in exI, rule_tac x = t2 in exI) + proof (rule ccontr, simp) + assume \\m. t2 \ \ (Y m)\ + then obtain m where G: \t2 \ \ (Y m)\ .. with E obtain n where "n \ m \ n \ (f t2)" by (meson finite_nat_set_iff_bounded_le nat_le_linear) - with D have "n \ m \ t2 \ \ (Y n)" by blast - with G C have False using le_approx1 po_class.chain_mono by blast - thus "\t1 t2. x = t1 @ t2 \ t1 \ \ S \ tickFree t1 \ front_tickFree t2" .. + with D have "n \ m \ t2 \ \ (Y n)" by blast + with G C show False using le_approx1 po_class.chain_mono by blast qed +qed + + lemma limproc_Seq_F2: "chain Y \ \ (S \<^bold>; lim_proc (range Y)) = \ (lim_proc (range (\i. S \<^bold>; Y i )))" apply(auto simp:Process_eq_spec D_Seq F_Seq T_Seq F_LUB D_LUB D_LUB_2 T_LUB T_LUB_2 chain_Seq2 del:conjI) - apply(auto)[1] - apply(auto)[1] - proof- - fix x X - assume A:"\t1. t1 @ [tick] \ \ S \ (\t2. x = t1 @ t2 \ (\m. (t2, X) \ \ (Y m)))" - and B: "\n. (\t1 t2. x = t1 @ t2 \ t1 @ [tick] \ \ S \ (t2, X) \ \ (Y n)) \ - (\t1 t2. x = t1 @ t2 \ t1 @ [tick] \ \ S \ t2 \ \ (Y n))" - and C: "chain Y" - hence D:"\n. (\t1 t2. x = t1 @ t2 \ t1 @ [tick] \ \ S \ (t2, X) \ \ (Y n))" by (meson NF_ND) - from A obtain f where D:"f = (\t2. {n. \t1. x = t1 @ t2 \ t1 @ [tick] \ \ S \ (t2, X) \ \ (Y n)})" - by simp - with D have "\n. n \ (\x\{t. \t1. x = t1 @ t}. f x)" using B NF_ND by fastforce - hence "infinite (\x\{t. \t1. x = t1 @ t}. f x)" by (simp add: Sup_set_def) - then obtain t2 where E:"t2\{t. \t1. x = t1 @ t} \ infinite (f t2)" using suffixes_fin by blast - then obtain t1 where F:"x = t1 @ t2 \ t1 @ [tick] \ \ S" using D not_finite_existsD by blast - from A F obtain m where G:"(t2, X) \ \ (Y m)" by blast - with E obtain n where "n \ m \ n \ (f t2)" by (meson finite_nat_set_iff_bounded_le nat_le_linear) - with D have "n \ m \ (t2, X) \ \ (Y n)" by blast - with G C have False using is_processT8 po_class.chain_mono proc_ord2a by blast - thus "(x, insert tick X) \ \ S \ tickFree x" .. - qed + apply(auto)[3] +proof (goal_cases) + fix s X + assume A:"\t1. t1 @ [tick] \ \ S \ (\t2. s = t1 @ t2 \ (\m. (t2, X) \ \ (Y m)))" + and B: "\n. \t1 t2. s = t1 @ t2 \ t1 @ [tick] \ \ S \ (t2, X) \ \ (Y n)" + and C: "chain Y" + from B have D:"\n. (\t1 t2. s = t1 @ t2 \ t1 @ [tick] \ \ S \ (t2, X) \ \ (Y n))" by (meson NF_ND) + from A obtain f where D:"f = (\t2. {n. \t1. s = t1 @ t2 \ t1 @ [tick] \ \ S \ (t2, X) \ \ (Y n)})" + by simp + with D have "\n. n \ (\x\{t. \t1. s = t1 @ t}. f x)" using B NF_ND by fastforce + hence "infinite (\x\{t. \t1. s = t1 @ t}. f x)" by (simp add: Sup_set_def) + then obtain t2 where E:"t2\{t. \t1. s = t1 @ t} \ infinite (f t2)" using suffixes_fin by blast + then obtain t1 where F:"s = t1 @ t2 \ t1 @ [tick] \ \ S" using D not_finite_existsD by blast + from A F obtain m where G:"(t2, X) \ \ (Y m)" by blast + with E obtain n where "n \ m \ n \ (f t2)" by (meson finite_nat_set_iff_bounded_le nat_le_linear) + with D have "n \ m \ (t2, X) \ \ (Y n)" by blast + with G C have False using is_processT8 po_class.chain_mono proc_ord2a by blast + thus \(s, insert tick X) \ \ S \ tickFree s\ by simp +qed lemma cont_right_D_Seq : "chain Y \ (S \<^bold>; (\ i. Y i)) = (\ i. (S \<^bold>; Y i))" by (simp add: Process_eq_spec chain_Seq2 limproc_Seq_D2 limproc_Seq_F2 limproc_is_thelub) lemma Seq_cont[simp]: assumes f:"cont f" and g:"cont g" shows "cont (\x. f x \<^bold>; g x)" proof - have A : "\x. cont g \ cont (\y. y \<^bold>; g x)" apply (rule contI2, rule monofunI) apply (auto) by (simp add: cont_left_D_Seq) have B : "\y. cont g \ cont (\x. y \<^bold>; g x)" apply (rule_tac c="(\ x. y \<^bold>; x)" in cont_compose) apply (rule contI2,rule monofunI) by (auto simp add: chain_Seq2 cont_right_D_Seq) show ?thesis using f g apply(rule_tac f="(\x. (\ f. f \<^bold>; g x))" in cont_apply) by(auto intro:contI2 monofunI simp:A B) qed end diff --git a/thys/HOL-CSP/Skip.thy b/thys/HOL-CSP/Skip.thy --- a/thys/HOL-CSP/Skip.thy +++ b/thys/HOL-CSP/Skip.thy @@ -1,80 +1,68 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * - * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * Author : Burkhart Wolff. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) section\The SKIP Process\ -theory Skip +theory Skip imports Process begin -definition SKIP :: "'a process" -where "SKIP \ Abs_process ({(s, X). s = [] \ tick \ X} \ {(s, X). s = [tick]}, {})" - -lemma is_process_REP_SKIP: -" is_process ({(s, X). s = [] \ tick \ X} \ {(s, X). s = [tick]}, {})" -apply(auto simp: FAILURES_def DIVERGENCES_def front_tickFree_def is_process_def) -apply(erule contrapos_np,drule neq_Nil_conv[THEN iffD1], auto) -done - -lemma is_process_REP_SKIP2: -"is_process ({[]} \ {X. tick \ X} \ {(s, X). s = [tick]}, {})" -using is_process_REP_SKIP by auto +lift_definition SKIP :: \'\ process\ + is \({(s, X). s = [] \ tick \ X} \ {(s, X). s = [tick]}, {})\ + unfolding is_process_def FAILURES_def DIVERGENCES_def + by (auto simp add: append_eq_Cons_conv) -lemmas process_prover = Rep_process Abs_process_inverse - FAILURES_def TRACES_def - DIVERGENCES_def is_process_REP_SKIP - lemma F_SKIP: -"\ SKIP = {(s, X). s = [] \ tick \ X} \ {(s, X). s = [tick]}" -by(simp add: process_prover SKIP_def Failures_def is_process_REP_SKIP2) + "\ SKIP = {(s, X). s = [] \ tick \ X} \ {(s, X). s = [tick]}" + by (simp add: FAILURES_def Failures.rep_eq SKIP.rep_eq) lemma D_SKIP: "\ SKIP = {}" -by(simp add: process_prover SKIP_def D_def is_process_REP_SKIP2) + by (simp add: DIVERGENCES_def Divergences.rep_eq SKIP.rep_eq) -lemma T_SKIP: "\ SKIP ={[],[tick]}" -by(auto simp: process_prover SKIP_def Traces_def is_process_REP_SKIP2) +lemma T_SKIP: "\ SKIP ={[], [tick]}" + by (auto simp add: Traces.rep_eq TRACES_def Failures.rep_eq[symmetric] F_SKIP) end diff --git a/thys/HOL-CSP/Stop.thy b/thys/HOL-CSP/Stop.thy --- a/thys/HOL-CSP/Stop.thy +++ b/thys/HOL-CSP/Stop.thy @@ -1,71 +1,78 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * - * Author : Burkhart Wolff, Safouan Taha, Lina Ye. + * Author : Burkhart Wolff. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) section\ The STOP Process \ theory Stop imports Process begin -definition STOP :: "'\ process" -where "STOP \ Abs_process ({(s, X). s = []}, {})" +lift_definition STOP :: \'\ process\ + is \({(s, X). s = []}, {})\ + unfolding is_process_def FAILURES_def DIVERGENCES_def by simp -lemma is_process_REP_STOP: "is_process ({(s, X). s = []},{})" -by(simp add: is_process_def FAILURES_def DIVERGENCES_def) - -lemma Rep_Abs_STOP : "Rep_process (Abs_process ({(s, X). s = []},{})) = ({(s, X). s = []},{})" -by(subst Abs_process_inverse, simp add: Rep_process is_process_REP_STOP, auto) lemma F_STOP : "\ STOP = {(s,X). s = []}" -by(simp add: STOP_def FAILURES_def Failures_def Rep_Abs_STOP) + by (simp add: FAILURES_def Failures.rep_eq STOP.rep_eq) lemma D_STOP: "\ STOP = {}" -by(simp add: STOP_def DIVERGENCES_def D_def Rep_Abs_STOP) + by (simp add: DIVERGENCES_def Divergences.rep_eq STOP.rep_eq) lemma T_STOP: "\ STOP = {[]}" -by(simp add: STOP_def TRACES_def FAILURES_def Traces_def Rep_Abs_STOP) + by (simp add: Traces.rep_eq TRACES_def Failures.rep_eq[symmetric] F_STOP) + + +lemma STOP_iff_T: \P = STOP \ \ P = {[]}\ + apply (intro iffI, simp add: T_STOP) + apply (subst Process_eq_spec, safe, simp_all add: F_STOP D_STOP) + by (use F_T in force, use is_processT5_S7 in fastforce) + (metis D_T append_Nil front_tickFree_single is_processT7_S + list.distinct(1) singletonD tickFree_Nil) + + + end diff --git a/thys/HOL-CSP/Sync.thy b/thys/HOL-CSP/Sync.thy --- a/thys/HOL-CSP/Sync.thy +++ b/thys/HOL-CSP/Sync.thy @@ -1,1893 +1,1868 @@ (*<*) \\ ******************************************************************** * Project : HOL-CSP - A Shallow Embedding of CSP in Isabelle/HOL * Version : 2.0 * * Author : Burkhart Wolff, Safouan Taha, Lina Ye. * (Based on HOL-CSP 1.0 by Haykal Tej and Burkhart Wolff) * * This file : A Combined CSP Theory * * Copyright (c) 2009 Université Paris-Sud, France * * All rights reserved. * * Redistribution and use in source and binary forms, with or without * modification, are permitted provided that the following conditions are * met: * * * Redistributions of source code must retain the above copyright * notice, this list of conditions and the following disclaimer. * * * Redistributions in binary form must reproduce the above * copyright notice, this list of conditions and the following * disclaimer in the documentation and/or other materials provided * with the distribution. * * * Neither the name of the copyright holders nor the names of its * contributors may be used to endorse or promote products derived * from this software without specific prior written permission. * * THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS * "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT * LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR * A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL THE COPYRIGHT * OWNER OR CONTRIBUTORS BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, * SPECIAL, EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT * LIMITED TO, PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, * DATA, OR PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY * THEORY OF LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT * (INCLUDING NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE * OF THIS SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE. ******************************************************************************\ (*>*) section\ The Synchronizing Operator \ theory Sync imports Process "HOL-Library.Infinite_Set" begin subsection\Basic Concepts\ fun setinterleaving:: "'a trace \ ('a event) set \ 'a trace \ ('a trace)set" where si_empty1: "setinterleaving([], X, []) = {[]}" | si_empty2: "setinterleaving([], X, (y # t)) = (if (y \ X) then {} else {z.\ u. z = (y # u) \ u \ setinterleaving ([], X, t)})" | si_empty3: "setinterleaving((x # s), X, []) = (if (x \ X) then {} else {z.\ u. z = (x # u) \ u \ setinterleaving (s, X, [])})" | si_neq : "setinterleaving((x # s), X, (y # t)) = (if (x \ X) then if (y \ X) then if (x = y) then {z.\ u. z = (x#u) \ u \ setinterleaving(s, X, t)} else {} else {z.\ u. z = (y#u) \ u \ setinterleaving ((x#s), X, t)} else if (y \ X) then {z.\ u. z = (x # u) \ u \ setinterleaving (s, X, (y # t))} \ {z.\ u. z = (y # u) \ u \ setinterleaving((x # s), X, t)} else {z.\ u. z = (x # u) \ u \ setinterleaving (s, X, (y # t))})" fun setinterleavingList::"'a trace \ ('a event) set \ 'a trace \ ('a trace)list" where si_empty1l: "setinterleavingList([], X, []) = [[]]" | si_empty2l: "setinterleavingList([], X, (y # t)) = (if (y \ X) then [] else map (\z. y#z) (setinterleavingList ([], X, t)))" | si_empty3l: "setinterleavingList((x # s), X, []) = (if (x \ X) then [] else map (\z. x#z) (setinterleavingList (s, X, [])))" | si_neql : "setinterleavingList((x # s), X, (y # t)) = (if (x \ X) then if (y \ X) then if (x = y) then map (\z. x#z) (setinterleavingList(s, X, t)) else [] else map (\z. y#z) (setinterleavingList ((x#s), X, t)) else if (y \ X) then map (\z. x#z) (setinterleavingList (s, X, (y # t))) @ map (\z. y#z) (setinterleavingList ((x#s), X, t)) else map (\z. x#z) (setinterleavingList (s, X, (y # t))))" lemma finiteSetinterleavingList: "finite (set (setinterleavingList(s, X, t)))" by auto lemma sym : "setinterleaving(s, X, t)= setinterleaving(t, X, s)" by (rule setinterleaving.induct[of "\(s,X,t). setinterleaving (s, X, t) = setinterleaving (t, X, s)" "(s, X, t)", simplified], auto) abbreviation "setinterleaves_syntax" ("_ setinterleaves '(()'(_, _')(), _')" [60,0,0,0]70) where "u setinterleaves ((s, t), X) == (u \ setinterleaving(s, X, t))" -subsection\Definition\ - -definition Sync :: "['a process,'a set,'a process] => 'a process" ("(3(_ \_\/ _))" [64,0,65] 64) - where "P \ A \ Q \ Abs_process({(s,R).\ t u X Y. (t,X) \ \ P \ (u,Y) \ \ Q \ - (s setinterleaves ((t,u),(ev`A) \ {tick})) \ - R = (X \ Y) \ ((ev`A) \ {tick}) \ X \ Y} \ - {(s,R).\ t u r v. front_tickFree v \ (tickFree r \ v=[]) \ - s = r@v \ - (r setinterleaves ((t,u),(ev`A) \ {tick})) \ - (t \ \ P \ u \ \ Q \ t \ \ Q \ u \ \ P)}, - - {s. \ t u r v. front_tickFree v \ (tickFree r \ v=[]) \ - s = r@v \ - (r setinterleaves ((t,u),(ev`A) \ {tick})) \ - (t \ \ P \ u \ \ Q \ t \ \ Q \ u \ \ P)})" - subsection\Consequences\ lemma emptyLeftProperty:"\s. s setinterleaves (([], t), A)\s=t" - apply(induct_tac t) - apply simp - by auto + by (induct t) auto + lemma emptyLeftSelf: " (\t1. t1\ set t\t1\A)\t setinterleaves (([], t), A) " - apply(induct_tac t) - apply simp - by auto + by (induct t) auto -lemma emptyLeftNonSync: "\s. s setinterleaves (([], t), A)\(\t1. t1\ set t\t1\A)" - apply(induct_tac t) - apply simp -proof- - fix a list - assume a: "\s. s setinterleaves (([], list), A) \ (\t1. t1 \ set list \ t1 \ A)" - thus "\s. s setinterleaves (([], a # list), A) \ (\t1. t1 \ set (a # list) \ t1 \ A)" - proof- - have th: "s setinterleaves (([], a # list), A)\a\A" by auto - obtain s1 where th1: "s setinterleaves (([], a # list), A)\s1 setinterleaves (([], list), A)" - using mem_Collect_eq th by fastforce - from a th1 have th2: "\s. s setinterleaves (([], a # list), A)\(\t. t\ set list\t \ A)" - by auto - with a show ?thesis - by (metis (no_types, lifting) empty_iff set_ConsD setinterleaving.simps(2)) - qed + +lemma emptyLeftNonSync: "s setinterleaves (([], t), A) \ \a. a \ set t \ a \ A" +proof (induct t arbitrary: s) + case Nil + show ?case by simp +next + case (Cons a t) + thus ?case by (cases s; simp split: if_split_asm) qed lemma ftf_Sync1: "a \ set(u) \ a \ set(t) \ s setinterleaves ((t, u), A) \ a \ set(s)" proof (induction "length t + length u" arbitrary:s t u rule:nat_less_induct) case 1 show ?case apply(cases t) using sym emptyLeftProperty apply blast apply(cases u) using sym emptyLeftProperty apply blast apply(simp split:if_splits, intro conjI impI, elim conjE disjE exE) apply (metis "1.hyps" add_less_le_mono length_Cons lessI order_refl set_ConsD) apply (metis "1.hyps" add_Suc_right length_Cons lessI set_ConsD) apply (metis "1.hyps" add_less_mono length_Cons lessI set_ConsD) apply (metis "1.hyps" add_Suc_right length_Cons lessI set_ConsD) apply (metis (no_types, opaque_lifting) "1.hyps" add.commute add_Suc_right insert_iff length_Cons lessI list.simps(15)) by (metis "1.hyps" add_less_le_mono length_Cons lessI order_refl set_ConsD) qed lemma addNonSyncEmpty: "sa setinterleaves (([], u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves (([y1], u), A) \ (sa @ [y1]) setinterleaves (([], u @ [y1]), A)" proof (induction "length u" arbitrary:sa u rule:nat_less_induct) case 1 then show ?case apply(cases u) apply simp proof- fix a list assume a: "\mx. m = length x \(\xa. xa setinterleaves (([], x), A) \ y1 \ A \ (xa @ [y1]) setinterleaves (([y1], x), A) \ (xa @ [y1]) setinterleaves (([], x @ [y1]), A))" and b: "u = a # list" from b have th: "sa setinterleaves (([], u), A)\(tl sa) setinterleaves (([], list), A)" by (metis emptyLeftNonSync emptyLeftProperty emptyLeftSelf list.distinct(1) list.sel(3) list.set_sel(2)) from a b th have th1: "sa setinterleaves (([], u), A) \ y1 \ A\ ((tl sa) @ [y1]) setinterleaves (([y1], list), A)\((tl sa)@[y1]) setinterleaves (([], list@[y1]), A)" by auto thus "sa setinterleaves (([], u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves (([y1], u), A) \ (sa @ [y1]) setinterleaves (([], u @ [y1]), A)" using b th by auto qed qed lemma addNonSync: "sa setinterleaves ((t,u),A)\ y1 \ A\ (sa@[y1]) setinterleaves ((t@[y1],u),A)\(sa@[y1]) setinterleaves ((t,u@[y1]),A)" proof (induction "length t + length u" arbitrary:sa t u rule:nat_less_induct) case 1 then show ?case apply(cases t) apply (simp add: addNonSyncEmpty) apply (cases u) apply (metis Sync.sym addNonSyncEmpty append_self_conv2) proof- fix a list aa lista assume a: "\mx xa. m = length x + length xa \ (\xb. xb setinterleaves ((x, xa), A) \ y1 \ A \(xb @ [y1]) setinterleaves ((x @ [y1], xa), A) \ (xb @ [y1]) setinterleaves ((x, xa @ [y1]), A))" and b: " t = a # list" and c: "u = aa # lista" thus " sa setinterleaves ((t, u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves ((t @ [y1], u), A) \ (sa @ [y1]) setinterleaves ((t, u @ [y1]), A)" proof- from b c have th0pre: "a=aa\aa\ A\sa setinterleaves ((t, u), A) \ (tl sa) setinterleaves ((list,lista), A)" by auto from th0pre a b c have th0pre1: "a=aa\aa\ A\sa setinterleaves ((t, u), A)\ y1 \ A\ ((tl sa) @ [y1]) setinterleaves ((list@ [y1], lista), A)\ ((tl sa) @ [y1]) setinterleaves ((list, lista@ [y1]), A)" by (metis add_less_mono length_Cons lessI) from th0pre th0pre1 b c have th0: "a=aa\aa\ A\sa setinterleaves ((t, u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves ((t @ [y1], u), A)\(sa @ [y1]) setinterleaves ((t, u @ [y1]), A)" by auto from b c have th1pre: "a\A\aa\ A\sa setinterleaves ((t, u), A) \ (tl sa) setinterleaves ((list,aa#lista), A)" by auto from th1pre a b c have th1pre1: "a\A\aa\ A\sa setinterleaves ((t, u), A)\ y1 \ A\ ((tl sa) @ [y1]) setinterleaves ((list@ [y1], aa#lista), A)\ ((tl sa) @ [y1]) setinterleaves ((list, aa#lista@ [y1]), A)" by (metis add_Suc append_Cons length_Cons lessI) from th1pre th1pre1 b c have th1: "a\A\aa\ A\sa setinterleaves ((t, u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves ((t @ [y1], u), A)\(sa @ [y1]) setinterleaves ((t, u @ [y1]), A)" by auto from b c have th2pre: "aa\A\a\ A\sa setinterleaves ((t, u), A) \ (tl sa) setinterleaves ((a#list,lista), A)" by auto from th2pre a b c have th2pre1: "aa\A\a\ A\sa setinterleaves ((t, u), A)\ y1 \ A\ ((tl sa) @ [y1]) setinterleaves ((a#list@ [y1], lista), A)\ ((tl sa) @ [y1]) setinterleaves ((a#list, lista@ [y1]), A)" by (metis add_Suc_right append_Cons length_Cons lessI) from th2pre th2pre1 b c have th2: "aa\A\a\ A\sa setinterleaves ((t, u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves ((t @ [y1], u), A)\(sa @ [y1]) setinterleaves ((t, u @ [y1]), A)" by auto from b c have th3pre: "aa\A\a\ A\sa setinterleaves ((t, u), A)\ (tl sa) setinterleaves ((a#list,lista), A)\(tl sa)setinterleaves ((list,aa#lista), A)" by auto from th3pre a b c have th3pre1: "aa\A\a\ A\sa setinterleaves ((t, u), A)\ (tl sa) setinterleaves ((a#list,lista), A)\ y1 \ A \ ((tl sa) @ [y1]) setinterleaves ((a#list@ [y1], lista ), A)\((tl sa) @ [y1]) setinterleaves ((a#list, lista@ [y1] ), A)" by (metis add_Suc_right append_Cons length_Cons lessI) from th3pre a b c have th3pre2: "aa\A\a\ A\sa setinterleaves ((t, u), A)\ (tl sa) setinterleaves ((list,aa#lista), A)\ y1 \ A \((tl sa) @ [y1]) setinterleaves ((list@ [y1], aa#lista ), A)\((tl sa) @ [y1]) setinterleaves ((list, aa#lista @ [y1]), A)" by (metis add_Suc append_Cons length_Cons lessI) from th3pre th3pre1 th3pre2 a b c have th3: "aa\A\a\ A\sa setinterleaves ((t, u), A) \ y1 \ A \(sa @ [y1]) setinterleaves ((t @ [y1], u), A) \ (sa @ [y1]) setinterleaves ((t, u @ [y1]), A)" by auto show ?thesis using b c th0 th1 th2 th3 by auto qed qed qed lemma addSyncEmpty: "sa setinterleaves (([], u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves (([y1], u @ [y1]), A)" proof (induction "length u" arbitrary:sa u rule:nat_less_induct) case 1 then show ?case apply(cases u) apply simp proof- fix a list assume a: "\mx. m = length x \(\xa. xa setinterleaves (([], x), A) \ y1 \ A \ (xa @ [y1]) setinterleaves (([y1], x @ [y1]), A))" and b: "u = a # list" from b have th: "sa setinterleaves (([], u), A)\(tl sa) setinterleaves (([], list), A)" by (metis emptyLeftNonSync emptyLeftProperty emptyLeftSelf list.distinct(1) list.sel(3) list.set_sel(2)) from a th have th1: "sa setinterleaves (([], u), A) \ y1 \ A\ ((tl sa) @ [y1]) setinterleaves (([y1], list @ [y1]), A)" using b by auto thus "sa setinterleaves (([], u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves (([y1], u @ [y1]), A)" using b th by auto qed qed lemma addSync: "sa setinterleaves ((t,u),A)\ y1 \ A\(sa@[y1]) setinterleaves ((t@[y1],u@[y1]),A)" + find_theorems \(@)\ name: cases proof (induction "length t + length u" arbitrary:sa t u rule:nat_less_induct) case 1 then show ?case apply(cases t) apply (simp add: addSyncEmpty) apply(cases u) apply (metis Sync.sym addSyncEmpty append.left_neutral) proof- fix a list aa lista assume a: "\mx xa. m = length x + length xa \ (\xb. xb setinterleaves ((x, xa), A) \ y1 \ A \ (xb @ [y1]) setinterleaves ((x @ [y1], xa @ [y1]), A))" and b: "t = a # list " and c: " u = aa # lista" thus "sa setinterleaves((t, u),A)\y1 \A\(sa @ [y1]) setinterleaves ((t @ [y1], u @ [y1]), A)" proof- from b c have th0pre: "a=aa\aa\ A\sa setinterleaves ((t, u), A) \ (tl sa) setinterleaves ((list,lista), A)" by auto from th0pre a b c have th0pre1: "a=aa\aa\ A\sa setinterleaves ((t, u), A)\ y1 \ A\ ((tl sa) @ [y1]) setinterleaves ((list@ [y1], lista @ [y1]), A)" by (metis add_less_mono length_Cons lessI) from th0pre th0pre1 b c have th0: "a=aa\a\A\sa setinterleaves ((t, u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves ((t @ [y1], u @ [y1]), A)" by auto from b c have th1pre: "a\A\aa\ A\sa setinterleaves ((t, u), A) \ (tl sa) setinterleaves ((list,aa#lista), A)" by auto from th1pre a b c have th1pre1: "a\A\aa\ A\sa setinterleaves ((t, u), A)\ y1 \ A\ ((tl sa) @ [y1]) setinterleaves ((list@ [y1], aa#lista @ [y1]), A)" by (metis add_Suc append_Cons length_Cons lessI) from th1pre th1pre1 b c have th1: "a\A\aa\ A\sa setinterleaves ((t, u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves ((t @ [y1], u @ [y1]), A)" by auto from b c have th2pre: "aa\A\a\ A\sa setinterleaves ((t, u), A) \ (tl sa) setinterleaves ((a#list,lista), A)" by auto from th2pre a b c have th2pre1: "aa\A\a\ A\sa setinterleaves ((t, u), A)\ y1 \ A\ ((tl sa) @ [y1]) setinterleaves ((a#list@ [y1], lista @ [y1]), A)" by (metis add_Suc_right append_Cons length_Cons lessI) from th2pre th2pre1 b c have th2: "aa\A\a\ A\sa setinterleaves ((t, u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves ((t @ [y1], u @ [y1]), A)" by auto from b c have th3pre: "aa\A\a\ A\sa setinterleaves ((t, u), A)\ (tl sa) setinterleaves((a#list,lista), A)\(tl sa)setinterleaves ((list,aa#lista), A)" by auto from th3pre a b c have th3pre1: "aa\A\a\ A\sa setinterleaves ((t, u), A)\ (tl sa) setinterleaves ((a#list,lista), A)\ y1 \ A \ ((tl sa) @ [y1]) setinterleaves ((a#list@ [y1], lista @ [y1]), A)" by (metis add_Suc_right append_Cons length_Cons lessI) from th3pre a b c have th3pre2: "aa\A\a\ A\sa setinterleaves ((t, u), A)\ (tl sa) setinterleaves ((list,aa#lista), A)\ y1 \ A \ ((tl sa) @ [y1]) setinterleaves ((list@ [y1], aa#lista @ [y1]), A)" by (metis add_Suc append_Cons length_Cons lessI) from th3pre th3pre1 th3pre2 a b c have th3: "aa\A\a\ A\sa setinterleaves ((t, u), A) \ y1 \ A \ (sa @ [y1]) setinterleaves ((t @ [y1], u @ [y1]), A)" by auto show ?thesis using b c th0 th1 th2 th3 by auto qed qed qed lemma doubleReverse: "s1 setinterleaves ((t, u), A)\(rev s1) setinterleaves ((rev t, rev u), A)" proof (induction "length t + length u" arbitrary:s1 t u rule:nat_less_induct) case 1 then show ?case apply(cases t) using emptyLeftNonSync apply (metis emptyLeftProperty emptyLeftSelf rev.simps(1) set_rev) apply(cases u) using sym apply (metis (no_types, lifting) emptyLeftNonSync emptyLeftProperty emptyLeftSelf rev.simps(1) set_rev) proof- fix a list aa lista assume a: "\mx xa. m = length x + length xa \ (\xb. xb setinterleaves ((x, xa), A) \rev xb setinterleaves ((rev x, rev xa), A))" and b: "t = a # list" and c: "u = aa # lista" thus "s1 setinterleaves ((t, u), A) \rev s1 setinterleaves ((rev t, rev u), A)" proof- from b c have th0pre: "a=aa\aa\ A\s1 setinterleaves ((t, u), A)\ (tl s1) setinterleaves ((list,lista), A)" by auto from th0pre a b c have th0pre1: "a=aa\aa\ A\s1 setinterleaves ((t, u), A)\ ((rev (tl s1)) setinterleaves ((rev list, rev lista), A))" by (metis add_less_mono length_Cons lessI) from th0pre th0pre1 b c have th0: "a=aa\aa\ A\s1 setinterleaves ((t, u), A) \ rev s1 setinterleaves ((rev t, rev u), A)" using addSync by fastforce from b c have th1pre: "a\A\aa\ A\s1 setinterleaves ((t, u), A)\ (tl s1) setinterleaves ((list,aa#lista), A)" by auto from th1pre a b c have th1pre1: "a\A\aa\ A\s1 setinterleaves ((t, u), A)\ ((rev (tl s1)) setinterleaves ((rev list, rev (aa#lista)), A))" by (metis add_less_mono1 length_Cons lessI) from th1pre th1pre1 b c have th1: "a\A\aa\ A\s1 setinterleaves ((t, u), A) \ rev s1 setinterleaves ((rev t, rev u), A)" using addNonSync by fastforce from b c have th2pre: "aa\A\a\ A\s1 setinterleaves ((t, u), A)\ (tl s1) setinterleaves ((a#list,lista), A)" by auto from th2pre a b c have th2pre1: "aa\A\a\ A\s1 setinterleaves ((t, u), A)\ ((rev (tl s1)) setinterleaves ((rev (a#list), rev lista), A))" by (metis add_Suc_right length_Cons lessI) from th2pre th2pre1 b c have th2: "aa\A\a\ A\s1 setinterleaves ((t, u), A) \ rev s1 setinterleaves ((rev t, rev u), A)" using addNonSync by fastforce from b c have th3pre: "aa\A\a\ A\s1 setinterleaves ((t, u), A)\ (tl s1) setinterleaves ((a#list,lista),A)\(tl s1) setinterleaves((list,aa#lista), A)" by auto from th3pre a b c have th3pre1: "aa\A\a\ A\s1 setinterleaves ((t, u), A)\ (tl s1) setinterleaves ((a#list,lista), A) \ ((rev (tl s1)) setinterleaves ((rev (a#list), rev (lista)), A))" by (metis add_Suc_right length_Cons lessI) from th3pre a b c have th3pre2: "aa\A\a\ A\s1 setinterleaves ((t, u), A)\ (tl s1) setinterleaves ((list,aa#lista), A) \ ((rev (tl s1)) setinterleaves ((rev (list), rev (aa#lista)), A))" by (metis add_less_mono1 length_Cons lessI) from th3pre1 have th3pre31: "aa\A\a\ A\s1 setinterleaves ((t, u), A)\ ((rev (tl s1)) setinterleaves ((rev (a#list), rev (lista)), A))\ ((rev (tl s1))@[aa]) setinterleaves ((rev (a#list), (rev (lista))@[aa]), A)" by (simp add: addNonSync) from th3pre2 have th3pre32: "aa\A\a\ A\s1 setinterleaves ((t, u), A)\ ((rev (tl s1)) setinterleaves ((rev (list), rev (aa#lista)), A))\ ((rev (tl s1))@[a]) setinterleaves (((rev (list)@[a]), (rev (aa#lista))), A)" by (simp add: addNonSync) from th3pre th3pre1 th3pre2 th3pre31 th3pre32 a b c have th3: "aa\A\a\ A\ s1 setinterleaves ((t, u), A) \rev s1 setinterleaves ((rev t, rev u), A)" by force show?thesis using b c th0 th1 th2 th3 by auto qed qed qed lemma ftf_Sync21: "(a\set(u)\a\set(t)\a\set(t)\a\set(u))\a\A\setinterleaving(u, A ,t)= {}" proof (induction "length t + length u" arbitrary: t u rule:nat_less_induct) case 1 then show ?case apply(cases t) using Sync.sym emptyLeftNonSync apply fastforce apply(cases u) apply auto[1] apply(simp split:if_splits, intro conjI impI, elim conjE disjE exE) apply blast apply auto[1] apply blast apply auto[1] apply auto[1] using less_SucI apply blast apply auto[1] apply auto[1] using list.simps(15) apply auto[1] by auto qed lemma ftf_Sync32: "u=u1@[tick]\t=t1@[tick]\s setinterleaves ((t, u), insert tick (ev ` A))\ \s1. s1 setinterleaves ((t1, u1), insert tick (ev ` A))\ (s=s1@[tick])" proof- assume h: "u=u1@[tick]\t=t1@[tick]\s setinterleaves ((t, u), insert tick (ev ` A))" thus "\s1. s1 setinterleaves ((t1, u1), insert tick (ev ` A))\ (s=s1@[tick])" proof- from h have a:"rev u=tick#rev u1" by auto from h have b:"rev t=tick#rev t1" by auto from h have ab: "(rev s) setinterleaves ((tick#rev t1, tick#rev u1), insert tick (ev ` A))" using doubleReverse by fastforce from h obtain ss where c: "ss setinterleaves ((rev t1, rev u1), insert tick (ev ` A))\ ss=tl(rev s)" using ab by auto from c have d: "(rev ss) setinterleaves (( t1, u1), insert tick (ev ` A))" using doubleReverse by fastforce from d have e: "rev s=tick#ss" using ab append_Cons_eq_iff c by auto show ?thesis using d e by blast qed qed lemma SyncWithTick_imp_NTF: "(s @ [tick]) setinterleaves ((t, u),insert tick (ev ` A)) \ t\ \ P \ u \ \ Q \ \t1 u1. t=t1@[tick] \ u=u1@[tick]\ s setinterleaves ((t1, u1), insert tick (ev ` A))" proof- assume h: "(s @ [tick]) setinterleaves ((t, u), insert tick (ev ` A))" and h1: "t\ \ P" and h2: "u\ \ Q" thus "\t1 u1. t=t1@[tick] \ u=u1@[tick]\ s setinterleaves ((t1, u1), insert tick (ev ` A))" proof- from h have a: "(tick#rev s) setinterleaves ((rev t, rev u), insert tick (ev ` A))" using doubleReverse by fastforce from h obtain tt uu where b: "t=tt@[tick]\u=uu@[tick]" by (metis T_nonTickFree_imp_decomp empty_iff ftf_Sync1 ftf_Sync21 h1 h2 insertI1 non_tickFree_tick tickFree_append tickFree_def) from h b have d: "s setinterleaves ((tt, uu), insert tick (ev ` A))" using ftf_Sync32 by blast show ?thesis using b d by blast qed qed lemma synPrefix1: " ta = [] \ \t1 u1. (s @ t) setinterleaves ((ta, u), A) \ t1 \ ta \ u1 \ u \ s setinterleaves ((t1, u1), A)" proof- assume a: "ta = []" obtain u1 where th: "u1=s" by blast from a have th2: "(s @ t) setinterleaves ((ta, u), A) \ (s @ t)= u" by (simp add: a emptyLeftProperty) from a have th3: "(s @ t) setinterleaves ((ta, u), A) \ (\t1. t1\set u\t1\A)" by (simp add: emptyLeftNonSync) from a th have th1: "(s @ t) setinterleaves ((ta, u), A) \ [] \ ta \ u1 \ u " using le_list_def th2 by blast from a th have thh1: "(s @ t) setinterleaves ((ta, u), A)\ (\t1. t1\set u1\t1\A) \ s setinterleaves (([], u1), A)" by (simp add: emptyLeftSelf) thus "\t1 u1. (s @ t) setinterleaves ((ta, u), A) \ t1 \ ta \ u1 \ u \ s setinterleaves ((t1, u1), A)" using th th1 th2 th3 thh1 by fastforce qed lemma synPrefix: "\t1 u1. (s @ t) setinterleaves ((ta, u), A)\ t1\ta\u1\u\s setinterleaves ((t1, u1), A)" proof (induction "length ta + length u" arbitrary: s t ta u rule:nat_less_induct) case 1 then show ?case apply(cases ta) using synPrefix1 apply fastforce apply(cases u) using sym synPrefix1 apply metis proof- fix a list aa lista s t assume a: " \mx xa. m = length x + length xa \ (\xb xc. \t1 u1. (xb @ xc) setinterleaves ((x, xa), A) \ t1 \ x \ u1 \ xa \ xb setinterleaves ((t1, u1), A))" and b: "ta = a # list" and c: "u = aa # lista " thus " \t1 u1. (s @ t) setinterleaves ((ta, u), A) \ t1 \ ta \ u1 \ u \ s setinterleaves ((t1, u1), A)" proof- from a have th0: "\xb xc. \t1 u1. (xb @ xc) setinterleaves ((list, lista), A) \ t1 \ list \ u1 \ lista \ xb setinterleaves ((t1, u1), A)" by (metis add_less_le_mono b c impossible_Cons le_cases not_le_imp_less) from b c obtain yb where thp: "a=aa\a\ A\(s@t) setinterleaves ((ta, u), A)\length s >1 \ yb=tl s" by blast with b c have thp4: "a=aa\a\ A\(s @ t) setinterleaves ((ta, u), A)\length s >1 \s=a#yb" by (auto, metis Cons_eq_append_conv list.sel(3) list.size(3) not_less0) have thp5: "a=aa\a\ A\(s @ t) setinterleaves ((ta, u), A)\length s >1 \ (yb @ t) setinterleaves ((list, lista), A)" using b c thp4 by auto from th0 obtain yt yu where thp1: "a = aa \ a \ A \ (s @ t) setinterleaves ((ta, u), A) \ 1 < length s\yb setinterleaves ((yt, yu), A)\yt\list\yu\lista" using thp5 by blast from thp thp1 have thp2: "a=aa\a\ A\ (s @ t) setinterleaves ((ta, u), A)\ length s >1\ s setinterleaves ((a#yt, aa#yu), A)" using thp4 by auto from b c have thp3: "a=aa\a\ A\ (s @ t) setinterleaves ((ta, u), A) \ length s=1\ s setinterleaves (([a], [aa]), A)" using append_eq_append_conv2[of s t "[aa]"] by (auto, metis append_Nil2 append_eq_append_conv length_Cons list.size(3)) have thp6: "a=aa\a\ A\ (s @ t) setinterleaves ((ta, u), A) \ length s=0\ s setinterleaves (([], []), A)" by auto from thp1 have thp7: "a=aa\a\ A\ (s @ t) setinterleaves ((ta, u), A)\ length s >1\ (a # yt)\ta \(aa # yu)\u" by (metis b c le_less less_cons) have th: "a=aa\a\ A\\t1 u1. (s @ t) setinterleaves ((ta, u), A) \ t1 \ ta \ u1 \ u \ s setinterleaves ((t1, u1), A)" proof - assume dd:"a=aa \ a\ A" consider (aa) "length s = 0" | (bb) "length s = 1" | (cc) "length s > 1" by linarith then show ?thesis proof cases case aa with thp6 show ?thesis by (rule_tac x="[]" in exI, rule_tac x="[]" in exI, simp) next case bb with dd thp3 b c show ?thesis by (rule_tac x="[a]" in exI, rule_tac x="[a]" in exI, auto simp add: le_list_def) next case cc with dd thp2 thp7 b c show ?thesis by (rule_tac x="a#yt" in exI, rule_tac x="a#yu" in exI, auto simp add: le_list_def) qed qed from b c have th1pre: "a\A\aa\ A\(s @ t) setinterleaves ((ta, u), A) \ tl (s @ t) setinterleaves ((list, u), A)\hd (s@t)=a" by auto from a b c obtain yt1 yu1 where th1pre1: "a\A\aa\ A\(s @ t) setinterleaves ((ta, u), A)\ length s>0\yt1\list\ yu1 \ u\ tl s setinterleaves ((yt1, yu1), A)" by (metis (no_types, lifting) length_Cons length_greater_0_conv lessI plus_nat.simps(2) th1pre tl_append2) from b have th1pre2: "yt1\list\a#yt1\ta" by (simp add: le_less less_cons) from b c have th1pre3: "a\A\aa\ A\tl s setinterleaves ((yt1, yu1), A)\ (a#(tl s)) setinterleaves ((a#yt1, yu1), A)" by (metis (mono_tags, lifting) addNonSync doubleReverse rev.simps(2) rev_rev_ident) from b c th1pre1 th1pre2 have th1pre4: "a\A\aa\ A\(s @ t) setinterleaves ((ta, u), A)\ length s>0 \a#yt1\ta\ yu1 \u\ s setinterleaves ((a#yt1, yu1), A)" using th1pre th1pre3 by fastforce have th1pre5:"a\A\aa\ A\(s @ t) setinterleaves ((ta, u), A)\length s=0 \[]\ta\ [] \u\ s setinterleaves (([], []), A)" by auto have th1: "a\A\aa\ A\\t1 u1. (s @ t) setinterleaves ((ta, u), A) \ t1 \ ta \ u1 \ u \ s setinterleaves ((t1, u1), A)" using th1pre4 th1pre5 by blast from b c have th2pre: "aa\A\a\ A\(s @ t) setinterleaves ((ta, u), A) \ tl (s @ t) setinterleaves ((ta, lista), A)\hd (s@t)=aa" by auto from a b c obtain yt2 yu2 where th2pre1: "aa\A\a\ A\(s @ t) setinterleaves ((ta, u), A)\ length s>0\yt2\ta\ yu2 \ lista\ (tl s) setinterleaves ((yt2, yu2), A)" by (metis (no_types, lifting) add_Suc_right length_Cons length_greater_0_conv lessI th2pre tl_append2) from c have th2pre2: "yu2\lista\aa#yu2\u" by (simp add: le_less less_cons) from b c have th2pre3: "aa\A\a\ A\tl s setinterleaves ((yt2, yu2), A)\ (aa#(tl s)) setinterleaves ((yt2, aa#yu2), A)" by (metis (mono_tags, lifting) addNonSync doubleReverse rev.simps(2) rev_rev_ident) from b c th2pre1 th2pre2 have th2pre4: "aa\A\a\ A\(s @ t) setinterleaves ((ta, u), A)\ length s>0 \yt2\ta\ aa#yu2 \u\ s setinterleaves ((yt2, aa#yu2), A)" using th2pre th2pre3 by fastforce have th2pre5:"aa\A\a\ A\(s @ t) setinterleaves ((ta, u), A)\length s=0 \ []\ta\ [] \u\ s setinterleaves (([], []), A)" by auto have th2: "aa\A\a\ A\\t1 u1. (s @ t) setinterleaves ((ta, u), A) \ t1 \ ta \ u1 \ u \ s setinterleaves ((t1, u1), A)" using th2pre4 th2pre5 by blast from b c have th3pre: "aa\A\a\ A\(s @ t) setinterleaves ((ta, u), A)\tl (s @ t) setinterleaves ((ta, lista), A)\hd (s@t)=aa\tl (s @ t)setinterleaves((list, u), A)\hd(s@t)=a" by auto from a b c th3pre obtain yt31 yu31 where th3pre1: "aa\A\a\ A\ (s @ t) setinterleaves ((ta, u), A)\tl (s @ t) setinterleaves ((ta, lista), A)\ hd (s@t)=aa\length s>0\yt31\ta\ yu31 \ lista\ (tl s) setinterleaves ((yt31, yu31), A)" by(metis(no_types,lifting)add_Suc_right length_Cons length_greater_0_conv lessI tl_append2) from a b c th3pre obtain yt32 yu32 where th3pre2: "aa\A\a\A\(s@t)setinterleaves ((ta, u),A)\ tl (s @ t) setinterleaves ((list, u), A)\hd (s@t)=a\length s>0\yt32\list\ yu32 \u\ (tl s) setinterleaves ((yt32, yu32), A)" by (metis add_less_mono1 impossible_Cons le_neq_implies_less length_greater_0_conv nat_le_linear tl_append2) from b c have th3pre3: "aa\A\a\ A\tl s setinterleaves ((yt31, yu31), A)\ (aa#(tl s)) setinterleaves ((yt31, aa#yu31), A)" by (metis (mono_tags, lifting) addNonSync doubleReverse rev.simps(2) rev_rev_ident) from b c th3pre1 th3pre2 have th3pre4: "aa\A\a\ A\(s @ t) setinterleaves ((ta, u), A)\ length s>0\tl (s @ t) setinterleaves ((ta, lista), A)\hd (s@t)=aa \yt31\ta\ aa#yu31 \u\ s setinterleaves ((yt31, aa#yu31), A)" by (metis hd_append2 le_less length_greater_0_conv less_cons list.exhaust_sel th3pre3) from b c have th3pre5: "aa\A\a\ A\tl s setinterleaves ((yt32, yu32), A)\ (a#(tl s)) setinterleaves ((a#yt32, yu32), A)" by (metis (mono_tags, lifting) addNonSync doubleReverse rev.simps(2) rev_rev_ident) from b c th3pre1 th3pre2 have th3pre6: "aa\A\a\ A\(s @ t) setinterleaves ((ta, u), A)\ length s>0\tl (s @ t) setinterleaves ((list, u), A)\hd (s@t)=a \a#yt32\ta\ yu32 \u\ s setinterleaves ((a#yt32, yu32), A)" using th3pre th3pre5 by (metis hd_append2 le_less length_greater_0_conv less_cons list.exhaust_sel) have th3pre7:"aa\A\a\ A\(s @ t) setinterleaves ((ta, u), A)\length s=0 \ []\ta\ [] \u\ s setinterleaves (([], []), A)" by auto have th3: "aa\A\a\ A\\t1 u1. (s @ t) setinterleaves ((ta, u), A) \ t1 \ ta \ u1 \ u \ s setinterleaves ((t1, u1), A)" using th3pre th3pre4 th3pre6 th3pre7 by blast with a b c th th1 th2 show ?thesis by fastforce qed qed qed lemma SyncWithTick_imp_NTF1: "(s @ [tick]) setinterleaves ((t, u), insert tick (ev ` A)) \ t\ \ P \ u\ \ Q \ \ t u Xa Y. (t, Xa) \ \ P \ ( (u, Y) \ \ Q \ s setinterleaves ((t, u), insert tick (ev ` A)) \ X - {tick} = (Xa \ Y) \ insert tick (ev ` A) \ Xa \ Y)" apply (drule SyncWithTick_imp_NTF) apply simp apply simp proof - assume A: "t \ \ P" and B: "u \ \ Q" and C: "\t1 u1. t = t1 @ [tick]\u = u1@[tick]\s setinterleaves ((t1, u1), insert tick (ev ` A))" from C obtain t1 u1 where D: "t = t1 @ [tick] \ u = u1 @ [tick] \ s setinterleaves ((t1, u1), insert tick (ev ` A))" by blast from A B D have E: "(t1, X-{tick}) \ \ P \ (u1, X-{tick}) \ \ Q" by (simp add: T_F process_charn) thus " \t u Xa Y. (t, Xa) \ \ P \ (u, Y) \ \ Q \ s setinterleaves ((t, u), insert tick (ev ` A)) \ X - {tick} = (Xa \ Y) \ insert tick (ev ` A) \ Xa \ Y" using A B C D E by blast qed lemma ftf_Sync: "front_tickFree u \ front_tickFree t \ s setinterleaves ((t, u),insert tick (ev ` A)) \ front_tickFree s" proof- assume A: "front_tickFree u" assume B: "front_tickFree t" assume C: "s setinterleaves ((t, u), insert tick (ev ` A))" thus "front_tickFree s" proof- from A obtain u1 where A1: "\ u2. u1\u\tickFree u1\ (u2\u\tickFree u2\u2\u1)" by (metis append.right_neutral append_eq_first_pref_spec front_tickFree_charn le_list_def tickFree_Nil) from A A1 have AA1: "u1=u\u=u1@[tick]" by (metis(no_types, lifting) antisym append_Nil2 append_eq_first_pref_spec append_is_Nil_conv front_tickFree_charn le_list_def less_list_def less_self nonTickFree_n_frontTickFree) from B obtain t1 where B1: "\ t2. t1\t\tickFree t1\ (t2\t\tickFree t2\t2\t1)" by (metis append.right_neutral append_eq_first_pref_spec front_tickFree_charn le_list_def tickFree_Nil) from B B1 have BB1: "t1=t\t=t1@[tick]" by (metis(no_types, lifting) antisym append_Nil2 append_eq_first_pref_spec append_is_Nil_conv front_tickFree_charn le_list_def less_list_def less_self nonTickFree_n_frontTickFree) from A1 B1 have C1: "\s1. s1 setinterleaves ((t1, u1), insert tick (ev ` A))\tickFree s1" by (meson ftf_Sync1 tickFree_def) from AA1 BB1 have CC1: "u1=u\t1=t\tickFree s" by (simp add: C C1) from AA1 BB1 have CC2: "u1=u\t=t1@[tick]\tickFree s" using ftf_Sync21 by (metis A1 C equals0D insertI1 non_tickFree_tick tickFree_append tickFree_def) from AA1 BB1 have CC3: "u=u1@[tick]\t1=t\tickFree s" using ftf_Sync21 by (metis B1 C insertI1 insert_not_empty mk_disjoint_insert non_tickFree_tick tickFree_append tickFree_def) from AA1 BB1 have CC4: "u=u1@[tick]\t=t1@[tick]\ \s1. s1 setinterleaves ((t1, u1), insert tick (ev ` A))\(s=s1@[tick])" using C ftf_Sync32 by blast from C1 CC4 have CC5: "front_tickFree s" using AA1 BB1 CC1 CC2 CC3 tickFree_implies_front_tickFree tickFree_implies_front_tickFree_single by auto with A1 B1 C1 AA1 BB1 show ?thesis using CC1 CC2 CC3 tickFree_implies_front_tickFree by auto qed qed lemma is_processT5_SYNC2: " \ sa Y t u Xa Ya. (t,Xa) \ \ P \ (u,Ya) \ \ Q \ (sa setinterleaves ((t,u),insert tick (ev ` A))) \\c. c \ Y \ (\ t1 u1. (sa@[c]) setinterleaves ((t1,u1),insert tick (ev ` A))\ (((t1,{}) \ \ P\(u1, {}) \ \ Q)\((t1,{}) \ \ Q\(u1, {}) \ \ P))) \ \t2 u2 Xb. (t2, Xb) \ \ P \ (\Yb. (u2, Yb) \ \ Q \ sa setinterleaves ((t2, u2), insert tick (ev ` A)) \ ( Xa \ Ya) \ insert tick (ev ` A) \ Xa \ Ya \ Y = (Xb \ Yb) \ insert tick (ev ` A) \ Xb \ Yb)" proof - fix sa Y t u Xa Ya assume A: "(t,Xa) \ \ P \ (u,Ya) \ \ Q \(sa setinterleaves ((t,u),insert tick (ev ` A)))" and B: " \c. c \ Y \ (\ t1 u1. (sa@[c]) setinterleaves ((t1,u1),insert tick (ev ` A)) \ (((t1,{}) \ \ P \ (u1, {}) \ \ Q) \ ((t1,{}) \ \ Q\(u1, {}) \ \ P)))" thus "\t2 u2 Xb. (t2, Xb) \ \ P \ (\Yb. (u2, Yb) \ \ Q \ sa setinterleaves ((t2, u2), insert tick (ev ` A)) \ (Xa \ Ya) \ insert tick (ev ` A) \ Xa \ Ya \ Y = (Xb \ Yb) \ insert tick (ev ` A) \ Xb \ Yb)" proof- from A B obtain Y1 Y2 where A1: "(Y1=(Y\insert tick (ev ` A)))\(Y2=(Y - insert tick (ev ` A)))" by blast from A1 have AA1: "Y1 \ Y2 = Y" by blast from A1 have newAA: "\y\Y1. y\ insert tick (ev ` A)" by blast from A1 have newAAA: "\y\Y2. y\ insert tick (ev ` A)" by blast from B A1 have AA2: "\z\Y1.(t@[z], {}) \ \ P \ (u@[z], {}) \ \ Q" proof(cases "\z \ Y1. (t@[z], {})\\ P \ (u@[z], {}) \ \ Q") case True from True obtain z1 where TrA: "z1\Y1\(t@[z1], {})\\ P \ (u@[z1], {}) \ \ Q" by blast from TrA A have TrB: "\sa setinterleaves ((t,u),insert tick(ev ` A)); z1\insert tick (ev ` A)\ \(sa@[z1]) setinterleaves ((t@[z1],(u@[z1])),insert tick (ev ` A))" proof- have a: "(rev sa) setinterleaves ((rev t,rev u),insert tick (ev ` A))" using local.A doubleReverse by blast have b: "(z1#rev sa) setinterleaves ((z1#rev t,z1# rev u),insert tick (ev ` A))" using TrA a newAA by auto show ?thesis using b doubleReverse by fastforce qed then show ?thesis using A1 B TrA local.A by blast next case False then show ?thesis by blast qed from A B A1 obtain Z1 Z2 where A2: "(Z1=Y1\{z.(t@[z], {})\\ P})\(Z2=Y1-{z.(t@[z], {})\\ P})" by blast from A2 have BA: "Y1=Z1\Z2" by blast from A2 have BAA: "\z\Z1. (t@[z], {}) \ \ P" by blast from A2 have BAAA: "\z\Z2. (u@[z], {})\ \ Q" using AA2 by blast from A2 BA BAA have A3: "(t, Z1)\ \ P" by (meson F_T NF_NT is_processT5_S7 local.A) from A A1 have A43: "\y\Y2. (t@[y], {})\ \ P \ (u@[y], {})\ \ Q" proof(cases "\y\Y2. ((t@[y], {})\ \ P)\((u@[y], {})\ \ Q)") case True from True obtain y1 where innAA: "y1\Y2\(((t@[y1], {})\ \ P)\((u@[y1], {})\ \ Q))" by blast from innAA have innAAA: "y1\ insert tick (ev ` A)" using newAAA by auto from innAA have innAA1: "\sa setinterleaves ((t,u),insert tick (ev ` A)); y1 \ insert tick (ev ` A)\ \ ((t@[y1], {})\ \ P \(sa@[y1]) setinterleaves ((t@[y1],u),insert tick (ev ` A))) \ ((u@[y1], {})\ \ Q\(sa@[y1]) setinterleaves((t,u@[y1]),insert tick (ev`A)))" by (simp add: addNonSync) with A B innAA1 innAA show ?thesis by (metis A1 DiffD1 innAAA is_processT4_empty) next case False then show ?thesis by blast qed from A1 A2 obtain Xbb where B1: "Xbb=(Xa\Z1\Y2)" by blast from A B1 A3 A43 have A5: "(t, Xbb)\ \ P" by (meson BAA is_processT5_S1) from A1 A2 obtain Ybb where B2: "Ybb=(Ya\Z2\Y2)" by blast from A B have A52: "(u, Ybb)\ \ Q" by (metis A43 B2 BAAA is_processT5_S1) from A1 A2 have A61: "(Xbb \ Ybb)\insert tick (ev ` A)=((Xa\Ya\Y1\Y2) \ insert tick (ev ` A))" using B1 B2 by blast from A1 have A62: "(Y1) \ insert tick (ev ` A)=Y1" using inf.orderE by auto from A1 have A63: "(Y2) \ insert tick (ev ` A)={}" by (simp add: AA1 Int_commute) from A61 A62 A63 have A64:"((Xa\Ya\Y1\Y2)\insert tick(ev `A))=((Xa\Ya)\insert tick(ev `A)\Y1)" by (simp add: Int_Un_distrib2) from AA1 BA B1 B2 have A65: "(Xbb\Ybb)\((Xa\Ya)\Y) " by auto from AA1 BA B1 B2 have A66:"(Xbb\Ybb)\((Xa\Ya)\Y2)" by auto from A1 A2 have A66: "((Xa \ Ya) \ insert tick (ev ` A) \ Xa \ Ya \ Y) = ((Xbb \ Ybb) \ insert tick (ev ` A) \ Xbb \ Ybb)" proof - have f1: "(Xa \ Ya) \ insert tick (ev ` A) \ Xa \ Ya \ Y = Y \ Xa \ Ya \ ((Xa \ Ya) \ insert tick (ev ` A) \ Xbb \ Ybb)"using A65 by auto have "Xa \ Ya \ Y2 \ Xbb \ Ybb = Xbb \ Ybb" by (meson A66 le_iff_sup) then have "(Xa \ Ya) \ insert tick (ev ` A) \ Xa \ Ya \ Y = Xbb \ Ybb \ ((Xa \ Ya) \ insert tick (ev ` A) \ Y1)" using f1 A1 by auto then show ?thesis using A61 A64 by force qed with A5 A52 A66 show ?thesis using local.A by blast qed qed lemma pt3: "ta \ \ P \ u \ \ Q \ (s @ t) setinterleaves ((ta, u), insert tick (ev ` A)) \ \t u X. (t, X) \ \ P \ (\Y. (u, Y) \ \ Q \ s setinterleaves ((t, u), insert tick (ev ` A)) \ tick \ X \ tick \ Y \ (X \ Y) \ ev ` A = {} \ X \ Y = {})" proof- assume a1: "ta \ \ P" assume a2: "u \ \ Q" assume a3: "(s @ t) setinterleaves ((ta, u), insert tick (ev ` A))" have aa: "ta \ \ P \u \ \ Q \ (s @ t) setinterleaves ((ta, u), insert tick (ev ` A))\ \t u. t \ \ P \(u \ \ Q \s setinterleaves ((t, u), insert tick (ev ` A)))" using is_processT3_ST_pref synPrefix by blast thus "\t u X. (t, X) \ \ P \(\Y. (u, Y) \ \ Q \ s setinterleaves ((t, u), insert tick (ev ` A)) \ tick \ X \ tick \ Y \ (X \ Y) \ ev ` A = {} \ X \ Y = {})" using NF_NT a1 a2 a3 by blast qed - -lemma Sync_maintains_is_process: - "is_process ({(s,R).\ t u X Y. (t,X) \ \ P \ (u,Y) \ \ Q \ - (s setinterleaves ((t,u),(ev`A) \ {tick})) \ - R = (X \ Y) \ ((ev`A) \ {tick}) \ X \ Y} \ - {(s,R).\ t u r v. front_tickFree v \ (tickFree r \ v=[]) \ - s = r@v \ - (r setinterleaves ((t,u),(ev`A) \ {tick})) \ - (t \ \ P \ u \ \ Q \ t \ \ Q \ u \ \ P)}, - {s. \ t u r v. front_tickFree v \ (tickFree r \ v=[]) \ + +subsection\The Sync Operator Definition\ + +lift_definition Sync :: "['a process,'a set,'a process] => 'a process" ("(3(_ \_\/ _))" [64,0,65] 64) + is "\P A Q. ({(s,R).\ t u X Y. (t,X) \ \ P \ (u,Y) \ \ Q \ + (s setinterleaves ((t,u),(ev`A) \ {tick})) \ + R = (X \ Y) \ ((ev`A) \ {tick}) \ X \ Y} \ + {(s,R).\ t u r v. front_tickFree v \ (tickFree r \ v=[]) \ + s = r@v \ + (r setinterleaves ((t,u),(ev`A) \ {tick})) \ + (t \ \ P \ u \ \ Q \ t \ \ Q \ u \ \ P)}, + + {s. \ t u r v. front_tickFree v \ (tickFree r \ v=[]) \ s = r@v \ (r setinterleaves ((t,u),(ev`A) \ {tick})) \ (t \ \ P \ u \ \ Q \ t \ \ Q \ u \ \ P)})" - (is "is_process(?f, ?d)") -proof(simp only: fst_conv snd_conv Failures_def is_process_def FAILURES_def DIVERGENCES_def, - fold FAILURES_def DIVERGENCES_def,fold Failures_def,intro conjI) - show "([], {}) \ ?f" - apply auto - by (metis Int_commute Int_empty_right Sync.si_empty1 Un_empty empty_iff insert_iff is_processT1) -next - show " \s X. (s, X) \ ?f \ front_tickFree s" - apply (auto simp:is_processT2 append_T_imp_tickFree front_tickFree_append D_imp_front_tickFree) - using ftf_Sync is_processT2 apply blast - using D_imp_front_tickFree ftf_Sync is_processT2_TR apply blast - using D_imp_front_tickFree ftf_Sync is_processT2_TR by blast -next - show "\s t. (s @ t, {}) \ ?f \ (s, {}) \ ?f" - apply auto - apply(drule F_T) - apply(drule F_T) - proof(goal_cases) - case (1 s t ta u X Y) +proof - + show "is_process ({(s,R).\ t u X Y. (t,X) \ \ P \ (u,Y) \ \ Q \ + (s setinterleaves ((t,u),(ev`(A :: 'a set)) \ {tick})) \ + R = (X \ Y) \ ((ev`A) \ {tick}) \ X \ Y} \ + {(s,R).\ t u r v. front_tickFree v \ (tickFree r \ v=[]) \ + s = r@v \ + (r setinterleaves ((t,u),(ev`A) \ {tick})) \ + (t \ \ P \ u \ \ Q \ t \ \ Q \ u \ \ P)}, + + {s. \ t u r v. front_tickFree v \ (tickFree r \ v=[]) \ + s = r@v \ + (r setinterleaves ((t,u),(ev`A) \ {tick})) \ + (t \ \ P \ u \ \ Q \ t \ \ Q \ u \ \ P)})" + (is "is_process(?f, ?d)") for P Q A + unfolding is_process_def FAILURES_def DIVERGENCES_def + proof (simp only: fst_conv snd_conv, intro conjI) + show "([], {}) \ ?f" + apply auto + by (metis Int_commute Int_empty_right Sync.si_empty1 Un_empty empty_iff insert_iff is_processT1) + next + show " \s X. (s, X) \ ?f \ front_tickFree s" + apply (auto simp:is_processT2 append_T_imp_tickFree front_tickFree_append D_imp_front_tickFree) + using ftf_Sync is_processT2 apply blast + using D_imp_front_tickFree ftf_Sync is_processT2_TR apply blast + using D_imp_front_tickFree ftf_Sync is_processT2_TR by blast + next + show "\s t. (s @ t, {}) \ ?f \ (s, {}) \ ?f" + apply auto + apply(drule F_T) + apply(drule F_T) + proof(goal_cases) + case (1 s t ta u X Y) + then show ?case + using pt3 by fastforce + next + case (2 s t ta u r v) + have aa: "front_tickFree v + \ s@t = r @ v \ r setinterleaves ((ta, u), insert tick (ev ` A)) + \ \t u r. r setinterleaves ((t,u),insert tick(ev `A)) \ + (\v. s=r@v\front_tickFree v \ + \tickFree r \ v\[]\(t \ \ P \ + u \ \ Q) \ (t \ \ Q \ u \ \ P)) + \ tickFree r \ ta \ \ P \ u \ \ Q \ s tickFree r \ ta \ \ P + \ u \ \ Q \ s \t u X. (t, X) \ \ P \ (\Y. (u, Y) \ \ Q \ + s setinterleaves ((t, u), insert tick (ev ` A)) \ + tick \ X \ tick \ Y \ (X \ Y) \ ev ` A = {} \ X \ Y = {})" + apply(drule D_T) + apply(simp add: le_list_def less_list_def) + using Sync.sym pt3 by blast + have aab: "\t u X. (t, X) \ \ P \(\Y. (u, Y) \ \ Q \ + s setinterleaves ((t, u), insert tick (ev ` A)) \ + tick \ X \ tick \ Y \ (X \ Y) \ ev ` A = {} \ X \ Y = {})" + using 2(1,2,3,4,5,6,7) aa aaa Sync.sym by auto + then show ?case by auto + next + case (3 s t ta u r v) + have aa: "front_tickFree v \s @ t = r @ v + \ r setinterleaves ((ta, u), insert tick (ev ` A)) + \ \t u r. r setinterleaves ((t, u), insert tick (ev ` A)) \ + (\v. s = r @ v \front_tickFree v \\ tickFree r \ v \ [] \ + (t \ \ P \ u \ \ Q) \ (t \ \ Q \ u \ \ P)) + \ tickFree r + \ ta \ \ Q \ u \ \ P \ s tickFree r \ ta \ \ Q \ u \ \ P \s \t u X. (t, X) \ \ P \ + (\Y. (u, Y) \ \ Q \s setinterleaves ((t, u), insert tick (ev ` A)) \ + tick \ X \ tick \ Y \ (X \ Y) \ ev ` A = {} \ X \ Y = {})" + apply(drule D_T) + apply(simp add: le_list_def less_list_def) + using Sync.sym pt3 by blast + from 3(1,2,3,4,5,6,7) Sync.sym aa have a1: "st u X. (t, X) \ \ P \ (\Y. (u, Y) \ \ Q \ + s setinterleaves ((t, u), insert tick (ev ` A)) \ + tick \ X \ tick \ Y \ (X \ Y) \ ev ` A = {} \ X \ Y = {})" + using 3(1,2,3,4,5,6,7) aa aaa Sync.sym by auto then show ?case - using pt3 by fastforce -next - case (2 s t ta u r v) - have aa: "front_tickFree v - \ s@t = r @ v \ r setinterleaves ((ta, u), insert tick (ev ` A)) - \ \t u r. r setinterleaves ((t,u),insert tick(ev `A)) \ - (\v. s=r@v\front_tickFree v \ - \tickFree r \ v\[]\(t \ \ P \ - u \ \ Q) \ (t \ \ Q \ u \ \ P)) - \ tickFree r \ ta \ \ P \ u \ \ Q \ s tickFree r \ ta \ \ P - \ u \ \ Q \ s \t u X. (t, X) \ \ P \ (\Y. (u, Y) \ \ Q \ - s setinterleaves ((t, u), insert tick (ev ` A)) \ - tick \ X \ tick \ Y \ (X \ Y) \ ev ` A = {} \ X \ Y = {})" - apply(drule D_T) - apply(simp add: le_list_def less_list_def) - using Sync.sym pt3 by blast - have aab: "\t u X. (t, X) \ \ P \(\Y. (u, Y) \ \ Q \ - s setinterleaves ((t, u), insert tick (ev ` A)) \ - tick \ X \ tick \ Y \ (X \ Y) \ ev ` A = {} \ X \ Y = {})" - using 2(1,2,3,4,5,6,7) aa aaa Sync.sym by auto - then show ?case by auto -next - case (3 s t ta u r v) - have aa: "front_tickFree v \s @ t = r @ v - \ r setinterleaves ((ta, u), insert tick (ev ` A)) - \ \t u r. r setinterleaves ((t, u), insert tick (ev ` A)) \ - (\v. s = r @ v \front_tickFree v \\ tickFree r \ v \ [] \ - (t \ \ P \ u \ \ Q) \ (t \ \ Q \ u \ \ P)) - \ tickFree r - \ ta \ \ Q \ u \ \ P \ s tickFree r \ ta \ \ Q \ u \ \ P \s \t u X. (t, X) \ \ P \ - (\Y. (u, Y) \ \ Q \s setinterleaves ((t, u), insert tick (ev ` A)) \ - tick \ X \ tick \ Y \ (X \ Y) \ ev ` A = {} \ X \ Y = {})" - apply(drule D_T) - apply(simp add: le_list_def less_list_def) - using Sync.sym pt3 by blast - from 3(1,2,3,4,5,6,7) Sync.sym aa have a1: "st u X. (t, X) \ \ P \ (\Y. (u, Y) \ \ Q \ - s setinterleaves ((t, u), insert tick (ev ` A)) \ - tick \ X \ tick \ Y \ (X \ Y) \ ev ` A = {} \ X \ Y = {})" - using 3(1,2,3,4,5,6,7) aa aaa Sync.sym by auto - then show ?case - by auto -next - case (4 s t ta u) - from 4(3) have a1: "ta \ \ P" - by (simp add: D_T) - then show ?case - using "4"(1) "4"(4) pt3 by blast -next - case (5 s t ta u) - from 5(3) have a1: "ta \ \ Q" - by (simp add: D_T) - then show ?case - using "5"(1) "5"(4) Sync.sym pt3 by blast -qed -next - show "\s X Y. (s, Y) \ ?f \ X \ Y \ (s, X) \ ?f" - apply auto - proof(goal_cases) - case (1 s X t u Xa Ya) - have aa: "X \ ((Xa \ Ya) \ insert tick (ev ` A) \ Xa \ Ya )\\x1 y1. x1\Xa \ y1\Ya \ - X=((x1 \ y1) \ insert tick (ev ` A) \ x1 \ y1)" - apply(simp add: Set.subset_iff_psubset_eq) - apply(erule disjE) - defer - apply blast - proof- - assume A: "X \ (Xa \ Ya) \ insert tick (ev ` A) \ Xa \ Ya" - from A obtain X1 where B: "X1=((Xa \ Ya) \ insert tick (ev ` A) \ Xa \ Ya)-X" - by blast - from A obtain x where C: "x=Xa-X1" - by blast - from A obtain y where D: "y=Ya-X1" - by blast - from A B C D have E: "X = (x \ y) \ insert tick (ev ` A) \ x \ y" - by blast - thus " \x1. (x1 \ Xa \ x1 = Xa) \ (\y1. (y1 \ Ya \ y1 = Ya) \ X = (x1 \ y1) \ - insert tick (ev ` A) \ x1 \ y1)" - using A B C D E by (metis Un_Diff_Int inf.strict_order_iff inf_sup_absorb) - qed - then show ?case using 1(1,2,3,4,5) - by (meson process_charn) + by auto + next + case (4 s t ta u) + from 4(3) have a1: "ta \ \ P" + by (simp add: D_T) + then show ?case + using "4"(1) "4"(4) pt3 by blast + next + case (5 s t ta u) + from 5(3) have a1: "ta \ \ Q" + by (simp add: D_T) + then show ?case + using "5"(1) "5"(4) Sync.sym pt3 by blast + qed + next + show "\s X Y. (s, Y) \ ?f \ X \ Y \ (s, X) \ ?f" + apply auto + proof(goal_cases) + case (1 s X t u Xa Ya) + have aa: "X \ ((Xa \ Ya) \ insert tick (ev ` A) \ Xa \ Ya )\\x1 y1. x1\Xa \ y1\Ya \ + X=((x1 \ y1) \ insert tick (ev ` A) \ x1 \ y1)" + apply(simp add: Set.subset_iff_psubset_eq) + apply(erule disjE) + defer + apply blast + proof- + assume A: "X \ (Xa \ Ya) \ insert tick (ev ` A) \ Xa \ Ya" + from A obtain X1 where B: "X1=((Xa \ Ya) \ insert tick (ev ` A) \ Xa \ Ya)-X" + by blast + from A obtain x where C: "x=Xa-X1" + by blast + from A obtain y where D: "y=Ya-X1" + by blast + from A B C D have E: "X = (x \ y) \ insert tick (ev ` A) \ x \ y" + by blast + thus " \x1. (x1 \ Xa \ x1 = Xa) \ (\y1. (y1 \ Ya \ y1 = Ya) \ X = (x1 \ y1) \ + insert tick (ev ` A) \ x1 \ y1)" + using A B C D E by (metis Un_Diff_Int inf.strict_order_iff inf_sup_absorb) + qed + then show ?case using 1(1,2,3,4,5) + by (meson process_charn) + qed + + next + let ?f1 = "{(s,R). \ t u X Y. (t,X) \ \ P \ (u,Y) \ \ Q \ + (s setinterleaves ((t,u),(ev`A) \ {tick})) \ + R = (X \ Y) \ ((ev`A) \ {tick}) \ X \ Y}" + have is_processT5_SYNC3: + "\sa X Y.(sa, X) \?f1 \(\c. c\Y\(sa@[c],{})\?f)\(sa, X\Y) \ ?f1" + apply(clarsimp) + apply(rule is_processT5_SYNC2[simplified]) + apply(auto simp:is_processT5) apply blast + by (metis Sync.sym Un_empty_right inf_sup_absorb inf_sup_aci(5) insert_absorb insert_not_empty) + show "\s X Y. (s, X) \ ?f \ (\c. c \ Y \ (s @ [c], {}) \ ?f) \ (s, X \ Y) \ ?f" + apply(intro allI impI, elim conjE UnE) + apply(rule rev_subsetD) + apply(rule is_processT5_SYNC3) + by(auto) + next + show "\s X. (s @ [tick], {}) \ ?f \ (s, X - {tick}) \ ?f" + apply auto + apply(drule F_T) + apply(drule F_T) + using SyncWithTick_imp_NTF1 apply fastforce + apply(metis append_assoc append_same_eq front_tickFree_dw_closed nonTickFree_n_frontTickFree + non_tickFree_tick tickFree_append) + apply(metis butlast_append butlast_snoc front_tickFree_charn non_tickFree_tick tickFree_append + tickFree_implies_front_tickFree) + apply (metis NT_ND append_Nil2 front_tickFree_Nil is_processT3_ST is_processT9_S_swap + SyncWithTick_imp_NTF) + by (metis NT_ND append_Nil2 front_tickFree_Nil is_processT3_ST is_processT9_S_swap + SyncWithTick_imp_NTF) + + next + show "\s t. s \ ?d \ tickFree s \ front_tickFree t \ s @ t \ ?d" + apply auto + using front_tickFree_append apply blast + using front_tickFree_append apply blast + apply blast + by blast + next + show "\s X. s \ ?d \ (s, X) \ ?f" + by blast + next + show "\s. s @ [tick] \ ?d \ s \ ?d" + apply auto + apply (metis butlast_append butlast_snoc front_tickFree_charn non_tickFree_tick + tickFree_append tickFree_implies_front_tickFree) + apply (metis butlast_append butlast_snoc front_tickFree_charn non_tickFree_tick + tickFree_append tickFree_implies_front_tickFree) + apply (metis D_T append.right_neutral front_tickFree_Nil is_processT3_ST is_processT9 + SyncWithTick_imp_NTF) + by (metis D_T append.right_neutral front_tickFree_Nil is_processT3_ST is_processT9 + SyncWithTick_imp_NTF) qed - -next - let ?f1 = "{(s,R). \ t u X Y. (t,X) \ \ P \ (u,Y) \ \ Q \ - (s setinterleaves ((t,u),(ev`A) \ {tick})) \ - R = (X \ Y) \ ((ev`A) \ {tick}) \ X \ Y}" - have is_processT5_SYNC3: - "\sa X Y.(sa, X) \?f1 \(\c. c\Y\(sa@[c],{})\?f)\(sa, X\Y) \ ?f1" - apply(clarsimp) - apply(rule is_processT5_SYNC2[simplified]) - apply(auto simp:is_processT5) apply blast - by (metis Sync.sym Un_empty_right inf_sup_absorb inf_sup_aci(5) insert_absorb insert_not_empty) - show "\s X Y. (s, X) \ ?f \ (\c. c \ Y \ (s @ [c], {}) \ ?f) \ (s, X \ Y) \ ?f" - apply(intro allI impI, elim conjE UnE) - apply(rule rev_subsetD) - apply(rule is_processT5_SYNC3) - by(auto) -next - show "\s X. (s @ [tick], {}) \ ?f \ (s, X - {tick}) \ ?f" - apply auto - apply(drule F_T) - apply(drule F_T) - using SyncWithTick_imp_NTF1 apply fastforce - apply(metis append_assoc append_same_eq front_tickFree_dw_closed nonTickFree_n_frontTickFree - non_tickFree_tick tickFree_append) - apply(metis butlast_append butlast_snoc front_tickFree_charn non_tickFree_tick tickFree_append - tickFree_implies_front_tickFree) - apply (metis NT_ND append_Nil2 front_tickFree_Nil is_processT3_ST is_processT9_S_swap - SyncWithTick_imp_NTF) - by (metis NT_ND append_Nil2 front_tickFree_Nil is_processT3_ST is_processT9_S_swap - SyncWithTick_imp_NTF) - -next - show "\s t. s \ ?d \ tickFree s \ front_tickFree t \ s @ t \ ?d" - apply auto - using front_tickFree_append apply blast - using front_tickFree_append apply blast - apply blast - by blast -next - show "\s X. s \ ?d \ (s, X) \ ?f" - by blast -next - show "\s. s @ [tick] \ ?d \ s \ ?d" - apply auto - apply (metis butlast_append butlast_snoc front_tickFree_charn non_tickFree_tick - tickFree_append tickFree_implies_front_tickFree) - apply (metis butlast_append butlast_snoc front_tickFree_charn non_tickFree_tick - tickFree_append tickFree_implies_front_tickFree) - apply (metis D_T append.right_neutral front_tickFree_Nil is_processT3_ST is_processT9 - SyncWithTick_imp_NTF) - by (metis D_T append.right_neutral front_tickFree_Nil is_processT3_ST is_processT9 - SyncWithTick_imp_NTF) qed -lemmas Rep_Abs_Sync[simp] = Abs_process_inverse[simplified, OF Sync_maintains_is_process] - -subsection\Projections\ +subsection\The Projections\ lemma F_Sync : "\ (P \ A \ Q) = {(s,R).\ t u X Y. (t,X) \ \ P \ (u,Y) \ \ Q \ s setinterleaves ((t,u),(ev`A) \ {tick}) \ R = (X \ Y) \ ((ev`A) \ {tick}) \ X \ Y} \ {(s,R).\ t u r v. front_tickFree v \ (tickFree r \ v=[]) \ s = r@v \ r setinterleaves ((t,u),(ev`A)\{tick}) \ (t \ \ P \ u \ \ Q \ t \ \ Q \ u \ \ P)}" - - - unfolding Sync_def - apply(subst Failures_def) - apply(simp only: Rep_Abs_Sync) - by(auto simp: FAILURES_def) - + by (simp add: Failures.rep_eq Sync.rep_eq FAILURES_def) lemma D_Sync : "\ (P \ A \ Q) = {s. \ t u r v. front_tickFree v \ (tickFree r \ v=[]) \ s = r@v \ r setinterleaves ((t,u),(ev`A) \ {tick}) \ (t \ \ P \ u \ \ Q \ t \ \ Q \ u \ \ P)}" - - unfolding Sync_def - apply(subst D_def) - apply(simp only: Rep_Abs_Sync) - by(simp add: DIVERGENCES_def) + by (simp add: Divergences.rep_eq Sync.rep_eq DIVERGENCES_def) lemma T_Sync : "\ (P \ A \ Q) = {s. \ t u. t \ \ P \ u \ \ Q \ s setinterleaves ((t,u),(ev`A) \ {tick})}\ {s. \ t u r v. front_tickFree v \ (tickFree r \ v=[]) \ s = r@v \ r setinterleaves ((t,u),(ev`A) \ {tick}) \ (t \ \ P \ u \ \ Q \ t \ \ Q \ u \ \ P)}" - - unfolding Sync_def - apply(subst Traces_def, simp only: Rep_Abs_Sync, simp add:TRACES_def FAILURES_def) - apply auto - apply (meson F_T) - using T_F by blast + by (simp add: Traces.rep_eq TRACES_def Failures.rep_eq[symmetric] F_Sync) blast + subsection\Syntax for Interleave and Parallel Operator \ abbreviation Inter_syntax ("(_|||_)" [72,73] 72) where "P ||| Q \ (P \ {} \ Q)" abbreviation Par_syntax ("(_||_)" [74,75] 74) where "P || Q \ (P \ UNIV \ Q)" subsection\ Continuity Rule \ lemma mono_Sync_D1: "P \ Q \ \ (Q \ A \ S) \ \ (P \ A \ S)" apply(auto simp: D_Sync) using le_approx1 apply blast using le_approx_lemma_T apply blast apply (metis append_Nil2 le_approx1 subsetCE tickFree_Nil tickFree_implies_front_tickFree) by (metis append_Nil2 le_approx_lemma_T subsetCE tickFree_Nil tickFree_implies_front_tickFree) lemma mono_Sync_D2: "P \ Q\(\ s. s \ \ (P \ A \ S) \ Ra (P \ A \ S) s = Ra (Q \ A \ S) s)" apply auto apply(simp add: Ra_def D_Sync F_Sync) apply (metis (no_types, lifting) F_T append_Nil2 front_tickFree_Nil le_approx2) apply(simp add: Ra_def D_Sync F_Sync) apply(insert le_approx_lemma_F[of P Q] le_approx_lemma_T[of P Q] le_approx1[of P Q] le_approx2T[of P Q], auto) apply blast apply blast apply blast using front_tickFree_Nil apply blast using front_tickFree_Nil by blast lemma interPrefixEmpty: "s setinterleaves ((t,[]),A)\t1\s1u1 s1. s setinterleaves((t,u),A)\t1u1\u\s1s1 setinterleaves((t1,u1),A)" proof (induction "length t + length u" arbitrary:s t u t1 rule:nat_less_induct) case 1 then show ?case apply(cases t) apply auto[1] apply(cases u) apply (meson interPrefixEmpty nil_le2) proof- fix a list aa lista assume a: " \mx xa. m = length x + length xa \ (\xb xc. \u1 s1. xb setinterleaves((x, xa),A)\xc < x \ u1 \ xa \ s1 < xb \ s1 setinterleaves ((xc, u1), A))" and b: "t = a # list" and c: "u = aa # lista" thus "\u1 s1. s setinterleaves((t, u),A)\t1u1 \u\s1s1 setinterleaves ((t1, u1), A)" proof- from b c have th0pre: "a=aa\aa\ A\s setinterleaves ((t, u), A) \ (tl s) setinterleaves ((list,lista), A)" by auto from a obtain yu ys where th0pre1: "a=aa\aa\ A \ s setinterleaves ((t, u), A)\ t1 < a # list \ length t1>0 \ yu \ lista \ ys ys setinterleaves ((tl t1, yu), A)" apply (simp add: le_list_def less_list_def append_eq_Cons_conv) by (metis add_less_mono b c length_Cons lessI list.sel(3) th0pre) from th0pre th0pre1 b c have th0pre2: "a=aa \ aa\A \ s setinterleaves ((a # list, u), A) \ t1length t1>0 \ (a#ys)setinterleaves((a#tl t1, a#yu), A)\a#ysa#yu\u\t1=a#tl t1" apply (simp add: le_list_def less_list_def Cons_eq_append_conv) by (metis append_Cons list.exhaust_sel list.inject list.sel(3)) have th0pre3: "a=aa\aa\A\s setinterleaves ((a # list, u), A)\t1 < a # list\length t1=0\ [][]\u\ [] setinterleaves ((t1, []), A)" apply (simp add: le_list_def less_list_def) using c by auto from th0pre2 th0pre3 c have th0 : "a=aa\aa\A\s setinterleaves ((a # list, u), A)\ t1 < a # list\\u1\u. \s1A\aa\ A\s setinterleaves ((a # list, u), A) \ (tl s) setinterleaves ((list,aa#lista), A)" by auto from a obtain yu1 ys1 where th1pre1: "a\A\aa\ A\s setinterleaves ((a # list, u), A)\ t1 < a # list\length t1>0\yu1\u\ys1ys1 setinterleaves ((tl t1, yu1), A)" apply (simp add: le_list_def less_list_def append_eq_Cons_conv) by (metis add_Suc b c length_Cons lessI list.sel(3) th1pre) from th1pre1 have th1pre2: "a\A\aa\ A \ ys1 setinterleaves ((tl t1, yu1), A)\ (a#ys1) setinterleaves ((a#tl t1, yu1), A)" apply simp by (metis (mono_tags, lifting) addNonSync doubleReverse rev.simps(2) rev_rev_ident) from th1pre th1pre1 b c have th1pre3: "a\A\aa\ A\s setinterleaves ((a # list, u), A)\ t1length t1>0\(a#ys1) setinterleaves ((a#tl t1, yu1), A)\a#ys1yu1\u\t1=a#tl t1" apply (simp add: le_list_def less_list_def ) by (metis append_Cons list.exhaust_sel list.inject list.sel(3) th1pre2) have th1preEmpty: "a\A\aa\ A\s setinterleaves ((a # list, u), A)\t1 < a # list\length t1=0 \[][]\u\ [] setinterleaves ((t1, []), A)" apply (simp add: le_list_def less_list_def) using c by auto from c have th1: "a\A\aa\ A\s setinterleaves ((a # list, u), A)\t1 < a # list\ \u1\u. \s1A\a\ A\s setinterleaves ((a # list, u), A) \ (tl s) setinterleaves ((a#list,lista), A)" by auto from a th2pre obtain yu2 ys2 where th2pre1: "aa\A\a\ A\s setinterleaves ((a # list, u), A)\ t1 < a # list\length t1>0\yu2\lista\ys2ys2 setinterleaves ((t1, yu2), A)" apply (simp add: le_list_def less_list_def) by (metis add_Suc_right b c length_Cons lessI) from th2pre1 have th2pre2: "aa\A\a\ A \ ys2 setinterleaves ((t1, yu2), A)\ (aa#ys2) setinterleaves ((t1, aa#yu2), A)" apply simp by (metis (mono_tags, lifting) addNonSync doubleReverse rev.simps(2) rev_rev_ident) from th2pre th2pre1 b c have th2pre3: "aa\A\a\ A \s setinterleaves ((a # list, u), A)\ t1 < a # list\length t1>0\(aa#ys2) setinterleaves ((t1, aa#yu2), A)\aa#ys2aa#yu2\u" apply (simp add: le_list_def less_list_def ) using th2pre2 by auto have th2preEmpty: "aa\A\a\ A\s setinterleaves ((a # list, u), A)\t1 < a # list\length t1=0 \[][]\u\ [] setinterleaves ((t1, []), A)" apply (simp add: le_list_def less_list_def) using c by auto from th2pre3 b c have th2: "aa\A\a\ A\s setinterleaves ((a # list, u), A)\t1 < a # list\ \u1\u. \s1A\a\ A\s setinterleaves ((a # list, u), A) \ (tl s) setinterleaves ((a#list,lista), A)\hd s=aa\(tl s) setinterleaves ((list,u), A)\hd s=a" by auto from a c b th3pre obtain yu3 ys3 where th3pre1: "(aa\A\a\ A\s setinterleaves ((a # list, u), A)\ t1 < a # list\length t1>0\(tl s) setinterleaves ((a#list,lista), A)\ yu3\lista\ys3ys3 setinterleaves((t1,yu3),A))" apply(simp add:le_list_def less_list_def) by (metis (no_types, lifting) add_Suc length_Cons lessI) have adsmallPrefix: "t1length t1>0\tl t1A\a\ A\ (tl s) setinterleaves ((list,u), A)\hd s=a\tl t1< list\yu30\u\ys30 ys30 setinterleaves ((tl t1, yu30), A))" apply (simp add: le_list_def less_list_def) by (metis (no_types, lifting) add_Suc_right length_Cons lessI) have th3pre1more1: "yu3\lista \aa#yu3\u " using c by (metis le_less less_cons) have th3pre1more2: "aa\A\a\ A\s setinterleaves ((a # list, aa#lista), A) \ (tl s) setinterleaves ((a#list,lista), A)\hd s=aa\ys3aa#ys3 A \a \ A \ ys3ys3 setinterleaves ((t1, yu3), A) \ hd s = aa \(aa # ys3) setinterleaves ((t1, aa # yu3), A)" by (metis (mono_tags, lifting) addNonSync doubleReverse rev.simps(2) rev_rev_ident) from th3pre1more1 th3pre1more2 th3pre1more21 th3pre1 have th3pre1more3: "(aa\A\a\ A\ s setinterleaves ((a # list, u), A)\ t1 < a # list\length t1>0\(tl s) setinterleaves ((a#list,lista), A)\hd s=aa\aa#yu3\u\aa#ys3(aa#ys3) setinterleaves ((t1, aa#yu3), A))" using c by blast have th3pre1more32: "aa\A\a\ A\s setinterleaves ((a # list, aa#lista), A) \ (tl s) setinterleaves ((list,aa#lista), A)\hd s=a\ys30a#ys30 A \a \ A \ ys30ys30 setinterleaves ((tl t1, yu30), A) \ hd s = a \t1(a # ys30) setinterleaves ((a#tl t1, yu30), A)" by (metis (mono_tags, lifting) addNonSync doubleReverse rev.simps(2) rev_rev_ident) from th3pre1more32 th3pre12 th3pre1more213 have th3pre1more33: "(aa\A\a\ A\s setinterleaves ((a # list, u), A)\ t1 < a # list\length t1>0\(tl s) setinterleaves ((list,aa#lista), A)\ hd s=a\yu30\u\a#ys30(a#ys30) setinterleaves ((t1, yu30), A))" by (metis adsmallPrefix c hd_append2 le_list_def length_greater_0_conv list.exhaust_sel list.sel(1) order.strict_implies_order) from th3pre1more3 th3pre1more33 b c th3pre have th3NoEmpty: "aa\A\a\ A\s setinterleaves ((a # list, u), A)\t1 < a # list\length t1>0\\u1\u. \s1A\a\ A\s setinterleaves ((a # list, u), A)\t1 < a # list\length t1=0 \[][]\u\ [] setinterleaves ((t1, []), A)" apply (simp add: le_list_def less_list_def) using c by auto from th3NoEmpty th3preEmpty have th3: "aa\A\a\ A\s setinterleaves ((a # list, u), A)\t1 < a # list\\u1\u. \s1min_elems(\ (P \A\ S)) \ \t u. r setinterleaves ((t, u),insert tick (ev ` A)) \ (t \ min_elems(\ P) \ u \ \ S \ t \ min_elems(\ S) \ u \ \ P \ u \ (\ P - (\ P-min_elems(\ P))))" proof- fix r assume A: "r \ min_elems(\ (P \A\ S))" thus "\t u. r setinterleaves ((t, u), insert tick (ev ` A))\(t \ min_elems(\ P) \ u \ \ S \ t \ min_elems(\ S) \ u \ \ P \ u \ (\ P - (\ P-min_elems(\ P))))" proof(cases "\t u r1. r1\r\ r1 setinterleaves((t, u),insert tick (ev ` A)) \ (t \ min_elems(\ P)) \ (t \ min_elems(\ S))") case True from A have AA: "r \ \ (P \A\ S)" using elem_min_elems by blast from AA obtain t1 u1 where A1: "\r1\r. r1 setinterleaves ((t1, u1), insert tick (ev ` A)) \ (t1 \ \ P \ t1 \ \ S)" apply(simp add: D_Sync) using le_list_def by blast from True A1 have A2: "(t1 \ min_elems(\ P)) \ (t1 \ min_elems(\ S))" by blast from A1 A2 min_elems5 obtain tm1 tm2 where tmA: "(t1\\ P \ (tm1\t1\tm1\ min_elems(\ P))) \ (t1\\ S\(tm2\t1\tm2\ min_elems(\ S)))" by metis from A2 tmA have AB1: "(t1\\ P\tm1(t1\\ S\tm2\ P\um1\u1 \ rm1 setinterleaves ((tm1, um1),insert tick (ev ` A)))\(t1\\ S\um2\u1 \ rm2 setinterleaves ((tm2, um2), insert tick (ev ` A)))" by (meson interPrefix) from A1 A2 True tmA AB1 have A3: "(t1\\ P\rm1rm1\\(P\A\S))\(t1\\ S\rm2rm2\ \(P\A\S))" by (meson dual_order.strict_implies_order interPrefix order_trans) from A3 AA have ATrue: "r \ min_elems(\ (P \A\ S))" using min_elems_def using A1 by blast with A ATrue show ?thesis by blast next case False from A have dsync: "r\\ (P \A\ S)" by (simp add: local.A elem_min_elems) from dsync obtain r1 t2 u2 where C: "r1\r\ r1 setinterleaves ((t2, u2), insert tick (ev ` A)) \ ((t2 \ \ P \ u2 \ \ S) \ (t2 \ \ S \ u2 \ \ P))" apply(simp add: D_Sync) using le_list_def by blast from C have r1D: "r1 \ \ (P \A\ S)" apply(simp add: D_Sync) using front_tickFree_Nil by blast from A C r1D have eq: "r1=r" apply(simp add: min_elems_def) using le_neq_trans by blast from A False have minAa: "(t2 \ \ P\u2\\ S)\(t2 \ min_elems(\ P))" proof(cases "t2 \ \ P \u2 \ \ S\ t2 \ min_elems(\ P)") case True from True obtain t3 where inA: "t3t3\min_elems(\ P)" by (metis le_imp_less_or_eq min_elems5) from inA obtain r3 u3 where inA1: "u3\u2\r3 setinterleaves((t3,u3),insert tick(ev ` A))\r3 \ S" using True is_processT3_ST_pref by blast from inA1 have inA2: "r3 \ \(P \A\ S)" apply (simp add: D_Sync) using elem_min_elems front_tickFree_Nil inA inA3 by blast with A eq inA1 show ?thesis by(simp add: min_elems_def) next case False then show ?thesis by blast qed from A False have minAb1: "(t2 \ \ S\u2\ \ P)\(t2 \ min_elems(\ S))" proof(cases "t2 \ \ S \ u2 \ \ P\ t2 \ min_elems(\ S)") case True from True obtain t3 r3 u3 where inA: "t3t3\min_elems(\ S)" and inA1: "u3\u2\r3 setinterleaves ((t3, u3), insert tick (ev ` A))\r3 \ P" using True is_processT3_ST_pref by blast from inA1 have inA2: "r3 \ \ (P \A\ S)" apply (simp add: D_Sync) using elem_min_elems front_tickFree_Nil inA inA3 by blast with A eq inA1 show ?thesis by(simp add: min_elems_def) next case False then show ?thesis by blast qed from A False have minAb2: "(t2 \ \ S\u2\ \ P)\(u2 \ (\ P-(\ P-min_elems(\ P))))" proof(cases "t2 \ \ S \u2 \ \ P\ u2 \(\ P-min_elems(\ P))") case True from True have inAAA: "u2\\ P\u2\min_elems(\ P)" by blast from inAAA obtain u3 where inA: "u3u3\min_elems(\ P)" by (metis le_neq_trans min_elems5) from inA obtain t3 r3 where inA1: "t3\t2\r3 setinterleaves((t3,u3),insert tick(ev ` A))\r3 \ P\t3 \ \ S" using D_T True elem_min_elems inA is_processT3_ST_pref by blast from inA1 have inA2: "r3 \ \ (P \A\ S)" apply (simp add: D_Sync) using Sync.sym front_tickFree_Nil inA3 by blast with A inA1 inA2 show ?thesis by(simp add: min_elems_def) next case False then show ?thesis by blast qed with A C show ?thesis using eq minAa minAb1 by blast qed qed lemma mono_Sync_D3: assumes ordered: "P \ Q" shows "min_elems (\ (P \A\ S)) \ \(Q \A\ S)" proof- have mono_sync2b:"P \ Q \ \r. r \ min_elems(\ (P\A\S))\r \ \(Q\A\S)" apply(frule impI) apply(auto simp: mono_Sync2a) proof - fix r assume A: "P \ Q" and B: "r \ min_elems(\(P\A\S))" from B obtain t u where E: "r setinterleaves ((t, u), insert tick (ev ` A)) \ (t \ min_elems(\ P) \ u \ \ S \ t \ min_elems(\ S) \ u \ (\ P-(\ P - min_elems(\ P))))" using mono_Sync2a by blast from E have F:"(t \ \ Q \ u \ \ S) \ (t \ \ S \ u \ \ Q)" using le_approx2T le_approx3 ordered by blast thus "r \ \ (Q \A\ S)" by (metis (no_types, lifting) E Sync.sym T_Sync UnCI Un_commute insert_is_Un mem_Collect_eq) qed show ?thesis apply(insert ordered) apply(frule mono_sync2b) by blast qed lemma mono_D_Sync: "\ (S \ A \ Q) = \ (Q \ A \ S)" by(auto simp: D_Sync) lemma mono_T_Sync: "\ (S \ A \ Q) = \ (Q \ A \ S)" apply(auto simp: T_Sync) using Sync.sym by fastforce+ lemma mono_F_Sync: "\ (S \ A \ Q) = \ (Q \ A \ S) " apply auto apply (simp add: F_Sync) using Sync.sym apply blast apply (simp add: F_Sync) using Sync.sym by blast lemma mono_Ra_Sync: "Ra (S \ A \ Q) s = Ra (Q \ A \ S) s" apply auto apply (auto simp: Ra_def) by (auto simp: mono_F_Sync) lemma mono_Sync[simp] : "P \ Q \ (P \ A \ S) \ (Q \ A \ S)" by(auto simp: le_approx_def mono_Sync_D1 mono_Sync_D2 mono_Sync_D3) lemma mono_Sync_sym [simp]: "P \ Q \ (S \ A \ P) \ (S \ A \ Q)" by (metis Process_eq_spec mono_D_Sync mono_F_Sync mono_Sync) lemma chain_Sync1: "chain Y \ chain (\i. Y i \ A \ S)" by(simp add: chain_def) lemma chain_Sync2: "chain Y \ chain (\i. S \ A \ Y i)" by(simp add: chain_def) lemma empty_setinterleaving : "[] setinterleaves ((t, u), A) \ t = []" by (cases t, cases u, auto, cases u, simp_all split:if_splits) lemma inters_fin_fund: "finite{(t, u). s setinterleaves ((t, u), A)}" proof (induction "length s" arbitrary:s rule:nat_less_induct) case 1 have A:"{(t, u). s setinterleaves((t, u), A)} \ {([],[])} \ {(t, u). s setinterleaves ((t, u), A) \ (\ a list. t = a#list \ a \ A) \ u = []} \ {(t, u). s setinterleaves ((t, u), A) \ (\ a list. u = a#list \ a \ A) \ t = []} \ {(t, u). s setinterleaves ((t, u), A) \ (\ a list aa lista. u = a#list \ t = aa#lista)}" (is "?A \ {([],[])} \ ?B \ ?C \ ?D") apply (rule subsetI, safe) apply(simp_all add: neq_Nil_conv) apply (metis Sync.si_empty2 Sync.sym empty_iff list.exhaust_sel) using list.exhaust_sel apply blast apply (metis Sync.sym emptyLeftNonSync list.exhaust_sel list.set_intros(1)) using list.exhaust_sel apply blast apply (metis emptyLeftNonSync list.exhaust_sel list.set_intros(1)) by (metis Sync.si_empty2 Sync.sym empty_iff list.exhaust_sel) have a1:"?B \ { ((hd s#list), [])| list. (tl s) setinterleaves ((list, []), A) \ (hd s) \ A}" (is "?B \ ?B1") by auto define f where a2:"f = (\a (t, (u::'a event list)). ((a::'a event)#t, ([]::'a event list)))" have a3:"?B1 \ {(hd s # list, []) |list. tl s setinterleaves ((list, []), A)} " (is "?B1 \ ?B2") by blast from a1 a3 have a13:"?B \ ?B2" by simp have AA: "finite ?B" proof (cases s) case Nil then show ?thesis using not_finite_existsD by fastforce next case (Cons a list) hence aa:"finite{(t,u).(tl s) setinterleaves((t, u), A)}"using 1[THEN spec, of "length (tl s)"] by (simp) have "{(hd s#list, [])|list. tl s setinterleaves((list,[]),A)}\(\(t, u).f(hd s)(t, u)) `{(t, u). tl s setinterleaves ((t, u), A)}" using a2 by auto hence "finite ?B2" using finite_imageI [of "{(t, u). (tl s) setinterleaves ((t, u), A)}" "\(t, u). f (hd s) (t, u)", OF aa] using rev_finite_subset by auto then show ?thesis using a13 by (meson rev_finite_subset) qed have a1:"?C \ { ([],(hd s#list))| list. (tl s) setinterleaves (([],list), A) \ (hd s) \ A}" (is "?C \ ?C1") by auto define f where a2:"f = (\a ((t::'a event list), u). (([]::'a event list), (a::'a event)#u))" have a3:"?C1 \ {([],hd s # list) |list. tl s setinterleaves (([],list), A)} " (is "?C1 \ ?C2") by blast from a1 a3 have a13:"?C \ ?C2" by simp have AAA:"finite ?C" proof (cases s) case Nil then show ?thesis using not_finite_existsD by fastforce next case (Cons a list) hence aa:"finite {(t,u).(tl s)setinterleaves((t, u), A)}"using 1[THEN spec, of "length (tl s)"] by (simp) have "{([], hd s # list) |list. tl s setinterleaves (([], list), A)} \ (\(t, u). f (hd s) (t, u)) ` {(t, u). tl s setinterleaves ((t, u), A)}" using a2 by auto hence "finite ?C2" using finite_imageI [of "{(t, u). (tl s) setinterleaves ((t, u), A)}" "\(t, u). f (hd s) (t, u)", OF aa] using rev_finite_subset by auto then show ?thesis using a13 by (meson rev_finite_subset) qed have dd0:"?D \ {(a#l, aa#la)|a aa l la. s setinterleaves ((a#l, aa#la), A)}" (is "?D \ ?D1") apply (rule subsetI, auto, simp split:if_splits) by (blast, blast, metis, blast) have AAAA:"finite ?D" proof (cases s) case Nil hence "?D \ {([],[])}" using empty_setinterleaving by auto then show ?thesis using infinite_super by auto next case (Cons a list) hence dd1:"?D1 \ {(a#l, u)| l u. list setinterleaves ((l, u), A)} \ {(t, a#la)| t la. list setinterleaves ((t, la), A)} \ {(a#l, a#la)|l la. list setinterleaves ((l, la), A)}"(is "?D1 \ ?D2 \ ?D3 \ ?D4") by (rule_tac subsetI, auto split:if_splits) with Cons have aa:"finite {(t,u). list setinterleaves ((t, u), A)}" using 1[THEN spec, of "length list"] by (simp) define f where a2:"f = (\ (t, (u::'a event list)). ((a::'a event)# t, u))" with Cons have "?D2 \ (f ` {(t, u). list setinterleaves ((t, u), A)})" using a2 by auto hence dd2:"finite ?D2" using finite_imageI [of "{(t, u). list setinterleaves ((t, u), A)}" f, OF aa] by (meson rev_finite_subset) define f where a2:"f = (\ ((t::'a event list),u). (t,(a::'a event)#u))" with Cons have "?D3 \ (f ` {(t, u). list setinterleaves ((t, u), A)})" using a2 by auto hence dd3:"finite ?D3" using finite_imageI [of "{(t, u). list setinterleaves ((t, u), A)}" f, OF aa] by (meson rev_finite_subset) define f where a2:"f = (\ (t,u). ((a::'a event)#t,(a::'a event)#u))" with Cons have "?D4 \ (f ` {(t, u). list setinterleaves ((t, u), A)})" using a2 by auto hence dd4:"finite ?D4" using finite_imageI [of "{(t, u). list setinterleaves ((t, u), A)}" f, OF aa] by (meson rev_finite_subset) with dd0 dd1 dd2 dd3 dd4 show ?thesis by (simp add: finite_subset) qed from A AA AAA AAAA show ?case by (simp add: finite_subset) qed lemma inters_fin: "finite{(t, u, r). r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. x = r @ v \ front_tickFree v \ (tickFree r \ v = []))}" (is "finite ?A") proof - have a:"?A\(\r \ {r. r \ x}. {a. \t u. a=(t,u,r)\r setinterleaves((t,u),insert tick(ev ` A))})" (is "?A \ (\r \ {r. r \ x}. ?P r)") using le_list_def by fastforce have b:"\(r::'a event list). finite (?P r)" proof(rule allI) fix r::"'a event list" define f where "f= (\((t::'a event list), (u::'a event list)). (t,u,r))" hence "?P r \ (f ` {(t, u). r setinterleaves ((t, u), insert tick (ev ` A) )})" by auto then show "finite (?P r)" using inters_fin_fund[of r "insert tick (ev ` A)"] finite_imageI[of "{(t, u). r setinterleaves ((t, u), insert tick (ev ` A))}" f] by (meson rev_finite_subset) qed have "{t. \t2. x = t @ t2} = {r. \ra. r @ ra = x}" by blast hence "finite {r. r \ x}" using prefixes_fin[of x, simplified Let_def, THEN conjunct1] by (auto simp add:le_list_def) with a b show ?thesis by (meson finite_UN_I infinite_super) qed lemma limproc_Sync_D: "chain Y\ \ (lim_proc(range Y)\A\ S) = \(lim_proc(range(\i. Y i \A\ S)))" apply(auto simp:Process_eq_spec F_Sync D_Sync T_Sync F_LUB D_LUB T_LUB chain_Sync1) apply blast apply blast using front_tickFree_Nil apply blast using front_tickFree_Nil apply blast proof - fix x assume A: "chain Y" and B: " \xa. \t u r v. front_tickFree v \ (tickFree r \ v = []) \ x = r @ v \ r setinterleaves ((t, u), insert tick (ev ` A)) \ (t \ \ (Y xa) \ u \ \ S \ t \ \ S \ u \ \ (Y xa))" thus "\t u r v. front_tickFree v \ (tickFree r \ v = []) \ x = r @ v \ r setinterleaves ((t, u), insert tick (ev ` A)) \ ((\x. t \ \ (Y x)) \ u \ \ S \ t \ \ S \ (\x. u \ \ (Y x)))" proof(cases "\t u r. r setinterleaves ((t, u), insert tick (ev ` A)) \ (r\x \ ((\x. t \ \ (Y x)) \ u \ \ S) \ (t \ \ S \ (\x. u \ \ (Y x))))") case True from A obtain tryunion where Ctryy: "tryunion={(t, u, r). r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v.(x = r @ v \ front_tickFree v \ (tickFree r \ v = [])))}" by simp from Ctryy inters_fin have tolfin: "finite tryunion" by auto obtain tryunion1 where Ctryy1: "tryunion1 = {n. \(t,u, r) \ tryunion. (t \ \ (Y n) \ u \ \ S \ t \ \ S \ u \ \ (Y n))}" by simp from B Ctryy1 Ctryy have Ctryy3: "\n. n\tryunion1" by blast from Ctryy3 have Ctryy2: "infinite tryunion1" using finite_nat_set_iff_bounded by auto from Ctryy2 Ctryy1 tolfin obtain r2 t2 u2 where Ee: "(t2,u2, r2) \ tryunion \ infinite {n. (t2 \ \ (Y n) \ u2 \ \ S \ t2 \ \ S \ u2 \ \ (Y n))}" by auto from True Ee Ctryy obtain x1 x2 where Ee1: "((t2 \ \ (Y x1)) \ u2 \ \ S) \ (t2 \ \ S \ (u2 \ \ (Y x2)))" apply (simp add: le_list_def) by blast from Ee Ee1 Ctryy obtain x3 where Ee2: "(x1\x3) \ (x2\x3) \ (t2 \ \ (Y x3) \ u2 \ \ S \ t2 \ \ S \ u2 \ \ (Y x3))" apply(simp add: infinite_nat_iff_unbounded_le) proof - assume a1: "r2 setinterleaves ((t2, u2), insert tick (ev ` A)) \ (\v. x = r2 @ v \ front_tickFree v \ (tickFree r2 \ v = [])) \ ((\m. \n\m. t2 \ \ (Y n) \ u2 \ \ S) \ (\m. \n\m. t2 \ \ S \ u2 \ \ (Y n)))" assume a2: "\x3. x1 \ x3 \ x2 \ x3\(t2 \ \ (Y x3) \ u2 \ \ S \ t2 \ \ S \ u2 \ \ (Y x3)) \ thesis" obtain nn :: "nat \ nat" where f3: "\x0. (\v1\x0. t2 \ \ S \ u2 \ \ (Y v1)) = (x0 \ nn x0 \ t2 \ \ S \ u2 \ \ (Y (nn x0)))" by moura obtain nna :: "nat \ nat" where "\x0. (\v1\x0. t2 \ \ (Y v1) \ u2 \ \ S) = (x0 \ nna x0 \ t2 \ \ (Y (nna x0)) \ u2 \ \ S)" by moura then have f4: "(\n. n \ nna n \ t2 \ \ (Y (nna n)) \ u2 \ \ S) \ (\n. n \ nn n \ t2 \ \ S \ u2 \ \ (Y (nn n)))" using f3 a1 by presburger have "(\n. \ n \ nn n \ t2 \ \ S \ u2 \ \ (Y (nn n))) \ (\n\x1. x2 \ n \ (t2 \ \ (Y n) \ u2 \ \ S \ t2 \ \ S \ u2 \ \ (Y n)))" by (metis le_cases order.trans) moreover { assume "\n. \ n \ nn n \ t2 \ \ S \ u2 \ \ (Y (nn n))" then have "\n. n \ nna n \ t2 \ \ (Y (nna n)) \ u2 \ \ S" using f4 by blast then have "\n\x1. x2 \ n \ (t2 \ \ (Y n) \ u2 \ \ S \ t2 \ \ S \ u2 \ \ (Y n))" by (metis (no_types) le_cases order.trans) } ultimately show ?thesis using a2 by blast qed from A have Abb1: "\n1 n2. n1\n2 \ Y n1 \ Y n2" by (simp add: po_class.chain_mono) from A Abb1 have Abb2: "\n1 n2. n1>n2 \ t \ \ (Y n2) \ t \ \ (Y n1)" using le_approx1 less_imp_le_nat by blast from A Abb1 have Abb3: "\n1 n2. n1>n2 \ t \ \ (Y n2) \ t \ \ (Y n1)" by (meson NT_ND le_approx2T less_imp_le_nat) from Abb2 Ee2 have Ab1: "t2 \ \ (Y x1)\t2 \ \ (Y x3)" by (meson Abb1 le_approx1 less_imp_le_nat subsetCE) from Abb3 have Ab2: "u2 \ \ (Y x2)\u2 \ \ (Y x3)" by (meson Abb1 Ee2 NT_ND le_approx2T less_imp_le_nat) from True Ee Ee1 Ee2 Ab1 Ab2 show ?thesis by blast next case False from A B False obtain t u r where Bb1: "r setinterleaves ((t, u), insert tick (ev ` A)) \ r \ x \ ((\x. t \ \ (Y x)) \ u \ \ S \ t \ \ S \ (\x. u \ \ (Y x)))" by auto from B have Bb21: "front_tickFree x" by (metis D_imp_front_tickFree append_Nil2 front_tickFree_append ftf_Sync is_processT2_TR) from B Bb1 have Bb2: "\v. front_tickFree v \ (tickFree r \ v = [])\ x = r @ v" by (metis Bb21 front_tickFree_Nil front_tickFree_mono le_list_def) then show ?thesis using Bb1 by blast qed qed lemma limproc_Sync_F_annex1: " chain Y \ \n. (a, b) \ \ (Y n \A\ S) \ \x. a \ \ (Y x \A\ S) \ \t u X. (\x. (t, X) \ \ (Y x)) \ (\Y. (u, Y) \ \ S \ a setinterleaves((t, u),insert tick (ev ` A)) \ b = (X \ Y) \ insert tick (ev ` A) \ X \ Y)" proof - fix a b assume A: "chain Y" assume B: "\n. (a, b) \ \ (Y n \A\ S)" assume C: "\x. a \ \ (Y x \A\ S)" thus "\t u X. (\x. (t, X) \ \ (Y x)) \ (\Y. (u, Y) \ \ S \ a setinterleaves ((t, u), insert tick (ev ` A)) \ b = (X \ Y) \ insert tick (ev ` A) \ X \ Y)" proof- from B C obtain x1 where D: "a \ \ (Y x1 \A\ S)" by blast from B D obtain t1 u1 X1 Y1 where E: "a setinterleaves ((t1, u1), insert tick (ev ` A)) \ (t1, X1) \ \ (Y x1) \ (u1, Y1) \ \ S \ b = (X1 \ Y1) \ insert tick (ev ` A) \ X1 \ Y1" apply(auto simp: D_Sync F_Sync) by fastforce from B D E have F: "t1 \ \ (Y x1) \ t1 \ \(Y x1)" apply(auto simp: D_Sync) using F_T front_tickFree_Nil apply blast by (simp add: F_T) from F have ee: "\i. (t1, X1) \ \ (Y i)\(u1, Y1)\ \ S\ a setinterleaves ((t1, u1), insert tick (ev ` A))\ b = (X1 \ Y1) \ insert tick (ev ` A) \ X1 \ Y1" using E chain_lemma is_processT8_S le_approx2 local.A by metis with A B C D E F ee show?thesis by blast qed qed lemma limproc_Sync_F_annex2: "chain Y \ \t u r. r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. a = r @ v \ front_tickFree v \ \ tickFree r \ v \ [] \ ((\x. t \ \ (Y x)) \ u \ \ S) \ (t \ \ S \ (\x. u \ \ (Y x)))) \ \x. a \ \(Y x \ A \ S)" apply(auto simp: D_Sync) proof- fix a assume A: "\t u r. r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. a = r @ v \ front_tickFree v \ \ tickFree r \ v \ [] \ ((\x. t \ \ (Y x)) \ u \ \ S) \ (t \ \ S \ (\x. u \ \ (Y x))))" assume B: "chain Y" thus "\x. \t u r. r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. a = r @ v \ front_tickFree v \ \ tickFree r\v\[] \ (t \ \ (Y x) \ u \ \ S) \ (t \ \ S \ u \ \ (Y x)))" proof- from B have D: "((\x. t \ \ (Y x)) \ u \ \ S) \ (t \ \ S \ (\x. u \ \ (Y x))) \ (\x. ((t \ \ (Y x) \ u \ \ S) \ (t \ \ S \ u \ \ (Y x))))" by (meson D_T chain_lemma le_approx1 le_approx2T subsetCE) from A obtain tryunion where Ctryy: "tryunion={(t, u, r). r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. a = r @ v \ front_tickFree v \ (tickFree r \ v = []))}" by simp from Ctryy have tolfin: "finite tryunion" using inters_fin by auto from B have Abb1: "\n1 n2. n1\n2 \ Y n1 \ Y n2" by (simp add: po_class.chain_mono) from A Abb1 have Abb2: "\n1 n2 t. n1>n2\t \ \ (Y n2)\t \ \ (Y n1)" using le_approx1 less_imp_le_nat by blast from A Abb1 have Abb3: "\n1 n2 t. n1>n2\t \ \ (Y n2)\t \ \ (Y n1)" by (meson NT_ND le_approx2T less_imp_le_nat) from Abb2 Abb3 have oneIndexPre: "\t u. ((\x. t \ \ (Y x)) \ u \ \ S) \ (t \ \ S \ (\x. u \ \ (Y x))) \ (\x. (t \ \ (Y x) \ u \ \ S) \ (t \ \ S \ u \ \ (Y x)))" by (meson B D_T ND_F_dir2' chain_lemma le_approx1 subsetCE) from A B oneIndexPre Abb2 Abb3 have oneIndex: "\t u r. r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. a = r @ v \front_tickFree v \ \ tickFree r\v\[] \ ( \x.( (( t \ \ (Y x)) \ u \ \ S) \ (t \ \ S \ (u \ \ (Y x))))))" by blast define proUnion1 where finiUnion: "proUnion1 = {n. \(t,u, r)\tryunion. (((t \ \ (Y n) \ u \ \ S) \ (t \ \ S \ u \ \ (Y n)))\ (\m. ( (t \ \ (Y m)\u \ \ S) \ (t \ \ S\u \ \ (Y m)))\n\m))}" from B have finiSet1: "\r t u. (t,u, r) \ tryunion \ finite {n1.(((t \ \ (Y n1)\u \ \ S) \ (t\\ S\u\\(Y n1)))\(\m1. ((t \ \ (Y m1) \ u \ \ S) \ (t \ \ S \ u \ \ (Y m1)))\n1\m1))}" by (metis (no_types, lifting) infinite_nat_iff_unbounded mem_Collect_eq not_less) from B tolfin finiUnion oneIndex finiSet1 have finiSet: "finite proUnion1" by auto from finiSet obtain proMax where ann: "proMax=Max(proUnion1)" by blast from ann Abb2 have ann1: "\num\proUnion1. \t. t \ \ (Y num)\t \ \ (Y proMax)" by (meson Abb1 Max_ge finiSet le_approx1 subsetCE) from ann Abb3 have ann2: "\num\proUnion1. \t. t \ \ (Y num)\t \ \ (Y proMax)" by (meson Abb1 D_T Max_ge finiSet le_approx2T) from finiUnion have ann3: "\num\proUnion1. \r t u. r setinterleaves((t, u),insert tick (ev ` A)) \ (\v. a = r @ v \front_tickFree v \ \ tickFree r \ v \ [] \(( ((t\\(Y num))\u \ \ S) \ (t \ \ S \ (u \ \ (Y num))))))" by auto obtain proUnion2 where ann0: "proUnion2 ={(t, u, r).\n. (t, u, r) \ tryunion \ ((t \ \ (Y n)\u \ \ S) \ (t \ \ S \ u \ \ (Y n)))}" by auto from ann0 Ctryy finiUnion have ann1: "\t u r. (t, u, r)\ proUnion2 \\num. num\proUnion1 \ ((t \ \ (Y num) \ u \ \ S) \ (t \ \ S \ u \ \ (Y num)))" proof- fix t u r assume a: "(t, u, r)\ proUnion2" define P where PP:"P = (\num. ((t \ \ (Y num) \ u \ \ S) \ (t \ \ S \ u \ \ (Y num))))" thus "\num. num\proUnion1\ P num" proof- from finiUnion obtain minIndexRB where ann1preB: "minIndexRB={n. (P n)}" by auto from B ann1preB obtain minmin where ab: "(minmin::nat) = Inf (minIndexRB)" by auto from ann0 ann1preB PP have ann1preNoEmpty1: "minIndexRB \ {}" using a by blast from ab ann1preNoEmpty1 have ann1pre0:"minmin \ minIndexRB" using Inf_nat_def wellorder_Least_lemma(1) by force have "minmin\ {n. \t u r. (t,u,r)\tryunion \ ((t \ \ (Y n) \ u \ \ S) \ (t \ \ S \ u \ \ (Y n))) \ (\m. (t \ \ (Y m) \ u \ \ S) \ (t \ \ S \ u \ \ (Y m)) \ n \ m)}" apply(rule CollectI,rule_tac x=t in exI,rule_tac x=u in exI,rule_tac x=r in exI,intro conjI) using a ann0 apply blast using PP ann1pre0 ann1preB apply blast using PP ann1pre0 ann1preB apply blast by (simp add: Inf_nat_def PP ab ann1preB wellorder_Least_lemma(2)) then show ?thesis using ann1pre0 ann1preB finiUnion by blast qed qed from ann1 have ann12: "\t u r. \num.(t, u, r) \ proUnion2 \ num\proUnion1 \ ((t \ \ (Y num) \ u \ \ S) \ (t \ \ S \ u \ \ (Y num)))" by auto from ann0 Ctryy have ann2: "\t u r. (t, u, r) \ proUnion2 \ (r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. a = r @ v \ front_tickFree v \ (tickFree r \ v = [])) \ (\nu1. (t \ \ (Y nu1) \ u \ \ S) \ (t \ \ S \ u \ \ (Y nu1))))" by simp have ann15: "\t u r. (r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. a = r @ v \ front_tickFree v \ (tickFree r \ v = [])) \ (\nu1. (t \ \ (Y nu1) \ u \ \ S) \ (t \ \ S \ u \ \ (Y nu1)))) \ (t, u, r) \ proUnion2" using Ctryy ann0 by blast from ann12 ann15 have ann01: "\t u r. ((r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. a = r @ v \ front_tickFree v \ (tickFree r \ v = [])) \ (\nu1. (t \ \ (Y nu1)\u \ \ S) \ (t \ \ S \ u \ \ (Y nu1))))\ (\num\proUnion1. ((t \ \ (Y num) \ u \ \ S) \ (t \ \ S \ u \ \ (Y num)))))" by blast from ann01 have annhelp: "\t u r. r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. a = r @ v \ front_tickFree v \(tickFree r \ v = [])\ (\num\proUnion1.((t \ \ (Y num) \ u \ \ S) \ (t \ \ S \ u \ \ (Y num)))))" by (metis oneIndex) from Abb1 Abb2 have annHelp2: "\nn t u. nn \ proUnion1 \ (t \ \ (Y nn) \ u \ \ S) \ (t \ \ S \ u \ \ (Y nn)) \ (t \ \ (Y proMax) \ u \ \ S) \ (t \ \ S \ u \ \ (Y proMax))" by (metis Abb3 Max_ge ann finiSet le_neq_trans) from annHelp2 annhelp have annHelp1: "\t u r. r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. a = r @ v \ front_tickFree v \ tickFree r \ v = [] \ (t \ \ (Y proMax) \ u \ \ S) \ (t \ \ S \ u \ \ (Y proMax)))" by blast with A B show ?thesis by blast qed qed lemma limproc_Sync_F: "chain Y \ \(lim_proc (range Y)\ A \S) = \(lim_proc (range (\i. Y i \ A \ S)))" apply(auto simp add: Process_eq_spec D_Sync F_Sync F_LUB D_LUB T_LUB chain_Sync1) apply blast apply blast apply blast using front_tickFree_Nil apply blast using front_tickFree_Nil apply blast proof- fix a b assume A1: " \x. (\t u X. (t, X) \ \ (Y x) \ (\Y. (u, Y) \ \ S \ a setinterleaves ((t, u), insert tick (ev ` A)) \ b = (X \ Y) \ insert tick (ev ` A) \ X \ Y)) \ (\t u r v. front_tickFree v \ (tickFree r \ v = []) \ a = r @ v \ r setinterleaves ((t, u), insert tick (ev ` A)) \ (t \ \ (Y x) \ u \ \ S \ t \ \ S \ u \ \ (Y x)))" and B: " \t u r. r setinterleaves ((t, u), insert tick (ev ` A)) \ (\v. a = r @ v \ front_tickFree v \ \ tickFree r \ v \ [] \ ((\x. t \ \ (Y x)) \ u \ \ S) \ (t \ \ S \ (\x. u \ \ (Y x))))" and C: "chain Y" thus "\t u X. (\x. (t, X) \ \ (Y x)) \(\Y. (u, Y) \ \ S \ a setinterleaves ((t, u), insert tick (ev ` A)) \ b = (X \ Y) \ insert tick (ev ` A) \ X \ Y)" proof (cases "\m. a \ \(Y m \ A \ S)") case True then obtain m where "a \ \(Y m \ A \ S)" by blast with A1 have D: "\n. (a, b)\ \ (Y n \ A \ S)" by (auto simp: F_Sync) with A1 B C show ?thesis apply(erule_tac x=m in allE) apply(frule limproc_Sync_F_annex2) apply simp by(simp add: limproc_Sync_F_annex1) next case False with False have E: "\n. a \ \ (Y n \ A \ S)" by blast with C E B show ?thesis apply auto apply(drule limproc_Sync_F_annex2) apply blast by fast qed qed lemma cont_left_Sync : "chain Y \ ((\ i. Y i) \ A \ S) = (\ i. (Y i \ A \ S))" by(simp add: Process_eq_spec chain_Sync1 limproc_Sync_D limproc_Sync_F limproc_is_thelub) lemma cont_right_Sync : "chain Y \ (S \ A \ (\ i. Y i)) = (\ i. (S \ A \ Y i))" by (metis (no_types, lifting) Process_eq_spec cont_left_Sync lub_eq mono_D_Sync mono_F_Sync) lemma Sync_cont[simp]: assumes f:"cont f" and g:"cont g" shows "cont (\x. (f x) \A\ (g x))" proof - have A : "\x. cont g \ cont (\y. y \A\ (g x))" apply (rule contI2, rule monofunI) apply(auto) by (simp add: cont_left_Sync) have B : "\y. cont g \ cont (\x. y \A\ g x)" apply(rule_tac c="(\ x. y \A\ x)" in cont_compose) apply(rule contI2, rule monofunI) apply (simp) apply (simp add: cont_right_Sync) by simp show ?thesis using f g apply(rule_tac f="(\x. (\ f. f \A\ g x))" in cont_apply) by(auto intro:contI2 simp:A B) qed end diff --git a/thys/HOL-CSP/document/root.tex b/thys/HOL-CSP/document/root.tex --- a/thys/HOL-CSP/document/root.tex +++ b/thys/HOL-CSP/document/root.tex @@ -1,64 +1,63 @@ \documentclass[11pt,a4paper]{book} -\usepackage[T1]{fontenc} \usepackage{isabelle,isabellesym} \usepackage{graphicx} \graphicspath {{figures/}} % further packages required for unusual symbols (see also % isabellesym.sty), use only when needed \usepackage{latexsym} \usepackage{amssymb} %for \, \, \, \, \, \, %\, \, \, \, \, %\, \, \ %\usepackage[greek,english]{babel} %option greek for \ %option english (default language) for \, \ -%\usepackage[utf8]{inputenc} +%\usepackage[latin1]{inputenc} %for \, \, \, \, %\, \, \ %\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} \begin{document} \title{HOL-CSP Version 2.0} -\author{ Safouan Taha \and Burkhart Wolff \and Lina Ye } +\author{ Benoit Ballenghien \and Safouan Taha \and Burkhart Wolff \and Lina Ye } \maketitle \tableofcontents % sane default for proof documents \parindent 0pt\parskip 0.5ex % generated text of all theories \input{session} % optional bibliography \bibliographystyle{abbrv} \bibliography{adb-long,root} \end{document} %%% Local Variables: %%% mode: latex %%% TeX-master: t %%% End: