diff --git a/thys/AutoFocus-Stream/AF_Stream_Exec.thy b/thys/AutoFocus-Stream/AF_Stream_Exec.thy --- a/thys/AutoFocus-Stream/AF_Stream_Exec.thy +++ b/thys/AutoFocus-Stream/AF_Stream_Exec.thy @@ -1,3106 +1,3106 @@ (* Title: AF_Stream_Exec.thy Date: Dec 2006 Author: David Trachtenherz *) section \Processing of message streams\ theory AF_Stream_Exec imports AF_Stream "List-Infinite.ListInf_Prefix" "List-Infinite.SetIntervalStep" begin subsection \Executing components with state transition functions\ subsubsection \Basic definitions\ text \ Function type for functions converting an input value to an input port message for a component\ type_synonym ('a, 'in) Port_Input_Value = "'a \ 'in message_af" text \ Function type for functions extracting the output value of a single output port from a component value\ type_synonym ('comp, 'out) Port_Output_Value = "'comp \ 'out message_af" text \ Function type for functions extracting the local state of a component from a component value\ type_synonym ('comp, 'state) Comp_Local_State = "'comp \ 'state" text \ Function type for transition functions computing the component's value after processing an input for a single time unit\ type_synonym ('comp, 'input) Comp_Trans_Fun = "'input \ 'comp \ 'comp" \ \Execute a component for all inputs in the input stream @{typ "'input list"}\ primrec f_Exec_Comp :: "('comp, 'input) Comp_Trans_Fun \ 'input list \ 'comp \ 'comp" where f_Exec_Nil: "f_Exec_Comp trans_fun [] c = c" | f_Exec_Cons: "f_Exec_Comp trans_fun (x#xs) c = f_Exec_Comp trans_fun xs (trans_fun x c)" \ \Execute the component for at most n steps\ definition f_Exec_Comp_N :: "('comp, 'input) Comp_Trans_Fun \ nat \ 'input list \ 'comp \ 'comp" where "f_Exec_Comp_N trans_fun n xs c \ f_Exec_Comp trans_fun (xs \ n) c" \ \Produce the component stream for all inputs in the input stream\ primrec f_Exec_Comp_Stream :: "('comp, 'input) Comp_Trans_Fun \ 'input list \ 'comp \ 'comp list" where f_Exec_Stream_Nil: "f_Exec_Comp_Stream trans_fun [] c = []" | f_Exec_Stream_Cons: "f_Exec_Comp_Stream trans_fun (x # xs) c = (trans_fun x c) # ( f_Exec_Comp_Stream trans_fun xs (trans_fun x c) )" primrec f_Exec_Comp_Stream_Init :: "('comp, 'input) Comp_Trans_Fun \ 'input list \ 'comp \ 'comp list" where f_Exec_Stream_Init_Nil: "f_Exec_Comp_Stream_Init trans_fun [] c = [c]" | f_Exec_Stream_Init_Cons: "f_Exec_Comp_Stream_Init trans_fun (x # xs) c = c # ( f_Exec_Comp_Stream_Init trans_fun xs (trans_fun x c) )" definition i_Exec_Comp_Stream :: "('comp, 'input) Comp_Trans_Fun \ 'input ilist \ 'comp \ 'comp ilist" where "i_Exec_Comp_Stream \ \trans_fun input c n. f_Exec_Comp trans_fun (input \ Suc n) c" definition i_Exec_Comp_Stream_Init :: "('comp, 'input) Comp_Trans_Fun \ 'input ilist \ 'comp \ 'comp ilist" where "i_Exec_Comp_Stream_Init \ \trans_fun input c n. f_Exec_Comp trans_fun (input \ n) c" subsubsection \Basic results\ lemma f_Exec_one: "f_Exec_Comp trans_fun [m] c = trans_fun m c" by simp lemma f_Exec_Stream_length[rule_format, simp]:" \c. length (f_Exec_Comp_Stream trans_fun xs c) = length xs" by (induct xs, simp_all) lemma f_Exec_Stream_empty_conv:" (f_Exec_Comp_Stream trans_fun xs c = []) = (xs = [])" by (simp add: length_0_conv[symmetric] del: length_0_conv) lemma f_Exec_Stream_not_empty_conv:" (f_Exec_Comp_Stream trans_fun xs c \ []) = (xs \ [])" by (simp add: f_Exec_Stream_empty_conv) lemma f_Exec_eq_f_Exec_Stream_last[rule_format]:" \c. f_Exec_Comp trans_fun xs c = last (c # (f_Exec_Comp_Stream trans_fun xs c))" by (induct xs, simp_all) corollary f_Exec_eq_f_Exec_Stream_last2[rule_format]: " xs \ [] \ f_Exec_Comp trans_fun xs c = last (f_Exec_Comp_Stream trans_fun xs c)" by (simp add: f_Exec_eq_f_Exec_Stream_last f_Exec_Stream_empty_conv[symmetric, of xs trans_fun c]) corollary f_Exec_eq_f_Exec_Stream_last_if: " f_Exec_Comp trans_fun xs c = (if xs = [] then c else last (f_Exec_Comp_Stream trans_fun xs c))" by (simp add: f_Exec_eq_f_Exec_Stream_last2) corollary f_Exec_take_eq_last_f_Exec_Stream_take:" \ xs \ []; 0 < n \ \ f_Exec_Comp trans_fun (xs \ n) c = last (f_Exec_Comp_Stream trans_fun (xs \ n) c)" by (simp add: f_Exec_eq_f_Exec_Stream_last2 take_not_empty_conv) corollary f_Exec_N_eq_last_f_Exec_Stream_take:" \ xs \ []; 0 < n \ \ f_Exec_Comp_N trans_fun n xs c = last (f_Exec_Comp_Stream trans_fun (xs \ n) c)" by (simp add: f_Exec_Comp_N_def f_Exec_take_eq_last_f_Exec_Stream_take) lemma f_Exec_Stream_nth: " \n c. n < length xs \ f_Exec_Comp_Stream trans_fun xs c ! n = f_Exec_Comp trans_fun (xs \ Suc n) c" apply (induct xs, simp) apply (simp add: nth_Cons') done lemma f_Exec_Stream_nth2: " n \ length xs \ (c # f_Exec_Comp_Stream trans_fun xs c) ! n = f_Exec_Comp trans_fun (xs \ n) c" by (simp add: nth_Cons' f_Exec_Stream_nth) lemma f_Exec_N_all:" length xs \ n \ f_Exec_Comp_N trans_fun n xs c = f_Exec_Comp trans_fun xs c" by (simp add: f_Exec_Comp_N_def) lemma f_Exec_Stream_append[rule_format]:"\c. f_Exec_Comp_Stream trans_fun (xs @ ys) c = (f_Exec_Comp_Stream trans_fun xs c) @ (f_Exec_Comp_Stream trans_fun ys (f_Exec_Comp trans_fun xs c))" by (induct xs, simp_all) corollary f_Exec_Stream_append_last_Cons[rule_format]:" f_Exec_Comp_Stream trans_fun (xs @ ys) c = (f_Exec_Comp_Stream trans_fun xs c) @ (f_Exec_Comp_Stream trans_fun ys (last (c # (f_Exec_Comp_Stream trans_fun xs c))))" by (simp add: f_Exec_Stream_append f_Exec_eq_f_Exec_Stream_last) corollary f_Exec_Stream_append_last[rule_format]:" xs \ [] \ f_Exec_Comp_Stream trans_fun (xs @ ys) c = (f_Exec_Comp_Stream trans_fun xs c) @ (f_Exec_Comp_Stream trans_fun ys (last (f_Exec_Comp_Stream trans_fun xs c)))" by (simp add: f_Exec_Stream_append_last_Cons f_Exec_Stream_empty_conv) corollary f_Exec_Stream_append_if:" f_Exec_Comp_Stream trans_fun (xs @ ys) c = (f_Exec_Comp_Stream trans_fun xs c) @ (f_Exec_Comp_Stream trans_fun ys ( if xs = [] then c else last (f_Exec_Comp_Stream trans_fun xs c)))" by (simp add: f_Exec_Stream_append f_Exec_eq_f_Exec_Stream_last_if) corollary f_Exec_append:" f_Exec_Comp trans_fun (xs @ ys) c = f_Exec_Comp trans_fun ys (f_Exec_Comp trans_fun xs c)" by (simp add: f_Exec_eq_f_Exec_Stream_last f_Exec_Stream_append_if f_Exec_Stream_empty_conv) corollary f_Exec_Stream_Cons_rev: " xs \ [] \ (trans_fun (hd xs) c) # f_Exec_Comp_Stream trans_fun (tl xs) (trans_fun (hd xs) c) = f_Exec_Comp_Stream trans_fun xs c" by (subst f_Exec_Stream_Cons[symmetric], simp) lemma f_Exec_Stream_snoc: " f_Exec_Comp_Stream trans_fun (xs @ [x]) c = f_Exec_Comp_Stream trans_fun xs c @ [trans_fun x (f_Exec_Comp trans_fun xs c)]" by (simp add: f_Exec_Stream_append) lemma f_Exec_snoc: " f_Exec_Comp trans_fun (xs @ [x]) c = trans_fun x (f_Exec_Comp trans_fun xs c)" by (simp add: f_Exec_append) lemma f_Exec_N_append[rule_format]:" f_Exec_Comp_N trans_fun (a + b) xs c = f_Exec_Comp_N trans_fun b (xs \ a) (f_Exec_Comp_N trans_fun a xs c)" apply (simp add: f_Exec_Comp_N_def f_Exec_append[symmetric]) apply (simp add: take_drop add.commute[of b]) apply (rule subst[of "xs \ (a + b) \ a" "xs \ a" ], simp add: min_eqL) apply (subst append_take_drop_id, simp) done corollary f_Exec_N_Suc[rule_format]:" f_Exec_Comp_N trans_fun (Suc n) xs c = f_Exec_Comp_N trans_fun (Suc 0) (xs \ n) (f_Exec_Comp_N trans_fun n xs c)" by (simp add: f_Exec_N_append[symmetric]) corollary f_Exec_N_Suc2[rule_format]:" n < length xs \ f_Exec_Comp_N trans_fun (Suc n) xs c = trans_fun (xs ! n) (f_Exec_Comp_N trans_fun n xs c)" by (simp add: f_Exec_Comp_N_def take_Suc_conv_app_nth f_Exec_append) theorem f_Exec_Stream_take:" (f_Exec_Comp_Stream trans_fun xs c) \ n = f_Exec_Comp_Stream trans_fun (xs \ n) c" apply (case_tac "length xs \ n", simp) apply (rule subst[OF append_take_drop_id, of _ n xs]) apply (simp add: f_Exec_Stream_append del: append_take_drop_id) done theorem f_Exec_Stream_drop:" (f_Exec_Comp_Stream trans_fun xs c) \ n = f_Exec_Comp_Stream trans_fun (xs \ n) (f_Exec_Comp trans_fun (xs \ n) c)" apply (case_tac "length xs \ n", simp) apply (rule subst[OF append_take_drop_id, of _ n xs]) apply (simp add: f_Exec_Stream_append del: append_take_drop_id) done lemma i_Exec_Stream_nth: " i_Exec_Comp_Stream trans_fun input c n = f_Exec_Comp trans_fun (input \ Suc n) c" by (simp add: i_Exec_Comp_Stream_def) lemma i_Exec_Stream_nth_Suc: " i_Exec_Comp_Stream trans_fun input c (Suc n) = trans_fun (input (Suc n)) (i_Exec_Comp_Stream trans_fun input c n)" by (simp add: i_Exec_Stream_nth i_take_Suc_conv_app_nth f_Exec_append) lemma i_Exec_Stream_nth_Suc_first: " i_Exec_Comp_Stream trans_fun input c (Suc n) = (i_Exec_Comp_Stream trans_fun (input \ Suc 0) (trans_fun (input 0) c) n)" by (simp add: i_Exec_Stream_nth i_take_Suc) lemma f_Exec_Stream_nth_eq_i_Exec_Stream_nth: " n < n' \ f_Exec_Comp_Stream trans_fun (input \ n') c ! n = i_Exec_Comp_Stream trans_fun input c n" by (simp add: f_Exec_Stream_nth i_Exec_Stream_nth min_eqR) lemma i_Exec_Stream_append: " i_Exec_Comp_Stream trans_fun (xs \ input) c = f_Exec_Comp_Stream trans_fun xs c \ i_Exec_Comp_Stream trans_fun input (f_Exec_Comp trans_fun xs c)" by (simp add: ilist_eq_iff i_Exec_Stream_nth f_Exec_Stream_nth f_Exec_append i_append_nth Suc_diff_le) lemma i_Exec_Stream_append_last_Cons: " i_Exec_Comp_Stream trans_fun (xs \ input) c = f_Exec_Comp_Stream trans_fun xs c \ i_Exec_Comp_Stream trans_fun input ( last (c # f_Exec_Comp_Stream trans_fun xs c))" by (simp add: f_Exec_eq_f_Exec_Stream_last i_Exec_Stream_append) lemma i_Exec_Stream_append_last: " xs \ [] \ i_Exec_Comp_Stream trans_fun (xs \ input) c = f_Exec_Comp_Stream trans_fun xs c \ i_Exec_Comp_Stream trans_fun input ( last (f_Exec_Comp_Stream trans_fun xs c))" by (simp add: f_Exec_Stream_empty_conv i_Exec_Stream_append_last_Cons) lemma i_Exec_Stream_append_if: " i_Exec_Comp_Stream trans_fun (xs \ input) c = f_Exec_Comp_Stream trans_fun xs c \ i_Exec_Comp_Stream trans_fun input ( if xs = [] then c else last (f_Exec_Comp_Stream trans_fun xs c))" by (simp add: i_Exec_Stream_append_last) corollary i_Exec_Stream_Cons: " i_Exec_Comp_Stream trans_fun ([x] \ input) c = [trans_fun x c] \ i_Exec_Comp_Stream trans_fun input (trans_fun x c)" by (simp add: i_Exec_Stream_append) corollary i_Exec_Stream_Cons_rev: " [trans_fun (input 0) c] \ i_Exec_Comp_Stream trans_fun (input \ Suc 0) (trans_fun (input 0) c) = i_Exec_Comp_Stream trans_fun input c" apply (insert i_Exec_Stream_append[of trans_fun "[input 0]" "input \ Suc 0" c]) apply (simp add: i_drop_Suc_conv_tl) done theorem i_Exec_Stream_take:" (i_Exec_Comp_Stream trans_fun input c) \ n = f_Exec_Comp_Stream trans_fun (input \ n) c" by (simp add: list_eq_iff f_Exec_Stream_nth i_Exec_Stream_nth min_eqR) theorem i_Exec_Stream_drop:" (i_Exec_Comp_Stream trans_fun input c) \ n = i_Exec_Comp_Stream trans_fun (input \ n) (f_Exec_Comp trans_fun (input \ n) c)" apply (rule subst[OF i_append_i_take_i_drop_id, of _ n input]) apply (simp add: i_Exec_Stream_append i_drop_def del: i_append_i_take_i_drop_id) done lemma f_Exec_Stream_expand_aggregate_map_take: " f_aggregate (map f (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c)) k ag \ n = f_aggregate (map f (f_Exec_Comp_Stream trans_fun ((xs \ n) \\<^sub>f k) c)) k ag" by (simp add: f_aggregate_take_mult[symmetric] take_map f_Exec_Stream_take f_expand_take_mult) corollary f_Exec_Stream_expand_aggregate_take: " f_aggregate (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c) k ag \ n = f_aggregate (f_Exec_Comp_Stream trans_fun ((xs \ n) \\<^sub>f k) c) k ag" by (insert f_Exec_Stream_expand_aggregate_map_take[of n id trans_fun xs k c ag], simp add: map_id) lemma i_Exec_Stream_expand_aggregate_map_take: " 0 < k \ i_aggregate (f \ (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c)) k ag \ n = f_aggregate (map f (f_Exec_Comp_Stream trans_fun ((input \ n) \\<^sub>f k) c)) k ag" by (simp add: i_aggregate_i_take_mult[symmetric] i_Exec_Stream_take i_expand_i_take_mult) corollary i_Exec_Stream_expand_aggregate_take: " 0 < k \ i_aggregate (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c) k ag \ n = f_aggregate (f_Exec_Comp_Stream trans_fun ((input \ n) \\<^sub>f k) c) k ag" by (drule i_Exec_Stream_expand_aggregate_map_take[of k n id trans_fun input c ag], simp add: map_id) lemma f_Exec_Stream_expand_aggregate_map_drop: " f_aggregate (map f (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c)) k ag \ n = f_aggregate (map f (f_Exec_Comp_Stream trans_fun ((xs \ n) \\<^sub>f k) ( f_Exec_Comp trans_fun ((xs \ n) \\<^sub>f k) c))) k ag" by (simp add: f_aggregate_drop_mult[symmetric] drop_map f_Exec_Stream_drop f_expand_take_mult f_expand_drop_mult) corollary f_Exec_Stream_expand_aggregate_drop: " f_aggregate (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c) k ag \ n = f_aggregate (f_Exec_Comp_Stream trans_fun ((xs \ n) \\<^sub>f k) ( f_Exec_Comp trans_fun ((xs \ n) \\<^sub>f k) c)) k ag" by (insert f_Exec_Stream_expand_aggregate_map_drop[of n id trans_fun xs k c ag], simp add: map_id) lemma i_Exec_Stream_expand_aggregate_map_drop: " 0 < k \ i_aggregate (f \ (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c)) k ag \ n = i_aggregate (f \ (i_Exec_Comp_Stream trans_fun ((input \ n) \\<^sub>i k) ( f_Exec_Comp trans_fun ((input \ n) \\<^sub>f k) c))) k ag" by (simp add: i_aggregate_i_drop_mult[symmetric] i_Exec_Stream_drop i_expand_i_take_mult i_expand_i_drop_mult) corollary i_Exec_Stream_expand_aggregate_drop: " 0 < k \ i_aggregate (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c) k ag \ n = i_aggregate (i_Exec_Comp_Stream trans_fun ((input \ n) \\<^sub>i k) ( f_Exec_Comp trans_fun ((input \ n) \\<^sub>f k) c)) k ag" by (drule i_Exec_Stream_expand_aggregate_map_drop[of k n id trans_fun input c ag], simp) lemma f_Exec_Stream_expand_aggregate_map_nth_eq_i_nth: " \ 0 < k; n < n' \ \ f_aggregate (map f (f_Exec_Comp_Stream trans_fun (input \ n' \\<^sub>f k) c)) k ag ! n = i_aggregate (f \ (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c)) k ag n" apply (simp add: f_aggregate_nth i_aggregate_nth f_Exec_Stream_take f_Exec_Stream_drop i_Exec_Stream_take i_Exec_Stream_drop drop_map take_map) apply (simp add: f_expand_take_mod i_expand_i_take_mod f_expand_drop_mod i_expand_i_drop_mod i_drop_i_take_1 drop_take_1 min_eqR) done corollary f_Exec_Stream_expand_aggregate_map_nth_eq_i_nth': " 0 < k \ f_aggregate (map f (f_Exec_Comp_Stream trans_fun (input \ Suc n \\<^sub>f k) c)) k ag ! n = i_aggregate (f \ (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c)) k ag n" by (simp add: f_Exec_Stream_expand_aggregate_map_nth_eq_i_nth) corollary f_Exec_Stream_expand_aggregate_nth_eq_i_nth: " \ 0 < k; n < n' \ \ f_aggregate (f_Exec_Comp_Stream trans_fun (input \ n' \\<^sub>f k) c) k ag ! n = i_aggregate (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c) k ag n" by (drule f_Exec_Stream_expand_aggregate_map_nth_eq_i_nth[where f=id], simp_all add: map_id) corollary f_Exec_Stream_expand_aggregate_nth_eq_i_nth': " 0 < k \ f_aggregate (f_Exec_Comp_Stream trans_fun (input \ Suc n \\<^sub>f k) c) k ag ! n = i_aggregate (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c) k ag n" by (simp add: f_Exec_Stream_expand_aggregate_nth_eq_i_nth) lemma f_Exec_Stream_expand_shrink_last_map_nth_eq_f_Exec_Comp: " \ 0 < k; n < length xs \ \ map f (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c) \
\<^bsub>fl\<^esub> k ! n = f (f_Exec_Comp trans_fun ((xs \ Suc n) \\<^sub>f k) c)" apply (simp add: f_shrink_last_map f_shrink_last_length f_shrink_last_nth) apply (subgoal_tac "n * k + k - Suc 0 < length xs * k") prefer 2 apply (drule Suc_leI[of n]) apply (drule mult_le_mono1[of _ _ k], simp) apply (simp add: f_Exec_Stream_nth add.commute[of k] f_expand_take_mult[symmetric]) done corollary f_Exec_Stream_expand_shrink_last_nth_eq_f_Exec_Comp: " \ 0 < k; n < length xs \ \ f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c \
\<^bsub>fl\<^esub> k ! n = f_Exec_Comp trans_fun ((xs \ Suc n) \\<^sub>f k) c" by (drule f_Exec_Stream_expand_shrink_last_map_nth_eq_f_Exec_Comp[where f=id], simp_all add: map_id) lemma f_Exec_Stream_expand_aggregate_map_nth: " \ 0 < k; n < length xs \ \ f_aggregate (map f (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c)) k ag ! n = ag (map f (f_Exec_Comp_Stream trans_fun (xs ! n # \\<^bsup>k - Suc 0\<^esup>) (f_Exec_Comp trans_fun (xs \ n \\<^sub>f k) c)))" apply (simp add: f_aggregate_nth take_map drop_map) apply (simp add: take_map drop_map f_Exec_Stream_drop f_Exec_Stream_take f_expand_take_mod f_expand_drop_mod drop_take_1) done corollary f_Exec_Stream_expand_aggregate_nth: " \ 0 < k; n < length xs \ \ f_aggregate (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c) k ag ! n = ag (f_Exec_Comp_Stream trans_fun (xs ! n # \\<^bsup>k - Suc 0\<^esup>) (f_Exec_Comp trans_fun (xs \ n \\<^sub>f k) c))" by (drule f_Exec_Stream_expand_aggregate_map_nth[where f=id], simp_all add: map_id) corollary f_Exec_Stream_expand_shrink_map_nth: " \ 0 < k; n < length xs \ \ (map f (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c)) \
\<^sub>f k ! n = last_message (map f (f_Exec_Comp_Stream trans_fun (xs ! n # \\<^bsup>k - Suc 0\<^esup>) (f_Exec_Comp trans_fun (xs \ n \\<^sub>f k) c)))" by (simp add: f_shrink_def f_Exec_Stream_expand_aggregate_map_nth) lemma i_Exec_Stream_expand_aggregate_map_nth: " 0 < k \ i_aggregate (f \ (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c)) k ag n = ag (map f (f_Exec_Comp_Stream trans_fun (input n # \\<^bsup>k - Suc 0\<^esup>) (f_Exec_Comp trans_fun ((input \ n) \\<^sub>f k) c)))" by (simp add: i_aggregate_nth i_Exec_Stream_drop i_Exec_Stream_take i_expand_i_take_mod i_expand_i_drop_mod i_drop_i_take_1) corollary i_Exec_Stream_expand_aggregate_nth: " 0 < k \ i_aggregate (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c) k ag n = ag (f_Exec_Comp_Stream trans_fun (input n # \\<^bsup>k - Suc 0\<^esup>) (f_Exec_Comp trans_fun ((input \ n) \\<^sub>f k) c))" by (drule i_Exec_Stream_expand_aggregate_map_nth[where f=id], simp add: map_id) corollary i_Exec_Stream_expand_shrink_map_nth: " 0 < k \ ((f \ (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c)) \
\<^sub>i k) n = last_message (map f (f_Exec_Comp_Stream trans_fun (input n # \\<^bsup>k - Suc 0\<^esup>) (f_Exec_Comp trans_fun (input \ n \\<^sub>f k) c)))" by (simp add: i_shrink_def i_Exec_Stream_expand_aggregate_map_nth) lemma f_Exec_Stream_expand_snoc: " \ 0 < k; n < length xs \ \ f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c \ (n * k) \ k = f_Exec_Comp_Stream trans_fun (xs ! n # \\<^bsup>k - Suc 0\<^esup>) (f_Exec_Comp trans_fun (xs \ n \\<^sub>f k) c)" by (simp add: f_Exec_Stream_drop f_Exec_Stream_take f_expand_take_mod f_expand_drop_mod drop_take_1) lemma f_Exec_Stream_expand_map_aggregate_append: " f_aggregate (map f (f_Exec_Comp_Stream trans_fun ((xs @ ys) \\<^sub>f k) c)) k ag = f_aggregate (map f (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c)) k ag @ f_aggregate (map f (f_Exec_Comp_Stream trans_fun (ys \\<^sub>f k) ( f_Exec_Comp trans_fun (xs \\<^sub>f k) c))) k ag" by (simp add: f_Exec_Stream_append f_aggregate_append_mod) lemma i_Exec_Stream_expand_map_aggregate_append: " i_aggregate (f \ (i_Exec_Comp_Stream trans_fun ((xs \ input) \\<^sub>i k) c)) k ag = f_aggregate (map f (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c)) k ag \ i_aggregate (f \ (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) ( f_Exec_Comp trans_fun (xs \\<^sub>f k) c))) k ag" by (simp add: i_expand_i_append i_Exec_Stream_append i_aggregate_i_append_mod) lemma f_Exec_Stream_expand_map_aggregate_Cons: " 0 < k \ f_aggregate (map f (f_Exec_Comp_Stream trans_fun ((x # xs) \\<^sub>f k) c)) k ag = ag (map f (f_Exec_Comp_Stream trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c)) # f_aggregate (map f (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) ( f_Exec_Comp trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c))) k ag" apply (subst append_eq_Cons[of x xs, symmetric]) apply (subst f_Exec_Stream_expand_map_aggregate_append) apply (simp add: f_aggregate_one) done lemma f_Exec_Stream_expand_map_aggregate_snoc: " 0 < k \ f_aggregate (map f (f_Exec_Comp_Stream trans_fun ((xs @ [x]) \\<^sub>f k) c)) k ag = f_aggregate (map f (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c)) k ag @ [ag (map f (f_Exec_Comp_Stream trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) ( f_Exec_Comp trans_fun (xs \\<^sub>f k) c)))]" apply (subst f_Exec_Stream_expand_map_aggregate_append) apply (simp add: f_aggregate_one) done lemma i_Exec_Stream_expand_map_aggregate_Cons: " 0 < k \ i_aggregate (f \ (i_Exec_Comp_Stream trans_fun (([x] \ input) \\<^sub>i k) c)) k ag = [ag (map f (f_Exec_Comp_Stream trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c))] \ i_aggregate (f \ (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) ( f_Exec_Comp trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c))) k ag" apply (subst i_Exec_Stream_expand_map_aggregate_append) apply (simp add: f_aggregate_one) done lemma f_Exec_N_eq_f_Exec_Stream_nth:" n \ length xs \ f_Exec_Comp_N trans_fun n xs c = (c # f_Exec_Comp_Stream trans_fun xs c) ! n" by (simp add: f_Exec_Comp_N_def f_Exec_Stream_nth2) theorem f_Exec_Stream_causal: " xs \ n = ys \ n \ (f_Exec_Comp_Stream trans_fun xs c) \ n = (f_Exec_Comp_Stream trans_fun ys c) \ n" by (simp add: f_Exec_Stream_take) theorem i_Exec_Stream_causal: " input1 \ n = input2 \ n \ (i_Exec_Comp_Stream trans_fun input1 c) \ n = (i_Exec_Comp_Stream trans_fun input2 c) \ n" by (simp add: i_Exec_Stream_take) text \Results for \f_Exec_Comp_Stream_Init\\ text \ \f_Exec_Comp_Stream_Init\ computes the execution stream of a component with the initial value of the component at the beginning of the result stream.\ lemma f_Exec_Stream_Init_length[rule_format, simp]:" \c. length (f_Exec_Comp_Stream_Init trans_fun xs c) = Suc (length xs)" by (induct xs, simp_all) lemma f_Exec_Stream_Init_not_empty:" (f_Exec_Comp_Stream_Init trans_fun xs c \ [])" by (simp add: length_0_conv[symmetric] del: length_0_conv) lemma f_Exec_eq_f_Exec_Stream_Init_last[rule_format]:" \c. f_Exec_Comp trans_fun xs c = last (f_Exec_Comp_Stream_Init trans_fun xs c)" by (induct xs, simp_all add: f_Exec_Stream_Init_not_empty) lemma f_Exec_Stream_Init_eq_f_Exec_Stream_Cons[rule_format]: " \c. f_Exec_Comp_Stream_Init trans_fun xs c = c # f_Exec_Comp_Stream trans_fun xs c" by (induct xs, simp_all) corollary f_Exec_Stream_Init_eq_f_Exec_Stream_Cons_output: " output_fun c = \ \ map output_fun (f_Exec_Comp_Stream_Init trans_fun xs c) = \ # map output_fun (f_Exec_Comp_Stream trans_fun xs c)" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons) corollary f_Exec_Stream_Init_tl_eq_f_Exec_Stream: " tl (f_Exec_Comp_Stream_Init trans_fun xs c) = f_Exec_Comp_Stream trans_fun xs c" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons) lemma f_Exec_N_eq_last_f_Exec_Stream_Init_take:" f_Exec_Comp_N trans_fun n xs c = last (f_Exec_Comp_Stream_Init trans_fun (xs \ n) c)" by (simp add: f_Exec_Comp_N_def f_Exec_eq_f_Exec_Stream_Init_last) lemma f_Exec_Stream_Init_nth: " n \ length xs \ f_Exec_Comp_Stream_Init trans_fun xs c ! n = f_Exec_Comp trans_fun (xs \ n) c" apply (subst f_Exec_Stream_Init_eq_f_Exec_Stream_Cons) apply (case_tac n, simp) apply (simp add: f_Exec_Stream_nth) done lemma f_Exec_Stream_Init_nth_0: "f_Exec_Comp_Stream_Init trans_fun xs c ! 0 = c" by (simp add: f_Exec_Stream_Init_nth) lemma f_Exec_Stream_Init_hd: "hd (f_Exec_Comp_Stream_Init trans_fun xs c) = c" by (simp add: hd_conv_nth f_Exec_Stream_Init_not_empty f_Exec_Stream_Init_nth_0) lemma f_Exec_Stream_Init_nth_Suc_eq_f_Exec_Stream_nth: " f_Exec_Comp_Stream_Init trans_fun xs c ! (Suc n) = f_Exec_Comp_Stream trans_fun xs c ! n" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons) lemma f_Exec_Stream_Init_append:" f_Exec_Comp_Stream_Init trans_fun (xs @ ys) c = (f_Exec_Comp_Stream_Init trans_fun xs c) @ tl (f_Exec_Comp_Stream_Init trans_fun ys (f_Exec_Comp trans_fun xs c))" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons f_Exec_Stream_append) corollary f_Exec_Stream_Init_append_last:" f_Exec_Comp_Stream_Init trans_fun (xs @ ys) c = (f_Exec_Comp_Stream_Init trans_fun xs c) @ tl (f_Exec_Comp_Stream_Init trans_fun ys (last (f_Exec_Comp_Stream_Init trans_fun xs c)))" by (simp add: f_Exec_Stream_Init_append f_Exec_eq_f_Exec_Stream_Init_last) lemma f_Exec_Stream_Init_f_Exec_Stream_append:" f_Exec_Comp_Stream_Init trans_fun (xs @ ys) c = (f_Exec_Comp_Stream_Init trans_fun xs c) @ (f_Exec_Comp_Stream trans_fun ys (f_Exec_Comp trans_fun xs c))" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons f_Exec_Stream_append) lemma f_Exec_Stream_Init_take:" (f_Exec_Comp_Stream_Init trans_fun xs c) \ Suc n = f_Exec_Comp_Stream_Init trans_fun (xs \ n) c" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons f_Exec_Stream_take) lemma f_Exec_Stream_Init_drop:" n \ length xs \ (f_Exec_Comp_Stream_Init trans_fun xs c) \ n = f_Exec_Comp_Stream_Init trans_fun (xs \ n) (f_Exec_Comp trans_fun (xs \ n) c)" apply (case_tac n, simp) apply (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons f_Exec_Stream_drop) apply (simp add: take_Suc_conv_app_nth f_Exec_append Cons_nth_drop_Suc[symmetric]) done lemma f_Exec_Stream_Init_drop_geq_not_valid:" length xs \ n \ (f_Exec_Comp_Stream_Init trans_fun xs c) \ Suc n \ f_Exec_Comp_Stream_Init trans_fun arbitrary_input arbitrary_comp" by (simp add: f_Exec_Stream_Init_not_empty[symmetric]) lemma i_Exec_Stream_Init_nth: " i_Exec_Comp_Stream_Init trans_fun input c n = f_Exec_Comp trans_fun (input \ n) c" by (simp add: i_Exec_Comp_Stream_Init_def) lemma i_Exec_Stream_Init_nth_0: " i_Exec_Comp_Stream_Init trans_fun input c 0 = c" by (simp add: i_Exec_Stream_Init_nth) lemma i_Exec_Stream_Init_nth_Suc_eq_i_Exec_Stream_nth: " i_Exec_Comp_Stream_Init trans_fun input c (Suc n) = i_Exec_Comp_Stream trans_fun input c n" by (simp add: i_Exec_Stream_Init_nth i_Exec_Stream_nth) lemma i_Exec_Stream_Init_eq_i_Exec_Stream_Cons: " i_Exec_Comp_Stream_Init trans_fun input c = [c] \ i_Exec_Comp_Stream trans_fun input c" by (simp add: ilist_eq_iff i_Exec_Stream_Init_nth i_append_nth i_Exec_Stream_nth) corollary i_Exec_Stream_Init_eq_i_Exec_Stream_Cons_output: " output_fun c = \ \ output_fun \ i_Exec_Comp_Stream_Init trans_fun input c = [\] \ (output_fun \ i_Exec_Comp_Stream trans_fun input c)" by (simp add: i_Exec_Stream_Init_eq_i_Exec_Stream_Cons) lemma i_Exec_Stream_Init_append:" i_Exec_Comp_Stream_Init trans_fun (input1 \ input2) c = (f_Exec_Comp_Stream_Init trans_fun input1 c) \ ((i_Exec_Comp_Stream_Init trans_fun input2 (f_Exec_Comp trans_fun input1 c)) \ Suc 0)" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons i_Exec_Stream_Init_eq_i_Exec_Stream_Cons i_Exec_Stream_append) corollary i_Exec_Stream_Init_append_last:" i_Exec_Comp_Stream_Init trans_fun (input1 \ input2) c = (f_Exec_Comp_Stream_Init trans_fun input1 c) \ ((i_Exec_Comp_Stream_Init trans_fun input2 (last (f_Exec_Comp_Stream_Init trans_fun input1 c))) \ Suc 0)" by (simp add: i_Exec_Stream_Init_append f_Exec_eq_f_Exec_Stream_Init_last) lemma i_Exec_Stream_Init_i_Exec_Stream_append:" i_Exec_Comp_Stream_Init trans_fun (input1 \ input2) c = (f_Exec_Comp_Stream_Init trans_fun input1 c) \ (i_Exec_Comp_Stream trans_fun input2 (f_Exec_Comp trans_fun input1 c))" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons i_Exec_Stream_Init_eq_i_Exec_Stream_Cons i_Exec_Stream_append) lemma i_Exec_Stream_Init_take:" (i_Exec_Comp_Stream_Init trans_fun input c) \ Suc n = f_Exec_Comp_Stream_Init trans_fun (input \ n) c" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons i_Exec_Stream_Init_eq_i_Exec_Stream_Cons i_Exec_Stream_take) lemma i_Exec_Stream_Init_drop:" (i_Exec_Comp_Stream_Init trans_fun input c) \ n = i_Exec_Comp_Stream_Init trans_fun (input \ n) (f_Exec_Comp trans_fun (input \ n) c)" apply (case_tac n, simp) apply (simp add: i_Exec_Stream_Init_eq_i_Exec_Stream_Cons i_Exec_Stream_drop) apply (simp add: ilist_eq_iff i_take_Suc_conv_app_nth f_Exec_append i_Exec_Stream_nth i_append_nth i_take_first i_take_drop_eq_map) apply (simp add: upt_conv_Cons) done theorem f_Exec_Stream_Init_strictly_causal: " xs \ n = ys \ n \ (f_Exec_Comp_Stream_Init trans_fun xs c) \ Suc n = (f_Exec_Comp_Stream_Init trans_fun ys c) \ Suc n" by (simp add: f_Exec_Stream_Init_take) theorem i_Exec_Stream_Init_strictly_causal: " input1 \ n = input2 \ n \ (i_Exec_Comp_Stream_Init trans_fun input1 c) \ Suc n = (i_Exec_Comp_Stream_Init trans_fun input2 c) \ Suc n" by (simp add: i_Exec_Stream_Init_take) theorem f_Exec_N_eq_f_Exec_Stream_Init_nth:" n \ length xs \ f_Exec_Comp_N trans_fun n xs c = f_Exec_Comp_Stream_Init trans_fun xs c ! n" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons f_Exec_N_eq_f_Exec_Stream_nth) text \Basic results for previous element functions\ text \ The functions \list_Previous\ and \ilist_Previous\ return the previous element of the list relatively to the specified position @{term n} or the initial element if @{term n} is 0,\ definition list_Previous :: "'value list \ 'value \ nat \ 'value" where "list_Previous xs init n \ case n of 0 \ init | Suc n' \ xs ! n'" definition ilist_Previous :: "'value ilist \ 'value \ nat \ 'value" where "ilist_Previous f init n \ case n of 0 \ init | Suc n' \ f n'" abbreviation "list_Previous'" :: "'value list \ 'value \ nat \ 'value" - ( "_\<^bsup>\' _\<^esup> _" [1000, 10, 100] 100) + ( "_\<^bsup>\'' _\<^esup> _" [1000, 10, 100] 100) where "xs\<^bsup>\' init\<^esup> n \ list_Previous xs init n" abbreviation "ilist_Previous'" :: "'value ilist \ 'value \ nat \ 'value" ( "_\<^bsup>\ _\<^esup> _" [1000, 10, 100] 100) where "f\<^bsup>\ init\<^esup> n \ ilist_Previous f init n" lemma list_Previous_nth: "xs\<^bsup>\' init\<^esup> n = (case n of 0 \ init | Suc n' \ xs ! n')" by (simp add: list_Previous_def) lemma ilist_Previous_nth: "f\<^bsup>\ init\<^esup> n = (case n of 0 \ init | Suc n' \ f n')" by (simp add: ilist_Previous_def) lemma list_Previous_nth_if: "xs\<^bsup>\' init\<^esup> n = (if n = 0 then init else xs ! (n - Suc 0))" by (case_tac n, simp_all add: list_Previous_nth) lemma ilist_Previous_nth_if: "f\<^bsup>\ init\<^esup> n = (if n = 0 then init else f (n - Suc 0))" by (case_tac n, simp_all add: ilist_Previous_nth) lemma list_Previous_Cons: "xs\<^bsup>\' init\<^esup> n = (init # xs) ! n" by (case_tac n, simp_all add: list_Previous_nth) lemma ilist_Previous_Cons: "f\<^bsup>\ init\<^esup> n = ([init] \ f) n" by (case_tac n, simp_all add: ilist_Previous_nth) lemma list_Previous_0: "xs\<^bsup>\' init\<^esup> 0 = init" by (simp add: list_Previous_def) lemma ilist_Previous_0: "f\<^bsup>\ init\<^esup> 0 = init" by (simp add: ilist_Previous_def) lemma list_Previous_gr0: "0 < n \ xs\<^bsup>\' init\<^esup> n = xs ! (n - Suc 0)" by (case_tac n, simp_all add: list_Previous_nth) lemma ilist_Previous_gr0: "0 < n \ f\<^bsup>\ init\<^esup> n = f (n - Suc 0)" by (case_tac n, simp_all add: ilist_Previous_nth) lemma list_Previous_Suc: "xs\<^bsup>\' init\<^esup> (Suc n) = xs ! n" by (simp add: list_Previous_def) lemma ilist_Previous_Suc: "f\<^bsup>\ init\<^esup> (Suc n) = f n" by (simp add: ilist_Previous_def) lemma f_Exec_Stream_Previous_f_Exec_Stream_Init: " f_Exec_Comp_Stream_Init trans_fun xs c ! n = (f_Exec_Comp_Stream trans_fun xs c)\<^bsup>\' c\<^esup> n" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons list_Previous_Cons) lemma i_Exec_Stream_Previous_i_Exec_Stream_Init: " i_Exec_Comp_Stream_Init trans_fun input c n = (i_Exec_Comp_Stream trans_fun input c)\<^bsup>\ c\<^esup> n" by (simp add: i_Exec_Stream_Init_eq_i_Exec_Stream_Cons ilist_Previous_Cons) lemma f_Exec_Stream_hd: " 0 < length xs \ hd (f_Exec_Comp_Stream trans_fun xs c) = trans_fun (hd xs) c" by (case_tac xs, simp+) lemma f_Exec_Stream_nth_0: " 0 < length xs \ (f_Exec_Comp_Stream trans_fun xs c) ! 0= trans_fun (xs ! 0) c" by (case_tac xs, simp+) text \ The calculation of the n-th result stream element from the previous result stream element and the current input stream element.\ lemma f_Exec_Stream_nth_gr0_calc: " \ n < length xs; 0 < n \ \ f_Exec_Comp_Stream trans_fun xs c ! n = trans_fun (xs ! n) (f_Exec_Comp_Stream trans_fun xs c ! (n - 1))" by (simp add: f_Exec_Stream_nth take_Suc_conv_app_nth f_Exec_append) lemma f_Exec_Stream_nth_calc_Previous: " n < length xs \ f_Exec_Comp_Stream trans_fun xs c ! n = trans_fun (xs ! n) ((f_Exec_Comp_Stream trans_fun xs c)\<^bsup>\' c\<^esup> n)" apply (case_tac n) apply (simp add: list_Previous_0 f_Exec_Stream_nth_0) apply (simp add: list_Previous_def f_Exec_Stream_nth_gr0_calc) done lemma i_Exec_Stream_nth_0: " (i_Exec_Comp_Stream trans_fun input c) 0 = trans_fun (input 0) c" by (simp add: i_Exec_Stream_nth i_take_first) lemma i_Exec_Stream_nth_gr0_calc: " 0 < n \ (i_Exec_Comp_Stream trans_fun input c) n = trans_fun (input n) ((i_Exec_Comp_Stream trans_fun input c) (n - 1))" by (simp add: i_Exec_Stream_nth i_take_Suc_conv_app_nth f_Exec_append) text \ The component state (and thus its output) at time point @{term "n"} is computed from the previous state (the state at time @{term "n-1"} for @{term "n > 0"} or the initial state for @{term "n = 0"}) and the input at time @{term "n"}.\ lemma i_Exec_Stream_nth_calc_Previous: " i_Exec_Comp_Stream trans_fun input c n = trans_fun (input n) ((i_Exec_Comp_Stream trans_fun input c)\<^bsup>\ c\<^esup> n)" by (simp add: i_Exec_Stream_nth ilist_Previous_nth_if i_take_first i_take_Suc_conv_app_nth f_Exec_append) lemma f_Exec_Stream_Init_nth_Suc_calc: " n < length xs \ f_Exec_Comp_Stream_Init trans_fun xs c ! Suc n = trans_fun (xs ! n) (f_Exec_Comp_Stream_Init trans_fun xs c ! n)" by (simp add: f_Exec_Stream_Init_eq_f_Exec_Stream_Cons f_Exec_Stream_nth nth_Cons' length_greater_0_conv[THEN iffD1, OF gr_implies_gr0] take_Suc_conv_app_nth f_Exec_append) lemma f_Exec_Stream_Init_nth_Plus1_calc: " n < length xs \ f_Exec_Comp_Stream_Init trans_fun xs c ! (n + 1)= trans_fun (xs ! n) (f_Exec_Comp_Stream_Init trans_fun xs c ! n)" by (simp add: f_Exec_Stream_Init_nth_Suc_calc) lemma f_Exec_Stream_Init_nth_gr0_calc: " \ n \ length xs; 0 < n \ \ f_Exec_Comp_Stream_Init trans_fun xs c ! n = trans_fun (xs ! (n - 1)) (f_Exec_Comp_Stream_Init trans_fun xs c ! (n - 1))" by (clarsimp simp: gr0_conv_Suc f_Exec_Stream_Init_nth_Suc_calc) text \ At the beginning, the component state (and thus its output) for the execution stream with initial state is represented by the initial state, contrary to the @{term "i_Exec_Comp_Stream"} that does not contain the initial state.\ text \ The component state (and thus its output) at time point @{term "n + 1"} for the execution stream with initial state is computed from the previous state (the state at time @{term "n"}) and the previous input (input at time @{term "n"}), contrary to the @{term "i_Exec_Comp_Stream"}, where each state at time @{term "n"} represents the resulting state after processing the input at time @{term "n"}.\ lemma i_Exec_Stream_Init_nth_Suc_calc: " i_Exec_Comp_Stream_Init trans_fun input c (Suc n) = trans_fun (input n) (i_Exec_Comp_Stream_Init trans_fun input c n)" by (simp add: i_Exec_Stream_Init_nth i_take_Suc_conv_app_nth f_Exec_append) lemma i_Exec_Stream_Init_nth_Plus1_calc: " i_Exec_Comp_Stream_Init trans_fun input c (n + 1) = trans_fun (input n) (i_Exec_Comp_Stream_Init trans_fun input c n)" by (simp add: i_Exec_Stream_Init_nth_Suc_calc) lemma i_Exec_Stream_Init_nth_gr0_calc: " 0 < n \ i_Exec_Comp_Stream_Init trans_fun input c n = trans_fun (input (n - 1)) (i_Exec_Comp_Stream_Init trans_fun input c (n - 1))" by (clarsimp simp: gr0_conv_Suc i_Exec_Stream_Init_nth_Suc_calc) text \Correlation between Pre/Post-Conditions for \f_Exec_Comp_Stream\ and \f_Exec_Comp_Stream_Init\\ lemma f_Exec_Stream_Pre_Post1: " \ n < length xs; c_n = (f_Exec_Comp_Stream trans_fun xs c)\<^bsup>\' c\<^esup> n; x_n = xs ! n \ \ (P1 x_n \ P2 c_n \ Q (f_Exec_Comp_Stream trans_fun xs c ! n)) = (P1 x_n \ P2 c_n \ Q (trans_fun x_n c_n))" by (simp add: f_Exec_Stream_nth_calc_Previous) text \Direct relation between input and result after transition\ lemma f_Exec_Stream_Pre_Post2: " \ n < length xs; c_n = (f_Exec_Comp_Stream trans_fun xs c)\<^bsup>\' c\<^esup> n; x_n = xs ! n \ \ (P c_n \ Q (xs ! n) (f_Exec_Comp_Stream trans_fun xs c ! n)) = (P c_n \ Q x_n (trans_fun x_n c_n))" by (simp add: f_Exec_Stream_nth_calc_Previous) lemma f_Exec_Stream_Pre_Post2_Suc: " \ Suc n < length xs; c_n = f_Exec_Comp_Stream trans_fun xs c ! n; x_n1 = xs ! Suc n \ \ (P c_n \ Q (xs ! Suc n) (f_Exec_Comp_Stream trans_fun xs c ! Suc n)) = (P c_n \ Q x_n1 (trans_fun x_n1 c_n))" by (simp add: f_Exec_Stream_nth_gr0_calc) lemma f_Exec_Stream_Init_Pre_Post1: " \ n < length xs; c_n = f_Exec_Comp_Stream_Init trans_fun xs c ! n; x_n = xs ! n \ \ (P1 x_n \ P2 c_n \ Q (f_Exec_Comp_Stream_Init trans_fun xs c ! Suc n)) = (P1 x_n \ P2 c_n \ Q (trans_fun x_n c_n))" by (simp add: f_Exec_Stream_Init_nth_Suc_calc) text \Direct relation between input and state before transition\ lemma f_Exec_Stream_Init_Pre_Post2: " \ n < length xs; c_n = f_Exec_Comp_Stream_Init trans_fun xs c ! n; x_n = xs ! n \ \ (P (xs ! n) (f_Exec_Comp_Stream_Init trans_fun xs c ! n) \ Q (f_Exec_Comp_Stream_Init trans_fun xs c ! Suc n)) = (P x_n c_n \ Q (trans_fun x_n c_n))" by (simp add: f_Exec_Stream_Init_nth_Suc_calc) lemma i_Exec_Stream_Pre_Post1: " \ c_n = (i_Exec_Comp_Stream trans_fun input c)\<^bsup>\ c\<^esup> n; x_n = input n \ \ (P1 x_n \ P2 c_n \ Q (i_Exec_Comp_Stream trans_fun input c n)) = (P1 x_n \ P2 c_n \ Q (trans_fun x_n c_n))" by (simp add: i_Exec_Stream_nth_calc_Previous) text \Direct relation between input and result after transition\ lemma i_Exec_Stream_Pre_Post2: " \ c_n = (i_Exec_Comp_Stream trans_fun input c)\<^bsup>\ c\<^esup> n; x_n = input n \ \ (P c_n \ Q (input n) (i_Exec_Comp_Stream trans_fun input c n)) = (P c_n \ Q x_n (trans_fun x_n c_n))" by (simp add: i_Exec_Stream_nth_calc_Previous) lemma i_Exec_Stream_Pre_Post2_Suc: " \ c_n = i_Exec_Comp_Stream trans_fun input c n; x_n1 = input (Suc n) \ \ (P c_n \ Q (input (Suc n)) (i_Exec_Comp_Stream trans_fun input c (Suc n))) = (P c_n \ Q x_n1 (trans_fun x_n1 c_n))" by (simp add: i_Exec_Stream_nth_gr0_calc) lemma i_Exec_Stream_Init_Pre_Post1: " \ c_n = i_Exec_Comp_Stream_Init trans_fun input c n; x_n = input n \ \ (P1 x_n \ P2 c_n \ Q (i_Exec_Comp_Stream_Init trans_fun input c (Suc n))) = (P1 x_n \ P2 c_n \ Q (trans_fun x_n c_n))" by (simp add: i_Exec_Stream_Init_nth_Suc_calc) text \Direct relation between input and state before transition\ lemma i_Exec_Stream_Init_Pre_Post2: " \ c_n = i_Exec_Comp_Stream_Init trans_fun input c n; x_n = input n \ \ (P (input n) (i_Exec_Comp_Stream_Init trans_fun input c n) \ Q (i_Exec_Comp_Stream_Init trans_fun input c (Suc n))) = (P x_n c_n \ Q (trans_fun x_n c_n))" by (simp add: i_Exec_Stream_Init_nth_Suc_calc) text \Basic results for stream prefices\ lemma f_Exec_Stream_prefix: " prefix xs ys \ prefix (f_Exec_Comp_Stream trans_fun xs c) (f_Exec_Comp_Stream trans_fun ys c)" by (clarsimp simp: prefix_def f_Exec_Stream_append) lemma i_Exec_Stream_prefix: " xs \ input \ f_Exec_Comp_Stream trans_fun xs c \ i_Exec_Comp_Stream trans_fun input c" by (simp add: iprefix_eq_iprefix_take i_Exec_Stream_take) lemma f_Exec_N_prefix: " \ n \ length xs; prefix xs ys \ \ f_Exec_Comp_N trans_fun n xs c = f_Exec_Comp_N trans_fun n ys c" by (simp add: f_Exec_Comp_N_def prefix_imp_take_eq) theorem f_Exec_Stream_prefix_causal[rule_format]:" n \ length (xs \ ys) \ f_Exec_Comp_Stream trans_fun xs c \ n = f_Exec_Comp_Stream trans_fun ys c \ n" by (rule f_Exec_Stream_causal, rule inf_prefix_take_correct) lemma f_Exec_Stream_Init_prefix:" prefix xs ys \ prefix (f_Exec_Comp_Stream_Init trans_fun xs c) (f_Exec_Comp_Stream_Init trans_fun ys c)" by (clarsimp simp: prefix_def f_Exec_Stream_Init_append) lemma i_Exec_Stream_Init_prefix: " xs \ input \ f_Exec_Comp_Stream_Init trans_fun xs c \ i_Exec_Comp_Stream_Init trans_fun input c" by (simp add: iprefix_eq_iprefix_take i_Exec_Stream_Init_take) theorem f_Exec_Stream_Init_prefix_strictly_causal[rule_format]:" n \ length (xs \ ys) \ f_Exec_Comp_Stream_Init trans_fun xs c \ Suc n = f_Exec_Comp_Stream_Init trans_fun ys c \ Suc n" by (rule f_Exec_Stream_Init_strictly_causal, rule inf_prefix_take_correct) text \ A predicate indicating whether a component is deterministically dependent on the local state extracted by the the given local state function.\ definition Deterministic_Trans_Fun :: "('comp, 'input) Comp_Trans_Fun \ ('comp, 'state) Comp_Local_State \ bool" where "Deterministic_Trans_Fun trans_fun localState \ \c1 c2 x. localState c1 = localState c2 \ trans_fun x c1 = trans_fun x c2" lemma Deterministic_f_Exec: " \ Deterministic_Trans_Fun trans_fun localState; localState c1 = localState c2; xs \ [] \ \ f_Exec_Comp trans_fun xs c1 = f_Exec_Comp trans_fun xs c2" apply (unfold Deterministic_Trans_Fun_def) apply (case_tac xs, simp) apply (rename_tac y ys) apply (drule_tac x=c1 in spec) apply (drule_tac x=c2 in spec) apply simp done lemma Deterministic_f_Exec_Stream: " \ Deterministic_Trans_Fun trans_fun localState; localState c1 = localState c2 \ \ f_Exec_Comp_Stream trans_fun xs c1 = f_Exec_Comp_Stream trans_fun xs c2" apply (clarsimp simp: list_eq_iff f_Exec_Stream_nth) apply (rule Deterministic_f_Exec) apply (simp add: length_greater_0_conv[THEN iffD1, OF gr_implies_gr0])+ done lemma Deterministic_i_Exec_Stream: " \ Deterministic_Trans_Fun trans_fun localState; localState c1 = localState c2 \ \ i_Exec_Comp_Stream trans_fun input c1 = i_Exec_Comp_Stream trans_fun input c2" apply (clarsimp simp: ilist_eq_iff i_Exec_Stream_nth) apply (rule Deterministic_f_Exec) apply simp+ done subsubsection \Connected streams\ text \ A predicate indicating for two message streams, that the ports, they correspond to, are connected. The predicate implies strict causality.\ definition f_Streams_Connected :: "'a fstream_af \ 'a fstream_af \ bool" where "f_Streams_Connected outS inS \ inS = \ # outS" definition i_Streams_Connected :: "'a istream_af \ 'a istream_af \ bool" where "i_Streams_Connected outS inS \ inS = [\] \ outS" lemmas Streams_Connected_defs = f_Streams_Connected_def i_Streams_Connected_def lemma f_Streams_Connected_imp_not_empty: "f_Streams_Connected outS inS \ inS \ []" by (simp add: f_Streams_Connected_def) lemma f_Streams_Connected_nth_conv: " f_Streams_Connected outS inS = (length inS = Suc (length outS) \ (\i \ | Suc k \ outS ! k)))" by (simp add: f_Streams_Connected_def list_eq_iff nth_Cons) lemma f_Streams_Connected_nth_conv_if: " f_Streams_Connected outS inS = (length inS = Suc (length outS) \ (\i else outS ! (i - Suc 0))))" apply (subst f_Streams_Connected_nth_conv) apply (rule conj_cong, simp) apply (rule all_imp_eqI, simp) apply (rename_tac i, case_tac i, simp+) done lemma i_Streams_Connected_nth_conv: " i_Streams_Connected outS inS = (\i. inS i = (case i of 0 \ \ | Suc k \ outS k))" by (simp add: i_Streams_Connected_def ilist_eq_iff i_append_nth_Cons) lemma i_Streams_Connected_nth_conv_if: " i_Streams_Connected outS inS = (\i. inS i = (if i = 0 then \ else outS (i - Suc 0)))" apply (subst i_Streams_Connected_nth_conv) apply (rule all_eqI) apply (rename_tac i, case_tac i, simp+) done lemma f_Exec_Stream_Init_eq_output_channel: " \ output_fun c = \; f_Streams_Connected (map output_fun (f_Exec_Comp_Stream trans_fun xs c)) channel \ \ map output_fun (f_Exec_Comp_Stream_Init trans_fun xs c) = channel" by (simp add: f_Streams_Connected_def f_Exec_Stream_Init_eq_f_Exec_Stream_Cons) lemma i_Exec_Stream_Init_eq_output_channel: " \ output_fun c = \; i_Streams_Connected (output_fun \ (i_Exec_Comp_Stream trans_fun input c)) channel \ \ output_fun \ (i_Exec_Comp_Stream_Init trans_fun input c) = channel" by (simp add: i_Streams_Connected_def i_Exec_Stream_Init_eq_i_Exec_Stream_Cons) lemma f_Exec_Stream_output_causal: " \ xs \ n = ys \ n; output1 = map output_fun (f_Exec_Comp_Stream trans_fun xs c); output2 = map output_fun (f_Exec_Comp_Stream trans_fun ys c) \ \ output1 \ n = output2 \ n" by (simp add: take_map f_Exec_Stream_causal[of n xs]) lemma f_Exec_Stream_Init_output_strictly_causal: " \ xs \ n = ys \ n; output1 = map output_fun (f_Exec_Comp_Stream_Init trans_fun xs c); output2 = map output_fun (f_Exec_Comp_Stream_Init trans_fun ys c) \ \ output1 \ Suc n = output2 \ Suc n" by (simp add: take_map f_Exec_Stream_Init_strictly_causal[of n xs]) lemma i_Exec_Stream_output_causal: " \ input1 \ n = input2 \ n; output1 = output_fun \ i_Exec_Comp_Stream trans_fun input1 c; output2 = output_fun \ i_Exec_Comp_Stream trans_fun input2 c \ \ output1 \ n = output2 \ n" by (simp add: i_Exec_Stream_causal[of n input1]) lemma i_Exec_Stream_Init_output_strictly_causal: " \ input1 \ n = input2 \ n; output1 = output_fun \ i_Exec_Comp_Stream_Init trans_fun input1 c; output2 = output_fun \ i_Exec_Comp_Stream_Init trans_fun input2 c \ \ output1 \ Suc n = output2 \ Suc n" by (simp add: i_Exec_Stream_Init_strictly_causal[of n input1]) lemma f_Exec_Stream_Connected_strictly_causal: " \ xs \ n = ys \ n; f_Streams_Connected (map output_fun (f_Exec_Comp_Stream trans_fun xs c)) channel1; f_Streams_Connected (map output_fun (f_Exec_Comp_Stream trans_fun ys c)) channel2 \ \ channel1 \ Suc n = channel2 \ Suc n" by (simp add: f_Streams_Connected_def take_map f_Exec_Stream_take) lemma i_Exec_Stream_Connected_strictly_causal: " \ input1 \ n = input2 \ n; i_Streams_Connected (portOutput \ (i_Exec_Comp_Stream trans_fun input1 c)) channel1; i_Streams_Connected (portOutput \ (i_Exec_Comp_Stream trans_fun input2 c)) channel2 \ \ channel1 \ Suc n = channel2 \ Suc n" by (simp add: i_Streams_Connected_def i_take_Suc_Cons i_Exec_Stream_take) text \ A predicate for the semantics with initial state in result stream indicating for two message streams that the ports, they correspond to, are connected.\ definition f_Streams_Connected_Init :: "'a fstream_af \ 'a fstream_af \ bool" where "f_Streams_Connected_Init outS inS \ inS = outS" definition i_Streams_Connected_Init :: "'a istream_af \ 'a istream_af \ bool" where "i_Streams_Connected_Init outS inS \ inS = outS" lemmas Streams_Connected_Init_defs = f_Streams_Connected_Init_def i_Streams_Connected_Init_def lemma f_Streams_Connected_Init_nth_conv: " f_Streams_Connected_Init outS inS = (length inS = length outS \ (\ii. inS i = outS i)" by (simp add: i_Streams_Connected_Init_def ilist_eq_iff) lemma f_Exec_Stream_Init_eq_output_channel2: " \ output_fun c = \; f_Streams_Connected_Init (map output_fun (f_Exec_Comp_Stream_Init trans_fun xs c)) channel \ \ map output_fun (f_Exec_Comp_Stream_Init trans_fun xs c) = channel" by (simp add: f_Streams_Connected_Init_def f_Exec_Stream_Init_eq_f_Exec_Stream_Cons) lemma i_Exec_Stream_Init_eq_output_channel2: " \ output_fun c = \; i_Streams_Connected_Init (output_fun \ (i_Exec_Comp_Stream_Init trans_fun input c)) channel \ \ output_fun \ (i_Exec_Comp_Stream_Init trans_fun input c) = channel" by (simp add: i_Streams_Connected_Init_def i_Exec_Stream_Init_eq_i_Exec_Stream_Cons) lemma f_Exec_Stream_Connected_Init_strictly_causal: " \ xs \ n = ys \ n; f_Streams_Connected_Init (map output_fun (f_Exec_Comp_Stream_Init trans_fun xs c)) channel1; f_Streams_Connected_Init (map output_fun (f_Exec_Comp_Stream_Init trans_fun ys c)) channel2 \ \ channel1 \ Suc n = channel2 \ Suc n" by (simp add: f_Streams_Connected_Init_def f_Exec_Stream_Init_eq_f_Exec_Stream_Cons take_map f_Exec_Stream_take) lemma i_Exec_Stream_Connected_Init_strictly_causal: " \ input1 \ n = input2 \ n; i_Streams_Connected_Init (portOutput \ (i_Exec_Comp_Stream_Init trans_fun input1 c)) channel1; i_Streams_Connected_Init (portOutput \ (i_Exec_Comp_Stream_Init trans_fun input2 c)) channel2 \ \ channel1 \ Suc n = channel2 \ Suc n" by (simp add: i_Streams_Connected_Init_def i_Exec_Stream_Init_eq_i_Exec_Stream_Cons i_Exec_Stream_take) subsubsection \Additional auxiliary results\ text \The following lemma shows that, if the system state is different at some time points with respect to a certain predicate @{term P}, then there exists a defined time point between these two, where the state change has taken place\ lemma f_State_Change_exists_set: " \ n1 \ n2; n1 \ I; n2 \ I; \ P (f_Exec_Comp trans_fun (input \ n1) c); P (f_Exec_Comp trans_fun (input \ n2) c) \ \ \n\I. n1 \ n \ n < n2 \ \ P (f_Exec_Comp trans_fun (input \ n) c) \ P (f_Exec_Comp trans_fun (input \ (inext n I)) c)" by (rule inext_predicate_change_exists) lemma f_State_Change_exists: " \ n1 \ n2; \ P (f_Exec_Comp trans_fun (input \ n1) c); P (f_Exec_Comp trans_fun (input \ n2) c) \ \ \n\n1. n < n2 \ \ P (f_Exec_Comp trans_fun (input \ n) c) \ P (f_Exec_Comp trans_fun (input \ (Suc n)) c)" by (rule nat_Suc_predicate_change_exists) lemma i_State_Change_exists_set: " \ n1 \ n2; n1 \ I; n2 \ I; \ P (i_Exec_Comp_Stream trans_fun input c n1); P (i_Exec_Comp_Stream trans_fun input c n2) \ \ \n\I. n1 \ n \ n < n2 \ \ P (i_Exec_Comp_Stream trans_fun input c n) \ P (i_Exec_Comp_Stream trans_fun input c (inext n I))" by (rule inext_predicate_change_exists) lemma i_State_Change_exists: " \ n1 \ n2; \ P (i_Exec_Comp_Stream trans_fun input c n1); P (i_Exec_Comp_Stream trans_fun input c n2) \ \ \n\n1. n < n2 \ \ P (i_Exec_Comp_Stream trans_fun input c n) \ P (i_Exec_Comp_Stream trans_fun input c (Suc n))" by (rule nat_Suc_predicate_change_exists) lemma i_State_Change_Init_exists_set: " \ n1 \ n2; n1 \ I; n2 \ I; \ P (i_Exec_Comp_Stream_Init trans_fun input c n1); P (i_Exec_Comp_Stream_Init trans_fun input c n2) \ \ \n\I. n1 \ n \ n < n2 \ \ P (i_Exec_Comp_Stream_Init trans_fun input c n) \ P (i_Exec_Comp_Stream_Init trans_fun input c (inext n I))" by (rule inext_predicate_change_exists) lemma i_State_Change_Init_exists: " \ n1 \ n2; \ P (i_Exec_Comp_Stream_Init trans_fun input c n1); P (i_Exec_Comp_Stream_Init trans_fun input c n2) \ \ \n\n1. n < n2 \ \ P (i_Exec_Comp_Stream_Init trans_fun input c n) \ P (i_Exec_Comp_Stream_Init trans_fun input c (Suc n))" by (rule nat_Suc_predicate_change_exists) subsection \Components with accelerated execution\ text \ This section deals with variable execution speed components. A component accelerated by a (clocking) factor @{term k} processes streams expanded by factor @{term k} and its output streams are compressed by factor @{term k}.\ subsubsection \Equivalence relation for executions\ text \ A predicate indicating for two components together with transition functions and a given equivalence predicate for their local states, that the components exhibit equivalent observable behaviour after expanding input streams and shrinking output streams by a constant factor, given that their local states are equivalent with respect to the specified equivalence relations.\ definition Equiv_Exec :: " 'input \ ('state1 \ 'state2 \ bool) \ \ \Equivalence predicate for local states\ ('comp1, 'state1) Comp_Local_State \ ('comp2, 'state2) Comp_Local_State \ ('input, 'input1) Port_Input_Value \ \ \Input adaptor for first component\ ('input, 'input2) Port_Input_Value \ \ \Input adaptor for second component\ ('comp1, 'output) Port_Output_Value \ ('comp2, 'output) Port_Output_Value \ ('comp1, 'input1 message_af) Comp_Trans_Fun \ ('comp2, 'input2 message_af) Comp_Trans_Fun \ nat \ nat \ 'comp1 \ 'comp2 \ bool" where "Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 \ equiv_states (localState1 c1) (localState2 c2) \ ( last_message (map output_fun1 ( f_Exec_Comp_Stream trans_fun1 (input_fun1 m # \\<^bsup>k1 - Suc 0\<^esup>) c1)) = last_message (map output_fun2 ( f_Exec_Comp_Stream trans_fun2 (input_fun2 m # \\<^bsup>k2 - Suc 0\<^esup>) c2)) \ equiv_states (localState1 (f_Exec_Comp trans_fun1 (input_fun1 m # \\<^bsup>k1 - Suc 0\<^esup>) c1)) (localState2 (f_Exec_Comp trans_fun2 (input_fun2 m # \\<^bsup>k2 - Suc 0\<^esup>) c2)))" text \ Predicate indicating for two components together with transition functions and a given equivalence predicate for their local states, that the equivalence predicate is stable with respect to component execution, i.e., it determines the equivalence of components' local states both for the initial states and after the components have processed an arbitrary input. The restricting version @{term "Equiv_Exec_stable_set"} guarantees stability only for inputs from a given restriction set, the not-restricting version guarantees stability for all inputs.\ definition Equiv_Exec_stable_set :: " 'input set \ ('state1 \ 'state2 \ bool) \ \ \Equivalence predicate for local states\ ('comp1, 'state1) Comp_Local_State \ ('comp2, 'state2) Comp_Local_State \ ('input, 'input1) Port_Input_Value \ \ \Input adaptor for first component\ ('input, 'input2) Port_Input_Value \ \ \Input adaptor for second component\ ('comp1, 'output) Port_Output_Value \ ('comp2, 'output) Port_Output_Value \ ('comp1, 'input1 message_af) Comp_Trans_Fun \ ('comp2, 'input2 message_af) Comp_Trans_Fun \ nat \ nat \ 'comp1 \ 'comp2 \ bool" where "Equiv_Exec_stable_set A equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 \ \input m. set input \ A \ m \ A \ Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input \\<^sub>f k2) c2)" definition Equiv_Exec_stable :: " ('state1 \ 'state2 \ bool) \ \ \Equivalence predicate for local states\ ('comp1, 'state1) Comp_Local_State \ ('comp2, 'state2) Comp_Local_State \ ('input, 'input1) Port_Input_Value \ \ \Input adaptor for first component\ ('input, 'input2) Port_Input_Value \ \ \Input adaptor for second component\ ('comp1, 'output) Port_Output_Value \ ('comp2, 'output) Port_Output_Value \ ('comp1, 'input1 message_af) Comp_Trans_Fun \ ('comp2, 'input2 message_af) Comp_Trans_Fun \ nat \ nat \ 'comp1 \ 'comp2 \ bool" where "Equiv_Exec_stable equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 \ \input m. Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input \\<^sub>f k2) c2)" lemma Equiv_Exec_equiv_statesI: " \ equiv_states (localState1 c1) (localState2 c2); Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 \ \ equiv_states (localState1 (f_Exec_Comp trans_fun1 (input_fun1 m # \\<^bsup>k1 - Suc 0\<^esup>) c1)) (localState2 (f_Exec_Comp trans_fun2 (input_fun2 m # \\<^bsup>k2 - Suc 0\<^esup>) c2))" by (simp add: Equiv_Exec_def) lemma Equiv_Exec_output_eqI: " \ equiv_states (localState1 c1) (localState2 c2); Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 \ \ last_message (map output_fun1 ( f_Exec_Comp_Stream trans_fun1 (input_fun1 m # \\<^bsup>k1 - Suc 0\<^esup>) c1)) = last_message (map output_fun2 ( f_Exec_Comp_Stream trans_fun2 (input_fun2 m # \\<^bsup>k2 - Suc 0\<^esup>) c2))" by (simp add: Equiv_Exec_def) lemma Equiv_Exec_equiv_statesI': " \ equiv_states (localState1 c1) (localState2 c2); Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 \ \ equiv_states (localState1 (f_Exec_Comp trans_fun1 NoMsg\<^bsup>k1 - Suc 0\<^esup> (trans_fun1 (input_fun1 m) c1))) (localState2 (f_Exec_Comp trans_fun2 NoMsg\<^bsup>k2 - Suc 0\<^esup> (trans_fun2 (input_fun2 m) c2)))" by (simp add: Equiv_Exec_def) lemma Equiv_Exec_le1: " \ k1 \ Suc 0; k2 \ Suc 0; equiv_states (localState1 c1) (localState2 c2); Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 \ \ output_fun1 (trans_fun1 (input_fun1 m) c1) = output_fun2 (trans_fun2 (input_fun2 m) c2) \ equiv_states (localState1 (trans_fun1 (input_fun1 m) c1)) (localState2 (trans_fun2 (input_fun2 m) c2))" by (simp add: Equiv_Exec_def) lemma Equiv_Exec_stable_set_UNIV: " Equiv_Exec_stable_set UNIV equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 = Equiv_Exec_stable equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2" by (simp add: Equiv_Exec_stable_set_def Equiv_Exec_stable_def) lemma Equiv_Exec_stable_setI: " \ Equiv_Exec_stable_set A equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2; set input \ A; m \ A \ \ Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input \\<^sub>f k2) c2)" by (simp add: Equiv_Exec_stable_set_def) lemma Equiv_Exec_stableI: " Equiv_Exec_stable equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 \ Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input \\<^sub>f k2) c2)" by (simp add: Equiv_Exec_stable_def) text \Reflexitity, symmetry and transitivity results for @{term "Equiv_Exec"}\ lemma Equiv_Exec_refl: " \ \c. equiv_states (localState c) (localState c) \ \ Equiv_Exec m equiv_states localState localState input_fun input_fun output_fun output_fun trans_fun trans_fun k k c c" by (simp add: Equiv_Exec_def) lemma Equiv_Exec_sym[rule_format]: " \ \c1 c2. equiv_states (localState1 c1) (localState2 c2) = equiv_states (localState2 c2) (localState1 c1) \ \ Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 = Equiv_Exec m equiv_states localState2 localState1 input_fun2 input_fun1 output_fun2 output_fun1 trans_fun2 trans_fun1 k2 k1 c2 c1" by (fastforce simp: Equiv_Exec_def) lemma Equiv_Exec_sym2: " \ equiv_states_sym = (\s1 s2. equiv_states s2 s1) \ \ Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 = Equiv_Exec m equiv_states_sym localState2 localState1 input_fun2 input_fun1 output_fun2 output_fun1 trans_fun2 trans_fun1 k2 k1 c2 c1" by (fastforce simp: Equiv_Exec_def) lemma Equiv_Exec_sym2_ex: " \equiv_states_sym. Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 = Equiv_Exec m equiv_states_sym localState2 localState1 input_fun2 input_fun1 output_fun2 output_fun1 trans_fun2 trans_fun1 k2 k1 c2 c1" by (rule exI, rule Equiv_Exec_sym2, simp) lemma Equiv_Exec_trans: " \ Equiv_Exec m equiv_states12 localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2; Equiv_Exec m equiv_states23 localState2 localState3 input_fun2 input_fun3 output_fun2 output_fun3 trans_fun2 trans_fun3 k2 k3 c2 c3; equiv_states13 = (\s1 s3. ( if s1 = localState1 c1 \ s3 = localState3 c3 then equiv_states12 s1 (localState2 c2) \ equiv_states23 (localState2 c2) s3 else equiv_states12 s1 ( localState2 (f_Exec_Comp trans_fun2 (input_fun2 m # \\<^bsup>k2 - Suc 0\<^esup>) c2))) \ equiv_states23 ( localState2 (f_Exec_Comp trans_fun2 (input_fun2 m # \\<^bsup>k2 - Suc 0\<^esup>) c2)) s3) \ \ Equiv_Exec m equiv_states13 localState1 localState3 input_fun1 input_fun3 output_fun1 output_fun3 trans_fun1 trans_fun3 k1 k3 c1 c3" by (fastforce simp: Equiv_Exec_def) lemma Equiv_Exec_trans_ex: " \ Equiv_Exec m equiv_states12 localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2; Equiv_Exec m equiv_states23 localState2 localState3 input_fun2 input_fun3 output_fun2 output_fun3 trans_fun2 trans_fun3 k2 k3 c2 c3 \ \ \equiv_states13. Equiv_Exec m equiv_states13 localState1 localState3 input_fun1 input_fun3 output_fun1 output_fun3 trans_fun1 trans_fun3 k1 k3 c1 c3" by (blast intro: Equiv_Exec_trans) text \A predicate indicating for a given local state extraction function and a given transition function, that components, whose states are equal with regard to the local state extraction function, are transformed into equal componenents, when the transition function is applied with the same input.\ definition Exec_Equal_State :: "('comp, 'state) Comp_Local_State \ ('comp, 'input message_af) Comp_Trans_Fun \ bool" where "Exec_Equal_State localState trans_fun \ \c1 c2 m. localState c1 = localState c2 \ trans_fun m c1 = trans_fun m c2" lemma Exec_Equal_StateD: " \ Exec_Equal_State localState trans_fun; localState c1 = localState c2 \ \ trans_fun m c1 = trans_fun m c2" by (unfold Exec_Equal_State_def, blast) lemma Exec_Equal_StateD': " Exec_Equal_State localState trans_fun \ \c1 c2 m. localState c1 = localState c2 \ trans_fun m c1 = trans_fun m c2" by (unfold Exec_Equal_State_def, blast) lemma Exec_Equal_StateI: " (\c1 c2 m. localState c1 = localState c2 \ trans_fun m c1 = trans_fun m c2) \ Exec_Equal_State localState trans_fun" by (unfold Exec_Equal_State_def, blast) lemma f_Exec_Equal_State: "\c1 c2. \ Exec_Equal_State localState trans_fun; localState c1 = localState c2; xs \ [] \ \ f_Exec_Comp trans_fun xs c1 = f_Exec_Comp trans_fun xs c2" apply (induct xs, simp) apply (case_tac "xs = []") apply simp apply (rule Exec_Equal_StateD, assumption+) apply (drule_tac x="trans_fun a c1" in meta_spec) apply (drule_tac x="trans_fun a c2" in meta_spec) apply (drule_tac ?c1.0=c1 and ?c2.0=c2 and m=a in Exec_Equal_StateD, assumption) apply simp done lemma f_Exec_Stream_Equal_State: " \ Exec_Equal_State localState trans_fun; localState c1 = localState c2 \ \ f_Exec_Comp_Stream trans_fun xs c1 = f_Exec_Comp_Stream trans_fun xs c2" apply (clarsimp simp: list_eq_iff f_Exec_Stream_nth) apply (drule gr_implies_gr0) apply (rule f_Exec_Equal_State) apply simp+ done lemma i_Exec_Stream_Equal_State: " \ Exec_Equal_State localState trans_fun; localState c1 = localState c2 \ \ i_Exec_Comp_Stream trans_fun input c1 = i_Exec_Comp_Stream trans_fun input c2" apply (clarsimp simp: ilist_eq_iff i_Exec_Stream_nth) apply (rule f_Exec_Equal_State) apply simp+ done subsubsection \Idle states\ definition State_Idle :: "('comp, 'state) Comp_Local_State \ ('comp \ 'output message_af) \ ('comp, 'input message_af) Comp_Trans_Fun \ 'state \ bool" where "State_Idle localState output_fun trans_fun state \ \c. localState c = state \ localState (trans_fun \ c) = state \ output_fun (trans_fun \ c) = \" lemma State_IdleD: " \ State_Idle localState output_fun trans_fun state; localState c = state \ \ localState (trans_fun \ c) = state \ output_fun (trans_fun \ c) = \" by (unfold State_Idle_def, blast) lemma State_IdleD': " State_Idle localState output_fun trans_fun state \ \c. localState c = state \ localState (trans_fun \ c) = state \ output_fun (trans_fun \ c) = \" by (unfold State_Idle_def, blast) lemma State_IdleI: " \ \c. localState c = state \ localState (trans_fun \ c) = state \ output_fun (trans_fun \ c) = \ \ \ State_Idle localState output_fun trans_fun state" by (unfold State_Idle_def, blast) lemma State_Idle_step[rule_format]: " \ State_Idle localState output_fun trans_fun (localState c) \ \ State_Idle localState output_fun trans_fun (localState (trans_fun \ c))" apply (frule State_IdleD[OF _ refl], erule conjE) apply (rule State_IdleI, rename_tac c0) apply (drule_tac c=c0 in State_IdleD) apply simp+ done lemma f_Exec_State_Idle_replicate_NoMsg_state[rule_format]: " \c. State_Idle localState output_fun trans_fun (localState c) \ localState (f_Exec_Comp trans_fun \\<^bsup>n\<^esup> c) = localState c" apply (induct n, simp) apply (frule State_Idle_step) apply (drule_tac c=c in State_IdleD, rule refl) apply simp done lemma f_Exec_State_Idle_replicate_NoMsg_gr0_output[rule_format]: "\c. \ State_Idle localState output_fun trans_fun (localState c); 0 < n \ \ output_fun (f_Exec_Comp trans_fun \\<^bsup>n\<^esup> c) = \" apply (induct n, simp) apply (case_tac "n = 0") apply simp apply (rule State_IdleD[THEN conjunct2], assumption, simp) apply (drule State_Idle_step) apply simp done lemma f_Exec_State_Idle_replicate_NoMsg_output[rule_format]: " \ State_Idle localState output_fun trans_fun (localState c); output_fun c = \ \ \ output_fun (f_Exec_Comp trans_fun \\<^bsup>n\<^esup> c) = \" apply (case_tac "n = 0", simp) apply (simp add: f_Exec_State_Idle_replicate_NoMsg_gr0_output) done lemma f_Exec_Stream_State_Idle_replicate_NoMsg_output[rule_format]: " \ State_Idle localState output_fun trans_fun (localState c) \ \ map output_fun (f_Exec_Comp_Stream trans_fun \\<^bsup>n\<^esup> c) = \\<^bsup>n\<^esup>" by (simp add: list_eq_iff f_Exec_Stream_nth min_eqL f_Exec_State_Idle_replicate_NoMsg_gr0_output del: replicate.simps) corollary f_Exec_State_Idle_append_replicate_NoMsg_state: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun xs c)) \ \ localState (f_Exec_Comp trans_fun (xs @ \\<^bsup>n\<^esup>) c) = localState (f_Exec_Comp trans_fun xs c)" by (simp add: f_Exec_append f_Exec_State_Idle_replicate_NoMsg_state) corollary f_Exec_State_Idle_append_replicate_NoMsg_ge_state: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun (xs @ \\<^bsup>m\<^esup>) c)); m \ n \ \ localState (f_Exec_Comp trans_fun (xs @ \\<^bsup>n\<^esup>) c) = localState (f_Exec_Comp trans_fun (xs @ \\<^bsup>m\<^esup>) c)" apply (rule_tac t=n and s="m + (n - m)" in subst, simp) apply (simp only: replicate_add append_assoc[symmetric]) apply (rule f_Exec_State_Idle_append_replicate_NoMsg_state, simp) done corollary f_Exec_State_Idle_replicate_NoMsg_ge_state: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun \\<^bsup>m\<^esup> c)); m \ n \ \ localState (f_Exec_Comp trans_fun \\<^bsup>n\<^esup> c) = localState (f_Exec_Comp trans_fun \\<^bsup>m\<^esup> c)" by (cut_tac f_Exec_State_Idle_append_replicate_NoMsg_ge_state[where xs="[]"], simp+) corollary f_Exec_State_Idle_append_replicate_NoMsg_gr0_output: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun xs c)); 0 < n \ \ output_fun (f_Exec_Comp trans_fun (xs @ \\<^bsup>n\<^esup>) c) = \" by (simp add: f_Exec_append f_Exec_State_Idle_replicate_NoMsg_gr0_output) corollary f_Exec_Stream_State_Idle_append_replicate_NoMsg_gr0_output: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun xs c)) \ \ map output_fun (f_Exec_Comp_Stream trans_fun (xs @ \\<^bsup>n\<^esup>) c) = map output_fun (f_Exec_Comp_Stream trans_fun xs c) @ \\<^bsup>n\<^esup>" by (simp add: f_Exec_Stream_append f_Exec_Stream_State_Idle_replicate_NoMsg_output) corollary f_Exec_State_Idle_append_replicate_NoMsg_gr_output: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun (xs @ \\<^bsup>m\<^esup>) c)); m < n \ \ output_fun (f_Exec_Comp trans_fun (xs @ \\<^bsup>n\<^esup>) c) = \" apply (rule_tac t=n and s="m + (n - m)" in subst, simp) apply (simp only: replicate_add append_assoc[symmetric]) apply (rule f_Exec_State_Idle_append_replicate_NoMsg_gr0_output, simp+) done corollary f_Exec_State_Idle_append_replicate_NoMsg_ge_output: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun (xs @ \\<^bsup>m\<^esup>) c)); output_fun (f_Exec_Comp trans_fun (xs @ \\<^bsup>m\<^esup>) c) = \; m \ n \ \ output_fun (f_Exec_Comp trans_fun (xs @ \\<^bsup>n\<^esup>) c) = \" by (fastforce simp: order_le_less f_Exec_State_Idle_append_replicate_NoMsg_gr_output) corollary f_Exec_State_Idle_replicate_NoMsg_gr_output: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun \\<^bsup>m\<^esup> c)); m < n \ \ output_fun (f_Exec_Comp trans_fun \\<^bsup>n\<^esup> c) = \" by (cut_tac xs="[]" in f_Exec_State_Idle_append_replicate_NoMsg_gr_output, simp+) corollary f_Exec_State_Idle_replicate_NoMsg_ge_output: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun \\<^bsup>m\<^esup> c)); output_fun (f_Exec_Comp trans_fun \\<^bsup>m\<^esup> c) = \; m \ n \ \ output_fun (f_Exec_Comp trans_fun \\<^bsup>n\<^esup> c) = \" by (fastforce simp: order_le_less f_Exec_State_Idle_replicate_NoMsg_gr_output) lemma State_Idle_append_replicate_NoMsg_output_last_message: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun xs c)) \ \ last_message (map output_fun (f_Exec_Comp_Stream trans_fun (xs @ \\<^bsup>n\<^esup>) c)) = last_message (map output_fun (f_Exec_Comp_Stream trans_fun xs c))" by (simp add: f_Exec_Stream_State_Idle_append_replicate_NoMsg_gr0_output last_message_append_replicate_NoMsg) lemma State_Idle_append_replicate_NoMsg_output_Msg_eq_last_message: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun xs c)); output_fun (f_Exec_Comp trans_fun xs c) \ \; xs \ [] \ \ last_message (map output_fun (f_Exec_Comp_Stream trans_fun (xs @ \\<^bsup>n\<^esup>) c)) = output_fun (f_Exec_Comp trans_fun xs c)" apply (simp add: State_Idle_append_replicate_NoMsg_output_last_message f_Exec_eq_f_Exec_Stream_last2 ) apply (subst last_message_Msg_eq_last) apply (simp add: map_last f_Exec_Stream_not_empty_conv)+ done corollary State_Idle_output_Msg_eq_last_message: " \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun xs c)); output_fun (f_Exec_Comp trans_fun xs c) \ \; xs \ [] \ \ last_message (map output_fun (f_Exec_Comp_Stream trans_fun xs c)) = output_fun (f_Exec_Comp trans_fun xs c)" by (rule_tac n=0 in subst[OF State_Idle_append_replicate_NoMsg_output_Msg_eq_last_message, rule_format], simp+) lemma State_Idle_imp_exists_state_change: " \ \ State_Idle localState output_fun trans_fun (localState c); State_Idle localState output_fun trans_fun (localState (f_Exec_Comp trans_fun \\<^bsup>n\<^esup> c)) \ \ \i State_Idle localState output_fun trans_fun (localState (f_Exec_Comp trans_fun \\<^bsup>i\<^esup> c)) \ ( \j\n. i < j \ State_Idle localState output_fun trans_fun (localState (f_Exec_Comp trans_fun \\<^bsup>j\<^esup> c))))" apply (cut_tac a=0 and b=n and P="\x. State_Idle localState output_fun trans_fun (localState (f_Exec_Comp trans_fun NoMsg\<^bsup>x\<^esup> c))" in nat_Suc_predicate_change_exists, simp+) apply (clarify, rename_tac n1) apply (rule_tac x=n1 in exI) apply clarsimp apply (rule_tac t="j" and s="Suc n1 + (j - Suc n1)" in subst, simp) apply (subst replicate_add) apply (simp add: replicate_add f_Exec_State_Idle_append_replicate_NoMsg_state) done lemma State_Idle_imp_exists_state_change2: " \ \ State_Idle localState output_fun trans_fun (localState c); State_Idle localState output_fun trans_fun (localState (f_Exec_Comp trans_fun \\<^bsup>n\<^esup> c)) \ \ \ij\i. \ State_Idle localState output_fun trans_fun (localState (f_Exec_Comp trans_fun \\<^bsup>i\<^esup> c))) \ (\j\n. i < j \ State_Idle localState output_fun trans_fun (localState (f_Exec_Comp trans_fun \\<^bsup>j\<^esup> c))))" apply (frule State_Idle_imp_exists_state_change, assumption) apply (clarify, rename_tac i) apply (rule_tac x=i in exI) apply simp done subsubsection \Basic definitions for accelerated execution\ text \Stream processing with accelerated components\ definition f_Exec_Comp_Stream_Acc_Output :: "nat \ \ \Acceleration factor\ ('comp \ 'output message_af) \ \ \Output extraction function\ ('comp, 'input message_af) Comp_Trans_Fun \ 'input fstream_af \ 'comp \ 'output fstream_af" where "f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c \ (map output_fun (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c)) \
\<^sub>f k" definition f_Exec_Comp_Stream_Acc_LocalState :: "nat \ \ \Acceleration factor\ ('comp \ 'state) \ \ \Local state extraction function\ ('comp, 'input message_af) Comp_Trans_Fun \ 'input fstream_af \ 'comp \ 'state list" where "f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c \ (map localState (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c)) \
\<^bsub>fl\<^esub> k" definition i_Exec_Comp_Stream_Acc_Output :: "nat \ \ \Acceleration factor\ ('comp \ 'output message_af) \ \ \Output extraction function\ ('comp, 'input message_af) Comp_Trans_Fun \ 'input istream_af \ 'comp \ 'output istream_af" where "i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c \ (output_fun \ (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c)) \
\<^sub>i k" definition i_Exec_Comp_Stream_Acc_LocalState :: "nat \ \ \Acceleration factor\ ('comp \ 'state) \ \ \Local state extraction function\ ('comp, 'input message_af) Comp_Trans_Fun \ 'input istream_af \ 'comp \ 'state ilist" where "i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input c \ (localState \ (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c)) \
\<^bsub>il\<^esub> k" definition f_Exec_Comp_Stream_Acc_Output_Init :: "nat \ \ \Acceleration factor\ ('comp \ 'output message_af) \ \ \Output extraction function\ ('comp, 'input message_af) Comp_Trans_Fun \ 'input fstream_af \ 'comp \ 'output fstream_af" where "f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun xs c \ (output_fun c) # f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c" definition f_Exec_Comp_Stream_Acc_LocalState_Init :: "nat \ \ \Acceleration factor\ ('comp \ 'state) \ \ \Local state extraction function\ ('comp, 'input message_af) Comp_Trans_Fun \ 'input fstream_af \ 'comp \ 'state list" where "f_Exec_Comp_Stream_Acc_LocalState_Init k localState trans_fun xs c \ (localState c) # f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c" definition i_Exec_Comp_Stream_Acc_Output_Init :: "nat \ \ \Acceleration factor\ ('comp \ 'output message_af) \ \ \Output extraction function\ ('comp, 'input message_af) Comp_Trans_Fun \ 'input istream_af \ 'comp \ 'output istream_af" where "i_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun input c \ [output_fun c] \ (i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c)" definition i_Exec_Comp_Stream_Acc_LocalState_Init :: "nat \ \ \Acceleration factor\ ('comp \ 'state) \ \ \Local state extraction function\ ('comp, 'input message_af) Comp_Trans_Fun \ 'input istream_af \ 'comp \ 'state ilist" where "i_Exec_Comp_Stream_Acc_LocalState_Init k localState trans_fun input c \ [localState c] \ (i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input c)" lemma f_Exec_Stream_Acc_Output_length[simp]: " 0 < k \ length (f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c) = length xs" by (simp add: f_Exec_Comp_Stream_Acc_Output_def f_shrink_length) lemma f_Exec_Stream_Acc_LocalState_length[simp]: " 0 < k \ length (f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c) = length xs" by (simp add: f_Exec_Comp_Stream_Acc_LocalState_def f_shrink_last_length) lemmas f_Exec_Stream_Acc_length = f_Exec_Stream_Acc_LocalState_length f_Exec_Stream_Acc_Output_length subsubsection \Basic results for accelerated execution\ lemma f_Exec_Stream_Acc_Output_Nil[simp]: " f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun [] c = []" by (simp add: f_Exec_Comp_Stream_Acc_Output_def) lemma f_Exec_Stream_Acc_LocalState_Nil[simp]: " f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun [] c = []" by (simp add: f_Exec_Comp_Stream_Acc_LocalState_def) lemmas f_Exec_Stream_Acc_Nil = f_Exec_Stream_Acc_LocalState_Nil f_Exec_Stream_Acc_Output_Nil lemma f_Exec_Stream_Acc_Output_0[simp]: " f_Exec_Comp_Stream_Acc_Output 0 output_fun trans_fun xs c = []" by (simp add: f_Exec_Comp_Stream_Acc_Output_def) lemma f_Exec_Stream_Acc_LocalState_0[simp]: " f_Exec_Comp_Stream_Acc_LocalState 0 localState trans_fun xs c = []" by (simp add: f_Exec_Comp_Stream_Acc_LocalState_def) lemmas f_Exec_Stream_Acc_0 = f_Exec_Stream_Acc_LocalState_0 f_Exec_Stream_Acc_Output_0 lemma f_Exec_Stream_Acc_Output_1[simp]: " f_Exec_Comp_Stream_Acc_Output (Suc 0) output_fun trans_fun xs c = map output_fun (f_Exec_Comp_Stream trans_fun xs c)" by (simp add: f_Exec_Comp_Stream_Acc_Output_def) lemma f_Exec_Stream_Acc_LocalState_1[simp]: " f_Exec_Comp_Stream_Acc_LocalState (Suc 0) localState trans_fun xs c = map localState (f_Exec_Comp_Stream trans_fun xs c)" by (simp add: f_Exec_Comp_Stream_Acc_LocalState_def) lemma i_Exec_Stream_Acc_Output_1[simp]: " i_Exec_Comp_Stream_Acc_Output (Suc 0) output_fun trans_fun input c = output_fun \ (i_Exec_Comp_Stream trans_fun input c)" by (simp add: i_Exec_Comp_Stream_Acc_Output_def) lemma i_Exec_Stream_Acc_LocalState_1[simp]: " i_Exec_Comp_Stream_Acc_LocalState (Suc 0) localState trans_fun input c = localState \ (i_Exec_Comp_Stream trans_fun input c)" by (simp add: i_Exec_Comp_Stream_Acc_LocalState_def) lemma f_Exec_Stream_Acc_Output_eq_last_message_hold: " f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c = (map output_fun (f_Exec_Comp_Stream trans_fun (xs \\<^sub>f k) c)) \\<^sub>f k \
\<^bsub>fl\<^esub> k" by (simp add: f_Exec_Comp_Stream_Acc_Output_def f_shrink_eq_f_last_message_hold_shrink_last) lemma i_Exec_Stream_Acc_Output_eq_last_message_hold: "0 < k \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c = (output_fun \ (i_Exec_Comp_Stream trans_fun (input \\<^sub>i k) c)) \\<^sub>i k \
\<^bsub>il\<^esub> k" by (simp add: i_Exec_Comp_Stream_Acc_Output_def i_shrink_eq_i_last_message_hold_shrink_last) lemma f_Exec_Stream_Acc_Output_take: " f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c \ n = f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun (xs \ n) c" by (simp add: f_Exec_Comp_Stream_Acc_Output_def f_shrink_def f_Exec_Stream_expand_aggregate_map_take) lemma f_Exec_Stream_Acc_Output_drop: " f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c \ n = f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun (xs \ n) ( f_Exec_Comp trans_fun (xs \ n \\<^sub>f k) c)" by (simp add: f_Exec_Comp_Stream_Acc_Output_def f_shrink_def f_Exec_Stream_expand_aggregate_map_drop) lemma i_Exec_Stream_Acc_Output_take: " 0 < k \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c \ n = f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun (input \ n) c" by (simp add: f_Exec_Comp_Stream_Acc_Output_def i_Exec_Comp_Stream_Acc_Output_def f_shrink_def i_shrink_def i_Exec_Stream_expand_aggregate_map_take) lemma i_Exec_Stream_Acc_Output_drop: " 0 < k \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c \ n = i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun (input \ n) ( f_Exec_Comp trans_fun (input \ n \\<^sub>f k) c)" by (simp add: i_Exec_Comp_Stream_Acc_Output_def i_shrink_def i_Exec_Stream_expand_aggregate_map_drop) lemma i_Exec_Stream_Acc_LocalState_take: " 0 < k \ i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input c \ n = f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun (input \ n) c" by (simp add: f_Exec_Comp_Stream_Acc_LocalState_def i_Exec_Comp_Stream_Acc_LocalState_def f_shrink_last_def i_shrink_last_def i_Exec_Stream_expand_aggregate_map_take) lemma i_Exec_Stream_Acc_LocalState_drop: " 0 < k \ i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input c \ n = i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun (input \ n) ( f_Exec_Comp trans_fun (input \ n \\<^sub>f k) c)" by (simp add: i_Exec_Comp_Stream_Acc_LocalState_def i_shrink_last_def i_Exec_Stream_expand_aggregate_map_drop) lemma f_Exec_Stream_Acc_Output_append: " f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun (xs @ ys) c = f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c @ f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun ys ( f_Exec_Comp trans_fun (xs \\<^sub>f k) c)" by (simp only: f_Exec_Comp_Stream_Acc_Output_def f_shrink_def f_Exec_Stream_expand_map_aggregate_append) lemma f_Exec_Stream_Acc_Output_Cons: " 0 < k \ f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun (x # xs) c = last_message (map output_fun (f_Exec_Comp_Stream trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c)) # f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs ( f_Exec_Comp trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c)" by (simp only: f_Exec_Comp_Stream_Acc_Output_def f_shrink_def f_Exec_Stream_expand_map_aggregate_Cons) lemma f_Exec_Stream_Acc_Output_one: " 0 < k \ f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun [x] c = [last_message (map output_fun (f_Exec_Comp_Stream trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c))]" by (simp add: f_Exec_Stream_Acc_Output_Cons) lemma f_Exec_Stream_Acc_Output_snoc: " 0 < k \ f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun (xs @ [x]) c = f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c @ [last_message (map output_fun (f_Exec_Comp_Stream trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) ( f_Exec_Comp trans_fun (xs \\<^sub>f k) c)))]" by (simp add: f_Exec_Stream_Acc_Output_append f_Exec_Stream_Acc_Output_one) lemma i_Exec_Stream_Acc_Output_append: " i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun (xs \ input) c = f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input ( f_Exec_Comp trans_fun (xs \\<^sub>f k) c)" by (simp add: f_Exec_Comp_Stream_Acc_Output_def i_Exec_Comp_Stream_Acc_Output_def f_shrink_def i_shrink_def i_Exec_Stream_expand_map_aggregate_append) lemma i_Exec_Stream_Acc_Output_Cons: " 0 < k \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun ([x] \ input) c = [last_message (map output_fun (f_Exec_Comp_Stream trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c))] \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input ( f_Exec_Comp trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c)" by (simp add: i_Exec_Stream_Acc_Output_append f_Exec_Stream_Acc_Output_one) lemma f_Exec_Stream_Acc_LocalState_append: " f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun (xs @ ys) c = f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c @ f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun ys ( f_Exec_Comp trans_fun (xs \\<^sub>f k) c)" by (simp only: f_Exec_Comp_Stream_Acc_LocalState_def f_shrink_last_def f_Exec_Stream_expand_map_aggregate_append) lemma f_Exec_Stream_Acc_LocalState_Cons: " 0 < k \ f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun (x # xs) c = localState (f_Exec_Comp trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c) # f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs ( f_Exec_Comp trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c)" apply (unfold f_Exec_Comp_Stream_Acc_LocalState_def) apply (simp only: f_shrink_last_map f_expand_Cons append_Cons[symmetric]) apply (simp add: f_Exec_Stream_append replicate_pred_Cons_length f_shrink_last_Cons del: f_Exec_Stream_Cons append_Cons) apply (simp add: f_Exec_eq_f_Exec_Stream_last2[symmetric] f_Exec_Stream_empty_conv) done lemma f_Exec_Stream_Acc_LocalState_one: " 0 < k \ f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun [x] c = [localState (f_Exec_Comp trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c)]" by (simp add: f_Exec_Stream_Acc_LocalState_Cons) lemma f_Exec_Stream_Acc_LocalState_snoc: " 0 < k \ f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun (xs @ [x]) c = f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c @ [localState (f_Exec_Comp trans_fun ((xs @ [x]) \\<^sub>f k) c)]" by (simp add: f_Exec_Stream_Acc_LocalState_append f_Exec_Stream_Acc_LocalState_Cons f_Exec_append) lemma i_Exec_Stream_Acc_LocalState_append: " i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun (xs \ input) c = f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c \ i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input ( f_Exec_Comp trans_fun (xs \\<^sub>f k) c)" by (simp add: f_Exec_Comp_Stream_Acc_LocalState_def i_Exec_Comp_Stream_Acc_LocalState_def f_shrink_last_def i_shrink_last_def i_Exec_Stream_expand_map_aggregate_append) lemma i_Exec_Stream_Acc_LocalState_Cons: " 0 < k \ i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun ([x] \ input) c = [localState (f_Exec_Comp trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c)] \ i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input ( f_Exec_Comp trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c)" by (simp add: i_Exec_Stream_Acc_LocalState_append f_Exec_Stream_Acc_LocalState_one f_expand_one) lemma f_Exec_Stream_Acc_Output_nth: " \ 0 < k; n < length xs \ \ f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c ! n = last_message (map output_fun ( f_Exec_Comp_Stream trans_fun (xs ! n # \\<^bsup>k - Suc 0\<^esup>) ( f_Exec_Comp trans_fun (xs \ n \\<^sub>f k) c)))" by (unfold f_Exec_Comp_Stream_Acc_Output_def f_shrink_def, rule f_Exec_Stream_expand_aggregate_map_nth) lemma f_Exec_Stream_Acc_Output_nth_eq_i_nth: " \ 0 < k; n < n' \ \ f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun (input \ n') c ! n = i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c n" by (unfold f_Exec_Comp_Stream_Acc_Output_def i_Exec_Comp_Stream_Acc_Output_def f_shrink_def i_shrink_def, rule f_Exec_Stream_expand_aggregate_map_nth_eq_i_nth) lemma i_Exec_Stream_Acc_Output_nth: " 0 < k \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c n = last_message (map output_fun ( f_Exec_Comp_Stream trans_fun (input n # \\<^bsup>k - Suc 0\<^esup>) ( f_Exec_Comp trans_fun (input \ n \\<^sub>f k) c)))" by (unfold i_Exec_Comp_Stream_Acc_Output_def i_shrink_def, rule i_Exec_Stream_expand_aggregate_map_nth) corollary i_Exec_Stream_Acc_Output_nth_f_nth: " 0 < k \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c n = f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun (input \ Suc n) c ! n" by (simp add: f_Exec_Stream_Acc_Output_nth_eq_i_nth) corollary i_Exec_Stream_Acc_Output_nth_f_last: " 0 < k \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c n = last (f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun (input \ Suc n) c)" by (simp add: i_Exec_Stream_Acc_Output_nth_f_nth last_nth length_greater_0_conv[THEN iffD1]) lemma f_Exec_Stream_Acc_LocalState_nth: " \ 0 < k; n < length xs \ \ f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c ! n = localState (f_Exec_Comp trans_fun (xs \ Suc n \\<^sub>f k) c)" apply (simp add: f_Exec_Comp_Stream_Acc_LocalState_def f_shrink_last_map) apply (simp add: f_shrink_last_nth' f_shrink_last_length del: mult_Suc) apply (simp add: f_Exec_Stream_nth less_imp_Suc_mult_pred_less f_expand_take_mod del: mult_Suc) done lemma f_Exec_Stream_Acc_LocalState_nth_eq_i_nth: " \ 0 < k; n < n' \ \ f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun (input \ n') c ! n = i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input c n" by (unfold f_Exec_Comp_Stream_Acc_LocalState_def i_Exec_Comp_Stream_Acc_LocalState_def f_shrink_last_def i_shrink_last_def, rule f_Exec_Stream_expand_aggregate_map_nth_eq_i_nth) corollary i_Exec_Stream_Acc_LocalState_nth_f_nth: " 0 < k \ i_Exec_Comp_Stream_Acc_LocalState k output_fun trans_fun input c n = f_Exec_Comp_Stream_Acc_LocalState k output_fun trans_fun (input \ Suc n) c ! n" by (simp add: f_Exec_Stream_Acc_LocalState_nth_eq_i_nth) corollary i_Exec_Stream_Acc_LocalState_nth_f_last: " 0 < k \ i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input c n = last (f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun (input \ Suc n) c)" by (simp add: i_Exec_Stream_Acc_LocalState_nth_f_nth last_nth length_greater_0_conv[THEN iffD1]) lemma i_Exec_Stream_Acc_LocalState_nth: " 0 < k \ i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input c n = localState (f_Exec_Comp trans_fun (input \ Suc n \\<^sub>f k) c)" by (simp add: i_Exec_Stream_Acc_LocalState_nth_f_nth f_Exec_Stream_Acc_LocalState_nth) lemma f_Exec_Stream_Acc_Output_causal: " xs \ n = ys \ n \ f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c \ n = f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun ys c \ n" by (simp add: f_Exec_Stream_Acc_Output_take) lemma i_Exec_Stream_Acc_Output_causal: " input1 \ n = input2 \ n \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input1 c \ n = i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input2 c \ n" apply (case_tac "k = 0") apply (simp add: i_Exec_Comp_Stream_Acc_Output_def) apply (simp add: i_Exec_Stream_Acc_Output_take) done lemma f_Exec_Stream_Acc_Output_Connected_strictly_causal: " \ xs \ n = ys \ n; f_Streams_Connected (f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c) channel1; f_Streams_Connected (f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun ys c) channel2 \ \ channel1 \ Suc n = channel2 \ Suc n" by (simp add: f_Streams_Connected_def f_Exec_Stream_Acc_Output_take) lemma i_Exec_Stream_Acc_Output_Connected_strictly_causal: " \ input1 \ n = input2 \ n; i_Streams_Connected (i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input1 c) channel1; i_Streams_Connected (i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input2 c) channel2 \ \ channel1 \ Suc n = channel2 \ Suc n" apply (unfold i_Streams_Connected_def) apply (case_tac "k = 0") apply (simp add: i_Exec_Comp_Stream_Acc_Output_def) apply (simp add: i_Exec_Stream_Acc_Output_take) done text \Complete execution cycles/steps of accelrated execution\ definition Acc_Trans_Fun_Step :: "nat \ \ \Acceleration factor\ ('comp, 'input message_af) Comp_Trans_Fun \ ('comp list \ 'comp) \ \ \Pointwise output shrink function\ 'input message_af \ 'comp \ 'comp" where "Acc_Trans_Fun_Step k trans_fun pointwise_shrink x c \ pointwise_shrink (f_Exec_Comp_Stream trans_fun (x # \\<^bsup>k - Suc 0\<^esup>) c)" definition is_Pointwise_Output_Shrink :: "('comp list \ 'comp) \ \ \Pointwise output shrink function\ ('comp \ 'output message_af) \ \ \Output extraction function for consideration\ bool" where "is_Pointwise_Output_Shrink pointwise_shrink output_fun \ \cs. output_fun (pointwise_shrink cs) = last_message (map output_fun cs)" primrec is_Pointwise_Output_Shrink_list :: "('comp list \ 'comp) \ \ \Pointwise output shrink function\ ('comp \ 'output message_af) list \ \ \List of output extraction functions for consideration\ bool" where "is_Pointwise_Output_Shrink_list pointwise_shrink [] = True" | "is_Pointwise_Output_Shrink_list pointwise_shrink (f # fs) = (is_Pointwise_Output_Shrink pointwise_shrink f \ is_Pointwise_Output_Shrink_list pointwise_shrink fs)" definition is_correct_localState_Pointwise_Output_Shrink :: "('comp list \ 'comp) \ \ \Pointwise output shrink function\ ('comp \ 'state) \ \ \Local state extraction function\ bool" where "is_correct_localState_Pointwise_Output_Shrink pointwise_shrink localState \ \cs. cs \ [] \ localState (pointwise_shrink cs) = localState (last cs)" lemma Deterministic_trans_fun_imp_acc_trans_fun: "Deterministic_Trans_Fun trans_fun localState \ Deterministic_Trans_Fun (Acc_Trans_Fun_Step k trans_fun pointwise_shrink) localState" apply (simp (no_asm) only: Deterministic_Trans_Fun_def Acc_Trans_Fun_Step_def) apply clarify apply (subst Deterministic_f_Exec_Stream, simp+) done lemma is_Pointwise_Output_Shrink_list_imp_is_Pointwise_Output_Shrink: "\ is_Pointwise_Output_Shrink_list pointwise_shrink fs; output_fun \ set fs \ \ is_Pointwise_Output_Shrink pointwise_shrink output_fun" apply (induct fs, simp) apply fastforce done lemma is_Pointwise_Output_Shrink_list_eq_is_Pointwise_Output_Shrink_all: "(is_Pointwise_Output_Shrink_list pointwise_shrink fs) = (\output_fun \ set fs. is_Pointwise_Output_Shrink pointwise_shrink output_fun)" apply (rule iffI) apply (rule ballI) apply (rule is_Pointwise_Output_Shrink_list_imp_is_Pointwise_Output_Shrink) apply (simp add: member_def)+ apply (induct fs, simp) apply simp done lemma is_Pointwise_Output_Shrink_subset: "\ is_Pointwise_Output_Shrink_list pointwise_shrink fs; set fs' \ set fs \ \ is_Pointwise_Output_Shrink_list pointwise_shrink fs'" by (fastforce simp: is_Pointwise_Output_Shrink_list_eq_is_Pointwise_Output_Shrink_all) lemma f_Exec_Stream_Acc_LocalState_eq_Acc_Trans_Fun_Step_LocalState: "\c. \ 0 < k; Deterministic_Trans_Fun trans_fun localState; is_correct_localState_Pointwise_Output_Shrink pointwise_shrink localState \ \ f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c = map localState (f_Exec_Comp_Stream (Acc_Trans_Fun_Step k trans_fun pointwise_shrink) xs c)" apply (drule Deterministic_trans_fun_imp_acc_trans_fun[of trans_fun localState k pointwise_shrink]) apply (clarsimp simp: list_eq_iff) apply (simp add: f_Exec_Stream_Acc_LocalState_nth f_Exec_Stream_nth) apply (induct xs, simp) apply (rename_tac x xs c i) apply (simp add: Acc_Trans_Fun_Step_def f_expand_Cons f_Exec_append) apply (case_tac i) apply simp apply (simp only: is_correct_localState_Pointwise_Output_Shrink_def) apply (drule_tac x="f_Exec_Comp_Stream trans_fun (x # NoMsg\<^bsup>k - Suc 0\<^esup>) c" in spec) apply (simp add: f_Exec_Stream_not_empty_conv f_Exec_eq_f_Exec_Stream_last) apply (rename_tac i2) apply (drule_tac x="f_Exec_Comp trans_fun \\<^bsup>k - Suc 0\<^esup> (trans_fun x c)" in meta_spec) apply (drule_tac x=i2 in meta_spec) apply (simp add: is_correct_localState_Pointwise_Output_Shrink_def) apply (drule_tac x="f_Exec_Comp_Stream trans_fun (x # NoMsg\<^bsup>k - Suc 0\<^esup>) c" in spec) apply (simp add: f_Exec_Stream_not_empty_conv) apply (rule arg_cong[where f=localState]) apply (rule Deterministic_f_Exec) apply assumption apply (simp add: f_Exec_eq_f_Exec_Stream_last) apply (simp add: length_greater_0_conv[symmetric] del: length_greater_0_conv) done lemma f_Exec_Stream_Acc_Output_eq_Acc_Trans_Fun_Step_Output: "\c. \ 0 < k; Deterministic_Trans_Fun trans_fun localState; is_correct_localState_Pointwise_Output_Shrink pointwise_shrink localState; is_Pointwise_Output_Shrink pointwise_shrink output_fun \ \ f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c = map output_fun (f_Exec_Comp_Stream (Acc_Trans_Fun_Step k trans_fun pointwise_shrink) xs c)" apply (drule Deterministic_trans_fun_imp_acc_trans_fun[of trans_fun localState k pointwise_shrink]) apply (clarsimp simp: list_eq_iff) apply (simp add: f_Exec_Stream_Acc_Output_nth f_Exec_Stream_nth del: f_Exec_Stream_Cons) apply (induct xs, simp) apply (rename_tac x xs c i) apply (simp add: Acc_Trans_Fun_Step_def del: f_Exec_Stream_Cons) apply (case_tac i) apply (simp add: is_Pointwise_Output_Shrink_def) apply (rename_tac i2) apply (simp add: f_Exec_append) apply (drule_tac x="f_Exec_Comp trans_fun \\<^bsup>k - Suc 0\<^esup> (trans_fun x c)" in meta_spec) apply (drule_tac x=i2 in meta_spec) apply (simp add: is_correct_localState_Pointwise_Output_Shrink_def) apply (drule_tac x="f_Exec_Comp_Stream trans_fun (x # NoMsg\<^bsup>k - Suc 0\<^esup>) c" in spec) apply (simp add: f_Exec_Stream_not_empty_conv) apply (rule arg_cong[where f=output_fun]) apply (rule Deterministic_f_Exec) apply assumption apply (simp add: f_Exec_eq_f_Exec_Stream_last) apply (simp add: length_greater_0_conv[symmetric] del: length_greater_0_conv) done lemma i_Exec_Stream_Acc_LocalState_eq_Acc_Trans_Fun_Step_LocalState: "\c. \ 0 < k; Deterministic_Trans_Fun trans_fun localState; is_correct_localState_Pointwise_Output_Shrink pointwise_shrink localState \ \ i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input c = localState \ (i_Exec_Comp_Stream (Acc_Trans_Fun_Step k trans_fun pointwise_shrink) input c)" apply (rule ilist_i_take_eq_conv[THEN iffD2], rule allI) apply (simp add: i_Exec_Stream_Acc_LocalState_take i_Exec_Stream_take f_Exec_Stream_Acc_LocalState_eq_Acc_Trans_Fun_Step_LocalState) done lemma i_Exec_Stream_Acc_Output_eq_Acc_Trans_Fun_Step_Output: "\c. \ 0 < k; Deterministic_Trans_Fun trans_fun localState; is_correct_localState_Pointwise_Output_Shrink pointwise_shrink localState; is_Pointwise_Output_Shrink pointwise_shrink output_fun \ \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c = output_fun \ (i_Exec_Comp_Stream (Acc_Trans_Fun_Step k trans_fun pointwise_shrink) input c)" apply (rule ilist_i_take_eq_conv[THEN iffD2], rule allI) apply (simp add: i_Exec_Stream_Acc_Output_take i_Exec_Stream_take f_Exec_Stream_Acc_Output_eq_Acc_Trans_Fun_Step_Output) done subsubsection \Basic results for accelerated execution with initial state in the resulting stream\ lemma f_Exec_Stream_Acc_Output_Init_length: " 0 < k \ length (f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun xs c) = Suc (length xs)" by (simp add: f_Exec_Comp_Stream_Acc_Output_Init_def) lemma f_Exec_Stream_Acc_LocalState_Init_length: " 0 < k \ length (f_Exec_Comp_Stream_Acc_LocalState_Init k localState trans_fun xs c) = Suc (length xs)" by (simp add: f_Exec_Comp_Stream_Acc_LocalState_Init_def) lemma f_Exec_Stream_Acc_Output_Init_Nil: " f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun [] c = [output_fun c]" by (simp add: f_Exec_Comp_Stream_Acc_Output_Init_def) lemma f_Exec_Stream_Acc_LocalState_Init_Nil: " f_Exec_Comp_Stream_Acc_LocalState_Init k localState trans_fun [] c = [localState c]" by (simp add: f_Exec_Comp_Stream_Acc_LocalState_Init_def) lemma f_Exec_Stream_Acc_Output_Init_1: " f_Exec_Comp_Stream_Acc_Output_Init (Suc 0) output_fun trans_fun xs c = map output_fun (f_Exec_Comp_Stream_Init trans_fun xs c)" by (simp add: f_Exec_Comp_Stream_Acc_Output_Init_def f_Exec_Stream_Init_eq_f_Exec_Stream_Cons) lemma f_Exec_Stream_Acc_LocalState_Init_1: " f_Exec_Comp_Stream_Acc_LocalState_Init (Suc 0) localState trans_fun xs c = map localState (f_Exec_Comp_Stream_Init trans_fun xs c)" by (simp add: f_Exec_Comp_Stream_Acc_LocalState_Init_def f_Exec_Stream_Init_eq_f_Exec_Stream_Cons) lemma i_Exec_Stream_Acc_Output_Init_1: " i_Exec_Comp_Stream_Acc_Output_Init (Suc 0) output_fun trans_fun input c = output_fun \ (i_Exec_Comp_Stream_Init trans_fun input c)" by (simp add: i_Exec_Comp_Stream_Acc_Output_Init_def i_Exec_Stream_Init_eq_i_Exec_Stream_Cons) lemma i_Exec_Stream_Acc_LocalState_Init_1: " i_Exec_Comp_Stream_Acc_LocalState_Init (Suc 0) localState trans_fun input c = localState \ (i_Exec_Comp_Stream_Init trans_fun input c)" by (simp add: i_Exec_Comp_Stream_Acc_LocalState_Init_def i_Exec_Stream_Init_eq_i_Exec_Stream_Cons) lemma f_Exec_Stream_Acc_Output_Init_take: " f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun xs c \ (Suc n) = f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun (xs \ n) c" by (simp add: f_Exec_Comp_Stream_Acc_Output_Init_def f_Exec_Stream_Acc_Output_take) lemma f_Exec_Stream_Acc_Output_Init_drop': " \ 0 < k; n < length xs \ \ f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun xs c \ Suc n = f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c \ n" by (simp add: f_Exec_Comp_Stream_Acc_Output_Init_def) lemma i_Exec_Stream_Acc_Output_Init_take: " 0 < k \ i_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun input c \ (Suc n) = f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun (input \ n) c" by (simp add: f_Exec_Comp_Stream_Acc_Output_Init_def i_Exec_Comp_Stream_Acc_Output_Init_def i_Exec_Stream_Acc_Output_take) lemma i_Exec_Stream_Acc_Output_Init_drop': " 0 < k \ i_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun xs c \ Suc n = i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c \ n" by (simp add: i_Exec_Comp_Stream_Acc_Output_Init_def) lemma f_Exec_Stream_Acc_Output_Init_strictly_causal: " xs \ n = ys \ n \ f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun xs c \ Suc n = f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun ys c \ Suc n" by (simp add: f_Exec_Comp_Stream_Acc_Output_Init_def, rule f_Exec_Stream_Acc_Output_causal) lemma i_Exec_Stream_Acc_Output_Init_strictly_causal: " input1 \ n = input2 \ n \ i_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun input1 c \ Suc n = i_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun input2 c \ Suc n" by (simp add: i_Exec_Comp_Stream_Acc_Output_Init_def, rule i_Exec_Stream_Acc_Output_causal) lemma f_Exec_Stream_Acc_Output_Init_eq_f_Exec_Stream_Acc_Output_Cons: " f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun xs c = output_fun c # f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c" by (simp add: f_Exec_Comp_Stream_Acc_Output_def f_Exec_Comp_Stream_Acc_Output_Init_def) lemma f_Exec_Stream_Acc_Output_Init_eq_f_Exec_Stream_Acc_Output_Cons_output: " output_fun c = \ \ f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun xs c = \ # f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c" by (simp add: f_Exec_Stream_Acc_Output_Init_eq_f_Exec_Stream_Acc_Output_Cons) lemma f_Exec_Stream__Acc_OutputInit_tl_eq_f_Exec_Stream_Acc_Output: " tl (f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun xs c) = f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c" by (simp add: f_Exec_Stream_Acc_Output_Init_eq_f_Exec_Stream_Acc_Output_Cons) lemma f_Exec_Stream_Previous_f_Exec_Stream_Acc_Output_Init: " f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun xs c ! n = (f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c)\<^bsup>\' output_fun c\<^esup> n" by (simp add: f_Exec_Stream_Acc_Output_Init_eq_f_Exec_Stream_Acc_Output_Cons list_Previous_nth_if nth_Cons') lemma f_Exec_Stream_Acc_Output_Init_eq_output_channel: " \ output_fun c = \; f_Streams_Connected (f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c) channel \ \ f_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun xs c = channel" by (simp add: f_Streams_Connected_def f_Exec_Stream_Acc_Output_Init_eq_f_Exec_Stream_Acc_Output_Cons_output) lemma i_Exec_Stream_Acc_Output_Init_eq_i_Exec_Stream_Acc_Output_Cons: " i_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun input c = [output_fun c] \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c" by (simp add: i_Exec_Comp_Stream_Acc_Output_def i_Exec_Comp_Stream_Acc_Output_Init_def) lemma i_Exec_Stream_Acc_Output_Init_eq_i_Exec_Stream_Acc_Output_Cons_output: " output_fun c = \ \ i_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun input c = [\ ] \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c" by (simp add: i_Exec_Stream_Acc_Output_Init_eq_i_Exec_Stream_Acc_Output_Cons) lemma i_Exec_Stream_Previous_i_Exec_Stream_Acc_Output_Init: " i_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun input c n = (i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c)\<^bsup>\ output_fun c\<^esup> n" by (simp add: i_Exec_Stream_Acc_Output_Init_eq_i_Exec_Stream_Acc_Output_Cons ilist_Previous_nth_if) lemma i_Exec_Stream_Acc_Output_Init_eq_output_channel: " \ output_fun c = \; i_Streams_Connected (i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c) channel \ \ i_Exec_Comp_Stream_Acc_Output_Init k output_fun trans_fun input c = channel" by (simp add: i_Streams_Connected_def i_Exec_Stream_Acc_Output_Init_eq_i_Exec_Stream_Acc_Output_Cons_output) subsubsection \Rules for proving execution equivalence\ text \ A required precondition is that the @{term equiv_states} relation, which indicates whether the local states of @{term c1} and @{term c2} are equivalent with respect to observable behaviour, is preserved also after executing an input stream, because the @{term equiv_states} relation should deliver valid results not only at the time point @{term 0} but at every time point.\ lemma f_Equiv_Exec_Stream_expand_shrink_equiv_state_set[rule_format]: " \c1 c2 i. \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); \input0. set input0 \ A \ (\m\A. Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input0 \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input0 \\<^sub>f k2) c2)); \ \\equiv_states\ relation implies equivalent executions\ \ \not only at the beginning but also after processing an input\ set input \ A; i < length input \ \ equiv_states (localState1 ((f_Exec_Comp_Stream trans_fun1 (map input_fun1 input \\<^sub>f k1) c1) \
\<^bsub>fl\<^esub> k1 ! i)) (localState2 ((f_Exec_Comp_Stream trans_fun2 (map input_fun2 input \\<^sub>f k2) c2) \
\<^bsub>fl\<^esub> k2 ! i))" apply (induct input, simp) apply (clarsimp simp: append_Cons[symmetric] f_Exec_Stream_append_if f_shrink_last_Cons nth_Cons simp del: last.simps f_Exec_Stream_Cons append_Cons) apply (case_tac i) apply (drule_tac x="[]" in spec) apply (drule mp, simp) apply (drule_tac x=a in bspec, assumption) apply (simp del: last.simps f_Exec_Stream_Cons) apply (subst f_Exec_eq_f_Exec_Stream_last2[symmetric], simp)+ apply (rule Equiv_Exec_equiv_statesI[of equiv_states localState1 _ localState2 _ _ input_fun1], assumption+) apply (rename_tac i') apply (subst f_Exec_eq_f_Exec_Stream_last2[symmetric], simp)+ apply (drule_tac x="f_Exec_Comp trans_fun1 (input_fun1 a # \\<^bsup>k1 - Suc 0\<^esup>) c1" in meta_spec) apply (drule_tac x="f_Exec_Comp trans_fun2 (input_fun2 a # \\<^bsup>k2 - Suc 0\<^esup>) c2" in meta_spec) apply (drule_tac x=i' in meta_spec) apply (drule meta_mp, simp)+ apply (drule_tac x="[]" in spec, simp) apply (drule_tac x=a in bspec, assumption) apply (rule Equiv_Exec_equiv_statesI'[of equiv_states localState1 _ localState2 _ _ input_fun1], simp+) apply clarsimp apply (drule meta_mp) apply clarify apply (drule_tac x="a # input0" in spec) apply (simp add: f_Exec_append) apply simp done corollary f_Equiv_Exec_Stream_expand_shrink_equiv_state: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); \input0 m. Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input0 \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input0 \\<^sub>f k2) c2); i \ equiv_states (localState1 ((f_Exec_Comp_Stream trans_fun1 (map input_fun1 input \\<^sub>f k1) c1) \
\<^bsub>fl\<^esub> k1 ! i)) (localState2 ((f_Exec_Comp_Stream trans_fun2 (map input_fun2 input \\<^sub>f k2) c2) \
\<^bsub>fl\<^esub> k2 ! i))" by (rule f_Equiv_Exec_Stream_expand_shrink_equiv_state_set[of k1 k2 equiv_states localState1 c1 localState2 c2 UNIV input_fun1 input_fun2 output_fun1 output_fun2], simp+) lemma f_Equiv_Exec_expand_shrink_equiv_state_set:" \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); \input0 m. \set input0 \ A; m \ A\ \ Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input0 \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input0 \\<^sub>f k2) c2); set input \ A \ \ equiv_states (localState1 (f_Exec_Comp trans_fun1 (map input_fun1 input \\<^sub>f k1) c1)) (localState2 (f_Exec_Comp trans_fun2 (map input_fun2 input \\<^sub>f k2) c2))" apply (case_tac "input = []", simp) apply (subgoal_tac "map input_fun1 input \\<^sub>f k1 \ [] \ map input_fun2 input \\<^sub>f k2 \ []") prefer 2 apply (simp add: length_greater_0_conv[symmetric] del: length_greater_0_conv) apply (simp add: f_Exec_eq_f_Exec_Stream_last2 last_nth f_Exec_Stream_not_empty_conv) apply (insert f_shrink_last_nth[of "length input - Suc 0" "f_Exec_Comp_Stream trans_fun1 (map input_fun1 input \\<^sub>f k1) c1" k1, symmetric]) apply (insert f_shrink_last_nth[of "length input - Suc 0" "f_Exec_Comp_Stream trans_fun2 (map input_fun2 input \\<^sub>f k2) c2" k2, symmetric]) apply (simp add: diff_mult_distrib gr0_imp_self_le_mult2) apply (rule f_Equiv_Exec_Stream_expand_shrink_equiv_state_set[of k1 k2 equiv_states localState1 _ localState2 _ A input_fun1 input_fun2 output_fun1 output_fun2]) apply simp+ done lemma f_Equiv_Exec_expand_shrink_equiv_state:" \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); \input0 m. Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input0 \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input0 \\<^sub>f k2) c2) \ \ equiv_states (localState1 (f_Exec_Comp trans_fun1 (map input_fun1 input \\<^sub>f k1) c1)) (localState2 (f_Exec_Comp trans_fun2 (map input_fun2 input \\<^sub>f k2) c2))" by (rule f_Equiv_Exec_expand_shrink_equiv_state_set[of k1 k2 equiv_states localState1 _ localState2 _ UNIV input_fun1 input_fun2 output_fun1 output_fun2], simp+) lemma i_Equiv_Exec_Stream_expand_shrink_equiv_state_set[rule_format]: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); \input0 m. \set input0 \ A; m \ A\ \ Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input0 \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input0 \\<^sub>f k2) c2); range input \ A \ \ equiv_states (localState1 ((i_Exec_Comp_Stream trans_fun1 ((input_fun1 \ input) \\<^sub>i k1) c1 \
\<^bsub>il\<^esub> k1) i)) (localState2 ((i_Exec_Comp_Stream trans_fun2 ((input_fun2 \ input) \\<^sub>i k2) c2 \
\<^bsub>il\<^esub> k2) i))" apply (simp add: i_shrink_last_nth i_Exec_Stream_nth i_expand_i_take_mod) apply (rule f_Equiv_Exec_expand_shrink_equiv_state_set[of k1 k2 equiv_states localState1 c1 localState2 c2 A input_fun1 input_fun2 output_fun1 output_fun2]) apply (simp add: subset_trans[OF set_i_take_subset])+ done lemma i_Equiv_Exec_Stream_expand_shrink_equiv_state: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); \input0 m. Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input0 \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input0 \\<^sub>f k2) c2) \ \ equiv_states (localState1 ((i_Exec_Comp_Stream trans_fun1 ((input_fun1 \ input) \\<^sub>i k1) c1 \
\<^bsub>il\<^esub> k1) i)) (localState2 ((i_Exec_Comp_Stream trans_fun2 ((input_fun2 \ input) \\<^sub>i k2) c2 \
\<^bsub>il\<^esub> k2) i))" by (rule i_Equiv_Exec_Stream_expand_shrink_equiv_state_set[of k1 k2 equiv_states localState1 c1 localState2 c2 UNIV input_fun1 input_fun2 output_fun1 output_fun2], simp+) lemma f_Equiv_Exec_Stream_expand_shrink_output_set_eq: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); \input0 m. \ set input0 \ A; m \ A \ \ Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input0 \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input0 \\<^sub>f k2) c2); set input \ A \ \ (map output_fun1 ( f_Exec_Comp_Stream trans_fun1 (map input_fun1 input \\<^sub>f k1) c1)) \
\<^sub>f k1 = (map output_fun2 ( f_Exec_Comp_Stream trans_fun2 (map input_fun2 input \\<^sub>f k2) c2)) \
\<^sub>f k2" apply (subst list_eq_iff) apply (clarsimp simp: f_shrink_length) apply (simp del: last.simps f_Exec_Stream_Cons add: f_shrink_nth take_map drop_map f_Exec_Stream_take f_Exec_Stream_drop f_expand_take_mod f_expand_drop_mod take_first) apply (frule_tac n=i in subset_trans[OF set_take_subset, rule_format]) apply (unfold atomize_all atomize_imp, intro allI impI) apply (frule_tac x="take i input" in spec) apply (drule_tac x="input ! i" in spec) apply (erule impE, assumption) apply (erule impE) apply (blast intro: nth_mem) apply (simp del: last.simps f_Exec_Stream_Cons) apply (rule Equiv_Exec_output_eqI[of equiv_states localState1 _ localState2 _ _ input_fun1 input_fun2]) apply (case_tac i, simp) apply (simp add: take_map[symmetric] f_Exec_Stream_expand_shrink_last_nth_eq_f_Exec_Comp[symmetric]) apply (frule Suc_lessD) apply (simp add: f_Equiv_Exec_Stream_expand_shrink_equiv_state_set[of k1 k2 equiv_states localState1 _ localState2 _ A input_fun1 input_fun2 output_fun1 output_fun2]) apply simp done lemma f_Equiv_Exec_Stream_expand_shrink_output_eq: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); \input0 m. Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input0 \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input0 \\<^sub>f k2) c2) \ \ (map output_fun1 ( f_Exec_Comp_Stream trans_fun1 (map input_fun1 input \\<^sub>f k1) c1)) \
\<^sub>f k1 = (map output_fun2 ( f_Exec_Comp_Stream trans_fun2 (map input_fun2 input \\<^sub>f k2) c2)) \
\<^sub>f k2" by (rule f_Equiv_Exec_Stream_expand_shrink_output_set_eq[of k1 k2 equiv_states localState1 _ localState2 _ UNIV], simp+) lemma i_Equiv_Exec_Stream_expand_shrink_output_set_eq: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); \input0 m. \ set input0 \ A; m \ A \ \ Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input0 \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input0 \\<^sub>f k2) c2); range input \ A \ \ (output_fun1 \ i_Exec_Comp_Stream trans_fun1 ((input_fun1 \ input) \\<^sub>i k1) c1) \
\<^sub>i k1 = (output_fun2 \ i_Exec_Comp_Stream trans_fun2 ((input_fun2 \ input) \\<^sub>i k2) c2) \
\<^sub>i k2" apply (clarsimp simp: ilist_eq_iff, rename_tac i) apply (simp del: last.simps f_Exec_Stream_Cons add: i_shrink_nth i_Exec_Stream_take i_Exec_Stream_drop i_expand_i_take_mod i_expand_i_drop_mod i_take_first map_one f_expand_one) apply (rule Equiv_Exec_output_eqI[of equiv_states localState1 _ localState2 _ _ input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2]) apply (rule f_Equiv_Exec_expand_shrink_equiv_state_set[of k1 k2 equiv_states localState1 _ localState2 _ A input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2]) apply (simp add: subset_trans[OF set_i_take_subset] subsetD[OF _ rangeI])+ done lemma i_Equiv_Exec_Stream_expand_shrink_output_eq: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); \input0 m. Equiv_Exec m equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 (f_Exec_Comp trans_fun1 (map input_fun1 input0 \\<^sub>f k1) c1) (f_Exec_Comp trans_fun2 (map input_fun2 input0 \\<^sub>f k2) c2) \ \ (output_fun1 \ i_Exec_Comp_Stream trans_fun1 ((input_fun1 \ input) \\<^sub>i k1) c1) \
\<^sub>i k1 = (output_fun2 \ i_Exec_Comp_Stream trans_fun2 ((input_fun2 \ input) \\<^sub>i k2) c2) \
\<^sub>i k2" apply (rule i_Equiv_Exec_Stream_expand_shrink_output_set_eq[of k1 k2 equiv_states localState1 c1 localState2 c2 UNIV input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2]) apply simp+ done lemma f_Equiv_Exec_Stream_Acc_LocalState_set: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); Equiv_Exec_stable_set A equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2; \ \\equiv_states\ relation implies equivalent executions\ \ \not only at the beginning but also after processing an input\ set input \ A; i < length input \ \ equiv_states (f_Exec_Comp_Stream_Acc_LocalState k1 localState1 trans_fun1 (map input_fun1 input) c1 ! i) (f_Exec_Comp_Stream_Acc_LocalState k2 localState2 trans_fun2 (map input_fun2 input) c2 ! i)" apply (unfold f_Exec_Comp_Stream_Acc_LocalState_def Equiv_Exec_stable_set_def) apply (simp add: f_shrink_last_map f_shrink_last_length) apply (rule f_Equiv_Exec_Stream_expand_shrink_equiv_state_set[of k1 k2 equiv_states localState1 c1 localState2 c2 A input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 input, rule_format]) apply simp+ done lemma f_Equiv_Exec_Stream_Acc_LocalState: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); Equiv_Exec_stable equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2; \ \\equiv_states\ relation implies equivalent executions\ \ \not only at the beginning but also after processing an input\ i < length input \ \ equiv_states (f_Exec_Comp_Stream_Acc_LocalState k1 localState1 trans_fun1 (map input_fun1 input) c1 ! i) (f_Exec_Comp_Stream_Acc_LocalState k2 localState2 trans_fun2 (map input_fun2 input) c2 ! i)" apply (rule f_Equiv_Exec_Stream_Acc_LocalState_set[where A=UNIV]) apply (simp add: Equiv_Exec_stable_set_UNIV)+ done lemma f_Equiv_Exec_Stream_Acc_Output_set_eq: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); Equiv_Exec_stable_set A equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2; set input \ A \ \ f_Exec_Comp_Stream_Acc_Output k1 output_fun1 trans_fun1 (map input_fun1 input) c1 = f_Exec_Comp_Stream_Acc_Output k2 output_fun2 trans_fun2 (map input_fun2 input) c2" apply (unfold f_Exec_Comp_Stream_Acc_Output_def Equiv_Exec_stable_set_def) apply (rule f_Equiv_Exec_Stream_expand_shrink_output_set_eq[of k1 k2 equiv_states localState1 c1 localState2 c2 A input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 input]) apply simp+ done lemma f_Equiv_Exec_Stream_Acc_Output_eq: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); Equiv_Exec_stable equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 \ \ f_Exec_Comp_Stream_Acc_Output k1 output_fun1 trans_fun1 (map input_fun1 input) c1 = f_Exec_Comp_Stream_Acc_Output k2 output_fun2 trans_fun2 (map input_fun2 input) c2" apply (rule f_Equiv_Exec_Stream_Acc_Output_set_eq[of k1 k2 equiv_states localState1 c1 localState2 c2 UNIV]) apply (simp add: Equiv_Exec_stable_set_UNIV)+ done lemma i_Equiv_Exec_Stream_Acc_LocalState_set: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); Equiv_Exec_stable_set A equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2; range input \ A \ \ equiv_states (i_Exec_Comp_Stream_Acc_LocalState k1 localState1 trans_fun1 (input_fun1 \ input) c1 i) (i_Exec_Comp_Stream_Acc_LocalState k2 localState2 trans_fun2 (input_fun2 \ input) c2 i)" apply (simp add: i_Exec_Stream_Acc_LocalState_nth_f_nth) apply (rule f_Equiv_Exec_Stream_Acc_LocalState_set) apply (simp add: subset_trans[OF set_i_take_subset])+ done lemma i_Equiv_Exec_Stream_Acc_LocalState: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); Equiv_Exec_stable equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 \ \ equiv_states (i_Exec_Comp_Stream_Acc_LocalState k1 localState1 trans_fun1 (input_fun1 \ input) c1 i) (i_Exec_Comp_Stream_Acc_LocalState k2 localState2 trans_fun2 (input_fun2 \ input) c2 i)" apply (rule i_Equiv_Exec_Stream_Acc_LocalState_set[where A=UNIV]) apply (simp add: Equiv_Exec_stable_set_UNIV)+ done lemma i_Equiv_Exec_Stream_Acc_Output_set_eq: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); Equiv_Exec_stable_set A equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2; range input \ A \ \ i_Exec_Comp_Stream_Acc_Output k1 output_fun1 trans_fun1 (input_fun1 \ input) c1 = i_Exec_Comp_Stream_Acc_Output k2 output_fun2 trans_fun2 (input_fun2 \ input) c2" apply (clarsimp simp: ilist_eq_iff i_Exec_Stream_Acc_Output_nth_f_nth, rename_tac i) apply (drule_tac n="Suc i" in subset_trans[OF set_i_take_subset, rule_format]) apply (simp add: f_Equiv_Exec_Stream_Acc_Output_set_eq[where equiv_states=equiv_states]) done lemma i_Equiv_Exec_Stream_Acc_Output_eq: " \ 0 < k1; 0 < k2; equiv_states (localState1 c1) (localState2 c2); Equiv_Exec_stable equiv_states localState1 localState2 input_fun1 input_fun2 output_fun1 output_fun2 trans_fun1 trans_fun2 k1 k2 c1 c2 \ \ i_Exec_Comp_Stream_Acc_Output k1 output_fun1 trans_fun1 (input_fun1 \ input) c1 = i_Exec_Comp_Stream_Acc_Output k2 output_fun2 trans_fun2 (input_fun2 \ input) c2" apply (rule i_Equiv_Exec_Stream_Acc_Output_set_eq[of k1 k2 equiv_states localState1 c1 localState2 c2 UNIV]) apply (simp add: Equiv_Exec_stable_set_UNIV)+ done subsubsection \Idle states and accelerated execution\ lemma f_Exec_Stream_Acc_LocalState__State_Idle_nth[rule_format]: " \c i. \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; \n\i. State_Idle localState output_fun trans_fun ( f_Exec_Comp_Stream_Acc_LocalState l localState trans_fun xs c ! n); i < length xs \ \ f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c ! i = f_Exec_Comp_Stream_Acc_LocalState l localState trans_fun xs c ! i" apply (frule length_greater_0_conv[THEN iffD1, OF gr_implies_gr0]) apply (simp only: f_Exec_Stream_Acc_LocalState_nth take_Suc_conv_app_nth) apply (simp only: f_expand_snoc f_Exec_append) apply (rule_tac s="\\<^bsup>l - Suc 0\<^esup> @ \\<^bsup>k-l\<^esup>" and t="\\<^bsup>k - Suc 0\<^esup>" in subst) apply (simp add: replicate_le_diff2) apply (subst append_Cons[symmetric]) apply (induct xs, simp) apply (case_tac i) apply (simp add: f_Exec_Stream_Acc_LocalState_Cons f_Exec_State_Idle_append_replicate_NoMsg_state) apply (rename_tac n) apply (drule_tac x="f_Exec_Comp trans_fun (a # \\<^bsup>l - Suc 0\<^esup>) c" in meta_spec) apply (drule_tac x=n in meta_spec) apply (simp del: f_Exec_Cons) apply (frule length_greater_imp_not_empty) apply (drule meta_mp) apply (simp add: f_Exec_Stream_Acc_LocalState_nth f_Exec_append) apply (simp add: append_Cons[symmetric] f_expand_Cons f_Exec_append del: append_Cons) apply (subgoal_tac " localState (f_Exec_Comp trans_fun (a # NoMsg\<^bsup>k - Suc 0\<^esup>) c) = localState (f_Exec_Comp trans_fun (a # NoMsg\<^bsup>l - Suc 0\<^esup>) c)") prefer 2 apply (drule_tac x=0 in spec) apply (simp add: f_Exec_Stream_Acc_LocalState_Cons) apply (subst replicate_le_diff2[OF Suc_leI, symmetric], assumption+) apply (simp add: append_Cons[symmetric] f_Exec_append del: append_Cons) apply (rule f_Exec_State_Idle_replicate_NoMsg_state, assumption) apply (case_tac "n = 0") apply (frule_tac ?c1.0="f_Exec_Comp trans_fun (a # NoMsg\<^bsup>k - Suc 0\<^esup>) c" and xs = "xs ! 0 # NoMsg\<^bsup>l - Suc 0\<^esup>" in f_Exec_Equal_State) apply simp+ apply (frule_tac ?c1.0="f_Exec_Comp trans_fun (a # NoMsg\<^bsup>k - Suc 0\<^esup>) c" and xs = "xs \ n \\<^sub>f k" in f_Exec_Equal_State) apply (simp add: f_expand_not_empty_conv)+ done corollary f_Exec_Stream_Acc_LocalState__State_Idle_eq[rule_format]: " \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; \n \ f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c = f_Exec_Comp_Stream_Acc_LocalState l localState trans_fun xs c" apply (clarsimp simp: list_eq_iff) apply (rule f_Exec_Stream_Acc_LocalState__State_Idle_nth) apply simp_all apply (drule_tac x=n in spec) apply simp done lemma i_Exec_Stream_Acc_LocalState__State_Idle_nth[rule_format]: " \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; \n\i. State_Idle localState output_fun trans_fun ( i_Exec_Comp_Stream_Acc_LocalState l localState trans_fun input c n) \ \ i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input c i = i_Exec_Comp_Stream_Acc_LocalState l localState trans_fun input c i" apply (simp only: f_Exec_Stream_Acc_LocalState_nth_eq_i_nth[of _ _ "Suc i", symmetric]) apply (rule f_Exec_Stream_Acc_LocalState__State_Idle_nth) apply simp_all apply (drule_tac x=n in spec) apply (simp add: f_Exec_Stream_Acc_LocalState_nth_eq_i_nth) done corollary i_Exec_Stream_Acc_LocalState__State_Idle_eq[rule_format]: " \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; \n. State_Idle localState output_fun trans_fun ( i_Exec_Comp_Stream_Acc_LocalState l localState trans_fun input c n) \ \ i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun input c = i_Exec_Comp_Stream_Acc_LocalState l localState trans_fun input c" apply (clarsimp simp: ilist_eq_iff) apply (rule i_Exec_Stream_Acc_LocalState__State_Idle_nth) apply simp_all apply (drule_tac x=n in spec) apply simp done lemma f_Exec_Stream_Acc_Output__State_Idle_nth[rule_format]: " \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; \n\i. State_Idle localState output_fun trans_fun ( f_Exec_Comp_Stream_Acc_LocalState l localState trans_fun xs c ! n); i < length xs \ \ f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c ! i = f_Exec_Comp_Stream_Acc_Output l output_fun trans_fun xs c ! i" apply (drule order_le_less[THEN iffD1], erule disjE) prefer 2 apply simp apply (frule zero_less_diff[of k l, THEN iffD2]) apply (frule length_greater_imp_not_empty) apply (simp add: f_Exec_Stream_Acc_Output_nth del: f_Exec_Stream_Cons) apply (subst replicate_le_diff2[OF Suc_leI, symmetric]) apply (simp del: f_Exec_Stream_Cons)+ apply (subst append_Cons[symmetric]) apply (case_tac i) apply (drule_tac x=0 in spec) apply (simp add: f_Exec_Stream_Acc_LocalState_nth take_first f_expand_one del: last.simps f_Exec_Cons f_Exec_Stream_Cons append_Cons replicate.simps) apply (simp only: f_Exec_Stream_append map_append last_message_append) apply (rule if_P') apply (clarsimp simp: last_message_NoMsg_conv f_Exec_Stream_nth min_eqL simp del: last.simps f_Exec_Comp.simps append_Cons replicate.simps) apply (rule f_Exec_State_Idle_replicate_NoMsg_gr0_output) apply (simp del: last.simps f_Exec_Comp_Stream.simps append_Cons)+ apply (rename_tac n) apply (simp only: f_Exec_Stream_append map_append last_message_append) apply (subgoal_tac " localState (f_Exec_Comp trans_fun (xs \ Suc n \\<^sub>f k) c) = localState (f_Exec_Comp trans_fun (xs \ Suc n \\<^sub>f l) c)") prefer 2 apply (simp add: f_Exec_Stream_Acc_LocalState_nth[symmetric]) apply (rule f_Exec_Stream_Acc_LocalState__State_Idle_nth) apply simp+ apply (rename_tac n, drule_tac x=n in spec, simp) apply simp apply (rule if_P') apply (simp add: last_message_NoMsg_conv f_Exec_Stream_nth min_eqL del: f_Exec_Comp.simps replicate.simps) apply (clarify, rename_tac j) apply (frule_tac x="Suc n" in spec) apply (simp only: f_Exec_Stream_Acc_LocalState_nth) apply (rule_tac ?c1.0="f_Exec_Comp trans_fun (xs \ Suc n \\<^sub>f l) c" and ?c2.0="f_Exec_Comp trans_fun (xs \ Suc n \\<^sub>f k) c" in subst[OF f_Exec_Equal_State, rule_format]) apply (simp del: f_Exec_Comp.simps replicate.simps)+ apply (simp only: take_Suc_conv_app_nth f_expand_snoc f_Exec_append) apply (rule f_Exec_State_Idle_replicate_NoMsg_gr0_output, assumption) apply simp apply (rule arg_cong[where f="\x. last_message (map output_fun x)"]) apply (rule f_Exec_Stream_Equal_State, assumption+) done lemma f_Exec_Stream_Acc_Output__State_Idle_eq[rule_format]: " \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; \n \ f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c = f_Exec_Comp_Stream_Acc_Output l output_fun trans_fun xs c" apply (clarsimp simp: list_eq_iff) apply (rule f_Exec_Stream_Acc_Output__State_Idle_nth) apply simp_all apply (drule_tac x=n in spec) apply simp done lemma i_Exec_Stream_Acc_Output__State_Idle_nth[rule_format]: " \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; \n\i. State_Idle localState output_fun trans_fun ( i_Exec_Comp_Stream_Acc_LocalState l localState trans_fun input c n) \ \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c i = i_Exec_Comp_Stream_Acc_Output l output_fun trans_fun input c i" apply (simp only: i_Exec_Stream_Acc_Output_nth_f_nth) apply (rule f_Exec_Stream_Acc_Output__State_Idle_nth) apply simp_all apply (drule_tac x=n in spec) apply (simp add: f_Exec_Stream_Acc_LocalState_nth_eq_i_nth) done lemma i_Exec_Stream_Acc_Output__State_Idle_eq[rule_format]: " \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; \n. State_Idle localState output_fun trans_fun ( i_Exec_Comp_Stream_Acc_LocalState l localState trans_fun input c n) \ \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c = i_Exec_Comp_Stream_Acc_Output l output_fun trans_fun input c" apply (clarsimp simp: ilist_eq_iff) apply (rule i_Exec_Stream_Acc_Output__State_Idle_nth) apply simp_all apply (drule_tac x=n in spec) apply simp done text \ When a certain number @{term l} of steps suffices to reach an idle state from any other idle state, than for any acceleration factor @{term "k \ l"} the accelerated processing of every input message will be finished in an idle state.\ lemma f_Exec_Stream_Acc_LocalState__State_Idle_all[rule_format]: " \c xs. \ 0 < l; l \ k; State_Idle localState output_fun trans_fun (localState c); \c m. State_Idle localState output_fun trans_fun (localState c) \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun (m # \\<^bsup>l - Suc 0\<^esup>) c)); i < length xs \ \ State_Idle localState output_fun trans_fun ( f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c ! i)" apply (frule length_greater_imp_not_empty) apply (subgoal_tac " State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun (hd xs # NoMsg\<^bsup>k - Suc 0\<^esup>) c))") prefer 2 apply (drule_tac x=c in spec, drule_tac x="hd xs" in spec) apply (rule subst[OF replicate_le_diff2[OF Suc_leI], of 0 l k], assumption+) apply (simp add: f_Exec_append f_Exec_State_Idle_replicate_NoMsg_state) apply (induct i) apply (simp add: f_Exec_Stream_Acc_LocalState_nth take_first hd_eq_first) apply (drule_tac x="f_Exec_Comp trans_fun (hd xs # NoMsg\<^bsup>k - Suc 0\<^esup>) c" in meta_spec) apply (drule_tac x="tl xs" in meta_spec) apply (subgoal_tac "i < length (tl xs) \ tl xs \ []", elim conjE) prefer 2 apply (simp add: length_greater_0_conv[symmetric] del: length_greater_0_conv) apply (simp add: f_Exec_Stream_Acc_LocalState_nth) apply (rule_tac n="Suc i" in ssubst[OF take_Suc, rule_format], assumption) apply (simp add: append_Cons[symmetric] f_Exec_append del: append_Cons) apply (drule meta_mp) apply (drule_tac x="f_Exec_Comp trans_fun (hd xs # NoMsg\<^bsup>k - Suc 0\<^esup>) c" in spec) apply (drule mp, simp) apply (drule_tac x="hd (tl xs)" in spec) apply (subst replicate_le_diff2[OF Suc_leI, of 0 l k, symmetric], simp+) apply (simp add: f_Exec_append f_Exec_State_Idle_replicate_NoMsg_state) apply (simp add: f_Exec_Stream_Acc_LocalState_nth) done lemma i_Exec_Stream_Acc_LocalState__State_Idle_all[rule_format]: " \ 0 < l; l \ k; State_Idle localState output_fun trans_fun (localState c); \c m. State_Idle localState output_fun trans_fun (localState c) \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun (m # \\<^bsup>l - Suc 0\<^esup>) c)) \ \ State_Idle localState output_fun trans_fun ( i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c i)" apply (simp only: i_Exec_Stream_Acc_LocalState_nth_f_nth) apply (rule f_Exec_Stream_Acc_LocalState__State_Idle_all) apply simp_all apply (rename_tac c' m, drule_tac x=c' in spec) apply simp done lemma f_Exec_Stream_Acc_Output__State_Idle_all_imp_eq[rule_format]: " \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; State_Idle localState output_fun trans_fun (localState c); \c m. State_Idle localState output_fun trans_fun (localState c) \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun (m # \\<^bsup>l - Suc 0\<^esup>) c)) \ \ f_Exec_Comp_Stream_Acc_Output k output_fun trans_fun xs c = f_Exec_Comp_Stream_Acc_Output l output_fun trans_fun xs c" apply (rule f_Exec_Stream_Acc_Output__State_Idle_eq, assumption+) apply (simp add: f_Exec_Stream_Acc_LocalState__State_Idle_all) done lemma i_Exec_Stream_Acc_Output__State_Idle_all_imp_eq[rule_format]: " \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; State_Idle localState output_fun trans_fun (localState c); \c m. State_Idle localState output_fun trans_fun (localState c) \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun (m # \\<^bsup>l - Suc 0\<^esup>) c)) \ \ i_Exec_Comp_Stream_Acc_Output k output_fun trans_fun input c = i_Exec_Comp_Stream_Acc_Output l output_fun trans_fun input c" apply (rule i_Exec_Stream_Acc_Output__State_Idle_eq, assumption+) apply (simp add: i_Exec_Stream_Acc_LocalState__State_Idle_all) done lemma f_Exec_Stream_Acc_LocalState__State_Idle_all_imp_eq[rule_format]: " \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; State_Idle localState output_fun trans_fun (localState c); \c m. State_Idle localState output_fun trans_fun (localState c) \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun (m # \\<^bsup>l - Suc 0\<^esup>) c)) \ \ f_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c = f_Exec_Comp_Stream_Acc_LocalState l localState trans_fun xs c" apply (rule f_Exec_Stream_Acc_LocalState__State_Idle_eq, assumption+) apply (rule f_Exec_Stream_Acc_LocalState__State_Idle_all) apply simp+ done lemma i_Exec_Stream_Acc_LocalState__State_Idle_all_imp_eq[rule_format]: " \ 0 < l; l \ k; Exec_Equal_State localState trans_fun; State_Idle localState output_fun trans_fun (localState c); \c m. State_Idle localState output_fun trans_fun (localState c) \ State_Idle localState output_fun trans_fun ( localState (f_Exec_Comp trans_fun (m # \\<^bsup>l - Suc 0\<^esup>) c)) \ \ i_Exec_Comp_Stream_Acc_LocalState k localState trans_fun xs c = i_Exec_Comp_Stream_Acc_LocalState l localState trans_fun xs c" apply (rule i_Exec_Stream_Acc_LocalState__State_Idle_eq, assumption+) apply (rule i_Exec_Stream_Acc_LocalState__State_Idle_all) apply simp+ done text \Converting inputs\ lemma f_Exec_input_map: "\c. f_Exec_Comp trans_fun (map f xs) c = f_Exec_Comp (trans_fun \ f) xs c" by (induct xs, simp+) lemma f_Exec_Stream_input_map: " f_Exec_Comp_Stream trans_fun (map f xs) c = f_Exec_Comp_Stream (trans_fun \ f) xs c" by (simp add: list_eq_iff f_Exec_Stream_nth take_map f_Exec_input_map) lemma i_Exec_Stream_input_map: " i_Exec_Comp_Stream trans_fun (f \ input) c = i_Exec_Comp_Stream (trans_fun \ f) input c" by (simp add: ilist_eq_iff i_Exec_Stream_nth f_Exec_input_map) end diff --git a/thys/Category2/Functors.thy b/thys/Category2/Functors.thy --- a/thys/Category2/Functors.thy +++ b/thys/Category2/Functors.thy @@ -1,507 +1,507 @@ (* Author: Alexander Katovsky *) section "Functor" theory Functors imports Category begin record ('o1, 'o2, 'm1, 'm2, 'a, 'b) Functor = CatDom :: "('o1,'m1,'a)Category_scheme" CatCod :: "('o2,'m2,'b)Category_scheme" MapM :: "'m1 \ 'm2" abbreviation FunctorMorApp :: "('o1, 'o2, 'm1, 'm2, 'a1, 'a2, 'a) Functor_scheme \ 'm1 \ 'm2" (infixr "##" 70) where "FunctorMorApp F m \ (MapM F) m" definition MapO :: "('o1, 'o2, 'm1, 'm2, 'a1, 'a2, 'a) Functor_scheme \ 'o1 \ 'o2" where "MapO F X \ THE Y . Y \ Obj(CatCod F) \ F ## (Id (CatDom F) X) = Id (CatCod F) Y" abbreviation FunctorObjApp :: "('o1, 'o2, 'm1, 'm2, 'a1, 'a2, 'a) Functor_scheme \ 'o1 \ 'o2" (infixr "@@" 70) where "FunctorObjApp F X \ (MapO F X)" locale PreFunctor = fixes F :: "('o1, 'o2, 'm1, 'm2, 'a1, 'a2, 'a) Functor_scheme" (structure) assumes FunctorComp: "f \>\<^bsub>CatDom F\<^esub> g \ F ## (f ;;\<^bsub>CatDom F\<^esub> g) = (F ## f) ;;\<^bsub>CatCod F\<^esub> (F ## g)" and FunctorId: "X \ obj\<^bsub>CatDom F\<^esub> \ \ Y \ obj\<^bsub>CatCod F\<^esub> . F ## (id\<^bsub>CatDom F\<^esub> X) = id\<^bsub>CatCod F\<^esub> Y" and CatDom[simp]: "Category(CatDom F)" and CatCod[simp]: "Category(CatCod F)" locale FunctorM = PreFunctor + assumes FunctorCompM: "f maps\<^bsub>CatDom F\<^esub> X to Y \ (F ## f) maps\<^bsub>CatCod F\<^esub> (F @@ X) to (F @@ Y)" locale FunctorExt = fixes F :: "('o1, 'o2, 'm1, 'm2, 'a1, 'a2, 'a) Functor_scheme" (structure) assumes FunctorMapExt: "(MapM F) \ extensional (Mor (CatDom F))" locale Functor = FunctorM + FunctorExt definition MakeFtor :: "('o1, 'o2, 'm1, 'm2, 'a, 'b,'r) Functor_scheme \ ('o1, 'o2, 'm1, 'm2, 'a, 'b,'r) Functor_scheme" where "MakeFtor F \ \ CatDom = CatDom F , CatCod = CatCod F , MapM = restrict (MapM F) (Mor (CatDom F)) , \ = Functor.more F \" lemma PreFunctorFunctor[simp]: "Functor F \ PreFunctor F" by (simp add: Functor_def FunctorM_def) lemmas functor_simps = PreFunctor.FunctorComp PreFunctor.FunctorId definition functor_abbrev ("Ftor _ : _ \ _" [81]) where "Ftor F : A \ B \ (Functor F) \ (CatDom F = A) \ (CatCod F = B)" lemma functor_abbrevE[elim]: "\Ftor F : A \ B ; \(Functor F) ; (CatDom F = A) ; (CatCod F = B)\ \ R\ \ R" by (auto simp add: functor_abbrev_def) definition functor_comp_def ("_ \>;;; _" [81]) where "functor_comp_def F G \ (Functor F) \ (Functor G) \ (CatDom G = CatCod F)" lemma functor_comp_def[elim]: "\F \>;;; G ; \Functor F ; Functor G ; CatDom G = CatCod F\ \ R\ \ R" by (auto simp add: functor_comp_def_def) lemma (in Functor) FunctorMapsTo: assumes "f \ mor\<^bsub>CatDom F\<^esub>" shows "F ## f maps\<^bsub>CatCod F\<^esub> (F @@ (dom\<^bsub>CatDom F\<^esub> f)) to (F @@ (cod\<^bsub>CatDom F\<^esub> f))" proof- have "f maps\<^bsub>CatDom F\<^esub> (dom\<^bsub>CatDom F\<^esub> f) to (cod\<^bsub>CatDom F\<^esub> f)" using assms by auto thus ?thesis by (simp add: FunctorCompM[of f "dom\<^bsub>CatDom F\<^esub> f" "cod\<^bsub>CatDom F\<^esub> f"]) qed lemma (in Functor) FunctorCodDom: assumes "f \ mor\<^bsub>CatDom F\<^esub>" shows "dom\<^bsub>CatCod F\<^esub>(F ## f) = F @@ (dom\<^bsub>CatDom F\<^esub> f)" and "cod\<^bsub>CatCod F\<^esub>(F ## f) = F @@ (cod\<^bsub>CatDom F\<^esub> f)" proof- have "F ## f maps\<^bsub>CatCod F\<^esub> (F @@ (dom\<^bsub>CatDom F\<^esub> f)) to (F @@ (cod\<^bsub>CatDom F\<^esub> f))" using assms by (simp add: FunctorMapsTo) thus "dom\<^bsub>CatCod F\<^esub>(F ## f) = F @@ (dom\<^bsub>CatDom F\<^esub> f)" and "cod\<^bsub>CatCod F\<^esub>(F ## f) = F @@ (cod\<^bsub>CatDom F\<^esub> f)" by auto qed lemma (in Functor) FunctorCompPreserved: "f \ mor\<^bsub>CatDom F\<^esub> \ F ## f \ mor\<^bsub>CatCod F\<^esub>" by (auto dest:FunctorMapsTo) lemma (in Functor) FunctorCompDef: assumes "f \>\<^bsub>CatDom F\<^esub> g" shows "(F ## f) \>\<^bsub>CatCod F\<^esub> (F ## g)" proof(auto simp add: CompDefined_def) show "F ## f \ mor\<^bsub>CatCod F\<^esub>" and "F ## g \ mor\<^bsub>CatCod F\<^esub>" using assms by (auto simp add: FunctorCompPreserved) have "f \ mor\<^bsub>CatDom F\<^esub>" and "g \ mor\<^bsub>CatDom F\<^esub>" using assms by auto hence a: "cod\<^bsub>CatCod F\<^esub> (F ## f) = F @@ (cod\<^bsub>CatDom F\<^esub> f)" and b: "dom\<^bsub>CatCod F\<^esub>(F ## g) = F @@ (dom\<^bsub>CatDom F\<^esub> g)" by (simp add: FunctorCodDom)+ have "cod\<^bsub>CatCod F\<^esub> (F ## f) = F @@ (dom\<^bsub>CatDom F\<^esub> g)" using assms a by auto also have "... = dom\<^bsub>CatCod F\<^esub> (F ## g)" using b by simp finally show "cod\<^bsub>CatCod F\<^esub> (F ## f) = dom\<^bsub>CatCod F\<^esub> (F ## g)" . qed lemma FunctorComp: "\Ftor F : A \ B ; f \>\<^bsub>A\<^esub> g\ \ F ## (f ;;\<^bsub>A\<^esub> g) = (F ## f) ;;\<^bsub>B\<^esub> (F ## g)" by (auto simp add: PreFunctor.FunctorComp) lemma FunctorCompDef: "\Ftor F : A \ B ; f \>\<^bsub>A\<^esub> g\ \ (F ## f) \>\<^bsub>B\<^esub> (F ## g)" by (auto simp add: Functor.FunctorCompDef) lemma FunctorMapsTo: assumes "Ftor F : A \ B" and "f maps\<^bsub>A\<^esub> X to Y" shows "(F ## f) maps\<^bsub>B\<^esub> (F @@ X) to (F @@ Y)" proof- have b: "CatCod F = B" and a: "CatDom F = A" and ff: "Functor F" using assms by auto have df: "(dom\<^bsub>CatDom F\<^esub> f) = X" and cf: "(cod\<^bsub>CatDom F\<^esub> f) = Y" using a assms by auto have "f \ mor\<^bsub>CatDom F\<^esub>" using assms by auto hence "F ## f maps\<^bsub>CatCod F\<^esub> (F @@ (dom\<^bsub>CatDom F\<^esub> f)) to (F @@ (cod\<^bsub>CatDom F\<^esub> f))" using ff by (simp add: Functor.FunctorMapsTo) thus ?thesis using df cf b by simp qed lemma (in PreFunctor) FunctorId2: assumes "X \ obj\<^bsub>CatDom F\<^esub>" shows "F @@ X \ obj\<^bsub>CatCod F\<^esub> \ F ## (id\<^bsub>CatDom F\<^esub> X) = id\<^bsub>CatCod F\<^esub> (F @@ X)" proof- let ?Q = "\ E Y . Y \ obj\<^bsub>CatCod F\<^esub> \ E = id\<^bsub>CatCod F\<^esub> Y" let ?P = "?Q (F ## (id\<^bsub>CatDom F\<^esub> X))" from assms FunctorId obtain Y where "?P Y" by auto moreover { fix y e z have "\?Q e y ; ?Q e z\ \ y = z" by (auto intro: Category.IdInj[of "CatCod F" y z]) } ultimately have "\! Z . ?P Z" by auto hence "?P (THE Y . ?P Y)" by (rule theI') thus ?thesis by (auto simp add: MapO_def) qed lemma FunctorId: assumes "Ftor F : C \ D" and "X \ Obj C" shows "F ## (Id C X) = Id D (F @@ X)" proof- have "CatDom F = C" and "CatCod F = D" and "PreFunctor F" using assms by auto thus ?thesis using assms PreFunctor.FunctorId2[of F X] by simp qed lemma (in Functor) DomFunctor: "f \ mor\<^bsub>CatDom F\<^esub> \ dom\<^bsub>CatCod F\<^esub> (F ## f) = F @@ (dom\<^bsub>CatDom F\<^esub> f)" by (simp add: FunctorCodDom) lemma (in Functor) CodFunctor: "f \ mor\<^bsub>CatDom F\<^esub> \ cod\<^bsub>CatCod F\<^esub> (F ## f) = F @@ (cod\<^bsub>CatDom F\<^esub> f)" by (simp add: FunctorCodDom) lemma (in Functor) FunctorId3Dom: assumes "f \ mor\<^bsub>CatDom F\<^esub>" shows "F ## (id\<^bsub>CatDom F\<^esub> (dom\<^bsub>CatDom F\<^esub> f)) = id\<^bsub>CatCod F\<^esub> (dom\<^bsub>CatCod F\<^esub> (F ## f))" proof- have "(dom\<^bsub>CatDom F\<^esub> f) \ obj\<^bsub>CatDom F\<^esub>" using assms by (simp add: Category.Cdom) hence "F ## (id\<^bsub>CatDom F\<^esub> (dom\<^bsub>CatDom F\<^esub> f)) = id\<^bsub>CatCod F\<^esub> (F @@ (dom\<^bsub>CatDom F\<^esub> f))" by (simp add: FunctorId2) also have "... = id\<^bsub>CatCod F\<^esub> (dom\<^bsub>CatCod F\<^esub> (F ## f))" using assms by (simp add: DomFunctor) finally show ?thesis by simp qed lemma (in Functor) FunctorId3Cod: assumes "f \ mor\<^bsub>CatDom F\<^esub>" shows "F ## (id\<^bsub>CatDom F\<^esub> (cod\<^bsub>CatDom F\<^esub> f)) = id\<^bsub>CatCod F\<^esub> (cod\<^bsub>CatCod F\<^esub> (F ## f))" proof- have "(cod\<^bsub>CatDom F\<^esub> f) \ obj\<^bsub>CatDom F\<^esub>" using assms by (simp add: Category.Ccod) hence "F ## (id\<^bsub>CatDom F\<^esub> (cod\<^bsub>CatDom F\<^esub> f)) = id\<^bsub>CatCod F\<^esub> (F @@ (cod\<^bsub>CatDom F\<^esub> f))" by (simp add: FunctorId2) also have "... = id\<^bsub>CatCod F\<^esub> (cod\<^bsub>CatCod F\<^esub> (F ## f))" using assms by (simp add: CodFunctor) finally show ?thesis by simp qed lemma (in PreFunctor) FmToFo: "\X \ obj\<^bsub>CatDom F\<^esub> ; Y \ obj\<^bsub>CatCod F\<^esub> ; F ## (id\<^bsub>CatDom F\<^esub> X) = id\<^bsub>CatCod F\<^esub> Y\ \ F @@ X = Y" by (auto simp add: FunctorId2 intro: Category.IdInj[of "CatCod F" "F @@ X" Y]) lemma MakeFtorPreFtor: assumes "PreFunctor F" shows "PreFunctor (MakeFtor F)" proof- { fix X assume a: "X \ obj\<^bsub>CatDom F\<^esub>" have "id\<^bsub>CatDom F \<^esub>X \ mor\<^bsub>CatDom F\<^esub>" proof- have "Category (CatDom F)" using assms by (simp add: PreFunctor_def) hence "id\<^bsub>CatDom F \<^esub>X maps\<^bsub>CatDom F\<^esub> X to X" using a by (simp add: Category.Cidm) thus ?thesis using a by (auto) qed } thus "PreFunctor (MakeFtor F)" using assms by(auto simp add: PreFunctor_def MakeFtor_def Category.MapsToMorDomCod) qed lemma MakeFtorMor: "f \ mor\<^bsub>CatDom F\<^esub> \ MakeFtor F ## f = F ## f" by(simp add: MakeFtor_def) lemma MakeFtorObj: assumes "PreFunctor F" and "X \ obj\<^bsub>CatDom F\<^esub>" shows "MakeFtor F @@ X = F @@ X" proof- have "X \ obj\<^bsub>CatDom (MakeFtor F)\<^esub>" using assms(2) by (simp add: MakeFtor_def) moreover have "(F @@ X) \ obj\<^bsub>CatCod (MakeFtor F)\<^esub>" using assms by (simp add: PreFunctor.FunctorId2 MakeFtor_def) moreover have "MakeFtor F ## id\<^bsub>CatDom (MakeFtor F) \<^esub>X = id\<^bsub>CatCod (MakeFtor F) \<^esub>(F @@ X)" proof- have "Category (CatDom F)" using assms(1) by (simp add: PreFunctor_def) hence "id\<^bsub>CatDom F \<^esub>X maps\<^bsub>CatDom F\<^esub> X to X" using assms(2) by (auto simp add: Category.Cidm) hence "id\<^bsub>CatDom F \<^esub>X \ mor\<^bsub>CatDom F\<^esub>" by auto hence "MakeFtor F ## id\<^bsub>CatDom (MakeFtor F) \<^esub>X = F ## id\<^bsub>CatDom F \<^esub>X" by (simp add: MakeFtor_def) also have "... = id\<^bsub>CatCod F \<^esub>(F @@ X)" using assms by (simp add: PreFunctor.FunctorId2) finally show ?thesis by (simp add: MakeFtor_def) qed moreover have "PreFunctor (MakeFtor F)" using assms(1) by (simp add: MakeFtorPreFtor) ultimately show ?thesis by (simp add: PreFunctor.FmToFo) qed lemma MakeFtor: assumes "FunctorM F" shows "Functor (MakeFtor F)" proof(intro_locales) show "PreFunctor (MakeFtor F)" using assms by (simp add: MakeFtorPreFtor FunctorM_def) show "FunctorM_axioms (MakeFtor F)" proof(auto simp add: FunctorM_axioms_def) { fix f X Y assume aa: "f maps\<^bsub>CatDom (MakeFtor F)\<^esub> X to Y" show "((MakeFtor F) ## f) maps\<^bsub>CatCod (MakeFtor F)\<^esub> ((MakeFtor F) @@ X) to ((MakeFtor F) @@ Y)" proof- have "((MakeFtor F) ## f) = F ## f" using aa by (auto simp add: MakeFtor_def) moreover have "((MakeFtor F) @@ X) = F @@ X" and "((MakeFtor F) @@ Y) = F @@ Y" proof- have "Category (CatDom F)" using assms by (simp add: FunctorM_def PreFunctor_def) hence "X \ obj\<^bsub>CatDom F\<^esub>" and "Y \ obj\<^bsub>CatDom F\<^esub>" using aa by (auto simp add: Category.MapsToObj MakeFtor_def) moreover have "PreFunctor F" using assms(1) by (simp add: FunctorM_def) ultimately show "((MakeFtor F) @@ X) = F @@ X" and "((MakeFtor F) @@ Y) = F @@ Y" by (simp add: MakeFtorObj)+ qed moreover have "F ## f maps\<^bsub>CatCod F\<^esub> (F @@ X) to (F @@ Y)" using assms(1) aa by (simp add: FunctorM.FunctorCompM MakeFtor_def) ultimately show ?thesis by (simp add: MakeFtor_def) qed } qed show "FunctorExt (MakeFtor F)" by(simp add: FunctorExt_def MakeFtor_def) qed definition - IdentityFunctor' :: "('o,'m,'a) Category_scheme \ ('o,'o,'m,'m,'a,'a) Functor" ("FId' _" [70]) where + IdentityFunctor' :: "('o,'m,'a) Category_scheme \ ('o,'o,'m,'m,'a,'a) Functor" ("FId'' _" [70]) where "IdentityFunctor' C \ \CatDom = C , CatCod = C , MapM = (\ f . f) \" definition IdentityFunctor ("FId _" [70]) where "IdentityFunctor C \ MakeFtor(IdentityFunctor' C)" lemma IdFtor'PreFunctor: "Category C \ PreFunctor (FId' C)" by(auto simp add: PreFunctor_def IdentityFunctor'_def) lemma IdFtor'Obj: assumes "Category C" and "X \ obj\<^bsub>CatDom (FId' C)\<^esub>" shows "(FId' C) @@ X = X" proof- have "(FId' C) ## id\<^bsub>CatDom (FId' C)\<^esub> X = id\<^bsub>CatCod (FId' C)\<^esub> X" by(simp add: IdentityFunctor'_def) moreover have "X \ obj\<^bsub>CatCod (FId' C)\<^esub>" using assms by (simp add: IdentityFunctor'_def) ultimately show ?thesis using assms by (simp add: PreFunctor.FmToFo IdFtor'PreFunctor) qed lemma IdFtor'FtorM: assumes "Category C" shows "FunctorM (FId' C)" proof(auto simp add: FunctorM_def IdFtor'PreFunctor assms FunctorM_axioms_def) { fix f X Y assume a: "f maps\<^bsub>CatDom (FId' C)\<^esub> X to Y" show "((FId' C) ## f) maps\<^bsub>CatCod (FId' C)\<^esub> ((FId' C) @@ X) to ((FId' C) @@ Y)" proof- have "X \ obj\<^bsub>CatDom (FId' C)\<^esub>" and "Y \ obj\<^bsub>CatDom (FId' C)\<^esub>" using a assms by (simp add: Category.MapsToObj IdentityFunctor'_def)+ moreover have "(FId' C) ## f = f" and "CatDom (FId' C) = CatCod (FId' C)" by (simp add: IdentityFunctor'_def)+ ultimately show ?thesis using assms a by(simp add: IdFtor'Obj) qed } qed lemma IdFtorFtor: "Category C \ Functor (FId C)" by (auto simp add: IdentityFunctor_def IdFtor'FtorM intro: MakeFtor) definition ConstFunctor' :: "('o1,'m1,'a) Category_scheme \ ('o2,'m2,'b) Category_scheme \ 'o2 \ ('o1,'o2,'m1,'m2,'a,'b) Functor" where "ConstFunctor' A B b \ \ CatDom = A , CatCod = B , MapM = (\ f . (Id B) b) \" definition "ConstFunctor A B b \ MakeFtor(ConstFunctor' A B b)" lemma ConstFtor' : assumes "Category A" "Category B" "b \ (Obj B)" shows "PreFunctor (ConstFunctor' A B b)" and "FunctorM (ConstFunctor' A B b)" proof- show "PreFunctor (ConstFunctor' A B b)" using assms apply (subst PreFunctor_def) apply (rule conjI)+ by (auto simp add: ConstFunctor'_def Category.Simps Category.CatIdCompId) moreover { fix X assume "X \ obj\<^bsub>A\<^esub>" "b \ obj\<^bsub>B\<^esub>" "PreFunctor (ConstFunctor' A B b)" hence "(ConstFunctor' A B b) @@ X = b" by (auto simp add: ConstFunctor'_def PreFunctor.FmToFo Category.Simps) } ultimately show "FunctorM (ConstFunctor' A B b)" using assms by (intro_locales, auto simp add: ConstFunctor'_def Category.Simps FunctorM_axioms_def) qed lemma ConstFtor: assumes "Category A" "Category B" "b \ (Obj B)" shows "Functor (ConstFunctor A B b)" by (auto simp add: assms ConstFtor' ConstFunctor_def MakeFtor) definition UnitFunctor :: "('o,'m,'a) Category_scheme \ ('o,unit,'m,unit,'a,unit) Functor" where "UnitFunctor C \ ConstFunctor C UnitCategory ()" lemma UnitFtor: assumes "Category C" shows "Functor(UnitFunctor C)" proof- have "() \ obj\<^bsub>UnitCategory\<^esub>" by (simp add: UnitCategory_def MakeCatObj) hence "Functor(ConstFunctor C UnitCategory ())" using assms by (simp add: ConstFtor) thus ?thesis by (simp add: UnitFunctor_def) qed definition FunctorComp' :: "('o1,'o2,'m1,'m2,'a1,'a2) Functor \ ('o2,'o3,'m2,'m3,'b1,'b2) Functor \ ('o1,'o3,'m1,'m3,'a1,'b2) Functor" (infixl ";;:" 71) where "FunctorComp' F G \ \ CatDom = CatDom F , CatCod = CatCod G , MapM = \ f . (MapM G)((MapM F) f) \" definition FunctorComp (infixl ";;;" 71) where "FunctorComp F G \ MakeFtor (FunctorComp' F G)" lemma FtorCompComp': assumes "f \>\<^bsub>CatDom F\<^esub> g" and "F \>;;; G" shows "G ## (F ## (f ;;\<^bsub>CatDom F\<^esub> g)) = (G ## (F ## f)) ;;\<^bsub>CatCod G\<^esub> (G ## (F ## g))" proof- have [simp]: "PreFunctor G \ PreFunctor F" using assms by auto have [simp]: "(F ## f) \>\<^bsub>CatDom G\<^esub> (F ## g)" using assms by (auto simp add: Functor.FunctorCompDef[of F f g]) have "F ## (f ;;\<^bsub>CatDom F\<^esub> g) = (F ## f) ;;\<^bsub>CatCod F\<^esub> (F ## g)" using assms by (auto simp add: PreFunctor.FunctorComp) hence "G ## (F ## (f ;;\<^bsub>CatDom F\<^esub> g)) = G ## ((F ## f) ;;\<^bsub>CatCod F\<^esub> (F ## g))" by simp also have "... = G ## ((F ## f) ;;\<^bsub>CatDom G\<^esub> (F ## g))" using assms by auto also have "... = (G ## (F ## f)) ;;\<^bsub>CatCod G\<^esub> (G ## (F ## g))" by (simp add: PreFunctor.FunctorComp[of G "(F ## f)" "(F ## g)"]) finally show ?thesis by simp qed lemma FtorCompId: assumes a: "X \ (Obj (CatDom F))" and "F \>;;; G" shows "G ## (F ## (id\<^bsub>CatDom F \<^esub>X)) = id\<^bsub>CatCod G\<^esub>(G @@ (F @@ X)) \ G @@ (F @@ X) \ (Obj (CatCod G))" proof- have [simp]: "PreFunctor G \ PreFunctor F" using assms by auto have b: "(F @@ X) \ obj\<^bsub>CatDom G\<^esub>" using assms by (auto simp add: PreFunctor.FunctorId2) have "G ## F ## (id\<^bsub>CatDom F \<^esub>X) = G ## (id\<^bsub>CatCod F\<^esub>(F @@ X))" using b a by (simp add: PreFunctor.FunctorId2[of F "X"]) also have "... = G ## (id\<^bsub>CatDom G\<^esub>(F @@ X))" using assms by auto also have "... = id\<^bsub>CatCod G\<^esub>(G @@ (F @@ X)) \ G @@ (F @@ X) \ (Obj (CatCod G))" using b by (simp add: PreFunctor.FunctorId2[of G "(F @@ X)"]) finally show ?thesis by simp qed lemma FtorCompIdDef: assumes a: "X \ (Obj (CatDom F))" and b: "PreFunctor (F ;;: G)" and "F \>;;; G" shows "(F ;;: G) @@ X = (G @@ (F @@ X))" proof- have "(F ;;: G) ## (id\<^bsub>CatDom (F ;;: G)\<^esub>(X)) = G ## (F ## (id\<^bsub>CatDom F\<^esub>(X)))" using assms by (simp add: FunctorComp'_def) also have "... = id\<^bsub>CatCod G\<^esub>(G @@ (F @@ (X)))" using assms a by(auto simp add: FtorCompId[of _ F G]) finally have "(F ;;: G) ## (id\<^bsub>CatDom (F ;;: G)\<^esub>(X)) = id\<^bsub>CatCod (F ;;: G)\<^esub>(G @@ (F @@ X))" using assms by (simp add: FunctorComp'_def) moreover have "G @@ (F @@ (X)) \ (Obj (CatCod (F ;;: G)))" using assms a by(auto simp add: FtorCompId[of _ F G] FunctorComp'_def) moreover have "X \ obj\<^bsub>CatDom (F ;;: G)\<^esub>" using a by (simp add: FunctorComp'_def) ultimately show ?thesis using b by (simp add: PreFunctor.FmToFo[of "F ;;: G" X "G @@ (F @@ X)"]) qed lemma FunctorCompMapsTo: assumes "f \ mor\<^bsub>CatDom (F ;;: G)\<^esub>" and "F \>;;; G" shows "(G ## (F ## f)) maps\<^bsub>CatCod G\<^esub> (G @@ (F @@ (dom\<^bsub>CatDom F\<^esub> f))) to (G @@ (F @@ (cod\<^bsub>CatDom F\<^esub> f)))" proof- have "f \ mor\<^bsub>CatDom F \<^esub>\ Functor F" using assms by (auto simp add: FunctorComp'_def) hence "(F ## f) maps\<^bsub>CatDom G\<^esub> (F @@ (dom\<^bsub>CatDom F\<^esub> f)) to (F @@ (cod\<^bsub>CatDom F\<^esub> f))" using assms by (auto simp add: Functor.FunctorMapsTo[of F f]) moreover have "FunctorM G" using assms by (auto simp add: FunctorComp_def Functor_def) ultimately show ?thesis by(simp add: FunctorM.FunctorCompM[of G "F ## f" "F @@ (dom\<^bsub>CatDom F\<^esub> f)" "F @@ (cod\<^bsub>CatDom F\<^esub> f)"]) qed lemma FunctorCompMapsTo2: assumes "f \ mor\<^bsub>CatDom (F ;;: G)\<^esub>" and "F \>;;; G" and "PreFunctor (F ;;: G)" shows "((F ;;: G) ## f) maps\<^bsub>CatCod (F ;;: G)\<^esub> ((F ;;: G) @@ (dom\<^bsub>CatDom (F ;;: G)\<^esub> f)) to ((F ;;: G) @@ (cod\<^bsub>CatDom (F ;;: G)\<^esub> f))" proof- have "Category (CatDom (F ;;: G))" using assms by (simp add: PreFunctor_def) hence 1: "(dom\<^bsub>CatDom (F ;;: G)\<^esub> f) \ obj\<^bsub>CatDom F \<^esub>\ (cod\<^bsub>CatDom (F ;;: G)\<^esub> f) \ obj\<^bsub>CatDom F\<^esub>" using assms by (auto simp add: Category.Simps FunctorComp'_def) have "(G ## (F ## f)) maps\<^bsub>CatCod G\<^esub> (G @@ (F @@ (dom\<^bsub>CatDom F\<^esub> f))) to (G @@ (F @@ (cod\<^bsub>CatDom F\<^esub> f)))" using assms by (auto simp add: FunctorCompMapsTo[of f F G]) moreover have "CatDom F = CatDom(F ;;: G) \ CatCod G = CatCod(F ;;: G) \ (G ## (F ## f)) = ((F ;;: G) ## f)" using assms by (simp add: FunctorComp'_def) moreover have "(F ;;: G) @@ (dom\<^bsub>CatDom (F ;;: G)\<^esub> f) = (G @@ (F @@ (dom\<^bsub>CatDom(F ;;: G)\<^esub> f))) \ (F ;;: G) @@ (cod\<^bsub>CatDom (F ;;: G)\<^esub> f) = (G @@ (F @@ (cod\<^bsub>CatDom(F ;;: G)\<^esub> f)))" by (auto simp add: FtorCompIdDef[of _ F G] 1 assms) ultimately show ?thesis by auto qed lemma FunctorCompMapsTo3: assumes "f maps\<^bsub>CatDom (F ;;: G)\<^esub> X to Y" and "F \>;;; G" and "PreFunctor (F ;;: G)" shows "F ;;: G ## f maps\<^bsub>CatCod (F ;;: G)\<^esub> F ;;: G @@ X to F ;;: G @@ Y" proof- have "f \ mor\<^bsub>CatDom (F ;;: G)\<^esub>" and "dom\<^bsub>CatDom (F ;;: G)\<^esub> f = X" and "cod\<^bsub>CatDom (F ;;: G)\<^esub> f = Y" using assms by auto thus ?thesis using assms by (auto intro: FunctorCompMapsTo2) qed lemma FtorCompPreFtor: assumes "F \>;;; G" shows "PreFunctor (F ;;: G)" proof- have 1: "PreFunctor G \ PreFunctor F" using assms by auto show "PreFunctor (F ;;: G)" using assms proof(auto simp add: PreFunctor_def FunctorComp'_def Category.Simps FtorCompId[of _ F G] intro:FtorCompComp') show "Category (CatDom F)" and "Category (CatCod G)" using assms 1 by (auto simp add: PreFunctor_def) qed qed lemma FtorCompM : assumes "F \>;;; G" shows "FunctorM (F ;;: G)" proof(auto simp only: FunctorM_def) show 1: "PreFunctor (F ;;: G)" using assms by (rule FtorCompPreFtor) { fix X Y f assume a: "f maps\<^bsub>CatDom (F ;;: G)\<^esub> X to Y" have "F ;;: G ## f maps\<^bsub>CatCod (F ;;: G)\<^esub> F ;;: G @@ X to F ;;: G @@ Y" using a assms 1 by (rule FunctorCompMapsTo3) } thus "FunctorM_axioms (F ;;: G)" by(auto simp add: 1 FunctorM_axioms_def) qed lemma FtorComp: assumes "F \>;;; G" shows "Functor (F ;;; G)" proof- have "FunctorM (F ;;: G)" using assms by (rule FtorCompM) thus ?thesis by (simp add: FunctorComp_def MakeFtor) qed lemma (in Functor) FunctorPreservesIso: assumes "ciso\<^bsub>CatDom F\<^esub> k" shows "ciso\<^bsub>CatCod F\<^esub> (F ## k)" proof- have [simp]: "k \ mor\<^bsub>CatDom F\<^esub>" using assms by (simp add: Category.IsoIsMor) have "cinv\<^bsub>CatCod F\<^esub> (F ## k) (F ## (Cinv\<^bsub>CatDom F\<^esub> k))" proof(rule Category.Inverse_relI) show "Category (CatCod F)" by simp show "(F ## k) \>\<^bsub>CatCod F\<^esub> (F ## (Cinv\<^bsub>CatDom F\<^esub> k))" by (rule FunctorCompDef, simp add: Category.IsoCompInv assms) show "(F ## k) ;;\<^bsub>CatCod F\<^esub> (F ## (Cinv\<^bsub>CatDom F\<^esub> k)) = id\<^bsub>CatCod F\<^esub> (dom\<^bsub>CatCod F\<^esub> (F ## k))" proof- have "(F ## k) ;;\<^bsub>CatCod F\<^esub> (F ## (Cinv\<^bsub>CatDom F\<^esub> k)) = F ## (k ;;\<^bsub>CatDom F\<^esub> (Cinv\<^bsub>CatDom F\<^esub> k))" using assms by(auto simp add: FunctorComp Category.IsoCompInv) also have "... = F ## (id\<^bsub>CatDom F\<^esub> (dom\<^bsub>CatDom F\<^esub> k))" using assms by (simp add: Category.IsoInvId2) also have "... = id\<^bsub>CatCod F\<^esub> (dom\<^bsub>CatCod F\<^esub> (F ## k))" by (simp add: FunctorId3Dom) finally show ?thesis by simp qed show "(F ## (Cinv\<^bsub>CatDom F\<^esub> k)) ;;\<^bsub>CatCod F\<^esub> (F ## k) = id\<^bsub>CatCod F\<^esub> (cod\<^bsub>CatCod F\<^esub> (F ## k))" proof- have "(F ## (Cinv\<^bsub>CatDom F\<^esub> k)) ;;\<^bsub>CatCod F\<^esub> (F ## k) = F ## ((Cinv\<^bsub>CatDom F\<^esub> k) ;;\<^bsub>CatDom F\<^esub> k)" using assms by(auto simp add: FunctorComp Category.InvCompIso) also have "... = F ## (id\<^bsub>CatDom F\<^esub> (cod\<^bsub>CatDom F\<^esub> k))" using assms by (simp add: Category.IsoInvId1) also have "... = id\<^bsub>CatCod F\<^esub> (cod\<^bsub>CatCod F\<^esub> (F ## k))" using assms by (simp add: FunctorId3Cod) finally show ?thesis by simp qed qed thus ?thesis by(auto simp add: isomorphism_def) qed declare PreFunctor.CatDom[simp] PreFunctor.CatCod [simp] lemma FunctorMFunctor[simp]: "Functor F \ FunctorM F" by (simp add: Functor_def) locale Equivalence = Functor + assumes Full: "\A \ Obj (CatDom F) ; B \ Obj (CatDom F) ; h maps\<^bsub>CatCod F\<^esub> (F @@ A) to (F @@ B)\ \ \ f . (f maps\<^bsub>CatDom F\<^esub> A to B) \ (F ## f = h)" and Faithful: "\f maps\<^bsub>CatDom F\<^esub> A to B ; g maps\<^bsub>CatDom F\<^esub> A to B ; F ## f = F ## g\ \ f = g" and IsoDense: "C \ Obj (CatCod F) \ \ A \ Obj (CatDom F) . ObjIso (CatCod F) (F @@ A) C" end diff --git a/thys/Category3/SetCategory.thy b/thys/Category3/SetCategory.thy --- a/thys/Category3/SetCategory.thy +++ b/thys/Category3/SetCategory.thy @@ -1,2443 +1,2443 @@ (* Title: SetCategory Author: Eugene W. Stark , 2016 Maintainer: Eugene W. Stark *) chapter SetCategory theory SetCategory imports Category Functor begin text\ This theory defines a locale \set_category\ that axiomatizes the notion ``category of all @{typ 'a}-sets and functions between them'' in the context of HOL. A primary reason for doing this is to make it possible to prove results (such as the Yoneda Lemma) that use such categories without having to commit to a particular element type @{typ 'a} and without having the results depend on the concrete details of a particular construction. The axiomatization given here is categorical, in the sense that if categories @{term S} and @{term S'} each interpret the \set_category\ locale, then a bijection between the sets of terminal objects of @{term S} and @{term S'} extends to an isomorphism of @{term S} and @{term S'} as categories. The axiomatization is based on the following idea: if, for some type @{typ 'a}, category @{term S} is the category of all @{typ 'a}-sets and functions between them, then the elements of type @{typ 'a} are in bijective correspondence with the terminal objects of category @{term S}. In addition, if @{term unity} is an arbitrarily chosen terminal object of @{term S}, then for each object @{term a}, the hom-set @{term "hom unity a"} (\emph{i.e.} the set of ``points'' or ``global elements'' of @{term a}) is in bijective correspondence with a subset of the terminal objects of @{term S}. By making a specific, but arbitrary, choice of such a correspondence, we can then associate with each object @{term a} of @{term S} a set @{term "set a"} that consists of all terminal objects @{term t} that correspond to some point @{term x} of @{term a}. Each arrow @{term f} then induces a function \Fun f \ set (dom f) \ set (cod f)\, defined on terminal objects of @{term S} by passing to points of @{term "dom f"}, composing with @{term f}, then passing back from points of @{term "cod f"} to terminal objects. Once we can associate a set with each object of @{term S} and a function with each arrow, we can force @{term S} to be isomorphic to the category of @{typ 'a}-sets by imposing suitable extensionality and completeness axioms. \ section "Some Lemmas about Restriction" text\ \sloppypar The development of the \set_category\ locale makes heavy use of the theory @{theory "HOL-Library.FuncSet"}. However, in some cases, I found that that theory did not provide results about restriction in the form that was most useful to me. I used the following additional results in various places. \ (* This variant of FuncSet.restrict_ext is sometimes useful. *) lemma restr_eqI: assumes "A = A'" and "\x. x \ A \ F x = F' x" shows "restrict F A = restrict F' A'" using assms by force (* This rule avoided a long proof in at least one instance where FuncSet.restrict_apply did not work. *) lemma restr_eqE [elim]: assumes "restrict F A = restrict F' A" and "x \ A" shows "F x = F' x" using assms restrict_def by metis (* This seems more useful than compose_eq in FuncSet. *) lemma compose_eq' [simp]: shows "compose A G F = restrict (G o F) A" unfolding compose_def restrict_def by auto section "Set Categories" text\ We first define the locale \set_category_data\, which sets out the basic data and definitions for the \set_category\ locale, without imposing any conditions other than that @{term S} is a category and that @{term img} is a function defined on the arrow type of @{term S}. The function @{term img} should be thought of as a mapping that takes a point @{term "x \ hom unity a"} to a corresponding terminal object @{term "img x"}. Eventually, assumptions will be introduced so that this is in fact the case. \ locale set_category_data = category S for S :: "'s comp" (infixr "\" 55) and img :: "'s \ 's" begin notation in_hom ("\_ : _ \ _\") text\ Call the set of all terminal objects of S the ``universe''. \ abbreviation Univ :: "'s set" where "Univ \ Collect terminal" text\ Choose an arbitrary element of the universe and call it @{term unity}. \ definition unity :: 's where "unity = (SOME t. terminal t)" text\ Each object @{term a} determines a subset @{term "set a"} of the universe, consisting of all those terminal objects @{term t} such that @{term "t = img x"} for some @{term "x \ hom unity a"}. \ definition set :: "'s \ 's set" where "set a = img ` hom unity a" text\ The inverse of the map @{term set} is a map @{term mkIde} that takes each subset of the universe to an identity of @{term[source=true] S}. \ definition mkIde :: "'s set \ 's" where "mkIde A = (if A \ Univ then inv_into (Collect ide) set A else null)" end text\ Next, we define a locale \set_category_given_img\ that augments the \set_category_data\ locale with assumptions that serve to define the notion of a set category with a chosen correspondence between points and terminal objects. The assumptions require that the universe be nonempty (so that the definition of @{term unity} makes sense), that the map @{term img} is a locally injective map taking points to terminal objects, that each terminal object @{term t} belongs to @{term "set t"}, that two objects of @{term S} are equal if they determine the same set, that two parallel arrows of @{term S} are equal if they determine the same function, that there is an object corresponding to each subset of the universe, and for any objects @{term a} and @{term b} and function @{term "F \ hom unity a \ hom unity b"} there is an arrow @{term "f \ hom a b"} whose action under the composition of @{term S} coincides with the function @{term F}. \ locale set_category_given_img = set_category_data S img for S :: "'s comp" (infixr "\" 55) and img :: "'s \ 's" + assumes nonempty_Univ: "Univ \ {}" and img_mapsto: "ide a \ img \ hom unity a \ Univ" and inj_img: "ide a \ inj_on img (hom unity a)" and stable_img: "terminal t \ t \ img ` hom unity t" and extensional_set: "\ ide a; ide b; set a = set b \ \ a = b" and extensional_arr: "\ par f f'; \x. \x : unity \ dom f\ \ f \ x = f' \ x \ \ f = f'" and set_complete: "A \ Univ \ \a. ide a \ set a = A" and fun_complete1: "\ ide a; ide b; F \ hom unity a \ hom unity b \ \ \f. \f : a \ b\ \ (\x. \x : unity \ dom f\ \ f \ x = F x)" begin text\ Each arrow @{term "f \ hom a b"} determines a function @{term "Fun f \ Univ \ Univ"}, by passing from @{term Univ} to @{term "hom a unity"}, composing with @{term f}, then passing back to @{term Univ}. \ definition Fun :: "'s \ 's \ 's" where "Fun f = restrict (img o S f o inv_into (hom unity (dom f)) img) (set (dom f))" lemma comp_arr_point: assumes "arr f" and "\x : unity \ dom f\" shows "f \ x = inv_into (hom unity (cod f)) img (Fun f (img x))" proof - have "\f \ x : unity \ cod f\" using assms by blast thus ?thesis using assms Fun_def inj_img set_def by simp qed text\ Parallel arrows that determine the same function are equal. \ lemma arr_eqI: assumes "par f f'" and "Fun f = Fun f'" shows "f = f'" using assms comp_arr_point extensional_arr by metis lemma terminal_unity: shows "terminal unity" using unity_def nonempty_Univ by (simp add: someI_ex) lemma ide_unity [simp]: shows "ide unity" using terminal_unity terminal_def by blast lemma set_subset_Univ [simp]: assumes "ide a" shows "set a \ Univ" using assms set_def img_mapsto by auto lemma inj_on_set: shows "inj_on set (Collect ide)" using extensional_set by (intro inj_onI, auto) text\ The mapping @{term mkIde}, which takes subsets of the universe to identities, and @{term set}, which takes identities to subsets of the universe, are inverses. \ lemma mkIde_set [simp]: assumes "ide a" shows "mkIde (set a) = a" using assms mkIde_def inj_on_set inv_into_f_f by simp lemma set_mkIde [simp]: assumes "A \ Univ" shows "set (mkIde A) = A" using assms mkIde_def set_complete someI_ex [of "\a. a \ Collect ide \ set a = A"] by (simp add: inv_into_def) lemma ide_mkIde [simp]: assumes "A \ Univ" shows "ide (mkIde A)" using assms mkIde_def mkIde_set set_complete by metis lemma arr_mkIde [iff]: shows "arr (mkIde A) \ A \ Univ" using not_arr_null mkIde_def ide_mkIde by auto lemma dom_mkIde [simp]: assumes "A \ Univ" shows "dom (mkIde A) = mkIde A" using assms ide_mkIde by simp lemma cod_mkIde [simp]: assumes "A \ Univ" shows "cod (mkIde A) = mkIde A" using assms ide_mkIde by simp text\ Each arrow @{term f} determines an extensional function from @{term "set (dom f)"} to @{term "set (cod f)"}. \ lemma Fun_mapsto: assumes "arr f" shows "Fun f \ extensional (set (dom f)) \ (set (dom f) \ set (cod f))" proof show "Fun f \ extensional (set (dom f))" using Fun_def by fastforce show "Fun f \ set (dom f) \ set (cod f)" proof fix t assume t: "t \ set (dom f)" have "Fun f t = img (f \ inv_into (hom unity (dom f)) img t)" using assms t Fun_def comp_def by simp moreover have "... \ set (cod f)" using assms t set_def inv_into_into [of t img "hom unity (dom f)"] by blast ultimately show "Fun f t \ set (cod f)" by auto qed qed text\ Identities of @{term[source=true] S} correspond to restrictions of the identity function. \ lemma Fun_ide [simp]: assumes "ide a" shows "Fun a = restrict (\x. x) (set a)" using assms Fun_def inj_img set_def comp_cod_arr by fastforce lemma Fun_mkIde [simp]: assumes "A \ Univ" shows "Fun (mkIde A) = restrict (\x. x) A" using assms by simp text\ Composition in @{term S} corresponds to extensional function composition. \ lemma Fun_comp [simp]: assumes "seq g f" shows "Fun (g \ f) = restrict (Fun g o Fun f) (set (dom f))" proof - have "restrict (img o S (g \ f) o (inv_into (hom unity (dom (g \ f))) img)) (set (dom (g \ f))) = restrict (Fun g o Fun f) (set (dom f))" proof - have 1: "set (dom (g \ f)) = set (dom f)" using assms by auto let ?img' = "\a. \t. inv_into (hom unity a) img t" have 2: "\t. t \ set (dom (g \ f)) \ (img o S (g \ f) o ?img' (dom (g \ f))) t = (Fun g o Fun f) t" proof - fix t assume "t \ set (dom (g \ f))" hence t: "t \ set (dom f)" by (simp add: 1) have 3: "\a x. x \ hom unity a \ ?img' a (img x) = x" using assms img_mapsto inj_img ide_cod inv_into_f_eq by (metis arrI in_homE mem_Collect_eq) have 4: "?img' (dom f) t \ hom unity (dom f)" using assms t inv_into_into [of t img "hom unity (dom f)"] set_def by simp have "(img o S (g \ f) o ?img' (dom (g \ f))) t = img (g \ f \ ?img' (dom f) t)" using assms dom_comp comp_assoc by simp also have "... = img (g \ ?img' (dom g) (Fun f t))" using assms t 3 Fun_def set_def comp_arr_point by auto also have "... = Fun g (Fun f t)" proof - have "Fun f t \ img ` hom unity (cod f)" using assms t Fun_mapsto set_def by fast thus ?thesis using assms by (auto simp add: set_def Fun_def) qed finally show "(img o S (g \ f) o ?img' (dom (g \ f))) t = (Fun g o Fun f) t" by auto qed show ?thesis using 1 2 by auto qed thus ?thesis using Fun_def by auto qed text\ The constructor @{term mkArr} is used to obtain an arrow given subsets @{term A} and @{term B} of the universe and a function @{term "F \ A \ B"}. \ definition mkArr :: "'s set \ 's set \ ('s \ 's) \ 's" where "mkArr A B F = (if A \ Univ \ B \ Univ \ F \ A \ B then (THE f. f \ hom (mkIde A) (mkIde B) \ Fun f = restrict F A) else null)" text\ Each function @{term "F \ set a \ set b"} determines a unique arrow @{term "f \ hom a b"}, such that @{term "Fun f"} is the restriction of @{term F} to @{term "set a"}. \ lemma fun_complete: assumes "ide a" and "ide b" and "F \ set a \ set b" shows "\!f. \f : a \ b\ \ Fun f = restrict F (set a)" proof - let ?P = "\f. \f : a \ b\ \ Fun f = restrict F (set a)" show "\!f. ?P f" proof have "\f. ?P f" proof - let ?F' = "\x. inv_into (hom unity b) img (F (img x))" have "?F' \ hom unity a \ hom unity b" proof fix x assume x: "x \ hom unity a" have "F (img x) \ set b" using assms(3) x set_def by auto thus "inv_into (hom unity b) img (F (img x)) \ hom unity b" using assms img_mapsto inj_img set_def by auto qed hence "\f. \f : a \ b\ \ (\x. \x : unity \ a\ \ f \ x = ?F' x)" using assms fun_complete1 by force from this obtain f where f: "\f : a \ b\ \ (\x. \x : unity \ a\ \ f \ x = ?F' x)" by blast let ?img' = "\a. \t. inv_into (hom unity a) img t" have "Fun f = restrict F (set a)" proof (unfold Fun_def, intro restr_eqI) show "set (dom f) = set a" using f by auto show "\t. t \ set (dom f) \ (img \ S f \ ?img' (dom f)) t = F t" proof - fix t assume t: "t \ set (dom f)" have "(img \ S f \ ?img' (dom f)) t = img (f \ ?img' (dom f) t)" by simp also have "... = img (?F' (?img' (dom f) t))" proof - have "?img' (dom f) t \ hom unity (dom f)" using t set_def inv_into_into by metis thus ?thesis using f by auto qed also have "... = img (?img' (cod f) (F t))" using f t set_def inj_img by auto also have "... = F t" proof - have "F t \ set (cod f)" using assms f t by auto thus ?thesis using f t set_def inj_img by auto qed finally show "(img \ S f \ ?img' (dom f)) t = F t" by auto qed qed thus ?thesis using f by blast qed thus F: "?P (SOME f. ?P f)" using someI_ex [of ?P] by fast show "\f'. ?P f' \ f' = (SOME f. ?P f)" using F arr_eqI by (metis (no_types, lifting) in_homE) qed qed lemma mkArr_in_hom: assumes "A \ Univ" and "B \ Univ" and "F \ A \ B" shows "\mkArr A B F : mkIde A \ mkIde B\" using assms mkArr_def fun_complete [of "mkIde A" "mkIde B" F] theI' [of "\f. f \ hom (mkIde A) (mkIde B) \ Fun f = restrict F A"] by simp text\ The ``only if'' direction of the next lemma can be achieved only if there exists a non-arrow element of type @{typ 's}, which can be used as the value of @{term "mkArr A B F"} in cases where @{term "F \ A \ B"}. Nevertheless, it is essential to have this, because without the ``only if'' direction, we can't derive any useful consequences from an assumption of the form @{term "arr (mkArr A B F)"}; instead we have to obtain @{term "F \ A \ B"} some other way. This is is usually highly inconvenient and it makes the theory very weak and almost unusable in practice. The observation that having a non-arrow value of type @{typ 's} solves this problem is ultimately what led me to incorporate @{term null} first into the definition of the \set_category\ locale and then, ultimately, into the definition of the \category\ locale. I believe this idea is critical to the usability of the entire development. \ (* TODO: This gets used as an introduction rule, but the conjunction on the right-hand side is not very convenient. *) lemma arr_mkArr [iff]: shows "arr (mkArr A B F) \ A \ Univ \ B \ Univ \ F \ A \ B" proof show "arr (mkArr A B F) \ A \ Univ \ B \ Univ \ F \ A \ B" using mkArr_def not_arr_null ex_un_null someI_ex [of "\f. \arr f"] by metis show "A \ Univ \ B \ Univ \ F \ A \ B \ arr (mkArr A B F)" using mkArr_in_hom by auto qed lemma Fun_mkArr': assumes "arr (mkArr A B F)" shows "\mkArr A B F : mkIde A \ mkIde B\" and "Fun (mkArr A B F) = restrict F A" proof - have 1: "A \ Univ \ B \ Univ \ F \ A \ B" using assms by fast have 2: "mkArr A B F \ hom (mkIde A) (mkIde B) \ Fun (mkArr A B F) = restrict F (set (mkIde A))" proof - have "\!f. f \ hom (mkIde A) (mkIde B) \ Fun f = restrict F (set (mkIde A))" using 1 fun_complete [of "mkIde A" "mkIde B" F] by simp thus ?thesis using 1 mkArr_def theI' by simp qed show "\mkArr A B F : mkIde A \ mkIde B\" using 1 2 by auto show "Fun (mkArr A B F) = restrict F A" using 1 2 by auto qed lemma mkArr_Fun [simp]: assumes "arr f" shows "mkArr (set (dom f)) (set (cod f)) (Fun f) = f" proof - have 1: "set (dom f) \ Univ \ set (cod f) \ Univ \ ide (dom f) \ ide (cod f) \ Fun f \ extensional (set (dom f)) \ (set (dom f) \ set (cod f))" using assms Fun_mapsto by force hence "\!f'. f' \ hom (dom f) (cod f) \ Fun f' = restrict (Fun f) (set (dom f))" using fun_complete by force moreover have "f \ hom (dom f) (cod f) \ Fun f = restrict (Fun f) (set (dom f))" using assms 1 extensional_restrict by force ultimately have "f = (THE f'. f' \ hom (dom f) (cod f) \ Fun f' = restrict (Fun f) (set (dom f)))" using theI' [of "\f'. f' \ hom (dom f) (cod f) \ Fun f' = restrict (Fun f) (set (dom f))"] by blast also have "... = mkArr (set (dom f)) (set (cod f)) (Fun f)" using assms 1 mkArr_def by simp finally show ?thesis by auto qed lemma dom_mkArr [simp]: assumes "arr (mkArr A B F)" shows "dom (mkArr A B F) = mkIde A" using assms Fun_mkArr' by auto lemma cod_mkArr [simp]: assumes "arr (mkArr A B F)" shows "cod (mkArr A B F) = mkIde B" using assms Fun_mkArr' by auto lemma Fun_mkArr [simp]: assumes "arr (mkArr A B F)" shows "Fun (mkArr A B F) = restrict F A" using assms Fun_mkArr' by auto text\ The following provides the basic technique for showing that arrows constructed using @{term mkArr} are equal. \ lemma mkArr_eqI [intro]: assumes "arr (mkArr A B F)" and "A = A'" and "B = B'" and "\x. x \ A \ F x = F' x" shows "mkArr A B F = mkArr A' B' F'" using assms arr_mkArr Fun_mkArr by (intro arr_eqI, auto simp add: Pi_iff) text\ This version avoids trivial proof obligations when the domain and codomain sets are identical from the context. \ lemma mkArr_eqI' [intro]: assumes "arr (mkArr A B F)" and "\x. x \ A \ F x = F' x" shows "mkArr A B F = mkArr A B F'" using assms mkArr_eqI by simp lemma mkArr_restrict_eq [simp]: assumes "arr (mkArr A B F)" shows "mkArr A B (restrict F A) = mkArr A B F" using assms by (intro mkArr_eqI', auto) lemma mkArr_restrict_eq': assumes "arr (mkArr A B (restrict F A))" shows "mkArr A B (restrict F A) = mkArr A B F" using assms by (intro mkArr_eqI', auto) lemma mkIde_as_mkArr [simp]: assumes "A \ Univ" shows "mkArr A A (\x. x) = mkIde A" using assms by (intro arr_eqI, auto) lemma comp_mkArr [simp]: assumes "arr (mkArr A B F)" and "arr (mkArr B C G)" shows "mkArr B C G \ mkArr A B F = mkArr A C (G \ F)" proof (intro arr_eqI) have 1: "seq (mkArr B C G) (mkArr A B F)" using assms by force have 2: "G o F \ A \ C" using assms by auto show "par (mkArr B C G \ mkArr A B F) (mkArr A C (G \ F))" using 1 2 by auto show "Fun (mkArr B C G \ mkArr A B F) = Fun (mkArr A C (G \ F))" using 1 2 by fastforce qed text\ The locale assumption @{prop stable_img} forces @{term "t \ set t"} in case @{term t} is a terminal object. This is very convenient, as it results in the characterization of terminal objects as identities @{term t} for which @{term "set t = {t}"}. However, it is not absolutely necessary to have this. The following weaker characterization of terminal objects can be proved without the @{prop stable_img} assumption. \ lemma terminal_char1: shows "terminal t \ ide t \ (\!x. x \ set t)" proof - have "terminal t \ ide t \ (\!x. x \ set t)" proof - assume t: "terminal t" have "ide t" using t terminal_def by auto moreover have "\!x. x \ set t" proof - have "\!x. x \ hom unity t" using t terminal_unity terminal_def by auto thus ?thesis using set_def by auto qed ultimately show "ide t \ (\!x. x \ set t)" by auto qed moreover have "ide t \ (\!x. x \ set t) \ terminal t" proof - assume t: "ide t \ (\!x. x \ set t)" from this obtain t' where "set t = {t'}" by blast hence t': "set t = {t'} \ {t'} \ Univ \ t = mkIde {t'}" using t set_subset_Univ mkIde_set by metis show "terminal t" proof show "ide t" using t by simp show "\a. ide a \ \!f. \f : a \ t\" proof - fix a assume a: "ide a" show "\!f. \f : a \ t\" proof show 1: "\mkArr (set a) {t'} (\x. t') : a \ t\" using a t t' mkArr_in_hom by (metis Pi_I' mkIde_set set_subset_Univ singletonD) show "\f. \f : a \ t\ \ f = mkArr (set a) {t'} (\x. t')" proof - fix f assume f: "\f : a \ t\" show "f = mkArr (set a) {t'} (\x. t')" proof (intro arr_eqI) show 1: "par f (mkArr (set a) {t'} (\x. t'))" using 1 f in_homE by metis show "Fun f = Fun (mkArr (set a) {t'} (\x. t'))" proof - have "Fun (mkArr (set a) {t'} (\x. t')) = (\x\set a. t')" using 1 Fun_mkArr by simp also have "... = Fun f" proof - have "\x. x \ set a \ Fun f x = t'" using f t' Fun_def mkArr_Fun arr_mkArr by (metis PiE in_homE singletonD) moreover have "\x. x \ set a \ Fun f x = undefined" using f Fun_def by auto ultimately show ?thesis by auto qed finally show ?thesis by force qed qed qed qed qed qed qed ultimately show ?thesis by blast qed text\ As stated above, in the presence of the @{prop stable_img} assumption we have the following stronger characterization of terminal objects. \ lemma terminal_char2: shows "terminal t \ ide t \ set t = {t}" proof assume t: "terminal t" show "ide t \ set t = {t}" proof show "ide t" using t terminal_char1 by auto show "set t = {t}" proof - have "\!x. x \ hom unity t" using t terminal_def terminal_unity by force moreover have "t \ img ` hom unity t" using t stable_img set_def by simp ultimately show ?thesis using set_def by auto qed qed next assume "ide t \ set t = {t}" thus "terminal t" using terminal_char1 by force qed end text\ At last, we define the \set_category\ locale by existentially quantifying out the choice of a particular @{term img} map. We need to know that such a map exists, but it does not matter which one we choose. \ locale set_category = category S for S :: "'s comp" (infixr "\" 55) + assumes ex_img: "\img. set_category_given_img S img" begin notation in_hom ("\_ : _ \ _\") definition some_img where "some_img = (SOME img. set_category_given_img S img)" end sublocale set_category \ set_category_given_img S some_img proof - have "\img. set_category_given_img S img" using ex_img by auto thus "set_category_given_img S some_img" using someI_ex [of "\img. set_category_given_img S img"] some_img_def by metis qed context set_category begin text\ The arbitrary choice of @{term img} induces a system of inclusions, which are arrows corresponding to inclusions of subsets. \ definition incl :: "'s \ bool" where "incl f = (arr f \ set (dom f) \ set (cod f) \ f = mkArr (set (dom f)) (set (cod f)) (\x. x))" lemma Fun_incl: assumes "incl f" shows "Fun f = (\x \ set (dom f). x)" using assms incl_def by (metis Fun_mkArr) lemma ex_incl_iff_subset: assumes "ide a" and "ide b" shows "(\f. \f : a \ b\ \ incl f) \ set a \ set b" proof show "\f. \f : a \ b\ \ incl f \ set a \ set b" using incl_def by auto show "set a \ set b \ \f. \f : a \ b\ \ incl f" proof assume 1: "set a \ set b" show "\mkArr (set a) (set b) (\x. x) : a \ b\ \ incl (mkArr (set a) (set b) (\x. x))" proof show "\mkArr (set a) (set b) (\x. x) : a \ b\" proof - have "(\x. x) \ set a \ set b" using 1 by auto thus ?thesis using assms mkArr_in_hom set_subset_Univ in_homI by auto qed thus "incl (mkArr (set a) (set b) (\x. x))" using 1 incl_def by force qed qed qed end section "Categoricity" text\ In this section we show that the \set_category\ locale completely characterizes the structure of its interpretations as categories, in the sense that for any two interpretations @{term S} and @{term S'}, a bijection between the universe of @{term S} and the universe of @{term S'} extends to an isomorphism of @{term S} and @{term S'}. \ locale two_set_categories_bij_betw_Univ = S: set_category S + S': set_category S' for S :: "'s comp" (infixr "\" 55) and S' :: "'t comp" (infixr "\\" 55) and \ :: "'s \ 't" + assumes bij_\: "bij_betw \ S.Univ S'.Univ" begin notation S.in_hom ("\_ : _ \ _\") - notation S'.in_hom ("\_ : _ \' _\") + notation S'.in_hom ("\_ : _ \'' _\") abbreviation \ where "\ \ inv_into S.Univ \" lemma \_\: assumes "t \ S.Univ" shows "\ (\ t) = t" using assms bij_\ bij_betw_inv_into_left by metis lemma \_\: assumes "t' \ S'.Univ" shows "\ (\ t') = t'" using assms bij_\ bij_betw_inv_into_right by metis lemma \_img_\_img: assumes "A \ S.Univ" shows "\ ` \ ` A = A" using assms bij_\ by (simp add: bij_betw_def) lemma \_img_\_img: assumes "A' \ S'.Univ" shows "\ ` \ ` A' = A'" using assms bij_\ by (simp add: bij_betw_def image_inv_into_cancel) text\ The object map @{term \o} of a functor from @{term[source=true] S} to @{term[source=true] S'}. \ definition \o where "\o = (\a \ Collect S.ide. S'.mkIde (\ ` S.set a))" lemma set_\o: assumes "S.ide a" shows "S'.set (\o a) = \ ` S.set a" proof - from assms have "S.set a \ S.Univ" by simp then show ?thesis using S'.set_mkIde \o_def assms bij_\ bij_betw_def image_mono mem_Collect_eq restrict_def by (metis (no_types, lifting)) qed lemma \o_preserves_ide: assumes "S.ide a" shows "S'.ide (\o a)" using assms S'.ide_mkIde S.set_subset_Univ bij_\ bij_betw_def image_mono restrict_apply' unfolding \o_def by (metis (mono_tags, lifting) mem_Collect_eq) text\ The map @{term \a} assigns to each arrow @{term f} of @{term[source=true] S} the function on the universe of @{term[source=true] S'} that is the same as the function induced by @{term f} on the universe of @{term[source=true] S}, up to the bijection @{term \} between the two universes. \ definition \a where "\a = (\f. \x' \ \ ` S.set (S.dom f). \ (S.Fun f (\ x')))" lemma \a_mapsto: assumes "S.arr f" shows "\a f \ S'.set (\o (S.dom f)) \ S'.set (\o (S.cod f))" proof - have "\a f \ \ ` S.set (S.dom f) \ \ ` S.set (S.cod f)" proof fix x assume x: "x \ \ ` S.set (S.dom f)" have "\ x \ S.set (S.dom f)" using assms x \_img_\_img [of "S.set (S.dom f)"] S.set_subset_Univ by auto hence "S.Fun f (\ x) \ S.set (S.cod f)" using assms S.Fun_mapsto by auto hence "\ (S.Fun f (\ x)) \ \ ` S.set (S.cod f)" by simp thus "\a f x \ \ ` S.set (S.cod f)" using x \a_def by auto qed thus ?thesis using assms set_\o \o_preserves_ide by auto qed text\ The map @{term \a} takes composition of arrows to extensional composition of functions. \ lemma \a_comp: assumes gf: "S.seq g f" shows "\a (g \ f) = restrict (\a g o \a f) (S'.set (\o (S.dom f)))" proof - have "\a (g \ f) = (\x' \ \ ` S.set (S.dom f). \ (S.Fun (S g f) (\ x')))" using gf \a_def by auto also have "... = (\x' \ \ ` S.set (S.dom f). \ (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\ x')))" using gf set_\o S.Fun_comp by simp also have "... = restrict (\a g o \a f) (S'.set (\o (S.dom f)))" proof - have "\x'. x' \ \ ` S.set (S.dom f) \ \ (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\ x')) = \a g (\a f x')" proof - fix x' assume X': "x' \ \ ` S.set (S.dom f)" hence 1: "\ x' \ S.set (S.dom f)" using gf \_img_\_img [of "S.set (S.dom f)"] S.set_subset_Univ S.ide_dom by blast hence "\ (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\ x')) = \ (S.Fun g (S.Fun f (\ x')))" using restrict_apply by auto also have "... = \ (S.Fun g (\ (\ (S.Fun f (\ x')))))" proof - have "S.Fun f (\ x') \ S.set (S.cod f)" using gf 1 S.Fun_mapsto by fast hence "\ (\ (S.Fun f (\ x'))) = S.Fun f (\ x')" using assms bij_\ S.set_subset_Univ bij_betw_def inv_into_f_f subsetCE S.ide_cod by (metis S.seqE) thus ?thesis by auto qed also have "... = \a g (\a f x')" proof - have "\a f x' \ \ ` S.set (S.cod f)" using gf S.ide_dom S.ide_cod X' \a_mapsto [of f] set_\o [of "S.dom f"] set_\o [of "S.cod f"] by blast thus ?thesis using gf X' \a_def by auto qed finally show "\ (restrict (S.Fun g o S.Fun f) (S.set (S.dom f)) (\ x')) = \a g (\a f x')" by auto qed thus ?thesis using assms set_\o by fastforce qed finally show ?thesis by auto qed text\ Finally, we use @{term \o} and @{term \a} to define a functor @{term \}. \ definition \ where "\ f = (if S.arr f then S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod f))) (\a f) else S'.null)" lemma \_in_hom: assumes "S.arr f" shows "\ f \ S'.hom (\o (S.dom f)) (\o (S.cod f))" proof - have "\\ f : S'.dom (\ f) \' S'.cod (\ f)\" using assms \_def \a_mapsto \o_preserves_ide by (intro S'.in_homI, auto) thus ?thesis using assms \_def \a_mapsto \o_preserves_ide by auto qed lemma \_ide [simp]: assumes "S.ide a" shows "\ a = \o a" proof - have "\ a = S'.mkArr (S'.set (\o a)) (S'.set (\o a)) (\x'. x')" proof - have "\\ a : \o a \' \o a\" using assms \_in_hom S.ide_in_hom by fastforce moreover have "\a a = restrict (\x'. x') (S'.set (\o a))" proof - have "\a a = (\x' \ \ ` S.set a. \ (S.Fun a (\ x')))" using assms \a_def restrict_apply by auto also have "... = (\x' \ S'.set (\o a). \ (\ x'))" proof - have "S.Fun a = (\x \ S.set a. x)" using assms S.Fun_ide by simp moreover have "\x'. x' \ \ ` S.set a \ \ x' \ S.set a" using assms bij_\ S.set_subset_Univ image_iff by (metis \_img_\_img) ultimately show ?thesis using assms set_\o by auto qed also have "... = restrict (\x'. x') (S'.set (\o a))" using assms S'.set_subset_Univ \o_preserves_ide \_\ by (meson restr_eqI subsetCE) ultimately show ?thesis by auto qed ultimately show ?thesis using assms \_def \o_preserves_ide S'.mkArr_restrict_eq' by (metis S'.arrI S.ide_char) qed thus ?thesis using assms S'.mkIde_as_mkArr \o_preserves_ide \_in_hom by simp qed lemma set_dom_\: assumes "S.arr f" shows "S'.set (S'.dom (\ f)) = \ ` (S.set (S.dom f))" using assms S.ide_dom \_in_hom \_ide set_\o by fastforce lemma \_comp: assumes "S.seq g f" shows "\ (g \ f) = \ g \\ \ f" proof - have "\ (g \ f) = S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod g))) (\a (S g f))" using \_def assms by auto also have "... = S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod g))) (restrict (\a g o \a f) (S'.set (\o (S.dom f))))" using assms \a_comp set_\o by force also have "... = S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod g))) (\a g o \a f)" proof - have "S'.arr (S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod g))) (\a g o \a f))" using assms \a_mapsto [of f] \a_mapsto [of g] \o_preserves_ide S'.arr_mkArr by (elim S.seqE, auto) thus ?thesis using assms S'.mkArr_restrict_eq by auto qed also have "... = S' (S'.mkArr (S'.set (\o (S.dom g))) (S'.set (\o (S.cod g))) (\a g)) (S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod f))) (\a f))" proof - have "S'.arr (S'.mkArr (S'.set (\o (S.dom f))) (S'.set (\o (S.cod f))) (\a f))" using assms \a_mapsto set_\o S.ide_dom S.ide_cod \o_preserves_ide S'.arr_mkArr S'.set_subset_Univ S.seqE by metis moreover have "S'.arr (S'.mkArr (S'.set (\o (S.dom g))) (S'.set (\o (S.cod g))) (\a g))" using assms \a_mapsto set_\o S.ide_dom S.ide_cod \o_preserves_ide S'.arr_mkArr S'.set_subset_Univ S.seqE by metis ultimately show ?thesis using assms S'.comp_mkArr by force qed also have "... = \ g \\ \ f" using assms \_def by force finally show ?thesis by fast qed interpretation \: "functor" S S' \ apply unfold_locales using \_def apply simp using \_in_hom \_comp by auto lemma \_is_functor: shows "functor S S' \" .. lemma Fun_\: assumes "S.arr f" and "x \ S.set (S.dom f)" shows "S'.Fun (\ f) (\ x) = \a f (\ x)" using assms \_def \.preserves_arr set_\o by auto lemma \_acts_elementwise: assumes "S.ide a" shows "S'.set (\ a) = \ ` S.set a" proof have 0: "S'.set (\ a) = \ ` S.set a" using assms \_ide set_\o by simp have 1: "\x. x \ S.set a \ \ x = \ x" proof - fix x assume x: "x \ S.set a" have 1: "S.terminal x" using assms x S.set_subset_Univ by blast hence 2: "S'.terminal (\ x)" by (metis CollectD CollectI bij_\ bij_betw_def image_iff) have "\ x = \o x" using assms x 1 \_ide S.terminal_def by auto also have "... = \ x" proof - have "\o x = S'.mkIde (\ ` S.set x)" using assms 1 x \o_def S.terminal_def by auto moreover have "S'.mkIde (\ ` S.set x) = \ x" using assms x 1 2 S.terminal_char2 S'.terminal_char2 S'.mkIde_set bij_\ by (metis image_empty image_insert) ultimately show ?thesis by auto qed finally show "\ x = \ x" by auto qed show "S'.set (\ a) \ \ ` S.set a" using 0 1 by force show "\ ` S.set a \ S'.set (\ a)" using 0 1 by force qed lemma \_preserves_incl: assumes "S.incl m" shows "S'.incl (\ m)" proof - have 1: "S.arr m \ S.set (S.dom m) \ S.set (S.cod m) \ m = S.mkArr (S.set (S.dom m)) (S.set (S.cod m)) (\x. x)" using assms S.incl_def by blast have "S'.arr (\ m)" using 1 by auto moreover have 2: "S'.set (S'.dom (\ m)) \ S'.set (S'.cod (\ m))" using 1 \.preserves_dom \.preserves_cod \_acts_elementwise by (metis (full_types) S.ide_cod S.ide_dom image_mono) moreover have "\ m = S'.mkArr (S'.set (S'.dom (\ m))) (S'.set (S'.cod (\ m))) (\x'. x')" proof - have "\ m = S'.mkArr (S'.set (\o (S.dom m))) (S'.set (\o (S.cod m))) (\a m)" using 1 \_def by simp also have "... = S'.mkArr (S'.set (S'.dom (\ m))) (S'.set (S'.cod (\ m))) (\a m)" using 1 \_ide by auto finally have 3: "\ m = S'.mkArr (S'.set (S'.dom (\ m))) (S'.set (S'.cod (\ m))) (\a m)" by auto also have "... = S'.mkArr (S'.set (S'.dom (\ m))) (S'.set (S'.cod (\ m))) (\x'. x')" proof - have 4: "S.Fun m = restrict (\x. x) (S.set (S.dom m))" using assms S.incl_def by (metis (full_types) S.Fun_mkArr) hence "\a m = restrict (\x'. x') (\ ` (S.set (S.dom m)))" proof - have 5: "\x'. x' \ \ ` S.set (S.dom m) \ \ (\ x') = x'" using 1 bij_\ bij_betw_def S'.set_subset_Univ S.ide_dom \o_preserves_ide f_inv_into_f set_\o by (metis subsetCE) have "\a m = restrict (\x'. \ (S.Fun m (\ x'))) (\ ` S.set (S.dom m))" using \a_def by simp also have "... = restrict (\x'. x') (\ ` S.set (S.dom m))" proof - have "\x. x \ \ ` (S.set (S.dom m)) \ \ (S.Fun m (\ x)) = x" proof - fix x assume x: "x \ \ ` (S.set (S.dom m))" hence "\ x \ S.set (S.dom m)" using 1 S.ide_dom S.set_subset_Univ \_img_\_img image_eqI by metis thus "\ (S.Fun m (\ x)) = x" using 1 4 5 x by simp qed thus ?thesis by auto qed finally show ?thesis by auto qed hence "\a m = restrict (\x'. x') (S'.set (S'.dom (\ m)))" using 1 set_dom_\ by auto thus ?thesis using 2 3 \S'.arr (\ m)\ S'.mkArr_restrict_eq S'.ide_cod S'.ide_dom S'.incl_def by (metis S'.arr_mkArr image_restrict_eq image_subset_iff_funcset) qed finally show ?thesis by auto qed ultimately show ?thesis using S'.incl_def by blast qed text\ Interchange the role of @{term \} and @{term \} to obtain a functor \\\ from @{term[source=true] S'} to @{term[source=true] S}. \ interpretation INV: two_set_categories_bij_betw_Univ S' S \ apply unfold_locales by (simp add: bij_\ bij_betw_inv_into) abbreviation \o where "\o \ INV.\o" abbreviation \a where "\a \ INV.\a" abbreviation \ where "\ \ INV.\" interpretation \: "functor" S' S \ using INV.\_is_functor by auto text\ The functors @{term \} and @{term \} are inverses. \ lemma Fun_\: assumes "S'.arr f'" and "x' \ S'.set (S'.dom f')" shows "S.Fun (\ f') (\ x') = \a f' (\ x')" using assms INV.Fun_\ by blast lemma \o_\o: assumes "S.ide a" shows "\o (\o a) = a" using assms \o_def INV.\o_def \_img_\_img \o_preserves_ide set_\o by force lemma \\: assumes "S.arr f" shows "\ (\ f) = f" proof (intro S.arr_eqI) show par: "S.par (\ (\ f)) f" using assms \o_preserves_ide \o_\o by auto show "S.Fun (\ (\ f)) = S.Fun f" proof - have "S.arr (\ (\ f))" using assms by auto moreover have "\ (\ f) = S.mkArr (S.set (S.dom f)) (S.set (S.cod f)) (\a (\ f))" using assms INV.\_def \_in_hom \o_\o by auto moreover have "\a (\ f) = (\x \ S.set (S.dom f). \ (S'.Fun (\ f) (\ x)))" proof - have "\a (\ f) = (\x \ \ ` S'.set (S'.dom (\ f)). \ (S'.Fun (\ f) (\ x)))" proof - have "\x. x \ \ ` S'.set (S'.dom (\ f)) \ INV.\ x = \ x" using assms S.ide_dom S.set_subset_Univ \.preserves_reflects_arr par bij_\ inv_into_inv_into_eq subsetCE INV.set_dom_\ by metis thus ?thesis using INV.\a_def by auto qed moreover have "\ ` S'.set (S'.dom (\ f)) = S.set (S.dom f)" using assms by (metis par \.preserves_reflects_arr INV.set_dom_\) ultimately show ?thesis by auto qed ultimately have 1: "S.Fun (\ (\ f)) = (\x \ S.set (S.dom f). \ (S'.Fun (\ f) (\ x)))" using S'.Fun_mkArr by simp show ?thesis proof fix x have "x \ S.set (S.dom f) \ S.Fun (\ (\ f)) x = S.Fun f x" using 1 assms extensional_def S.Fun_mapsto S.Fun_def by auto moreover have "x \ S.set (S.dom f) \ S.Fun (\ (\ f)) x = S.Fun f x" proof - assume x: "x \ S.set (S.dom f)" have "S.Fun (\ (\ f)) x = \ (\ (S.Fun f (\ (\ x))))" using assms x 1 Fun_\ bij_\ \a_def by auto also have "... = S.Fun f x" proof - have 2: "\x. x \ S.Univ \ \ (\ x) = x" using bij_\ bij_betw_inv_into_left by fast have "S.Fun f (\ (\ x)) = S.Fun f x" using assms x 2 by (metis S.ide_dom S.set_subset_Univ subsetCE) moreover have "S.Fun f x \ S.Univ" using x assms S.Fun_mapsto S.set_subset_Univ S.ide_cod by blast ultimately show ?thesis using 2 by auto qed finally show ?thesis by auto qed ultimately show "S.Fun (\ (\ f)) x = S.Fun f x" by auto qed qed qed lemma \o_\o: assumes "S'.ide a'" shows "\o (\o a') = a'" using assms \o_def INV.\o_def \_img_\_img INV.\o_preserves_ide \_\ INV.set_\o by force lemma \\: assumes "S'.arr f'" shows "\ (\ f') = f'" proof (intro S'.arr_eqI) show par: "S'.par (\ (\ f')) f'" using assms \.preserves_ide \.preserves_ide \_ide INV.\_ide \o_\o by auto show "S'.Fun (\ (\ f')) = S'.Fun f'" proof - have "S'.arr (\ (\ f'))" using assms by blast moreover have "\ (\ f') = S'.mkArr (S'.set (S'.dom f')) (S'.set (S'.cod f')) (\a (\ f'))" using assms \_def INV.\_in_hom \o_\o by simp moreover have "\a (\ f') = (\x' \ S'.set (S'.dom f'). \ (S.Fun (\ f') (\ x')))" unfolding \a_def using assms par \.preserves_arr set_dom_\ by metis ultimately have 1: "S'.Fun (\ (\ f')) = (\x' \ S'.set (S'.dom f'). \ (S.Fun (\ f') (\ x')))" using S'.Fun_mkArr by simp show ?thesis proof fix x' have "x' \ S'.set (S'.dom f') \ S'.Fun (\ (\ f')) x' = S'.Fun f' x'" using 1 assms S'.Fun_mapsto extensional_def by (simp add: S'.Fun_def) moreover have "x' \ S'.set (S'.dom f') \ S'.Fun (\ (\ f')) x' = S'.Fun f' x'" proof - assume x': "x' \ S'.set (S'.dom f')" have "S'.Fun (\ (\ f')) x' = \ (S.Fun (\ f') (\ x'))" using x' 1 by auto also have "... = \ (\a f' (\ x'))" using Fun_\ x' assms S'.set_subset_Univ bij_\ by metis also have "... = \ (\ (S'.Fun f' (\ (\ x'))))" proof - have "\ (\a f' (\ x')) = \ (\ (S'.Fun f' x'))" proof - have "x' \ S'.Univ" by (meson S'.ide_dom S'.set_subset_Univ assms subsetCE x') thus ?thesis by (simp add: INV.\a_def INV.\_\ x') qed also have "... = \ (\ (S'.Fun f' (\ (\ x'))))" using assms x' \_\ S'.set_subset_Univ S'.ide_dom by (metis subsetCE) finally show ?thesis by auto qed also have "... = S'.Fun f' x'" proof - have 2: "\x'. x' \ S'.Univ \ \ (\ x') = x'" using bij_\ bij_betw_inv_into_right by fast have "S'.Fun f' (\ (\ x')) = S'.Fun f' x'" using assms x' 2 S'.set_subset_Univ S'.ide_dom by (metis subsetCE) moreover have "S'.Fun f' x' \ S'.Univ" using x' assms S'.Fun_mapsto S'.set_subset_Univ S'.ide_cod by blast ultimately show ?thesis using 2 by auto qed finally show ?thesis by auto qed ultimately show "S'.Fun (\ (\ f')) x' = S'.Fun f' x'" by auto qed qed qed lemma inverse_functors_\_\: shows "inverse_functors S S' \ \" proof - interpret \\: composite_functor S S' S \ \ .. have inv: "\ o \ = S.map" using \\ S.map_def \\.is_extensional by auto interpret \\: composite_functor S' S S' \ \ .. have inv': "\ o \ = S'.map" using \\ S'.map_def \\.is_extensional by auto show ?thesis using inv inv' by (unfold_locales, auto) qed lemma are_isomorphic: shows "\\. invertible_functor S S' \ \ (\m. S.incl m \ S'.incl (\ m))" proof - interpret inverse_functors S S' \ \ using inverse_functors_\_\ by auto have 1: "inverse_functors S S' \ \" .. interpret invertible_functor S S' \ apply unfold_locales using 1 by auto have "invertible_functor S S' \" .. thus ?thesis using \_preserves_incl by auto qed end (* * The main result: set_category is categorical, in the following (logical) sense: * If S and S' are two "set categories", and if the sets of terminal objects of S and S' * are in bijective correspondence, then S and S' are isomorphic as categories, * via a functor that preserves inclusion maps, hence the inclusion relation between sets. *) theorem set_category_is_categorical: assumes "set_category S" and "set_category S'" and "bij_betw \ (set_category_data.Univ S) (set_category_data.Univ S')" shows "\\. invertible_functor S S' \ \ (\m. set_category.incl S m \ set_category.incl S' (\ m))" proof - interpret S: set_category S using assms(1) by auto interpret S': set_category S' using assms(2) by auto interpret two_set_categories_bij_betw_Univ S S' \ apply (unfold_locales) using assms(3) by auto show ?thesis using are_isomorphic by auto qed section "Further Properties of Set Categories" text\ In this section we further develop the consequences of the \set_category\ axioms, and establish characterizations of a number of standard category-theoretic notions for a \set_category\. \ context set_category begin abbreviation Dom where "Dom f \ set (dom f)" abbreviation Cod where "Cod f \ set (cod f)" subsection "Initial Object" text\ The object corresponding to the empty set is an initial object. \ definition empty where "empty = mkIde {}" lemma initial_empty: shows "initial empty" proof show 0: "ide empty" using empty_def by auto show "\b. ide b \ \!f. \f : empty \ b\" proof - fix b assume b: "ide b" show "\!f. \f : empty \ b\" proof show 1: "\mkArr {} (set b) (\x. x) : empty \ b\" using b empty_def mkArr_in_hom mkIde_set set_subset_Univ by (metis 0 Pi_empty UNIV_I arr_mkIde) show "\f. \f : empty \ b\ \ f = mkArr {} (set b) (\x. x)" proof - fix f assume f: "\f : empty \ b\" show "f = mkArr {} (set b) (\x. x)" proof (intro arr_eqI) show 1: "par f (mkArr {} (set b) (\x. x))" using 1 f by force show "Fun f = Fun (mkArr {} (set b) (\x. x))" using empty_def 1 f Fun_mapsto by fastforce qed qed qed qed qed subsection "Identity Arrows" text\ Identity arrows correspond to restrictions of the identity function. \ lemma ide_char: assumes "arr f" shows "ide f \ Dom f = Cod f \ Fun f = (\x \ Dom f. x)" using assms mkIde_as_mkArr mkArr_Fun Fun_ide in_homE ide_cod mkArr_Fun mkIde_set by (metis ide_char) lemma ideI: assumes "arr f" and "Dom f = Cod f" and "\x. x \ Dom f \ Fun f x = x" shows "ide f" proof - have "Fun f = (\x \ Dom f. x)" using assms Fun_def by auto thus ?thesis using assms ide_char by blast qed subsection "Inclusions" lemma ide_implies_incl: assumes "ide a" shows "incl a" proof - have "arr a \ Dom a \ Cod a" using assms by auto moreover have "a = mkArr (Dom a) (Cod a) (\x. x)" using assms by simp ultimately show ?thesis using incl_def by simp qed definition incl_in :: "'s \ 's \ bool" where "incl_in a b = (ide a \ ide b \ set a \ set b)" abbreviation incl_of where "incl_of a b \ mkArr (set a) (set b) (\x. x)" lemma elem_set_implies_set_eq_singleton: assumes "a \ set b" shows "set a = {a}" proof - have "ide b" using assms set_def by auto thus ?thesis using assms set_subset_Univ terminal_char2 by (metis mem_Collect_eq subsetCE) qed lemma elem_set_implies_incl_in: assumes "a \ set b" shows "incl_in a b" proof - have b: "ide b" using assms set_def by auto hence "set b \ Univ" by simp hence "a \ Univ \ set a \ set b" using assms elem_set_implies_set_eq_singleton by auto hence "ide a \ set a \ set b" using b terminal_char1 by simp thus ?thesis using b incl_in_def by simp qed lemma incl_incl_of [simp]: assumes "incl_in a b" shows "incl (incl_of a b)" and "\incl_of a b : a \ b\" proof - show "\incl_of a b : a \ b\" using assms incl_in_def mkArr_in_hom by (metis image_ident image_subset_iff_funcset mkIde_set set_subset_Univ) thus "incl (incl_of a b)" using assms incl_def incl_in_def by fastforce qed text\ There is at most one inclusion between any pair of objects. \ lemma incls_coherent: assumes "par f f'" and "incl f" and "incl f'" shows "f = f'" using assms incl_def fun_complete by auto text\ The set of inclusions is closed under composition. \ lemma incl_comp [simp]: assumes "incl f" and "incl g" and "cod f = dom g" shows "incl (g \ f)" proof - have 1: "seq g f" using assms incl_def by auto moreover have "Dom (g \ f) \ Cod (g \ f)" using assms 1 incl_def by auto moreover have "g \ f = mkArr (Dom f) (Cod g) (restrict (\x. x) (Dom f))" using assms 1 Fun_comp incl_def Fun_mkArr mkArr_Fun Fun_ide comp_cod_arr ide_dom dom_comp cod_comp by metis ultimately show ?thesis using incl_def by force qed subsection "Image Factorization" text\ The image of an arrow is the object that corresponds to the set-theoretic image of the domain set under the function induced by the arrow. \ abbreviation Img where "Img f \ Fun f ` Dom f" definition img where "img f = mkIde (Img f)" lemma ide_img [simp]: assumes "arr f" shows "ide (img f)" proof - have "Fun f ` Dom f \ Cod f" using assms Fun_mapsto by blast moreover have "Cod f \ Univ" using assms by simp ultimately show ?thesis using img_def by simp qed lemma set_img [simp]: assumes "arr f" shows "set (img f) = Img f" proof - have "Fun f ` set (dom f) \ set (cod f) \ set (cod f) \ Univ" using assms Fun_mapsto by auto hence "Fun f ` set (dom f) \ Univ" by auto thus ?thesis using assms img_def set_mkIde by auto qed lemma img_point_in_Univ: assumes "\x : unity \ a\" shows "img x \ Univ" proof - have "set (img x) = {Fun x unity}" using assms img_def terminal_unity terminal_char2 image_empty image_insert mem_Collect_eq set_img by force thus "img x \ Univ" using assms terminal_char1 by auto qed lemma incl_in_img_cod: assumes "arr f" shows "incl_in (img f) (cod f)" proof (unfold img_def) have 1: "Img f \ Cod f \ Cod f \ Univ" using assms Fun_mapsto by auto hence 2: "ide (mkIde (Img f))" by fastforce moreover have "ide (cod f)" using assms by auto moreover have "set (mkIde (Img f)) \ Cod f" using 1 2 by force ultimately show "incl_in (mkIde (Img f)) (cod f)" using incl_in_def by blast qed lemma img_point_elem_set: assumes "\x : unity \ a\" shows "img x \ set a" proof - have "incl_in (img x) a" using assms incl_in_img_cod by auto hence "set (img x) \ set a" using incl_in_def by blast moreover have "img x \ set (img x)" using assms img_point_in_Univ terminal_char2 by simp ultimately show ?thesis by auto qed text\ The corestriction of an arrow @{term f} is the arrow @{term "corestr f \ hom (dom f) (img f)"} that induces the same function on the universe as @{term f}. \ definition corestr where "corestr f = mkArr (Dom f) (Img f) (Fun f)" lemma corestr_in_hom: assumes "arr f" shows "\corestr f : dom f \ img f\" proof - have "Fun f \ Dom f \ Fun f ` Dom f \ Dom f \ Univ" using assms by auto moreover have "Fun f ` Dom f \ Univ" proof - have "Fun f ` Dom f \ Cod f \ Cod f \ Univ" using assms Fun_mapsto by auto thus ?thesis by blast qed ultimately have "mkArr (Dom f) (Fun f ` Dom f) (Fun f) \ hom (dom f) (img f)" using assms img_def mkArr_in_hom [of "Dom f" "Fun f ` Dom f" "Fun f"] by simp thus ?thesis using corestr_def by fastforce qed text\ Every arrow factors as a corestriction followed by an inclusion. \ lemma img_fact: assumes "arr f" shows "S (incl_of (img f) (cod f)) (corestr f) = f" proof (intro arr_eqI) have 1: "\corestr f : dom f \ img f\" using assms corestr_in_hom by blast moreover have 2: "\incl_of (img f) (cod f) : img f \ cod f\" using assms incl_in_img_cod incl_incl_of by fast ultimately show P: "par (incl_of (img f) (cod f) \ corestr f) f" using assms in_homE by blast show "Fun (incl_of (img f) (cod f) \ corestr f) = Fun f" proof - have "Fun (incl_of (img f) (cod f) \ corestr f) = restrict (Fun (incl_of (img f) (cod f)) o Fun (corestr f)) (Dom f)" using Fun_comp 1 2 P by auto also have "... = restrict (restrict (\x. x) (Img f) o restrict (Fun f) (Dom f)) (Dom f)" proof - have "Fun (corestr f) = restrict (Fun f) (Dom f)" using assms corestr_def Fun_mkArr corestr_in_hom by force moreover have "Fun (incl_of (img f) (cod f)) = restrict (\x. x) (Img f)" proof - have "arr (incl_of (img f) (cod f))" using incl_incl_of P by blast moreover have "incl_of (img f) (cod f) = mkArr (Img f) (Cod f) (\x. x)" using assms by fastforce ultimately show ?thesis using assms img_def Fun_mkArr by metis qed ultimately show ?thesis by argo qed also have "... = Fun f" proof fix x show "restrict (restrict (\x. x) (Img f) o restrict (Fun f) (Dom f)) (Dom f) x = Fun f x" using assms extensional_restrict Fun_mapsto extensional_arb [of "Fun f" "Dom f" x] by (cases "x \ Dom f", auto) qed finally show ?thesis by auto qed qed lemma Fun_corestr: assumes "arr f" shows "Fun (corestr f) = Fun f" proof - have 1: "f = incl_of (img f) (cod f) \ corestr f" using assms img_fact by auto hence 2: "Fun f = restrict (Fun (incl_of (img f) (cod f)) o Fun (corestr f)) (Dom f)" using assms by (metis Fun_comp dom_comp) also have "... = restrict (Fun (corestr f)) (Dom f)" using assms by (metis 1 2 Fun_mkArr seqE mkArr_Fun corestr_def) also have "... = Fun (corestr f)" using assms 1 by (metis Fun_def dom_comp extensional_restrict restrict_extensional) finally show ?thesis by auto qed subsection "Points and Terminal Objects" text\ To each element @{term t} of @{term "set a"} is associated a point @{term "mkPoint a t \ hom unity a"}. The function induced by such a point is the constant-@{term t} function on the set @{term "{unity}"}. \ definition mkPoint where "mkPoint a t \ mkArr {unity} (set a) (\_. t)" lemma mkPoint_in_hom: assumes "ide a" and "t \ set a" shows "\mkPoint a t : unity \ a\" using assms mkArr_in_hom by (metis Pi_I mkIde_set set_subset_Univ terminal_char2 terminal_unity mkPoint_def) lemma Fun_mkPoint: assumes "ide a" and "t \ set a" shows "Fun (mkPoint a t) = (\_ \ {unity}. t)" using assms mkPoint_def terminal_unity by force text\ For each object @{term a} the function @{term "mkPoint a"} has as its inverse the restriction of the function @{term img} to @{term "hom unity a"} \ lemma mkPoint_img: shows "img \ hom unity a \ set a" and "\x. \x : unity \ a\ \ mkPoint a (img x) = x" proof - show "img \ hom unity a \ set a" using img_point_elem_set by simp show "\x. \x : unity \ a\ \ mkPoint a (img x) = x" proof - fix x assume x: "\x : unity \ a\" show "mkPoint a (img x) = x" proof (intro arr_eqI) have 0: "img x \ set a" using x img_point_elem_set by metis hence 1: "mkPoint a (img x) \ hom unity a" using x mkPoint_in_hom by force thus 2: "par (mkPoint a (img x)) x" using x by fastforce have "Fun (mkPoint a (img x)) = (\_ \ {unity}. img x)" using 1 mkPoint_def by auto also have "... = Fun x" proof fix z have "z \ unity \ (\_ \ {unity}. img x) z = Fun x z" using x Fun_mapsto Fun_def restrict_apply singletonD terminal_char2 terminal_unity by auto moreover have "(\_ \ {unity}. img x) unity = Fun x unity" using x 0 elem_set_implies_set_eq_singleton set_img terminal_char2 terminal_unity by (metis 2 image_insert in_homE restrict_apply singletonI singleton_insert_inj_eq) ultimately show "(\_ \ {unity}. img x) z = Fun x z" by auto qed finally show "Fun (mkPoint a (img x)) = Fun x" by auto qed qed qed lemma img_mkPoint: assumes "ide a" shows "mkPoint a \ set a \ hom unity a" and "\t. t \ set a \ img (mkPoint a t) = t" proof - show "mkPoint a \ set a \ hom unity a" using assms(1) mkPoint_in_hom by simp show "\t. t \ set a \ img (mkPoint a t) = t" proof - fix t assume t: "t \ set a" show "img (mkPoint a t) = t" proof - have 1: "arr (mkPoint a t)" using assms t mkPoint_in_hom by auto have "Fun (mkPoint a t) ` {unity} = {t}" using 1 mkPoint_def by simp thus ?thesis by (metis 1 t elem_set_implies_incl_in elem_set_implies_set_eq_singleton img_def incl_in_def dom_mkArr mkIde_set terminal_char2 terminal_unity mkPoint_def) qed qed qed text\ For each object @{term a} the elements of @{term "hom unity a"} are therefore in bijective correspondence with @{term "set a"}. \ lemma bij_betw_points_and_set: assumes "ide a" shows "bij_betw img (hom unity a) (set a)" proof (intro bij_betwI) show "img \ hom unity a \ set a" using assms mkPoint_img by auto show "mkPoint a \ set a \ hom unity a" using assms img_mkPoint by auto show "\x. x \ hom unity a \ mkPoint a (img x) = x" using assms mkPoint_img by auto show "\t. t \ set a \ img (mkPoint a t) = t" using assms img_mkPoint by auto qed text\ The function on the universe induced by an arrow @{term f} agrees, under the bijection between @{term "hom unity (dom f)"} and @{term "Dom f"}, with the action of @{term f} by composition on @{term "hom unity (dom f)"}. \ lemma Fun_point: assumes "\x : unity \ a\" shows "Fun x = (\_ \ {unity}. img x)" using assms mkPoint_img img_mkPoint Fun_mkPoint [of a "img x"] img_point_elem_set by auto lemma comp_arr_mkPoint: assumes "arr f" and "t \ Dom f" shows "f \ mkPoint (dom f) t = mkPoint (cod f) (Fun f t)" proof (intro arr_eqI) have 0: "seq f (mkPoint (dom f) t)" using assms mkPoint_in_hom [of "dom f" t] by auto have 1: "\f \ mkPoint (dom f) t : unity \ cod f\" using assms mkPoint_in_hom [of "dom f" t] by auto show "par (f \ mkPoint (dom f) t) (mkPoint (cod f) (Fun f t))" proof - have "\mkPoint (cod f) (Fun f t) : unity \ cod f\" using assms Fun_mapsto mkPoint_in_hom [of "cod f" "Fun f t"] by auto thus ?thesis using 1 by fastforce qed show "Fun (f \ mkPoint (dom f) t) = Fun (mkPoint (cod f) (Fun f t))" proof - have "Fun (f \ mkPoint (dom f) t) = restrict (Fun f o Fun (mkPoint (dom f) t)) {unity}" using assms 0 1 Fun_comp terminal_char2 terminal_unity by auto also have "... = (\_ \ {unity}. Fun f t)" using assms Fun_mkPoint by auto also have "... = Fun (mkPoint (cod f) (Fun f t))" using assms Fun_mkPoint [of "cod f" "Fun f t"] Fun_mapsto by fastforce finally show ?thesis by auto qed qed lemma comp_arr_point: assumes "arr f" and "\x : unity \ dom f\" shows "f \ x = mkPoint (cod f) (Fun f (img x))" proof - have "x = mkPoint (dom f) (img x)" using assms mkPoint_img by simp thus ?thesis using assms comp_arr_mkPoint [of f "img x"] by (simp add: img_point_elem_set) qed text\ This agreement allows us to express @{term "Fun f"} in terms of composition. \ lemma Fun_in_terms_of_comp: assumes "arr f" shows "Fun f = restrict (img o S f o mkPoint (dom f)) (Dom f)" proof fix t have "t \ Dom f \ Fun f t = restrict (img o S f o mkPoint (dom f)) (Dom f) t" using assms by (simp add: Fun_def) moreover have "t \ Dom f \ Fun f t = restrict (img o S f o mkPoint (dom f)) (Dom f) t" proof - assume t: "t \ Dom f" have 1: "f \ mkPoint (dom f) t = mkPoint (cod f) (Fun f t)" using assms t comp_arr_mkPoint by simp hence "img (f \ mkPoint (dom f) t) = img (mkPoint (cod f) (Fun f t))" by simp thus ?thesis proof - have "Fun f t \ Cod f" using assms t Fun_mapsto by auto thus ?thesis using assms t 1 img_mkPoint by auto qed qed ultimately show "Fun f t = restrict (img o S f o mkPoint (dom f)) (Dom f) t" by auto qed text\ We therefore obtain a rule for proving parallel arrows equal by showing that they have the same action by composition on points. \ (* * TODO: Find out why attempts to use this as the main rule for a proof loop * unless the specific instance is given. *) lemma arr_eqI': assumes "par f f'" and "\x. \x : unity \ dom f\ \ f \ x = f' \ x" shows "f = f'" using assms Fun_in_terms_of_comp mkPoint_in_hom by (intro arr_eqI, auto) text\ An arrow can therefore be specified by giving its action by composition on points. In many situations, this is more natural than specifying it as a function on the universe. \ definition mkArr' where "mkArr' a b F = mkArr (set a) (set b) (img o F o mkPoint a)" lemma mkArr'_in_hom: assumes "ide a" and "ide b" and "F \ hom unity a \ hom unity b" shows "\mkArr' a b F : a \ b\" proof - have "img o F o mkPoint a \ set a \ set b" proof fix t assume t: "t \ set a" thus "(img o F o mkPoint a) t \ set b" using assms mkPoint_in_hom img_point_elem_set [of "F (mkPoint a t)" b] by auto qed thus ?thesis using assms mkArr'_def mkArr_in_hom [of "set a" "set b"] by simp qed lemma comp_point_mkArr': assumes "ide a" and "ide b" and "F \ hom unity a \ hom unity b" shows "\x. \x : unity \ a\ \ mkArr' a b F \ x = F x" proof - fix x assume x: "\x : unity \ a\" have "Fun (mkArr' a b F) (img x) = img (F x)" unfolding mkArr'_def using assms x Fun_mkArr arr_mkArr img_point_elem_set mkPoint_img mkPoint_in_hom by (simp add: Pi_iff) hence "mkArr' a b F \ x = mkPoint b (img (F x))" using assms x mkArr'_in_hom [of a b F] comp_arr_point by auto thus "mkArr' a b F \ x = F x" using assms x mkPoint_img(2) by auto qed text\ A third characterization of terminal objects is as those objects whose set of points is a singleton. \ lemma terminal_char3: assumes "\!x. \x : unity \ a\" shows "terminal a" proof - have a: "ide a" using assms ide_cod mem_Collect_eq by blast hence 1: "bij_betw img (hom unity a) (set a)" using assms bij_betw_points_and_set by auto hence "img ` (hom unity a) = set a" by (simp add: bij_betw_def) moreover have "hom unity a = {THE x. x \ hom unity a}" using assms theI' [of "\x. x \ hom unity a"] by auto ultimately have "set a = {img (THE x. x \ hom unity a)}" by (metis image_empty image_insert) thus ?thesis using a terminal_char1 by simp qed text\ The following is an alternative formulation of functional completeness, which says that any function on points uniquely determines an arrow. \ lemma fun_complete': assumes "ide a" and "ide b" and "F \ hom unity a \ hom unity b" shows "\!f. \f : a \ b\ \ (\x. \x : unity \ a\ \ f \ x = F x)" proof have 1: "\mkArr' a b F : a \ b\" using assms mkArr'_in_hom by auto moreover have 2: "\x. \x : unity \ a\ \ mkArr' a b F \ x = F x" using assms comp_point_mkArr' by auto ultimately show "\mkArr' a b F : a \ b\ \ (\x. \x : unity \ a\ \ mkArr' a b F \ x = F x)" by blast fix f assume f: "\f : a \ b\ \ (\x. \x : unity \ a\ \ f \ x = F x)" show "f = mkArr' a b F" using f 1 2 by (intro arr_eqI' [of f "mkArr' a b F"], fastforce, auto) qed subsection "The `Determines Same Function' Relation on Arrows" text\ An important part of understanding the structure of a category of sets and functions is to characterize when it is that two arrows ``determine the same function''. The following result provides one answer to this: two arrows with a common domain determine the same function if and only if they can be rendered equal by composing with a cospan of inclusions. \ lemma eq_Fun_iff_incl_joinable: assumes "span f f'" shows "Fun f = Fun f' \ (\m m'. incl m \ incl m' \ seq m f \ seq m' f' \ m \ f = m' \ f')" proof assume ff': "Fun f = Fun f'" let ?b = "mkIde (Cod f \ Cod f')" let ?m = "incl_of (cod f) ?b" let ?m' = "incl_of (cod f') ?b" have "incl ?m" using assms incl_incl_of [of "cod f" ?b] incl_in_def by simp have "incl ?m'" using assms incl_incl_of [of "cod f'" ?b] incl_in_def by simp have m: "?m = mkArr (Cod f) (Cod f \ Cod f') (\x. x)" by (simp add: assms) have m': "?m' = mkArr (Cod f') (Cod f \ Cod f') (\x. x)" by (simp add: assms) have seq: "seq ?m f \ seq ?m' f'" using assms m m' by simp have "?m \ f = ?m' \ f'" proof (intro arr_eqI) show par: "par (?m \ f) (?m' \ f')" using assms m m' by simp show "Fun (?m \ f) = Fun (?m' \ f')" using assms seq par ff' Fun_mapsto Fun_comp seqE by (metis Fun_ide Fun_mkArr comp_cod_arr ide_cod) qed hence "incl ?m \ incl ?m' \ seq ?m f \ seq ?m' f' \ ?m \ f = ?m' \ f'" using seq \incl ?m\ \incl ?m'\ by simp thus "\m m'. incl m \ incl m' \ seq m f \ seq m' f' \ m \ f = m' \ f'" by auto next assume ff': "\m m'. incl m \ incl m' \ seq m f \ seq m' f' \ m \ f = m' \ f'" show "Fun f = Fun f'" proof - from ff' obtain m m' where mm': "incl m \ incl m' \ seq m f \ seq m' f' \ m \ f = m' \ f'" by blast show ?thesis using ff' mm' Fun_incl seqE by (metis Fun_comp Fun_ide comp_cod_arr ide_cod) qed qed text\ Another answer to the same question: two arrows with a common domain determine the same function if and only if their corestrictions are equal. \ lemma eq_Fun_iff_eq_corestr: assumes "span f f'" shows "Fun f = Fun f' \ corestr f = corestr f'" using assms corestr_def Fun_corestr by metis subsection "Retractions, Sections, and Isomorphisms" text\ An arrow is a retraction if and only if its image coincides with its codomain. \ lemma retraction_if_Img_eq_Cod: assumes "arr g" and "Img g = Cod g" shows "retraction g" and "ide (g \ mkArr (Cod g) (Dom g) (inv_into (Dom g) (Fun g)))" proof - let ?F = "inv_into (Dom g) (Fun g)" let ?f = "mkArr (Cod g) (Dom g) ?F" have f: "arr ?f" proof have "Cod g \ Univ \ Dom g \ Univ" using assms by auto moreover have "?F \ Cod g \ Dom g" proof fix y assume y: "y \ Cod g" let ?P = "\x. x \ Dom g \ Fun g x = y" have "\x. ?P x" using y assms by force hence "?P (SOME x. ?P x)" using someI_ex [of ?P] by fast hence "?P (?F y)" using Hilbert_Choice.inv_into_def by metis thus "?F y \ Dom g" by auto qed ultimately show "Cod g \ Univ \ Dom g \ Univ \ ?F \ Cod g \ Dom g" by auto qed show "ide (g \ ?f)" proof - have "g = mkArr (Dom g) (Cod g) (Fun g)" using assms by auto hence "g \ ?f = mkArr (Cod g) (Cod g) (Fun g o ?F)" using assms(1) f comp_mkArr by metis moreover have "mkArr (Cod g) (Cod g) (\y. y) = ..." proof (intro mkArr_eqI') show "arr (mkArr (Cod g) (Cod g) (\y. y))" using assms arr_cod_iff_arr by auto show "\y. y \ Cod g \ y = (Fun g o ?F) y" using assms by (simp add: f_inv_into_f) qed ultimately show ?thesis using assms f by auto qed thus "retraction g" by auto qed lemma retraction_char: shows "retraction g \ arr g \ Img g = Cod g" proof assume G: "retraction g" show "arr g \ Img g = Cod g" proof show "arr g" using G by blast show "Img g = Cod g" proof - from G obtain f where f: "ide (g \ f)" by blast have "restrict (Fun g o Fun f) (Cod g) = restrict (\x. x) (Cod g)" using f Fun_comp Fun_ide ide_compE by metis hence "Fun g ` Fun f ` Cod g = Cod g" by (metis image_comp image_ident image_restrict_eq) moreover have "Fun f ` Cod g \ Dom g" using f Fun_mapsto arr_mkArr mkArr_Fun funcset_image by (metis seqE ide_compE ide_compE) moreover have "Img g \ Cod g" using f Fun_mapsto by blast ultimately show ?thesis by blast qed qed next assume "arr g \ Img g = Cod g" thus "retraction g" using retraction_if_Img_eq_Cod by blast qed text\ Every corestriction is a retraction. \ lemma retraction_corestr: assumes "arr f" shows "retraction (corestr f)" using assms retraction_char Fun_corestr corestr_in_hom by fastforce text\ An arrow is a section if and only if it induces an injective function on its domain, except in the special case that it has an empty domain set and a nonempty codomain set. \ lemma section_if_inj: assumes "arr f" and "inj_on (Fun f) (Dom f)" and "Dom f = {} \ Cod f = {}" shows "section f" and "ide (mkArr (Cod f) (Dom f) (\y. if y \ Img f then SOME x. x \ Dom f \ Fun f x = y else SOME x. x \ Dom f) \ f)" proof - let ?P= "\y. \x. x \ Dom f \ Fun f x = y" let ?G = "\y. if y \ Img f then SOME x. ?P y x else SOME x. x \ Dom f" let ?g = "mkArr (Cod f) (Dom f) ?G" have g: "arr ?g" proof - have 1: "Cod f \ Univ" using assms by simp have 2: "Dom f \ Univ" using assms by simp have 3: "?G \ Cod f \ Dom f" proof fix y assume Y: "y \ Cod f" show "?G y \ Dom f" proof (cases "y \ Img f") assume "y \ Img f" hence "(\x. ?P y x) \ ?G y = (SOME x. ?P y x)" using Y by auto hence "?P y (?G y)" using someI_ex [of "?P y"] by argo thus "?G y \ Dom f" by auto next assume "y \ Img f" hence "(\x. x \ Dom f) \ ?G y = (SOME x. x \ Dom f)" using assms Y by auto thus "?G y \ Dom f" using someI_ex [of "\x. x \ Dom f"] by argo qed qed show ?thesis using 1 2 3 by simp qed show "ide (?g \ f)" proof - have "f = mkArr (Dom f) (Cod f) (Fun f)" using assms by auto hence "?g \ f = mkArr (Dom f) (Dom f) (?G o Fun f)" using assms(1) g comp_mkArr [of "Dom f" "Cod f" "Fun f" "Dom f" ?G] by argo moreover have "mkArr (Dom f) (Dom f) (\x. x) = ..." proof (intro mkArr_eqI') show "arr (mkArr (Dom f) (Dom f) (\x. x))" using assms by auto show "\x. x \ Dom f \ x = (?G o Fun f) x" proof - fix x assume x: "x \ Dom f" have "Fun f x \ Img f" using x by blast hence *: "(\x'. ?P (Fun f x) x') \ ?G (Fun f x) = (SOME x'. ?P (Fun f x) x')" by auto then have "?P (Fun f x) (?G (Fun f x))" using someI_ex [of "?P (Fun f x)"] by argo with * have "x = ?G (Fun f x)" using assms x inj_on_def [of "Fun f" "Dom f"] by simp thus "x = (?G o Fun f) x" by simp qed qed ultimately show ?thesis using assms by auto qed thus "section f" by auto qed lemma section_char: shows "section f \ arr f \ (Dom f = {} \ Cod f = {}) \ inj_on (Fun f) (Dom f)" proof assume f: "section f" from f obtain g where g: "ide (g \ f)" using section_def by blast show "arr f \ (Dom f = {} \ Cod f = {}) \ inj_on (Fun f) (Dom f)" proof - have "arr f" using f by blast moreover have "Dom f = {} \ Cod f = {}" proof - have "Cod f \ {} \ Dom f \ {}" proof assume "Cod f \ {}" from this obtain y where "y \ Cod f" by blast hence "Fun g y \ Dom f" using g Fun_mapsto by (metis seqE ide_compE image_eqI retractionI retraction_char) thus "Dom f \ {}" by blast qed thus ?thesis by auto qed moreover have "inj_on (Fun f) (Dom f)" proof - have "restrict (Fun g o Fun f) (Dom f) = Fun (g \ f)" using g Fun_comp by (metis Fun_comp ide_compE) also have "... = restrict (\x. x) (Dom f)" using g Fun_ide by auto finally have "restrict (Fun g o Fun f) (Dom f) = restrict (\x. x) (Dom f)" by auto thus ?thesis using inj_onI inj_on_imageI2 inj_on_restrict_eq by metis qed ultimately show ?thesis by auto qed next assume F: "arr f \ (Dom f = {} \ Cod f = {}) \ inj_on (Fun f) (Dom f)" thus "section f" using section_if_inj by auto qed text\ Section-retraction pairs can also be characterized by an inverse relationship between the functions they induce. \ lemma section_retraction_char: shows "ide (g \ f) \ antipar f g \ compose (Dom f) (Fun g) (Fun f) = (\x \ Dom f. x)" proof show "ide (g \ f) \ antipar f g \ compose (Dom f) (Fun g) (Fun f) = (\x \ Dom f. x)" proof - assume fg: "ide (g \ f)" have 1: "antipar f g" using fg by force moreover have "compose (Dom f) (Fun g) (Fun f) = (\x \ Dom f. x)" proof fix x have "x \ Dom f \ compose (Dom f) (Fun g) (Fun f) x = (\x \ Dom f. x) x" by (simp add: compose_def) moreover have "x \ Dom f \ compose (Dom f) (Fun g) (Fun f) x = (\x \ Dom f. x) x" using fg 1 Fun_comp by (metis Fun_comp Fun_ide compose_eq' ide_compE) ultimately show "compose (Dom f) (Fun g) (Fun f) x = (\x \ Dom f. x) x" by auto qed ultimately show ?thesis by auto qed show "antipar f g \ compose (Dom f) (Fun g) (Fun f) = (\x \ Dom f. x) \ ide (g \ f)" proof - assume fg: "antipar f g \ compose (Dom f) (Fun g) (Fun f) = (\x \ Dom f. x)" show "ide (g \ f)" proof - have 1: "arr (g \ f)" using fg by auto moreover have "Dom (g \ f) = Cod (S g f)" using fg 1 by force moreover have "Fun (g \ f) = (\x \ Dom (g \ f). x)" using fg 1 by force ultimately show ?thesis using 1 ide_char by blast qed qed qed text\ Antiparallel arrows @{term f} and @{term g} are inverses if the functions they induce are inverses. \ lemma inverse_arrows_char: shows "inverse_arrows f g \ antipar f g \ compose (Dom f) (Fun g) (Fun f) = (\x \ Dom f. x) \ compose (Dom g) (Fun f) (Fun g) = (\y \ Dom g. y)" using section_retraction_char by blast text\ An arrow is an isomorphism if and only if the function it induces is a bijection. \ lemma iso_char: shows "iso f \ arr f \ bij_betw (Fun f) (Dom f) (Cod f)" proof - have "iso f \ section f \ retraction f" using iso_iff_section_and_retraction by auto also have "... \ arr f \ inj_on (Fun f) (Dom f) \ Img f = Cod f" using section_char retraction_char by force also have "... \ arr f \ bij_betw (Fun f) (Dom f) (Cod f)" using inj_on_def bij_betw_def [of "Fun f" "Dom f" "Cod f"] by meson finally show ?thesis by auto qed text\ The inverse of an isomorphism is constructed by inverting the induced function. \ lemma inv_char: assumes "iso f" shows "inv f = mkArr (Cod f) (Dom f) (inv_into (Dom f) (Fun f))" proof - let ?g = "mkArr (Cod f) (Dom f) (inv_into (Dom f) (Fun f))" have "ide (f \ ?g)" using assms iso_is_retraction retraction_char retraction_if_Img_eq_Cod by simp moreover have "ide (?g \ f)" proof - let ?g' = "mkArr (Cod f) (Dom f) (\y. if y \ Img f then SOME x. x \ Dom f \ Fun f x = y else SOME x. x \ Dom f)" have 1: "ide (?g' \ f)" using assms iso_is_section section_char section_if_inj by simp moreover have "?g' = ?g" proof show "arr ?g'" using 1 ide_compE by blast show "\y. y \ Cod f \ (if y \ Img f then SOME x. x \ Dom f \ Fun f x = y else SOME x. x \ Dom f) = inv_into (Dom f) (Fun f) y" proof - fix y assume "y \ Cod f" hence "y \ Img f" using assms iso_is_retraction retraction_char by metis thus "(if y \ Img f then SOME x. x \ Dom f \ Fun f x = y else SOME x. x \ Dom f) = inv_into (Dom f) (Fun f) y" using inv_into_def by metis qed qed ultimately show ?thesis by auto qed ultimately have "inverse_arrows f ?g" by auto thus ?thesis using inverse_unique by blast qed lemma Fun_inv: assumes "iso f" shows "Fun (inv f) = restrict (inv_into (Dom f) (Fun f)) (Cod f)" using assms inv_in_hom inv_char iso_inv_iso iso_is_arr Fun_mkArr by metis subsection "Monomorphisms and Epimorphisms" text\ An arrow is a monomorphism if and only if the function it induces is injective. \ lemma mono_char: shows "mono f \ arr f \ inj_on (Fun f) (Dom f)" proof assume f: "mono f" hence "arr f" using mono_def by auto moreover have "inj_on (Fun f) (Dom f)" proof (intro inj_onI) have 0: "inj_on (S f) (hom unity (dom f))" proof - have "hom unity (dom f) \ {g. seq f g}" using f mono_def arrI by auto hence "\A. hom unity (dom f) \ A \ inj_on (S f) A" using f mono_def by auto thus ?thesis by (meson subset_inj_on) qed fix x x' assume x: "x \ Dom f" and x': "x' \ Dom f" and xx': "Fun f x = Fun f x'" have 1: "mkPoint (dom f) x \ hom unity (dom f) \ mkPoint (dom f) x' \ hom unity (dom f)" using x x' \arr f\ mkPoint_in_hom by simp have "f \ mkPoint (dom f) x = f \ mkPoint (dom f) x'" using \arr f\ x x' xx' comp_arr_mkPoint by simp hence "mkPoint (dom f) x = mkPoint (dom f) x'" using 0 1 inj_onD [of "S f" "hom unity (dom f)" "mkPoint (dom f) x"] by simp thus "x = x'" using \arr f\ x x' img_mkPoint(2) img_mkPoint(2) ide_dom by metis qed ultimately show "arr f \ inj_on (Fun f) (Dom f)" by auto next assume f: "arr f \ inj_on (Fun f) (Dom f)" show "mono f" proof show "arr f" using f by auto show "\g g'. seq f g \ seq f g' \ f \ g = f \ g' \ g = g'" proof - fix g g' assume gg': "seq f g \ seq f g' \ f \ g = f \ g'" show "g = g'" proof (intro arr_eqI) show par: "par g g'" using gg' dom_comp by (metis seqE) show "Fun g = Fun g'" proof fix x have "x \ Dom g \ Fun g x = Fun g' x" using gg' by (simp add: par Fun_def) moreover have "x \ Dom g \ Fun g x = Fun g' x" proof - assume x: "x \ Dom g" have "Fun f (Fun g x) = Fun (f \ g) x" using gg' x Fun_comp [of f g] by auto also have "... = Fun f (Fun g' x)" using par f gg' x monoE by simp finally have "Fun f (Fun g x) = Fun f (Fun g' x)" by auto moreover have "Fun g x \ Dom f \ Fun g' x \ Dom f" using par gg' x Fun_mapsto by fastforce ultimately show "Fun g x = Fun g' x" using f gg' inj_onD [of "Fun f" "Dom f" "Fun g x" "Fun g' x"] by simp qed ultimately show "Fun g x = Fun g' x" by auto qed qed qed qed qed text\ Inclusions are monomorphisms. \ lemma mono_imp_incl: assumes "incl f" shows "mono f" using assms incl_def Fun_incl mono_char by auto text\ A monomorphism is a section, except in case it has an empty domain set and a nonempty codomain set. \ lemma mono_imp_section: assumes "mono f" and "Dom f = {} \ Cod f = {}" shows "section f" using assms mono_char section_char by auto text\ An arrow is an epimorphism if and only if either its image coincides with its codomain, or else the universe has only a single element (in which case all arrows are epimorphisms). \ lemma epi_char: shows "epi f \ arr f \ (Img f = Cod f \ (\t t'. t \ Univ \ t' \ Univ \ t = t'))" proof assume epi: "epi f" show "arr f \ (Img f = Cod f \ (\t t'. t \ Univ \ t' \ Univ \ t = t'))" proof - have f: "arr f" using epi epi_implies_arr by auto moreover have "\(\t t'. t \ Univ \ t' \ Univ \ t = t') \ Img f = Cod f" proof - assume "\(\t t'. t \ Univ \ t' \ Univ \ t = t')" from this obtain tt and ff where B: "tt \ Univ \ ff \ Univ \ tt \ ff" by blast show "Img f = Cod f" proof show "Img f \ Cod f" using f Fun_mapsto by auto show "Cod f \ Img f" proof let ?g = "mkArr (Cod f) {ff, tt} (\y. tt)" let ?g' = "mkArr (Cod f) {ff, tt} (\y. if \x. x \ Dom f \ Fun f x = y then tt else ff)" let ?b = "mkIde {ff, tt}" have g: "\?g : cod f \ ?b\ \ Fun ?g = (\y \ Cod f. tt)" using f B in_homI [of ?g] by simp have g': "?g' \ hom (cod f) ?b \ Fun ?g' = (\y \ Cod f. if \x. x \ Dom f \ Fun f x = y then tt else ff)" using f B in_homI [of ?g'] by simp have "?g \ f = ?g' \ f" proof (intro arr_eqI) show "par (?g \ f) (?g' \ f)" using f g g' by auto show "Fun (?g \ f) = Fun (?g' \ f)" using f g g' Fun_comp comp_mkArr by force qed hence gg': "?g = ?g'" using epi f g g' epiE [of f ?g ?g'] by fastforce fix y assume y: "y \ Cod f" have "Fun ?g' y = tt" using gg' g y by simp hence "(if \x. x \ Dom f \ Fun f x = y then tt else ff) = tt" using g' y by simp hence "\x. x \ Dom f \ Fun f x = y" using B by argo thus "y \ Img f" by blast qed qed qed ultimately show "arr f \ (Img f = Cod f \ (\t t'. t \ Univ \ t' \ Univ \ t = t'))" by fast qed next show "arr f \ (Img f = Cod f \ (\t t'. t \ Univ \ t' \ Univ \ t = t')) \ epi f" proof - have "arr f \ Img f = Cod f \ epi f" proof - assume f: "arr f \ Img f = Cod f" show "epi f" using f arr_eqI' epiE retractionI retraction_if_Img_eq_Cod retraction_is_epi by meson qed moreover have "arr f \ (\t t'. t \ Univ \ t' \ Univ \ t = t') \ epi f" proof - assume f: "arr f \ (\t t'. t \ Univ \ t' \ Univ \ t = t')" have "\f f'. par f f' \ f = f'" proof - fix f f' assume ff': "par f f'" show "f = f'" proof (intro arr_eqI) show "par f f'" using ff' by simp have "\t t'. t \ Cod f \ t' \ Cod f \ t = t'" using f ff' set_subset_Univ ide_cod subsetD by blast thus "Fun f = Fun f'" using ff' Fun_mapsto [of f] Fun_mapsto [of f'] extensional_arb [of "Fun f" "Dom f"] extensional_arb [of "Fun f'" "Dom f"] by fastforce qed qed moreover have "\g g'. par (g \ f) (g' \ f) \ par g g'" by force ultimately show "epi f" using f by (intro epiI; metis) qed ultimately show "arr f \ (Img f = Cod f \ (\t t'. t \ Univ \ t' \ Univ \ t = t')) \ epi f" by auto qed qed text\ An epimorphism is a retraction, except in the case of a degenerate universe with only a single element. \ lemma epi_imp_retraction: assumes "epi f" and "\t t'. t \ Univ \ t' \ Univ \ t \ t'" shows "retraction f" using assms epi_char retraction_char by auto text\ Retraction/inclusion factorization is unique (not just up to isomorphism -- remember that the notion of inclusion is not categorical but depends on the arbitrarily chosen @{term img}). \ lemma unique_retr_incl_fact: assumes "seq m e" and "seq m' e'" and "m \ e = m' \ e'" and "incl m" and "incl m'" and "retraction e" and "retraction e'" shows "m = m'" and "e = e'" proof - have 1: "cod m = cod m' \ dom e = dom e'" using assms(1-3) by (metis dom_comp cod_comp) hence 2: "span e e'" using assms(1-2) by blast hence 3: "Fun e = Fun e'" using assms eq_Fun_iff_incl_joinable by meson hence "img e = img e'" using assms 1 img_def by auto moreover have "img e = cod e \ img e' = cod e'" using assms(6-7) retraction_char img_def by simp ultimately have "par e e'" using 2 by simp thus "e = e'" using 3 arr_eqI by blast hence "par m m'" using assms(1) assms(2) 1 by fastforce thus "m = m'" using assms(4) assms(5) incls_coherent by blast qed end section "Concrete Set Categories" text\ The \set_category\ locale is useful for stating results that depend on a category of @{typ 'a}-sets and functions, without having to commit to a particular element type @{typ 'a}. However, in applications we often need to work with a category of sets and functions that is guaranteed to contain sets corresponding to the subsets of some extrinsically given type @{typ 'a}. A \emph{concrete set category} is a set category \S\ that is equipped with an injective function @{term \} from type @{typ 'a} to \S.Univ\. The following locale serves to facilitate some of the technical aspects of passing back and forth between elements of type @{typ 'a} and the elements of \S.Univ\. \ locale concrete_set_category = set_category S for S :: "'s comp" (infixr "\\<^sub>S" 55) and U :: "'a set" and \ :: "'a \ 's" + assumes \_mapsto: "\ \ U \ Univ" and inj_\: "inj_on \ U" begin abbreviation \ where "\ \ inv_into U \" lemma \_mapsto: shows "\ \ \ ` U \ U" by (simp add: inv_into_into) lemma \_\ [simp]: assumes "x \ U" shows "\ (\ x) = x" using assms inj_\ inv_into_f_f by simp lemma \_\ [simp]: assumes "t \ \ ` U" shows "\ (\ t) = t" using assms o_def inj_\ by auto end end diff --git a/thys/CoreC++/Progress.thy b/thys/CoreC++/Progress.thy --- a/thys/CoreC++/Progress.thy +++ b/thys/CoreC++/Progress.thy @@ -1,1042 +1,1042 @@ (* Title: CoreC++ Author: Daniel Wasserrab Maintainer: Daniel Wasserrab Based on the Jinja theory J/Progress.thy by Tobias Nipkow *) section \Progress of Small Step Semantics\ theory Progress imports Equivalence DefAss Conform begin subsection \Some pre-definitions\ lemma final_refE: "\ P,E,h \ e : Class C; final e; \r. e = ref r \ Q; \r. e = Throw r \ Q \ \ Q" by (simp add:final_def,auto,case_tac v,auto) lemma finalRefE: "\ P,E,h \ e : T; is_refT T; final e; e = null \ Q; \r. e = ref r \ Q; \r. e = Throw r \ Q\ \ Q" apply (cases T) apply (simp add:is_refT_def)+ apply (simp add:final_def) apply (erule disjE) apply clarsimp apply (erule exE)+ apply fastforce apply (auto simp:final_def is_refT_def) apply (case_tac v) apply auto done lemma subE: "\ P \ T \ T'; is_type P T'; wf_prog wf_md P; \ T = T'; \C. T \ Class C \ \ Q; \C D. \ T = Class C; T' = Class D; P \ Path C to D unique \ \ Q; \C. \ T = NT; T' = Class C \ \ Q \ \ Q" apply(cases T') apply auto apply(drule_tac T = "T" in widen_Class) apply auto done lemma assumes wf:"wf_prog wf_md P" and typeof:" P \ typeof\<^bsub>h\<^esub> v = Some T'" and type:"is_type P T" shows sub_casts:"P \ T' \ T \ \v'. P \ T casts v to v'" proof(erule subE) from type show "is_type P T" . next from wf show "wf_prog wf_md P" . next assume "T' = T" and "\C. T' \ Class C" thus "\v'. P \ T casts v to v'" by(fastforce intro:casts_prim) next fix C D assume T':"T' = Class C" and T:"T = Class D" and path_unique:"P \ Path C to D unique" from T' typeof obtain a Cs where v:"v = Ref(a,Cs)" and last:"last Cs = C" by(auto dest!:typeof_Class_Subo) from last path_unique obtain Cs' where "P \ Path last Cs to D via Cs'" by(auto simp:path_unique_def path_via_def) hence "P \ Class D casts Ref(a,Cs) to Ref(a,Cs@\<^sub>pCs')" by -(rule casts_ref,simp_all) with T v show "\v'. P \ T casts v to v'" by auto next fix C assume "T' = NT" and T:"T = Class C" with typeof have "v = Null" by simp with T show "\v'. P \ T casts v to v'" by(fastforce intro:casts_null) qed text\Derivation of new induction scheme for well typing:\ inductive WTrt' :: "[prog,env,heap,expr, ty ] \ bool" - ("_,_,_ \ _ :' _" [51,51,51]50) + ("_,_,_ \ _ :'' _" [51,51,51]50) and WTrts':: "[prog,env,heap,expr list,ty list] \ bool" ("_,_,_ \ _ [:''] _" [51,51,51]50) for P :: prog where "is_class P C \ P,E,h \ new C :' Class C" | "\is_class P C; P,E,h \ e :' T; is_refT T\ \ P,E,h \ Cast C e :' Class C" | "\is_class P C; P,E,h \ e :' T; is_refT T\ \ P,E,h \ \C\e :' Class C" | "P \ typeof\<^bsub>h\<^esub> v = Some T \ P,E,h \ Val v :' T" | "E V = Some T \ P,E,h \ Var V :' T" | "\ P,E,h \ e\<^sub>1 :' T\<^sub>1; P,E,h \ e\<^sub>2 :' T\<^sub>2; case bop of Eq \ T = Boolean | Add \ T\<^sub>1 = Integer \ T\<^sub>2 = Integer \ T = Integer \ \ P,E,h \ e\<^sub>1 \bop\ e\<^sub>2 :' T" | "\ P,E,h \ Var V :' T; P,E,h \ e :' T' \<^cancel>\V \ This\; P \ T' \ T \ \ P,E,h \ V:=e :' T" | "\P,E,h \ e :' Class C; Cs \ []; P \ C has least F:T via Cs\ \ P,E,h \ e\F{Cs} :' T" | "P,E,h \ e :' NT \ P,E,h \ e\F{Cs} :' T" | "\P,E,h \ e\<^sub>1 :' Class C; Cs \ []; P \ C has least F:T via Cs; P,E,h \ e\<^sub>2 :' T'; P \ T' \ T \ \ P,E,h \ e\<^sub>1\F{Cs}:=e\<^sub>2 :' T" | "\ P,E,h \ e\<^sub>1:'NT; P,E,h \ e\<^sub>2 :' T'; P \ T' \ T \ \ P,E,h \ e\<^sub>1\F{Cs}:=e\<^sub>2 :' T" | "\ P,E,h \ e :' Class C; P \ C has least M = (Ts,T,m) via Cs; P,E,h \ es [:'] Ts'; P \ Ts' [\] Ts \ \ P,E,h \ e\M(es) :' T" | "\ P,E,h \ e :' Class C'; P \ Path C' to C unique; P \ C has least M = (Ts,T,m) via Cs; P,E,h \ es [:'] Ts'; P \ Ts' [\] Ts \ \ P,E,h \ e\(C::)M(es) :' T" | "\P,E,h \ e :' NT; P,E,h \ es [:'] Ts\ \ P,E,h \ Call e Copt M es :' T" | "\ P \ typeof\<^bsub>h\<^esub> v = Some T'; P,E(V\T),h \ e\<^sub>2 :' T\<^sub>2; P \ T' \ T; is_type P T \ \ P,E,h \ {V:T := Val v; e\<^sub>2} :' T\<^sub>2" | "\ P,E(V\T),h \ e :' T'; \ assigned V e; is_type P T \ \ P,E,h \ {V:T; e} :' T'" | "\ P,E,h \ e\<^sub>1 :' T\<^sub>1; P,E,h \ e\<^sub>2 :' T\<^sub>2 \ \ P,E,h \ e\<^sub>1;;e\<^sub>2 :' T\<^sub>2" | "\ P,E,h \ e :' Boolean; P,E,h \ e\<^sub>1:' T; P,E,h \ e\<^sub>2:' T \ \ P,E,h \ if (e) e\<^sub>1 else e\<^sub>2 :' T" | "\ P,E,h \ e :' Boolean; P,E,h \ c:' T \ \ P,E,h \ while(e) c :' Void" | "\ P,E,h \ e :' T'; is_refT T'\ \ P,E,h \ throw e :' T" | "P,E,h \ [] [:'] []" | "\ P,E,h \ e :' T; P,E,h \ es [:'] Ts \ \ P,E,h \ e#es [:'] T#Ts" lemmas WTrt'_induct = WTrt'_WTrts'.induct [split_format (complete)] and WTrt'_inducts = WTrt'_WTrts'.inducts [split_format (complete)] inductive_cases WTrt'_elim_cases[elim!]: "P,E,h \ V :=e :' T" text\... and some easy consequences:\ lemma [iff]: "P,E,h \ e\<^sub>1;;e\<^sub>2 :' T\<^sub>2 = (\T\<^sub>1. P,E,h \ e\<^sub>1:' T\<^sub>1 \ P,E,h \ e\<^sub>2:' T\<^sub>2)" apply(rule iffI) apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) done lemma [iff]: "P,E,h \ Val v :' T = (P \ typeof\<^bsub>h\<^esub> v = Some T)" apply(rule iffI) apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) done lemma [iff]: "P,E,h \ Var V :' T = (E V = Some T)" apply(rule iffI) apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) done lemma wt_wt': "P,E,h \ e : T \ P,E,h \ e :' T" and wts_wts': "P,E,h \ es [:] Ts \ P,E,h \ es [:'] Ts" proof (induct rule:WTrt_inducts) case (WTrtBlock E V T h e T') thus ?case apply(case_tac "assigned V e") apply(auto intro:WTrt'_WTrts'.intros simp add:fun_upd_same assigned_def simp del:fun_upd_apply) done qed(auto intro:WTrt'_WTrts'.intros simp del:fun_upd_apply) lemma wt'_wt: "P,E,h \ e :' T \ P,E,h \ e : T" and wts'_wts: "P,E,h \ es [:'] Ts \ P,E,h \ es [:] Ts" apply (induct rule:WTrt'_inducts) apply (fastforce intro: WTrt_WTrts.intros)+ done corollary wt'_iff_wt: "(P,E,h \ e :' T) = (P,E,h \ e : T)" by(blast intro:wt_wt' wt'_wt) corollary wts'_iff_wts: "(P,E,h \ es [:'] Ts) = (P,E,h \ es [:] Ts)" by(blast intro:wts_wts' wts'_wts) lemmas WTrt_inducts2 = WTrt'_inducts [unfolded wt'_iff_wt wts'_iff_wts, case_names WTrtNew WTrtDynCast WTrtStaticCast WTrtVal WTrtVar WTrtBinOp WTrtLAss WTrtFAcc WTrtFAccNT WTrtFAss WTrtFAssNT WTrtCall WTrtStaticCall WTrtCallNT WTrtInitBlock WTrtBlock WTrtSeq WTrtCond WTrtWhile WTrtThrow WTrtNil WTrtCons, consumes 1] subsection\The theorem \progress\\ lemma mdc_leq_dyn_type: "P,E,h \ e : T \ \C a Cs D S. T = Class C \ e = ref(a,Cs) \ h a = Some(D,S) \ P \ D \\<^sup>* C" and "P,E,h \ es [:] Ts \ \T Ts' e es' C a Cs D S. Ts = T#Ts' \ es = e#es' \ T = Class C \ e = ref(a,Cs) \ h a = Some(D,S) \ P \ D \\<^sup>* C" proof (induct rule:WTrt_inducts2) case (WTrtVal h v T E) have type:"P \ typeof\<^bsub>h\<^esub> v = Some T" by fact { fix C a Cs D S assume "T = Class C" and "Val v = ref(a,Cs)" and "h a = Some(D,S)" with type have "Subobjs P D Cs" and "C = last Cs" by (auto split:if_split_asm) hence "P \ D \\<^sup>* C" by simp (rule Subobjs_subclass) } thus ?case by blast qed auto lemma appendPath_append_last: assumes notempty:"Ds \ []" shows"(Cs @\<^sub>p Ds) @\<^sub>p [last Ds] = (Cs @\<^sub>p Ds)" proof - have "last Cs = hd Ds \ last (Cs @ tl Ds) = last Ds" proof(cases "tl Ds = []") case True assume last:"last Cs = hd Ds" with True notempty have "Ds = [last Cs]" by (fastforce dest:hd_Cons_tl) hence "last Ds = last Cs" by simp with True show ?thesis by simp next case False assume last:"last Cs = hd Ds" from notempty False have "last (tl Ds) = last Ds" by -(drule hd_Cons_tl,drule_tac x="hd Ds" in last_ConsR,simp) with False show ?thesis by simp qed thus ?thesis by(simp add:appendPath_def) qed theorem assumes wf: "wwf_prog P" shows progress: "P,E,h \ e : T \ (\l. \ P \ h \; P \ E \; \ e \dom l\; \ final e \ \ \e' s'. P,E \ \e,(h,l)\ \ \e',s'\)" and "P,E,h \ es [:] Ts \ (\l. \ P \ h \; P \ E \; \s es \dom l\; \ finals es \ \ \es' s'. P,E \ \es,(h,l)\ [\] \es',s'\)" proof (induct rule:WTrt_inducts2) case (WTrtNew C E h) show ?case proof cases assume "\a. h a = None" with WTrtNew show ?thesis by (fastforce del:exE intro!:RedNew simp:new_Addr_def) next assume "\(\a. h a = None)" with WTrtNew show ?thesis by(fastforce intro:RedNewFail simp add:new_Addr_def) qed next case (WTrtDynCast C E h e T) have wte: "P,E,h \ e : T" and refT: "is_refT T" and "class": "is_class P C" and IH: "\l. \P \ h \; P \ E \; \ e \dom l\; \ final e\ \ \e' s'. P,E \ \e,(h,l)\ \ \e',s'\" and D: "\ (Cast C e) \dom l\" and hconf: "P \ h \" and envconf:"P \ E \" by fact+ from D have De: "\ e \dom l\" by auto show ?case proof cases assume "final e" with wte refT show ?thesis proof (rule finalRefE) assume "e = null" thus ?case by(fastforce intro:RedDynCastNull) next fix r assume "e = ref r" then obtain a Cs where ref:"e = ref(a,Cs)" by (cases r) auto with wte obtain D S where h:"h a = Some(D,S)" by auto show ?thesis proof (cases "P \ Path D to C unique") case True then obtain Cs' where path:"P \ Path D to C via Cs'" by (fastforce simp:path_via_def path_unique_def) then obtain Ds where "Ds = appendPath Cs Cs'" by simp with h path True ref show ?thesis by (fastforce intro:RedDynCast) next case False hence path_not_unique:"\ P \ Path D to C unique" . show ?thesis proof(cases "P \ Path last Cs to C unique") case True then obtain Cs' where "P \ Path last Cs to C via Cs'" by(auto simp:path_via_def path_unique_def) with True ref show ?thesis by(fastforce intro:RedStaticUpDynCast) next case False hence path_not_unique':"\ P \ Path last Cs to C unique" . thus ?thesis proof(cases "C \ set Cs") case False then obtain Ds Ds' where "Cs = Ds@[C]@Ds'" by (auto simp:in_set_conv_decomp) with ref show ?thesis by(fastforce intro:RedStaticDownDynCast) next case True with path_not_unique path_not_unique' h ref show ?thesis by (fastforce intro:RedDynCastFail) qed qed qed next fix r assume "e = Throw r" thus ?thesis by(blast intro!:red_reds.DynCastThrow) qed next assume nf: "\ final e" from IH[OF hconf envconf De nf] show ?thesis by (blast intro:DynCastRed) qed next case (WTrtStaticCast C E h e T) have wte: "P,E,h \ e : T" and refT: "is_refT T" and "class": "is_class P C" and IH: "\l. \P \ h \; P \ E \; \ e \dom l\; \ final e\ \ \e' s'. P,E \ \e,(h,l)\ \ \e',s'\" and D: "\ (\C\e) \dom l\" and hconf: "P \ h \" and envconf:"P \ E \" by fact+ from D have De: "\ e \dom l\" by auto show ?case proof cases assume "final e" with wte refT show ?thesis proof (rule finalRefE) assume "e = null" with "class" show ?case by(fastforce intro:RedStaticCastNull) next fix r assume "e = ref r" then obtain a Cs where ref:"e = ref(a,Cs)" by (cases r) auto with wte wf have "class":"is_class P (last Cs)" by (auto intro:Subobj_last_isClass split:if_split_asm) show ?thesis proof(cases "P \ (last Cs) \\<^sup>* C") case True with "class" wf obtain Cs' where "P \ Path last Cs to C via Cs'" by(fastforce dest:leq_implies_path) with True ref show ?thesis by(fastforce intro:RedStaticUpCast) next case False have notleq:"\ P \ last Cs \\<^sup>* C" by fact thus ?thesis proof(cases "C \ set Cs") case False then obtain Ds Ds' where "Cs = Ds@[C]@Ds'" by (auto simp:in_set_conv_decomp) with ref show ?thesis by(fastforce intro:RedStaticDownCast) next case True with ref notleq show ?thesis by (fastforce intro:RedStaticCastFail) qed qed next fix r assume "e = Throw r" thus ?thesis by(blast intro!:red_reds.StaticCastThrow) qed next assume nf: "\ final e" from IH[OF hconf envconf De nf] show ?thesis by (blast intro:StaticCastRed) qed next case WTrtVal thus ?case by(simp add:final_def) next case WTrtVar thus ?case by(fastforce intro:RedVar simp:hyper_isin_def) next case (WTrtBinOp E h e1 T1 e2 T2 bop T') have bop:"case bop of Eq \ T' = Boolean | Add \ T1 = Integer \ T2 = Integer \ T' = Integer" and wte1:"P,E,h \ e1 : T1" and wte2:"P,E,h \ e2 : T2" by fact+ show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v1 assume e1 [simp]:"e1 = Val v1" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v2 assume e2 [simp]:"e2 = Val v2" show ?thesis proof (cases bop) assume "bop = Eq" thus ?thesis using WTrtBinOp by(fastforce intro:RedBinOp) next assume Add:"bop = Add" with e1 e2 wte1 wte2 bop obtain i1 i2 where "v1 = Intg i1" and "v2 = Intg i2" by (auto dest!:typeof_Integer) with Add obtain v where "binop(bop,v1,v2) = Some v" by simp with e1 e2 show ?thesis by (fastforce intro:RedBinOp) qed next fix a assume "e2 = Throw a" thus ?thesis by(auto intro:red_reds.BinOpThrow2) qed next assume "\ final e2" with WTrtBinOp show ?thesis by simp (fast intro!:BinOpRed2) qed next fix r assume "e1 = Throw r" thus ?thesis by simp (fast intro:red_reds.BinOpThrow1) qed next assume "\ final e1" with WTrtBinOp show ?thesis by simp (fast intro:BinOpRed1) qed next case (WTrtLAss E h V T e T') have wte:"P,E,h \ e : T'" and wtvar:"P,E,h \ Var V : T" and sub:"P \ T' \ T" and envconf:"P \ E \" by fact+ from envconf wtvar have type:"is_type P T" by(auto simp:envconf_def) show ?case proof cases assume fin:"final e" from fin show ?case proof (rule finalE) fix v assume e:"e = Val v" from sub type wf show ?case proof(rule subE) assume eq:"T' = T" and "\C. T' \ Class C" hence "P \ T casts v to v" by simp(rule casts_prim) with wte wtvar eq e show ?thesis by(auto intro!:RedLAss) next fix C D assume T':"T' = Class C" and T:"T = Class D" and path_unique:"P \ Path C to D unique" from wte e T' obtain a Cs where ref:"e = ref(a,Cs)" and last:"last Cs = C" by (auto dest!:typeof_Class_Subo) from path_unique obtain Cs' where path_via:"P \ Path C to D via Cs'" by(auto simp:path_unique_def path_via_def) with last have "P \ Class D casts Ref(a,Cs) to Ref(a,Cs@\<^sub>pCs')" by (fastforce intro:casts_ref simp:path_via_def) with wte wtvar T ref show ?thesis by(auto intro!:RedLAss) next fix C assume T':"T' = NT" and T:"T = Class C" with wte e have null:"e = null" by auto have "P \ Class C casts Null to Null" by -(rule casts_null) with wte wtvar T null show ?thesis by(auto intro!:RedLAss) qed next fix r assume "e = Throw r" thus ?thesis by(fastforce intro:red_reds.LAssThrow) qed next assume "\ final e" with WTrtLAss show ?thesis by simp (fast intro:LAssRed) qed next case (WTrtFAcc E h e C Cs F T) have wte: "P,E,h \ e : Class C" and field: "P \ C has least F:T via Cs" and notemptyCs:"Cs \ []" and hconf: "P \ h \" by fact+ show ?case proof cases assume "final e" with wte show ?thesis proof (rule final_refE) fix r assume e: "e = ref r" then obtain a Cs' where ref:"e = ref(a,Cs')" by (cases r) auto with wte obtain D S where h:"h a = Some(D,S)" and suboD:"Subobjs P D Cs'" and last:"last Cs' = C" by (fastforce split:if_split_asm) from field obtain Bs fs ms where "class": "class P (last Cs) = Some(Bs,fs,ms)" and fs:"map_of fs F = Some T" by (fastforce simp:LeastFieldDecl_def FieldDecls_def) obtain Ds where Ds:"Ds = Cs'@\<^sub>pCs" by simp with notemptyCs "class" have class':"class P (last Ds) = Some(Bs,fs,ms)" by (drule_tac Cs'="Cs'" in appendPath_last) simp from field suboD last Ds wf have subo:"Subobjs P D Ds" by(fastforce intro:Subobjs_appendPath simp:LeastFieldDecl_def FieldDecls_def) with hconf h have "P,h \ (D,S) \" by (auto simp:hconf_def) with class' subo obtain fs' where S:"(Ds,fs') \ S" and "P,h \ fs' (:\) map_of fs" apply (auto simp:oconf_def) apply (erule_tac x="Ds" in allE) apply auto apply (erule_tac x="Ds" in allE) apply (erule_tac x="fs'" in allE) apply auto done with fs obtain v where "fs' F = Some v" by (fastforce simp:fconf_def) with h last Ds S have "P,E \ \(ref (a,Cs'))\F{Cs}, (h,l)\ \ \Val v,(h,l)\" by (fastforce intro:RedFAcc) with ref show ?thesis by blast next fix r assume "e = Throw r" thus ?thesis by(fastforce intro:red_reds.FAccThrow) qed next assume "\ final e" with WTrtFAcc show ?thesis by(fastforce intro!:FAccRed) qed next case (WTrtFAccNT E h e F Cs T) show ?case proof cases assume "final e" \ \@{term e} is @{term null} or @{term throw}\ with WTrtFAccNT show ?thesis by(fastforce simp:final_def intro: RedFAccNull red_reds.FAccThrow dest!:typeof_NT) next assume "\ final e" \ \@{term e} reduces by IH\ with WTrtFAccNT show ?thesis by simp (fast intro:FAccRed) qed next case (WTrtFAss E h e\<^sub>1 C Cs F T e\<^sub>2 T') have wte1:"P,E,h \ e\<^sub>1 : Class C" and wte2:"P,E,h \ e\<^sub>2 : T'" and field:"P \ C has least F:T via Cs" and notemptyCs:"Cs \ []" and sub:"P \ T' \ T" and hconf:"P \ h \" by fact+ from field wf have type:"is_type P T" by(rule least_field_is_type) show ?case proof cases assume "final e\<^sub>1" with wte1 show ?thesis proof (rule final_refE) fix r assume e1: "e\<^sub>1 = ref r" show ?thesis proof cases assume "final e\<^sub>2" thus ?thesis proof (rule finalE) fix v assume e2:"e\<^sub>2 = Val v" from e1 obtain a Cs' where ref:"e\<^sub>1 = ref(a,Cs')" by (cases r) auto with wte1 obtain D S where h:"h a = Some(D,S)" and suboD:"Subobjs P D Cs'" and last:"last Cs' = C" by (fastforce split:if_split_asm) from field obtain Bs fs ms where "class": "class P (last Cs) = Some(Bs,fs,ms)" and fs:"map_of fs F = Some T" by (fastforce simp:LeastFieldDecl_def FieldDecls_def) obtain Ds where Ds:"Ds = Cs'@\<^sub>pCs" by simp with notemptyCs "class" have class':"class P (last Ds) = Some(Bs,fs,ms)" by (drule_tac Cs'="Cs'" in appendPath_last) simp from field suboD last Ds wf have subo:"Subobjs P D Ds" by(fastforce intro:Subobjs_appendPath simp:LeastFieldDecl_def FieldDecls_def) with hconf h have "P,h \ (D,S) \" by (auto simp:hconf_def) with class' subo obtain fs' where S:"(Ds,fs') \ S" by (auto simp:oconf_def) from sub type wf show ?thesis proof(rule subE) assume eq:"T' = T" and "\C. T' \ Class C" hence "P \ T casts v to v" by simp(rule casts_prim) with h last field Ds notemptyCs S eq have "P,E \ \(ref (a,Cs'))\F{Cs}:=(Val v), (h,l)\ \ \Val v, (h(a \ (D,insert (Ds,fs'(F\v)) (S - {(Ds,fs')}))),l)\" by (fastforce intro:RedFAss) with ref e2 show ?thesis by blast next fix C' D' assume T':"T' = Class C'" and T:"T = Class D'" and path_unique:"P \ Path C' to D' unique" from wte2 e2 T' obtain a' Cs'' where ref2:"e\<^sub>2 = ref(a',Cs'')" and last':"last Cs'' = C'" by (auto dest!:typeof_Class_Subo) from path_unique obtain Ds' where "P \ Path C' to D' via Ds'" by(auto simp:path_via_def path_unique_def) with last' have casts:"P \ Class D' casts Ref(a',Cs'') to Ref(a',Cs''@\<^sub>pDs')" by (fastforce intro:casts_ref simp:path_via_def) obtain v' where "v' = Ref(a',Cs''@\<^sub>pDs')" by simp with h last field Ds notemptyCs S ref e2 ref2 T casts have "P,E \ \(ref (a,Cs'))\F{Cs}:=(Val v), (h,l)\ \ \Val v',(h(a \ (D,insert (Ds,fs'(F\v'))(S-{(Ds,fs')}))),l)\" by (fastforce intro:RedFAss) with ref e2 show ?thesis by blast next fix C' assume T':"T' = NT" and T:"T = Class C'" from e2 wte2 T' have null:"e\<^sub>2 = null" by auto have casts:"P \ Class C' casts Null to Null" by -(rule casts_null) obtain v' where "v' = Null" by simp with h last field Ds notemptyCs S ref e2 null T casts have "P,E \ \(ref (a,Cs'))\F{Cs}:=(Val v), (h,l)\ \ \Val v', (h(a \ (D,insert (Ds,fs'(F\v')) (S - {(Ds,fs')}))),l)\" by (fastforce intro:RedFAss) with ref e2 show ?thesis by blast qed next fix r assume "e\<^sub>2 = Throw r" thus ?thesis using e1 by(fastforce intro:red_reds.FAssThrow2) qed next assume "\ final e\<^sub>2" with WTrtFAss e1 show ?thesis by simp (fast intro!:FAssRed2) qed next fix r assume "e\<^sub>1 = Throw r" thus ?thesis by(fastforce intro:red_reds.FAssThrow1) qed next assume "\ final e\<^sub>1" with WTrtFAss show ?thesis by simp (blast intro!:FAssRed1) qed next case (WTrtFAssNT E h e\<^sub>1 e\<^sub>2 T' T F Cs) show ?case proof cases assume e1: "final e\<^sub>1" \ \@{term e\<^sub>1} is @{term null} or @{term throw}\ show ?thesis proof cases assume "final e\<^sub>2" \ \@{term e\<^sub>2} is @{term Val} or @{term throw}\ with WTrtFAssNT e1 show ?thesis by(fastforce simp:final_def intro:RedFAssNull red_reds.FAssThrow1 red_reds.FAssThrow2 dest!:typeof_NT) next assume "\ final e\<^sub>2" \ \@{term e\<^sub>2} reduces by IH\ with WTrtFAssNT e1 show ?thesis by (fastforce simp:final_def intro!:red_reds.FAssRed2 red_reds.FAssThrow1) qed next assume "\ final e\<^sub>1" \ \@{term e\<^sub>1} reduces by IH\ with WTrtFAssNT show ?thesis by (fastforce intro:FAssRed1) qed next case (WTrtCall E h e C M Ts T pns body Cs es Ts') have wte: "P,E,h \ e : Class C" and "method":"P \ C has least M = (Ts, T, pns, body) via Cs" and wtes: "P,E,h \ es [:] Ts'"and sub: "P \ Ts' [\] Ts" and IHes: "\l. \P \ h \; P \ E \; \s es \dom l\; \ finals es\ \ \es' s'. P,E \ \es,(h,l)\ [\] \es',s'\" and hconf: "P \ h \" and envconf:"P \ E \" and D: "\ (e\M(es)) \dom l\" by fact+ show ?case proof cases assume final:"final e" with wte show ?thesis proof (rule final_refE) fix r assume ref: "e = ref r" show ?thesis proof cases assume es: "\vs. es = map Val vs" from ref obtain a Cs' where ref:"e = ref(a,Cs')" by (cases r) auto with wte obtain D S where h:"h a = Some(D,S)" and suboD:"Subobjs P D Cs'" and last:"last Cs' = C" by (fastforce split:if_split_asm) from wte ref h have subcls:"P \ D \\<^sup>* C" by -(drule mdc_leq_dyn_type,auto) from "method" have has:"P \ C has M = (Ts,T,pns,body) via Cs" by(rule has_least_method_has_method) from es obtain vs where vs:"es = map Val vs" by auto obtain Cs'' Ts'' T' pns' body' where ass:"P \ (D,Cs'@\<^sub>pCs) selects M = (Ts'',T',pns',body') via Cs'' \ length Ts'' = length pns' \ length vs = length pns' \ P \ T' \ T" proof (cases "\Ts'' T' pns' body' Ds. P \ D has least M = (Ts'',T',pns',body') via Ds") case True then obtain Ts'' T' pns' body' Cs'' where least:"P \ D has least M = (Ts'',T',pns',body') via Cs''" by auto hence select:"P \ (D,Cs'@\<^sub>pCs) selects M = (Ts'',T',pns',body') via Cs''" by(rule dyn_unique) from subcls least wf has have "Ts = Ts''" and leq:"P \ T' \ T" by -(drule leq_method_subtypes,simp_all,blast)+ hence "length Ts = length Ts''" by (simp add:list_all2_iff) with sub have "length Ts' = length Ts''" by (simp add:list_all2_iff) with WTrts_same_length[OF wtes] vs have length:"length vs = length Ts''" by simp from has_least_wf_mdecl[OF wf least] have lengthParams:"length Ts'' = length pns'" by (simp add:wf_mdecl_def) with length have "length vs = length pns'" by simp with select lengthParams leq show ?thesis using that by blast next case False hence non_dyn:"\Ts'' T' pns' body' Ds. \ P \ D has least M = (Ts'',T',pns',body') via Ds" by auto from suboD last have path:"P \ Path D to C via Cs'" by(simp add:path_via_def) from "method" have notempty:"Cs \ []" by(fastforce intro!:Subobjs_nonempty simp:LeastMethodDef_def MethodDefs_def) from suboD have "class": "is_class P D" by(rule Subobjs_isClass) from suboD last have path:"P \ Path D to C via Cs'" by(simp add:path_via_def) with "method" wf have "P \ D has M = (Ts,T,pns,body) via Cs'@\<^sub>pCs" by(auto intro:has_path_has has_least_method_has_method) with "class" wf obtain Cs'' Ts'' T' pns' body' where overrider: "P \ (D,Cs'@\<^sub>pCs) has overrider M = (Ts'',T',pns',body') via Cs''" by(auto dest!:class_wf simp:is_class_def wf_cdecl_def,blast) with non_dyn have select:"P \ (D,Cs'@\<^sub>pCs) selects M = (Ts'',T',pns',body') via Cs''" by-(rule dyn_ambiguous,simp_all) from notempty have eq:"(Cs' @\<^sub>p Cs) @\<^sub>p [last Cs] = (Cs' @\<^sub>p Cs)" by(rule appendPath_append_last) from "method" wf have "P \ last Cs has least M = (Ts,T,pns,body) via [last Cs]" by(auto dest:Subobj_last_isClass intro:Subobjs_Base subobjs_rel simp:LeastMethodDef_def MethodDefs_def) with notempty have "P \ last(Cs'@\<^sub>pCs) has least M = (Ts,T,pns,body) via [last Cs]" by -(drule_tac Cs'="Cs'" in appendPath_last,simp) with overrider wf eq have "(Cs'',(Ts'',T',pns',body')) \ MinimalMethodDefs P D M" and "P,D \ Cs'' \ Cs'@\<^sub>pCs" by(auto simp:FinalOverriderMethodDef_def OverriderMethodDefs_def) (drule wf_sees_method_fun,auto) with subcls wf notempty has path have "Ts = Ts''" and leq:"P \ T' \ T" by -(drule leq_methods_subtypes,simp_all,blast)+ hence "length Ts = length Ts''" by (simp add:list_all2_iff) with sub have "length Ts' = length Ts''" by (simp add:list_all2_iff) with WTrts_same_length[OF wtes] vs have length:"length vs = length Ts''" by simp from select_method_wf_mdecl[OF wf select] have lengthParams:"length Ts'' = length pns'" by (simp add:wf_mdecl_def) with length have "length vs = length pns'" by simp with select lengthParams leq show ?thesis using that by blast qed obtain new_body where "case T of Class D \ new_body = \D\blocks(this#pns',Class(last Cs'')#Ts'',Ref(a,Cs'')#vs,body') | _ \ new_body = blocks(this#pns',Class(last Cs'')#Ts'',Ref(a,Cs'')#vs,body')" by(cases T) auto with h "method" last ass ref vs show ?thesis by (auto intro!:exI RedCall) next assume "\(\vs. es = map Val vs)" hence not_all_Val: "\(\e \ set es. \v. e = Val v)" by(simp add:ex_map_conv) let ?ves = "takeWhile (\e. \v. e = Val v) es" let ?rest = "dropWhile (\e. \v. e = Val v) es" let ?ex = "hd ?rest" let ?rst = "tl ?rest" from not_all_Val have nonempty: "?rest \ []" by auto hence es: "es = ?ves @ ?ex # ?rst" by simp have "\e \ set ?ves. \v. e = Val v" by(fastforce dest:set_takeWhileD) then obtain vs where ves: "?ves = map Val vs" using ex_map_conv by blast show ?thesis proof cases assume "final ?ex" moreover from nonempty have "\(\v. ?ex = Val v)" by(auto simp:neq_Nil_conv simp del:dropWhile_eq_Nil_conv) (simp add:dropWhile_eq_Cons_conv) ultimately obtain r' where ex_Throw: "?ex = Throw r'" by(fast elim!:finalE) show ?thesis using ref es ex_Throw ves by(fastforce intro:red_reds.CallThrowParams) next assume not_fin: "\ final ?ex" have "finals es = finals(?ves @ ?ex # ?rst)" using es by(rule arg_cong) also have "\ = finals(?ex # ?rst)" using ves by simp finally have "finals es = finals(?ex # ?rst)" . hence "\ finals es" using not_finals_ConsI[OF not_fin] by blast thus ?thesis using ref D IHes[OF hconf envconf] by(fastforce intro!:CallParams) qed qed next fix r assume "e = Throw r" with WTrtCall.prems show ?thesis by(fast intro!:red_reds.CallThrowObj) qed next assume "\ final e" with WTrtCall show ?thesis by simp (blast intro!:CallObj) qed next case (WTrtStaticCall E h e C' C M Ts T pns body Cs es Ts') have wte: "P,E,h \ e : Class C'" and path_unique:"P \ Path C' to C unique" and "method":"P \ C has least M = (Ts, T, pns, body) via Cs" and wtes: "P,E,h \ es [:] Ts'"and sub: "P \ Ts' [\] Ts" and IHes: "\l. \P \ h \; envconf P E; \s es \dom l\; \ finals es\ \ \es' s'. P,E \ \es,(h,l)\ [\] \es',s'\" and hconf: "P \ h \" and envconf:"envconf P E" and D: "\ (e\(C::)M(es)) \dom l\" by fact+ show ?case proof cases assume final:"final e" with wte show ?thesis proof (rule final_refE) fix r assume ref: "e = ref r" show ?thesis proof cases assume es: "\vs. es = map Val vs" from ref obtain a Cs' where ref:"e = ref(a,Cs')" by (cases r) auto with wte have last:"last Cs' = C'" by (fastforce split:if_split_asm) with path_unique obtain Cs'' where path_via:"P \ Path (last Cs') to C via Cs''" by (auto simp add:path_via_def path_unique_def) obtain Ds where Ds:"Ds = (Cs'@\<^sub>pCs'')@\<^sub>pCs" by simp from es obtain vs where vs:"es = map Val vs" by auto from sub have "length Ts' = length Ts" by (simp add:list_all2_iff) with WTrts_same_length[OF wtes] vs have length:"length vs = length Ts" by simp from has_least_wf_mdecl[OF wf "method"] have lengthParams:"length Ts = length pns" by (simp add:wf_mdecl_def) with "method" last path_unique path_via Ds length ref vs show ?thesis by (auto intro!:exI RedStaticCall) next assume "\(\vs. es = map Val vs)" hence not_all_Val: "\(\e \ set es. \v. e = Val v)" by(simp add:ex_map_conv) let ?ves = "takeWhile (\e. \v. e = Val v) es" let ?rest = "dropWhile (\e. \v. e = Val v) es" let ?ex = "hd ?rest" let ?rst = "tl ?rest" from not_all_Val have nonempty: "?rest \ []" by auto hence es: "es = ?ves @ ?ex # ?rst" by simp have "\e \ set ?ves. \v. e = Val v" by(fastforce dest:set_takeWhileD) then obtain vs where ves: "?ves = map Val vs" using ex_map_conv by blast show ?thesis proof cases assume "final ?ex" moreover from nonempty have "\(\v. ?ex = Val v)" by(auto simp:neq_Nil_conv simp del:dropWhile_eq_Nil_conv) (simp add:dropWhile_eq_Cons_conv) ultimately obtain r' where ex_Throw: "?ex = Throw r'" by(fast elim!:finalE) show ?thesis using ref es ex_Throw ves by(fastforce intro:red_reds.CallThrowParams) next assume not_fin: "\ final ?ex" have "finals es = finals(?ves @ ?ex # ?rst)" using es by(rule arg_cong) also have "\ = finals(?ex # ?rst)" using ves by simp finally have "finals es = finals(?ex # ?rst)" . hence "\ finals es" using not_finals_ConsI[OF not_fin] by blast thus ?thesis using ref D IHes[OF hconf envconf] by(fastforce intro!:CallParams) qed qed next fix r assume "e = Throw r" with WTrtStaticCall.prems show ?thesis by(fast intro!:red_reds.CallThrowObj) qed next assume "\ final e" with WTrtStaticCall show ?thesis by simp (blast intro!:CallObj) qed next case (WTrtCallNT E h e es Ts Copt M T) show ?case proof cases assume "final e" moreover { fix v assume e: "e = Val v" hence "e = null" using WTrtCallNT by simp have ?case proof cases assume "finals es" moreover { fix vs assume "es = map Val vs" with WTrtCallNT e have ?thesis by(fastforce intro: RedCallNull dest!:typeof_NT) } moreover { fix vs a es' assume "es = map Val vs @ Throw a # es'" with WTrtCallNT e have ?thesis by(fastforce intro: CallThrowParams) } ultimately show ?thesis by(fastforce simp:finals_def) next assume "\ finals es" \ \@{term es} reduces by IH\ with WTrtCallNT e show ?thesis by(fastforce intro: CallParams) qed } moreover { fix r assume "e = Throw r" with WTrtCallNT have ?case by(fastforce intro: CallThrowObj) } ultimately show ?thesis by(fastforce simp:final_def) next assume "\ final e" \ \@{term e} reduces by IH\ with WTrtCallNT show ?thesis by (fastforce intro:CallObj) qed next case (WTrtInitBlock h v T' E V T e\<^sub>2 T\<^sub>2) have IH2: "\l. \P \ h \; P \ E(V \ T) \; \ e\<^sub>2 \dom l\; \ final e\<^sub>2\ \ \e' s'. P,E(V \ T) \ \e\<^sub>2,(h,l)\ \ \e',s'\" and typeof:"P \ typeof\<^bsub>h\<^esub> v = Some T'" and type:"is_type P T" and sub:"P \ T' \ T" and hconf: "P \ h \" and envconf:"P \ E \" and D: "\ {V:T := Val v; e\<^sub>2} \dom l\" by fact+ from wf typeof type sub obtain v' where casts:"P \ T casts v to v'" by(auto dest:sub_casts) show ?case proof cases assume fin:"final e\<^sub>2" with casts show ?thesis by(fastforce elim:finalE intro:RedInitBlock red_reds.InitBlockThrow) next assume not_fin2: "\ final e\<^sub>2" from D have D2: "\ e\<^sub>2 \dom(l(V\v'))\" by (auto simp:hyperset_defs) from envconf type have "P \ E(V \ T) \" by(auto simp:envconf_def) from IH2[OF hconf this D2 not_fin2] obtain h' l' e' where red2: "P,E(V \ T) \ \e\<^sub>2,(h, l(V\v'))\ \ \e',(h', l')\" by auto from red_lcl_incr[OF red2] have "V \ dom l'" by auto with red2 casts show ?thesis by(fastforce intro:InitBlockRed) qed next case (WTrtBlock E V T h e T') have IH: "\l. \P \ h \; P \ E(V \ T) \; \ e \dom l\; \ final e\ \ \e' s'. P,E(V \ T) \ \e,(h,l)\ \ \e',s'\" and unass: "\ assigned V e" and type:"is_type P T" and hconf: "P \ h \" and envconf:"P \ E \" and D: "\ {V:T; e} \dom l\" by fact+ show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume "e = Val v" with type show ?thesis by(fast intro:RedBlock) next fix r assume "e = Throw r" with type show ?thesis by(fast intro:red_reds.BlockThrow) qed next assume not_fin: "\ final e" from D have De: "\ e \dom(l(V:=None))\" by(simp add:hyperset_defs) from envconf type have "P \ E(V \ T) \" by(auto simp:envconf_def) from IH[OF hconf this De not_fin] obtain h' l' e' where red: "P,E(V \ T) \ \e,(h,l(V:=None))\ \ \e',(h',l')\" by auto show ?thesis proof (cases "l' V") assume "l' V = None" with red unass show ?thesis by(blast intro: BlockRedNone) next fix v assume "l' V = Some v" with red unass type show ?thesis by(blast intro: BlockRedSome) qed qed next case (WTrtSeq E h e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2) show ?case proof cases assume "final e\<^sub>1" thus ?thesis by(fast elim:finalE intro:intro:RedSeq red_reds.SeqThrow) next assume "\ final e\<^sub>1" with WTrtSeq show ?thesis by simp (blast intro:SeqRed) qed next case (WTrtCond E h e e\<^sub>1 T e\<^sub>2) have wt: "P,E,h \ e : Boolean" by fact show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume val: "e = Val v" then obtain b where v: "v = Bool b" using wt by (fastforce dest:typeof_Boolean) show ?thesis proof (cases b) case True with val v show ?thesis by(auto intro:RedCondT) next case False with val v show ?thesis by(auto intro:RedCondF) qed next fix r assume "e = Throw r" thus ?thesis by(fast intro:red_reds.CondThrow) qed next assume "\ final e" with WTrtCond show ?thesis by simp (fast intro:CondRed) qed next case WTrtWhile show ?case by(fast intro:RedWhile) next case (WTrtThrow E h e T' T) show ?case proof cases assume "final e" \ \Then @{term e} must be @{term throw} or @{term null}\ with WTrtThrow show ?thesis by(fastforce simp:final_def is_refT_def intro:red_reds.ThrowThrow red_reds.RedThrowNull dest!:typeof_NT typeof_Class_Subo) next assume "\ final e" \ \Then @{term e} must reduce\ with WTrtThrow show ?thesis by simp (blast intro:ThrowRed) qed next case WTrtNil thus ?case by simp next case (WTrtCons E h e T es Ts) have IHe: "\l. \P \ h \; P \ E \; \ e \dom l\; \ final e\ \ \e' s'. P,E \ \e,(h,l)\ \ \e',s'\" and IHes: "\l. \P \ h \; P \ E \; \s es \dom l\; \ finals es\ \ \es' s'. P,E \ \es,(h,l)\ [\] \es',s'\" and hconf: "P \ h \" and envconf:"P \ E \" and D: "\s (e#es) \dom l\" and not_fins: "\ finals(e # es)" by fact+ have De: "\ e \dom l\" and Des: "\s es (\dom l\ \ \ e)" using D by auto show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume e: "e = Val v" hence Des': "\s es \dom l\" using De Des by auto have not_fins_tl: "\ finals es" using not_fins e by simp show ?thesis using e IHes[OF hconf envconf Des' not_fins_tl] by (blast intro!:ListRed2) next fix r assume "e = Throw r" hence False using not_fins by simp thus ?thesis .. qed next assume "\ final e" from IHe[OF hconf envconf De this] show ?thesis by(fast intro!:ListRed1) qed qed end diff --git a/thys/Differential_Dynamic_Logic/Syntax.thy b/thys/Differential_Dynamic_Logic/Syntax.thy --- a/thys/Differential_Dynamic_Logic/Syntax.thy +++ b/thys/Differential_Dynamic_Logic/Syntax.thy @@ -1,358 +1,358 @@ theory Syntax imports Complex_Main "Ids" begin section \Syntax\ text \ We define the syntax of dL terms, formulas and hybrid programs. As in CADE'15, the syntax allows arbitrarily nested differentials. However, the semantics of such terms is very surprising (e.g. (x')' is zero in every state), so we define predicates dfree and dsafe to describe terms with no differentials and no nested differentials, respectively. In keeping with the CADE'15 presentation we currently make the simplifying assumption that all terms are smooth, and thus division and arbitrary exponentiation are absent from the syntax. Several other standard logical constructs are implemented as derived forms to reduce the soundness burden. The types of formulas and programs are parameterized by three finite types ('a, 'b, 'c) used as identifiers for function constants, context constants, and everything else, respectively. These type variables are distinct because some substitution operations affect one type variable while leaving the others unchanged. Because these types will be finite in practice, it is more useful to think of them as natural numbers that happen to be represented as types (due to HOL's lack of dependent types). The types of terms and ODE systems follow the same approach, but have only two type variables because they cannot contain contexts. \ datatype ('a, 'c) trm = \ \Real-valued variables given meaning by the state and modified by programs.\ Var 'c \ \N.B. This is technically more expressive than true dL since most reals\ \ \can't be written down.\ | Const real \ \A function (applied to its arguments) consists of an identifier for the function\ \ \and a function \'c \ ('a, 'c) trm\ (where \'c\ is a finite type) which specifies one\ \ \argument of the function for each element of type \'c\. To simulate a function with\ \ \less than \'c\ arguments, set the remaining arguments to a constant, such as \Const 0\\ | Function 'a "'c \ ('a, 'c) trm" ("$f") | Plus "('a, 'c) trm" "('a, 'c) trm" | Times "('a, 'c) trm" "('a, 'c) trm" \ \A (real-valued) variable standing for a differential, such as \x'\, given meaning by the state\ \ \and modified by programs.\ -| DiffVar 'c ("$'") +| DiffVar 'c ("$''") \ \The differential of an arbitrary term \(\)'\\ | Differential "('a, 'c) trm" datatype('a, 'c) ODE = \ \Variable standing for an ODE system, given meaning by the interpretation\ OVar 'c \ \Singleton ODE defining \x' = \\, where \\\ may or may not contain \x\\ \ \(but must not contain differentials)\ | OSing 'c "('a, 'c) trm" \ \The product \OProd ODE1 ODE2\ composes two ODE systems in parallel, e.g.\ \ \\OProd (x' = y) (y' = -x)\ is the system \{x' = y, y' = -x}\\ | OProd "('a, 'c) ODE" "('a, 'c) ODE" datatype ('a, 'b, 'c) hp = \ \Variables standing for programs, given meaning by the interpretation.\ Pvar 'c ("$\") \ \Assignment to a real-valued variable \x := \\\ | Assign 'c "('a, 'c) trm" (infixr ":=" 10) \ \Assignment to a differential variable\ | DiffAssign 'c "('a, 'c) trm" \ \Program \?\\ succeeds iff \\\ holds in current state.\ | Test "('a, 'b, 'c) formula" ("?") \ \An ODE program is an ODE system with some evolution domain.\ | EvolveODE "('a, 'c) ODE" "('a, 'b, 'c) formula" \ \Non-deterministic choice between two programs \a\ and \b\\ | Choice "('a, 'b, 'c) hp" "('a, 'b, 'c) hp" (infixl "\\" 10) \ \Sequential composition of two programs \a\ and \b\\ | Sequence "('a, 'b, 'c) hp" "('a, 'b, 'c) hp" (infixr ";;" 8) \ \Nondeterministic repetition of a program \a\, zero or more times.\ | Loop "('a, 'b, 'c) hp" ("_**") and ('a, 'b, 'c) formula = Geq "('a, 'c) trm" "('a, 'c) trm" | Prop 'c "'c \ ('a, 'c) trm" ("$\") | Not "('a, 'b, 'c) formula" ("!") | And "('a, 'b, 'c) formula" "('a, 'b, 'c) formula" (infixl "&&" 8) | Exists 'c "('a, 'b, 'c) formula" \ \\\\\\\ iff exists run of \\\ where \\\ is true in end state\ | Diamond "('a, 'b, 'c) hp" "('a, 'b, 'c) formula" ("(\ _ \ _)" 10) \ \Contexts \C\ are symbols standing for functions from (the semantics of) formulas to\ \ \(the semantics of) formulas, thus \C(\)\ is another formula. While not necessary\ \ \in terms of expressiveness, contexts allow for more efficient reasoning principles.\ | InContext 'b "('a, 'b, 'c) formula" \ \Derived forms\ definition Or :: "('a, 'b, 'c) formula \ ('a, 'b, 'c) formula \ ('a, 'b, 'c) formula" (infixl "||" 7) where "Or P Q = Not (And (Not P) (Not Q))" definition Implies :: "('a, 'b, 'c) formula \ ('a, 'b, 'c) formula \ ('a, 'b, 'c) formula" (infixr "\" 10) where "Implies P Q = Or Q (Not P)" definition Equiv :: "('a, 'b, 'c) formula \ ('a, 'b, 'c) formula \ ('a, 'b, 'c) formula" (infixl "\" 10) where "Equiv P Q = Or (And P Q) (And (Not P) (Not Q))" definition Forall :: "'c \ ('a, 'b, 'c) formula \ ('a, 'b, 'c) formula" where "Forall x P = Not (Exists x (Not P))" definition Equals :: "('a, 'c) trm \ ('a, 'c) trm \ ('a, 'b, 'c) formula" where "Equals \ \' = ((Geq \ \') && (Geq \' \))" definition Greater :: "('a, 'c) trm \ ('a, 'c) trm \ ('a, 'b, 'c) formula" where "Greater \ \' = ((Geq \ \') && (Not (Geq \' \)))" definition Box :: "('a, 'b, 'c) hp \ ('a, 'b, 'c) formula \ ('a, 'b, 'c) formula" ("([[_]]_)" 10) where "Box \ P = Not (Diamond \ (Not P))" definition TT ::"('a,'b,'c) formula" where "TT = Geq (Const 0) (Const 0)" definition FF ::"('a,'b,'c) formula" where "FF = Geq (Const 0) (Const 1)" type_synonym ('a,'b,'c) sequent = "('a,'b,'c) formula list * ('a,'b,'c) formula list" \ \Rule: assumptions, then conclusion\ type_synonym ('a,'b,'c) rule = "('a,'b,'c) sequent list * ('a,'b,'c) sequent" \ \silliness to enable proving disequality lemmas\ primrec sizeF::"('sf,'sc, 'sz) formula \ nat" and sizeP::"('sf,'sc, 'sz) hp \ nat" where "sizeP (Pvar a) = 1" | "sizeP (Assign x \) = 1" | "sizeP (DiffAssign x \) = 1" | "sizeP (Test \) = Suc (sizeF \)" | "sizeP (EvolveODE ODE \) = Suc (sizeF \)" | "sizeP (Choice \ \) = Suc (sizeP \ + sizeP \)" | "sizeP (Sequence \ \) = Suc (sizeP \ + sizeP \)" | "sizeP (Loop \) = Suc (sizeP \)" | "sizeF (Geq p q) = 1" | "sizeF (Prop p args) = 1" | "sizeF (Not p) = Suc (sizeF p)" | "sizeF (And p q) = sizeF p + sizeF q" | "sizeF (Exists x p) = Suc (sizeF p)" | "sizeF (Diamond p q) = Suc (sizeP p + sizeF q)" | "sizeF (InContext C \) = Suc (sizeF \)" lemma sizeF_diseq:"sizeF p \ sizeF q \ p \ q" by auto named_theorems "expr_diseq" "Structural disequality rules for expressions" lemma [expr_diseq]:"p \ And p q" by(induction p, auto) lemma [expr_diseq]:"q \ And p q" by(induction q, auto) lemma [expr_diseq]:"p \ Not p" by(induction p, auto) lemma [expr_diseq]:"p \ Or p q" by(rule sizeF_diseq, auto simp add: Or_def) lemma [expr_diseq]:"q \ Or p q" by(rule sizeF_diseq, auto simp add: Or_def) lemma [expr_diseq]:"p \ Implies p q" by(rule sizeF_diseq, auto simp add: Implies_def Or_def) lemma [expr_diseq]:"q \ Implies p q" by(rule sizeF_diseq, auto simp add: Implies_def Or_def) lemma [expr_diseq]:"p \ Equiv p q" by(rule sizeF_diseq, auto simp add: Equiv_def Or_def) lemma [expr_diseq]:"q \ Equiv p q" by(rule sizeF_diseq, auto simp add: Equiv_def Or_def) lemma [expr_diseq]:"p \ Exists x p" by(induction p, auto) lemma [expr_diseq]:"p \ Diamond a p" by(induction p, auto) lemma [expr_diseq]:"p \ InContext C p" by(induction p, auto) \ \A predicational is like a context with no argument, i.e. a variable standing for a\ \ \state-dependent formula, given meaning by the interpretation. This differs from a predicate\ \ \because predicates depend only on their arguments (which might then indirectly depend on the state).\ \ \We encode a predicational as a context applied to a formula whose truth value is constant with\ \ \respect to the state (specifically, always true)\ fun Predicational :: "'b \ ('a, 'b, 'c) formula" ("Pc") where "Predicational P = InContext P (Geq (Const 0) (Const 0))" \ \Abbreviations for common syntactic constructs in order to make axiom definitions, etc. more\ \ \readable.\ context ids begin \ \"Empty" function argument tuple, encoded as tuple where all arguments assume a constant value.\ definition empty::" 'b \ ('a, 'b) trm" where "empty \ \i.(Const 0)" \ \Function argument tuple with (effectively) one argument, where all others have a constant value.\ fun singleton :: "('a, 'sz) trm \ ('sz \ ('a, 'sz) trm)" where "singleton t i = (if i = vid1 then t else (Const 0))" lemma expand_singleton:"singleton t = (\i. (if i = vid1 then t else (Const 0)))" by auto \ \Function applied to one argument\ definition f1::"'sf \ 'sz \ ('sf,'sz) trm" where "f1 f x = Function f (singleton (Var x))" \ \Function applied to zero arguments (simulates a constant symbol given meaning by the interpretation)\ definition f0::"'sf \ ('sf,'sz) trm" where "f0 f = Function f empty" \ \Predicate applied to one argument\ definition p1::"'sz \ 'sz \ ('sf, 'sc, 'sz) formula" where "p1 p x = Prop p (singleton (Var x))" \ \Predicational\ definition P::"'sc \ ('sf, 'sc, 'sz) formula" where "P p = Predicational p" end subsection \Well-Formedness predicates\ inductive dfree :: "('a, 'c) trm \ bool" where dfree_Var: "dfree (Var i)" | dfree_Const: "dfree (Const r)" | dfree_Fun: "(\i. dfree (args i)) \ dfree (Function i args)" | dfree_Plus: "dfree \\<^sub>1 \ dfree \\<^sub>2 \ dfree (Plus \\<^sub>1 \\<^sub>2)" | dfree_Times: "dfree \\<^sub>1 \ dfree \\<^sub>2 \ dfree (Times \\<^sub>1 \\<^sub>2)" inductive dsafe :: "('a, 'c) trm \ bool" where dsafe_Var: "dsafe (Var i)" | dsafe_Const: "dsafe (Const r)" | dsafe_Fun: "(\i. dsafe (args i)) \ dsafe (Function i args)" | dsafe_Plus: "dsafe \\<^sub>1 \ dsafe \\<^sub>2 \ dsafe (Plus \\<^sub>1 \\<^sub>2)" | dsafe_Times: "dsafe \\<^sub>1 \ dsafe \\<^sub>2 \ dsafe (Times \\<^sub>1 \\<^sub>2)" | dsafe_Diff: "dfree \ \ dsafe (Differential \)" | dsafe_DiffVar: "dsafe ($' i)" \ \Explictly-written variables that are bound by the ODE. Needed to compute whether\ \ \ODE's are valid (e.g. whether they bind the same variable twice)\ fun ODE_dom::"('a, 'c) ODE \ 'c set" where "ODE_dom (OVar c) = {}" | "ODE_dom (OSing x \) = {x}" | "ODE_dom (OProd ODE1 ODE2) = ODE_dom ODE1 \ ODE_dom ODE2" inductive osafe:: "('a, 'c) ODE \ bool" where osafe_Var:"osafe (OVar c)" | osafe_Sing:"dfree \ \ osafe (OSing x \)" | osafe_Prod:"osafe ODE1 \ osafe ODE2 \ ODE_dom ODE1 \ ODE_dom ODE2 = {} \ osafe (OProd ODE1 ODE2)" \ \Programs/formulas without any differential terms. This definition not currently used but may\ \ \be useful in the future.\ inductive hpfree:: "('a, 'b, 'c) hp \ bool" and ffree:: "('a, 'b, 'c) formula \ bool" where "hpfree (Pvar x)" | "dfree e \ hpfree (Assign x e)" \ \Differential programs allowed but not differential terms\ | "dfree e \ hpfree (DiffAssign x e)" | "ffree P \ hpfree (Test P)" \ \Differential programs allowed but not differential terms\ | "osafe ODE \ ffree P \ hpfree (EvolveODE ODE P)" | "hpfree a \ hpfree b \ hpfree (Choice a b )" | "hpfree a \ hpfree b \ hpfree (Sequence a b)" | "hpfree a \ hpfree (Loop a)" | "ffree f \ ffree (InContext C f)" | "(\arg. arg \ range args \ dfree arg) \ ffree (Prop p args)" | "ffree p \ ffree (Not p)" | "ffree p \ ffree q \ ffree (And p q)" | "ffree p \ ffree (Exists x p)" | "hpfree a \ ffree p \ ffree (Diamond a p)" | "ffree (Predicational P)" | "dfree t1 \ dfree t2 \ ffree (Geq t1 t2)" inductive hpsafe:: "('a, 'b, 'c) hp \ bool" and fsafe:: "('a, 'b, 'c) formula \ bool" where hpsafe_Pvar:"hpsafe (Pvar x)" | hpsafe_Assign:"dsafe e \ hpsafe (Assign x e)" | hpsafe_DiffAssign:"dsafe e \ hpsafe (DiffAssign x e)" | hpsafe_Test:"fsafe P \ hpsafe (Test P)" | hpsafe_Evolve:"osafe ODE \ fsafe P \ hpsafe (EvolveODE ODE P)" | hpsafe_Choice:"hpsafe a \ hpsafe b \ hpsafe (Choice a b )" | hpsafe_Sequence:"hpsafe a \ hpsafe b \ hpsafe (Sequence a b)" | hpsafe_Loop:"hpsafe a \ hpsafe (Loop a)" | fsafe_Geq:"dsafe t1 \ dsafe t2 \ fsafe (Geq t1 t2)" | fsafe_Prop:"(\i. dsafe (args i)) \ fsafe (Prop p args)" | fsafe_Not:"fsafe p \ fsafe (Not p)" | fsafe_And:"fsafe p \ fsafe q \ fsafe (And p q)" | fsafe_Exists:"fsafe p \ fsafe (Exists x p)" | fsafe_Diamond:"hpsafe a \ fsafe p \ fsafe (Diamond a p)" | fsafe_InContext:"fsafe f \ fsafe (InContext C f)" \ \Auto-generated simplifier rules for safety predicates\ inductive_simps dfree_Plus_simps[simp]: "dfree (Plus a b)" and dfree_Times_simps[simp]: "dfree (Times a b)" and dfree_Var_simps[simp]: "dfree (Var x)" and dfree_DiffVar_simps[simp]: "dfree (DiffVar x)" and dfree_Differential_simps[simp]: "dfree (Differential x)" and dfree_Fun_simps[simp]: "dfree (Function i args)" and dfree_Const_simps[simp]: "dfree (Const r)" inductive_simps dsafe_Plus_simps[simp]: "dsafe (Plus a b)" and dsafe_Times_simps[simp]: "dsafe (Times a b)" and dsafe_Var_simps[simp]: "dsafe (Var x)" and dsafe_DiffVar_simps[simp]: "dsafe (DiffVar x)" and dsafe_Fun_simps[simp]: "dsafe (Function i args)" and dsafe_Diff_simps[simp]: "dsafe (Differential a)" and dsafe_Const_simps[simp]: "dsafe (Const r)" inductive_simps osafe_OVar_simps[simp]:"osafe (OVar c)" and osafe_OSing_simps[simp]:"osafe (OSing x \)" and osafe_OProd_simps[simp]:"osafe (OProd ODE1 ODE2)" inductive_simps hpsafe_Pvar_simps[simp]: "hpsafe (Pvar a)" and hpsafe_Sequence_simps[simp]: "hpsafe (a ;; b)" and hpsafe_Loop_simps[simp]: "hpsafe (a**)" and hpsafe_ODE_simps[simp]: "hpsafe (EvolveODE ODE p)" and hpsafe_Choice_simps[simp]: "hpsafe (a \\ b)" and hpsafe_Assign_simps[simp]: "hpsafe (Assign x e)" and hpsafe_DiffAssign_simps[simp]: "hpsafe (DiffAssign x e)" and hpsafe_Test_simps[simp]: "hpsafe (? p)" and fsafe_Geq_simps[simp]: "fsafe (Geq t1 t2)" and fsafe_Prop_simps[simp]: "fsafe (Prop p args)" and fsafe_Not_simps[simp]: "fsafe (Not p)" and fsafe_And_simps[simp]: "fsafe (And p q)" and fsafe_Exists_simps[simp]: "fsafe (Exists x p)" and fsafe_Diamond_simps[simp]: "fsafe (Diamond a p)" and fsafe_Context_simps[simp]: "fsafe (InContext C p)" definition Ssafe::"('sf,'sc,'sz) sequent \ bool" where "Ssafe S \((\i. i \ 0 \ i < length (fst S) \ fsafe (nth (fst S) i)) \(\i. i \ 0 \ i < length (snd S) \ fsafe (nth (snd S) i)))" definition Rsafe::"('sf,'sc,'sz) rule \ bool" where "Rsafe R \ ((\i. i \ 0 \ i < length (fst R) \ Ssafe (nth (fst R) i)) \ Ssafe (snd R))" \ \Basic reasoning principles about syntactic constructs, including inductive principles\ lemma dfree_is_dsafe: "dfree \ \ dsafe \" by (induction rule: dfree.induct) (auto intro: dsafe.intros) lemma hp_induct [case_names Var Assign DiffAssign Test Evolve Choice Compose Star]: "(\x. P ($\ x)) \ (\x1 x2. P (x1 := x2)) \ (\x1 x2. P (DiffAssign x1 x2)) \ (\x. P (? x)) \ (\x1 x2. P (EvolveODE x1 x2)) \ (\x1 x2. P x1 \ P x2 \ P (x1 \\ x2)) \ (\x1 x2. P x1 \ P x2 \ P (x1 ;; x2)) \ (\x. P x \ P x**) \ P hp" by(induction rule: hp.induct) (auto) lemma fml_induct: "(\t1 t2. P (Geq t1 t2)) \ (\p args. P (Prop p args)) \ (\p. P p \ P (Not p)) \ (\p q. P p \ P q \ P (And p q)) \ (\x p. P p \ P (Exists x p)) \ (\a p. P p \ P (Diamond a p)) \ (\C p. P p \ P (InContext C p)) \ P \" by (induction rule: formula.induct) (auto) context ids begin lemma proj_sing1:"(singleton \ vid1) = \" by (auto) lemma proj_sing2:"vid1 \ y \ (singleton \ y) = (Const 0)" by (auto) end end diff --git a/thys/Hoare_Time/QuantK_Hoare.thy b/thys/Hoare_Time/QuantK_Hoare.thy --- a/thys/Hoare_Time/QuantK_Hoare.thy +++ b/thys/Hoare_Time/QuantK_Hoare.thy @@ -1,794 +1,794 @@ section \Quantitative Hoare Logic (big-O style)\ theory QuantK_Hoare imports Big_StepT Complex_Main "HOL-Library.Extended_Nat" begin abbreviation "eq a b == (And (Not (Less a b)) (Not (Less b a)))" type_synonym lvname = string type_synonym assn = "state \ bool" (* time bound *) type_synonym qassn = "state \ enat" (* time bound *) text \The support of an assn2\ abbreviation state_subst :: "state \ aexp \ vname \ state" ("_[_'/_]" [1000,0,0] 999) where "s[a/x] == s(x := aval a s)" fun emb :: "bool \ enat" ("\") where "emb False = \" | "emb True = 0" subsection "Definition of Validity" (* this definition refines the normal Hoare Triple Validity *) definition hoare2o_valid :: "qassn \ com \ qassn \ bool" - ("\\<^sub>2\<^sub>' {(1_)}/ (_)/ {(1_)}" 50) where + ("\\<^sub>2\<^sub>'' {(1_)}/ (_)/ {(1_)}" 50) where "\\<^sub>2\<^sub>' {P} c {Q} \ (\k>0. (\s. P s < \ \ (\t p. ((c,s) \ p \ t) \ enat k * P s \ p + enat k * Q t)))" subsection "Hoare Rules" inductive - hoareQ :: "qassn \ com \ qassn \ bool" ("\\<^sub>2\<^sub>' ({(1_)}/ (_)/ {(1_)})" 50) + hoareQ :: "qassn \ com \ qassn \ bool" ("\\<^sub>2\<^sub>'' ({(1_)}/ (_)/ {(1_)})" 50) where Skip: "\\<^sub>2\<^sub>' {%s. eSuc (P s)} SKIP {P}" | Assign: "\\<^sub>2\<^sub>' {\s. eSuc (P (s[a/x]))} x::=a { P }" | If: "\ \\<^sub>2\<^sub>' {\s. P s + \( bval b s)} c\<^sub>1 { Q}; \\<^sub>2\<^sub>' {\s. P s + \(\ bval b s)} c\<^sub>2 { Q} \ \ \\<^sub>2\<^sub>' {\s. eSuc (P s)} IF b THEN c\<^sub>1 ELSE c\<^sub>2 { Q }" | Seq: "\ \\<^sub>2\<^sub>' { P\<^sub>1 } c\<^sub>1 { P\<^sub>2 }; \\<^sub>2\<^sub>' {P\<^sub>2} c\<^sub>2 { P\<^sub>3 }\ \ \\<^sub>2\<^sub>' {P\<^sub>1} c\<^sub>1;;c\<^sub>2 {P\<^sub>3}" | While: "\ \\<^sub>2\<^sub>' { %s. I s + \(bval b s) } c { %t. I t + 1 } \ \ \\<^sub>2\<^sub>' {\s. I s + 1 } WHILE b DO c {\s. I s + \(\ bval b s) }" | conseq: "\ \\<^sub>2\<^sub>' {P}c{Q} ; \s. P s \ enat k * P' s ; \s. enat k * Q' s \ Q s; k>0 \ \ \\<^sub>2\<^sub>' {P'}c{ Q'}" text \Derived Rules\ lemma const: "\ \\<^sub>2\<^sub>' {\s. enat k * P s}c{\s. enat k * Q s}; k>0 \ \ \\<^sub>2\<^sub>' {P}c{ Q}" apply(rule conseq) by auto inductive hoareQ' :: "qassn \ com \ qassn \ bool" ("\\<^sub>Z ({(1_)}/ (_)/ {(1_)})" 50) where ZSkip: "\\<^sub>Z {%s. eSuc (P s)} SKIP {P}" | ZAssign: "\\<^sub>Z {\s. eSuc (P (s[a/x]))} x::=a { P }" | ZIf: "\ \\<^sub>Z {\s. P s + \( bval b s)} c\<^sub>1 { Q}; \\<^sub>Z {\s. P s + \(\ bval b s)} c\<^sub>2 { Q} \ \ \\<^sub>Z {\s. eSuc (P s)} IF b THEN c\<^sub>1 ELSE c\<^sub>2 { Q }" | ZSeq: "\ \\<^sub>Z { P\<^sub>1 } c\<^sub>1 { P\<^sub>2 }; \\<^sub>Z {P\<^sub>2} c\<^sub>2 { P\<^sub>3 }\ \ \\<^sub>Z {P\<^sub>1} c\<^sub>1;;c\<^sub>2 {P\<^sub>3}" | ZWhile: "\ \\<^sub>Z { %s. I s + \(bval b s) } c { %t. I t + 1 } \ \ \\<^sub>Z {\s. I s + 1 } WHILE b DO c {\s. I s + \(\ bval b s) }" | Zconseq': "\ \\<^sub>Z {P}c{Q} ; \s. P s \ P' s ; \s. Q' s \ Q s \ \ \\<^sub>Z {P'}c{ Q'}" | Zconst: "\ \\<^sub>Z {\s. enat k * P s}c{\s. enat k * Q s}; k>0 \ \ \\<^sub>Z {P}c{ Q}" lemma Zconseq: "\ \\<^sub>Z {P}c{Q} ; \s. P s \ enat k * P' s ; \s. enat k * Q' s \ Q s; k>0 \ \ \\<^sub>Z {P'}c{ Q'}" apply(rule Zconst[of k P' c Q']) apply(rule Zconseq'[where P=P and Q=Q]) by auto lemma ZQ: " \\<^sub>Z { P } c { Q } \ \\<^sub>2\<^sub>' { P } c { Q }" apply(induct rule: hoareQ'.induct) apply (auto simp: hoareQ.Skip hoareQ.Assign hoareQ.If hoareQ.Seq hoareQ.While) subgoal using conseq[where k=1] using one_enat_def by auto subgoal for k P c Q using const by auto done lemma QZ: " \\<^sub>2\<^sub>' { P } c { Q } \ \\<^sub>Z { P } c { Q }" apply(induct rule: hoareQ.induct) apply (auto simp: ZSkip ZAssign ZIf ZSeq ZWhile ) using Zconseq by blast lemma QZ_iff: "\\<^sub>2\<^sub>' { P } c { Q } \ \\<^sub>Z { P } c { Q }" using ZQ QZ by metis subsection "Soundness" lemma enatSuc0[simp]: "enat (Suc 0) * x = x" using one_enat_def by auto theorem hoareQ_sound: "\\<^sub>2\<^sub>' {P}c{ Q} \ \\<^sub>2\<^sub>' {P}c{ Q}" apply(unfold hoare2o_valid_def) proof( induction rule: hoareQ.induct) case (Skip P) show ?case apply(rule exI[where x=1]) apply(auto) subgoal for s apply(rule exI[where x=s]) apply(rule exI[where x="Suc 0"]) apply safe apply fast by (metis add.left_neutral add.right_neutral eSuc_enat iadd_Suc le_iff_add zero_enat_def) done next case (Assign P a x) show ?case apply(rule exI[where x=1]) apply(auto) subgoal for s apply(rule exI[where x="s[a/x]"]) apply(rule exI[where x="Suc 0"]) apply safe apply fast by (metis add.left_neutral add.right_neutral eSuc_enat iadd_Suc le_iff_add zero_enat_def) done next case (Seq P1 C1 P2 C2 P3) from Seq(3) obtain k1 where Seq3: "\s. P1 s < \ \ (\t p. (C1, s) \ p \ t \ enat p + enat k1 * P2 t \ enat k1 * P1 s)" and 10: "k1>0" by blast from Seq(4) obtain k2 where Seq4: "\s. P2 s < \ \ (\t p. (C2, s) \ p \ t \ enat p + enat k2 * P3 t \ enat k2 * P2 s)" and 20: "k2>0" by blast let ?k = "lcm k1 k2" (* or k1*k2 *) show ?case apply(rule exI[where x="?k"]) proof (safe) from 10 20 show "lcm k1 k2>0" by (auto simp: lcm_pos_nat) fix s assume ninfP1: "P1 s < \" with Seq3 obtain t1 p1 where 1: "(C1, s) \ p1 \ t1" and q1: "enat p1 + k1 * P2 t1 \ k1 * P1 s" by blast with ninfP1 have ninfP2: "P2 t1 < \" using not_le 10 by fastforce with Seq4 obtain t2 p2 where 2: "(C2, t1) \ p2 \ t2" and q2: "enat p2 + k2 * P3 t2 \ k2 * P2 t1" by blast with ninfP2 have ninfP3: "P3 t2 < \" using not_le 20 by fastforce then obtain u2 where u2: "P3 t2 = enat u2" by auto from ninfP2 obtain u1 where u1: "P2 t1 = enat u1" by auto from ninfP1 obtain u0 where u0: "P1 s = enat u0" by auto from Big_StepT.Seq[OF 1 2] have 12: "(C1;; C2, s) \ p1 + p2 \ t2" by simp have i: "(C1;; C2, s) \ p1+p2 \ t2" using 1 and 2 by auto from 10 20 have p: "k1 div gcd k1 k2 > 0" "k2 div gcd k1 k2 > 0" by (simp_all add: Divides.div_positive) have za: "?k = (k1 div gcd k1 k2) * k2" apply(simp only: lcm_nat_def) by (simp add: dvd_div_mult) have za2: "?k = (k2 div gcd k1 k2) * k1" apply(simp only: lcm_nat_def) by (metis dvd_div_mult gcd_dvd2 mult.commute) from q1[unfolded u1 u2 u0] have z: "p1 + k1 * u1 \ k1 * u0" by auto from q2[unfolded u1 u2 u0] have y: " p2 + k2 * u2 \ k2 * u1" by auto have "p1+p2 + ?k * u2 \ p1 + (k1 div gcd k1 k2)*p2 + ?k * u2 " using p by simp also have "\ \ (k2 div gcd k1 k2)*p1 + (k1 div gcd k1 k2)*p2 + ?k * u2 " using p by simp also have "\ = (k2 div gcd k1 k2)*p1 + (k1 div gcd k1 k2)*(p2 + k2* u2)" apply(simp only: za) by algebra also have "\ \ (k2 div gcd k1 k2)*p1 + (k1 div gcd k1 k2)*(k2 * u1)" using y by (metis add_left_mono distrib_left le_iff_add) also have "\ = (k2 div gcd k1 k2)*p1 + ?k * u1" by(simp only: za) also have "\ = (k2 div gcd k1 k2)*p1 + (k2 div gcd k1 k2) *(k1* u1)" by(simp only: za2) also have "\ \ (k2 div gcd k1 k2)*(p1 + k1*u1)" by (simp add: distrib_left) also have "\ \ (k2 div gcd k1 k2)*(k1 * u0)" using z by fastforce also have "\ \ ?k * u0" by(simp only: za2) finally have "p1+p2 + ?k * u2 \ ?k * u0" . then have ii: "enat (p1+p2) + ?k * P3 t2 \ ?k * P1 s" unfolding u0 u2 by auto from i ii show "\t p. (C1;; C2, s) \ p \ t \ enat p + ?k * P3 t \ ?k * P1 s " by blast qed next case (If P b c1 Q c2) from If(3) obtain kT where If3: "\s. P s + \ (bval b s) < \ \ (\t p. (c1, s) \ p \ t \ enat p + enat kT * Q t \ enat kT * (P s + \ (bval b s))) " and T: "kT > 0" by blast from If(4) obtain kF where If4: "\s. P s + \ (\ bval b s) < \ \ (\t p. (c2, s) \ p \ t \ enat p + enat kF * Q t \ enat kF * (P s + \ (\ bval b s)))" and F: "kF > 0" by blast show ?case apply(rule exI[where x="kT*kF"]) proof (safe) from T F show "0 < kT * kF" by auto fix s assume "eSuc (P s) < \" then have i: "P s < \" using enat_ord_simps(4) by fastforce then obtain u0 where u0: "P s = enat u0" by auto show "\t p. (IF b THEN c1 ELSE c2, s) \ p \ t \ enat p + enat (kT * kF) * Q t \ enat (kT * kF) * eSuc (P s)" proof(cases "bval b s") case True with i have "P s + emb (bval b s) < \" by simp with If3 obtain p t where 1: "(c1, s) \ p \ t" and q: "enat p + enat kT * Q t \ enat kT * (P s + emb (bval b s))" by blast from Big_StepT.IfTrue[OF True 1] have 2: "(IF b THEN c1 ELSE c2, s) \ p + 1 \ t" by simp from q have "Q t < \" using i T True using less_irrefl by fastforce then obtain u1 where u1: "Q t = enat u1" by auto from q True have q': "p + kT * u1 \ kT * u0" unfolding u0 u1 by auto have "(p+1) + (kT * kF) * u1 \ kF*(p+1) + (kT * kF) * u1" using F by (simp add: mult_eq_if) also have "\ \ kF*(p+1 + kT * u1)" by (simp add: add_mult_distrib2) also have "\ \ kF*(1 + kT * u0)" using q' by auto also have "\ \ (kT * kF) * Suc u0" using T by simp finally have " (p+1) + (kT * kF) * u1 \ (kT * kF) * Suc u0" . then have 1: "enat (p+1) + enat (kT * kF) * Q t \ enat (kT * kF) * eSuc (P s)" unfolding u1 u0 by (simp add: eSuc_enat) from 1 2 show ?thesis by metis next case False with i have "P s + emb (~bval b s) < \" by simp with If4 obtain p t where 1: "(c2, s) \ p \ t" and q: "enat p + enat kF * Q t \ enat kF * (P s + emb (~bval b s))" by blast from Big_StepT.IfFalse[OF False 1] have 2: "(IF b THEN c1 ELSE c2, s) \ p + 1 \ t" by simp from q have "Q t < \" using i F False using less_irrefl by fastforce then obtain u1 where u1: "Q t = enat u1" by auto from q False have q': "p + kF * u1 \ kF * u0" unfolding u0 u1 by auto have "(p+1) + (kF * kT) * u1 \ kT*(p+1) + (kF * kT) * u1" using T by (simp add: mult_eq_if) also have "\ \ kT*(p+1 + kF * u1)" by (simp add: add_mult_distrib2) also have "\ \ kT*(1 + kF * u0)" using q' by auto also have "\ \ (kF * kT) * Suc u0" using F by simp finally have " (p+1) + (kT * kF) * u1 \ (kT * kF) * Suc u0" by (simp add: mult.commute) then have 1: "enat (p+1) + enat (kT * kF) * Q t \ enat (kT * kF) * eSuc (P s)" unfolding u1 u0 by (simp add: eSuc_enat) from 1 2 show ?thesis by metis qed qed next case (conseq P c Q k1 P' Q') from conseq(5) obtain k where c4: "\s. P s < \ \ (\t p. (c, s) \ p \ t \ enat p + enat k * Q t \ enat k * P s)" and 0: "k>0" by blast show ?case apply(rule exI[where x="k*k1"]) proof (safe) show "k*k1>0" using 0 conseq(4) by auto fix s assume "P' s < \" with conseq(2,4) have "P s < \" using le_less_trans by (metis enat.distinct(2) enat_ord_simps(4) imult_is_infinity) with c4 obtain p t where 1: "(c, s) \ p \ t" and 2: "enat p + enat k * Q t \ enat k * P s" by blast have "enat p + enat (k*k1) * Q' t = enat p + enat (k) * ( (enat k1) * Q' t)" by (metis mult.assoc times_enat_simps(1)) also have "\ \ enat p + enat (k) * Q t" using conseq(3) by (metis add_left_mono distrib_left le_iff_add) also have "\ \ enat k * P s" using 2 by auto also have "\ \ enat (k*k1) * P' s" using conseq(2) by (metis mult.assoc mult_left_mono not_less not_less_zero times_enat_simps(1)) finally have 2: "enat p + enat (k*k1) * Q' t \ enat (k*k1) * P' s" by auto from 1 2 show "\t p. (c, s) \ p \ t \ enat p + (k*k1) * Q' t \ (k*k1) * P' s" by auto qed next case (While INV b c) then obtain k where W2: "\s. INV s + \ (bval b s) < \ \ (\t p. (c, s) \ p \ t \ enat p + enat k * (INV t + 1) \ enat k * (INV s + \ (bval b s)))" and g0: "k>0" by blast show ?case apply(rule exI[where x=k]) proof (safe) show "0" then have f: "INV s < \" using enat_ord_simps(4) by fastforce then obtain n where i: "INV s = enat n" using not_infinity_eq by auto have "INV s = enat n \ \t p. (WHILE b DO c, s) \ p \ t \ enat p + enat k * (INV t + emb (\ bval b t)) \ enat k * (INV s + 1)" proof (induct n arbitrary: s rule: less_induct) case (less n) then show ?case proof (cases "bval b s") case False show ?thesis apply(rule exI[where x="s"]) apply(rule exI[where x="Suc 0"]) apply safe apply (fact WhileFalse[OF False]) using False apply (simp add: one_enat_def) using g0 by (metis One_nat_def Suc_ile_eq add.commute add_left_mono distrib_left enat_0_iff(2) mult.right_neutral not_gr_zero one_enat_def) next case True with less(2) W2 have "(\t p. (c, s) \ p \ t \ enat p + enat k * (INV t + 1) \ enat k * INV s )" by force then obtain t p where o: "(c, s) \ p \ t" and q: "enat p + enat k * (INV t + 1) \ enat k * INV s " by auto from o bigstep_progress have p: "p > 0" by blast from q have pf: "enat k * (INV t + 1) \ enat k * INV s" using dual_order.trans by fastforce then have "INV t < \" using less(2) using g0 not_le by fastforce then obtain invt where invt: "INV t = enat invt" by auto from pf g0 have g: "INV t < INV s" unfolding less(2) invt by (metis (full_types) Suc_ile_eq add.commute eSuc_enat enat_ord_simps(1) nat_mult_le_cancel_disj plus_1_eSuc(1) times_enat_simps(1)) then have ninfINVt: "INV t < \" using less(2) using enat_ord_simps(4) by fastforce then obtain n' where i: "INV t = enat n'" using not_infinity_eq by auto with less(2) have ii: "n' < n" using g by auto from i ii less(1) obtain t2 p2 where o2: "(WHILE b DO c, t) \ p2 \ t2" and q2: "enat p2 + enat k * (INV t2 + emb (\ bval b t2)) \ enat k * ( INV t + 1)" by blast have ende: "~ bval b t2" apply(rule ccontr) apply(simp) using q2 g0 ninfINVt by (simp add: i one_enat_def) from WhileTrue[OF True o o2] have "(WHILE b DO c, s) \ 1 + p + p2 \ t2" by simp from ende q2 have q2': "enat p2 + enat k * INV t2 \ enat k * (INV t + 1)" by simp show ?thesis apply(rule exI[where x=t2]) apply(rule exI[where x= "1 + p + p2"]) apply(safe) apply(fact) using ende apply(simp) proof - have "enat (Suc (p + p2)) + enat k * INV t2 = enat (Suc p) + enat p2 + enat k * INV t2" by fastforce also have "\ \ enat (Suc p) + enat k * (INV t + 1)" using q2' by (metis ab_semigroup_add_class.add_ac(1) add_left_mono) also have "\ \ 1 + enat k * (INV s)" using q by (metis (no_types, hide_lams) add.commute add_left_mono eSuc_enat iadd_Suc plus_1_eSuc(1)) also have "\ \ enat k + enat k * (INV s)" using g0 by (simp add: Suc_leI one_enat_def) also have "\ \ enat k * (INV s + 1)" by (simp add: add.commute distrib_left) finally show "enat (Suc (p + p2)) + enat k * INV t2 \ enat k * (INV s + 1)" . qed qed qed from this[OF i] show "\t p. (WHILE b DO c, s) \ p \ t \ enat p + enat k * (INV t + emb (\ bval b t)) \ enat k * (INV s + 1)" . qed qed lemma conseq': "\ \\<^sub>2\<^sub>' {P} c {Q} ; \s. P s \ P' s; \s. Q' s \ Q s \ \ \\<^sub>2\<^sub>' {P'} c {Q'}" apply(rule conseq[where k=1]) by auto lemma strengthen_pre: "\ \s. P s \ P' s; \\<^sub>2\<^sub>' {P} c {Q} \ \ \\<^sub>2\<^sub>' {P'} c {Q}" apply(rule conseq[where k=1 and Q'=Q and Q=Q]) by auto lemma weaken_post: "\ \\<^sub>2\<^sub>' {P} c {Q}; \s. Q s \ Q' s \ \ \\<^sub>2\<^sub>' {P} c {Q'}" apply(rule conseq[where k=1]) by auto lemma Assign': "\s. P s \ eSuc ( Q(s[a/x])) \ \\<^sub>2\<^sub>' {P} x ::= a {Q}" by (simp add: strengthen_pre[OF _ Assign]) subsection "Completeness" lemma bigstep_det: "(c1, s) \ p1 \ t1 \ (c1, s) \ p \ t \ p1=p \ t1=t" using big_step_t_determ2 by simp lemma bigstepT_the_cost: "(c, s) \ P \ T \ (THE n. \a. (c, s) \ n \ a) = P" using bigstep_det by blast lemma bigstepT_the_state: "(c, s) \ P \ T \ (THE a. \n. (c, s) \ n \ a) = T" using bigstep_det by blast lemma SKIPnot: "(\ (SKIP, s) \ p \ t) = (s\t \ p\Suc 0)" by blast lemma SKIPp: "(THE p. \t. (SKIP, s) \ p \ t) = Suc 0" apply(rule the_equality) apply fast apply auto done lemma SKIPt: "(THE t. \p. (SKIP, s) \ p \ t) = s" apply(rule the_equality) apply fast apply auto done lemma ASSp: "(THE p. Ex (big_step_t (x ::= e, s) p)) = Suc 0" apply(rule the_equality) apply fast apply auto done lemma ASSt: "(THE t. \p. (x ::= e, s) \ p \ t) = s(x := aval e s)" apply(rule the_equality) apply fast apply auto done lemma ASSnot: "( \ (x ::= e, s) \ p \ t ) = (p\Suc 0 \ t\s(x := aval e s))" apply auto done text\ The completeness proof proceeds along the same lines as the one for partial correctness. First we have to strengthen our notion of weakest precondition to take termination into account:\ definition wpQ :: "com \ qassn \ qassn" ("wp\<^sub>Q") where "wp\<^sub>Q c Q = (\s. (if (\t p. (c,s) \ p \ t \ Q t < \) then enat (THE p. \t. (c,s) \ p \ t) + Q (THE t. \p. (c,s) \ p \ t) else \))" lemma wpQ_skip[simp]: "wp\<^sub>Q SKIP Q = (%s. eSuc (Q s))" apply(auto intro!: ext simp: wpQ_def) prefer 2 apply(simp only: SKIPnot) apply(simp) apply(simp only: SKIPp SKIPt) using one_enat_def plus_1_eSuc(1) by auto lemma wpQ_ass[simp]: "wp\<^sub>Q (x ::= e) Q = (\s. eSuc (Q (s(x := aval e s))))" by (auto intro!: ext simp: wpQ_def ASSp ASSt ASSnot eSuc_enat) lemma wpt_Seq[simp]: "wp\<^sub>Q (c\<^sub>1;;c\<^sub>2) Q = wp\<^sub>Q c\<^sub>1 (wp\<^sub>Q c\<^sub>2 Q)" unfolding wpQ_def proof (rule, case_tac "\t p. (c\<^sub>1;; c\<^sub>2, s) \ p \ t \ Q t < \", goal_cases) case (1 s) then obtain u p where ter: "(c\<^sub>1;; c\<^sub>2, s) \ p \ u" and Q: "Q u < \" by blast then obtain t p1 p2 where i: "(c\<^sub>1 , s) \ p1 \ t" and ii: "(c\<^sub>2 , t) \ p2 \ u" and p: "p1 + p2 = p" by blast from bigstepT_the_state[OF i] have t: "(THE t. \p. (c\<^sub>1, s) \ p \ t) = t" by blast from bigstepT_the_state[OF ii] have t2: "(THE u. \p. (c\<^sub>2, t) \ p \ u) = u" by blast from bigstepT_the_cost[OF i] have firstcost: "(THE p. \t. (c\<^sub>1, s) \ p \ t) = p1" by blast from bigstepT_the_cost[OF ii] have secondcost: "(THE p. \u. (c\<^sub>2, t) \ p \ u) = p2" by blast have totalcost: "(THE p. Ex (big_step_t (c\<^sub>1;; c\<^sub>2, s) p)) = p1 + p2" using bigstepT_the_cost[OF ter] p by auto have totalstate: "(THE t. \p. (c\<^sub>1;; c\<^sub>2, s) \ p \ t) = u" using bigstepT_the_state[OF ter] by auto have c2: "\ta p. (c\<^sub>2, t) \ p \ ta \ Q ta < \" apply(rule exI[where x= u]) apply(rule exI[where x= p2]) apply safe apply fact+ done have C: "\t p. (c\<^sub>1, s) \ p \ t \ (if \ta p. (c\<^sub>2, t) \ p \ ta \ Q ta < \ then enat (THE p. Ex (big_step_t (c\<^sub>2, t) p)) + Q (THE ta. \p. (c\<^sub>2, t) \ p \ ta) else \) < \" apply(rule exI[where x=t]) apply(rule exI[where x=p1]) apply safe apply fact apply(simp only: c2 if_True) using Q bigstepT_the_state ii by auto show ?case apply(simp only: 1 if_True t t2 c2 C totalcost totalstate firstcost secondcost) by fastforce next case (2 s) show ?case apply(simp only: 2 if_False) apply auto using 2 by force qed lemma wpQ_If[simp]: "wp\<^sub>Q (IF b THEN c\<^sub>1 ELSE c\<^sub>2) Q = (\s. eSuc (wp\<^sub>Q (if bval b s then c\<^sub>1 else c\<^sub>2) Q s))" apply (auto simp: wpQ_def fun_eq_iff) subgoal for x t p i ta ia xa apply(simp only: IfTrue[THEN bigstepT_the_state]) apply(simp only: IfTrue[THEN bigstepT_the_cost]) apply(simp only: bigstepT_the_cost bigstepT_the_state) by (simp add: eSuc_enat) apply(simp only: bigstepT_the_state bigstepT_the_cost) apply force apply(simp only: bigstepT_the_state bigstepT_the_cost) proof(goal_cases) case (1 x t p i ta ia xa) note f= IfFalse[THEN bigstepT_the_state, of b x c\<^sub>2 xa ta "Suc xa" c\<^sub>1, simplified, OF 1(4) 1(5)] note f2= IfFalse[THEN bigstepT_the_cost, of b x c\<^sub>2 xa ta "Suc xa" c\<^sub>1, simplified, OF 1(4) 1(5)] note g= bigstep_det[OF 1(1) 1(5)] show ?case apply(simp only: f f2) using 1 g by (simp add: eSuc_enat) next case 2 then show ?case apply(simp only: bigstepT_the_state bigstepT_the_cost) apply force done qed lemma hoareQ_inf: "\\<^sub>2\<^sub>' {%s. \} c { Q}" apply (induction c arbitrary: Q) apply(auto intro: hoareQ.Skip hoareQ.Assign hoareQ.Seq hoareQ.conseq) subgoal apply(rule hoareQ.conseq) apply(rule hoareQ.If[where P="%s. \"]) by(auto intro: hoareQ.If hoareQ.conseq) subgoal apply(rule hoareQ.conseq) apply(rule hoareQ.While[where I="%s. \"]) apply(rule hoareQ.conseq) by auto done lemma assumes b: "bval b s" shows wpQ_WhileTrue: " wp\<^sub>Q c (wp\<^sub>Q (WHILE b DO c) Q) s + 1 \ wp\<^sub>Q (WHILE b DO c) Q s" proof (cases "\t p. (WHILE b DO c, s) \ p \ t \ Q t < \") case True then obtain t p where w: "(WHILE b DO c, s) \ p \ t" and q: "Q t < \" by blast from b w obtain p1 p2 t1 where c: "(c, s) \ p1 \ t1" and w': "(WHILE b DO c, t1) \ p2 \ t" and sum: "1 + p1 + p2 = p" by auto have g: "\ta p. (WHILE b DO c, t1) \ p \ ta \ Q ta < \" apply(rule exI[where x="t"]) apply(rule exI[where x="p2"]) apply safe apply fact+ done have h: "\t p. (c, s) \ p \ t \ (if \ta p. (WHILE b DO c, t) \ p \ ta \ Q ta < \ then enat (THE p. Ex (big_step_t (WHILE b DO c, t) p)) + Q (THE ta. \p. (WHILE b DO c, t) \ p \ ta) else \) < \" apply(rule exI[where x="t1"]) apply(rule exI[where x="p1"]) apply safe apply fact apply(simp only: g if_True) using bigstepT_the_state bigstepT_the_cost w' q by(auto) have "wp\<^sub>Q c (wp\<^sub>Q (WHILE b DO c) Q) s + 1 = enat p + Q t" unfolding wpQ_def apply(simp only: h if_True) apply(simp only: bigstepT_the_state[OF c] bigstepT_the_cost[OF c] g if_True bigstepT_the_state[OF w'] bigstepT_the_cost[OF w']) using sum by (metis One_nat_def ab_semigroup_add_class.add_ac(1) add.commute add.right_neutral eSuc_enat plus_1_eSuc(2) plus_enat_simps(1)) also have "\ = wp\<^sub>Q (WHILE b DO c) Q s" unfolding wpQ_def apply(simp only: True if_True) using bigstepT_the_state bigstepT_the_cost w apply(simp) done finally show ?thesis by simp next case False have "wp\<^sub>Q (WHILE b DO c) Q s = \" unfolding wpQ_def apply(simp only: False if_False) done then show ?thesis by auto qed lemma assumes b: "~ bval b s" shows wpQ_WhileFalse: " Q s + 1 \ wp\<^sub>Q (WHILE b DO c) Q s" proof (cases "\t p. (WHILE b DO c, s) \ p \ t \ Q t < \") case True with b obtain t p where w: "(WHILE b DO c, s) \ p \ t" and "Q t < \" by blast with b have c: "s=t" "p=Suc 0" by auto have " wp\<^sub>Q (WHILE b DO c) Q s = Q s + 1" unfolding wpQ_def apply(simp only: True if_True) using w c bigstepT_the_cost bigstepT_the_state by(auto simp add: one_enat_def) then show ?thesis by auto next case False have "wp\<^sub>Q (WHILE b DO c) Q s = \" unfolding wpQ_def apply(simp only: False if_False) done then show ?thesis by auto qed lemma wpQ_is_pre: "\\<^sub>2\<^sub>' {wp\<^sub>Q c Q} c { Q}" proof (induction c arbitrary: Q) case SKIP show ?case apply (auto intro: hoareQ.Skip) done next case Assign show ?case apply (auto intro:hoareQ.Assign) done next case Seq thus ?case by (auto intro:hoareQ.Seq) next case (If x1 c1 c2 Q) thus ?case apply (auto intro!: hoareQ.If ) apply(rule hoareQ.conseq) apply(auto) apply(rule hoareQ.conseq) apply(auto) done next case (While b c) show ?case apply(rule conseq[where k=1]) apply(rule hoareQ.While[where I="%s. (if bval b s then wp\<^sub>Q c (wp\<^sub>Q (WHILE b DO c) Q) s else Q s)"]) apply(rule conseq[where k=1]) apply(rule While[of "wp\<^sub>Q (WHILE b DO c) Q"]) apply(case_tac "bval b s") apply(simp) apply(simp) subgoal for s apply(cases "bval b s") using wpQ_WhileTrue apply simp using wpQ_WhileFalse apply simp done apply simp subgoal for s apply(cases "bval b s") using wpQ_WhileTrue apply simp using wpQ_WhileFalse apply simp done apply(case_tac "bval b s") apply(simp) apply(simp) apply simp done qed lemma wpQ_is_pre': "\\<^sub>2\<^sub>' {wp\<^sub>Q c (%s. enat k * Q s )} c {(%s. enat k * Q s )}" using wpQ_is_pre by blast lemma wpQ_is_weakestprePotential1: "\\<^sub>2\<^sub>' {P}c{Q} \ (\k>0. \s. wp\<^sub>Q c (%s. enat k* Q s) s \ enat k * P s)" apply(auto simp: hoare2o_valid_def wpQ_def) proof (goal_cases) case (1 k) show ?case proof (rule exI[where x=k], safe) show "0 p \ t" "enat k * Q t = enat i" show "enat (\\<^sub>t (c, s)) + enat k * Q (\\<^sub>s (c, s)) \ enat k * P s" proof (cases "P s < \") case True with 1 obtain t p' where i: "(c, s) \ p' \ t" and ii: "enat p' + enat k * Q t \ enat k * P s" by auto show ?thesis by(simp add: bigstepT_the_state[OF i] bigstepT_the_cost[OF i] ii) next case False then show ?thesis using 1 by auto qed next fix s assume "\t. (\p. \ (c, s) \ p \ t) \ enat k * Q t = \" then show "enat k * P s = \" using 1 by force qed qed theorem hoareQ_complete: "\\<^sub>2\<^sub>' {P}c{Q} \ \\<^sub>2\<^sub>' {P}c{ Q}" proof - assume "\\<^sub>2\<^sub>' {P}c{Q}" with wpQ_is_weakestprePotential1 obtain k where "k>0" and 1: "\s. wp\<^sub>Q c (\s. enat k * Q s) s \ enat k * P s" by blast show "\\<^sub>2\<^sub>' {P}c{Q}" apply(rule conseq[OF wpQ_is_pre']) apply(fact 1) apply simp by fact qed theorem hoareQ_complete': "\\<^sub>2\<^sub>' {P}c{Q} \ \\<^sub>2\<^sub>' {P}c{ Q}" unfolding hoare2o_valid_def proof - assume "\k>0. \s. P s < \ \ (\t p. (c, s) \ p \ t \ enat p + enat k * Q t \ enat k * P s)" then obtain k where f: "\s. P s < \ \ (\t p. (c, s) \ p \ t \ enat p + enat k * Q t \ enat k * P s)" and k: "k>0" by auto show "\\<^sub>2\<^sub>' {P}c{ Q}" apply(rule conseq[OF wpQ_is_pre', where Q'=Q, simplified, where k1=k and k=k and Q1=Q]) unfolding wpQ_def subgoal for s proof(cases "P s < \") case True with f obtain t p' where i: "(c, s) \ p' \ t" and ii: "enat p' + enat k * Q t \ enat k * P s" by auto from ii k True have iii: "enat k * Q t < \" using imult_is_infinity by fastforce have kla: "\t p. (c, s) \ p \ t \ enat k * Q t < \" using iii i by auto show ?thesis unfolding bigstepT_the_state[OF i] unfolding bigstepT_the_cost[OF i] apply(simp only: kla) using ii by simp next case False then show ?thesis using k by auto qed subgoal by auto using k by auto qed corollary hoareQ_sound_complete: " \\<^sub>2\<^sub>' {P}c{Q} \ \\<^sub>2\<^sub>' {P}c{ Q}" by (metis hoareQ_sound hoareQ_complete) subsection "Example" lemma fixes X::int assumes "0 < X" shows Z: "eSuc (enat (nat (2 * X) * nat (2 * X))) \ enat (5 * (nat (X * X)))" proof - from assms have nn: "0 \ X" by auto from assms have "0 < nat X" by auto then have "0 < enat (nat X)" by (simp add: zero_enat_def) then have A: "eSuc 0 \ enat (nat X)" using ileI1 by blast have " (nat X) \ (nat (X*X))" using nn nat_mult_distrib by auto then have D: "enat (nat X) \ enat (nat (X*X))" by auto have C: "(enat (nat (2 * X) * nat (2 * X))) = 4* enat (nat (X * X))" using nn nat_mult_distrib by (simp add: numeral_eq_enat) have "eSuc (enat (nat (2 * X) * nat (2 * X))) = eSuc 0 + (enat (nat (2 * X) * nat (2 * X)))" using one_eSuc plus_1_eSuc(1) by auto also have "\ \ enat (nat X) + (enat (nat (2 * X) * nat (2 * X)))" using A add_right_mono by blast also have "\ \ enat (nat X) + 4* enat (nat (X * X))" using C by auto also have "\ \ enat (nat (X * X)) + 4* enat (nat (X * X))" using D by auto also have "\ = 5* enat (nat (X * X))" by (metis eSuc_numeral mult_eSuc semiring_norm(5)) also have "\ = enat ( 5* nat (X * X))" by (simp add: numeral_eq_enat) finally show ?thesis . qed lemma weakenpre: "\ \\<^sub>2\<^sub>' {P}c{Q} ; (\s. P s \ P' s) \ \ \\<^sub>2\<^sub>' {P'}c{ Q}" using conseq[where Q'=Q and k=1] by auto lemma whileDecr: "\\<^sub>2\<^sub>' { %s. enat (nat (s ''x'')) + 1} WHILE (Less (N 0) (V ''x'')) DO (SKIP;; SKIP;; ''x'' ::= Plus (V ''x'') (N (-1))) { %s. enat 0}" apply(rule conseq[where k=4]) apply(rule While[where I="%s. enat 4 * (enat (nat (s ''x'')))"]) prefer 2 subgoal for s apply(simp only: one_enat_def plus_enat_simps times_enat_simps enat_ord_code(1)) by presburger apply(rule Seq[where P\<^sub>2="wp\<^sub>Q (''x'' ::= Plus (V ''x'') (N (-1))) (\t. enat 4 * enat (nat (t ''x'')) + 1)"]) apply(simp) apply(rule Seq[where P\<^sub>2="wp\<^sub>Q (SKIP) (\s. eSuc (enat (4 * nat (s ''x'' - 1)) + 1))"]) apply simp subgoal apply(rule weakenpre) apply(rule Skip) apply auto subgoal for s apply(cases "s ''x'' > 0") apply auto apply(simp only: one_enat_def plus_enat_simps times_enat_simps enat_ord_code(1) eSuc_enat) done done subgoal apply simp apply(rule Skip) done subgoal apply simp apply(rule weakenpre) apply(rule Assign) by simp apply simp subgoal for s apply(cases "s ''x'' > 0") by auto by simp lemma whileDecrIf: "\\<^sub>2\<^sub>' { %s. enat (nat (s ''x'')) + 1} WHILE (Less (N 0) (V ''x'')) DO ( (IF Less (N 0) (V ''z'') THEN SKIP;; SKIP ELSE SKIP );; ''x'' ::= Plus (V ''x'') (N (-1))) { %s. enat 0}" apply(rule conseq[OF While, where k=6 and I1="%s. enat 6 * (enat (nat (s ''x'')))"]) prefer 2 subgoal for s apply(simp only: one_enat_def plus_enat_simps times_enat_simps enat_ord_code(1)) by presburger apply(rule Seq[where P\<^sub>2="wp\<^sub>Q (''x'' ::= Plus (V ''x'') (N (-1))) (\t. enat 6 * enat (nat (t ''x'')) + 1)"]) apply(simp) apply(rule weakenpre) apply(rule If[where P="wp\<^sub>Q (IF Less (N 0) (V ''z'') THEN SKIP;; SKIP ELSE SKIP ) (\s. eSuc (enat (6 * nat (s ''x'' - 1)) + 1))"]) subgoal apply simp apply(rule Seq[where P\<^sub>2="wp\<^sub>Q (SKIP) (\s. eSuc (enat (6 * nat (s ''x'' - 1)) + 1))"]) subgoal apply(rule weakenpre) apply(rule Skip) by auto subgoal apply(rule weakenpre) apply(rule Skip) by auto done subgoal apply simp subgoal apply(rule weakenpre) apply(rule Skip) by auto done subgoal apply auto subgoal for s apply(cases "s ''x'' > 0") apply auto apply(simp only: one_enat_def plus_enat_simps times_enat_simps enat_ord_code(1) eSuc_enat) done subgoal for s apply(cases "s ''x'' > 0") apply auto apply(simp only: one_enat_def plus_enat_simps times_enat_simps enat_ord_code(1) eSuc_enat) done done subgoal apply simp apply(rule weakenpre) apply(rule Assign) by simp apply simp subgoal for s apply(cases "s ''x'' > 0") by auto by simp lemma whileDecrIf2: "\\<^sub>2\<^sub>' { %s. enat (nat (s ''x'')) + 1} WHILE (Less (N 0) (V ''x'')) DO ( (IF Less (N 0) (V ''z'') THEN SKIP;; SKIP ELSE SKIP );; ''x'' ::= Plus (V ''x'') (N (-1))) { %s. enat 0}" apply(rule conseq[OF While, where k=6 and I1="%s. enat 6 * (enat (nat (s ''x'')))"]) apply(rule Seq[where P\<^sub>2="wp\<^sub>Q (''x'' ::= Plus (V ''x'') (N (-1))) (\t. enat 6 * enat (nat (t ''x'')) + 1)"]) apply(simp) apply(rule weakenpre) apply(rule If[where P="wp\<^sub>Q (IF Less (N 0) (V ''z'') THEN SKIP;; SKIP ELSE SKIP ) (\s. eSuc (enat (6 * nat (s ''x'' - 1)) + 1))"]) subgoal apply simp apply(rule Seq[where P\<^sub>2="wp\<^sub>Q (SKIP) (\s. eSuc (enat (6 * nat (s ''x'' - 1)) + 1))"]) subgoal apply(rule weakenpre) apply(rule Skip) by auto subgoal apply(rule weakenpre) apply(rule Skip) by auto done subgoal apply simp subgoal apply(rule weakenpre) apply(rule Skip) by auto done prefer 2 subgoal apply simp apply(rule weakenpre) apply(rule Assign) by simp subgoal apply auto subgoal for s apply(cases "s ''x'' > 0") apply auto apply(simp only: one_enat_def plus_enat_simps times_enat_simps enat_ord_code(1) eSuc_enat) done subgoal for s apply(cases "s ''x'' > 0") apply auto apply(simp only: one_enat_def plus_enat_simps times_enat_simps enat_ord_code(1) eSuc_enat) done done subgoal for s apply(simp only: one_enat_def plus_enat_simps times_enat_simps enat_ord_code(1)) by presburger subgoal for s apply(cases "s ''x'' > 0") by auto by simp end \ No newline at end of file diff --git a/thys/Hoare_Time/SepLogK_Hoare.thy b/thys/Hoare_Time/SepLogK_Hoare.thy --- a/thys/Hoare_Time/SepLogK_Hoare.thy +++ b/thys/Hoare_Time/SepLogK_Hoare.thy @@ -1,1186 +1,1186 @@ section "Hoare Logic based on Separation Logic and Time Credits (big-O style)" theory SepLogK_Hoare imports Big_StepT Partial_Evaluation Big_StepT_Partial begin subsection \Definition of Validity\ definition hoare3o_valid :: "assn2 \ com \ assn2 \ bool" - ("\\<^sub>3\<^sub>' {(1_)}/ (_)/ { (1_)}" 50) where + ("\\<^sub>3\<^sub>'' {(1_)}/ (_)/ { (1_)}" 50) where "\\<^sub>3\<^sub>' { P } c { Q } \ (\ k>0. (\ps n. P (ps,n) \ (\ps' ps'' m e e'. ((c,ps) \\<^sub>A m \ ps' + ps'') \ ps' ## ps'' \ k * n = k * e + e' + m \ Q (ps',e))))" subsection "Hoare Rules" inductive hoare3a :: "assn2 \ com \ assn2 \ bool" ("\\<^sub>3\<^sub>a ({(1_)}/ (_)/ { (1_)})" 50) where Skip: "\\<^sub>3\<^sub>a {$1} SKIP { $0}" | Assign4: "\\<^sub>3\<^sub>a { (\(ps,t). x\dom ps \ vars a \ dom ps \ Q (ps(x\(paval a ps)),t) ) ** $1} x::=a { Q }" | If: "\ \\<^sub>3\<^sub>a { \(s,n). P (s,n) \ lmaps_to_axpr b True s } c\<^sub>1 { Q }; \\<^sub>3\<^sub>a { \(s,n). P (s,n) \ lmaps_to_axpr b False s } c\<^sub>2 { Q } \ \ \\<^sub>3\<^sub>a { (\(s,n). P (s,n) \ vars b \ dom s) ** $1} IF b THEN c\<^sub>1 ELSE c\<^sub>2 { Q}" | Frame: "\ \\<^sub>3\<^sub>a { P } C { Q } \ \ \\<^sub>3\<^sub>a { P ** F } C { Q ** F } " | Seq: "\ \\<^sub>3\<^sub>a { P } C\<^sub>1 { Q } ; \\<^sub>3\<^sub>a { Q } C\<^sub>2 { R } \ \ \\<^sub>3\<^sub>a { P } C\<^sub>1 ;; C\<^sub>2 { R } " | While: "\ \\<^sub>3\<^sub>a { (\(s,n). P (s,n) \ lmaps_to_axpr b True s) } C { (\(s,n). P (s,n) \ vars b \ dom s) ** $1 } \ \ \\<^sub>3\<^sub>a { (\(s,n). P (s,n) \ vars b \ dom s) ** $1 } WHILE b DO C { \(s,n). P (s,n) \ lmaps_to_axpr b False s } " | conseqS: "\ \\<^sub>3\<^sub>a {P}c{Q} ; \s n. P' (s,n) \ P (s,n) ; \s n. Q (s,n) \ Q' (s,n) \ \ \\<^sub>3\<^sub>a {P'}c{ Q'}" inductive hoare3b :: "assn2 \ com \ assn2 \ bool" ("\\<^sub>3\<^sub>b ({(1_)}/ (_)/ { (1_)})" 50) where import: "\\<^sub>3\<^sub>a {P} c { Q} \ \\<^sub>3\<^sub>b {P} c { Q}" | conseq: "\ \\<^sub>3\<^sub>b {P}c{Q} ; \s n. P' (s,n) \ P (s,k*n) ; \s n. Q (s,n) \ Q' (s,n div k); k>0 \ \ \\<^sub>3\<^sub>b {P'}c{ Q'}" inductive - hoare3' :: "assn2 \ com \ assn2 \ bool" ("\\<^sub>3\<^sub>' ({(1_)}/ (_)/ { (1_)})" 50) + hoare3' :: "assn2 \ com \ assn2 \ bool" ("\\<^sub>3\<^sub>'' ({(1_)}/ (_)/ { (1_)})" 50) where Skip: "\\<^sub>3\<^sub>' {$1} SKIP { $0}" | Assign: "\\<^sub>3\<^sub>' {lmaps_to_expr_x x a v ** $1} x::=a { (%(s,c). dom s = vars a - {x} \ c = 0) ** x \ v }" | Assign': "\\<^sub>3\<^sub>' {pointsto x v' ** ( pointsto x v \* Q) ** $1} x::= N v { Q }" | Assign2: "\\<^sub>3\<^sub>' {\v . ( ((\v'. pointsto x v') ** ( pointsto x v \* Q) ** $1) and sep_true ** (%(ps,n). vars a \ dom ps \ paval a ps = v ) )} x::= a { Q }" | If: "\ \\<^sub>3\<^sub>' { \(s,n). P (s,n) \ lmaps_to_axpr b True s } c\<^sub>1 { Q }; \\<^sub>3\<^sub>' { \(s,n). P (s,n) \ lmaps_to_axpr b False s } c\<^sub>2 { Q } \ \ \\<^sub>3\<^sub>' { (\(s,n). P (s,n) \ vars b \ dom s) ** $1} IF b THEN c\<^sub>1 ELSE c\<^sub>2 { Q}" | Frame: "\ \\<^sub>3\<^sub>' { P } C { Q } \ \ \\<^sub>3\<^sub>' { P ** F } C { Q ** F } " | Seq: "\ \\<^sub>3\<^sub>' { P } C\<^sub>1 { Q } ; \\<^sub>3\<^sub>' { Q } C\<^sub>2 { R } \ \ \\<^sub>3\<^sub>' { P } C\<^sub>1 ;; C\<^sub>2 { R } " | While: "\ \\<^sub>3\<^sub>' { (\(s,n). P (s,n) \ lmaps_to_axpr b True s) } C { (\(s,n). P (s,n) \ vars b \ dom s) ** $1 } \ \ \\<^sub>3\<^sub>' { (\(s,n). P (s,n) \ vars b \ dom s) ** $1 } WHILE b DO C { \(s,n). P (s,n) \ lmaps_to_axpr b False s } " | conseq: "\ \\<^sub>3\<^sub>' {P}c{Q} ; \s n. P' (s,n) \ P (s,k*n) ; \s n. Q (s,n) \ Q' (s,n div k); k>0 \ \ \\<^sub>3\<^sub>' {P'}c{ Q'}" | normalize: "\ \\<^sub>3\<^sub>' { P ** $m } C { Q ** $n }; n\m \ \ \\<^sub>3\<^sub>' { P ** $(m-n) } C { Q } " | Assign''': "\\<^sub>3\<^sub>' { $1 ** (x \ ds) } x ::= (N v) { (x \ v) }" | Assign'''': "\ symeval P a v; \\<^sub>3\<^sub>' {P} x ::= (N v) {Q'} \ \ \\<^sub>3\<^sub>' {P} x ::= a {Q'}" | Assign4: "\\<^sub>3\<^sub>' { (\(ps,t). x\dom ps \ vars a \ dom ps \ Q (ps(x\(paval a ps)),t) ) ** $1} x::=a { Q }" | False: "\\<^sub>3\<^sub>' { \(ps,n). False } c { Q }" | pureI: "( P \ \\<^sub>3\<^sub>' { Q} c { R}) \ \\<^sub>3\<^sub>' {\P ** Q} c { R}" (* which is more general Assign4 or Assign2 ? *) definition A4 :: "vname \ aexp \ assn2 \ assn2" where "A4 x a Q = ((\(ps,t). x\dom ps \ vars a \ dom ps \ Q (ps(x\(paval a ps)),t) ) ** $1)" definition A2 :: "vname \ aexp \ assn2 \ assn2" where "A2 x a Q = (\v . ( ((\v'. pointsto x v') ** ( pointsto x v \* Q) ** $1) and sep_true ** (%(ps,n). vars a \ dom ps \ paval a ps = v ) ))" lemma "A4 x a Q (ps,n) \ A2 x a Q (ps,n)" unfolding A4_def A2_def sep_conj_def dollar_def sep_impl_def pointsto_def apply auto apply(rule exI[where x="paval a ps"]) apply safe subgoal for n v apply(rule exI[where x="[x \ v]::partstate"]) apply(rule exI[where x=0]) apply auto apply(rule exI[where x="ps(x:=None)"]) apply auto unfolding sep_disj_fun_def domain_conv apply auto unfolding plus_fun_conv apply auto by (simp add: map_add_upd_left map_upd_triv) subgoal for n v apply(rule exI[where x=0]) apply(rule exI[where x=n]) apply(rule exI[where x=ps]) by auto done lemma "A2 x a Q (ps,n) \ A4 x a Q (ps,n)" unfolding A4_def A2_def sep_conj_def dollar_def sep_impl_def pointsto_def apply (auto simp: sep_disj_commute) subgoal for aa ba ab ac bc xa bd apply(rule exI[where x=bd]) by (auto simp: sep_add_ac domain_conv sep_disj_fun_def) subgoal for aa ba ab ac bc xa bd apply(rule exI[where x=bd]) apply (auto simp: sep_add_ac) subgoal apply (auto simp: domain_conv sep_disj_fun_def) by (metis fun_upd_same none_def plus_fun_def) subgoal by (metis domD map_add_dom_app_simps(1) plus_fun_conv subsetCE) subgoal proof - assume a: "ab + [x \ xa] = aa + ac" assume b: "ps = aa + ac" and o: "aa ## ac" then have b': "ps = ac + aa" by(simp add: sep_add_ac) assume vars: "vars a \ dom ac" have pa: "paval a ps = paval a ac" unfolding b' apply(rule paval_extend) using o vars by (simp_all add: sep_add_ac) have f: "\f. (ab + [x \ xa])(x \ f) = ab + [x \ f]" by (simp add: plus_fun_conv) assume "Q (ab + [x \ paval a ac], bd)" thus "Q ((aa + ac)(x \ paval a (aa + ac)), bd)" unfolding b[symmetric] pa unfolding b a[symmetric] pa f by auto qed done done lemma E_extendsR: "\\<^sub>3\<^sub>a { P } c { F } \ \\<^sub>3\<^sub>' { P } c { F }" apply (induct rule: hoare3a.induct) apply(intro hoare3'.Skip) apply(intro hoare3'.Assign4) subgoal using hoare3'.If by auto subgoal using hoare3'.Frame by auto subgoal using hoare3'.Seq by auto subgoal using hoare3'.While by auto subgoal using hoare3'.conseq[where k=1] by simp done lemma E_extendsS: "\\<^sub>3\<^sub>b { P } c { F } \ \\<^sub>3\<^sub>' { P } c { F }" apply (induct rule: hoare3b.induct) apply(erule E_extendsR) using hoare3'.conseq by blast lemma Skip': "P = (F ** $1) \ \\<^sub>3\<^sub>' { P } SKIP { F }" apply(rule conseq[where k=1]) apply(rule Frame[where F=F]) apply(rule Skip) by (auto simp: sep_conj_ac) subsubsection "experiments with explicit and implicit GarbageCollection" lemma "( (\ps n. P (ps,n) \ (\ps' ps'' m e e'. ((c,ps) \\<^sub>A m \ ps' + ps'') \ ps' ## ps'' \ n = e + e' + m \ Q (ps',e)))) \ (\ps n. P (ps,n) \ (\ps' m e . ((c,ps) \\<^sub>A m \ ps') \ n = e + m \ (Q ** (\_. True)) (ps',e)))" proof (safe) fix ps n assume "\ps n. P (ps, n) \ (\ps' ps'' m e e'. (c, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ n = e + e' + m \ Q (ps', e))" "P (ps, n)" then obtain ps' ps'' m e e' where C: "(c, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ n = e + e' + m \ Q (ps', e)" by blast show "\ps' m e. (c, ps) \\<^sub>A m \ ps' \ n = e + m \ (Q ** (\_. True)) (ps',e)" unfolding sep_conj_def apply(rule exI[where x="ps' + ps''"]) apply(rule exI[where x="m"]) apply(rule exI[where x="e+e'"]) using C by auto next fix ps n assume "\ps n. P (ps,n) \ (\ps' m e . ((c,ps) \\<^sub>A m \ ps') \ n = e + m \ (Q ** (\_. True)) (ps',e))" "P (ps, n)" then obtain ps' m e where C: "((c,ps) \\<^sub>A m \ ps') \ n = e + m" and Q: "(Q ** (\_. True)) (ps',e)" by blast from Q obtain ps1 ps2 e1 e2 where Q': "Q (ps1,e1)" "ps'=ps1+ps2" "ps1##ps2" "e=e1+e2" unfolding sep_conj_def by auto show "\ps' ps'' m e e'. (c, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ n = e + e' + m \ Q (ps', e)" apply(rule exI[where x="ps1"]) apply(rule exI[where x="ps2"]) apply(rule exI[where x="m"]) apply(rule exI[where x="e1"]) apply(rule exI[where x="e2"]) using C Q' by auto qed subsection \Soundness\ theorem hoareT_sound2_part: assumes "\\<^sub>3\<^sub>' { P }c{ Q }" shows "\\<^sub>3\<^sub>' { P } c { Q }" using assms proof(induction rule: hoare3'.induct) case (conseq P c Q P' k1 Q') then obtain k where p: "\ps n. P (ps, n) \ (\ps' ps'' m e e'. ((c,ps) \\<^sub>A m \ ps' + ps'') \ ps' ## ps'' \ k * n = k * e + e' + m \ Q (ps',e))" and gt0: "k>0" unfolding hoare3o_valid_def by blast show ?case unfolding hoare3o_valid_def apply(rule exI[where x="k*k1"]) apply safe using gt0 conseq(4) apply simp proof - fix ps n assume "P' (ps,n)" with conseq(2) have "P (ps, k1*n)" by simp with p obtain ps' ps'' m e e' where pB: "(c, ps) \\<^sub>A m \ ps' + ps''" and orth: "ps' ## ps''" and m: "k * (k1 * n) = k*e + e' + m" and Q: "Q (ps', e)" by blast from Q conseq(3) have Q': "Q' (ps', e div k1)" by auto have "k * k1 * n = k*e + e' + m" using m by auto also have "\ = k*(k1 * (e div k1) + e mod k1) + e' + m" using mod_mult_div_eq by simp also have "\ = k*k1*(e div k1) + (k*(e mod k1) + e') + m" by (metis add.assoc distrib_left mult.assoc) finally have "k * k1 * n = k * k1 * (e div k1) + (k * (e mod k1) + e') + m" . show "\ps' ps'' m e e'. (c, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k * k1 * n = k * k1 * e + e' + m \ Q' (ps', e)" apply(rule exI[where x="ps'"]) apply(rule exI[where x="ps''"]) apply(rule exI[where x=m]) apply(rule exI[where x="e div k1"]) apply(rule exI[where x="k * (e mod k1) + e'"]) apply safe apply fact apply fact apply fact apply fact done qed next case (Frame P c Q F) from Frame(2)[unfolded hoare3o_valid_def] obtain k where hyp: "\ps n. P (ps, n) \ (\ps' ps'' m e e'. ((c,ps) \\<^sub>A m \ ps' + ps'') \ ps' ## ps'' \ k * n = k * e + e' + m \ Q (ps',e))" and k: "k>0" unfolding hoare3o_valid_def by blast show ?case unfolding hoare3o_valid_def apply(rule exI[where x=k]) using k apply simp proof(safe) fix ps n assume "(P \* F) (ps, n)" then obtain ps1 ps2 where orth: "ps1 ## ps2" and add: "(ps, n) = ps1 + ps2" and P: "P ps1" and F: "F ps2" unfolding sep_conj_def by blast from hyp P have "(\ps' ps'' m e e'. ((c,fst ps1) \\<^sub>A m \ ps' + ps'') \ ps' ## ps'' \ k * snd ps1 = k * e + e' + m \ Q (ps',e))" by simp then obtain ps' ps'' m e e' where a: "(c, fst ps1) \\<^sub>A m \ ps' + ps''" and orth2[simp]: "ps' ## ps''" and m: "k * snd ps1 = k * e + e' + m" and Q: "Q (ps', e)" by blast from big_step_t3_post_dom_conv[OF a] have dom: "dom (ps' + ps'') = dom (fst ps1)" by auto from add have g: "ps = fst ps1 + fst ps2" and h: "n = snd ps1 + snd ps2" by (auto simp add: plus_prod_def) from orth have [simp]: "fst ps2 ## ps'" "fst ps2 ## ps''" apply (metis dom map_convs(1) orth2 sep_disj_addD1 sep_disj_commuteI sep_disj_fun_def sep_disj_prod_def) by (metis dom map_convs(1) orth orth2 sep_add_commute sep_disj_addD1 sep_disj_commuteI sep_disj_fun_def sep_disj_prod_def) then have e: "ps' ## fst ps2" unfolding sep_disj_fun_def using dom unfolding domain_conv by blast have 3: "(Q \* F) (ps'+fst ps2, e+snd ps2)" unfolding sep_conj_def apply(rule exI[where x="(ps',e)"]) apply(rule exI[where x="ps2"]) apply safe subgoal using orth unfolding sep_disj_prod_def apply (auto simp: sep_disj_nat_def) apply(rule e) done subgoal unfolding plus_prod_def apply auto done apply fact apply fact done show "\ps' ps'' m. (c, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ (\e. (\e'. k * n = k * e + e' + m) \ (Q \* F) (ps', e))" apply(rule exI[where x="ps'+fst ps2"]) apply(rule exI[where x="ps''"]) apply(rule exI[where x=m]) proof safe show "(c, ps) \\<^sub>A m \ ps' + fst ps2 + ps''" apply(rule Framer2[OF _ _ g]) apply (fact a) using orth apply (auto simp: sep_disj_prod_def) by (metis \fst ps2 ## ps''\ \fst ps2 ## ps'\ orth2 sep_add_assoc sep_add_commute sep_disj_commuteI) next show "ps' + fst ps2 ## ps''" by (metis dom map_convs(1) orth orth2 sep_add_disjI1 sep_disj_fun_def sep_disj_prod_def) next show "\e. (\e'. k * n = k * e + e' + m) \ (Q \* F) (ps' + fst ps2, e)" apply(rule exI[where x="e+snd ps2"]) apply safe subgoal proof(rule exI[where x=e']) have "k * n = k * snd ps1 + k * snd ps2" unfolding h by (simp add: distrib_left) also have "\ = k * e + e' + m + k* snd ps2" unfolding m by auto finally show "k * n = k * (e + snd ps2) + e' + m" by algebra qed apply fact done qed qed next case (False c Q) then show ?case by (auto simp: hoare3o_valid_def) next case (Assign2 x Q a) show ?case unfolding hoare3o_valid_def apply (rule exI[where x=1], safe) apply auto proof - fix ps n v assume A: "((\s. \xa. (x \ xa) s) \* (x \ v \* Q) \* $ (Suc 0)) (ps, n)" assume B: "((\s. True) \* (\(ps, n). vars a \ dom ps \ paval a ps = v)) (ps, n)" from A obtain ps1 ps2 n1 n2 v' where "ps1 ## ps2" and add1: "ps = ps1 + ps2" and n: "n = n1 + n2" and 1: "(\xaa. (x \ xaa) (ps1,n1))" and 2: "((x \ v \* Q) \* $ (Suc 0)) (ps2,n2)" unfolding sep_conj_def by fastforce from 2 obtain ps2a ps2b n2a n2b where "ps2a ## ps2b" and add2: "ps2 = ps2a + ps2b" and n2: "n2 = n2a + n2b" and Q: "(x \ v \* Q) (ps2a,n2a)" and ps2b: "ps2b=0" and n2b: "n2b=1" unfolding dollar_def sep_conj_def by fastforce from 1 obtain v' where n1: "n1=0" and p: "ps1 = ([x \ v']::partstate)" and x: "x : dom ps1" by (auto simp: pointsto_def) from x add1 have x: "x : dom ps" by (simp add: plus_fun_conv subset_iff) have f: "([x \ v'] + ps2a)(x \ v) = ps2a + [x \ v]" by (smt \\thesis. (\v'. \n1 = 0; ps1 = [x \ v']; x \ dom ps1\ \ thesis) \ thesis\ \ps1 ## ps2\ add2 disjoint_iff_not_equal dom_fun_upd domain_conv fun_upd_upd map_add_upd_left option.distinct(1) plus_fun_conv ps2b sep_add_commute sep_add_zero sep_disj_fun_def) let ?n' = "n2a + n1" from n n2 n2b have n': "n=1+?n'" by simp have Q': "Q (ps(x \ v), ?n')" using Q n1 unfolding sep_impl_def apply auto unfolding pointsto_def apply auto subgoal by (metis \ps1 ## ps2\ \ps2 = ps2a + ps2b\ \ps2b = 0\ dom_fun_upd domain_conv option.distinct(1) p sep_add_zero sep_disj_commute sep_disj_fun_def) subgoal unfolding add1 p add2 ps2b by (auto simp: f) done from B obtain ps1 ps2 n1 n2 where orth: "ps1 ## ps2" and add: "ps = ps2 + ps1" and n: "n=n1+n2" and vars: "vars a \ dom ps2" and v: "paval a ps2 = v" unfolding sep_conj_def by (auto simp: sep_add_ac) from vars add have a: "vars a \ dom ps" by (simp add: plus_fun_conv subset_iff) from a x have "vars a \ {x} \ dom ps" by auto have "paval a ps = v" unfolding add apply(subst paval_extend) using orth vars v by(auto simp: sep_disj_commute) show "\ps' ps'' m. (x ::= a, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ (\e. (\e'. n = e + e' + m) \ Q (ps', e))" apply(rule exI[where x="ps(x\v)"]) apply(rule exI[where x=0]) apply(rule exI[where x="Suc 0"]) apply auto apply(rule big_step_t_part.Assign) apply fact+ apply simp apply(rule exI[where x="?n'"]) apply safe apply(rule exI[where x=0]) using n' apply simp using Q' by auto qed next case Skip then show ?case by (auto simp: hoare3o_valid_def dollar_def) next case (Assign4 x a Q) then show ?case apply (auto simp: dollar_def sep_conj_def hoare3o_valid_def ) apply(rule exI[where x=1]) apply auto subgoal for ps b y apply(rule exI[where x="ps(x \ paval a ps)"]) apply(rule exI[where x="0"]) apply(rule exI[where x="Suc 0"]) apply auto apply(rule exI[where x=b]) by auto done next case (Assign' x v' v Q ) have "\aa. aa ## [x \ v'] \ \ aa ## [x \ v] \ False" unfolding sep_disj_fun_def domain_def apply auto by (smt Collect_conj_eq Collect_empty_eq) have f: "\v'. domain [x \ v'] = {x}" unfolding domain_conv by auto { fix ps assume u: "ps ## [x \ v']" have 2: "[x \ v'] + ps = ps + [x \ v']" "[x \ v] + ps = ps + [x \ v]" subgoal apply (subst sep_add_commute) using u by (auto simp: sep_add_ac) subgoal apply (subst sep_add_commute) using u apply (auto simp: sep_add_ac) unfolding sep_disj_fun_def f by auto done have "(x ::= N v, [x \ v'] + ps) \\<^sub>A Suc 0 \ [x \ v] + ps" apply(rule Framer[OF big_step_t_part.Assign]) apply simp_all using u by (auto simp: sep_add_ac) then have "(x ::= N v, ps + [x \ v']) \\<^sub>A Suc 0 \ ps + [x \ v]" by (simp only: 2) } note f2 = this from Assign' show ?case apply (auto simp: dollar_def sep_conj_def pointsto_def sep_impl_def hoare3o_valid_def ) apply(rule exI[where x=1]) apply (auto simp: sep_add_ac) subgoal unfolding sep_disj_fun_def f by auto subgoal for ps n apply(rule exI[where x="ps+[x \ v]"]) apply(rule exI[where x="0"]) apply(rule exI[where x="Suc 0"]) apply safe subgoal using f2 by auto subgoal by auto subgoal by force done done next case (Assign x a v) then show ?case unfolding hoare3o_valid_def apply(rule exI[where x=1]) apply auto apply (auto simp: dollar_def ) subgoal for ps n apply (subst (asm) separate_othogonal) apply auto apply(rule exI[where x="ps(x:=Some v)"]) apply(rule exI[where x="0"]) apply(rule exI[where x="1"]) apply auto apply (auto simp: pointsto_def) unfolding sep_conj_def subgoal apply(rule exI[where x="((%y. if y=x then None else ps y) , 0)"]) apply(rule exI[where x="((%y. if y = x then Some (paval a ps) else None),0)"]) apply (auto simp: sep_disj_prod_def sep_disj_fun_def plus_fun_def) apply (smt domIff domain_conv) apply (metis domI insertE option.simps(3)) using domIff by fastforce done done next case (If P b c\<^sub>1 Q c\<^sub>2) from If(3)[unfolded hoare3o_valid_def] obtain k1 where T: "\ps n. P (ps, n) \ vars b \ dom ps \ pbval b ps \ (\ps' ps'' m e e'. (c\<^sub>1, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k1 * n = k1 * e + e' + m \ Q (ps', e))" and k1: "k1 > 0" by force from If(4)[unfolded hoare3o_valid_def] obtain k2 where F: "\ps n. P (ps, n) \ vars b \ dom ps \ \ pbval b ps \ (\ps' ps'' m e e'. (c\<^sub>2, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k2 * n = k2 * e + e' + m \ Q (ps', e))" and k2: "k2 > 0" by force show ?case unfolding hoare3o_valid_def apply auto apply (auto simp: dollar_def) apply(rule exI[where x="k1 * k2"]) using k1 k2 apply auto proof (goal_cases) case (1 ps n) then obtain n' where P: "P (ps, n')" and dom: "vars b \ dom ps" and Suc: "n = Suc n'" unfolding sep_conj_def by force show ?case proof(cases "pbval b ps") case True with T[OF P dom] obtain ps' ps'' m e e' where d: "(c\<^sub>1, ps) \\<^sub>A m \ ps' + ps''" and orth: "ps' ## ps''" and m1: "k1 * n' = k1 * e + e' + m" and Q: "Q (ps', e)" by blast from big_step_t3_post_dom_conv[OF d] have klong: "dom (ps' + ps'') = dom ps" . from k1 obtain k1' where k1': "k1 = Suc k1'" using gr0_implies_Suc by blast from k2 obtain k2' where k2': "k2 = Suc k2'" using gr0_implies_Suc by blast let ?e1 = "(k2' * (e' + m + k1) + e' + k1')" show ?thesis apply(rule exI[where x=ps']) apply(rule exI[where x=ps'']) apply(rule exI[where x="m+1"]) apply safe apply(rule big_step_t_part.IfTrue) apply (rule dom) apply fact apply (rule True) apply (rule d) apply simp apply fact subgoal apply(rule exI[where x="e"]) apply safe subgoal proof (rule exI[where x="?e1"]) have "k1 * k2 * n = k2 * (k1*n)" by auto also have "\ = k2 * (k1*n' + k1)" unfolding Suc by auto also have "\ = k2 * (k1 * e + e' + m + k1)" unfolding m1 by auto also have "\ = k1 * k2 * e + k2* (e' + m + k1)" by algebra also have "\ = k1 * k2 * e + k2' * (e' + m + k1) + (e' + m + k1)" unfolding k2' by simp also have "\ = k1 * k2 * e + k2' * (e' + m + k1) + (e' + k1' + m + 1)" unfolding k1' by simp also have "\ = k1 *k2*e + (k2' * (e' + m + k1) + e' + k1') + (m+1)" by algebra finally show "k1 * k2 * n = k1 * k2 * e + ?e1 + (m + 1)" . qed using Q by force done next case False with F[OF P dom] obtain ps' ps'' m e e' where d: "(c\<^sub>2, ps) \\<^sub>A m \ ps' + ps''" and orth: "ps' ## ps''" and m2: "k2 * n' = k2 * e + e' + m" and Q: "Q (ps', e)" by blast from big_step_t3_post_dom_conv[OF d] have klong: "dom (ps' + ps'') = dom ps" . from k1 obtain k1' where k1': "k1 = Suc k1'" using gr0_implies_Suc by blast from k2 obtain k2' where k2': "k2 = Suc k2'" using gr0_implies_Suc by blast let ?e2 = "(k1' * (e' + m + k2) + e' + k2')" show ?thesis apply(rule exI[where x=ps']) apply(rule exI[where x=ps'']) apply(rule exI[where x="m+1"]) apply safe apply(rule big_step_t_part.IfFalse) apply (rule dom) apply fact apply (rule False) apply (rule d) apply simp apply fact subgoal apply(rule exI[where x="e"]) apply safe subgoal proof (rule exI[where x="?e2"]) have "k1 * k2 * n = k1 * (k2*n)" by auto also have "\ = k1 * (k2*n' + k2)" unfolding Suc by auto also have "\ = k1 * (k2 * e + e' + m + k2)" unfolding m2 by auto also have "\ = k1 * k2 * e + k1* (e' + m + k2)" by algebra also have "\ = k1 * k2 * e + k1' * (e' + m + k2) + (e' + m + k2)" unfolding k1' by simp also have "\ = k1 * k2 * e + k1' * (e' + m + k2) + (e' + k2' + m + 1)" unfolding k2' by simp also have "\ = k1 *k2*e + (k1' * (e' + m + k2) + e' + k2') + (m+1)" by algebra finally show "k1 * k2 * n = k1 * k2 * e + ?e2 + (m + 1)" . qed using Q by force done qed qed next case (Seq P C\<^sub>1 Q C\<^sub>2 R) from Seq(3)[unfolded hoare3o_valid_def] obtain k1 where 1: "(\ps n. P (ps, n) \ (\ps' ps'' m e e'. (C\<^sub>1, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k1 * n = k1 * e + e' + m \ Q (ps', e)))" and k10: "k1 > 0" by blast from Seq(4)[unfolded hoare3o_valid_def] obtain k2 where 2: " (\ps n. Q (ps, n) \ (\ps' ps'' m e e'. (C\<^sub>2, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k2 * n = k2 * e + e' + m \ R (ps', e)))" and k20: "k2 > 0" by blast from k10 obtain k1' where k1': "k1 = Suc k1'" using gr0_implies_Suc by blast from k20 obtain k2' where k2': "k2 = Suc k2'" using gr0_implies_Suc by blast show ?case unfolding hoare3o_valid_def apply(rule exI[where x="k2*k1"]) proof safe fix ps n assume "P (ps, n)" with 1 obtain ps1' ps1'' m1 e1 e1' where C1: "(C\<^sub>1, ps) \\<^sub>A m1 \ ps1' + ps1''" and orth: "ps1' ## ps1''" and m1: "k1 * n = k1 * e1 + e1' + m1"and Q: "Q (ps1', e1)" by blast from Q and 2 obtain ps2' ps2'' m2 e2 e2' where C2: "(C\<^sub>2, ps1') \\<^sub>A m2 \ ps2' + ps2''" and orth2: "ps2' ## ps2''" and m2: "k2 * e1 = k2 * e2 + e2' + m2" and R: "R (ps2', e2)" by blast let ?ee = "(k1 *e2' + k2*e1' +k2'*m1+ k1'*m2)" show "\ps' ps'' m e e'. (C\<^sub>1;; C\<^sub>2, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k2 * k1 * n = k2 * k1 * e + e' + m \ R (ps', e)" apply(rule exI[where x=ps2']) apply(rule exI[where x="ps2'' + ps1''"]) apply(rule exI[where x="m1+m2"]) apply(rule exI[where x="e2"]) apply(rule exI[where x="?ee"]) proof safe have C2': "(C\<^sub>2, ps1' + ps1'') \\<^sub>A m2 \ ps2' + (ps2'' + ps1'')" apply(rule Framer2[OF C2, of ps1'']) apply fact apply simp using sep_add_assoc by (metis C2 big_step_t3_post_dom_conv map_convs(1) orth orth2 sep_add_commute sep_disj_addD1 sep_disj_commuteI sep_disj_fun_def) show "(C\<^sub>1;; C\<^sub>2, ps) \\<^sub>A m1 + m2 \ ps2' + (ps2'' + ps1'')" using C1 C2' by auto next show "ps2' ## ps2'' + ps1''" by (metis C2 big_step_t3_post_dom_conv map_convs(1) orth orth2 sep_disj_addI3 sep_disj_fun_def) next have "k2 * k1 * n = k2 * (k1 * n)" by auto also have "\ = k2 * (k1 * e1 + e1' + m1)" using m1 by auto also have "\ = k1 * k2 * e1 + k2 * (e1' + m1)" by algebra also have "\ = k1 * (k2 * e2 + e2' + m2) + k2 * (e1' + m1)" using m2 by auto also have "\ = k2 * k1 * e2 + (k1 *e2' + k2*e1' +k2*m1+ k1*m2)" by algebra also have "\ = k2 * k1 * e2 + (k1 *e2' + k2*e1' +k2'*m1+m1+ k1'*m2+m2)" unfolding k1' k2' by auto also have "\ = k2 * k1 * e2 + (k1 *e2' + k2*e1' +k2'*m1+ k1'*m2)+(m1+m2)" by auto finally show "k2 * k1 * n = k2 * k1 * e2 + ?ee + (m1 + m2)" . qed fact qed (simp add: k10 k20) next case (While P b C) { assume "\k>0. \ps n. (case (ps, n) of (s, n) \ P (s, n) \ lmaps_to_axpr b True s) \ (\ps' ps'' m e e'. (C, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k * n = k * e + e' + m \ ((\(s, n). P (s, n) \ vars b \ dom s) \* $ 1) (ps', e))" then obtain k where While2: "\ps n. (case (ps, n) of (s, n) \ P (s, n) \ lmaps_to_axpr b True s) \ (\ps' ps'' m e e'. (C, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k * n = k * e + e' + m \ ((\(s, n). P (s, n) \ vars b \ dom s) \* $ 1) (ps', e))" and k: "k>0" by blast from k obtain k' where k': "k = Suc k'" using gr0_implies_Suc by blast have "\k>0. \ps n. ((\(s, n). P (s, n) \ vars b \ dom s) \* $ 1) (ps, n) \ (\ps' ps'' m e e'. (WHILE b DO C, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k * n = k * e + e' + m \ (case (ps', e) of (s, n) \ P (s, n) \ lmaps_to_axpr b False s))" (* have "\k>0. \ps n. ((\(s, n). P (s, n) \ vars b \ dom s) \* $ 1) (ps, n) \ (\ps' m. (WHILE b DO C, ps) \\<^sub>A m \ ps' \ m \ k * n \ (case (ps', n - m) of (s, n) \ P (s, n) \ lmaps_to_axpr b False s))" *) proof (rule exI[where x=k], safe, goal_cases) case (2 ps n) from 2 show ?case proof(induct n arbitrary: ps rule: less_induct) case (less x ps3) show ?case proof(cases "pbval b ps3") case True from less(2) obtain x' where P: "P (ps3, x')" and dom: "vars b \ dom ps3" and Suc: "x = Suc x'" unfolding sep_conj_def dollar_def by auto from P dom True have g: "((\(s, n). P (s, n) \ lmaps_to_axpr b True s)) (ps3, x')" unfolding dollar_def by auto from While2 g obtain ps3' ps3'' m e e' where C: "(C, ps3) \\<^sub>A m \ ps3' + ps3''" and orth: "ps3' ## ps3''" and x: "k * x' = k * e + e' + m" and P': "((\(s, n). P (s, n) \ vars b \ dom s) \* $ 1) (ps3', e)" by blast then obtain x''' where P'': "P (ps3', x''')" and domb: "vars b \ dom ps3'" and Suc'': "e = Suc x'''" unfolding sep_conj_def dollar_def by auto from C big_step_t3_post_dom_conv have "dom ps3 = dom (ps3' + ps3'')" by simp with dom have dom': "vars b \ dom (ps3' + ps3'')" by auto from C big_step_t3_gt0 have gt0: "m > 0" by auto have "e < x" using x Suc gt0 by (metis k le_add1 le_less_trans less_SucI less_add_same_cancel1 nat_mult_less_cancel1) (*have "\ps' m. (WHILE b DO C, ps3') \\<^sub>A m \ ps' \ m \ k * (x - (1 + m)) \ P (ps', (x - (1 + m)) - m) \ vars b \ dom ps' \ \ pbval b ps'" *) have "\ps' ps'' m e2 e2'. (WHILE b DO C, ps3') \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k * e = k * e2 + e2' + m \ P (ps', e2) \ lmaps_to_axpr b False ps'" apply(rule less(1)) apply fact by fact then obtain ps4' ps4'' mt et et' where w: "((WHILE b DO C, ps3') \\<^sub>A mt \ ps4' + ps4'')" and ortho: "ps4' ## ps4''" and m'': "k * e = k * et + et' + mt" and P'': "P (ps4', et)" and dom'': "vars b \ dom ps4'" and b'': "\ pbval b ps4'" by auto have "ps4'' ## ps3''" and "ps4' ## ps3''" by (metis big_step_t3_post_dom_conv domain_conv orth ortho sep_add_disjD sep_disj_fun_def w)+ show ?thesis apply(rule exI[where x="ps4'"]) apply(rule exI[where x="(ps4'' + ps3'')"]) apply(rule exI[where x="1 + m + mt"]) apply(rule exI[where x="et"]) apply(rule exI[where x=" et' + k' + e'"]) proof (safe) have "(WHILE b DO C, ps3' + ps3'') \\<^sub>A mt \ ps4' + (ps4'' + ps3'')" apply(rule Framer2[OF w, of ps3'']) apply fact apply simp apply(rule sep_add_assoc[symmetric]) by fact+ show "(WHILE b DO C, ps3) \\<^sub>A 1 + m + mt \ ps4' + (ps4'' + ps3'')" apply(rule WhileTrue) apply fact apply fact apply (fact C) apply fact by auto next show "ps4' ## ps4'' + ps3''" by (metis big_step_t3_post_dom_conv domain_conv orth ortho sep_disj_addI3 sep_disj_fun_def w) next have "k * x = k * x' + k" unfolding Suc by auto also have "\ = k * e + e' + m + k" unfolding x by simp also have "\ = k * et + et' + mt + e' + m + k" using m'' by simp also have "\ = k * et + et' + mt + e' + m + 1 + k'" using k' by simp also have "\ = k * et + ( et' + k' + e') + (1 + m + mt)" using k' by simp finally show "k * x = k * et + ( et' + k' + e') + (1 + m + mt)" by simp next show "P (ps4', et)" by fact next show "lmaps_to_axpr b False ps4'" apply simp using dom'' b'' .. qed next case False from less(2) obtain x' where P: "P (ps3, x')" and dom: "vars b \ dom ps3" and Suc: "x = Suc x'" unfolding dollar_def sep_conj_def by force show ?thesis apply(rule exI[where x=ps3]) apply(rule exI[where x=0]) apply(rule exI[where x="Suc 0"]) apply(rule exI[where x="x'"]) apply(rule exI[where x="k'"]) apply safe apply simp apply (rule big_step_t_part.WhileFalse) subgoal using dom by simp apply fact apply simp using Suc k k' apply simp using P Suc apply simp using dom apply auto using False apply auto done qed qed qed (fact) } with While(2) show ?case unfolding hoare3o_valid_def by simp next case (Assign''' x ds v) then show ?case unfolding hoare3o_valid_def apply auto apply(rule exI[where x=1]) apply auto subgoal for ps n apply(rule exI[where x="ps(x\v)"]) apply(rule exI[where x="0"]) apply(rule exI[where x="Suc 0"]) apply safe apply(rule big_step_t_part.Assign) apply (auto) subgoal apply(subst (asm) separate_othogonal_commuted') by(auto simp: dollar_def pointsto_def) subgoal apply(subst (asm) separate_othogonal_commuted') by(auto simp: dollar_def pointsto_def) done done next case (Assign'''' P a v x Q') from Assign''''(3)[unfolded hoare3o_valid_def] obtain k where k: "k>0" and A: " \ps n. P (ps, n) \ (\ps' ps'' m e e'. (x ::= N v, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k * n = k * e + e' + m \ Q' (ps', e)) " by auto show ?case unfolding hoare3o_valid_def apply auto apply(rule exI[where x=k]) using k apply auto proof (goal_cases) case (1 ps n) with A obtain ps' ps'' m e e' where " (x ::= N v, ps) \\<^sub>A m \ ps' + ps''" and orth: "ps' ## ps''" and m: "k * n = k * e + e' + m" and Q: "Q' (ps', e)" by metis from 1(2) Assign''''(1)[unfolded symeval_def] have "paval' a ps = Some v" by auto show ?case apply(rule exI[where x=ps']) apply(rule exI[where x=ps'']) apply(rule exI[where x=m]) apply safe apply(rule avalDirekt3_correct) apply fact+ apply(rule exI[where x=e]) apply safe apply(rule exI[where x=e']) apply fact apply fact done qed next case (pureI P Q c R) show ?case proof (cases "P") case True with pureI(2)[unfolded hoare3o_valid_def] obtain k where k: "k>0" and thing: " \ps n. Q (ps, n) \ (\ps' ps'' m e e'. (c, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k * n = k * e + e' + m \ R (ps', e))" by auto show ?thesis unfolding hoare3o_valid_def apply(rule exI[where x=k]) apply safe apply fact using thing by fastforce next case False show ?thesis unfolding hoare3o_valid_def apply(rule exI[where x=1]) using False by auto qed next case (normalize P m C Q n) then show ?case unfolding hoare3o_valid_def apply(safe) subgoal for k apply(rule exI[where x=k]) apply safe proof (goal_cases) case (1 ps N) have Q2: "P (ps, N - (m - n))" apply(rule stardiff) by fact have mn: "m - n \ N" apply(rule stardiff(2)) by fact have P: "(P \* $ m) (ps, N - (m - n) + m)" unfolding sep_conj_def dollar_def apply(rule exI[where x="(ps,N - (m - n))"]) apply(rule exI[where x="(0,m)"]) apply(auto simp: sep_disj_prod_def sep_disj_nat_def) by fact have N: " N - (m - n) + m = N +n" using normalize(2) using mn by auto from P N have P': "(P \* $ m) (ps, N +n)" by auto from P' 1(4) obtain ps' ps'' m' e e' where "(C, ps) \\<^sub>A m' \ ps' + ps''" and orth: "ps' ## ps''" and m': "k * (N +n) = k * e + e' + m'" and Q: "(Q \* $ n) (ps', e)" by blast have Q2: "Q (ps', e - n)" apply(rule stardiff) by fact have en: "e\n" using Q using stardiff(2) by blast show ?case apply(rule exI[where x="ps'"]) apply(rule exI[where x="ps''"]) apply(rule exI[where x="m'"]) apply(rule exI[where x="e - n"]) apply(rule exI[where x="e'"]) proof (safe) show "(C, ps) \\<^sub>A m' \ ps' + ps''" by fact next show "ps' ## ps''" by fact next have "k * N = k * ( (N + n)- n)" by auto also have "\ = k*(N + n) - k*n" using right_diff_distrib' by blast also have "\ = (k * e + e' + m') - k*n" using m' by auto also have "\ = k * e - k*n + e' + m'" using en by (metis Nat.add_diff_assoc2 ab_semigroup_add_class.add_ac(1) distrib_left le_add1 le_add_diff_inverse) also have "\ = k * (e-n) + e' + m'" by (simp add: diff_mult_distrib2) finally show "k * N = k * (e - n) + e' + m'" . next show "Q (ps', e - n)" by fact qed qed done qed thm hoareT_sound2_part E_extendsR lemma hoareT_sound2_partR: "\\<^sub>3\<^sub>a {P} c { Q} \ \\<^sub>3\<^sub>' {P} c {Q}" using hoareT_sound2_part E_extendsR by blast subsubsection \nice example\ lemma Frame_R: assumes "\\<^sub>3\<^sub>' { P } C { Q }" "Frame P' P F" shows "\\<^sub>3\<^sub>' { P' } C { Q ** F } " apply(rule conseq[where k=1]) apply(rule Frame) apply(rule assms(1)) using assms(2) unfolding Frame_def by auto lemma strengthen_post: assumes "\\<^sub>3\<^sub>' {P}c{Q}" "\s. Q s \ Q' s" shows "\\<^sub>3\<^sub>' {P}c{ Q'}" apply(rule conseq[where k=1]) apply (rule assms(1)) apply simp apply simp apply fact apply simp done lemmas strongAssign = Assign''''[OF _ strengthen_post, OF _ Frame_R, OF _ Assign'''] lemma weakenpre: "\ \\<^sub>3\<^sub>' {P}c{Q} ; \s. P' s \ P s \ \ \\<^sub>3\<^sub>' {P'}c{ Q}" using conseq[where k=1] by auto lemma weakenpreR: "\ \\<^sub>3\<^sub>a {P}c{Q} ; \s. P' s \ P s \ \ \\<^sub>3\<^sub>a {P'}c{ Q}" using hoare3a.conseqS by auto subsection "Completeness" -definition wp3' :: "com \ assn2 \ assn2" ("wp\<^sub>3\<^sub>'") where +definition wp3' :: "com \ assn2 \ assn2" ("wp\<^sub>3\<^sub>''") where "wp\<^sub>3\<^sub>' c Q = (\(s,n). \t m. n\m \ (c,s) \\<^sub>A m \ t \ Q (t,n-m))" lemma wp3Skip[simp]: "wp\<^sub>3\<^sub>' SKIP Q = (Q ** $1)" apply (auto intro!: ext simp: wp3'_def) unfolding sep_conj_def dollar_def sep_disj_prod_def sep_disj_nat_def apply auto apply force subgoal for t n apply(rule exI[where x=t]) apply(rule exI[where x="Suc 0"]) using big_step_t_part.Skip by auto done lemma wp3Assign[simp]: "wp\<^sub>3\<^sub>' (x ::= e) Q = ((\(ps,t). vars e \ {x} \ dom ps \ Q (ps(x \ paval e ps),t)) ** $1)" apply (auto intro!: ext simp: wp3'_def ) unfolding sep_conj_def apply (auto simp: sep_disj_prod_def sep_disj_nat_def dollar_def) apply force by fastforce lemma wpt_Seq[simp]: "wp\<^sub>3\<^sub>' (c\<^sub>1;;c\<^sub>2) Q = wp\<^sub>3\<^sub>' c\<^sub>1 (wp\<^sub>3\<^sub>' c\<^sub>2 Q)" apply (auto simp: wp3'_def fun_eq_iff ) subgoal for a b t m1 s2 m2 apply(rule exI[where x="s2"]) apply(rule exI[where x="m1"]) apply simp apply(rule exI[where x="t"]) apply(rule exI[where x="m2"]) apply simp done subgoal for s m t' m1 t m2 apply(rule exI[where x="t"]) apply(rule exI[where x="m1+m2"]) apply (auto simp: big_step_t_part.Seq) done done lemma wp3If[simp]: "wp\<^sub>3\<^sub>' (IF b THEN c\<^sub>1 ELSE c\<^sub>2) Q = ((\(ps,t). vars b \ dom ps \ wp\<^sub>3\<^sub>' (if pbval b ps then c\<^sub>1 else c\<^sub>2) Q (ps,t)) ** $1)" apply (auto simp: wp3'_def fun_eq_iff) unfolding sep_conj_def apply (auto simp: sep_disj_prod_def sep_disj_nat_def dollar_def) subgoal for a ba t x apply(rule exI[where x="ba - 1"]) apply auto apply(rule exI[where x=t]) apply(rule exI[where x=x]) apply auto done subgoal for a ba t x apply(rule exI[where x="ba - 1"]) apply auto apply(rule exI[where x=t]) apply(rule exI[where x=x]) apply auto done subgoal for a ba t m apply(rule exI[where x=t]) apply(rule exI[where x="Suc m"]) apply auto apply(cases "pbval b a") subgoal apply simp apply(subst big_step_t_part.IfTrue) using big_step_t3_post_dom_conv by auto subgoal apply simp apply(subst big_step_t_part.IfFalse) using big_step_t3_post_dom_conv by auto done done lemma sFTrue: assumes "pbval b ps" "vars b \ dom ps" shows "wp\<^sub>3\<^sub>' (WHILE b DO c) Q (ps, n) = ((\(ps, n). vars b \ dom ps \ (if pbval b ps then wp\<^sub>3\<^sub>' c (wp\<^sub>3\<^sub>' (WHILE b DO c) Q) (ps, n) else Q (ps, n))) \* $ 1) (ps, n)" (is "?wp = (?I \* $ 1) _") proof assume "wp\<^sub>3\<^sub>' (WHILE b DO c) Q (ps, n)" from this[unfolded wp3'_def] obtain ps'' tt where tn: "tt \ n" and w1: "(WHILE b DO c, ps) \\<^sub>A tt \ ps''" and Q: "Q (ps'', n - tt) " by blast with assms obtain t t' ps' where w2: "(WHILE b DO c, ps') \\<^sub>A t' \ ps''" and c: "(c, ps) \\<^sub>A t \ ps'" and tt: "tt=1+t+t'" by auto from tn obtain k where n: "n=tt+k" using le_Suc_ex by blast from assms show "(?I \* $ 1) (ps,n)" unfolding sep_conj_def dollar_def wp3'_def apply auto apply(rule exI[where x="t+t'+k"]) apply safe subgoal using n tt by auto apply(rule exI[where x="ps'"]) apply(rule exI[where x="t"]) using c apply auto apply(rule exI[where x="ps''"]) apply(rule exI[where x="t'"]) using w2 Q n by auto next assume "(?I \* $ 1) (ps,n)" with assms have Q: "wp\<^sub>3\<^sub>' c (wp\<^sub>3\<^sub>' (WHILE b DO c) Q) (ps, n-1)" and n: "n\1" unfolding dollar_def sep_conj_def by auto then obtain t ps' t' ps'' where t: "t \ n - 1" and c: "(c, ps) \\<^sub>A t \ ps'" and t': "t' \ (n-1) - t" and w: "(WHILE b DO c, ps') \\<^sub>A t' \ ps''" and Q: "Q (ps'', ((n-1) - t) - t')" unfolding wp3'_def by auto show "?wp" unfolding wp3'_def apply simp apply(rule exI[where x="ps''"]) apply(rule exI[where x="1+t+t'"]) apply safe subgoal using t t' n by simp subgoal using c w assms by auto subgoal using Q t t' n by simp done qed lemma sFFalse: assumes "~ pbval b ps" "vars b \ dom ps" shows "wp\<^sub>3\<^sub>' (WHILE b DO c) Q (ps, n) = ((\(ps, n). vars b \ dom ps \ (if pbval b ps then wp\<^sub>3\<^sub>' c (wp\<^sub>3\<^sub>' (WHILE b DO c) Q) (ps, n) else Q (ps, n))) \* $ 1) (ps, n)" (is "?wp = (?I \* $ 1) _") proof assume "wp\<^sub>3\<^sub>' (WHILE b DO c) Q (ps, n)" from this[unfolded wp3'_def] obtain ps' t where tn: "t \ n" and w1: "(WHILE b DO c, ps) \\<^sub>A t \ ps'" and Q: "Q (ps', n - t) " by blast from assms have w2: "(WHILE b DO c, ps) \\<^sub>A 1 \ ps" by auto from w1 w2 big_step_t_determ2 have t1: "t=1" and pps: "ps=ps'" by auto from assms show "(?I \* $ 1) (ps,n)" unfolding sep_conj_def dollar_def using t1 tn Q pps apply auto apply(rule exI[where x="n-1"]) by auto next assume "(?I \* $ 1) (ps,n)" with assms have Q: "Q(ps,n-1)" "n\1" unfolding dollar_def sep_conj_def by auto from assms have w2: "(WHILE b DO c, ps) \\<^sub>A 1 \ ps" by auto show "?wp" unfolding wp3'_def apply auto apply(rule exI[where x=ps]) apply(rule exI[where x=1]) using Q w2 by auto qed lemma sF': "wp\<^sub>3\<^sub>' (WHILE b DO c) Q (ps,n) = ((\(ps, n). vars b \ dom ps \ (if pbval b ps then wp\<^sub>3\<^sub>' c (wp\<^sub>3\<^sub>' (WHILE b DO c) Q) (ps, n) else Q (ps, n))) \* $ 1) (ps,n)" apply(cases "vars b \ dom ps") subgoal apply(cases "pbval b ps") using sFTrue sFFalse by auto subgoal by (auto simp add: dollar_def wp3'_def sep_conj_def) done lemma sF: "wp\<^sub>3\<^sub>' (WHILE b DO c) Q s = ((\(ps, n). vars b \ dom ps \ (if pbval b ps then wp\<^sub>3\<^sub>' c (wp\<^sub>3\<^sub>' (WHILE b DO c) Q) (ps, n) else Q (ps, n))) \* $ 1) s" using sF' by (metis (mono_tags, lifting) prod.case_eq_if prod.collapse sep_conj_impl1) lemma strengthen_postR: assumes "\\<^sub>3\<^sub>a {P}c{Q}" "\s. Q s \ Q' s" shows "\\<^sub>3\<^sub>a {P}c{ Q'}" apply(rule hoare3a.conseqS) apply (rule assms(1)) apply simp by (fact assms(2)) lemma assumes "\Q. \\<^sub>3\<^sub>a {wp\<^sub>3\<^sub>' c Q} c {Q}" shows WhileWpisPre: "\\<^sub>3\<^sub>a {wp\<^sub>3\<^sub>' (WHILE b DO c) Q} WHILE b DO c { Q}" proof - define I where "I \ (\(ps, n). vars b \ dom ps \ (if pbval b ps then wp\<^sub>3\<^sub>' c (wp\<^sub>3\<^sub>' (WHILE b DO c) Q) (ps, n) else Q (ps, n)))" define I' where "I' \ (\(ps, n). (if pbval b ps then wp\<^sub>3\<^sub>' c (wp\<^sub>3\<^sub>' (WHILE b DO c) Q) (ps, n) else Q (ps, n)))" have I': "I = (\(ps, n). vars b \ dom ps \ I' (ps, n))" unfolding I_def I'_def by auto from assms[where Q="(wp\<^sub>3\<^sub>' (WHILE b DO c) Q)"] have c: "\\<^sub>3\<^sub>a {wp\<^sub>3\<^sub>' c (wp\<^sub>3\<^sub>' (WHILE b DO c) Q)} c {(wp\<^sub>3\<^sub>' (WHILE b DO c) Q)}" . have c': "\\<^sub>3\<^sub>a { (\(s,n). I (s,n) \ lmaps_to_axpr b True s) } c { I ** $1 }" apply(rule hoare3a.conseqS) apply(rule c) subgoal apply auto unfolding I_def by auto subgoal unfolding I_def using sF by auto done have c'': "\\<^sub>3\<^sub>a { (\(s,n). I (s,n) \ lmaps_to_axpr b True s) } c { (\(s, n). I (s, n) \ vars b \ dom s) ** $1 }" apply(rule strengthen_postR[OF c']) unfolding I' by (smt R case_prod_beta prod.sel(1) prod.sel(2)) have ka: "(\(s, n). I (s, n) \ vars b \ dom s) = I" apply rule unfolding I' by auto from hoare3a.While[where P="I"] c'' have w: "\\<^sub>3\<^sub>a { (\(s,n). I (s,n) \ vars b \ dom s) ** $1 } WHILE b DO c { \(s,n). I (s,n) \ lmaps_to_axpr b False s }" . show "\\<^sub>3\<^sub>a {wp\<^sub>3\<^sub>' (WHILE b DO c) Q} WHILE b DO c { Q}" apply(rule hoare3a.conseqS) apply(rule w) subgoal unfolding ka using sF I_def by simp subgoal unfolding I_def by auto done qed lemma wpT_is_pre: "\\<^sub>3\<^sub>a {wp\<^sub>3\<^sub>' c Q} c { Q}" proof (induction c arbitrary: Q) case SKIP then show ?case apply auto using hoare3a.Frame[where F=Q and Q="$0" and P="$1", OF hoare3a.Skip] by (auto simp: sep.add_ac) next case (Assign x1 x2) then show ?case using hoare3a.Assign4 by simp next case (Seq c1 c2) then show ?case apply auto apply(subst hoare3a.Seq[rotated]) by auto next case (If x1 c1 c2) then show ?case apply auto apply(rule weakenpreR[OF hoare3a.If, where P1="%(ps,n). wp\<^sub>3\<^sub>' (if pbval x1 ps then c1 else c2) Q (ps,n)"]) apply auto subgoal apply(rule hoare3a.conseqS[where P="wp\<^sub>3\<^sub>' c1 Q" and Q=Q]) by auto subgoal apply(rule hoare3a.conseqS[where P="wp\<^sub>3\<^sub>' c2 Q" and Q=Q]) by auto proof - fix a b assume "((\(ps, t). vars x1 \ dom ps \ wp\<^sub>3\<^sub>' (if pbval x1 ps then c1 else c2) Q (ps, t)) \* $ (Suc 0)) (a, b)" then show "((\(ps, t). wp\<^sub>3\<^sub>' (if pbval x1 ps then c1 else c2) Q (ps, t) \ vars x1 \ dom ps) \* $ (Suc 0)) (a, b)" unfolding sep_conj_def apply auto apply(case_tac "pbval x1 aa") apply auto done (* strange! *) qed next case (While b c) with WhileWpisPre show ?case . qed lemma hoare3o_valid_alt: "\\<^sub>3\<^sub>' { P } c { Q } \ (\ k>0. (\ps n. P (ps,n div k) \ (\ps' ps'' m e e'. ((c,ps) \\<^sub>A m \ ps' + ps'') \ ps' ## ps'' \ n = e + e' + m \ Q (ps',e div k))))" proof - assume "\\<^sub>3\<^sub>' {P} c { Q}" from this[unfolded hoare3o_valid_def] obtain k where k0: "k>0" and P: "\ps n. P (ps, n) \ (\ps' ps'' m e e'. (c, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k * n = k * e + e' + m \ Q (ps', e)) " by blast show ?thesis apply(rule exI[where x=k]) apply safe apply fact proof - fix ps n assume "P (ps, n div k)" with P obtain ps' ps'' m e e' where 1: "(c, ps) \\<^sub>A m \ ps' + ps''" "ps' ## ps''" and e: "k * (n div k) = k * e + e' + m" and Q: "Q (ps', e)" by blast have "k * (n div k) \ n" using k0 by simp then obtain e'' where "n = k * (n div k) + e''" using le_Suc_ex by blast also have "\ = k * e + e' + e'' + m" using e by auto finally have "n = k * e + (e'+e'') + m" and "Q (ps', (k*e) div k)" using Q k0 by auto with 1 show " \ps' ps'' m e e'. (c, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ n = e + e' + m \ Q (ps', e div k)" by blast qed qed lemma valid_alternative_with_GC: assumes "(\ps n. P (ps,n) \ (\ps' ps'' m e e'. ((c,ps) \\<^sub>A m \ ps' + ps'') \ ps' ## ps'' \ n = e + e' + m \ Q (ps',e))) " shows "(\ps n. P (ps,n) \ (\ps' m e . ((c,ps) \\<^sub>A m \ ps' ) \ n = e + m \ (Q ** sep_true) (ps',e)))" proof safe fix ps n assume P: "P (ps,n)" with assms obtain ps' ps'' m e e' where c: "(c,ps) \\<^sub>A m \ ps' + ps''" and ps: "ps' ## ps''" and n: "n = e + e' + m" and Q: "Q (ps',e)" by blast show "\ps' m e. (c, ps) \\<^sub>A m \ ps' \ n = e + m \ (Q \* (\s. True)) (ps', e)" apply(rule exI[where x="ps' + ps''"]) apply(rule exI[where x=m]) apply(rule exI[where x="e+e'"]) apply safe apply fact apply fact unfolding sep_conj_def apply simp apply(rule exI[where x=ps']) apply(rule exI[where x=e]) apply(rule exI[where x=ps'']) apply safe apply fact apply(rule exI[where x=e']) apply simp apply fact done qed lemma hoare3o_valid_GC: "\\<^sub>3\<^sub>' {P} c { Q } \ \\<^sub>3\<^sub>' {P} c { Q ** sep_true}" proof - assume "\\<^sub>3\<^sub>' {P} c { Q }" then obtain k where "k>0" and P: "\ps n. P (ps, n) \ (\ps' ps'' m e e'. (c, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k * n = k * e + e' + m \ Q (ps', e))" unfolding hoare3o_valid_def by blast show "\\<^sub>3\<^sub>' {P} c { Q ** sep_true}" unfolding hoare3o_valid_def apply(rule exI[where x=k]) apply safe apply fact proof - fix ps n assume "P (ps, n)" with P obtain ps' ps'' m e e' where "(c, ps) \\<^sub>A m \ ps' + ps''"" ps' ## ps'' "" k * n = k * e + e' + m " and Q: " Q (ps', e)" by blast show "\ps' ps'' m e e'. (c, ps) \\<^sub>A m \ ps' + ps'' \ ps' ## ps'' \ k * n = k * e + e' + m \ (Q \* (\s. True)) (ps', e)" apply(rule exI[where x=ps']) apply(rule exI[where x=ps'']) apply(rule exI[where x=m]) apply(rule exI[where x=e]) apply(rule exI[where x=e']) apply safe apply fact+ unfolding sep_conj_def apply(rule exI[where x="(ps', e)"]) apply(rule exI[where x="0"]) using Q by simp qed qed lemma hoare3a_sound_GC: "\\<^sub>3\<^sub>a {P} c { Q} \ \\<^sub>3\<^sub>' {P} c { Q ** sep_true}" using hoare3o_valid_GC hoareT_sound2_partR by auto lemma valid_wp: "\\<^sub>3\<^sub>' {P} c { Q} \ (\k>0. \s n. P (s, n) \ wp\<^sub>3\<^sub>' c (\(ps, n). (Q ** sep_true) (ps, n div k)) (s, k * n))" proof - let ?P = "\k (ps,n). P (ps,n div k)" let ?Q = "\k (ps,n). Q (ps,n div k)" let ?QG = "\k (ps,n). (Q ** sep_true) (ps,n div k)" assume "\\<^sub>3\<^sub>' {P} c { Q}" then obtain k where k[simp]: "k>0" and "(\ps n. P (ps,n div k) \ (\ps' ps'' m e e'. ((c,ps) \\<^sub>A m \ ps' + ps'') \ ps' ## ps'' \ n = e + e' + m \ Q (ps',e div k)))" using hoare3o_valid_alt by blast then have "(\ps n. ?P k (ps,n) \ (\ps' ps'' m e e'. ((c,ps) \\<^sub>A m \ ps' + ps'') \ ps' ## ps'' \ n = e + e' + m \ ?Q k (ps',e)))" by auto then have f: "(\ps n. ?P k (ps,n) \ (\ps' m e . ((c,ps) \\<^sub>A m \ ps' ) \ n = e + m \ (?Q k ** sep_true) (ps',e)))" apply(rule valid_alternative_with_GC) done have "\s n. P (s, n) \ wp\<^sub>3\<^sub>' c (\(ps, n). (Q ** sep_true) (ps, n div k)) (s, k * n)" unfolding wp3'_def apply auto proof - fix ps n assume "P (ps,n)" then have "P (ps,(k*n) div k)" apply simp done with f obtain ps' m e where "((c,ps) \\<^sub>A m \ ps' )" and z: "k * n = e + m" and Q: "(?Q k ** sep_true) (ps',e)" by blast from z have e: "e = k * n -m " by auto from Q[unfolded e sep_conj_def] obtain ps1 ps2 e1 e2 where "ps1 ## ps2" "(ps' = ps1 +ps2)" and eq: "k * n - m = e1 + e2" and Q: "Q (ps1, e1 div k)" by force let ?f = "(e1 + e2) div k - (e1 div k + (e2 div k))" have kl: "(e1 + e2) div k \ (e1 div k + (e2 div k))" using k using div_add1_eq le_iff_add by blast show "\t m. m \ k * n \ (c, ps) \\<^sub>A m \ t \ (Q \* (\s. True)) (t, (k * n - m) div k)" apply(rule exI[where x=ps']) apply(rule exI[where x=m]) apply safe using z apply simp apply fact unfolding e unfolding sep_conj_def apply(rule exI[where x="(ps1,e1 div k)"]) apply(rule exI[where x="(ps2,e2 div k+?f)"]) apply auto apply fact+ unfolding eq using kl apply force using Q by auto qed then show "(\k>0. \s n. P (s, n) \ wp\<^sub>3\<^sub>' c (\(ps, n). (Q ** sep_true) (ps, n div k)) (s, k * n))" using k by metis qed theorem completeness: "\\<^sub>3\<^sub>' {P} c { Q} \ \\<^sub>3\<^sub>b {P} c { Q ** sep_true}" proof - let ?P = "\k (ps,n). P (ps,n div k)" let ?Q = "\k (ps,n). Q (ps,n div k)" let ?QG = "\k (ps,n). (Q ** sep_true) (ps,n div k)" assume "\\<^sub>3\<^sub>' {P} c { Q}" then obtain k where k[simp]: "k>0" and P: "\s n. P (s, n) \ wp\<^sub>3\<^sub>' c (\(ps, n). (Q ** sep_true) (ps, n div k)) (s, k * n)" using valid_wp by blast from wpT_is_pre have R: "\\<^sub>3\<^sub>a {wp\<^sub>3\<^sub>' c (?QG k)} c {?QG k}" by auto show "\\<^sub>3\<^sub>b {P} c { Q ** sep_true}" apply(rule hoare3b.conseq[OF hoare3b.import[OF R], where k=k]) subgoal for s n by (fact P) apply simp by (fact k) qed thm E_extendsR completeness lemma completenessR: "\\<^sub>3\<^sub>' {P} c { Q} \ \\<^sub>3\<^sub>' {P} c { Q ** sep_true}" using E_extendsS completeness by metis end diff --git a/thys/Hybrid_Logic/Hybrid_Logic.thy b/thys/Hybrid_Logic/Hybrid_Logic.thy --- a/thys/Hybrid_Logic/Hybrid_Logic.thy +++ b/thys/Hybrid_Logic/Hybrid_Logic.thy @@ -1,4820 +1,4820 @@ theory Hybrid_Logic imports "HOL-Library.Countable" begin section \Syntax\ datatype ('a, 'b) fm = Pro 'a | Nom 'b | Neg \('a, 'b) fm\ (\\<^bold>\ _\ [40] 40) | Dis \('a, 'b) fm\ \('a, 'b) fm\ (infixr \\<^bold>\\ 30) | Dia \('a, 'b) fm\ (\\<^bold>\ _\ 10) | Sat 'b \('a, 'b) fm\ (\\<^bold>@ _ _\ 10) text \We can give other connectives as abbreviations.\ abbreviation Top (\\<^bold>\\) where \\<^bold>\ \ (undefined \<^bold>\ \<^bold>\ undefined)\ abbreviation Con (infixr \\<^bold>\\ 35) where \p \<^bold>\ q \ \<^bold>\ (\<^bold>\ p \<^bold>\ \<^bold>\ q)\ abbreviation Imp (infixr \\<^bold>\\ 25) where \p \<^bold>\ q \ \<^bold>\ (p \<^bold>\ \<^bold>\ q)\ abbreviation Box (\\<^bold>\ _\ 10) where \\<^bold>\ p \ \<^bold>\ (\<^bold>\ \<^bold>\ p)\ primrec nominals :: \('a, 'b) fm \ 'b set\ where \nominals (Pro x) = {}\ | \nominals (Nom i) = {i}\ | \nominals (\<^bold>\ p) = nominals p\ | \nominals (p \<^bold>\ q) = nominals p \ nominals q\ | \nominals (\<^bold>\ p) = nominals p\ | \nominals (\<^bold>@ i p) = {i} \ nominals p\ primrec sub :: \('b \ 'c) \ ('a, 'b) fm \ ('a, 'c) fm\ where \sub _ (Pro x) = Pro x\ | \sub f (Nom i) = Nom (f i)\ | \sub f (\<^bold>\ p) = (\<^bold>\ sub f p)\ | \sub f (p \<^bold>\ q) = (sub f p \<^bold>\ sub f q)\ | \sub f (\<^bold>\ p) = (\<^bold>\ sub f p)\ | \sub f (\<^bold>@ i p) = (\<^bold>@ (f i) (sub f p))\ lemma sub_nominals: \nominals (sub f p) = f ` nominals p\ by (induct p) auto lemma sub_id: \sub id p = p\ by (induct p) simp_all lemma sub_upd_fresh: \i \ nominals p \ sub (f(i := j)) p = sub f p\ by (induct p) auto section \Semantics\ text \ Type variable \'w\ stands for the set of worlds and \'a\ for the set of propositional symbols. The accessibility relation is given by \R\ and the valuation by \V\. The mapping from nominals to worlds is an extra argument \g\ to the semantics.\ datatype ('w, 'a) model = Model (R: \'w \ 'w set\) (V: \'w \ 'a \ bool\) primrec semantics :: \('w, 'a) model \ ('b \ 'w) \ 'w \ ('a, 'b) fm \ bool\ (\_, _, _ \ _\ [50, 50, 50] 50) where \(M, _, w \ Pro x) = V M w x\ | \(_, g, w \ Nom i) = (w = g i)\ | \(M, g, w \ \<^bold>\ p) = (\ M, g, w \ p)\ | \(M, g, w \ (p \<^bold>\ q)) = ((M, g, w \ p) \ (M, g, w \ q))\ | \(M, g, w \ \<^bold>\ p) = (\v \ R M w. M, g, v \ p)\ | \(M, g, _ \ \<^bold>@ i p) = (M, g, g i \ p)\ lemma \M, g, w \ \<^bold>\\ by simp lemma semantics_fresh: \i \ nominals p \ (M, g, w \ p) = (M, g(i := v), w \ p)\ by (induct p arbitrary: w) auto subsection \Examples\ abbreviation is_named :: \('w, 'b) model \ bool\ where \is_named M \ \w. \a. V M a = w\ abbreviation reflexive :: \('w, 'b) model \ bool\ where \reflexive M \ \w. w \ R M w\ abbreviation irreflexive :: \('w, 'b) model \ bool\ where \irreflexive M \ \w. w \ R M w\ abbreviation symmetric :: \('w, 'b) model \ bool\ where \symmetric M \ \v w. w \ R M v \ v \ R M w\ abbreviation asymmetric :: \('w, 'b) model \ bool\ where \asymmetric M \ \v w. \ (w \ R M v \ v \ R M w)\ abbreviation transitive :: \('w, 'b) model \ bool\ where \transitive M \ \v w x. w \ R M v \ x \ R M w \ x \ R M v\ abbreviation universal :: \('w, 'b) model \ bool\ where \universal M \ \v w. v \ R M w\ lemma \irreflexive M \ M, g, w \ \<^bold>@ i \<^bold>\ (\<^bold>\ Nom i)\ proof - assume \irreflexive M\ then have \g i \ R M (g i)\ by simp then have \\ (\v \ R M (g i). g i = v)\ by simp then have \\ M, g, g i \ \<^bold>\ Nom i\ by simp then have \M, g, g i \ \<^bold>\ (\<^bold>\ Nom i)\ by simp then show \M, g, w \ \<^bold>@ i \<^bold>\ (\<^bold>\ Nom i)\ by simp qed text \We can automatically show some characterizations of frames by pure axioms.\ lemma \irreflexive M = (\g w. M, g, w \ \<^bold>@ i \<^bold>\ (\<^bold>\ Nom i))\ by auto lemma \asymmetric M = (\g w. M, g, w \ \<^bold>@ i (\<^bold>\ \<^bold>\ (\<^bold>\ Nom i)))\ by auto lemma \universal M = (\g w. M, g, w \ \<^bold>\ Nom i)\ by auto section \Tableau\ text \ A block is defined as a list of formulas and an opening nominal. The opening nominal is not necessarily in the list. A branch is a list of blocks. All blocks have an opening nominal; there is no possibility of an unnamed initial block. \ type_synonym ('a, 'b) block = \('a, 'b) fm list \ 'b\ type_synonym ('a, 'b) branch = \('a, 'b) block list\ abbreviation member_list :: \'a \ 'a list \ bool\ (\_ \. _\ [51, 51] 50) where \x \. xs \ x \ set xs\ text \The predicate \on\ accounts for the opening nominal when checking if a formula is on a block.\ primrec on :: \('a, 'b) fm \ ('a, 'b) block \ bool\ (\_ on _\ [51, 51] 50) where \p on (ps, i) = (p \. ps \ p = Nom i)\ syntax "_Ballon" :: \pttrn \ 'a set \ bool \ bool\ (\(3\(_/on_)./ _)\ [0, 0, 10] 10) "_Bexon" :: \pttrn \ 'a set \ bool \ bool\ (\(3\(_/on_)./ _)\ [0, 0, 10] 10) translations "\p on A. P" \ "\p. p on A \ P" "\p on A. P" \ "\p. p on A \ P" abbreviation list_nominals :: \('a, 'b) fm list \ 'b set\ where \list_nominals ps \ (\p \ set ps. nominals p)\ primrec block_nominals :: \('a, 'b) block \ 'b set\ where \block_nominals (ps, i) = {i} \ list_nominals ps\ definition branch_nominals :: \('a, 'b) branch \ 'b set\ where \branch_nominals branch \ (\block \ set branch. block_nominals block)\ abbreviation at_in :: \('a, 'b) fm \ 'b \ ('a, 'b) branch \ bool\ (\_ at _ in _\ [51, 51, 51] 50) where \p at a in branch \ \ps. (ps, a) \. branch \ p on (ps, a)\ definition new :: \('a, 'b) fm \ 'b \ ('a, 'b) branch \ bool\ where \new p a branch \ \ p at a in branch\ definition witnessed :: \('a, 'b) fm \ 'b \ ('a, 'b) branch \ bool\ where \witnessed p a branch \ \i. (\<^bold>@ i p) at a in branch \ (\<^bold>\ Nom i) at a in branch\ text \ A branch has a closing tableau iff it is contained in the following inductively defined set. In that case I call the branch closeable. The number in front of the turnstile is the number of available coins to spend on GoTo. \ inductive ST :: \nat \ ('a, 'b) branch \ bool\ (\_ \ _\ [50, 50] 50) where Close: \p at i in branch \ (\<^bold>\ p) at i in branch \ n \ branch\ | Neg: \(\<^bold>\ \<^bold>\ p) at a in (ps, a) # branch \ new p a ((ps, a) # branch) \ Suc n \ (p # ps, a) # branch \ n \ (ps, a) # branch\ | DisP: \(p \<^bold>\ q) at a in (ps, a) # branch \ new p a ((ps, a) # branch) \ new q a ((ps, a) # branch) \ Suc n \ (p # ps, a) # branch \ Suc n \ (q # ps, a) # branch \ n \ (ps, a) # branch\ | DisN: \(\<^bold>\ (p \<^bold>\ q)) at a in (ps, a) # branch \ new (\<^bold>\ p) a ((ps, a) # branch) \ new (\<^bold>\ q) a ((ps, a) # branch) \ Suc n \ ((\<^bold>\ q) # (\<^bold>\ p) # ps, a) # branch \ n \ (ps, a) # branch\ | DiaP: \(\<^bold>\ p) at a in (ps, a) # branch \ i \ branch_nominals ((ps, a) # branch) \ \a. p = Nom a \ \ witnessed p a ((ps, a) # branch) \ Suc n \ ((\<^bold>@ i p) # (\<^bold>\ Nom i) # ps, a) # branch \ n \ (ps, a) # branch\ | DiaN: \(\<^bold>\ (\<^bold>\ p)) at a in (ps, a) # branch \ (\<^bold>\ Nom i) at a in (ps, a) # branch \ new (\<^bold>\ (\<^bold>@ i p)) a ((ps, a) # branch) \ Suc n \ ((\<^bold>\ (\<^bold>@ i p)) # ps, a) # branch \ n \ (ps, a) # branch\ | SatP: \(\<^bold>@ a p) at b in (ps, a) # branch \ new p a ((ps, a) # branch) \ Suc n \ (p # ps, a) # branch \ n \ (ps, a) # branch\ | SatN: \(\<^bold>\ (\<^bold>@ a p)) at b in (ps, a) # branch \ new (\<^bold>\ p) a ((ps, a) # branch) \ Suc n \ ((\<^bold>\ p) # ps, a) # branch \ n \ (ps, a) # branch\ | GoTo: \i \ branch_nominals branch \ n \ ([], i) # branch \ Suc n \ branch\ | Nom: \p at b in (ps, a) # branch \ Nom i at b in (ps, a) # branch \ Nom i at a in (ps, a) # branch \ new p a ((ps, a) # branch) \ Suc n \ (p # ps, a) # branch \ n \ (ps, a) # branch\ abbreviation ST_ex :: \('a, 'b) branch \ bool\ (\\ _\ [50] 50) where \\ branch \ \n. n \ branch\ lemma ST_Suc: \n \ branch \ Suc n \ branch\ by (induct n branch rule: ST.induct) (simp_all add: ST.intros) text \A verified derivation in the calculus.\ lemma fixes i defines \p \ \<^bold>\ (\<^bold>@ i (Nom i))\ shows \Suc n \ [([p], a)]\ proof - have \i \ branch_nominals [([p], a)]\ unfolding p_def branch_nominals_def by simp then have ?thesis if \n \ [([], i), ([p], a)]\ using that GoTo by fast moreover have \new (\<^bold>\ Nom i) i [([], i), ([p], a)]\ unfolding p_def new_def by auto moreover have \(\<^bold>\ (\<^bold>@ i (Nom i))) at a in [([], i), ([p], a)]\ unfolding p_def by fastforce ultimately have ?thesis if \Suc n \ [([\<^bold>\ Nom i], i), ([p], a)]\ using that SatN by fast then show ?thesis by (meson Close list.set_intros(1) on.simps) qed section \Soundness\ text \ An \i\-block is satisfied by a model \M\ and assignment \g\ if all formulas on the block are true under \M\ at the world \g i\ A branch is satisfied by a model and assignment if all blocks on it are. \ primrec block_sat :: \('w, 'a) model \ ('b \ 'w) \ ('a, 'b) block \ bool\ (\_, _ \\<^sub>B _\ [50, 50] 50) where \(M, g \\<^sub>B (ps, i)) = (\p on (ps, i). M, g, g i \ p)\ abbreviation branch_sat :: \('w, 'a) model \ ('b \ 'w) \ ('a, 'b) branch \ bool\ (\_, _ \\<^sub>\ _\ [50, 50] 50) where \M, g \\<^sub>\ branch \ \(ps, i) \ set branch. M, g \\<^sub>B (ps, i)\ lemma block_nominals: \p on block \ i \ nominals p \ i \ block_nominals block\ by (induct block) auto lemma block_sat_fresh: assumes \M, g \\<^sub>B block\ \i \ block_nominals block\ shows \M, g(i := v) \\<^sub>B block\ using assms proof (induct block) case (Pair ps a) then have \\p on (ps, a). i \ nominals p\ using block_nominals by fast moreover have \i \ a\ using calculation by simp ultimately have \\p on (ps, a). M, g(i := v), (g(i := v)) a \ p\ using Pair semantics_fresh by fastforce then show ?case by (meson block_sat.simps) qed lemma branch_sat_fresh: assumes \M, g \\<^sub>\ branch\ \i \ branch_nominals branch\ shows \M, g(i := v) \\<^sub>\ branch\ using assms using block_sat_fresh unfolding branch_nominals_def by fast text \If a branch has a derivation then it cannot be satisfied.\ lemma soundness': \n \ branch \ M, g \\<^sub>\ branch \ False\ proof (induct branch arbitrary: g rule: ST.induct) case (Close p i branch) then have \M, g, g i \ p\ \M, g, g i \ \<^bold>\ p\ by fastforce+ then show ?case by simp next case (Neg p a ps branch) have \M, g, g a \ p\ using Neg(1, 5) by fastforce then have \M, g \\<^sub>\ (p # ps, a) # branch\ using Neg(5) by simp then show ?case using Neg(4) by blast next case (DisP p q a ps branch) consider \M, g, g a \ p\ | \M, g, g a \ q\ using DisP(1, 8) by fastforce then consider \M, g \\<^sub>\ (p # ps, a) # branch\ | \M, g \\<^sub>\ (q # ps, a) # branch\ using DisP(8) by auto then show ?case using DisP(5, 7) by metis next case (DisN p q a ps branch) have \M, g, g a \ \<^bold>\ p\ \M, g, g a \ \<^bold>\ q\ using DisN(1, 5) by fastforce+ then have \M, g \\<^sub>\ ((\<^bold>\ q) # (\<^bold>\ p) # ps, a) # branch\ using DisN(5) by simp then show ?case using DisN(4) by blast next case (DiaP p a ps branch i) then have *: \M, g \\<^sub>B (ps, a)\ by simp have \i \ nominals p\ using DiaP(1-2) unfolding branch_nominals_def by fastforce have \M, g, g a \ \<^bold>\ p\ using DiaP(1, 7) by fastforce then obtain v where \v \ R M (g a)\ \M, g, v \ p\ by auto then have \M, g(i := v), v \ p\ using \i \ nominals p\ semantics_fresh by metis then have \M, g(i := v), g a \ \<^bold>@ i p\ by simp moreover have \M, g(i := v), g a \ \<^bold>\ Nom i\ using \v \ R M (g a)\ by simp moreover have \M, g(i := v) \\<^sub>\ (ps, a) # branch\ using DiaP(2, 7) branch_sat_fresh by fast moreover have \i \ block_nominals (ps, a)\ using DiaP(2) unfolding branch_nominals_def by simp then have \\p on (ps, a). M, g(i := v), g a \ p\ using * semantics_fresh by fastforce ultimately have \M, g(i := v) \\<^sub>\ ((\<^bold>@ i p) # (\<^bold>\ Nom i) # ps, a) # branch\ by auto then show ?case using DiaP by blast next case (DiaN p a ps branch i) have \M, g, g a \ \<^bold>\ (\<^bold>\ p)\ \M, g, g a \ \<^bold>\ Nom i\ using DiaN(1-2, 6) by fastforce+ then have \M, g, g a \ \<^bold>\ (\<^bold>@ i p)\ by simp then have \M, g \\<^sub>\ ((\<^bold>\ (\<^bold>@ i p)) # ps, a) # branch\ using DiaN(6) by simp then show ?thesis using DiaN(5) by blast next case (SatP a p b ps branch) have \M, g, g a \ p\ using SatP(1, 5) by fastforce then have \M, g \\<^sub>\ (p # ps, a) # branch\ using SatP(5) by simp then show ?case using SatP(4) by blast next case (SatN a p b ps branch) have \M, g, g a \ \<^bold>\ p\ using SatN(1, 5) by fastforce then have \M, g \\<^sub>\ ((\<^bold>\ p) # ps, a) # branch\ using SatN(5) by simp then show ?case using SatN(4) by blast next case (GoTo i branch) then show ?case by auto next case (Nom p b ps a branch i g) have \M, g, g b \ p\ \M, g, g b \ Nom i\ using Nom(1-2, 7) by fastforce+ then have \M, g, g i \ p\ by simp moreover have \M, g, g a \ Nom i\ using Nom(3, 7) by fastforce moreover have \M, g \\<^sub>B (ps, a)\ using Nom(7) by simp ultimately have \M, g \\<^sub>B (p # ps, a)\ by simp then have \M, g \\<^sub>\ (p # ps, a) # branch\ using Nom(7) by simp then show ?case using Nom(6) by blast qed lemma block_sat: \\p on block. M, g, w \ p \ M, g \\<^sub>B block\ by (induct block) auto lemma branch_sat: assumes \\(ps, i) \ set branch. \p on (ps, i). M, g, w \ p\ shows \M, g \\<^sub>\ branch\ using assms block_sat by fast lemma soundness: assumes \n \ branch\ shows \\block \ set branch. \p on block. \ M, g, w \ p\ using assms soundness' branch_sat by fast corollary \\ n \ []\ using soundness by fastforce theorem soundness_fresh: assumes \n \ [([\<^bold>\ p], i)]\ \i \ nominals p\ shows \M, g, w \ p\ proof - from assms(1) have \M, g, g i \ p\ for g using soundness by fastforce then have \M, g(i := w), (g(i := w)) i \ p\ by blast then have \M, g(i := w), w \ p\ by simp then have \M, g(i := g i), w \ p\ using assms(2) semantics_fresh by metis then show ?thesis by simp qed section \No Detours\ text \ We only need to spend initial coins when we apply GoTo twice in a row. Otherwise another rule will have been applied in-between that justifies the GoTo. Therefore by avoiding detours we can close any closeable branch starting from 1 coin. \ primrec nonempty :: \('a, 'b) block \ bool\ where \nonempty (ps, i) = (ps \ [])\ lemma nonempty_Suc: assumes \n \ (ps, a) # filter nonempty left @ right\ \q at a in (ps, a) # filter nonempty left @ right\ \q \ Nom a\ shows \Suc n \ filter nonempty ((ps, a) # left) @ right\ proof (cases ps) case Nil then have \a \ branch_nominals (filter nonempty left @ right)\ unfolding branch_nominals_def using assms(2-3) by fastforce then show ?thesis using assms(1) Nil GoTo by auto next case Cons then show ?thesis using assms(1) ST_Suc by auto qed lemma ST_nonempty: \n \ left @ right \ Suc m \ filter nonempty left @ right\ proof (induct n \left @ right\ arbitrary: left right rule: ST.induct) case (Close p i n) have \(\<^bold>\ p) at i in filter nonempty left @ right\ using Close(2) by fastforce moreover from this have \p at i in filter nonempty left @ right\ using Close(1) by fastforce ultimately show ?case using ST.Close by fast next case (Neg p a ps branch n) then show ?case proof (cases left) case Nil then have \Suc m \ (p # ps, a) # branch\ using Neg(4) by fastforce then have \m \ (ps, a) # branch\ using Neg(1-2) ST.Neg by fast then show ?thesis using Nil Neg(5) ST_Suc by auto next case (Cons _ left') then have \Suc m \ (p # ps, a) # filter nonempty left' @ right\ using Neg(4)[where left=\_ # left'\] Neg(5) by fastforce moreover have *: \(\<^bold>\ \<^bold>\ p) at a in (ps, a) # filter nonempty left' @ right\ using Cons Neg(1, 5) by fastforce moreover have \new p a ((ps, a) # filter nonempty left' @ right)\ using Cons Neg(2, 5) unfolding new_def by auto ultimately have \m \ (ps, a) # filter nonempty left' @ right\ using ST.Neg by fast then have \Suc m \ filter nonempty ((ps, a) # left') @ right\ using * nonempty_Suc by fast then show ?thesis using Cons Neg(5) by auto qed next case (DisP p q a ps branch n) then show ?case proof (cases left) case Nil then have \Suc m \ (p # ps, a) # branch\ \Suc m \ (q # ps, a) # branch\ using DisP(5, 7) by fastforce+ then have \m \ (ps, a) # branch\ using DisP(1-3) ST.DisP by fast then show ?thesis using Nil DisP(8) ST_Suc by auto next case (Cons _ left') then have \Suc m \ (p # ps, a) # filter nonempty left' @ right\ \Suc m \ (q # ps, a) # filter nonempty left' @ right\ using DisP(5, 7)[where left=\_ # left'\] DisP(8) by fastforce+ moreover have *: \(p \<^bold>\ q) at a in (ps, a) # filter nonempty left' @ right\ using Cons DisP(1, 8) by fastforce moreover have \new p a ((ps, a) # filter nonempty left' @ right)\ \new q a ((ps, a) # filter nonempty left' @ right)\ using Cons DisP(2-3, 8) unfolding new_def by auto ultimately have \m \ (ps, a) # filter nonempty left' @ right\ using ST.DisP by fast then have \Suc m \ filter nonempty ((ps, a) # left') @ right\ using * nonempty_Suc by fast then show ?thesis using Cons DisP(8) by auto qed next case (DisN p q a ps branch n) then show ?case proof (cases left) case Nil then have \Suc m \ ((\<^bold>\ q) # (\<^bold>\ p) # ps, a) # branch\ using DisN(4) by fastforce then have \m \ (ps, a) # branch\ using DisN(1-2) ST.DisN by fast then show ?thesis using Nil DisN(5) ST_Suc by auto next case (Cons _ left') then have \Suc m \ ((\<^bold>\ q) # (\<^bold>\ p) # ps, a) # filter nonempty left' @ right\ using DisN(4)[where left=\_ # left'\] DisN(5) by fastforce moreover have *: \(\<^bold>\ (p \<^bold>\ q)) at a in (ps, a) # filter nonempty left' @ right\ using Cons DisN(1, 5) by fastforce moreover consider \new (\<^bold>\ p) a ((ps, a) # filter nonempty left' @ right)\ | \new (\<^bold>\ q) a ((ps, a) # filter nonempty left' @ right)\ using Cons DisN(2, 5) unfolding new_def by auto ultimately have \m \ (ps, a) # filter nonempty left' @ right\ using ST.DisN by metis then have \Suc m \ filter nonempty ((ps, a) # left') @ right\ using * nonempty_Suc by fast then show ?thesis using Cons DisN(5) by auto qed next case (DiaP p a ps branch i n) then show ?case proof (cases left) case Nil then have \Suc m \ ((\<^bold>@ i p) # (\<^bold>\ Nom i) # ps, a) # branch\ using DiaP(6) by fastforce then have \m \ (ps, a) # branch\ using DiaP(1-4) ST.DiaP by fast then show ?thesis using Nil DiaP(7) ST_Suc by auto next case (Cons _ left') then have \Suc m \ ((\<^bold>@ i p) # (\<^bold>\ Nom i) # ps, a) # filter nonempty left' @ right\ using DiaP(6)[where left=\_ # left'\] DiaP(7) by fastforce moreover have *: \(\<^bold>\ p) at a in (ps, a) # filter nonempty left' @ right\ using Cons DiaP(1, 7) by fastforce moreover have \i \ branch_nominals ((ps, a) # filter nonempty left' @ right)\ using Cons DiaP(2, 7) unfolding branch_nominals_def by auto moreover have \\ witnessed p a ((ps, a) # filter nonempty left' @ right)\ using Cons DiaP(4, 7) unfolding witnessed_def by auto ultimately have \m \ (ps, a) # filter nonempty left' @ right\ using DiaP(3) ST.DiaP by fast then have \Suc m \ filter nonempty ((ps, a) # left') @ right\ using * nonempty_Suc by fast then show ?thesis using Cons DiaP(7) by auto qed next case (DiaN p a ps branch i n) then show ?case proof (cases left) case Nil then have \Suc m \ ((\<^bold>\ (\<^bold>@ i p)) # ps, a) # branch\ using DiaN(5) by fastforce then have \m \ (ps, a) # branch\ using DiaN(1-3) ST.DiaN by fast then show ?thesis using Nil DiaN(6) ST_Suc by auto next case (Cons _ left') then have \Suc m \ ((\<^bold>\ (\<^bold>@ i p)) # ps, a) # filter nonempty left' @ right\ using DiaN(5)[where left=\_ # left'\] DiaN(6) by fastforce moreover have *: \(\<^bold>\ (\<^bold>\ p)) at a in (ps, a) # filter nonempty left' @ right\ using Cons DiaN(1, 6) by fastforce moreover have *: \(\<^bold>\ Nom i) at a in (ps, a) # filter nonempty left' @ right\ using Cons DiaN(2, 6) by fastforce moreover have \new (\<^bold>\ (\<^bold>@ i p)) a ((ps, a) # filter nonempty left' @ right)\ using Cons DiaN(3, 6) unfolding new_def by auto ultimately have \m \ (ps, a) # filter nonempty left' @ right\ using ST.DiaN by fast then have \Suc m \ filter nonempty ((ps, a) # left') @ right\ using * nonempty_Suc by fast then show ?thesis using Cons DiaN(6) by auto qed next case (SatP a p b ps branch n) then show ?case proof (cases left) case Nil then have \Suc m \ (p # ps, a) # branch\ using SatP(4) by fastforce then have \m \ (ps, a) # branch\ using SatP(1-2) ST.SatP by fast then show ?thesis using Nil SatP(5) ST_Suc by auto next case (Cons _ left') then have \Suc m \ (p # ps, a) # filter nonempty left' @ right\ using SatP(4)[where left=\_ # left'\] SatP(5) by fastforce moreover have \(\<^bold>@ a p) at b in (ps, a) # filter nonempty left' @ right\ using Cons SatP(1, 5) by fastforce moreover have \new p a ((ps, a) # filter nonempty left' @ right)\ using Cons SatP(2, 5) unfolding new_def by auto ultimately have *: \m \ (ps, a) # filter nonempty left' @ right\ using ST.SatP by fast then have \Suc m \ filter nonempty ((ps, a) # left') @ right\ proof (cases ps) case Nil then have \a \ branch_nominals (filter nonempty left' @ right)\ unfolding branch_nominals_def using SatP(1, 5) Cons by fastforce then show ?thesis using * Nil GoTo by fastforce next case Cons then show ?thesis using * ST_Suc by auto qed then show ?thesis using Cons SatP(5) by auto qed next case (SatN a p b ps branch n) then show ?case proof (cases left) case Nil then have \Suc m \ ((\<^bold>\ p) # ps, a) # branch\ using SatN(4) by fastforce then have \m \ (ps, a) # branch\ using SatN(1-2) ST.SatN by fast then show ?thesis using Nil SatN(5) ST_Suc by auto next case (Cons _ left') then have \Suc m \ ((\<^bold>\ p) # ps, a) # filter nonempty left' @ right\ using SatN(4)[where left=\_ # left'\] SatN(5) by fastforce moreover have \(\<^bold>\ (\<^bold>@ a p)) at b in (ps, a) # filter nonempty left' @ right\ using Cons SatN(1, 5) by fastforce moreover have \new (\<^bold>\ p) a ((ps, a) # filter nonempty left' @ right)\ using Cons SatN(2, 5) unfolding new_def by auto ultimately have *: \m \ (ps, a) # filter nonempty left' @ right\ using ST.SatN by fast then have \Suc m \ filter nonempty ((ps, a) # left') @ right\ proof (cases ps) case Nil then have \a \ branch_nominals (filter nonempty left' @ right)\ unfolding branch_nominals_def using SatN(1, 5) Cons by fastforce then show ?thesis using * Nil GoTo by fastforce next case Cons then show ?thesis using * ST_Suc by auto qed then show ?thesis using Cons SatN(5) by auto qed next case (GoTo i n) show ?case using GoTo(3)[where left=\([], i) # left\] by simp next case (Nom p b ps a branch i n) then show ?case proof (cases left) case Nil then have \Suc m \ (p # ps, a) # branch\ using Nom(6) by fastforce then have \m \ (ps, a) # branch\ using Nom(1-4) ST.Nom by fast then show ?thesis using Nil Nom(7) ST_Suc by auto next case (Cons _ left') then have \Suc m \ (p # ps, a) # filter nonempty left' @ right\ using Nom(6)[where left=\_ # left'\] Nom(7) by fastforce moreover have \p at b in (ps, a) # filter nonempty left' @ right\ \Nom i at b in (ps, a) # filter nonempty left' @ right\ \Nom i at a in (ps, a) # filter nonempty left' @ right\ using Cons Nom(1-4, 7) unfolding new_def by simp_all (metis empty_iff empty_set)+ moreover have \new p a ((ps, a) # filter nonempty left' @ right)\ using Cons Nom(4, 7) unfolding new_def by auto ultimately have *: \m \ (ps, a) # filter nonempty left' @ right\ using ST.Nom by fast then have \Suc m \ filter nonempty ((ps, a) # left') @ right\ proof (cases ps) case Nil moreover have \a \ b\ using Nom(1, 4) unfolding new_def by blast ultimately have \a \ branch_nominals (filter nonempty left' @ right)\ proof (cases \i = a\) case True then have \Nom a at b in left' @ right\ using \a \ b\ Nom(2, 7) Cons by auto then show ?thesis unfolding branch_nominals_def using \a \ b\ by fastforce next case False then have \Nom i at a in left' @ right\ using Nom(3, 7) Nil Cons by auto then show ?thesis unfolding branch_nominals_def using False by fastforce qed then show ?thesis using * Nil GoTo by auto next case Cons then show ?thesis using * ST_Suc by auto qed then show ?thesis using Cons Nom(7) by auto qed qed theorem ST_coins: \n \ branch \ Suc m \ branch\ using ST_nonempty[where left=\[]\] by auto corollary ST_one: \n \ branch \ 1 \ branch\ using ST_coins by auto subsection \Free GoTo\ text \The above result allows us to derive a version of GoTo that works "for free."\ lemma GoTo': assumes \Suc n \ ([], i) # branch\ \i \ branch_nominals branch\ shows \Suc n \ branch\ using assms GoTo ST_coins by fast section \Indexed Mapping\ text \This section contains some machinery for deriving rules.\ subsection \Indexing\ text \ We use pairs of natural numbers to index into the branch. The first component specifies the block and the second specifies the formula in that block. We index from the back to ensure that indices are stable under the addition of new formulas and blocks. \ primrec rev_nth :: \'a list \ nat \ 'a option\ (infixl \!.\ 100) where \[] !. v = None\ | \(x # xs) !. v = (if length xs = v then Some x else xs !. v)\ lemma rev_nth_last: \xs !. 0 = Some x \ last xs = x\ by (induct xs) auto lemma rev_nth_zero: \(xs @ [x]) !. 0 = Some x\ by (induct xs) auto lemma rev_nth_snoc: \(xs @ [x]) !. Suc v = Some y \ xs !. v = Some y\ by (induct xs) auto lemma rev_nth_Suc: \(xs @ [x]) !. Suc v = xs !. v\ by (induct xs) auto lemma rev_nth_bounded: \v < length xs \ \x. xs !. v = Some x\ by (induct xs) simp_all lemma rev_nth_Cons: \xs !. v = Some y \ (x # xs) !. v = Some y\ proof (induct xs arbitrary: v rule: rev_induct) case (snoc a xs) then show ?case proof (induct v) case (Suc v) then have \xs !. v = Some y\ using rev_nth_snoc by fast then have \(x # xs) !. v = Some y\ using Suc(2) by blast then show ?case using Suc(3) by auto qed simp qed simp lemma rev_nth_append: \xs !. v = Some y \ (ys @ xs) !. v = Some y\ using rev_nth_Cons[where xs=\_ @ xs\] by (induct ys) simp_all lemma rev_nth_mem: \block \. branch \ (\v. branch !. v = Some block)\ proof assume \block \. branch\ then show \\v. branch !. v = Some block\ proof (induct branch) case (Cons block' branch) then show ?case proof (cases \block = block'\) case False then have \\v. branch !. v = Some block\ using Cons by simp then show ?thesis using rev_nth_Cons by fast qed auto qed simp next assume \\v. branch !. v = Some block\ then show \block \. branch\ proof (induct branch) case (Cons block' branch) then show ?case by simp (metis option.sel) qed simp qed lemma rev_nth_on: \p on (ps, i) \ (\v. ps !. v = Some p) \ p = Nom i\ by (simp add: rev_nth_mem) lemma rev_nth_Some: \xs !. v = Some y \ v < length xs\ proof (induct xs arbitrary: v rule: rev_induct) case (snoc x xs) then show ?case by (induct v) (simp_all, metis rev_nth_snoc) qed simp lemma index_Cons: assumes \((ps, a) # branch) !. v = Some (qs, b)\ \qs !. v' = Some q\ shows \\qs'. ((p # ps, a) # branch) !. v = Some (qs', b) \ qs' !. v' = Some q\ proof - have \((p # ps, a) # branch) !. v = Some (qs, b) \ ((p # ps, a) # branch) !. v = Some (p # qs, b)\ using assms(1) by auto moreover have \qs !. v' = Some q\ \(p # qs) !. v' = Some q\ using assms(2) rev_nth_Cons by fast+ ultimately show ?thesis by fastforce qed subsection \Mapping\ primrec mapi :: \(nat \ 'a \ 'b) \ 'a list \ 'b list\ where \mapi f [] = []\ | \mapi f (x # xs) = f (length xs) x # mapi f xs\ primrec mapi_block :: \(nat \ ('a, 'b) fm \ ('a, 'b) fm) \ (('a, 'b) block \ ('a, 'b) block)\ where \mapi_block f (ps, i) = (mapi f ps, i)\ definition mapi_branch :: \(nat \ nat \ ('a, 'b) fm \ ('a, 'b) fm) \ (('a, 'b) branch \ ('a, 'b) branch)\ where \mapi_branch f branch \ mapi (\v. mapi_block (f v)) branch\ abbreviation mapper :: \(('a, 'b) fm \ ('a, 'b) fm) \ (nat \ nat) set \ nat \ nat \ ('a, 'b) fm \ ('a, 'b) fm\ where \mapper f xs v v' p \ (if (v, v') \ xs then f p else p)\ lemma mapi_block_add_oob: assumes \length ps \ v'\ shows \mapi_block (mapper f ({(v, v')} \ xs) v) (ps, i) = mapi_block (mapper f xs v) (ps, i)\ using assms by (induct ps) simp_all lemma mapi_branch_add_oob: assumes \length branch \ v\ shows \mapi_branch (mapper f ({(v, v')} \ xs)) branch = mapi_branch (mapper f xs) branch\ unfolding mapi_branch_def using assms by (induct branch) simp_all lemma mapi_branch_head_add_oob: \mapi_branch (mapper f ({(length branch, length ps)} \ xs)) ((ps, a) # branch) = mapi_branch (mapper f xs) ((ps, a) # branch)\ using mapi_branch_add_oob[where branch=branch] unfolding mapi_branch_def using mapi_block_add_oob[where ps=ps] by simp lemma mapi_branch_mem: assumes \(ps, i) \. branch\ shows \\v. (mapi (f v) ps, i) \. mapi_branch f branch\ unfolding mapi_branch_def using assms by (induct branch) auto lemma rev_nth_mapi_branch: assumes \branch !. v = Some (ps, a)\ shows \(mapi (f v) ps, a) \. mapi_branch f branch\ unfolding mapi_branch_def using assms by (induct branch) (simp_all, metis mapi_block.simps option.inject) lemma rev_nth_mapi_block: assumes \ps !. v' = Some p\ shows \f v' p on (mapi f ps, a)\ using assms by (induct ps) (simp_all, metis option.sel) lemma mapi_append: \mapi f (xs @ ys) = mapi (\v. f (v + length ys)) xs @ mapi f ys\ by (induct xs) simp_all lemma mapi_block_id: \mapi_block (mapper f {} v) (ps, i) = (ps, i)\ by (induct ps) auto lemma mapi_branch_id: \mapi_branch (mapper f {}) branch = branch\ unfolding mapi_branch_def using mapi_block_id by (induct branch) auto lemma length_mapi: \length (mapi f xs) = length xs\ by (induct xs) auto lemma mapi_rev_nth: assumes \xs !. v = Some x\ shows \mapi f xs !. v = Some (f v x)\ using assms proof (induct xs arbitrary: v) case (Cons y xs) have *: \mapi f (y # xs) = f (length xs) y # mapi f xs\ by simp show ?case proof (cases \v = length xs\) case True then have \mapi f (y # xs) !. v = Some (f (length xs) y)\ using length_mapi * by (metis rev_nth.simps(2)) then show ?thesis using Cons.prems True by auto next case False then show ?thesis using * Cons length_mapi by (metis rev_nth.simps(2)) qed qed simp section \Duplicate Formulas\ subsection \Removable indices\ abbreviation \proj \ Equiv_Relations.proj\ definition all_is :: \('a, 'b) fm \ ('a, 'b) fm list \ nat set \ bool\ where \all_is p ps xs \ \v \ xs. ps !. v = Some p\ definition is_at :: \('a, 'b) fm \ 'b \ ('a, 'b) branch \ nat \ nat \ bool\ where \is_at p i branch v v' \ \ps. branch !. v = Some (ps, i) \ ps !. v' = Some p\ text \This definition is slightly complicated by the inability to index the opening nominal.\ definition is_elsewhere :: \('a, 'b) fm \ 'b \ ('a, 'b) branch \ (nat \ nat) set \ bool\ where \is_elsewhere p i branch xs \ \w w' ps. (w, w') \ xs \ branch !. w = Some (ps, i) \ (p = Nom i \ ps !. w' = Some p)\ definition Dup :: \('a, 'b) fm \ 'b \ ('a, 'b) branch \ (nat \ nat) set \ bool\ where \Dup p i branch xs \ \(v, v') \ xs. is_at p i branch v v' \ is_elsewhere p i branch xs\ lemma Dup_all_is: assumes \Dup p i branch xs\ \branch !. v = Some (ps, a)\ shows \all_is p ps (proj xs v)\ using assms unfolding Dup_def is_at_def all_is_def proj_def by auto lemma Dup_branch: \Dup p i branch xs \ Dup p i (extra @ branch) xs\ unfolding Dup_def is_at_def is_elsewhere_def using rev_nth_append by fast lemma Dup_block: assumes \Dup p i ((ps, a) # branch) xs\ shows \Dup p i ((ps' @ ps, a) # branch) xs\ unfolding Dup_def proof safe fix v v' assume \(v, v') \ xs\ then show \is_at p i ((ps' @ ps, a) # branch) v v'\ using assms rev_nth_append unfolding Dup_def is_at_def by fastforce next fix v v' assume \(v, v') \ xs\ then obtain w w' qs where \(w, w') \ xs\ \((ps, a) # branch) !. w = Some (qs, i)\ \p = Nom i \ qs !. w' = Some p\ using assms unfolding Dup_def is_elsewhere_def by blast then have \\qs. ((ps' @ ps, a) # branch) !. w = Some (qs, i) \ (p = Nom i \ qs !. w' = Some p)\ using rev_nth_append by fastforce then show \is_elsewhere p i ((ps' @ ps, a) # branch) xs\ unfolding is_elsewhere_def using \(w, w') \ xs\ by blast qed lemma Dup_opening: assumes \Dup p i branch xs\ \(v, v') \ xs\ \branch !. v = Some (ps, a)\ shows \i = a\ using assms unfolding Dup_def is_at_def by auto lemma Dup_head: \Dup p i ((ps, a) # branch) xs \ Dup p i ((q # ps, a) # branch) xs\ using Dup_block[where ps'=\[_]\] by simp lemma Dup_head_oob': assumes \Dup p i ((ps, a) # branch) xs\ shows \(length branch, k + length ps) \ xs\ using assms rev_nth_Some unfolding Dup_def is_at_def by fastforce lemma Dup_head_oob: assumes \Dup p i ((ps, a) # branch) xs\ shows \(length branch, length ps) \ xs\ using assms Dup_head_oob'[where k=0] by fastforce subsection \Omitting formulas\ primrec omit :: \nat set \ ('a, 'b) fm list \ ('a, 'b) fm list\ where \omit xs [] = []\ | \omit xs (p # ps) = (if length ps \ xs then omit xs ps else p # omit xs ps)\ primrec omit_block :: \nat set \ ('a, 'b) block \ ('a, 'b) block\ where \omit_block xs (ps, a) = (omit xs ps, a)\ definition omit_branch :: \(nat \ nat) set \ ('a, 'b) branch \ ('a, 'b) branch\ where \omit_branch xs branch \ mapi (\v. omit_block (proj xs v)) branch\ lemma omit_mem: \ps !. v = Some p \ v \ xs \ p \. omit xs ps\ proof (induct ps) case (Cons q ps) then show ?case by (cases \v = length ps\) simp_all qed simp lemma omit_id: \omit {} ps = ps\ by (induct ps) auto lemma omit_block_id: \omit_block {} block = block\ using omit_id by (cases block) simp lemma omit_branch_id: \omit_branch {} branch = branch\ unfolding omit_branch_def proj_def using omit_block_id by (induct branch) fastforce+ lemma omit_branch_mem_diff_opening: assumes \Dup p i branch xs\ \(ps, a) \. branch\ \i \ a\ shows \(ps, a) \. omit_branch xs branch\ proof - obtain v where v: \branch !. v = Some (ps, a)\ using assms(2) rev_nth_mem by fast then have \omit_branch xs branch !. v = Some (omit (proj xs v) ps, a)\ unfolding omit_branch_def by (simp add: mapi_rev_nth) then have *: \(omit (proj xs v) ps, a) \. omit_branch xs branch\ using rev_nth_mem by fast moreover have \proj xs v = {}\ unfolding proj_def using assms(1, 3) v Dup_opening by fast then have \omit (proj xs v) ps = ps\ using omit_id by auto ultimately show ?thesis by simp qed lemma omit_branch_mem_same_opening: assumes \Dup p i branch xs\ \p at i in branch\ shows \p at i in omit_branch xs branch\ proof - obtain ps where ps: \(ps, i) \. branch\ \p on (ps, i)\ using assms(2) by blast then obtain v where v: \branch !. v = Some (ps, i)\ using rev_nth_mem by fast then have \omit_branch xs branch !. v = Some (omit (proj xs v) ps, i)\ unfolding omit_branch_def by (simp add: mapi_rev_nth) then have *: \(omit (proj xs v) ps, i) \. omit_branch xs branch\ using rev_nth_mem by fast consider v' where \ps !. v' = Some p\ \(v, v') \ xs\ | v' where \ps !. v' = Some p\ \(v, v') \ xs\ | \p = Nom i\ using ps v rev_nth_mem by fastforce then show ?thesis proof cases case (1 v') then obtain qs w w' where qs: \(w, w') \ xs\ \branch !. w = Some (qs, i)\ \p = Nom i \ qs !. w' = Some p\ using assms(1) unfolding Dup_def is_elsewhere_def by blast then have \omit_branch xs branch !. w = Some (omit (proj xs w) qs, i)\ unfolding omit_branch_def by (simp add: mapi_rev_nth) then have \(omit (proj xs w) qs, i) \. omit_branch xs branch\ using rev_nth_mem by fast moreover have \p on (omit (proj xs w) qs, i)\ unfolding proj_def using qs(1, 3) omit_mem by fastforce ultimately show ?thesis by blast next case (2 v') then show ?thesis using * omit_mem unfolding proj_def by (metis Image_singleton_iff on.simps) next case 3 then show ?thesis using * by auto qed qed lemma omit_del: assumes \p \. ps\ \p \ set (omit xs ps)\ shows \\v. ps !. v = Some p \ v \ xs\ using assms omit_mem rev_nth_mem by metis lemma omit_all_is: assumes \all_is p ps xs\ \q \. ps\ \q \ set (omit xs ps)\ shows \q = p\ using assms omit_del unfolding all_is_def by fastforce lemma omit_branch_mem_diff_formula: assumes \Dup p i branch xs\ \q at i in branch\ \p \ q\ shows \q at i in omit_branch xs branch\ proof - obtain ps where ps: \(ps, i) \. branch\ \q on (ps, i)\ using assms(2) by blast then obtain v where v: \branch !. v = Some (ps, i)\ using rev_nth_mem by fast then have \omit_branch xs branch !. v = Some (omit (proj xs v) ps, i)\ unfolding omit_branch_def by (simp add: mapi_rev_nth) then have *: \(omit (proj xs v) ps, i) \. omit_branch xs branch\ using rev_nth_mem by fast moreover have \all_is p ps (proj xs v)\ using assms(1) Dup_all_is v by fast then have \q on (omit (proj xs v) ps, i)\ using ps assms(3) omit_all_is by auto ultimately show ?thesis by blast qed lemma omit_branch_mem: assumes \Dup p i branch xs\ \q at a in branch\ shows \q at a in omit_branch xs branch\ using assms omit_branch_mem_diff_opening omit_branch_mem_same_opening omit_branch_mem_diff_formula by fast lemma omit_set: \set (omit xs ps) \ set ps\ by (induct ps) auto lemma on_omit: \p on (omit xs ps, i) \ p on (ps, i)\ using omit_set by auto lemma Dup_set: assumes \all_is p ps xs\ shows \{p} \ set (omit xs ps) = {p} \ set ps\ using assms omit_all_is omit_set unfolding all_is_def by fast lemma Dup_list_nominals: assumes \all_is p ps xs\ shows \nominals p \ list_nominals (omit xs ps) = nominals p \ list_nominals ps\ using assms Dup_set by fastforce lemma Dup_block_nominals: assumes \all_is p ps xs\ shows \nominals p \ block_nominals (omit xs ps, i) = nominals p \ block_nominals (ps, i)\ using assms by (simp add: Dup_list_nominals) lemma Dup_branch_nominals': assumes \Dup p i branch xs\ shows \nominals p \ branch_nominals (omit_branch xs branch) = nominals p \ branch_nominals branch\ proof - have \\(v, v') \ xs. v < length branch \ is_at p i branch v v'\ using assms unfolding Dup_def by auto then show ?thesis proof (induct branch) case Nil then show ?case unfolding omit_branch_def by simp next case (Cons block branch) then show ?case proof (cases block) case (Pair ps a) have \\(v, v') \ xs. v < length branch \ is_at p i branch v v'\ using Cons(2) rev_nth_Cons unfolding is_at_def by auto then have \nominals p \ branch_nominals (omit_branch xs branch) = nominals p \ branch_nominals branch\ using Cons(1) by blast then have \nominals p \ branch_nominals (omit_branch xs ((ps, a) # branch)) = nominals p \ block_nominals (omit (proj xs (length branch)) ps, a) \ branch_nominals branch\ unfolding branch_nominals_def omit_branch_def by auto moreover have \all_is p ps (proj xs (length branch))\ using Cons(2) Pair unfolding proj_def all_is_def is_at_def by auto then have \nominals p \ block_nominals (omit (proj xs (length branch)) ps, a) = nominals p \ block_nominals (ps, a)\ using Dup_block_nominals by fast then have \nominals p \ block_nominals (omit_block (proj xs (length branch)) (ps, a)) = nominals p \ block_nominals (ps, a)\ by simp ultimately have \nominals p \ branch_nominals (omit_branch xs ((ps, a) # branch)) = nominals p \ block_nominals (ps, a) \ branch_nominals branch\ by auto then show ?thesis unfolding branch_nominals_def using Pair by auto qed qed qed lemma Dup_branch_nominals: assumes \Dup p i branch xs\ shows \branch_nominals (omit_branch xs branch) = branch_nominals branch\ proof (cases \xs = {}\) case True then show ?thesis using omit_branch_id by metis next case False with assms obtain ps w w' where \(w, w') \ xs\ \branch !. w = Some (ps, i)\ \p = Nom i \ ps !. w' = Some p\ unfolding Dup_def is_elsewhere_def by fast then have *: \(ps, i) \. branch\ \p on (ps, i)\ using rev_nth_mem rev_nth_on by fast+ then have \nominals p \ branch_nominals branch\ unfolding branch_nominals_def using block_nominals by fast moreover obtain ps' where \(ps', i) \. omit_branch xs branch\ \p on (ps', i)\ using assms * omit_branch_mem by fast then have \nominals p \ branch_nominals (omit_branch xs branch)\ unfolding branch_nominals_def using block_nominals by fast moreover have \nominals p \ branch_nominals (omit_branch xs branch) = nominals p \ branch_nominals branch\ using assms Dup_branch_nominals' by fast ultimately show ?thesis by blast qed lemma omit_branch_mem_dual: assumes \p at i in omit_branch xs branch\ shows \p at i in branch\ proof - obtain ps where ps: \(ps, i) \. omit_branch xs branch\ \p on (ps, i)\ using assms(1) by blast then obtain v where v: \omit_branch xs branch !. v = Some (ps, i)\ using rev_nth_mem unfolding omit_branch_def by fast then have \v < length (omit_branch xs branch)\ using rev_nth_Some by fast then have \v < length branch\ unfolding omit_branch_def using length_mapi by metis then obtain ps' i' where ps': \branch !. v = Some (ps', i')\ using rev_nth_bounded by (metis surj_pair) then have \omit_branch xs branch !. v = Some (omit (proj xs v) ps', i')\ unfolding omit_branch_def by (simp add: mapi_rev_nth) then have \ps = omit (proj xs v) ps'\ \i = i'\ using v by simp_all then have \p on (ps', i)\ using ps omit_set by auto moreover have \(ps', i) \. branch\ using ps' \i = i'\ rev_nth_mem by fast ultimately show ?thesis using \ps = omit (proj xs v) ps'\ by blast qed lemma witnessed_omit_branch: assumes \witnessed p a (omit_branch xs branch)\ shows \witnessed p a branch\ proof - obtain ps qs i where ps: \(ps, a) \. omit_branch xs branch\ \(\<^bold>@ i p) on (ps, a)\ and qs: \(qs, a) \. omit_branch xs branch\ \(\<^bold>\ Nom i) on (qs, a)\ using assms unfolding witnessed_def by blast from ps obtain ps' where \(ps', a) \. branch\ \(\<^bold>@ i p) on (ps', a)\ using omit_branch_mem_dual by fast moreover from qs obtain qs' where \(qs', a) \. branch\ \(\<^bold>\ Nom i) on (qs', a)\ using omit_branch_mem_dual by fast ultimately show ?thesis unfolding witnessed_def by blast qed lemma new_omit_branch: assumes \new p a branch\ shows \new p a (omit_branch xs branch)\ using assms omit_branch_mem_dual unfolding new_def by fast lemma omit_oob: assumes \length ps \ v\ shows \omit ({v} \ xs) ps = omit xs ps\ using assms by (induct ps) simp_all lemma omit_branch_oob: assumes \length branch \ v\ shows \omit_branch ({(v, v')} \ xs) branch = omit_branch xs branch\ using assms proof (induct branch) case Nil then show ?case unfolding omit_branch_def by simp next case (Cons block branch) let ?xs = \({(v, v')} \ xs)\ show ?case proof (cases block) case (Pair ps a) then have \omit_branch ?xs ((ps, a) # branch) = (omit (proj ?xs (length branch)) ps, a) # omit_branch xs branch\ using Cons unfolding omit_branch_def by simp moreover have \proj ?xs (length branch) = proj xs (length branch)\ using Cons(2) unfolding proj_def by auto ultimately show ?thesis unfolding omit_branch_def by simp qed qed subsection \Induction\ lemma ST_Dup: assumes \n \ branch\ \Dup q i branch xs\ shows \n \ omit_branch xs branch\ using assms proof (induct n branch) case (Close p i' branch n) have \p at i' in omit_branch xs branch\ using Close(1, 3) omit_branch_mem by fast moreover have \(\<^bold>\ p) at i' in omit_branch xs branch\ using Close(2, 3) omit_branch_mem by fast ultimately show ?case using ST.Close by fast next case (Neg p a ps branch n) have \Suc n \ omit_branch xs ((p # ps, a) # branch)\ using Neg(4-) Dup_head by fast moreover have \(length branch, length ps) \ xs\ using Neg(5) Dup_head_oob by fast ultimately have \Suc n \ (p # omit (proj xs (length branch)) ps, a) # omit_branch xs branch\ unfolding omit_branch_def proj_def by simp moreover have \(\<^bold>\ \<^bold>\ p) at a in omit_branch xs ((ps, a) # branch)\ using Neg(1, 5) omit_branch_mem by fast moreover have \new p a (omit_branch xs ((ps, a) # branch))\ using Neg(2) new_omit_branch by fast ultimately show ?case by (simp add: omit_branch_def ST.Neg) next case (DisP p q a ps branch n) have \Suc n \ omit_branch xs ((p # ps, a) # branch)\ \Suc n \ omit_branch xs ((q # ps, a) # branch)\ using DisP(4-) Dup_head by fast+ moreover have \(length branch, length ps) \ xs\ using DisP(8) Dup_head_oob by fast ultimately have \Suc n \ (p # omit (proj xs (length branch)) ps, a) # omit_branch xs branch\ \Suc n \ (q # omit (proj xs (length branch)) ps, a) # omit_branch xs branch\ unfolding omit_branch_def proj_def by simp_all moreover have \(p \<^bold>\ q) at a in omit_branch xs ((ps, a) # branch)\ using DisP(1, 8) omit_branch_mem by fast moreover have \new p a (omit_branch xs ((ps, a) # branch))\ using DisP(2) new_omit_branch by fast moreover have \new q a (omit_branch xs ((ps, a) # branch))\ using DisP(3) new_omit_branch by fast ultimately show ?case by (simp add: omit_branch_def ST.DisP) next case (DisN p q a ps branch n) have \Suc n \ omit_branch xs (((\<^bold>\ q) # (\<^bold>\ p) # ps, a) # branch)\ using DisN(4-) Dup_block[where ps'=\[_, _]\] by fastforce moreover have \(length branch, length ps) \ xs\ using DisN(5) Dup_head_oob by fast moreover have \(length branch, 1 + length ps) \ xs\ using DisN(5) Dup_head_oob' by fast ultimately have \Suc n \ ((\<^bold>\ q) # (\<^bold>\ p) # omit (proj xs (length branch)) ps, a) # omit_branch xs branch\ unfolding omit_branch_def proj_def by simp moreover have \(\<^bold>\ (p \<^bold>\ q)) at a in omit_branch xs ((ps, a) # branch)\ using DisN(1, 5) omit_branch_mem by fast moreover have \new (\<^bold>\ p) a (omit_branch xs ((ps, a) # branch)) \ new (\<^bold>\ q) a (omit_branch xs ((ps, a) # branch))\ using DisN(2) new_omit_branch by fast ultimately show ?case by (simp add: omit_branch_def ST.DisN) next case (DiaP p a ps branch i n) have \Suc n \ omit_branch xs (((\<^bold>@ i p) # (\<^bold>\ Nom i) # ps, a) # branch)\ using DiaP(4-) Dup_block[where ps'=\[_, _]\] by fastforce moreover have \(length branch, length ps) \ xs\ using DiaP(7) Dup_head_oob by fast moreover have \(length branch, 1+ length ps) \ xs\ using DiaP(7) Dup_head_oob' by fast ultimately have \Suc n \ ((\<^bold>@ i p) # (\<^bold>\ Nom i) # omit (proj xs (length branch)) ps, a) # omit_branch xs branch\ unfolding omit_branch_def proj_def by simp moreover have \(\<^bold>\ p) at a in omit_branch xs ((ps, a) # branch)\ using DiaP(1, 7) omit_branch_mem by fast moreover have \i \ branch_nominals (omit_branch xs ((ps, a) # branch))\ using DiaP(2, 7) Dup_branch_nominals by fast moreover have \\ witnessed p a (omit_branch xs ((ps, a) # branch))\ using DiaP(4) witnessed_omit_branch by fast ultimately show ?case using DiaP(3) by (simp add: omit_branch_def ST.DiaP) next case (DiaN p a ps branch i n) have \Suc n \ omit_branch xs (((\<^bold>\ (\<^bold>@ i p)) # ps, a) # branch)\ using DiaN(4-) Dup_head by fast moreover have \(length branch, length ps) \ xs\ using DiaN(6) Dup_head_oob by fast ultimately have \Suc n \ ((\<^bold>\ (\<^bold>@ i p)) # omit (proj xs (length branch)) ps, a) # omit_branch xs branch\ unfolding omit_branch_def proj_def by simp moreover have \(\<^bold>\ (\<^bold>\ p)) at a in omit_branch xs ((ps, a) # branch)\ using DiaN(1, 6) omit_branch_mem by fast moreover have \(\<^bold>\ Nom i) at a in omit_branch xs ((ps, a) # branch)\ using DiaN(2, 6) omit_branch_mem by fast moreover have \new (\<^bold>\ (\<^bold>@ i p)) a (omit_branch xs ((ps, a) # branch))\ using DiaN(3) new_omit_branch by fast ultimately show ?case by (simp add: omit_branch_def ST.DiaN) next case (SatP a p b ps branch n) have \Suc n \ omit_branch xs ((p # ps, a) # branch)\ using SatP(4-) Dup_head by fast moreover have \(length branch, length ps) \ xs\ using SatP(5) Dup_head_oob by fast ultimately have \Suc n \ (p # omit (proj xs (length branch)) ps, a) # omit_branch xs branch\ unfolding omit_branch_def proj_def by simp moreover have \(\<^bold>@ a p) at b in omit_branch xs ((ps, a) # branch)\ using SatP(1, 5) omit_branch_mem by fast moreover have \new p a (omit_branch xs ((ps, a) # branch))\ using SatP(2) new_omit_branch by fast ultimately show ?case by (simp add: omit_branch_def ST.SatP) next case (SatN a p b ps branch n) have \Suc n \ omit_branch xs (((\<^bold>\ p) # ps, a) # branch)\ using SatN(4-) Dup_head by fast moreover have \(length branch, length ps) \ xs\ using SatN(5) Dup_head_oob by fast ultimately have \Suc n \ ((\<^bold>\ p) # omit (proj xs (length branch)) ps, a) # omit_branch xs branch\ unfolding omit_branch_def proj_def by simp moreover have \(\<^bold>\ (\<^bold>@ a p)) at b in omit_branch xs ((ps, a) # branch)\ using SatN(1, 5) omit_branch_mem by fast moreover have \new (\<^bold>\ p) a (omit_branch xs ((ps, a) # branch))\ using SatN(2) new_omit_branch by fast ultimately show ?case by (simp add: omit_branch_def ST.SatN) next case (GoTo i branch n) then have \n \ omit_branch xs (([], i) # branch)\ using Dup_branch[where extra=\[([], i)]\] by fastforce then have \n \ ([], i) # omit_branch xs branch\ unfolding omit_branch_def by simp moreover have \i \ branch_nominals (omit_branch xs branch)\ using GoTo(1, 4) Dup_branch_nominals by fast ultimately show ?case unfolding omit_branch_def by (simp add: ST.GoTo) next case (Nom p b ps a branch i' n) have \Suc n \ omit_branch xs ((p # ps, a) # branch)\ using Nom(4-) Dup_head by fast moreover have \(length branch, length ps) \ xs\ using Nom(7) Dup_head_oob by fast ultimately have \Suc n \ (p # omit (proj xs (length branch)) ps, a) # omit_branch xs branch\ unfolding omit_branch_def proj_def by simp moreover have \p at b in omit_branch xs ((ps, a) # branch)\ using Nom(1, 7) omit_branch_mem by fast moreover have \Nom i' at b in omit_branch xs ((ps, a) # branch)\ using Nom(2, 7) omit_branch_mem by fast moreover have \Nom i' at a in omit_branch xs ((ps, a) # branch)\ using Nom(3, 7) omit_branch_mem by fast moreover have \new p a (omit_branch xs ((ps, a) # branch))\ using Nom(4) new_omit_branch by fast ultimately show ?case by (simp add: omit_branch_def ST.Nom) qed theorem Dup: assumes \n \ (p # ps, a) # branch\ \\ new p a ((ps, a) # branch)\ shows \n \ (ps, a) # branch\ proof - obtain qs where qs: \(qs, a) \. (ps, a) # branch\ \p on (qs, a)\ using assms(2) unfolding new_def by blast let ?xs = \{(length branch, length ps)}\ have *: \is_at p a ((p # ps, a) # branch) (length branch) (length ps)\ unfolding is_at_def by simp have \Dup p a ((p # ps, a) # branch) ?xs\ proof (cases \p = Nom a\) case True moreover have \((p # ps, a) # branch) !. length branch = Some (p # ps, a)\ by simp moreover have \p on (p # ps, a)\ by simp ultimately have \is_elsewhere p a ((p # ps, a) # branch) ?xs\ unfolding is_elsewhere_def using assms(2) rev_nth_Some by (metis (mono_tags, lifting) Pair_inject less_le singletonD) then show ?thesis unfolding Dup_def using * by blast next case false: False then show ?thesis proof (cases \ps = qs\) case True then obtain w' where w': \qs !. w' = Some p\ using qs(2) false rev_nth_mem by fastforce then have \(p # ps) !. w' = Some p\ using True rev_nth_Cons by fast moreover have \((p # ps, a) # branch) !. length branch = Some (p # ps, a)\ by simp moreover have \(length branch, w') \ ?xs\ using True w' rev_nth_Some by fast ultimately have \is_elsewhere p a ((p # ps, a) # branch) ?xs\ unfolding is_elsewhere_def by fast then show ?thesis unfolding Dup_def using * by fast next case False then obtain w where w: \branch !. w = Some (qs, a)\ using qs(1) rev_nth_mem by fastforce moreover obtain w' where w': \qs !. w' = Some p\ using qs(2) false rev_nth_mem by fastforce moreover have \(w, w') \ ?xs\ using rev_nth_Some w by fast ultimately have \is_elsewhere p a ((p # ps, a) # branch) ?xs\ unfolding is_elsewhere_def using rev_nth_Cons by fast then show ?thesis unfolding Dup_def using * by fast qed qed then have \n \ omit_branch ?xs ((p # ps, a) # branch)\ using assms(1) ST_Dup by fast then have \n \ (omit (proj ?xs (length branch)) ps, a) # omit_branch ?xs branch\ unfolding omit_branch_def proj_def by simp moreover have \omit_branch ?xs branch = omit_branch {} branch\ using omit_branch_oob by auto then have \omit_branch ?xs branch = branch\ using omit_branch_id by simp moreover have \proj ?xs (length branch) = {length ps}\ unfolding proj_def by blast then have \omit (proj ?xs (length branch)) ps = omit {} ps\ using omit_oob by auto then have \omit (proj ?xs (length branch)) ps = ps\ using omit_id by simp ultimately show ?thesis by simp qed subsection \Unrestricted rules\ lemma ST_add: \n \ branch \ m + n \ branch\ using ST_Suc by (induct m) auto lemma ST_le: \n \ branch \ n \ m \ m \ branch\ using ST_add by (metis le_add_diff_inverse2) lemma Neg': assumes \(\<^bold>\ \<^bold>\ p) at a in (ps, a) # branch\ \n \ (p # ps, a) # branch\ shows \n \ (ps, a) # branch\ using assms Neg Dup ST_Suc by metis lemma DisP': assumes \(p \<^bold>\ q) at a in (ps, a) # branch\ \n \ (p # ps, a) # branch\ \n \ (q # ps, a) # branch\ shows \n \ (ps, a) # branch\ proof (cases \new p a ((ps, a) # branch) \ new q a ((ps, a) # branch)\) case True then show ?thesis using assms DisP ST_Suc by fast next case False then show ?thesis using assms Dup by fast qed lemma DisP'': assumes \(p \<^bold>\ q) at a in (ps, a) # branch\ \n \ (p # ps, a) # branch\ \m \ (q # ps, a) # branch\ shows \max n m \ (ps, a) # branch\ proof (cases \n \ m\) case True then have \m \ (p # ps, a) # branch\ using assms(2) ST_le by blast then show ?thesis using assms True by (simp add: DisP' max.absorb2) next case False then have \n \ (q # ps, a) # branch\ using assms(3) ST_le by fastforce then show ?thesis using assms False by (simp add: DisP' max.absorb1) qed lemma DisN': assumes \(\<^bold>\ (p \<^bold>\ q)) at a in (ps, a) # branch\ \n \ ((\<^bold>\ q) # (\<^bold>\ p) # ps, a) # branch\ shows \n \ (ps, a) # branch\ proof (cases \new (\<^bold>\ q) a ((ps, a) # branch) \ new (\<^bold>\ p) a ((ps, a) # branch)\) case True then show ?thesis using assms DisN ST_Suc by fast next case False then show ?thesis using assms Dup by (metis (no_types, lifting) list.set_intros(1-2) new_def on.simps set_ConsD) qed lemma DiaN': assumes \(\<^bold>\ (\<^bold>\ p)) at a in (ps, a) # branch\ \(\<^bold>\ Nom i) at a in (ps, a) # branch\ \n \ ((\<^bold>\ (\<^bold>@ i p)) # ps, a) # branch\ shows \n \ (ps, a) # branch\ using assms DiaN Dup ST_Suc by fast lemma SatP': assumes \(\<^bold>@ a p) at b in (ps, a) # branch\ \n \ (p # ps, a) # branch\ shows \n \ (ps, a) # branch\ using assms SatP Dup ST_Suc by fast lemma SatN': assumes \(\<^bold>\ (\<^bold>@ a p)) at b in (ps, a) # branch\ \n \ ((\<^bold>\ p) # ps, a) # branch\ shows \n \ (ps, a) # branch\ using assms SatN Dup ST_Suc by fast lemma Nom': assumes \p at b in (ps, a) # branch\ \Nom i at b in (ps, a) # branch\ \Nom i at a in (ps, a) # branch\ \n \ (p # ps, a) # branch\ shows \n \ (ps, a) # branch\ using assms Nom Dup ST_Suc by fast section \Substitution\ lemma finite_nominals: \finite (nominals p)\ by (induct p) simp_all lemma finite_block_nominals: \finite (block_nominals block)\ using finite_nominals by (induct block) auto lemma finite_branch_nominals: \finite (branch_nominals branch)\ unfolding branch_nominals_def by (induct branch) (auto simp: finite_block_nominals) abbreviation sub_list :: \('b \ 'c) \ ('a, 'b) fm list \ ('a, 'c) fm list\ where \sub_list f ps \ map (sub f) ps\ primrec sub_block :: \('b \ 'c) \ ('a, 'b) block \ ('a, 'c) block\ where \sub_block f (ps, i) = (sub_list f ps, f i)\ definition sub_branch :: \('b \ 'c) \ ('a, 'b) branch \ ('a, 'c) branch\ where \sub_branch f blocks \ map (sub_block f) blocks\ lemma sub_block_mem: \p on block \ sub f p on sub_block f block\ by (induct block) auto lemma sub_branch_mem: assumes \(ps, i) \. branch\ shows \(sub_list f ps, f i) \. sub_branch f branch\ unfolding sub_branch_def using assms image_iff by fastforce lemma sub_block_nominals: \block_nominals (sub_block f block) = f ` block_nominals block\ by (induct block) (auto simp: sub_nominals) lemma sub_branch_nominals: \branch_nominals (sub_branch f branch) = f ` branch_nominals branch\ unfolding branch_nominals_def sub_branch_def by (induct branch) (auto simp: sub_block_nominals) lemma sub_list_id: \sub_list id ps = ps\ using sub_id by (induct ps) auto lemma sub_block_id: \sub_block id block = block\ using sub_list_id by (induct block) auto lemma sub_branch_id: \sub_branch id branch = branch\ unfolding sub_branch_def using sub_block_id by (induct branch) auto lemma sub_block_upd_fresh: assumes \i \ block_nominals block\ shows \sub_block (f(i := j)) block = sub_block f block\ using assms by (induct block) (auto simp add: sub_upd_fresh) lemma sub_branch_upd_fresh: assumes \i \ branch_nominals branch\ shows \sub_branch (f(i := j)) branch = sub_branch f branch\ using assms unfolding branch_nominals_def sub_branch_def by (induct branch) (auto simp: sub_block_upd_fresh) lemma sub_comp: \sub f (sub g p) = sub (f o g) p\ by (induct p) simp_all lemma sub_list_comp: \sub_list f (sub_list g ps) = sub_list (f o g) ps\ using sub_comp by (induct ps) auto lemma sub_block_comp: \sub_block f (sub_block g block) = sub_block (f o g) block\ using sub_list_comp by (induct block) simp_all lemma sub_branch_comp: \sub_branch f (sub_branch g branch) = sub_branch (f o g) branch\ unfolding sub_branch_def using sub_block_comp by (induct branch) fastforce+ lemma swap_id: \(id(i := j, j := i)) o (id(i := j, j := i)) = id\ by auto lemma at_in_sub_branch: assumes \p at i in (ps, a) # branch\ shows \sub f p at f i in (sub_list f ps, f a) # sub_branch f branch\ using assms sub_branch_mem by fastforce text \ If a branch has a closing tableau then so does any branch obtained by renaming nominals. Since some formulas on the renamed branch may no longer be new, they do not contribute any fuel and so we existentially quantify over the fuel needed to close the new branch. \ lemma ST_sub': fixes f :: \'b \ 'c\ assumes \\(f :: 'b \ 'c) i A. finite A \ i \ A \ \j. j \ f ` A\ shows \n \ branch \ \ sub_branch f branch\ proof (induct branch arbitrary: f rule: ST.induct) case (Close p i branch) have \sub f p at f i in sub_branch f branch\ using Close(1) sub_branch_mem by fastforce moreover have \(\<^bold>\ sub f p) at f i in sub_branch f branch\ using Close(2) sub_branch_mem by force ultimately show ?case using ST.Close by fast next case (Neg p a ps branch n) then have \\ (sub f p # sub_list f ps, f a) # sub_branch f branch\ unfolding sub_branch_def by simp moreover have \(\<^bold>\ \<^bold>\ sub f p) at f a in (sub_list f ps, f a) # sub_branch f branch\ using Neg(1) at_in_sub_branch by (metis (no_types, hide_lams) sub.simps(3)) ultimately have \\ (sub_list f ps, f a) # sub_branch f branch\ using Neg' by fast then show ?case unfolding sub_branch_def by simp next case (DisP p q a ps branch n) then have \\ (sub f p # sub_list f ps, f a) # sub_branch f branch\ \\ (sub f q # sub_list f ps, f a) # sub_branch f branch\ unfolding sub_branch_def by simp_all moreover have \(sub f p \<^bold>\ sub f q) at f a in (sub_list f ps, f a) # sub_branch f branch\ using DisP(1) at_in_sub_branch by (metis (no_types, hide_lams) sub.simps(4)) ultimately have \\ (sub_list f ps, f a) # sub_branch f branch\ using DisP'' by fast then show ?case unfolding sub_branch_def by simp next case (DisN p q a ps branch n) then have \\ ((\<^bold>\ sub f q) # (\<^bold>\ sub f p) # sub_list f ps, f a) # sub_branch f branch\ unfolding sub_branch_def by simp moreover have \(\<^bold>\ (sub f p \<^bold>\ sub f q)) at f a in (sub_list f ps, f a) # sub_branch f branch\ using DisN(1) at_in_sub_branch by (metis (no_types, hide_lams) sub.simps(3-4)) ultimately have \\ (sub_list f ps, f a) # sub_branch f branch\ using DisN' by fast then show ?case unfolding sub_branch_def by simp next case (DiaP p a ps branch i n) show ?case proof (cases \witnessed (sub f p) (f a) (sub_branch f ((ps, a) # branch))\) case True then obtain i' where rs: \(\<^bold>@ i' (sub f p)) at f a in (sub_list f ps, f a) # sub_branch f branch\ and ts: \(\<^bold>\ Nom i') at f a in (sub_list f ps, f a) # sub_branch f branch\ unfolding sub_branch_def witnessed_def by auto from rs have rs': \(\<^bold>@ i' (sub f p)) at f a in ((\<^bold>\ Nom i') # sub_list f ps, f a) # sub_branch f branch\ by fastforce let ?f = \f(i := i')\ let ?branch = \sub_branch ?f branch\ have \sub_branch ?f ((ps, a) # branch) = sub_branch f ((ps, a) # branch)\ using DiaP(2) sub_branch_upd_fresh by fast then have **: \sub_list ?f ps = sub_list f ps\ \?f a = f a\ \?branch = sub_branch f branch\ unfolding sub_branch_def by simp_all have p: \sub ?f p = sub f p\ using DiaP(1-2) sub_upd_fresh unfolding branch_nominals_def by fastforce have \\ ((\<^bold>@ (?f i) (sub ?f p)) # (\<^bold>\ Nom (?f i)) # sub_list ?f ps, ?f a) # ?branch\ using DiaP(6)[where f=\?f\] unfolding sub_branch_def by simp then have \\ ((\<^bold>@ i' (sub f p)) # (\<^bold>\ Nom i') # sub_list f ps, f a) # sub_branch f branch\ using p ** by simp then have \\ ((\<^bold>\ Nom i') # sub_list f ps, f a) # sub_branch f branch\ using rs' by (meson Nom' on.simps) then have \\ (sub_list f ps, f a) # sub_branch f branch\ using ts by (meson Nom' on.simps) then show ?thesis unfolding sub_branch_def by auto next case False have \finite (branch_nominals ((ps, a) # branch))\ by (simp add: finite_branch_nominals) then obtain j where *: \j \ f ` branch_nominals ((ps, a) # branch)\ using DiaP(2) assms by blast let ?f = \f(i := j)\ let ?branch = \sub_branch ?f branch\ have **: \sub_branch ?f ((ps, a) # branch) = sub_branch f ((ps, a) # branch)\ using DiaP(2) sub_branch_upd_fresh by fast then have ***: \sub_list ?f ps = sub_list f ps\ \?f a = f a\ \?branch = sub_branch f branch\ unfolding sub_branch_def by simp_all moreover have p: \sub ?f p = sub f p\ using DiaP(1-2) sub_upd_fresh unfolding branch_nominals_def by fastforce ultimately have \\ witnessed (sub ?f p) (?f a) (sub_branch ?f ((ps, a) # branch))\ using False ** by simp then have w: \\ witnessed (sub ?f p) (?f a) ((sub_list ?f ps, ?f a) # ?branch)\ unfolding sub_branch_def by simp have \\ ((\<^bold>@ (?f i) (sub ?f p)) # (\<^bold>\ Nom (?f i)) # sub_list ?f ps, ?f a) # ?branch\ using DiaP(6)[where f=\?f\] ST_Suc unfolding sub_branch_def by simp moreover have \sub ?f (\<^bold>\ p) at ?f a in (sub_list ?f ps, ?f a) # sub_branch ?f branch\ using DiaP(1) at_in_sub_branch by fast then have \(\<^bold>\ sub ?f p) at ?f a in (sub_list ?f ps, ?f a) # sub_branch ?f branch\ by simp moreover have \\a. sub ?f p = Nom a\ using DiaP(3) by (cases p) simp_all moreover have \?f i \ branch_nominals ((sub_list ?f ps, ?f a) # ?branch)\ using * ** sub_branch_nominals unfolding sub_branch_def by (metis fun_upd_same list.simps(9) sub_block.simps) ultimately have \\ (sub_list ?f ps, ?f a) # ?branch\ using w ST.DiaP ST_Suc by fast then show ?thesis using *** unfolding sub_branch_def by simp qed next case (DiaN p a ps branch i n) then have \\ ((\<^bold>\ (\<^bold>@ (f i) (sub f p))) # sub_list f ps, f a) # sub_branch f branch\ unfolding sub_branch_def by simp moreover have \(\<^bold>\ (\<^bold>\ sub f p)) at f a in (sub_list f ps, f a) # sub_branch f branch\ using DiaN(1) at_in_sub_branch by (metis (no_types, hide_lams) sub.simps(3, 5)) moreover have \(\<^bold>\ Nom (f i)) at f a in (sub_list f ps, f a) # sub_branch f branch\ using DiaN(2) at_in_sub_branch by (metis (no_types, hide_lams) sub.simps(2, 5)) ultimately have \\ (sub_list f ps, f a) # sub_branch f branch\ using DiaN' by fast then show ?case unfolding sub_branch_def by simp next case (SatP a p b ps branch n) then have \\ (sub f p # sub_list f ps, f a) # sub_branch f branch\ unfolding sub_branch_def by simp moreover have \(\<^bold>@ (f a) (sub f p)) at f b in (sub_list f ps, f a) # sub_branch f branch\ using SatP(1) at_in_sub_branch by (metis (no_types, hide_lams) sub.simps(6)) ultimately have \\ (sub_list f ps, f a) # sub_branch f branch\ using SatP' by fast then show ?case unfolding sub_branch_def by simp next case (SatN a p b ps branch n) then have \\ ((\<^bold>\ sub f p) # sub_list f ps, f a) # sub_branch f branch\ unfolding sub_branch_def by simp moreover have \(\<^bold>\ (\<^bold>@ (f a) (sub f p))) at f b in (sub_list f ps, f a) # sub_branch f branch\ using SatN(1) at_in_sub_branch by (metis (no_types, hide_lams) sub.simps(3, 6)) ultimately have \\ (sub_list f ps, f a) # sub_branch f branch\ using SatN' by fast then show ?case unfolding sub_branch_def by simp next case (GoTo i branch n) then have \\ ([], f i) # sub_branch f branch\ unfolding sub_branch_def by simp moreover have \f i \ branch_nominals (sub_branch f branch)\ using GoTo(1) sub_branch_nominals by fast ultimately show ?case using ST.GoTo by fast next case (Nom p b ps a branch i n) then have \\ (sub f p # sub_list f ps, f a) # sub_branch f branch\ unfolding sub_branch_def by simp moreover have \sub f p at f b in (sub_list f ps, f a) # sub_branch f branch\ using Nom(1) at_in_sub_branch by fast moreover have \Nom (f i) at f b in (sub_list f ps, f a) # sub_branch f branch\ using Nom(2) at_in_sub_branch by (metis (no_types, hide_lams) sub.simps(2)) moreover have \Nom (f i) at f a in (sub_list f ps, f a) # sub_branch f branch\ using Nom(3) at_in_sub_branch by (metis (no_types, hide_lams) sub.simps(2)) ultimately have \\ (sub_list f ps, f a) # sub_branch f branch\ using Nom' by fast then show ?case unfolding sub_branch_def by simp qed lemma ex_fresh_gt: fixes f :: \'b \ 'c\ assumes \\g :: 'c \ 'b. surj g\ \finite A\ \i \ A\ shows \\j. j \ f ` A\ proof (rule ccontr) assume \\j. j \ f ` A\ moreover obtain g :: \'c \ 'b\ where \surj g\ using assms(1) by blast ultimately show False using assms(2-3) by (metis UNIV_I UNIV_eq_I card_image_le card_seteq finite_imageI image_comp subsetI) qed corollary ST_sub_gt: fixes f :: \'b \ 'c\ assumes \\g :: 'c \ 'b. surj g\ \\ branch\ shows \\ sub_branch f branch\ using assms ex_fresh_gt ST_sub' by metis corollary ST_sub_inf: fixes f :: \'b \ 'c\ assumes \infinite (UNIV :: 'c set)\ \n \ branch\ shows \\ sub_branch f branch\ proof - have \finite A \ \j. j \ f ` A\ for A and f :: \'b \ 'c\ using assms(1) ex_new_if_finite by blast then show ?thesis using assms(2) ST_sub' by metis qed corollary ST_sub: fixes f :: \'b \ 'b\ assumes \\ branch\ shows \\ sub_branch f branch\ proof - have \finite A \ i \ A \ \j. j \ f ` A\ for i A and f :: \'b \ 'b\ by (metis card_image_le card_seteq finite_imageI subsetI) then show ?thesis using assms ST_sub' by metis qed subsection \Unrestricted \(\<^bold>\)\ rule\ lemma DiaP': assumes \(\<^bold>\ p) at a in (ps, a) # branch\ \\a. p = Nom a\ \i \ branch_nominals ((ps, a) # branch)\ \\ ((\<^bold>@ i p) # (\<^bold>\ Nom i) # ps, a) # branch\ shows \\ (ps, a) # branch\ proof (cases \witnessed p a ((ps, a) # branch)\) case True then obtain i' where rs: \(\<^bold>@ i' p) at a in (ps, a) # branch\ and ts: \(\<^bold>\ Nom i') at a in (ps, a) # branch\ unfolding witnessed_def by blast then have rs': \(\<^bold>@ i' p) at a in ((\<^bold>\ Nom i') # ps, a) # branch\ by fastforce let ?f = \id(i := i')\ have \\ sub_branch ?f (((\<^bold>@ i p) # (\<^bold>\ Nom i) # ps, a) # branch)\ using assms(4) ST_sub by blast then have \\ ((\<^bold>@ i' (sub ?f p)) # (\<^bold>\ Nom i') # sub_list ?f ps, ?f a) # sub_branch ?f branch\ unfolding sub_branch_def by simp moreover have \i \ nominals p\ \i \ list_nominals ps\ \i \ a\ \i \ branch_nominals branch\ using assms(1-3) unfolding branch_nominals_def by fastforce+ then have \sub ?f p = p\ by (simp add: sub_id sub_upd_fresh) moreover have \sub_list ?f ps = ps\ using \i \ list_nominals ps\ by (simp add: map_idI sub_id sub_upd_fresh) moreover have \?f a = a\ using \i \ a\ by simp moreover have \sub_branch ?f branch = branch\ using \i \ branch_nominals branch\ by (simp add: sub_branch_id sub_branch_upd_fresh) ultimately have \\ ((\<^bold>@ i' p) # (\<^bold>\ Nom i') # ps, a) # branch\ by simp then have \\ ((\<^bold>\ Nom i') # ps, a) # branch\ using rs' by (meson Nom' on.simps) then show ?thesis using ts by (meson Nom' on.simps) next case False then show ?thesis using assms DiaP ST_Suc by fast qed section \Structural Properties\ lemma block_nominals_branch: assumes \block \. branch\ shows \block_nominals block \ branch_nominals branch\ unfolding branch_nominals_def using assms by blast lemma sub_block_fresh: assumes \i \ branch_nominals branch\ \block \. branch\ shows \sub_block (f(i := j)) block = sub_block f block\ using assms block_nominals_branch sub_block_upd_fresh by fast lemma list_down_induct [consumes 1, case_names Start Cons]: assumes \\y \ set ys. Q y\ \P (ys @ xs)\ \\y xs. Q y \ P (y # xs) \ P xs\ shows \P xs\ using assms by (induct ys) auto text \ If the last block on a branch has opening nominal \a\ and the last formulas on that block occur on another block alongside nominal \a\, then we can drop those formulas. \ lemma ST_drop_prefix: assumes \set ps \ set qs\ \(qs, a) \. branch\ \n \ (ps @ ps', a) # branch\ shows \n \ (ps', a) # branch\ proof - have \\p \ set ps. p on (qs, a)\ using assms(1) by auto then show ?thesis proof (induct ps' rule: list_down_induct) case Start then show ?case using assms(3) . next case (Cons p ps) then show ?case using assms(2) by (meson Nom' list.set_intros(2) on.simps) qed qed text \We can drop a block if it is subsumed by another block.\ lemma ST_drop_block: assumes \set ps \ set ps'\ \(ps', a) \. branch\ \n \ (ps, a) # branch\ shows \Suc n \ branch\ using assms proof (induct branch) case Nil then show ?case by simp next case (Cons block branch) then show ?case proof (cases block) case (Pair qs b) then have \n \ ([], a) # (qs, b) # branch\ using Cons(2-4) ST_drop_prefix[where branch=\(qs, b) # branch\] by simp moreover have \a \ branch_nominals ((qs, b) # branch)\ unfolding branch_nominals_def using Cons(3) Pair by fastforce ultimately have \Suc n \ (qs, b) # branch\ by (simp add: GoTo) then show ?thesis using Pair Dup by fast qed qed lemma ST_drop_block': assumes \n \ (ps, a) # branch\ \(ps, a) \. branch\ shows \Suc n \ branch\ using assms ST_drop_block by fastforce lemma sub_branch_image: \set (sub_branch f branch) = sub_block f ` set branch\ unfolding sub_branch_def by simp lemma sub_block_repl: assumes \j \ block_nominals block\ shows \i \ block_nominals (sub_block (id(i := j, j := i)) block)\ using assms by (simp add: image_iff sub_block_nominals) lemma sub_branch_repl: assumes \j \ branch_nominals branch\ shows \i \ branch_nominals (sub_branch (id(i := j, j := i)) branch)\ using assms by (simp add: image_iff sub_branch_nominals) text \If a finite set of blocks has a closing tableau then so does any finite superset.\ lemma ST_struct: fixes branch :: \('a, 'b) branch\ assumes inf: \infinite (UNIV :: 'b set)\ and \n \ branch\ \set branch \ set branch'\ shows \\ branch'\ using assms(2-3) proof (induct n branch arbitrary: branch' rule: ST.induct) case (Close p i branch) then show ?case using ST.Close by fast next case (Neg p a ps branch n) have \\ (p # ps, a) # branch'\ using Neg(4-5) by (simp add: subset_code(1)) moreover have \(\<^bold>\ \<^bold>\ p) at a in (ps, a) # branch'\ using Neg(1, 5) by auto ultimately have \\ (ps, a) # branch'\ using Neg' by fast moreover have \(ps, a) \. branch'\ using Neg(5) by simp ultimately show ?case using ST_drop_block' by fast next case (DisP p q a ps branch) have \\ (p # ps, a) # branch'\ \\ (q # ps, a) # branch'\ using DisP(5, 7-8) by (simp_all add: subset_code(1)) moreover have \(p \<^bold>\ q) at a in (ps, a) # branch'\ using DisP(1, 8) by auto ultimately have \\ (ps, a) # branch'\ using DisP'' by fast moreover have \(ps, a) \. branch'\ using DisP(8) by simp ultimately show ?case using ST_drop_block' by fast next case (DisN p q a ps branch n) have \\ ((\<^bold>\ q) # (\<^bold>\ p) # ps, a)# branch'\ using DisN(4-5) by (simp add: subset_code(1)) moreover have \(\<^bold>\ (p \<^bold>\ q)) at a in (ps, a) # branch'\ using DisN(1, 5) by auto ultimately have \\ (ps, a) # branch'\ using DisN' by fast moreover have \(ps, a) \. branch'\ using DisN(5) by simp ultimately show ?case using ST_drop_block' by fast next case (DiaP p a ps branch i n) have \finite (branch_nominals branch')\ by (simp add: finite_branch_nominals) then obtain j where j: \j \ branch_nominals branch'\ using assms ex_new_if_finite by blast then have j': \j \ branch_nominals ((ps, a) # branch)\ using DiaP(7) unfolding branch_nominals_def by blast let ?f = \id(i := j, j := i)\ let ?branch' = \sub_branch ?f branch'\ have branch': \sub_branch ?f ?branch' = branch'\ using sub_branch_comp sub_branch_id swap_id by metis have branch: \sub_branch ?f ((ps, a) # branch) = (ps, a) # branch\ using DiaP(2) j' sub_branch_id sub_branch_upd_fresh by metis moreover have \set (sub_branch ?f ((ps, a) # branch)) \ set ?branch'\ using DiaP(7) sub_branch_image by blast ultimately have *: \set ((ps, a) # branch) \ set ?branch'\ unfolding sub_branch_def by auto have \i \ block_nominals (ps, a)\ using DiaP unfolding branch_nominals_def by simp moreover have \i \ branch_nominals ?branch'\ using j sub_branch_repl by metis ultimately have i: \i \ branch_nominals ((ps, a) # ?branch')\ unfolding branch_nominals_def by simp have \\ ((\<^bold>@ i p) # (\<^bold>\ Nom i) # ps, a) # ?branch'\ using DiaP(6) * by (metis (no_types, lifting) subset_code(1) insert_mono list.set(2) set_subset_Cons) moreover have \(\<^bold>\ p) at a in (ps, a) # ?branch'\ using DiaP(1, 7) * by (meson set_subset_Cons subset_code(1)) ultimately have \\ (ps, a) # ?branch'\ using inf DiaP(3) i DiaP' by metis then have \\ sub_branch ?f ((ps, a) # ?branch')\ using ST_sub by blast then have \\ (ps, a) # branch'\ using branch' branch unfolding sub_branch_def by simp moreover have \(ps, a) \. branch'\ using \set ((ps, a) # branch) \ set branch'\ by simp ultimately show ?case using ST_drop_block' by fast next case (DiaN p a ps branch i n) have \\ ((\<^bold>\ (\<^bold>@ i p)) # ps, a) # branch'\ using DiaN(5-6) by (simp add: subset_code(1)) moreover have \(\<^bold>\ (\<^bold>\ p)) at a in (ps, a) # branch'\ \(\<^bold>\ Nom i) at a in (ps, a) # branch'\ using DiaN(1-2, 6) by auto ultimately have \\ (ps, a) # branch'\ using DiaN' by fast moreover have \(ps, a) \. branch'\ using DiaN(6) by simp ultimately show ?case using ST_drop_block' by fast next case (SatP a p b ps branch n) have \\ (p # ps, a) # branch'\ using SatP(4-5) by (simp add: subset_code(1)) moreover have \(\<^bold>@ a p) at b in (ps, a) # branch'\ using SatP(1, 5) by auto ultimately have \\ (ps, a) # branch'\ using SatP' by fast moreover have \(ps, a) \. branch'\ using SatP(5) by simp ultimately show ?case using ST_drop_block' by fast next case (SatN a p b ps branch n) have\\ ((\<^bold>\ p) # ps, a) # branch'\ using SatN(4-5) by (simp add: subset_code(1)) moreover have \(\<^bold>\ (\<^bold>@ a p)) at b in (ps, a) # branch'\ using SatN(1, 5) by auto ultimately have \\ (ps, a) # branch'\ using SatN' by fast moreover have \(ps, a) \. branch'\ using SatN(5) by simp ultimately show ?case using ST_drop_block' by fast next case (GoTo i branch n) then have \\ ([], i) # branch'\ by (simp add: subset_code(1)) moreover have \i \ branch_nominals branch'\ using GoTo(1, 4) unfolding branch_nominals_def by auto ultimately show ?case using GoTo(2) ST.GoTo by fast next case (Nom p b ps a branch i n) have \\ (p # ps, a) # branch'\ using Nom(6-7) by (simp add: subset_code(1)) moreover have \p at b in (ps, a) # branch'\ using Nom(1, 7) by auto moreover have \Nom i at b in (ps, a) # branch'\ using Nom(2, 7) by auto moreover have \Nom i at a in (ps, a) # branch'\ using Nom(3, 7) by auto ultimately have \\ (ps, a) # branch'\ using Nom' by fast moreover have \(ps, a) \. branch'\ using Nom(7) by simp ultimately show ?case using ST_drop_block' by fast qed text \ If a branch has a closing tableau then we can replace the formulas of the last block on that branch with any finite superset and still obtain a closing tableau. \ lemma ST_struct_block: fixes branch :: \('a, 'b) branch\ assumes inf: \infinite (UNIV :: 'b set)\ and \n \ (ps, a) # branch\ \set ps \ set ps'\ shows \\ (ps', a) # branch\ using assms(2-3) proof (induct \(ps, a) # branch\ arbitrary: ps ps' rule: ST.induct) case (Close p i n ts ts') then have \p at i in (ts', a) # branch\ \(\<^bold>\ p) at i in (ts', a) # branch\ by auto then show ?case using ST.Close by fast next case (Neg p ps n) then have \(\<^bold>\ \<^bold>\ p) at a in (ps', a) # branch\ by auto moreover have \\ (p # ps', a) # branch\ using Neg(4-5) by (simp add: subset_code(1)) ultimately show ?case using Neg' by fast next case (DisP p q ps n) then have \(p \<^bold>\ q) at a in (ps', a) # branch\ by auto moreover have \\ (p # ps', a) # branch\ \\ (q # ps', a) # branch\ using DisP(5, 7-8) by (simp_all add: subset_code(1)) ultimately show ?case using DisP'' by fast next case (DisN p q ps n) then have \(\<^bold>\ (p \<^bold>\ q)) at a in (ps', a) # branch\ by auto moreover have \\ ((\<^bold>\ q) # (\<^bold>\ p) # ps', a) # branch\ using DisN(4-5) by (simp add: subset_code(1)) ultimately show ?case using DisN' by fast next case (DiaP p ps i n) have \finite (branch_nominals ((ps', a) # branch))\ using finite_branch_nominals by blast then obtain j where j: \j \ branch_nominals ((ps', a) # branch)\ using assms ex_new_if_finite by blast then have j': \j \ block_nominals (ps, a)\ using DiaP.prems unfolding branch_nominals_def by auto let ?f = \id(i := j, j := i)\ let ?ps' = \sub_list ?f ps'\ have ps': \sub_list ?f ?ps' = ps'\ using sub_list_comp sub_list_id swap_id by metis have \i \ block_nominals (ps, a)\ using DiaP(1-2) unfolding branch_nominals_def by simp then have ps: \sub_block ?f (ps, a) = (ps, a)\ using j' sub_block_id sub_block_upd_fresh by metis moreover have \set (sub_list ?f ps) \ set (sub_list ?f ps')\ using \set ps \ set ps'\ by auto ultimately have *: \set ps \ set ?ps'\ by simp have \i \ branch_nominals branch\ using DiaP unfolding branch_nominals_def by simp moreover have \j \ branch_nominals branch\ using j unfolding branch_nominals_def by simp ultimately have branch: \sub_branch ?f branch = branch\ using sub_branch_id sub_branch_upd_fresh by metis have \i \ a\ \j \ a\ using DiaP j unfolding branch_nominals_def by simp_all then have \?f a = a\ by simp moreover have \j \ block_nominals (ps', a)\ using j unfolding branch_nominals_def by simp ultimately have \i \ block_nominals (?ps', a)\ using sub_block_repl[where block=\(ps', a)\ and i=i and j=j] by simp have \(\<^bold>\ p) at a in (?ps', a) # branch\ using DiaP(1) * by fastforce moreover have \\ ((\<^bold>@ i p) # (\<^bold>\ Nom i) # ?ps', a) # branch\ using * DiaP(6) by (simp add: subset_code(1)) moreover have \i \ branch_nominals ((?ps', a) # branch)\ using DiaP(2) \i \ block_nominals (?ps', a)\ unfolding branch_nominals_def by simp ultimately have \\ (?ps', a) # branch\ using DiaP(3) DiaP' by fast then have \\ sub_branch ?f ((?ps', a) # branch)\ using ST_sub by blast then have \\ (sub_list ?f ?ps', ?f a) # sub_branch ?f branch\ unfolding sub_branch_def by simp then show ?case using \?f a = a\ ps' branch by simp next case (DiaN p ps i n) then have \(\<^bold>\ (\<^bold>\ p)) at a in (ps', a) # branch\ \(\<^bold>\ Nom i) at a in (ps', a) # branch\ by auto moreover have \\ ((\<^bold>\ (\<^bold>@ i p)) # ps', a) # branch\ using DiaN(5-6) by (simp add: subset_code(1)) ultimately show ?case using DiaN' by fast next case (SatP p b ps n) then have \(\<^bold>@ a p) at b in (ps', a) # branch\ by auto moreover have \\ (p # ps', a) # branch\ using SatP(4-5) by (simp add: subset_code(1)) ultimately show ?case using SatP' by fast next case (SatN p b ps n) then have \(\<^bold>\ (\<^bold>@ a p)) at b in (ps', a) # branch\ by auto moreover have \\ ((\<^bold>\ p) # ps', a) # branch\ using SatN(4-5) by (simp add: subset_code(1)) ultimately show ?case using SatN' by fast next case (GoTo i n ps) then have \Suc n \ (ps, a) # branch\ using ST.GoTo by fast then obtain m where \m \ (ps, a) # (ps', a) # branch\ using inf ST_struct[where branch'=\(ps, a) # _ # _\] by fastforce then have \Suc m \ (ps', a) # branch\ using GoTo(4) by (simp add: ST_drop_block[where a=a]) then show ?case by blast next case (Nom p b ps i) have \p at b in (ps', a) # branch\ using Nom(1, 7) by auto moreover have \Nom i at b in (ps', a) # branch\ using Nom(2, 7) by auto moreover have \Nom i at a in (ps', a) # branch\ using Nom(3, 7) by auto moreover have \\ (p # ps', a) # branch\ using Nom(6-7) by (simp add: subset_code(1)) ultimately show ?case using Nom' by fast qed section \Unrestricted \(\<^bold>@)\ and \(\<^bold>\ \<^bold>@)\ rules\ text \We can derive more general versions of the \(\<^bold>@)\ and \(\<^bold>\ \<^bold>@)\ rules using weakening and \Nom\.\ lemma SatP'': fixes i :: 'b assumes inf: \infinite (UNIV :: 'b set)\ and \(\<^bold>@ i p) at b in (ps, a) # branch\ \Nom i at a in (ps, a) # branch\ \\ (p # ps, a) # branch\ shows \\ (ps, a) # branch\ proof - have \i \ branch_nominals ((ps, a) # branch)\ using assms(2) unfolding branch_nominals_def by fastforce then have ?thesis if \\ ([], i) # (ps, a) # branch\ using that GoTo by fast then have ?thesis if \\ ([Nom a], i) # (ps, a) # branch\ using that assms(3) by (meson Nom' list.set_intros(1) on.simps set_subset_Cons subsetD) then have ?thesis if \\ ([p, Nom a], i) # (ps, a) # branch\ using that assms(2) by (meson SatP' set_subset_Cons subsetD) then have ?thesis if \\ (ps, a) # ([p, Nom a], i) # branch\ using that inf ST_struct[where branch'=\([p, Nom a], i) # (ps, a) # branch\] by fastforce then have ?thesis if \\ (p # ps, a) # ([p, Nom a], i) # branch\ using that by (meson Nom' list.set_intros(1) on.simps set_subset_Cons subsetD) then show ?thesis using inf assms(4) ST_struct[where branch'=\(p # ps, a) # ([p, Nom a], i) # branch\] by fastforce qed lemma SatN'': fixes i :: 'b assumes inf: \infinite (UNIV :: 'b set)\ and \(\<^bold>\ (\<^bold>@ i p)) at b in (ps, a) # branch\ \Nom i at a in (ps, a) # branch\ \\ ((\<^bold>\ p) # ps, a) # branch\ shows \\ (ps, a) # branch\ proof - have \i \ branch_nominals ((ps, a) # branch)\ using assms(2) unfolding branch_nominals_def by fastforce then have ?thesis if \\ ([], i) # (ps, a) # branch\ using that GoTo by fast then have ?thesis if \\ ([Nom a], i) # (ps, a) # branch\ using that assms(3) by (meson Nom' list.set_intros(1) on.simps set_subset_Cons subsetD) then have ?thesis if \\ ([\<^bold>\ p, Nom a], i) # (ps, a) # branch\ using that assms(2) by (meson SatN' set_subset_Cons subsetD) then have ?thesis if \\ (ps, a) # ([\<^bold>\ p, Nom a], i) # branch\ using that inf ST_struct[where branch'=\([\<^bold>\ p, Nom a], i) # (ps, a) # branch\] by fastforce then have ?thesis if \\ ((\<^bold>\ p) # ps, a) # ([\<^bold>\ p, Nom a], i) # branch\ using that by (meson Nom' list.set_intros(1) on.simps set_subset_Cons subsetD) then show ?thesis using inf assms(4) ST_struct[where branch'=\((\<^bold>\ p) # ps, a) # ([\<^bold>\ p, Nom a], i) # branch\] by fastforce qed section \Bridge\ text \ We are going to define a \descendants k i branch\ relation on sets of indices. The sets are built on the index of a \\<^bold>\ Nom k\ on an \i\-block in \branch\ and can be extended by indices of formula occurrences that can be thought of as descending from that \\<^bold>\ Nom k\ by application of either the \(\<^bold>\ \<^bold>\)\ or \Nom\ rule. We will show that if we have nominals \j\ and \k\ on the same block in a closeable branch, then the branch obtained by the following transformation is also closeable: For every index \v\, if the formula at \v\ is \\<^bold>\ Nom k\, replace it by \\<^bold>\ Nom j\ and if it is \\<^bold>\ (\<^bold>@ k p)\ replace it by \\<^bold>\ (\<^bold>@ j p)\. There are no other cases. From this transformation we can derive the Bridge rule. \ subsection \Replacing\ abbreviation bridge' :: \'b \ 'b \ ('a, 'b) fm \ ('a, 'b) fm\ where \bridge' k j p \ case p of (\<^bold>\ Nom k') \ (if k = k' then (\<^bold>\ Nom j) else (\<^bold>\ Nom k')) | (\<^bold>\ (\<^bold>@ k' q)) \ (if k = k' then (\<^bold>\ (\<^bold>@ j q)) else (\<^bold>\ (\<^bold>@ k' q))) | p \ p\ abbreviation bridge :: \'b \ 'b \ (nat \ nat) set \ nat \ nat \ ('a, 'b) fm \ ('a, 'b) fm\ where \bridge k j \ mapper (bridge' k j)\ lemma bridge_on_Nom: \Nom i on (ps, a) \ Nom i on (mapi (bridge k j xs v) ps, a)\ by (induct ps) auto lemma bridge'_nominals: \nominals (bridge' k j p) \ {k, j} = nominals p \ {k, j}\ proof (induct p) case (Neg p) then show ?case by (cases p) auto next case (Dia p) then show ?case by (cases p) auto qed auto lemma bridge_nominals: \nominals (bridge k j xs v v' p) \ {k, j} = nominals p \ {k, j}\ proof (cases \(v, v') \ xs\) case True then have \nominals (bridge k j xs v v' p) = nominals (bridge' k j p)\ by simp then show ?thesis using bridge'_nominals by metis qed simp lemma bridge_block_nominals: \block_nominals (mapi_block (bridge k j xs v) (ps, a)) \ {k, j} = block_nominals (ps, a) \ {k, j}\ proof (induct ps) case Nil then show ?case by simp next case (Cons p ps) have \?case \ (nominals (bridge k j xs v (length ps) p)) \ (block_nominals (mapi_block (bridge k j xs v) (ps, a)) \ {k, j}) = (nominals p) \ (block_nominals (ps, a) \ {k, j})\ by simp also have \\ \ (nominals (bridge k j xs v (length ps) p) \ {k, j}) \ (block_nominals (mapi_block (bridge k j xs v) (ps, a)) \ {k, j}) = (nominals p \ {k, j}) \ (block_nominals (ps, a) \ {k, j})\ by blast moreover have \nominals (bridge k j xs v (length ps) p) \ {k, j} = nominals p \ {k, j}\ using bridge_nominals by metis moreover note Cons ultimately show ?case by argo qed lemma bridge_branch_nominals: \branch_nominals (mapi_branch (bridge k j xs) branch) \ {k, j} = branch_nominals branch \ {k, j}\ proof (induct branch) case Nil then show ?case unfolding branch_nominals_def mapi_branch_def by simp next case (Cons block branch) have \?case \ (block_nominals (mapi_block (bridge k j xs (length branch)) block)) \ (branch_nominals (mapi_branch (bridge k j xs) branch) \ {k, j}) = (block_nominals block) \ (branch_nominals branch \ {k, j})\ unfolding branch_nominals_def mapi_branch_def by simp also have \\ \ (block_nominals (mapi_block (bridge k j xs (length branch)) block) \ {k, j}) \ (branch_nominals (mapi_branch (bridge k j xs) branch) \ {k, j}) = (block_nominals block \ {k, j}) \ (branch_nominals branch \ {k, j})\ by blast moreover have \block_nominals (mapi_block (bridge k j xs (length branch)) block) \ {k, j} = block_nominals block \ {k, j}\ using bridge_block_nominals[where ps=\fst block\ and a=\snd block\] by simp ultimately show ?case using Cons by argo qed lemma at_in_mapi_branch: assumes \p at a in branch\ \p \ Nom a\ shows \\v v'. f v v' p at a in mapi_branch f branch\ using assms by (meson mapi_branch_mem rev_nth_mapi_block rev_nth_on) lemma nom_at_in_bridge: fixes k j xs defines \f \ bridge k j xs\ assumes \Nom i at a in branch\ shows \Nom i at a in mapi_branch f branch\ proof - obtain qs where qs: \(qs, a) \. branch\ \Nom i on (qs, a)\ using assms(2) by blast then obtain l where \(mapi (f l) qs, a) \. mapi_branch f branch\ using mapi_branch_mem by fast moreover have \Nom i on (mapi (f l) qs, a)\ unfolding f_def using qs(2) by (induct qs) auto ultimately show ?thesis by blast qed lemma nominals_mapi_branch_bridge: assumes \Nom j at a in branch\ \Nom k at a in branch\ shows \branch_nominals (mapi_branch (bridge k j xs) branch) = branch_nominals branch\ proof - let ?f = \bridge k j xs\ have \Nom j at a in mapi_branch ?f branch\ using assms(1) nom_at_in_bridge by fast then have \j \ branch_nominals (mapi_branch ?f branch)\ unfolding branch_nominals_def by fastforce moreover have \Nom k at a in mapi_branch ?f branch\ using assms(2) nom_at_in_bridge by fast then have \k \ branch_nominals (mapi_branch ?f branch)\ unfolding branch_nominals_def by fastforce moreover have \j \ branch_nominals branch\ \k \ branch_nominals branch\ using assms(1-2) unfolding branch_nominals_def by fastforce+ moreover have \branch_nominals (mapi_branch ?f branch) \ {k, j} = branch_nominals branch \ {k, j}\ using bridge_branch_nominals by metis ultimately show ?thesis by blast qed lemma bridge_proper_dia: assumes \\a. p = Nom a\ shows \bridge k j xs v v' (\<^bold>\ p) = (\<^bold>\ p)\ using assms by (induct p) simp_all lemma bridge_compl_cases: fixes k j xs v v' w w' p defines \q \ bridge k j xs v v' p\ and \q' \ bridge k j xs w w' (\<^bold>\ p)\ shows \(q = (\<^bold>\ Nom j) \ q' = (\<^bold>\ (\<^bold>\ Nom k))) \ (\r. q = (\<^bold>\ (\<^bold>@ j r)) \ q' = (\<^bold>\ \<^bold>\ (\<^bold>@ k r))) \ (\r. q = (\<^bold>@ k r) \ q' = (\<^bold>\ (\<^bold>@ j r))) \ (q = p \ q' = (\<^bold>\ p))\ proof (cases p) case (Neg p) then show ?thesis by (cases p) (simp_all add: q_def q'_def) next case (Dia p) then show ?thesis by (cases p) (simp_all add: q_def q'_def) qed (simp_all add: q_def q'_def) subsection \Descendants\ inductive descendants :: \'b \ 'b \ ('a, 'b) branch \ (nat \ nat) set \ bool\ where Initial: \branch !. v = Some (qs, i) \ qs !. v' = Some (\<^bold>\ Nom k) \ descendants k i branch {(v, v')}\ | Derived: \branch !. v = Some (qs, a) \ qs !. v' = Some (\<^bold>\ (\<^bold>@ k p)) \ descendants k i branch xs \ (w, w') \ xs \ branch !. w = Some (rs, a) \ rs !. w' = Some (\<^bold>\ Nom k) \ descendants k i branch ({(v, v')} \ xs)\ | Copied: \branch !. v = Some (qs, a) \ qs !. v' = Some p \ descendants k i branch xs \ (w, w') \ xs \ branch !. w = Some (rs, b) \ rs !. w' = Some p \ Nom j at b in branch \ Nom j at a in branch \ descendants k i branch ({(v, v')} \ xs)\ lemma descendants_initial: assumes \descendants k i branch xs\ shows \\(v, v') \ xs. \ps. branch !. v = Some (ps, i) \ ps !. v' = Some (\<^bold>\ Nom k)\ using assms by (induct k i branch xs rule: descendants.induct) simp_all lemma descendants_bounds_fst: assumes \descendants k i branch xs\ \(v, v') \ xs\ shows \v < length branch\ using assms rev_nth_Some by (induct k i branch xs rule: descendants.induct) fast+ lemma descendants_bounds_snd: assumes \descendants k i branch xs\ \(v, v') \ xs\ \branch !. v = Some (ps, a)\ shows \v' < length ps\ using assms by (induct k i branch xs rule: descendants.induct) (auto simp: rev_nth_Some) lemma descendants_branch: \descendants k i branch xs \ descendants k i (extra @ branch) xs\ proof (induct k i branch xs rule: descendants.induct) case (Initial branch v qs i v' k) then show ?case using rev_nth_append descendants.Initial by fast next case (Derived branch v qs a v' k p i xs w w' rs) then have \(extra @ branch) !. v = Some (qs, a)\ \(extra @ branch) !. w = Some (rs, a)\ using rev_nth_append by fast+ then show ?case using Derived(2, 4-5, 7) descendants.Derived by fast next case (Copied branch v qs a v' p k i xs w w' rs b j) then have \(extra @ branch) !. v = Some (qs, a)\ \(extra @ branch) !. w = Some (rs, b)\ using rev_nth_append by fast+ moreover have \Nom j at b in (extra @ branch)\ \Nom j at a in (extra @ branch)\ using Copied(8-9) by auto ultimately show ?case using Copied(2-4, 5-7, 9) descendants.Copied by fast qed lemma descendants_block: assumes \descendants k i ((ps, a) # branch) xs\ shows \descendants k i ((ps' @ ps, a) # branch) xs\ using assms proof (induct k i \(ps, a) # branch\ xs arbitrary: ps a branch rule: descendants.induct) case (Initial v qs i v' k) have \((ps' @ ps, a) # branch) !. v = Some (qs, i) \ ((ps' @ ps, a) # branch) !. v = Some (ps' @ qs, i)\ using Initial(1) by auto moreover have \qs !. v' = Some (\<^bold>\ Nom k)\ \(ps' @ qs) !. v' = Some (\<^bold>\ Nom k)\ using Initial(2) rev_nth_append by simp_all ultimately show ?case using descendants.Initial by fast next case (Derived v qs a' v' k p i xs w w' rs) have \((ps' @ ps, a) # branch) !. v = Some (qs, a') \ ((ps' @ ps, a) # branch) !. v = Some (ps' @ qs, a')\ using Derived(1) by auto moreover have \qs !. v' = Some (\<^bold>\ (\<^bold>@ k p))\ \(ps' @ qs) !. v' = Some (\<^bold>\ (\<^bold>@ k p))\ using Derived(2) rev_nth_append by simp_all moreover have \((ps' @ ps, a) # branch) !. w = Some (rs, a') \ ((ps' @ ps, a) # branch) !. w = Some (ps' @ rs, a')\ using \((ps, a) # branch) !. w = Some (rs, a')\ by auto moreover have \rs !. w' = Some (\<^bold>\ Nom k)\ \(ps' @ rs) !. w' = Some (\<^bold>\ Nom k)\ using Derived(7) rev_nth_append by simp_all ultimately show ?case using Derived(4-5) descendants.Derived by fast next case (Copied v qs a' v' p k i xs w w' rs b j) have \((ps' @ ps, a) # branch) !. v = Some (qs, a') \ ((ps' @ ps, a) # branch) !. v = Some (ps' @ qs, a')\ using Copied(1) by auto moreover have \qs !. v' = Some p\ \(ps' @ qs) !. v' = Some p\ using Copied(2) rev_nth_append by simp_all moreover have \((ps' @ ps, a) # branch) !. w = Some (rs, b) \ ((ps' @ ps, a) # branch) !. w = Some (ps' @ rs, b)\ using Copied(6) by auto moreover have \rs !. w' = Some p\ \(ps' @ rs) !. w' = Some p\ using Copied(7) rev_nth_append by simp_all moreover have \((ps' @ ps, a) # branch) !. w = Some (rs, b) \ ((ps' @ ps, a) # branch) !. w = Some (ps' @ rs, b)\ using Copied(6) by auto moreover have \rs !. w' = Some p\ \(ps' @ rs) !. w' = Some p\ using Copied(7) rev_nth_append by simp_all moreover have \Nom j at b in (ps' @ ps, a) # branch\ using Copied(8) by fastforce moreover have \Nom j at a' in (ps' @ ps, a) # branch\ using Copied(9) by fastforce ultimately show ?case using Copied(4-5) descendants.Copied[where branch=\(ps' @ ps, a) # branch\] by blast qed lemma descendants_no_head: assumes \descendants k i ((ps, a) # branch) xs\ shows \descendants k i ((p # ps, a) # branch) xs\ using assms descendants_block[where ps'=\[_]\] by simp lemma descendants_types: assumes \descendants k i branch xs\ \(v, v') \ xs\ \branch !. v = Some (ps, a)\ \ps !. v' = Some p\ shows \p = (\<^bold>\ Nom k) \ (\q. p = (\<^bold>\ (\<^bold>@ k q)))\ using assms by (induct k i branch xs arbitrary: v v' ps a) fastforce+ lemma descendants_oob_head': assumes \descendants k i ((ps, a) # branch) xs\ shows \(length branch, m + length ps) \ xs\ using assms descendants_bounds_snd by fastforce lemma descendants_oob_head: assumes \descendants k i ((ps, a) # branch) xs\ shows \(length branch, length ps) \ xs\ using assms descendants_oob_head'[where m=0] by fastforce subsection \Induction\ text \ We induct over an arbitrary set of indices. That way, we can determine in each case whether the extension gets replaced or not by manipulating the set before applying the induction hypothesis. \ lemma ST_bridge': fixes a :: 'b assumes inf: \infinite (UNIV :: 'b set)\ and \n \ (ps, a) # branch\ \descendants k i ((ps, a) # branch) xs\ \Nom j at c in branch\ \Nom k at c in branch\ shows \\ mapi_branch (bridge k j xs) ((ps, a) # branch)\ using assms(2-) proof (induct n \(ps, a) # branch\ arbitrary: ps a branch xs rule: ST.induct) case (Close p i') let ?f = \bridge k j xs\ let ?branch = \mapi_branch ?f ((ps, a) # branch)\ obtain qs where qs: \(qs, i') \. (ps, a) # branch\ \p on (qs, i')\ using Close(1) by blast obtain rs where rs: \(rs, i') \. (ps, a) # branch\ \(\<^bold>\ p) on (rs, i')\ using Close(2) by blast obtain v where v: \(mapi (?f v) qs, i') \. ?branch\ using qs mapi_branch_mem by fast obtain w where w: \(mapi (?f w) rs, i') \. ?branch\ using rs mapi_branch_mem by fast have j: \Nom j at c in ?branch\ using Close(4) nom_at_in_bridge unfolding mapi_branch_def by fastforce have k: \Nom k at c in ?branch\ using Close(5) nom_at_in_bridge unfolding mapi_branch_def by fastforce show ?case proof (cases \\a. p = Nom a\) case True then have \p on (mapi (?f v) qs, i')\ using qs bridge_on_Nom by fast moreover have \(\<^bold>\ p) on (mapi (?f w) rs, i')\ using rs(2) True by (induct rs) auto ultimately show ?thesis using v w ST.Close by fast next case False then obtain v' where \qs !. v' = Some p\ using qs rev_nth_on by fast then have qs': \(?f v v' p) on (mapi (?f v) qs, i')\ using rev_nth_mapi_block by fast then obtain w' where \rs !. w' = Some (\<^bold>\ p)\ using rs rev_nth_on by fast then have rs': \(?f w w' (\<^bold>\ p)) on (mapi (?f w) rs, i')\ using rev_nth_mapi_block by fast obtain q q' where q: \?f v v' p = q\ and q': \?f w w' (\<^bold>\ p) = q'\ by simp_all then consider (dia) \q = (\<^bold>\ Nom j)\ \q' = (\<^bold>\ (\<^bold>\ Nom k))\ | (satn)\\r. q = (\<^bold>\ (\<^bold>@ j r)) \ q' = (\<^bold>\ \<^bold>\ (\<^bold>@ k r))\ | (sat) \\r. q = (\<^bold>@ k r) \ q' = (\<^bold>\ (\<^bold>@ j r))\ | (old) \q = p\ \q' = (\<^bold>\ p)\ using bridge_compl_cases by fast then show ?thesis proof cases case dia then have *: \(\<^bold>\ Nom j) on (mapi (?f v) qs, i')\ \(\<^bold>\ (\<^bold>\ Nom k)) on (mapi (?f w) rs, i')\ using q qs' q' rs' by simp_all have \i' \ branch_nominals ?branch\ unfolding branch_nominals_def using v by fastforce then have ?thesis if \\ ([], i') # ?branch\ using that GoTo by fast moreover have \(mapi (?f v) qs, i') \. ([], i') # ?branch\ using v by simp moreover have \(mapi (?f w) rs, i') \. ([], i') # ?branch\ using w by simp ultimately have ?thesis if \\ ([\<^bold>\ (\<^bold>@ j (Nom k))], i') # ?branch\ using that * by (meson DiaN') moreover have \j \ branch_nominals (([\<^bold>\ (\<^bold>@ j (Nom k))], i') # ?branch)\ unfolding branch_nominals_def by simp ultimately have ?thesis if \\ ([], j) # ([\<^bold>\ (\<^bold>@ j (Nom k))], i') # ?branch\ using that GoTo by fast moreover have \Nom j at c in ([], j) # ([\<^bold>\ (\<^bold>@ j (Nom k))], i') # ?branch\ using j by auto moreover have \Nom k at c in ([], j) # ([\<^bold>\ (\<^bold>@ j (Nom k))], i') # ?branch\ using k by auto ultimately have ?thesis if \\ ([Nom k], j) # ([\<^bold>\ (\<^bold>@ j (Nom k))], i') # ?branch\ using that by (meson Nom' list.set_intros(1) on.simps) moreover have \(\<^bold>\ (\<^bold>@ j (Nom k))) at i' in ([Nom k], j) # ([\<^bold>\ (\<^bold>@ j (Nom k))], i') # ?branch\ by fastforce ultimately have ?thesis if \\ ([\<^bold>\ Nom k, Nom k], j) # ([\<^bold>\ (\<^bold>@ j (Nom k))], i') # ?branch\ using that SatN' by fast moreover have \Nom k at j in ([\<^bold>\ Nom k, Nom k], j) # ([\<^bold>\ (\<^bold>@ j (Nom k))], i') # ?branch\ by fastforce moreover have \(\<^bold>\ Nom k) at j in ([\<^bold>\ Nom k, Nom k], j) # ([\<^bold>\ (\<^bold>@ j (Nom k))], i') # ?branch\ by fastforce ultimately show ?thesis using ST.Close by fast next case satn then obtain r where *: \(\<^bold>\ (\<^bold>@ j r)) on (mapi (?f v) qs, i')\ \(\<^bold>\ \<^bold>\ (\<^bold>@ k r)) on (mapi (?f w) rs, i')\ using q qs' q' rs' by auto have \i' \ branch_nominals ?branch\ unfolding branch_nominals_def using v by fastforce then have ?thesis if \\ ([], i') # ?branch\ using that GoTo by fast moreover have \(mapi (?f w) rs, i') \. ([], i') # ?branch\ using w by simp ultimately have ?thesis if \\ ([\<^bold>@ k r], i') # ?branch\ using that *(2) by (meson Neg') moreover have \j \ branch_nominals (([\<^bold>@ k r], i') # ?branch)\ unfolding branch_nominals_def using j by fastforce ultimately have ?thesis if \\ ([], j) # ([\<^bold>@ k r], i') # ?branch\ using that j GoTo by fast moreover have \Nom j at c in ([], j) # ([\<^bold>@ k r], i') # ?branch\ \Nom k at c in ([], j) # ([\<^bold>@ k r], i') # ?branch\ using j k by auto ultimately have ?thesis if \\ ([Nom k], j) # ([\<^bold>@ k r], i') # ?branch\ using that by (meson Nom' list.set_intros(1) on.simps) then have ?thesis if \\ ([r, Nom k], j) # ([\<^bold>@ k r], i') # ?branch\ using that inf by (meson SatP'' list.set_intros(1) list.set_intros(2) on.simps) moreover have \(mapi (?f v) qs, i') \. ([r, Nom k], j) # ([\<^bold>@ k r], i') # ?branch\ using v by simp ultimately have ?thesis if \\ ([\<^bold>\ r, r, Nom k], j) # ([\<^bold>@ k r], i') # ?branch\ using that *(1) by (meson SatN') moreover have \r at j in ([\<^bold>\ r, r, Nom k], j) # ([\<^bold>@ k r], i') # ?branch\ by fastforce moreover have \(\<^bold>\ r) at j in ([\<^bold>\ r, r, Nom k], j) # ([\<^bold>@ k r], i') # ?branch\ by fastforce ultimately show ?thesis using ST.Close by fast next case sat then obtain r where *: \(\<^bold>@ k r) on (mapi (?f v) qs, i')\ \(\<^bold>\ (\<^bold>@ j r)) on (mapi (?f w) rs, i')\ using q qs' q' rs' by auto have \j \ branch_nominals ?branch\ unfolding branch_nominals_def using j by fastforce then have ?thesis if \\ ([], j) # ?branch\ using that GoTo by fast moreover have \Nom j at c in ([], j) # ?branch\ \Nom k at c in ([], j) # ?branch\ using j k by auto ultimately have ?thesis if \\ ([Nom k], j) # ?branch\ using that by (meson Nom' list.set_intros(1) on.simps) moreover have \(mapi (?f v) qs, i') \. ([Nom k], j) # ?branch\ using v by simp ultimately have ?thesis if \\ ([r, Nom k], j) # ?branch\ using that *(1) inf by (meson SatP'' list.set_intros(1) on.simps) moreover have \(mapi (?f w) rs, i') \. ([r, Nom k], j) # ?branch\ using w by simp ultimately have ?thesis if \\ ([\<^bold>\ r, r, Nom k], j) # ?branch\ using that *(2) by (meson SatN') moreover have \r at j in ([\<^bold>\ r, r, Nom k], j) # ?branch\ by fastforce moreover have \(\<^bold>\ r) at j in ([\<^bold>\ r, r, Nom k], j) # ?branch\ by fastforce ultimately show ?thesis using ST.Close by fast next case old then have \p on (mapi (?f v) qs, i')\ \(\<^bold>\ p) on (mapi (?f w) rs, i')\ using q qs' q' rs' by simp_all then show ?thesis using v w ST.Close[where p=p and i=i'] by fast qed qed next case (Neg p a ps branch) let ?f = \bridge k j xs\ have p: \?f l l' (\<^bold>\ \<^bold>\ p) = (\<^bold>\ \<^bold>\ p)\ for l l' by simp have \descendants k i ((p # ps, a) # branch) xs\ using Neg(5) descendants_no_head by fast then have \\ mapi_branch ?f ((p # ps, a) # branch)\ using Neg(4-) by blast moreover have \(length branch, length ps) \ xs\ using Neg(5) descendants_oob_head by fast ultimately have \\ (p # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \\l l'. ?f l l' (\<^bold>\ \<^bold>\ p) at a in mapi_branch ?f ((ps, a) # branch)\ using Neg(1) at_in_mapi_branch by fast then have \(\<^bold>\ \<^bold>\ p) at a in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def using p by simp ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using Neg' by fast then show ?case unfolding mapi_branch_def by auto next case (DisP p q a ps branch) let ?f = \bridge k j xs\ have p: \?f l l' (p \<^bold>\ q) = (p \<^bold>\ q)\ for l l' by simp have \descendants k i ((p # ps, a) # branch) xs\ \descendants k i ((q # ps, a) # branch) xs\ using DisP(8) descendants_no_head by fast+ then have \\ mapi_branch ?f ((p # ps, a) # branch)\ \\ mapi_branch ?f ((q # ps, a) # branch)\ using DisP(5-) by blast+ moreover have \(length branch, length ps) \ xs\ using DisP(8) descendants_oob_head by fast ultimately have \\ (p # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ \\ (q # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp_all moreover have \\l l'. ?f l l' (p \<^bold>\ q) at a in mapi_branch ?f ((ps, a) # branch)\ using DisP(1) at_in_mapi_branch by fast then have \(p \<^bold>\ q) at a in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def using p by simp ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using DisP'' by fast then show ?case unfolding mapi_branch_def by auto next case (DisN p q a ps branch) let ?f = \bridge k j xs\ have p: \?f l l' (\<^bold>\ (p \<^bold>\ q)) = (\<^bold>\ (p \<^bold>\ q))\ for l l' by simp have \descendants k i (((\<^bold>\ p) # ps, a) # branch) xs\ using DisN(5) descendants_no_head by fast then have \descendants k i (((\<^bold>\ q) # (\<^bold>\ p) # ps, a) # branch) xs\ using descendants_no_head by fast then have \\ mapi_branch ?f (((\<^bold>\ q) # (\<^bold>\ p) # ps, a) # branch)\ using DisN(4-) by blast moreover have \(length branch, length ps) \ xs\ using DisN(5) descendants_oob_head by fast moreover have \(length branch, 1 + length ps) \ xs\ using DisN(5) descendants_oob_head' by fast ultimately have \\ ((\<^bold>\ q) # (\<^bold>\ p) # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \\l l'. ?f l l' (\<^bold>\ (p \<^bold>\ q)) at a in mapi_branch ?f ((ps, a) # branch)\ using DisN(1) at_in_mapi_branch by fast then have \(\<^bold>\ (p \<^bold>\ q)) at a in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def using p by simp ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using DisN' by fast then show ?case unfolding mapi_branch_def by auto next case (DiaP p a ps branch i') let ?f = \bridge k j xs\ have p: \?f l l' (\<^bold>\ p) = (\<^bold>\ p)\ for l l' using DiaP(3) bridge_proper_dia by fast have \branch_nominals (mapi_branch ?f ((ps, a) # branch)) = branch_nominals ((ps, a) # branch)\ using DiaP(8-) nominals_mapi_branch_bridge[where j=j and k=k and branch=\(ps, a) # branch\] by auto then have i': \i' \ branch_nominals ((mapi (?f (length branch)) ps, a) # mapi_branch ?f branch)\ unfolding mapi_branch_def using DiaP(2) by simp have 1: \?f (length branch) (1 + length ps) (\<^bold>@ i' p) = (\<^bold>@ i' p)\ by simp have \i' \ k\ using DiaP(2, 8-) unfolding branch_nominals_def by fastforce then have 2: \?f (length branch) (length ps) (\<^bold>\ Nom i') = (\<^bold>\ Nom i')\ by simp have \descendants k i (((\<^bold>@ i' p) # (\<^bold>\ Nom i') # ps, a) # branch) xs\ using DiaP(7) descendants_block[where ps'=\[_, _]\] by fastforce then have \\ mapi_branch ?f (((\<^bold>@ i' p) # (\<^bold>\ Nom i') # ps, a) # branch)\ using DiaP(4-) by blast then have \\ ((\<^bold>@ i' p) # (\<^bold>\ Nom i') # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def using 1 by (simp add: 2) moreover have \\l l'. ?f l l' (\<^bold>\ p) at a in mapi_branch ?f ((ps, a) # branch)\ using DiaP(1) at_in_mapi_branch by fast then have \(\<^bold>\ p) at a in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def using p by simp ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using i' DiaP(3) DiaP' by fast then show ?case unfolding mapi_branch_def by simp next case (DiaN p a ps branch i') have p: \bridge k j xs l l' (\<^bold>\ (\<^bold>\ p)) = (\<^bold>\ (\<^bold>\ p))\ for xs l l' by simp obtain rs where rs: \(rs, a) \. (ps, a) # branch\ \(\<^bold>\ Nom i') on (rs, a)\ using DiaN(2) by fast obtain v where v: \((ps, a) # branch) !. v = Some (rs, a)\ using rs(1) rev_nth_mem by fast obtain v' where v': \rs !. v' = Some (\<^bold>\ Nom i')\ using rs(2) rev_nth_on by fast show ?case proof (cases \(v, v') \ xs\) case True then have \i' = k\ using DiaN(6) v v' descendants_types by fast let ?xs = \{(length branch, length ps)} \ xs\ let ?f = \bridge k j ?xs\ let ?branch = \((\<^bold>\ (\<^bold>@ i' p)) # ps, a) # branch\ obtain rs' where \(((\<^bold>\ (\<^bold>@ k p)) # ps, a) # branch) !. v = Some (rs', a)\ \rs' !. v' = Some (\<^bold>\ Nom i')\ using v v' index_Cons by fast moreover have \descendants k i (((\<^bold>\ (\<^bold>@ k p)) # ps, a) # branch) xs\ using DiaN(6) descendants_block[where ps'=\[_]\] by fastforce moreover have \?branch !. length branch = Some ((\<^bold>\ (\<^bold>@ k p)) # ps, a)\ using \i' = k\ by simp moreover have \((\<^bold>\ (\<^bold>@ k p)) # ps) !. length ps = Some (\<^bold>\ (\<^bold>@ k p))\ by simp ultimately have \descendants k i (((\<^bold>\ (\<^bold>@ k p)) # ps, a) # branch) ?xs\ using True \i' = k\ Derived[where branch=\_ # branch\] by simp then have \\ mapi_branch ?f (((\<^bold>\ (\<^bold>@ k p)) # ps, a) # branch)\ using \i' = k\ DiaN(5-) by blast then have \\ ((\<^bold>\ (\<^bold>@ j p)) # mapi (?f (length branch)) ps, a) # mapi_branch (bridge k j ?xs) branch\ unfolding mapi_branch_def using \i' = k\ by simp moreover have \\l l'. ?f l l' (\<^bold>\ (\<^bold>\ p)) at a in mapi_branch ?f ((ps, a) # branch)\ using DiaN(1) at_in_mapi_branch by fast then have \(\<^bold>\ (\<^bold>\ p)) at a in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def using p[where xs=\?xs\] by simp moreover have \(mapi (?f v) rs, a) \. mapi_branch ?f ((ps, a) # branch)\ using v rev_nth_mapi_branch by fast then have \(mapi (?f v) rs, a) \ set ((mapi (?f (length branch)) ps, a) # mapi_branch ?f branch)\ unfolding mapi_branch_def by simp moreover have \?f v v' (\<^bold>\ Nom i') on (mapi (?f v) rs, a)\ using v' rev_nth_mapi_block by fast then have \(\<^bold>\ Nom j) on (mapi (?f v) rs, a)\ using True \i' = k\ by simp ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ by (meson DiaN') then have \\ (mapi (bridge k j xs (length branch)) ps, a) # mapi_branch (bridge k j xs) branch\ using mapi_branch_head_add_oob[where branch=branch and ps=ps] unfolding mapi_branch_def by simp then show ?thesis unfolding mapi_branch_def by simp next case False let ?f = \bridge k j xs\ have \descendants k i (((\<^bold>\ (\<^bold>@ i' p)) # ps, a) # branch) xs\ using DiaN(6) descendants_no_head by fast then have \\ mapi_branch ?f (((\<^bold>\ (\<^bold>@ i' p)) # ps, a) # branch)\ using DiaN(5-) by blast moreover have \(length branch, length ps) \ xs\ using DiaN(6) descendants_oob_head by fast ultimately have \\ ((\<^bold>\ (\<^bold>@ i' p)) # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \\l l'. ?f l l' (\<^bold>\ (\<^bold>\ p)) at a in mapi_branch ?f ((ps, a) # branch)\ using DiaN(1) at_in_mapi_branch by fast then have \(\<^bold>\ (\<^bold>\ p)) at a in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def using p[where xs=\xs\] by simp moreover have \(mapi (?f v) rs, a) \. mapi_branch ?f ((ps, a) # branch)\ using v rev_nth_mapi_branch by fast then have \(mapi (?f v) rs, a) \ set ((mapi (?f (length branch)) ps, a) # mapi_branch ?f branch)\ unfolding mapi_branch_def by simp moreover have \?f v v' (\<^bold>\ Nom i') on (mapi (?f v) rs, a)\ using v' rev_nth_mapi_block by fast then have \(\<^bold>\ Nom i') on (mapi (?f v) rs, a)\ using False by simp ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ by (meson DiaN') then show ?thesis unfolding mapi_branch_def by simp qed next case (SatP a p b ps branch) let ?f = \bridge k j xs\ have p: \?f l l' (\<^bold>@ a p) = (\<^bold>@ a p)\ for l l' by simp have \descendants k i ((p # ps, a) # branch) xs\ using SatP(5) descendants_no_head by fast then have \\ mapi_branch ?f ((p # ps, a) # branch)\ using SatP(4-) by blast moreover have \(length branch, length ps) \ xs\ using SatP(5) descendants_oob_head by fast ultimately have \\ (p # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \\l l'. ?f l l' (\<^bold>@ a p) at b in mapi_branch ?f ((ps, a) # branch)\ using SatP(1) at_in_mapi_branch by fast then have \(\<^bold>@ a p) at b in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def using p by simp ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using SatP' by fast then show ?case unfolding mapi_branch_def by simp next case (SatN a p b ps branch) obtain qs where qs: \(qs, b) \. (ps, a) # branch\ \(\<^bold>\ (\<^bold>@ a p)) on (qs, b)\ using SatN(1) by fast obtain v where v: \((ps, a) # branch) !. v = Some (qs, b)\ using qs(1) rev_nth_mem by fast obtain v' where v': \qs !. v' = Some (\<^bold>\ (\<^bold>@ a p))\ using qs(2) rev_nth_on by fast show ?case proof (cases \(v, v') \ xs\) case True then have \a = k\ using SatN(5) v v' descendants_types by fast let ?f = \bridge k j xs\ let ?branch = \((\<^bold>\ p) # ps, a) # branch\ have p: \?f v v' (\<^bold>\ (\<^bold>@ k p)) = (\<^bold>\ (\<^bold>@ j p))\ using True by simp obtain rs' where \?branch !. v = Some (rs', b)\ \rs' !. v' = Some (\<^bold>\ (\<^bold>@ k p))\ using v v' \a = k\ index_Cons by fast have \descendants k i ?branch xs\ using SatN(5) descendants_no_head by fast then have \\ mapi_branch ?f ?branch\ using \a = k\ SatN(4-) by blast moreover have \(length branch, length ps) \ xs\ using SatN(5) descendants_oob_head by fast ultimately have \\ ((\<^bold>\ p) # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def using \a = k\ by simp then have *: \\ ((\<^bold>\ p) # Nom j # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using inf ST_struct_block[where ps'=\(\<^bold>\ p) # Nom j # _\] by fastforce have \(mapi (?f v) qs, b) \. mapi_branch ?f ((ps, a) # branch)\ using v rev_nth_mapi_branch by fast then have \(mapi (?f v) qs, b) \ set ((mapi (?f (length branch)) ps, a) # mapi_branch ?f branch)\ unfolding mapi_branch_def by simp moreover have \?f v v' (\<^bold>\ (\<^bold>@ k p)) on (mapi (?f v) qs, b)\ using v' \a = k\ rev_nth_mapi_block by fast then have \(\<^bold>\ (\<^bold>@ j p)) on (mapi (?f v) qs, b)\ using p by simp ultimately obtain qs' where \(qs', b) \. (Nom j # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ \(\<^bold>\ (\<^bold>@ j p)) on (qs', b)\ by fastforce then have \\ (Nom j # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using * inf by (meson SatN'' list.set_intros(1) on.simps) moreover have \Nom j at c in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using SatN(6) nom_at_in_bridge unfolding mapi_branch_def by fastforce moreover have \Nom k at c in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using SatN(7) nom_at_in_bridge unfolding mapi_branch_def by fastforce ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using \a = k\ by (meson Nom' list.set_intros(1) on.simps) then show ?thesis unfolding mapi_branch_def by simp next case False let ?f = \bridge k j xs\ have \descendants k i (((\<^bold>\ p) # ps, a) # branch) xs\ using SatN(5) descendants_no_head by fast then have \\ mapi_branch (bridge k j xs) (((\<^bold>\ p) # ps, a) # branch)\ using SatN(4-) by blast moreover have \(length branch, length ps) \ xs\ using SatN(5) descendants_oob_head by fast ultimately have \\ ((\<^bold>\ p) # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \(mapi (?f v) qs, b) \. mapi_branch ?f ((ps, a) # branch)\ using v rev_nth_mapi_branch by fast then have \(mapi (?f v) qs, b) \ set ((mapi (?f (length branch)) ps, a) # mapi_branch ?f branch)\ unfolding mapi_branch_def by simp moreover have \?f v v' (\<^bold>\ (\<^bold>@ a p)) on (mapi (?f v) qs, b)\ using v' rev_nth_mapi_block by fast then have \(\<^bold>\ (\<^bold>@ a p)) on (mapi (?f v) qs, b)\ using False by simp ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ by (meson SatN') then show ?thesis unfolding mapi_branch_def by simp qed next case (GoTo i' n ps a branch) let ?f = \bridge k j xs\ have \descendants k i (([], i') # (ps, a) # branch) xs\ using GoTo(4) descendants_branch[where extra=\[_]\] by simp then have \\ mapi_branch ?f (([], i') # (ps, a) # branch)\ using GoTo(3, 5-) by auto then have \\ ([], i') # mapi_branch ?f ((ps, a) # branch)\ unfolding mapi_branch_def by simp moreover have \branch_nominals (mapi_branch ?f ((ps, a) # branch)) = branch_nominals ((ps, a) # branch)\ using GoTo(5-) nominals_mapi_branch_bridge[where j=j and k=k and branch=\(ps, a) # branch\] by auto then have \i' \ branch_nominals (mapi_branch (bridge k j xs) ((ps, a) # branch))\ using GoTo(1) by blast ultimately show ?case using ST.GoTo by fast next case (Nom p b ps a branch i') show ?case proof (cases \\j. p = Nom j\) case True let ?f = \bridge k j xs\ have \descendants k i ((p # ps, a) # branch) xs\ using Nom(7) descendants_block[where ps'=\[p]\] by simp then have \\ mapi_branch ?f ((p # ps, a) # branch)\ using Nom(6-) by blast moreover have \?f (length branch) (length ps) p = p\ using True by auto ultimately have \\ (p # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \p at b in mapi_branch ?f ((ps, a) # branch)\ using Nom(1) True nom_at_in_bridge by fast then have \p at b in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \Nom i' at b in mapi_branch ?f ((ps, a) # branch)\ using Nom(2) True nom_at_in_bridge by fast then have \Nom i' at b in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \Nom i' at a in mapi_branch ?f ((ps, a) # branch)\ using Nom(3) True nom_at_in_bridge by fast then have \Nom i' at a in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using Nom' by fast then show ?thesis unfolding mapi_branch_def by simp next case False obtain qs where qs: \(qs, b) \. (ps, a) # branch\ \p on (qs, b)\ using Nom(1) by blast obtain v where v: \((ps, a) # branch) !. v = Some (qs, b)\ using qs(1) rev_nth_mem by fast obtain v' where v': \qs !. v' = Some p\ using qs(2) False rev_nth_on by fast show ?thesis proof (cases \(v, v') \ xs\) case True let ?xs = \{(length branch, length ps)} \ xs\ let ?f = \bridge k j ?xs\ let ?p = \bridge' k j p\ have p: \?f v v' p = ?p\ using True by simp obtain qs' where \((p # ps, a) # branch) !. v = Some (qs', b)\ \qs' !. v' = Some p\ using v v' index_Cons by fast moreover have \Nom i' at b in (p # ps, a) # branch\ using Nom(2) by fastforce moreover have \descendants k i ((p # ps, a) # branch) xs\ using Nom(7) descendants_block[where ps'=\[p]\] by simp moreover have \((p # ps, a) # branch) !. length branch = Some (p # ps, a)\ \(p # ps) !. length ps = Some p\ by simp_all moreover have \Nom i' at a in (p # ps, a) # branch\ using Nom(3) by fastforce ultimately have \descendants k i ((p # ps, a) # branch) ?xs\ using True Copied by fast then have \\ mapi_branch ?f ((p # ps, a) # branch)\ using Nom(6-) by blast then have \\ (?p # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \(mapi (?f v) qs, b) \. mapi_branch ?f ((ps, a) # branch)\ using v rev_nth_mapi_branch by fast then have \(mapi (?f v) qs, b) \. (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \?f v v' p on (mapi (?f v) qs, b)\ using v v' rev_nth_mapi_block by fast then have \?p on (mapi (?f v) qs, b)\ using p by simp moreover have \Nom i' at b in mapi_branch ?f ((ps, a) # branch)\ using Nom(2) nom_at_in_bridge by fast then have \Nom i' at b in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \Nom i' at a in mapi_branch ?f ((ps, a) # branch)\ using Nom(3) nom_at_in_bridge by fast then have \Nom i' at a in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using Nom' by fast then have \\ (mapi (bridge k j xs (length branch)) ps, a) # mapi_branch (bridge k j xs) branch\ using mapi_branch_head_add_oob[where branch=branch and ps=ps] unfolding mapi_branch_def by simp then show ?thesis unfolding mapi_branch_def by simp next case False let ?f = \bridge k j xs\ have \descendants k i ((p # ps, a) # branch) xs\ using Nom(7) descendants_no_head by fast then have \\ mapi_branch ?f ((p # ps, a) # branch)\ using Nom(6-) by blast moreover have \(length branch, length ps) \ xs\ using Nom(7) descendants_oob_head by fast ultimately have \\ (p # mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \(mapi (?f v) qs, b) \. mapi_branch ?f ((ps, a) # branch)\ using v rev_nth_mapi_branch by fast then have \(mapi (?f v) qs, b) \. (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \?f v v' p on (mapi (?f v) qs, b)\ using v v' rev_nth_mapi_block by fast then have \p on (mapi (?f v) qs, b)\ using False by simp moreover have \Nom i' at b in mapi_branch ?f ((ps, a) # branch)\ using Nom(2) nom_at_in_bridge by fast then have \Nom i' at b in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp moreover have \Nom i' at a in mapi_branch ?f ((ps, a) # branch)\ using Nom(3) nom_at_in_bridge by fast then have \Nom i' at a in (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ unfolding mapi_branch_def by simp ultimately have \\ (mapi (?f (length branch)) ps, a) # mapi_branch ?f branch\ using Nom' by fast then show ?thesis unfolding mapi_branch_def by simp qed qed qed lemma ST_bridge: fixes i :: 'b assumes inf: \infinite (UNIV :: 'b set)\ and \\ branch\ \descendants k i branch xs\ \Nom j at c in branch\ \Nom k at c in branch\ shows \\ mapi_branch (bridge k j xs) branch\ proof - have \\ ([], j) # branch\ using inf assms(2) ST_struct[where branch'=\([], j) # branch\] by auto moreover have \descendants k i (([], j) # branch) xs\ using assms(3) descendants_branch[where extra=\[_]\] by fastforce ultimately have \\ mapi_branch (bridge k j xs) (([], j) # branch)\ using ST_bridge' inf assms(3-) by fast then have *: \\ ([], j) # mapi_branch (bridge k j xs) branch\ unfolding mapi_branch_def by simp have \branch_nominals (mapi_branch (bridge k j xs) branch) = branch_nominals branch\ using nominals_mapi_branch_bridge assms(4-) by fast moreover have \j \ branch_nominals branch\ using assms(4) unfolding branch_nominals_def by fastforce ultimately have \j \ branch_nominals (mapi_branch (bridge k j xs) branch)\ by simp then show ?thesis using * GoTo by fast qed subsection \Derivation\ theorem Bridge: fixes i :: 'b assumes inf: \infinite (UNIV :: 'b set)\ and \Nom i at b in branch\ \(\<^bold>\ Nom j) at b in branch\ \Nom i at a in branch\ \Nom j at c in branch\ \Nom k at c in branch\ \\ ((\<^bold>\ Nom k) # ps, a) # branch\ shows \\ (ps, a) # branch\ proof - let ?xs = \{(length branch, length ps)}\ have \descendants k a (((\<^bold>\ Nom k) # ps, a) # branch) ?xs\ using Initial by force moreover have \Nom j at c in ((\<^bold>\ Nom k) # ps, a) # branch\ \Nom k at c in ((\<^bold>\ Nom k) # ps, a) # branch\ using assms(5-6) by auto ultimately have \\ mapi_branch (bridge k j ?xs) (((\<^bold>\ Nom k) # ps, a) # branch)\ using ST_bridge inf assms(7) by fast then have \\ ((\<^bold>\ Nom j) # mapi (bridge k j ?xs (length branch)) ps, a) # mapi_branch (bridge k j ?xs) branch\ unfolding mapi_branch_def by simp moreover have \mapi_branch (bridge k j {(length branch, length ps)}) branch = mapi_branch (bridge k j {}) branch\ using mapi_branch_add_oob[where xs=\{}\] by fastforce moreover have \mapi (bridge k j ?xs (length branch)) ps = mapi (bridge k j {} (length branch)) ps\ using mapi_block_add_oob[where xs=\{}\ and ps=ps] by simp ultimately have \\ ((\<^bold>\ Nom j) # ps, a) # branch\ using mapi_block_id[where ps=ps] mapi_branch_id[where branch=branch] by simp then show ?thesis using assms(2-5) by (meson Nom' set_subset_Cons subsetD) qed section \Completeness\ subsection \Hintikka\ abbreviation at_in' :: \('a, 'b) fm \ 'b \ ('a, 'b) block set \ bool\ - (\_ at _ in' _\ [51, 51, 51] 50) where + (\_ at _ in'' _\ [51, 51, 51] 50) where \p at a in' branch \ \ps. (ps, a) \ branch \ p on (ps, a)\ text \ A set of blocks is Hintikka if it satisfies the following requirements. Intuitively, if it corresponds to an exhausted open branch. \ definition hintikka :: \('a, 'b) block set \ bool\ where \hintikka H \ (\x i j. Nom j at i in' H \ Pro x at j in' H \ \ (\<^bold>\ Pro x) at i in' H) \ (\a i. Nom a at i in' H \ \ (\<^bold>\ Nom a) at i in' H) \ (\i j. (\<^bold>\ Nom j) at i in' H \ \ (\<^bold>\ (\<^bold>\ Nom j)) at i in' H) \ (\p i. i \ nominals p \ (\block \ H. p on block) \ (\ps. (ps, i) \ H)) \ (\i j. Nom j at i in' H \ Nom i at j in' H) \ (\i j k. Nom j at i in' H \ Nom k at j in' H \ Nom k at i in' H) \ (\i j k. (\<^bold>\ Nom j) at i in' H \ Nom k at j in' H \ (\<^bold>\ Nom k) at i in' H) \ (\i j k. (\<^bold>\ Nom j) at i in' H \ Nom k at i in' H \ (\<^bold>\ Nom j) at k in' H) \ (\p q i. (p \<^bold>\ q) at i in' H \ p at i in' H \ q at i in' H) \ (\p q i. (\<^bold>\ (p \<^bold>\ q)) at i in' H \ (\<^bold>\ p) at i in' H \ (\<^bold>\ q) at i in' H) \ (\p i. (\<^bold>\ \<^bold>\ p) at i in' H \ p at i in' H) \ (\p i a. (\<^bold>@ i p) at a in' H \ p at i in' H) \ (\p i a. (\<^bold>\ (\<^bold>@ i p)) at a in' H \ (\<^bold>\ p) at i in' H) \ (\p i. (\a. p = Nom a) \ (\<^bold>\ p) at i in' H \ (\j. (\<^bold>\ Nom j) at i in' H \ (\<^bold>@ j p) at i in' H)) \ (\p i j. (\<^bold>\ (\<^bold>\ p)) at i in' H \ (\<^bold>\ Nom j) at i in' H \ (\<^bold>\ (\<^bold>@ j p)) at i in' H)\ text \ Two nominals \i\ and \j\ are equivalent in respect to a Hintikka set \H\ if \H\ contains an \i\-block with \j\ on it. This is shown to be an equivalence relation. \ definition hequiv :: \('a, 'b) block set \ 'b \ 'b \ bool\ where \hequiv H i j \ Nom j at i in' H\ abbreviation hequiv_rel :: \('a, 'b) block set \ ('b \ 'b) set\ where \hequiv_rel H \ {(i, j) |i j. hequiv H i j}\ definition names :: \('a, 'b) block set \ 'b set\ where \names H \ {i |ps i. (ps, i) \ H}\ lemma hequiv_refl: \hintikka H \ i \ names H \ hequiv H i i\ unfolding hintikka_def hequiv_def names_def by auto lemma hequiv_refl': \hintikka H \ (ps, i) \ H \ hequiv H i i\ using hequiv_refl unfolding names_def by fast lemma hequiv_sym: \hintikka H \ hequiv H i j = hequiv H j i\ unfolding hintikka_def hequiv_def by meson lemma hequiv_trans: \hintikka H \ hequiv H i j \ hequiv H j k \ hequiv H i k\ unfolding hintikka_def hequiv_def by meson lemma hequiv_names: \hequiv H i j \ i \ names H\ unfolding hequiv_def names_def by blast lemma hequiv_names_rel: \hintikka H \ hequiv_rel H \ names H \ names H\ using hequiv_sym hequiv_names by fast lemma hequiv_refl_rel: \hintikka H \ refl_on (names H) (hequiv_rel H)\ unfolding refl_on_def using hequiv_refl hequiv_names_rel by fast lemma hequiv_sym_rel: \hintikka H \ sym (hequiv_rel H)\ unfolding sym_def using hequiv_sym by fast lemma hequiv_trans_rel: \hintikka H \ trans (hequiv_rel H)\ unfolding trans_def using hequiv_trans by fast lemma hequiv_rel: \hintikka H \ equiv (names H) (hequiv_rel H)\ using hequiv_refl_rel hequiv_sym_rel hequiv_trans_rel by (rule equivI) subsubsection \Named model\ text \ Given a Hintikka set \H\ and a formula \p\ on a block in \H\ we will construct a model that satisfies \p\. The worlds of our model are sets of equivalent nominals and nominals are assigned to their equivalence class. From a world \is\, we can reach a world \js\ iff there is an \i \ is\ and a \j \ js\ s.t. there is an \i\-block in \H\ with \\<^bold>\ Nom j\ on it. A propositional symbol \p\ is true in a world \is\ if there exists an \i \ is\ s.t. \p\ occurs on an \i\-block in \H\. \ definition assign :: \('a, 'b) block set \ 'b \ 'b set\ where \assign H i \ proj (hequiv_rel H) i\ definition reach :: \('a, 'b) block set \ 'b set \ 'b set set\ where \reach H is \ {assign H j |i j. i \ is \ (\<^bold>\ Nom j) at i in' H}\ definition val :: \('a, 'b) block set \ 'b set \ 'a \ bool\ where \val H is x \ \i \ is. Pro x at i in' H\ lemma hequiv_assign: \hintikka H \ hequiv H i j \ assign H i = assign H j\ unfolding proj_def assign_def using equiv_class_eq hequiv_rel by fast lemma hequiv_model: assumes \hintikka H\ \hequiv H i j\ shows \(Model (reach H) (val H), assign H, assign H i \ p) = (Model (reach H) (val H), assign H, assign H j \ p)\ using assms hequiv_assign by fastforce lemma assign_refl: \hintikka H \ i \ names H \ i \ assign H i\ unfolding proj_def assign_def using hequiv_refl by fastforce lemma assign_refl': \hintikka H \ (ps, i) \ H \ i \ assign H i\ using assign_refl unfolding names_def by fast lemma assign_named: \hintikka H \ i \ names H \ \j \ assign H i. \ps. (ps, j) \ H\ using hequiv_sym unfolding proj_def assign_def hequiv_def by fast lemma nominal_in_names: assumes \hintikka H\ \\block \ H. i \ block_nominals block\ shows \i \ names H\ proof - have \(\p. i \ nominals p \ (\block \ H. p on block) \ (\br. (br, i) \ H))\ using \hintikka H\ unfolding hintikka_def by meson then show ?thesis unfolding names_def using assms(2) by fastforce qed lemma assign_sym: \hintikka H \ j \ assign H i \ i \ assign H j\ unfolding proj_def assign_def using hequiv_sym by fast lemma hintikka_model: assumes \hintikka H\ shows \p at i in' H \ Model (reach H) (val H), assign H, assign H i \ p\ \(\<^bold>\ p) at i in' H \ \ Model (reach H) (val H), assign H, assign H i \ p\ proof (induct p arbitrary: i) fix i case (Pro x) assume \Pro x at i in' H\ then show \Model (reach H) (val H), assign H, assign H i \ Pro x\ using assms(1) assign_refl' unfolding val_def by fastforce next fix i case (Pro x) assume \(\<^bold>\ Pro x) at i in' H\ then have \\ Pro x at j in' H\ if \Nom j at i in' H\ for j using that assms(1) unfolding hintikka_def by meson then have \\ Pro x at j in' H\ if \hequiv H i j\ for j using that unfolding hequiv_def by simp then have \\ val H (assign H i) x\ unfolding proj_def val_def assign_def by blast then show \\ Model (reach H) (val H), assign H, assign H i \ Pro x\ by simp next fix i case (Nom a) assume \Nom a at i in' H\ then have \assign H a = assign H i\ using assms(1) hequiv_assign hequiv_sym unfolding hequiv_def by fast then show \Model (reach H) (val H), assign H, assign H i \ Nom a\ by simp next fix i case (Nom a) assume \(\<^bold>\ Nom a) at i in' H\ then have \\ Nom a at i in' H\ using assms(1) unfolding hintikka_def by meson then have \\ hequiv H i a\ unfolding hequiv_def by blast then have \\ hequiv H a i\ using assms(1) hequiv_sym by fast moreover have \hequiv H i i\ using assms(1) \(\<^bold>\ Nom a) at i in' H\ hequiv_refl' by fast ultimately have \assign H a \ assign H i\ unfolding proj_def assign_def by blast then show \\ Model (reach H) (val H), assign H, assign H i \ Nom a\ by simp next fix i case (Neg p) moreover assume \(\<^bold>\ p) at i in' H\ ultimately show \Model (reach H) (val H), assign H, assign H i \ \<^bold>\ p\ using assms(1) by simp next fix i case (Neg p) moreover assume \(\<^bold>\ \<^bold>\ p) at i in' H\ then have \p at i in' H\ using assms(1) unfolding hintikka_def by meson ultimately show \\ Model (reach H) (val H), assign H, assign H i \ \<^bold>\ p\ using assms(1) by auto next fix i case (Dis p q) moreover assume \(p \<^bold>\ q) at i in' H\ then have \p at i in' H \ q at i in' H\ using assms(1) unfolding hintikka_def by meson ultimately show \Model (reach H) (val H), assign H, assign H i \ (p \<^bold>\ q)\ by (meson semantics.simps(4)) next fix i case (Dis p q) moreover assume \(\<^bold>\ (p \<^bold>\ q)) at i in' H\ then have \(\<^bold>\ p) at i in' H\ \(\<^bold>\ q) at i in' H\ using assms(1) unfolding hintikka_def by meson+ ultimately show \\ Model (reach H) (val H), assign H, assign H i \ (p \<^bold>\ q)\ by auto next fix i case (Dia p) moreover assume \(\<^bold>\ p) at i in' H\ ultimately show \Model (reach H) (val H), assign H, assign H i \ \<^bold>\ p\ proof (cases \\j. p = Nom j\) case True then obtain j where \p = Nom j\ by blast have \i \ assign H i\ using assms(1) \(\<^bold>\ p) at i in' H\ assign_refl' by fast moreover have \j \ nominals (\<^bold>\ p)\ using \p = Nom j\ by simp then have \(\block \ H. (\<^bold>\ p) on block) \ (\qs. (qs, j) \ H)\ using assms(1) unfolding hintikka_def by meson then have \\bl. (bl, j) \ H\ using \(\<^bold>\ p) at i in' H\ by blast then have \j \ assign H j\ using assms(1) \(\<^bold>\ p) at i in' H\ assign_refl' by fast moreover have \(\<^bold>\ Nom j) at i in' H\ using \p = Nom j\ \(\<^bold>\ p) at i in' H\ by blast ultimately have \assign H j \ reach H (assign H i)\ using \(\<^bold>\ p) at i in' H\ unfolding reach_def by auto then show ?thesis using \p = Nom j\ by simp next case False then have \\j. (\<^bold>\ Nom j) at i in' H \ (\<^bold>@ j p) at i in' H\ using assms \(\<^bold>\ p) at i in' H\ unfolding hintikka_def by blast then obtain j where *: \(\<^bold>\ Nom j) at i in' H\ \(\<^bold>@ j p) at i in' H\ by blast from *(2) have \p at j in' H\ using assms(1) unfolding hintikka_def by blast then have \Model (reach H) (val H), assign H, assign H j \ p\ using Dia by blast have \i \ assign H i\ using assms(1) assign_refl' \(\<^bold>\ p) at i in' H\ by fast moreover have \j \ assign H j\ using assms(1) assign_refl' \p at j in' H\ by fast ultimately have \assign H j \ reach H (assign H i)\ using *(1) unfolding reach_def by auto then show ?thesis using \Model (reach H) (val H), assign H, assign H j \ p\ by auto qed next fix i case (Dia p) assume \(\<^bold>\ (\<^bold>\ p)) at i in' H\ show \\ Model (reach H) (val H), assign H, assign H i \ \<^bold>\ p\ proof assume \Model (reach H) (val H), assign H, assign H i \ \<^bold>\ p\ then obtain i' j where \Model (reach H) (val H), assign H, assign H j \ p\ \i' \ assign H i\ \(\<^bold>\ Nom j) at i' in' H\ unfolding reach_def by auto have \Nom i' at i in' H\ using \i' \ assign H i\ \(\<^bold>\ Nom j) at i' in' H\ unfolding hequiv_def proj_def assign_def by auto then have \Nom i at i' in' H\ using assms(1) unfolding hintikka_def by meson then have \(\<^bold>\ Nom j) at i in' H\ using assms(1) \(\<^bold>\ Nom j) at i' in' H\ unfolding hintikka_def by meson then have \(\<^bold>\ (\<^bold>@ j p)) at i in' H\ using assms(1) \(\<^bold>\ (\<^bold>\ p)) at i in' H\ unfolding hintikka_def by meson moreover have \(\<^bold>\ p) at j in' H\ if \\a. (\<^bold>\ (\<^bold>@ j p)) at a in' H\ using that assms(1) unfolding hintikka_def by blast ultimately have \(\<^bold>\ p) at j in' H\ by blast then have \\ Model (reach H) (val H), assign H, assign H j \ p\ using Dia by blast then show False using \Model (reach H) (val H), assign H, assign H j \ p\ by blast qed next fix i case (Sat j p) assume \(\<^bold>@ j p) at i in' H\ moreover have \p at j in' H\ if \\a. (\<^bold>@ j p) at a in' H\ using that assms(1) unfolding hintikka_def by meson ultimately have \p at j in' H\ by blast then show \Model (reach H) (val H), assign H, assign H i \ \<^bold>@ j p\ using Sat by simp next fix i case (Sat j p) assume \(\<^bold>\ (\<^bold>@ j p)) at i in' H\ moreover have \(\<^bold>\ p) at j in' H\ if \\a. (\<^bold>\ (\<^bold>@ j p)) at a in' H\ using that assms(1) unfolding hintikka_def by meson ultimately have \(\<^bold>\ p) at j in' H\ by blast then show \\ Model (reach H) (val H), assign H, assign H i \ \<^bold>@ j p\ using Sat by simp qed subsection \Lindenbaum-Henkin\ text \ A set of blocks is consistent if no finite subset can be derived. Given a consistent set of blocks we are going to extend it to be saturated and maximally consistent and show that is then Hintikka. \ definition consistent :: \('a, 'b) block set \ bool\ where \consistent S \ \S'. set S' \ S \ \ S'\ instance fm :: (countable, countable) countable by countable_datatype definition proper_dia :: \('a, 'b) fm \ ('a, 'b) fm option\ where \proper_dia p \ case p of (\<^bold>\ p) \ (if \a. p = Nom a then Some p else None) | _ \ None\ lemma proper_dia: \proper_dia p = Some q \ p = (\<^bold>\ q) \ (\a. q = Nom a)\ unfolding proper_dia_def by (cases p) (simp_all, metis option.discI option.inject) text \The following function witnesses each \\<^bold>\ p\ in a fresh world.\ primrec witness_list :: \('a, 'b) fm list \ 'b set \ ('a, 'b) fm list\ where \witness_list [] _ = []\ | \witness_list (p # ps) used = (case proper_dia p of None \ witness_list ps used | Some q \ let i = SOME i. i \ used in (\<^bold>@ i q) # (\<^bold>\ Nom i) # witness_list ps ({i} \ used))\ primrec witness :: \('a, 'b) block \ 'b set \ ('a, 'b) block\ where \witness (ps, a) used = (witness_list ps used, a)\ lemma witness_list: \proper_dia p = Some q \ witness_list (p # ps) used = (let i = SOME i. i \ used in (\<^bold>@ i q) # (\<^bold>\ Nom i) # witness_list ps ({i} \ used))\ by simp primrec extend :: \('a, 'b) block set \ (nat \ ('a, 'b) block) \ nat \ ('a, 'b) block set\ where \extend S f 0 = S\ | \extend S f (Suc n) = (if \ consistent ({f n} \ extend S f n) then extend S f n else let used = (\block \ {f n} \ extend S f n. block_nominals block) in {f n, witness (f n) used} \ extend S f n)\ definition Extend :: \('a, 'b) block set \ (nat \ ('a, 'b) block) \ ('a, 'b) block set\ where \Extend S f \ (\n. extend S f n)\ lemma extend_chain: \extend S f n \ extend S f (Suc n)\ by auto lemma extend_mem: \S \ extend S f n\ by (induct n) auto lemma Extend_mem: \S \ Extend S f\ unfolding Extend_def using extend_mem by blast subsubsection \Consistency\ lemma split_list: \set A \ {x} \ X \ x \. A \ \B. set (x # B) = set A \ x \ set B\ by simp (metis Diff_insert_absorb mk_disjoint_insert set_removeAll) lemma consistent_drop_single: fixes a :: 'b assumes inf: \infinite (UNIV :: 'b set)\ and cons: \consistent ({(p # ps, a)} \ S)\ shows \consistent ({(ps, a)} \ S)\ unfolding consistent_def proof assume \\S'. set S' \ {(ps, a)} \ S \ (\n. n \ S')\ then obtain S' n where \set S' \ {(ps, a)} \ S\ \(ps, a) \. S'\ \n \ S'\ using assms unfolding consistent_def by blast then obtain S'' where \set ((ps, a) # S'') = set S'\ \(ps, a) \ set S''\ using split_list by metis then have \\ (ps, a) # S''\ using inf ST_struct \n \ S'\ by blast then have \\ (p # ps, a) # S''\ using inf ST_struct_block[where ps'=\p # ps\] by fastforce moreover have \set ((p # ps, a) # S'') \ {(p # ps, a)} \ S\ using \(ps, a) \ set S''\ \set ((ps, a) # S'') = set S'\ \set S' \ {(ps, a)} \ S\ by auto ultimately show False using cons unfolding consistent_def by blast qed lemma consistent_drop_block: \consistent ({block} \ S) \ consistent S\ unfolding consistent_def by blast lemma inconsistent_weaken: \\ consistent S \ S \ S' \ \ consistent S'\ unfolding consistent_def by blast lemma finite_nominals_set: \finite S \ finite (\block \ S. block_nominals block)\ by (induct S rule: finite_induct) (simp_all add: finite_block_nominals) lemma witness_list_used: fixes i :: 'b assumes inf: \infinite (UNIV :: 'b set)\ and \finite used\ \i \ list_nominals ps\ shows \i \ list_nominals (witness_list ps ({i} \ used))\ using assms(2-) proof (induct ps arbitrary: used) case (Cons p ps) then show ?case proof (cases \proper_dia p\) case (Some q) let ?j = \SOME j. j \ {i} \ used\ have \finite ({i} \ used)\ using \finite used\ by simp then have \\j. j \ {i} \ used\ using inf ex_new_if_finite by metis then have j: \?j \ {i} \ used\ using someI_ex by metis have \witness_list (p # ps) ({i} \ used) = (\<^bold>@ ?j q) # (\<^bold>\ Nom ?j) # witness_list ps ({?j} \ ({i} \ used))\ using Some witness_list by metis then have *: \list_nominals (witness_list (p # ps) ({i} \ used)) = {?j} \ nominals q \ list_nominals (witness_list ps ({?j} \ ({i} \ used)))\ by simp have \finite ({?j} \ used)\ using \finite used\ by simp moreover have \i \ list_nominals ps\ using \i \ list_nominals (p # ps)\ by simp ultimately have \i \ list_nominals (witness_list ps ({i} \ ({?j} \ used)))\ using Cons by metis moreover have \{i} \ ({?j} \ used) = {?j} \ ({i} \ used)\ by blast moreover have \i \ ?j\ using j by auto ultimately have \i \ list_nominals (witness_list (p # ps) ({i} \ used)) \ i \ nominals q\ using * by simp moreover have \i \ nominals q\ using Cons(3) Some proper_dia by fastforce ultimately show ?thesis by blast qed simp qed simp lemma witness_used: fixes i :: 'b assumes inf: \infinite (UNIV :: 'b set)\ and \finite used\ \i \ block_nominals block\ shows \i \ block_nominals (witness block ({i} \ used))\ using assms witness_list_used by (induct block) fastforce lemma consistent_witness_list: fixes a :: 'b assumes inf: \infinite (UNIV :: 'b set)\ and \consistent S\ \(ps, a) \ S\ \finite used\ \(\ (block_nominals ` S)) \ used\ shows \consistent ({(witness_list ps used, a)} \ S)\ using assms(2-) proof (induct ps arbitrary: used S) case Nil then have \{(witness_list [] used, a)} \ S = S\ by auto then show ?case using \consistent S\ by simp next case (Cons p ps) have \{(p # ps, a)} \ S = S\ using \(p # ps, a) \ S\ by blast then have \consistent ({(p # ps, a)} \ S)\ using \consistent S\ by simp then have \consistent ({(ps, a)} \ S)\ using inf consistent_drop_single by fast moreover have \(ps, a) \ {(ps, a)} \ S\ by simp moreover have \\ (block_nominals ` ({(ps, a)} \ S)) \ extra \ used\ for extra using \(p # ps, a) \ S\ \\ (block_nominals ` S) \ used\ by fastforce moreover have \finite (extra \ used)\ if \finite extra\ for extra using that \finite used\ by blast ultimately have cons: \consistent ({(witness_list ps (extra \ used), a)} \ ({(ps, a)} \ S))\ if \finite extra\ for extra using that Cons by metis show ?case proof (cases \proper_dia p\) case None then have \witness_list (p # ps) used = witness_list ps used\ by auto moreover have \consistent ({(witness_list ps used, a)} \ ({(ps, a)} \ S))\ using cons[where extra=\{}\] by simp then have \consistent ({(witness_list ps used, a)} \ S)\ using consistent_drop_block[where block=\(ps, a)\] by simp ultimately show ?thesis by simp next case (Some q) let ?i = \SOME i. i \ used\ have \\i. i \ used\ using \finite used\ inf ex_new_if_finite by metis then have \?i \ used\ using someI_ex by metis then have i: \?i \ \ (block_nominals ` S)\ using Cons by blast then have \?i \ block_nominals (p # ps, a)\ using Cons by blast let ?tail = \witness_list ps ({?i} \ used)\ have \consistent ({(?tail, a)} \ ({(ps, a)} \ S))\ using cons[where extra=\{?i}\] by simp then have *: \consistent ({(?tail, a)} \ S)\ using consistent_drop_block[where block=\(ps, a)\] by simp have \witness_list (p # ps) used = (\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail\ using Some witness_list by metis moreover have \consistent ({((\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail, a)} \ S)\ unfolding consistent_def proof assume \\S'. set S' \ {((\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail, a)} \ S \ (\n. n \ S')\ then obtain S' n where \n \ S'\ and S': \set S' \ {((\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail, a)} \ S\ \((\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail, a) \. S'\ using * unfolding consistent_def by blast then obtain S'' where S'': \set (((\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail, a) # S'') = set S'\ \((\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail, a) \ set S''\ using split_list[where x=\((\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail, a)\] by blast then have \\ ((\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail, a) # S''\ using inf ST_struct \n \ S'\ by blast moreover have \set (((\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail, a) # S'') \ set (((\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail, a) # (p # ps, a) # S'')\ by auto ultimately have **: \\ ((\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # ?tail, a) # (p # ps, a) # S''\ using inf ST_struct by blast have \?i \ block_nominals (?tail, a)\ using inf \finite used\ \?i \ block_nominals (p # ps, a)\ witness_used by fastforce moreover have \?i \ branch_nominals S''\ unfolding branch_nominals_def using i S' S'' by auto ultimately have \?i \ branch_nominals ((?tail, a) # (p # ps, a) # S'')\ using \?i \ block_nominals (p # ps, a)\ unfolding branch_nominals_def by simp moreover have \\a. q = Nom a\ using Some proper_dia by blast moreover have \(p # ps, a) \. (?tail, a) # (p # ps, a) # S''\ by simp moreover have \p = (\<^bold>\ q)\ using Some proper_dia by blast then have \(\<^bold>\ q) on (p # ps, a)\ by simp ultimately have \\ (?tail, a) # (p # ps, a) # S''\ using ** DiaP' by fast moreover have \set ((p # ps, a) # S'') \ S\ using Cons S' S'' by auto ultimately show False using * unfolding consistent_def by (simp add: subset_Un_eq) qed ultimately show ?thesis by simp qed qed lemma consistent_witness: fixes block :: \('a, 'b) block\ assumes \infinite (UNIV :: 'b set)\ \consistent S\ \finite (\ (block_nominals ` S))\ \block \ S\ shows \consistent ({witness block (\ (block_nominals ` S))} \ S)\ using assms consistent_witness_list by (cases block) fastforce lemma consistent_extend: fixes S :: \('a, 'b) block set\ assumes inf: \infinite (UNIV :: 'b set)\ and \consistent (extend S f n)\ \finite (\ (block_nominals ` extend S f n))\ shows \consistent (extend S f (Suc n))\ proof - consider (inconsistent) \\ consistent ({f n} \ extend S f n)\ | (consistent) \consistent ({f n} \ extend S f n)\ by blast then show ?thesis proof cases case inconsistent then show ?thesis using assms by simp next case consistent let ?used = \\block \ {f n} \ extend S f n. block_nominals block\ have *: \extend S f (n + 1) = {f n, witness (f n) ?used} \ extend S f n\ using consistent by simp have \consistent ({f n} \ extend S f n)\ using consistent by simp moreover have \finite ((\ (block_nominals ` ({f n} \ extend S f n))))\ using \finite (\ (block_nominals ` extend S f n))\ finite_nominals_set by force moreover have \f n \ {f n} \ extend S f n\ by simp ultimately have \consistent ({witness (f n) ?used} \ ({f n} \ extend S f n))\ using inf consistent_witness by blast then show ?thesis using * by simp qed qed lemma finite_nominals_extend: assumes \finite (\ (block_nominals ` S))\ shows \finite (\ (block_nominals ` extend S f n))\ using assms by (induct n) (simp_all add: finite_block_nominals) lemma consistent_extend': fixes S :: \('a, 'b) block set\ assumes inf: \infinite (UNIV :: 'b set)\ and \consistent S\ \finite (\ (block_nominals ` S))\ shows \consistent (extend S f n)\ using assms by (induct n) (simp, metis consistent_extend finite_nominals_extend) lemma UN_finite_bound: assumes \finite A\ \A \ (\n. f n)\ shows \\m :: nat. A \ (\n \ m. f n)\ using assms proof (induct A rule: finite_induct) case (insert x A) then obtain m where \A \ (\n \ m. f n)\ by fast then have \A \ (\n \ (m + k). f n)\ for k by fastforce moreover obtain m' where \x \ f m'\ using insert(4) by blast ultimately have \{x} \ A \ (\n \ m + m'. f n)\ by auto then show ?case by blast qed simp lemma extend_bound: \(\n \ m. extend S f n) = extend S f m\ proof (induct m) case (Suc m) have \\ (extend S f ` {..Suc m}) = \ (extend S f ` {..m}) \ extend S f (Suc m)\ using atMost_Suc by auto also have \\ = extend S f m \ extend S f (Suc m)\ using Suc by blast also have \\ = extend S f (Suc m)\ using extend_chain by blast finally show ?case by simp qed simp lemma consistent_Extend: fixes S :: \('a, 'b) block set\ assumes inf: \infinite (UNIV :: 'b set)\ and \consistent S\ \finite (\ (block_nominals ` S))\ shows \consistent (Extend S f)\ unfolding Extend_def proof (rule ccontr) assume \\ consistent (\ (range (extend S f)))\ then obtain S' n where \n \ S'\ \set S' \ (\n. extend S f n)\ unfolding consistent_def by blast moreover have \finite (set S')\ by simp ultimately obtain m where \set S' \ (\n \ m. extend S f n)\ using UN_finite_bound by metis then have \set S' \ extend S f m\ using extend_bound by blast moreover have \consistent (extend S f m)\ using assms consistent_extend' by blast ultimately show False unfolding consistent_def using \n \ S'\ by blast qed subsubsection \Maximality\ text \A set of blocks is maximally consistent if any proper extension makes it inconsistent.\ definition maximal :: \('a, 'b) block set \ bool\ where \maximal S \ consistent S \ (\block. block \ S \ \ consistent ({block} \ S))\ lemma extend_not_mem: \f n \ extend S f (Suc n) \ \ consistent ({f n} \ extend S f n)\ by (metis Un_insert_left extend.simps(2) insertI1) lemma maximal_Extend: fixes S :: \('a, 'b) block set\ assumes inf: \infinite (UNIV :: 'b set)\ and \consistent S\ \finite (\ (block_nominals ` S))\ \surj f\ shows \maximal (Extend S f)\ proof (rule ccontr) assume \\ maximal (Extend S f)\ then obtain block where \block \ Extend S f\ \consistent ({block} \ Extend S f)\ unfolding maximal_def using assms consistent_Extend by blast obtain n where n: \f n = block\ using \surj f\ unfolding surj_def by metis then have \block \ extend S f (Suc n)\ using \block \ Extend S f\ extend_chain unfolding Extend_def by blast then have \\ consistent ({block} \ extend S f n)\ using n extend_not_mem by blast moreover have \block \ extend S f n\ using \block \ extend S f (Suc n)\ extend_chain by auto then have \{block} \ extend S f n \ {block} \ Extend S f\ unfolding Extend_def by blast ultimately have \\ consistent ({block} \ Extend S f)\ using inconsistent_weaken by blast then show False using \consistent ({block} \ Extend S f)\ by blast qed subsubsection \Saturation\ text \A set of blocks is saturated if every \\<^bold>\ p\ is witnessed.\ definition saturated :: \('a, 'b) block set \ bool\ where \saturated S \ \p i. (\<^bold>\ p) at i in' S \ (\a. p = Nom a) \ (\j. (\<^bold>@ j p) at i in' S \ (\<^bold>\ Nom j) at i in' S)\ lemma witness_list_append: \\extra. witness_list (ps @ qs) used = witness_list ps used @ witness_list qs (extra \ used)\ proof (induct ps arbitrary: used) case Nil then show ?case by (metis Un_absorb append_self_conv2 witness_list.simps(1)) next case (Cons p ps) show ?case proof (cases \\q. proper_dia p = Some q\) case True let ?i = \SOME i. i \ used\ from True obtain q where q: \proper_dia p = Some q\ by blast moreover have \(p # ps) @ qs = p # (ps @ qs)\ by simp ultimately have \witness_list ((p # ps) @ qs) used = (\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # witness_list (ps @ qs) ({?i} \ used)\ using witness_list by metis then have \\extra. witness_list ((p # ps) @ qs) used = (\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # witness_list ps ({?i} \ used) @ witness_list qs (extra \ ({?i} \ used))\ using Cons by metis moreover have \(\<^bold>@ ?i q) # (\<^bold>\ Nom ?i) # witness_list ps ({?i} \ used) = witness_list (p # ps) used\ using q witness_list by metis ultimately have \\extra. witness_list ((p # ps) @ qs) used = witness_list (p # ps) used @ witness_list qs (extra \ ({?i} \ used))\ by (metis append_Cons) then have \\extra. witness_list ((p # ps) @ qs) used = witness_list (p # ps) used @ witness_list qs (({?i} \ extra) \ used)\ by simp then show ?thesis by blast qed (simp add: Cons) qed lemma ex_witness_list: assumes \p \. ps\ \proper_dia p = Some q\ shows \\i. {\<^bold>@ i q, \<^bold>\ Nom i} \ set (witness_list ps used)\ using \p \. ps\ proof (induct ps arbitrary: used) case (Cons a ps) then show ?case proof (induct \a = p\) case True then have \\i. witness_list (a # ps) used = (\<^bold>@ i q) # (\<^bold>\ Nom i) # witness_list ps ({i} \ used)\ using \proper_dia p = Some q\ witness_list by metis then show ?case by auto next case False then have \\i. {\<^bold>@ i q, \<^bold>\ Nom i} \ set (witness_list ps (extra \ used))\ for extra by simp moreover have \\extra. witness_list (a # ps) used = witness_list [a] used @ witness_list ps (extra \ used)\ using witness_list_append[where ps=\[_]\] by simp ultimately show ?case by fastforce qed qed simp lemma saturated_Extend: fixes S :: \('a, 'b) block set\ assumes inf: \infinite (UNIV :: 'b set)\ and \consistent S\ \finite (\ (block_nominals ` S))\ \surj f\ shows \saturated (Extend S f)\ unfolding saturated_def proof safe fix ps i p assume \(ps, i) \ Extend S f\ \(\<^bold>\ p) on (ps, i)\ \\a. p = Nom a\ obtain n where n: \f n = (ps, i)\ using \surj f\ unfolding surj_def by metis let ?used = \(\block \ {f n} \ extend S f n. block_nominals block)\ have \extend S f n \ Extend S f\ unfolding Extend_def by auto moreover have \consistent (Extend S f)\ using assms consistent_Extend by blast ultimately have \consistent ({(ps, i)} \ extend S f n)\ using \(ps, i) \ Extend S f\ inconsistent_weaken by blast then have \extend S f (Suc n) = {f n, witness (f n) ?used} \ extend S f n\ using n \(\<^bold>\ p) on (ps, i)\ by auto then have \witness (f n) ?used \ Extend S f\ unfolding Extend_def by blast then have *: \(witness_list ps ?used, i) \ Extend S f\ using n by simp have \(\<^bold>\ p) \. ps\ using \(\<^bold>\ p) on (ps, i)\ by simp moreover have \proper_dia (\<^bold>\ p) = Some p\ unfolding proper_dia_def using \\a. p = Nom a\ by simp ultimately have \\j. (\<^bold>@ j p) on (witness_list ps ?used, i) \ (\<^bold>\ Nom j) on (witness_list ps ?used, i)\ using ex_witness_list by fastforce then show \\j. (\qs. (qs, i) \ Extend S f \ (\<^bold>@ j p) on (qs, i)) \ (\rs. (rs, i) \ Extend S f \ (\<^bold>\ Nom j) on (rs, i))\ using * by blast qed subsection \Smullyan-Fitting\ lemma hintikka_Extend: fixes S :: \('a, 'b) block set\ assumes inf: \infinite (UNIV :: 'b set)\ and \maximal S\ \consistent S\ \saturated S\ shows \hintikka S\ unfolding hintikka_def proof safe fix x i j ps qs rs assume ps: \(ps, i) \ S\ \Nom j on (ps, i)\ and qs: \(qs, j) \ S\ \Pro x on (qs, j)\ and rs: \(rs, i) \ S\ \(\<^bold>\ Pro x) on (rs, i)\ then have \\ n \ [(ps, i), (qs, j), (rs, i)]\ for n using \consistent S\ unfolding consistent_def by simp moreover have \n \ [(Pro x # ps, i), (qs, j), (rs, i)]\ for n using ps(2) rs(2) by (meson Close list.set_intros(1) on.simps set_subset_Cons subsetD) then have \n \ [(ps, i), (qs, j), (rs, i)]\ for n using ps qs rev_nth_on by (meson Nom' list.set_intros(1) rev_nth_on set_subset_Cons subsetD) ultimately show False by blast next fix a i ps qs assume ps: \(ps, i) \ S\ \Nom a on (ps, i)\ and qs: \(qs, i) \ S\ \(\<^bold>\ Nom a) on (qs, i)\ then have \\ n \ [(ps, i), (qs, i)]\ for n using \consistent S\ unfolding consistent_def by simp moreover have \n \ [(ps, i), (qs, i)]\ for n using ps(2) qs(2) by (meson Close list.set_intros(1) set_subset_Cons subset_code(1)) ultimately show False by blast next fix i j ps qs assume ps: \(ps, i) \ S\ \(\<^bold>\ Nom j) on (ps, i)\ and qs: \(qs, i) \ S\ \(\<^bold>\ (\<^bold>\ Nom j)) on (qs, i)\ then have \\ n \ [(ps, i), (qs, i)]\ for n using \consistent S\ unfolding consistent_def by simp moreover have \n \ [(ps, i), (qs, i)]\ for n using ps(2) qs(2) Close[where p=\\<^bold>\ Nom j\ and i=i] by force ultimately show False by blast next fix p i ps a assume i: \i \ nominals p\ and ps: \(ps, a) \ S\ \p on (ps, a)\ show \\qs. (qs, i) \ S\ proof (rule ccontr) assume \\qs. (qs, i) \ S\ then obtain S' n where \n \ S'\ and S': \set S' \ {([], i)} \ S\ and \([], i) \. S'\ using \maximal S\ unfolding maximal_def consistent_def by (metis insert_is_Un subset_insert) then obtain S'' where S'': \set (([], i) # S'') = set S'\ \([], i) \ set S''\ using split_list[where x=\([], i)\] by blast then have \\ ([], i) # (ps, a) # S''\ using inf ST_struct[where branch'=\([], i) # (ps, a) # S''\] \n \ S'\ by fastforce moreover have \i \ branch_nominals ((ps, a) # S'')\ using i ps unfolding branch_nominals_def by auto ultimately have \\ (ps, a) # S''\ using GoTo by fast moreover have \set ((ps, a) # S'') \ S\ using S' S'' ps by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed next fix i j ps assume ps: \(ps, i) \ S\ \Nom j on (ps, i)\ show \\qs. (qs, j) \ S \ Nom i on (qs, j)\ proof (rule ccontr) assume \\qs. (qs, j) \ S \ Nom i on (qs, j)\ then obtain S' n where \n \ S'\ and S': \set S' \ {([Nom i], j)} \ S\ and \([Nom i], j) \. S'\ using \maximal S\ unfolding maximal_def consistent_def by (metis insert_is_Un list.set_intros(1) on.simps subset_insert) then obtain S'' where S'': \set (([Nom i], j) # S'') = set S'\ \([Nom i], j) \ set S''\ using split_list[where x=\([Nom i], j)\] by blast then have \\ ([Nom i], j) # (ps, i) # S''\ using inf ST_struct[where branch'=\([Nom i], j) # (ps, i) # S''\] \n \ S'\ by fastforce then have \\ ([], j) # (ps, i) # S''\ using \Nom j on (ps, i)\ by (meson Nom' list.set_intros(1) list.set_intros(2) on.simps) moreover have \j \ branch_nominals ((ps, i) # S'')\ using \Nom j on (ps, i)\ unfolding branch_nominals_def by fastforce ultimately have \\ (ps, i) # S''\ using GoTo by fast moreover have \set ((ps, i) # S'') \ S\ using S' S'' ps by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed next fix i j k ps qs assume ps: \(ps, i) \ S\ \Nom j on (ps, i)\ and qs: \(qs, j) \ S\ \Nom k on (qs, j)\ show \\rs. (rs, i) \ S \ Nom k on (rs, i)\ proof (rule ccontr) assume \\rs. (rs, i) \ S \ Nom k on (rs, i)\ then obtain S' n where \n \ S'\ and S': \set S' \ {([Nom k], i)} \ S\ and \([Nom k], i) \. S'\ using \maximal S\ unfolding maximal_def consistent_def by (metis insert_is_Un list.set_intros(1) on.simps subset_insert) then obtain S'' where S'': \set (([Nom k], i) # S'') = set S'\ \([Nom k], i) \ set S''\ using split_list[where x=\([Nom k], i)\] by blast then have \\ ([Nom k], i) # (Nom k # ps, i) # (qs, j) # S''\ using inf ST_struct[where branch'=\([Nom k], i) # (Nom k # ps, i) # (qs, j) # S''\] \n \ S'\ by fastforce then have \\ ([], i) # (Nom k # ps, i) # (qs, j) # S''\ by (meson Nom' list.set_intros(1-2) on.simps) moreover have \i \ branch_nominals ((Nom k # ps, i) # (qs, j) # S'')\ unfolding branch_nominals_def by simp ultimately have \\ (Nom k # ps, i) # (qs, j) # S''\ using GoTo by fast then have \\ (ps, i) # (qs, j) # S''\ using ps qs by (meson Nom' list.set_intros(1-2) on.simps) moreover have \set ((ps, i) # (qs, j) # S'') \ S\ using S' S'' ps qs by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed next fix i j k ps qs assume ps: \(ps, i) \ S\ \(\<^bold>\ Nom j) on (ps, i)\ and qs: \(qs, j) \ S\ \Nom k on (qs, j)\ show \\rs. (rs, i) \ S \ (\<^bold>\ Nom k) on (rs, i)\ proof (rule ccontr) assume \\rs. (rs, i) \ S \ (\<^bold>\ Nom k) on (rs, i)\ then obtain S' n where \n \ S'\ and S': \set S' \ {([\<^bold>\ Nom k], i)} \ S\ and \([\<^bold>\ Nom k], i) \. S'\ using \maximal S\ unfolding maximal_def consistent_def by (metis insert_is_Un list.set_intros(1) on.simps subset_insert) then obtain S'' where S'': \set (([\<^bold>\ Nom k], i) # S'') = set S'\ \([\<^bold>\ Nom k], i) \ set S''\ using split_list[where x=\([\<^bold>\ Nom k], i)\] by blast then have \\ ([\<^bold>\ Nom k], i) # (ps, i) # (qs, j) # S''\ using inf ST_struct[where branch'=\([\<^bold>\ Nom k], i) # (ps, i) # (qs, j) # S''\] \n \ S'\ by fastforce then have \\ ([], i) # (ps, i) # (qs, j) # S''\ using ps(2) qs(2) inf by (meson Bridge list.set_intros(1) on.simps set_subset_Cons subset_iff) moreover have \i \ branch_nominals ((ps, i) # (qs, j) # S'')\ using ps unfolding branch_nominals_def by fastforce ultimately have \\ (ps, i) # (qs, j) # S''\ using GoTo by fast moreover have \set ((ps, i) # (qs, j) # S'') \ S\ using S' S'' ps qs by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed next fix i j k ps qs assume ps: \(ps, i) \ S\ \(\<^bold>\ Nom j) on (ps, i)\ and qs: \(qs, i) \ S\ \Nom k on (qs, i)\ show \\rs. (rs, k) \ S \ (\<^bold>\ Nom j) on (rs, k)\ proof (rule ccontr) assume \\rs. (rs, k) \ S \ (\<^bold>\ Nom j) on (rs, k)\ then obtain S' n where \n \ S'\ and S': \set S' \ {([\<^bold>\ Nom j], k)} \ S\ and \([\<^bold>\ Nom j], k) \. S'\ using \maximal S\ unfolding maximal_def consistent_def by (metis insert_is_Un list.set_intros(1) on.simps subset_insert) then obtain S'' where S'': \set (([\<^bold>\ Nom j], k) # S'') = set S'\ \([\<^bold>\ Nom j], k) \ set S''\ using split_list[where x=\([\<^bold>\ Nom j], k)\] by blast then have \\ ([\<^bold>\ Nom j], k) # (Nom k # ps, i) # (qs, i) # S''\ using inf ST_struct[where branch'=\([\<^bold>\ Nom j], k) # (Nom k # ps, i) # (qs, i) # S''\] \n \ S'\ by fastforce then have \\ ([], k) # (Nom k # ps, i) # (qs, i) # S''\ using ps(2) by (meson Nom' fm.simps(22) list.set_intros(1-2) on.simps) moreover have \k \ branch_nominals ((Nom k # ps, i) # (qs, i) # S'')\ unfolding branch_nominals_def by simp ultimately have \\ (Nom k # ps, i) # (qs, i) # S''\ using GoTo by fast then have \\ (ps, i) # (qs, i) # S''\ using ps(2) qs(2) by (meson Nom' list.set_intros(1-2) on.simps) moreover have \set ((ps, i) # (qs, i) # S'') \ S\ using S' S'' ps qs by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed next fix p q i ps assume ps: \(ps, i) \ S\ \(p \<^bold>\ q) on (ps, i)\ and *: \\ q at i in' S\ show \p at i in' S\ proof (rule ccontr) assume \\ p at i in' S\ then obtain Sp' np where \np \ Sp'\ and Sp': \set Sp' \ {(p # ps, i)} \ S\ and \(p # ps, i) \. Sp'\ using \maximal S\ unfolding maximal_def consistent_def by (metis insert_is_Un list.set_intros(1) on.simps subset_insert) then obtain Sp'' where Sp'': \set ((p # ps, i) # Sp'') = set Sp'\ \(p # ps, i) \ set Sp''\ using split_list[where x=\(p # ps, i)\] by blast then have \\ (p # ps, i) # Sp''\ using \np \ Sp'\ inf ST_struct by blast obtain Sq' nq where \nq \ Sq'\ and Sq': \set Sq' \ {(q # ps, i)} \ S\ and \(q # ps, i) \. Sq'\ using * \maximal S\ unfolding maximal_def consistent_def by (metis insert_is_Un list.set_intros(1) on.simps subset_insert) then obtain Sq'' where Sq'': \set ((q # ps, i) # Sq'') = set Sq'\ \(q # ps, i) \ set Sq''\ using split_list[where x=\(q # ps, i)\] by blast then have \\ (q # ps, i) # Sq''\ using \nq \ Sq'\ inf ST_struct by blast obtain S'' where S'': \set S'' = set Sp'' \ set Sq''\ by (meson set_union) then have \set ((p # ps, i) # Sp'') \ set ((p # ps, i) # S'')\ \set ((q # ps, i) # Sq'') \ set ((q # ps, i) # S'')\ by auto then have \\ (p # ps, i) # S''\ \\ (q # ps, i) # S''\ using \\ (p # ps, i) # Sp''\ \\ (q # ps, i) # Sq''\ inf ST_struct by blast+ then have \\ (ps, i) # S''\ using ps by (meson DisP'' list.set_intros(1)) moreover have \set ((ps, i) # S'') \ S\ using ps Sp' Sp'' Sq' Sq'' S'' by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed next fix p q i ps assume ps: \(ps, i) \ S\ \(\<^bold>\ (p \<^bold>\ q)) on (ps, i)\ show \(\<^bold>\ p) at i in' S\ proof (rule ccontr) assume \\ (\<^bold>\ p) at i in' S\ then obtain S' where \\ S'\ and S': \set S' \ {((\<^bold>\ q) # (\<^bold>\ p) # ps, i)} \ S\ and \((\<^bold>\ q) # (\<^bold>\ p) # ps, i) \. S'\ using \maximal S\ unfolding maximal_def consistent_def by (metis (mono_tags, lifting) insert_is_Un insert_subset list.simps(15) on.simps set_subset_Cons subset_insert) then obtain S'' where S'': \set (((\<^bold>\ q) # (\<^bold>\ p) # ps, i) # S'') = set S'\ \((\<^bold>\ q) # (\<^bold>\ p) # ps, i) \ set S''\ using split_list[where x=\((\<^bold>\ q) # (\<^bold>\ p) # ps, i)\] by blast then have \\ ((\<^bold>\ q) # (\<^bold>\ p) # ps, i) # S''\ using inf ST_struct \\ S'\ by blast then have \\ (ps, i) # S''\ using ps by (meson DisN' list.set_intros(1)) moreover have \set ((ps, i) # S'') \ S\ using S' S'' ps by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed next fix p q i ps assume ps: \(ps, i) \ S\ \(\<^bold>\ (p \<^bold>\ q)) on (ps, i)\ show \(\<^bold>\ q) at i in' S\ proof (rule ccontr) assume \\ (\<^bold>\ q) at i in' S\ then obtain S' where \\ S'\ and S': \set S' \ {((\<^bold>\ q) # (\<^bold>\ p) # ps, i)} \ S\ and \((\<^bold>\ q) # (\<^bold>\ p) # ps, i) \. S'\ using \maximal S\ unfolding maximal_def consistent_def by (metis (mono_tags, lifting) insert_is_Un insert_subset list.simps(15) on.simps set_subset_Cons subset_insert) then obtain S'' where S'': \set (((\<^bold>\ q) # (\<^bold>\ p) # ps, i) # S'') = set S'\ \((\<^bold>\ q) # (\<^bold>\ p) # ps, i) \ set S''\ using split_list[where x=\((\<^bold>\ q) # (\<^bold>\ p) # ps, i)\] by blast then have \\ ((\<^bold>\ q) # (\<^bold>\ p) # ps, i) # S''\ using inf ST_struct \\ S'\ by blast then have \\ (ps, i) # S''\ using ps by (meson DisN' list.set_intros(1)) moreover have \set ((ps, i) # S'') \ S\ using S' S'' ps by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed next fix p i ps assume ps: \(ps, i) \ S\ \(\<^bold>\ \<^bold>\ p) on (ps, i)\ show \\qs. (qs, i) \ S \ p on (qs, i)\ proof (rule ccontr) assume \\qs. (qs, i) \ S \ p on (qs, i)\ then obtain S' n where \n \ S'\ and S': \set S' \ {(p # ps, i)} \ S\ and \(p # ps, i) \. S'\ using \maximal S\ unfolding maximal_def consistent_def by (metis insert_is_Un list.set_intros(1) on.simps subset_insert) then obtain S'' where S'': \set ((p # ps, i) # S'') = set S'\ \(p # ps, i) \ set S''\ using split_list[where x=\(p # ps, i)\] by blast then have \\ (p # ps, i) # S''\ using inf ST_struct \n \ S'\ by blast then have \\ (ps, i) # S''\ using ps by (meson Neg' list.set_intros(1)) moreover have \set ((ps, i) # S'') \ S\ using S' S'' ps by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed next fix p i ps a assume ps: \(ps, a) \ S\ \(\<^bold>@ i p) on (ps, a)\ show \\qs. (qs, i) \ S \ p on (qs, i)\ proof (rule ccontr) assume \\qs. (qs, i) \ S \ p on (qs, i)\ then obtain S' n where \n \ S'\ and S': \set S' \ {([p], i)} \ S\ and \([p], i) \. S'\ using \maximal S\ unfolding maximal_def consistent_def by (metis insert_is_Un list.set_intros(1) on.simps subset_insert) then obtain S'' where S'': \set (([p], i) # S'') = set S'\ \([p], i) \ set S''\ using split_list[where x=\([p], i)\] by blast then have \\ ([p], i) # S''\ using inf ST_struct \n \ S'\ by blast moreover have \set (([p], i) # S'') \ set (([p], i) # (ps, a) # S'')\ by auto ultimately have \\ ([p], i) # (ps, a) # S''\ using inf ST_struct \n \ S'\ by blast then have \\ ([], i) # (ps, a) # S''\ using ps by (metis SatP' insert_iff list.simps(15)) moreover have \i \ branch_nominals ((ps, a) # S'')\ using ps unfolding branch_nominals_def by fastforce ultimately have \\ (ps, a) # S''\ using GoTo by fast moreover have \set ((ps, a) # S'') \ S\ using S' S'' ps by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed next fix p i ps a assume ps: \(ps, a) \ S\ \(\<^bold>\ (\<^bold>@ i p)) on (ps, a)\ show \\qs. (qs, i) \ S \ (\<^bold>\ p) on (qs, i)\ proof (rule ccontr) assume \\qs. (qs, i) \ S \ (\<^bold>\ p) on (qs, i)\ then obtain S' n where \n \ S'\ and S': \set S' \ {([\<^bold>\ p], i)} \ S\ and \([\<^bold>\ p], i) \. S'\ using \maximal S\ unfolding maximal_def consistent_def by (metis insert_is_Un list.set_intros(1) on.simps subset_insert) then obtain S'' where S'': \set (([\<^bold>\ p], i) # S'') = set S'\ \([\<^bold>\ p], i) \ set S''\ using split_list[where x=\([\<^bold>\ p], i)\] by blast then have \\ ([\<^bold>\ p], i) # S''\ using inf ST_struct \n \ S'\ by blast then have \\ ([\<^bold>\ p], i) # (ps, a) # S''\ using inf ST_struct[where branch'=\([\<^bold>\ p], i) # _ # S''\] \n \ S'\ by fastforce then have \\ ([], i) # (ps, a) # S''\ using ps by (metis SatN' insert_iff list.simps(15)) moreover have \i \ branch_nominals ((ps, a) # S'')\ using ps unfolding branch_nominals_def by fastforce ultimately have \\ (ps, a) # S''\ using GoTo by fast moreover have \set ((ps, a) # S'') \ S\ using S' S'' ps by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed next fix p i ps assume \\a. p = Nom a\ and ps: \(ps, i) \ S\ \(\<^bold>\ p) on (ps, i)\ then show \\j. (\qs. (qs, i) \ S \ (\<^bold>\ Nom j) on (qs, i)) \ (\rs. (rs, i) \ S \ (\<^bold>@ j p) on (rs, i))\ using \saturated S\ unfolding saturated_def by blast next fix p i j ps qs assume ps: \(ps, i) \ S\ \(\<^bold>\ (\<^bold>\ p)) on (ps, i)\ and qs: \(qs, i) \ S\ \(\<^bold>\ Nom j) on (qs, i)\ show \\rs. (rs, i) \ S \ (\<^bold>\ (\<^bold>@ j p)) on (rs, i)\ proof (rule ccontr) assume \\qs. (qs, i) \ S \ (\<^bold>\ (\<^bold>@ j p)) on (qs, i)\ then obtain S' n where \n\ S'\ and S': \set S' \ {([\<^bold>\ (\<^bold>@ j p)], i)} \ S\ and \([\<^bold>\ (\<^bold>@ j p)], i) \. S'\ using \maximal S\ unfolding maximal_def consistent_def by (metis insert_is_Un list.set_intros(1) on.simps subset_insert) then obtain S'' where S'': \set (([\<^bold>\ (\<^bold>@ j p)], i) # S'') = set S'\ \([\<^bold>\ (\<^bold>@ j p)], i) \ set S''\ using split_list[where x=\([\<^bold>\ (\<^bold>@ j p)], i)\] by blast then have \\ ([\<^bold>\ (\<^bold>@ j p)], i) # S''\ using inf ST_struct \n \ S'\ by blast then have \\ ([\<^bold>\ (\<^bold>@ j p)], i) # (ps, i) # (qs, i) # S''\ using inf ST_struct[where branch'=\([_], _) # (ps, i) # (qs, i) # S''\] \n \ S'\ by fastforce then have \\ ([], i) # (ps, i) # (qs, i) # S''\ using ps(2) qs(2) by (meson DiaN' list.set_intros(1) set_subset_Cons subset_iff) moreover have \i \ branch_nominals ((ps, i) # (qs, i) # S'')\ unfolding branch_nominals_def by simp ultimately have \\ (ps, i) # (qs, i) # S''\ using GoTo by fast moreover have \set ((ps, i) # (qs, i) # S'') \ S\ using S' S'' ps qs by auto ultimately show False using \consistent S\ unfolding consistent_def by blast qed qed subsection \Result\ theorem completeness: fixes p :: \('a :: countable, 'b :: countable) fm\ assumes inf: \infinite (UNIV :: 'b set)\ and valid: \\(M :: ('b set, 'a) model) g w. M, g, w \ p\ shows \1 \ [([\<^bold>\ p], i)]\ proof - have \\ [([\<^bold>\ p], i)]\ proof (rule ccontr) assume \\ \ [([\<^bold>\ p], i)]\ then have *: \consistent {([\<^bold>\ p], i)}\ unfolding consistent_def using ST_struct inf by (metis empty_set list.simps(15)) let ?S = \Extend {([\<^bold>\ p], i)} from_nat\ have \finite {([\<^bold>\ p], i)}\ by simp then have fin: \finite (\ (block_nominals ` {([\<^bold>\ p], i)}))\ using finite_nominals_set by blast have \consistent ?S\ using consistent_Extend inf * fin by blast moreover have \maximal ?S\ using maximal_Extend inf * fin by fastforce moreover have \saturated ?S\ using saturated_Extend inf * fin by fastforce ultimately have \hintikka ?S\ using hintikka_Extend inf by blast moreover have \([\<^bold>\ p], i) \ ?S\ using Extend_mem by blast moreover have \(\<^bold>\ p) on ([\<^bold>\ p], i)\ by simp ultimately have \\ Model (reach ?S) (val ?S), assign ?S, assign ?S i \ p\ using hintikka_model by fast then show False using valid by blast qed then show ?thesis using ST_one by blast qed text \ We arbitrarily fix nominal and propositional symbols to be natural numbers (any countably infinite type suffices) and define validity as truth in all models with sets of natural numbers as worlds. We show below that this implies validity for any type of worlds. \ abbreviation \valid p \ \(M :: (nat set, nat) model) (g :: nat \ _) w. M, g, w \ p\ text \A formula is valid iff its negation has a closing tableau from a fresh world.\ theorem main: assumes \i \ nominals p\ shows \valid p \ 1 \ [([\<^bold>\ p], i)]\ proof assume \valid p\ then show \1 \ [([\<^bold>\ p], i)]\ using completeness by blast next assume \1 \ [([\<^bold>\ p], i)]\ then show \valid p\ using assms soundness_fresh by fast qed text \The restricted validity implies validity in general.\ theorem valid_semantics: \valid p \ M, g, w \ p\ proof assume \valid p\ then have \i \ nominals p \ \ [([\<^bold>\ p], i)]\ for i using main by blast moreover have \\i. i \ nominals p\ by (simp add: finite_nominals ex_new_if_finite) ultimately show \M, g, w \ p\ using soundness_fresh by fast qed end diff --git a/thys/Iptables_Semantics/Semantics.thy b/thys/Iptables_Semantics/Semantics.thy --- a/thys/Iptables_Semantics/Semantics.thy +++ b/thys/Iptables_Semantics/Semantics.thy @@ -1,1226 +1,1226 @@ theory Semantics imports Main Firewall_Common "Common/List_Misc" "HOL-Library.LaTeXsugar" begin section\Big Step Semantics\ text\ The assumption we apply in general is that the firewall does not alter any packets. \ text\A firewall ruleset is a map of chain names (e.g., INPUT, OUTPUT, FORWARD, arbitrary-user-defined-chain) to a list of rules. The list of rules is processed sequentially.\ type_synonym 'a ruleset = "string \ 'a rule list" text\A matcher (parameterized by the type of primitive @{typ 'a} and packet @{typ 'p}) is a function which just tells whether a given primitive and packet matches.\ type_synonym ('a, 'p) matcher = "'a \ 'p \ bool" text\Example: Assume a network packet only has a destination street number (for simplicity, of type @{typ "nat"}) and we only support the following match expression: Is the packet's street number within a certain range? The type for the primitive could then be @{typ "nat \ nat"} and a possible implementation for @{typ "(nat \ nat, nat) matcher"} could be @{term "match_street_number (a,b) p \ p \ {a .. b}"}. Usually, the primitives are a datatype which supports interfaces, IP addresses, protocols, ports, payload, ...\ text\Given an @{typ "('a, 'p) matcher"} and a match expression, does a packet of type @{typ 'p} match the match expression?\ fun matches :: "('a, 'p) matcher \ 'a match_expr \ 'p \ bool" where "matches \ (MatchAnd e1 e2) p \ matches \ e1 p \ matches \ e2 p" | "matches \ (MatchNot me) p \ \ matches \ me p" | "matches \ (Match e) p \ \ e p" | "matches _ MatchAny _ \ True" (*Note: "matches \ (MatchNot me) p \ \ matches \ me p" does not work for ternary logic. Here, we have Boolean logic and everything is fine.*) inductive iptables_bigstep :: "'a ruleset \ ('a, 'p) matcher \ 'p \ 'a rule list \ state \ state \ bool" ("_,_,_\ \_, _\ \ _" [60,60,60,20,98,98] 89) for \ and \ and p where skip: "\,\,p\ \[], t\ \ t" | accept: "matches \ m p \ \,\,p\ \[Rule m Accept], Undecided\ \ Decision FinalAllow" | drop: "matches \ m p \ \,\,p\ \[Rule m Drop], Undecided\ \ Decision FinalDeny" | reject: "matches \ m p \ \,\,p\ \[Rule m Reject], Undecided\ \ Decision FinalDeny" | log: "matches \ m p \ \,\,p\ \[Rule m Log], Undecided\ \ Undecided" | (*empty does not do anything to the packet. It could update the internal firewall state, e.g. marking a packet for later-on rate limiting*) empty: "matches \ m p \ \,\,p\ \[Rule m Empty], Undecided\ \ Undecided" | nomatch: "\ matches \ m p \ \,\,p\ \[Rule m a], Undecided\ \ Undecided" | decision: "\,\,p\ \rs, Decision X\ \ Decision X" | seq: "\\,\,p\ \rs\<^sub>1, Undecided\ \ t; \,\,p\ \rs\<^sub>2, t\ \ t'\ \ \,\,p\ \rs\<^sub>1@rs\<^sub>2, Undecided\ \ t'" | call_return: "\ matches \ m p; \ chain = Some (rs\<^sub>1@[Rule m' Return]@rs\<^sub>2); matches \ m' p; \,\,p\ \rs\<^sub>1, Undecided\ \ Undecided \ \ \,\,p\ \[Rule m (Call chain)], Undecided\ \ Undecided" | call_result: "\ matches \ m p; \ chain = Some rs; \,\,p\ \rs, Undecided\ \ t \ \ \,\,p\ \[Rule m (Call chain)], Undecided\ \ t" text\ The semantic rules again in pretty format: \begin{center} @{thm[mode=Axiom] skip [no_vars]}\\[1ex] @{thm[mode=Rule] accept [no_vars]}\\[1ex] @{thm[mode=Rule] drop [no_vars]}\\[1ex] @{thm[mode=Rule] reject [no_vars]}\\[1ex] @{thm[mode=Rule] log [no_vars]}\\[1ex] @{thm[mode=Rule] empty [no_vars]}\\[1ex] @{thm[mode=Rule] nomatch [no_vars]}\\[1ex] @{thm[mode=Rule] decision [no_vars]}\\[1ex] @{thm[mode=Rule] seq [no_vars]} \\[1ex] @{thm[mode=Rule] call_return [no_vars]}\\[1ex] @{thm[mode=Rule] call_result [no_vars]} \end{center} \ (*future work: Add abstraction function for unknown actions. At the moment, only the explicitly listed actions are supported. This would also require a @{text "Decision FinalUnknown"} state Problem: An unknown action may modify a packet. Assume that we have a firewall which accepts the packets A->B and rewrites the header to A->C. After that firewall, there is another firewall which only accepts packets for A->C. A can send through both firewalls. If our model says that the firewall accepts packets A->B but does not consider packet modification, A might not be able to pass the second firewall with this model. Luckily, our model is correct for the filtering behaviour and explicitly does not support any actions with packet modification. Thus, the described scenario is not a counterexample that our model is wrong but a hint for future features we may want to support. Luckily, we introduced the @{term "Decision state"}, which should make adding packet modification states easy. *) lemma deny: "matches \ m p \ a = Drop \ a = Reject \ iptables_bigstep \ \ p [Rule m a] Undecided (Decision FinalDeny)" by (auto intro: drop reject) lemma seq_cons: assumes "\,\,p\ \[r],Undecided\ \ t" and "\,\,p\ \rs,t\ \ t'" shows "\,\,p\ \r#rs, Undecided\ \ t'" proof - from assms have "\,\,p\ \[r] @ rs, Undecided\ \ t'" by (rule seq) thus ?thesis by simp qed lemma iptables_bigstep_induct [case_names Skip Allow Deny Log Nomatch Decision Seq Call_return Call_result, induct pred: iptables_bigstep]: "\ \,\,p\ \rs,s\ \ t; \t. P [] t t; \m a. matches \ m p \ a = Accept \ P [Rule m a] Undecided (Decision FinalAllow); \m a. matches \ m p \ a = Drop \ a = Reject \ P [Rule m a] Undecided (Decision FinalDeny); \m a. matches \ m p \ a = Log \ a = Empty \ P [Rule m a] Undecided Undecided; \m a. \ matches \ m p \ P [Rule m a] Undecided Undecided; \rs X. P rs (Decision X) (Decision X); \rs rs\<^sub>1 rs\<^sub>2 t t'. rs = rs\<^sub>1 @ rs\<^sub>2 \ \,\,p\ \rs\<^sub>1,Undecided\ \ t \ P rs\<^sub>1 Undecided t \ \,\,p\ \rs\<^sub>2,t\ \ t' \ P rs\<^sub>2 t t' \ P rs Undecided t'; \m a chain rs\<^sub>1 m' rs\<^sub>2. matches \ m p \ a = Call chain \ \ chain = Some (rs\<^sub>1 @ [Rule m' Return] @ rs\<^sub>2) \ matches \ m' p \ \,\,p\ \rs\<^sub>1,Undecided\ \ Undecided \ P rs\<^sub>1 Undecided Undecided \ P [Rule m a] Undecided Undecided; \m a chain rs t. matches \ m p \ a = Call chain \ \ chain = Some rs \ \,\,p\ \rs,Undecided\ \ t \ P rs Undecided t \ P [Rule m a] Undecided t \ \ P rs s t" by (induction rule: iptables_bigstep.induct) auto lemma skipD: "\,\,p\ \r, s\ \ t \ r = [] \ s = t" by (induction rule: iptables_bigstep.induct) auto lemma decisionD: "\,\,p\ \r, s\ \ t \ s = Decision X \ t = Decision X" by (induction rule: iptables_bigstep_induct) auto context notes skipD[dest] list_app_singletonE[elim] begin lemma acceptD: "\,\,p\ \r, s\ \ t \ r = [Rule m Accept] \ matches \ m p \ s = Undecided \ t = Decision FinalAllow" by (induction rule: iptables_bigstep.induct) auto lemma dropD: "\,\,p\ \r, s\ \ t \ r = [Rule m Drop] \ matches \ m p \ s = Undecided \ t = Decision FinalDeny" by (induction rule: iptables_bigstep.induct) auto lemma rejectD: "\,\,p\ \r, s\ \ t \ r = [Rule m Reject] \ matches \ m p \ s = Undecided \ t = Decision FinalDeny" by (induction rule: iptables_bigstep.induct) auto lemma logD: "\,\,p\ \r, s\ \ t \ r = [Rule m Log] \ matches \ m p \ s = Undecided \ t = Undecided" by (induction rule: iptables_bigstep.induct) auto lemma emptyD: "\,\,p\ \r, s\ \ t \ r = [Rule m Empty] \ matches \ m p \ s = Undecided \ t = Undecided" by (induction rule: iptables_bigstep.induct) auto lemma nomatchD: "\,\,p\ \r, s\ \ t \ r = [Rule m a] \ s = Undecided \ \ matches \ m p \ t = Undecided" by (induction rule: iptables_bigstep.induct) auto lemma callD: assumes "\,\,p\ \r, s\ \ t" "r = [Rule m (Call chain)]" "s = Undecided" "matches \ m p" "\ chain = Some rs" obtains "\,\,p\ \rs,s\ \ t" | rs\<^sub>1 rs\<^sub>2 m' where "rs = rs\<^sub>1 @ Rule m' Return # rs\<^sub>2" "matches \ m' p" "\,\,p\ \rs\<^sub>1,s\ \ Undecided" "t = Undecided" using assms proof (induction r s t arbitrary: rs rule: iptables_bigstep.induct) case (seq rs\<^sub>1) thus ?case by (cases rs\<^sub>1) auto qed auto end lemmas iptables_bigstepD = skipD acceptD dropD rejectD logD emptyD nomatchD decisionD callD lemma seq': assumes "rs = rs\<^sub>1 @ rs\<^sub>2" "\,\,p\ \rs\<^sub>1,s\ \ t" "\,\,p\ \rs\<^sub>2,t\ \ t'" shows "\,\,p\ \rs,s\ \ t'" using assms by (cases s) (auto intro: seq decision dest: decisionD) lemma seq'_cons: "\,\,p\ \[r],s\ \ t \ \,\,p\ \rs,t\ \ t' \ \,\,p\ \r#rs, s\ \ t'" by (metis decision decisionD state.exhaust seq_cons) lemma seq_split: assumes "\,\,p\ \rs, s\ \ t" "rs = rs\<^sub>1@rs\<^sub>2" obtains t' where "\,\,p\ \rs\<^sub>1,s\ \ t'" "\,\,p\ \rs\<^sub>2,t'\ \ t" using assms proof (induction rs s t arbitrary: rs\<^sub>1 rs\<^sub>2 thesis rule: iptables_bigstep_induct) case Allow thus ?case by (cases rs\<^sub>1) (auto intro: iptables_bigstep.intros) next case Deny thus ?case by (cases rs\<^sub>1) (auto intro: iptables_bigstep.intros) next case Log thus ?case by (cases rs\<^sub>1) (auto intro: iptables_bigstep.intros) next case Nomatch thus ?case by (cases rs\<^sub>1) (auto intro: iptables_bigstep.intros) next case (Seq rs rsa rsb t t') hence rs: "rsa @ rsb = rs\<^sub>1 @ rs\<^sub>2" by simp note List.append_eq_append_conv_if[simp] from rs show ?case proof (cases rule: list_app_eq_cases) case longer with Seq have t1: "\,\,p\ \take (length rsa) rs\<^sub>1, Undecided\ \ t" by simp from Seq longer obtain t2 where t2a: "\,\,p\ \drop (length rsa) rs\<^sub>1,t\ \ t2" and rs2_t2: "\,\,p\ \rs\<^sub>2,t2\ \ t'" by blast with t1 rs2_t2 have "\,\,p\ \take (length rsa) rs\<^sub>1 @ drop (length rsa) rs\<^sub>1,Undecided\ \ t2" by (blast intro: iptables_bigstep.seq) with Seq rs2_t2 show ?thesis by simp next case shorter with rs have rsa': "rsa = rs\<^sub>1 @ take (length rsa - length rs\<^sub>1) rs\<^sub>2" by (metis append_eq_conv_conj length_drop) from shorter rs have rsb': "rsb = drop (length rsa - length rs\<^sub>1) rs\<^sub>2" by (metis append_eq_conv_conj length_drop) from Seq rsa' obtain t1 where t1a: "\,\,p\ \rs\<^sub>1,Undecided\ \ t1" and t1b: "\,\,p\ \take (length rsa - length rs\<^sub>1) rs\<^sub>2,t1\ \ t" by blast from rsb' Seq.hyps have t2: "\,\,p\ \drop (length rsa - length rs\<^sub>1) rs\<^sub>2,t\ \ t'" by blast with seq' t1b have "\,\,p\ \rs\<^sub>2,t1\ \ t'" by fastforce with Seq t1a show ?thesis by fast qed next case Call_return hence "\,\,p\ \rs\<^sub>1, Undecided\ \ Undecided" "\,\,p\ \rs\<^sub>2, Undecided\ \ Undecided" by (case_tac [!] rs\<^sub>1) (auto intro: iptables_bigstep.skip iptables_bigstep.call_return) thus ?case by fact next case (Call_result _ _ _ _ t) show ?case proof (cases rs\<^sub>1) case Nil with Call_result have "\,\,p\ \rs\<^sub>1, Undecided\ \ Undecided" "\,\,p\ \rs\<^sub>2, Undecided\ \ t" by (auto intro: iptables_bigstep.intros) thus ?thesis by fact next case Cons with Call_result have "\,\,p\ \rs\<^sub>1, Undecided\ \ t" "\,\,p\ \rs\<^sub>2, t\ \ t" by (auto intro: iptables_bigstep.intros) thus ?thesis by fact qed qed (auto intro: iptables_bigstep.intros) lemma seqE: assumes "\,\,p\ \rs\<^sub>1@rs\<^sub>2, s\ \ t" obtains ti where "\,\,p\ \rs\<^sub>1,s\ \ ti" "\,\,p\ \rs\<^sub>2,ti\ \ t" using assms by (force elim: seq_split) lemma seqE_cons: assumes "\,\,p\ \r#rs, s\ \ t" obtains ti where "\,\,p\ \[r],s\ \ ti" "\,\,p\ \rs,ti\ \ t" using assms by (metis append_Cons append_Nil seqE) lemma nomatch': assumes "\r. r \ set rs \ \ matches \ (get_match r) p" shows "\,\,p\ \rs, s\ \ s" proof(cases s) case Undecided have "\r\set rs. \ matches \ (get_match r) p \ \,\,p\ \rs, Undecided\ \ Undecided" proof(induction rs) case Nil thus ?case by (fast intro: skip) next case (Cons r rs) hence "\,\,p\ \[r], Undecided\ \ Undecided" by (cases r) (auto intro: nomatch) with Cons show ?case by (fastforce intro: seq_cons) qed with assms Undecided show ?thesis by simp qed (blast intro: decision) text\there are only two cases when there can be a Return on top-level: \<^item> the firewall is in a Decision state \<^item> the return does not match In both cases, it is not applied! \ lemma no_free_return: assumes "\,\,p\ \[Rule m Return], Undecided\ \ t" and "matches \ m p" shows "False" proof - { fix a s have no_free_return_hlp: "\,\,p\ \a,s\ \ t \ matches \ m p \ s = Undecided \ a = [Rule m Return] \ False" proof (induction rule: iptables_bigstep.induct) case (seq rs\<^sub>1) thus ?case by (cases rs\<^sub>1) (auto dest: skipD) qed simp_all } with assms show ?thesis by blast qed (* seq_split is elim, seq_progress is dest *) lemma seq_progress: "\,\,p\ \rs, s\ \ t \ rs = rs\<^sub>1@rs\<^sub>2 \ \,\,p\ \rs\<^sub>1, s\ \ t' \ \,\,p\ \rs\<^sub>2, t'\ \ t" proof(induction arbitrary: rs\<^sub>1 rs\<^sub>2 t' rule: iptables_bigstep_induct) case Allow thus ?case by (cases "rs\<^sub>1") (auto intro: iptables_bigstep.intros dest: iptables_bigstepD) next case Deny thus ?case by (cases "rs\<^sub>1") (auto intro: iptables_bigstep.intros dest: iptables_bigstepD) next case Log thus ?case by (cases "rs\<^sub>1") (auto intro: iptables_bigstep.intros dest: iptables_bigstepD) next case Nomatch thus ?case by (cases "rs\<^sub>1") (auto intro: iptables_bigstep.intros dest: iptables_bigstepD) next case Decision thus ?case by (cases "rs\<^sub>1") (auto intro: iptables_bigstep.intros dest: iptables_bigstepD) next case(Seq rs rsa rsb t t' rs\<^sub>1 rs\<^sub>2 t'') hence rs: "rsa @ rsb = rs\<^sub>1 @ rs\<^sub>2" by simp note List.append_eq_append_conv_if[simp] (* TODO larsrh custom case distinction rule *) from rs show "\,\,p\ \rs\<^sub>2,t''\ \ t'" proof(cases rule: list_app_eq_cases) case longer have "rs\<^sub>1 = take (length rsa) rs\<^sub>1 @ drop (length rsa) rs\<^sub>1" by auto with Seq longer show ?thesis by (metis append_Nil2 skipD seq_split) next case shorter with Seq(7) Seq.hyps(3) Seq.IH(1) rs show ?thesis by (metis seq' append_eq_conv_conj) qed next case(Call_return m a chain rsa m' rsb) have xx: "\,\,p\ \[Rule m (Call chain)], Undecided\ \ t' \ matches \ m p \ \ chain = Some (rsa @ Rule m' Return # rsb) \ matches \ m' p \ \,\,p\ \rsa, Undecided\ \ Undecided \ t' = Undecided" apply(erule callD) apply(simp_all) apply(erule seqE) apply(erule seqE_cons) by (metis Call_return.IH no_free_return self_append_conv skipD) show ?case proof (cases rs\<^sub>1) case (Cons r rs) thus ?thesis using Call_return apply(case_tac "[Rule m a] = rs\<^sub>2") apply(simp) apply(simp) using xx by blast next case Nil moreover hence "t' = Undecided" by (metis Call_return.hyps(1) Call_return.prems(2) append.simps(1) decision no_free_return seq state.exhaust) moreover have "\m. \,\,p\ \[Rule m a], Undecided\ \ Undecided" by (metis (no_types) Call_return(2) Call_return.hyps(3) Call_return.hyps(4) Call_return.hyps(5) call_return nomatch) ultimately show ?thesis using Call_return.prems(1) by auto qed next case(Call_result m a chain rs t) thus ?case proof (cases rs\<^sub>1) case Cons thus ?thesis using Call_result apply(auto simp add: iptables_bigstep.skip iptables_bigstep.call_result dest: skipD) apply(drule callD, simp_all) apply blast by (metis Cons_eq_appendI append_self_conv2 no_free_return seq_split) qed (fastforce intro: iptables_bigstep.intros dest: skipD) qed (auto dest: iptables_bigstepD) theorem iptables_bigstep_deterministic: assumes "\,\,p\ \rs, s\ \ t" and "\,\,p\ \rs, s\ \ t'" shows "t = t'" proof - { fix r1 r2 m t assume a1: "\,\,p\ \r1 @ Rule m Return # r2, Undecided\ \ t" and a2: "matches \ m p" and a3: "\,\,p\ \r1,Undecided\ \ Undecided" have False proof - from a1 a3 have "\,\,p\ \Rule m Return # r2, Undecided\ \ t" by (blast intro: seq_progress) hence "\,\,p\ \[Rule m Return] @ r2, Undecided\ \ t" by simp from seqE[OF this] obtain ti where "\,\,p\ \[Rule m Return], Undecided\ \ ti" by blast with no_free_return a2 show False by fast (*by (blast intro: no_free_return elim: seq_split)*) qed } note no_free_return_seq=this from assms show ?thesis proof (induction arbitrary: t' rule: iptables_bigstep_induct) case Seq thus ?case by (metis seq_progress) next case Call_result thus ?case by (metis no_free_return_seq callD) next case Call_return thus ?case by (metis append_Cons callD no_free_return_seq) qed (auto dest: iptables_bigstepD) qed lemma iptables_bigstep_to_undecided: "\,\,p\ \rs, s\ \ Undecided \ s = Undecided" by (metis decisionD state.exhaust) lemma iptables_bigstep_to_decision: "\,\,p\ \rs, Decision Y\ \ Decision X \ Y = X" by (metis decisionD state.inject) lemma Rule_UndecidedE: assumes "\,\,p\ \[Rule m a], Undecided\ \ Undecided" obtains (nomatch) "\ matches \ m p" | (log) "a = Log \ a = Empty" | (call) c where "a = Call c" "matches \ m p" using assms proof (induction "[Rule m a]" Undecided Undecided rule: iptables_bigstep_induct) case Seq thus ?case by (metis append_eq_Cons_conv append_is_Nil_conv iptables_bigstep_to_undecided) qed simp_all lemma Rule_DecisionE: assumes "\,\,p\ \[Rule m a], Undecided\ \ Decision X" obtains (call) chain where "matches \ m p" "a = Call chain" | (accept_reject) "matches \ m p" "X = FinalAllow \ a = Accept" "X = FinalDeny \ a = Drop \ a = Reject" using assms proof (induction "[Rule m a]" Undecided "Decision X" rule: iptables_bigstep_induct) case (Seq rs\<^sub>1) thus ?case by (cases rs\<^sub>1) (auto dest: skipD) qed simp_all lemma log_remove: assumes "\,\,p\ \rs\<^sub>1 @ [Rule m Log] @ rs\<^sub>2, s\ \ t" shows "\,\,p\ \rs\<^sub>1 @ rs\<^sub>2, s\ \ t" proof - from assms obtain t' where t': "\,\,p\ \rs\<^sub>1, s\ \ t'" "\,\,p\ \[Rule m Log] @ rs\<^sub>2, t'\ \ t" by (blast elim: seqE) hence "\,\,p\ \Rule m Log # rs\<^sub>2, t'\ \ t" by simp then obtain t'' where "\,\,p\ \[Rule m Log], t'\ \ t''" "\,\,p\ \rs\<^sub>2, t''\ \ t" by (blast elim: seqE_cons) with t' show ?thesis by (metis state.exhaust iptables_bigstep_deterministic decision log nomatch seq) qed lemma empty_empty: assumes "\,\,p\ \rs\<^sub>1 @ [Rule m Empty] @ rs\<^sub>2, s\ \ t" shows "\,\,p\ \rs\<^sub>1 @ rs\<^sub>2, s\ \ t" proof - from assms obtain t' where t': "\,\,p\ \rs\<^sub>1, s\ \ t'" "\,\,p\ \[Rule m Empty] @ rs\<^sub>2, t'\ \ t" by (blast elim: seqE) hence "\,\,p\ \Rule m Empty # rs\<^sub>2, t'\ \ t" by simp then obtain t'' where "\,\,p\ \[Rule m Empty], t'\ \ t''" "\,\,p\ \rs\<^sub>2, t''\ \ t" by (blast elim: seqE_cons) with t' show ?thesis by (metis state.exhaust iptables_bigstep_deterministic decision empty nomatch seq) qed lemma Unknown_actions_False: "\,\,p\ \r # rs, Undecided\ \ t \ r = Rule m a \ matches \ m p \ a = Unknown \ (\chain. a = Goto chain) \ False" proof - have 1: "\,\,p\ \[Rule m Unknown], Undecided\ \ t \ matches \ m p \ False" by (induction "[Rule m Unknown]" Undecided t rule: iptables_bigstep.induct) (auto elim: list_app_singletonE dest: skipD) { fix chain have "\,\,p\ \[Rule m (Goto chain)], Undecided\ \ t \ matches \ m p \ False" by (induction "[Rule m (Goto chain)]" Undecided t rule: iptables_bigstep.induct) (auto elim: list_app_singletonE dest: skipD) }note 2=this show "\,\,p\ \r # rs, Undecided\ \ t \ r = Rule m a \ matches \ m p \ a = Unknown \ (\chain. a = Goto chain) \ False" apply(erule seqE_cons) apply(case_tac ti) apply(simp_all) using Rule_UndecidedE apply fastforce by (metis "1" "2" decision iptables_bigstep_deterministic) qed text\ The notation we prefer in the paper. The semantics are defined for fixed \\\ and \\\ \ locale iptables_bigstep_fixedbackground = fixes \::"'a ruleset" and \::"('a, 'p) matcher" begin inductive iptables_bigstep' :: "'p \ 'a rule list \ state \ state \ bool" - ("_\' \_, _\ \ _" [60,20,98,98] 89) + ("_\'' \_, _\ \ _" [60,20,98,98] 89) for p where skip: "p\' \[], t\ \ t" | accept: "matches \ m p \ p\' \[Rule m Accept], Undecided\ \ Decision FinalAllow" | drop: "matches \ m p \ p\' \[Rule m Drop], Undecided\ \ Decision FinalDeny" | reject: "matches \ m p \ p\' \[Rule m Reject], Undecided\ \ Decision FinalDeny" | log: "matches \ m p \ p\' \[Rule m Log], Undecided\ \ Undecided" | empty: "matches \ m p \ p\' \[Rule m Empty], Undecided\ \ Undecided" | nomatch: "\ matches \ m p \ p\' \[Rule m a], Undecided\ \ Undecided" | decision: "p\' \rs, Decision X\ \ Decision X" | seq: "\p\' \rs\<^sub>1, Undecided\ \ t; p\' \rs\<^sub>2, t\ \ t'\ \ p\' \rs\<^sub>1@rs\<^sub>2, Undecided\ \ t'" | call_return: "\ matches \ m p; \ chain = Some (rs\<^sub>1@[Rule m' Return]@rs\<^sub>2); matches \ m' p; p\' \rs\<^sub>1, Undecided\ \ Undecided \ \ p\' \[Rule m (Call chain)], Undecided\ \ Undecided" | call_result: "\ matches \ m p; p\' \the (\ chain), Undecided\ \ t \ \ p\' \[Rule m (Call chain)], Undecided\ \ t" definition wf_\:: "'a rule list \ bool" where "wf_\ rs \ \rsg \ ran \ \ {rs}. (\r \ set rsg. \ chain. get_action r = Call chain \ \ chain \ None)" lemma wf_\_append: "wf_\ (rs1@rs2) \ wf_\ rs1 \ wf_\ rs2" by(simp add: wf_\_def, blast) lemma wf_\_tail: "wf_\ (r # rs) \ wf_\ rs" by(simp add: wf_\_def) lemma wf_\_Call: "wf_\ [Rule m (Call chain)] \ wf_\ (the (\ chain)) \ (\rs. \ chain = Some rs)" apply(simp add: wf_\_def) by (metis option.collapse ranI) lemma "wf_\ rs \ p\' \rs, s\ \ t \ \,\,p\ \rs, s\ \ t" apply(rule iffI) apply(rotate_tac 1) apply(induction rs s t rule: iptables_bigstep'.induct) apply(auto intro: iptables_bigstep.intros simp: wf_\_append dest!: wf_\_Call)[11] apply(rotate_tac 1) apply(induction rs s t rule: iptables_bigstep.induct) apply(auto intro: iptables_bigstep'.intros simp: wf_\_append dest!: wf_\_Call)[11] done end text\Showing that semantics are defined. For rulesets which can be loaded by the Linux kernel. The kernel does not allow loops.\ text\ We call a ruleset well-formed (wf) iff all @{const Call}s are into actually existing chains. \ definition wf_chain :: "'a ruleset \ 'a rule list \ bool" where "wf_chain \ rs \ (\r \ set rs. \ chain. get_action r = Call chain \ \ chain \ None)" lemma wf_chain_append: "wf_chain \ (rs1@rs2) \ wf_chain \ rs1 \ wf_chain \ rs2" by(simp add: wf_chain_def, blast) lemma wf_chain_fst: "wf_chain \ (r # rs) \ wf_chain \ (rs)" by(simp add: wf_chain_def) text\This is what our tool will check at runtime\ definition sanity_wf_ruleset :: "(string \ 'a rule list) list \ bool" where "sanity_wf_ruleset \ \ distinct (map fst \) \ (\ rs \ ran (map_of \). (\r \ set rs. case get_action r of Accept \ True | Drop \ True | Reject \ True | Log \ True | Empty \ True | Call chain \ chain \ dom (map_of \) | Goto chain \ chain \ dom (map_of \) | Return \ True | _ \ False))" lemma sanity_wf_ruleset_wf_chain: "sanity_wf_ruleset \ \ rs \ ran (map_of \) \ wf_chain (map_of \) rs" apply(simp add: sanity_wf_ruleset_def wf_chain_def) by fastforce lemma sanity_wf_ruleset_start: "sanity_wf_ruleset \ \ chain_name \ dom (map_of \) \ default_action = Accept \ default_action = Drop \ wf_chain (map_of \) [Rule MatchAny (Call chain_name), Rule MatchAny default_action]" apply(simp add: sanity_wf_ruleset_def wf_chain_def) apply(safe) apply(simp_all) apply blast+ done lemma [code]: "sanity_wf_ruleset \ = (let dom = map fst \; ran = map snd \ in distinct dom \ (\ rs \ set ran. (\r \ set rs. case get_action r of Accept \ True | Drop \ True | Reject \ True | Log \ True | Empty \ True | Call chain \ chain \ set dom | Goto chain \ chain \ set dom | Return \ True | _ \ False)))" proof - have set_map_fst: "set (map fst \) = dom (map_of \)" by (simp add: dom_map_of_conv_image_fst) have set_map_snd: "distinct (map fst \) \ set (map snd \) = ran (map_of \)" by (simp add: ran_distinct) show ?thesis unfolding sanity_wf_ruleset_def Let_def apply(subst set_map_fst)+ apply(rule iffI) apply(elim conjE) apply(subst set_map_snd) apply(simp) apply(simp) apply(elim conjE) apply(subst(asm) set_map_snd) apply(simp_all) done qed lemma semantics_bigstep_defined1: assumes "\rsg \ ran \ \ {rs}. wf_chain \ rsg" and "\rsg \ ran \ \ {rs}. \ r \ set rsg. (\chain. get_action r \ Goto chain) \ get_action r \ Unknown" and "\ r \ set rs. get_action r \ Return" (*no toplevel return*) and "(\name \ dom \. \t. \,\,p\ \the (\ name), Undecided\ \ t)" (*defined for all chains in the background ruleset*) shows "\t. \,\,p\ \rs, s\ \ t" using assms proof(induction rs) case Nil thus ?case apply(rule_tac x=s in exI) by(simp add: skip) next case (Cons r rs) from Cons.prems Cons.IH obtain t' where t': "\,\,p\ \rs, s\ \ t'" apply simp apply(elim conjE) apply(simp add: wf_chain_fst) by blast obtain m a where r: "r = Rule m a" by(cases r) blast show ?case proof(cases "matches \ m p") case False hence "\,\,p\ \[r], s\ \ s" apply(cases s) apply(simp add: nomatch r) by(simp add: decision) thus ?thesis apply(rule_tac x=t' in exI) apply(rule_tac t=s in seq'_cons) apply assumption using t' by(simp) next case True show ?thesis proof(cases s) case (Decision X) thus ?thesis apply(rule_tac x="Decision X" in exI) by(simp add: decision) next case Undecided have "\t. \,\,p\ \Rule m a # rs, Undecided\ \ t" proof(cases a) case Accept with True show ?thesis apply(rule_tac x="Decision FinalAllow" in exI) apply(rule_tac t="Decision FinalAllow" in seq'_cons) by(auto intro: iptables_bigstep.intros) next case Drop with True show ?thesis apply(rule_tac x="Decision FinalDeny" in exI) apply(rule_tac t="Decision FinalDeny" in seq'_cons) by(auto intro: iptables_bigstep.intros) next case Log with True t' Undecided show ?thesis apply(rule_tac x=t' in exI) apply(rule_tac t=Undecided in seq'_cons) by(auto intro: iptables_bigstep.intros) next case Reject with True show ?thesis apply(rule_tac x="Decision FinalDeny" in exI) apply(rule_tac t="Decision FinalDeny" in seq'_cons) by(auto intro: iptables_bigstep.intros)[2] next case Return with Cons.prems(3)[simplified r] show ?thesis by simp next case Goto with Cons.prems(2)[simplified r] show ?thesis by auto next case (Call chain_name) from Call Cons.prems(1) obtain rs' where 1: "\ chain_name = Some rs'" by(simp add: r wf_chain_def) blast with Cons.prems(4) obtain t'' where 2: "\,\,p\ \the (\ chain_name), Undecided\ \ t''" by blast from 1 2 True have "\,\,p\ \[Rule m (Call chain_name)], Undecided\ \ t''" by(auto dest: call_result) with Call t' Undecided show ?thesis apply(simp add: r) apply(cases t'') apply simp apply(rule_tac x=t' in exI) apply(rule_tac t=Undecided in seq'_cons) apply(auto intro: iptables_bigstep.intros)[2] apply(simp) apply(rule_tac x=t'' in exI) apply(rule_tac t=t'' in seq'_cons) apply(auto intro: iptables_bigstep.intros) done next case Empty with True t' Undecided show ?thesis apply(rule_tac x=t' in exI) apply(rule_tac t=Undecided in seq'_cons) by(auto intro: iptables_bigstep.intros) next case Unknown with Cons.prems(2)[simplified r] show ?thesis by(simp) qed thus ?thesis unfolding r Undecided by simp qed qed qed text\Showing the main theorem\ context begin private lemma iptables_bigstep_defined_if_singleton_rules: "\ r \ set rs. (\t. \,\,p\ \[r], s\ \ t) \ \t. \,\,p\ \rs, s\ \ t" proof(induction rs arbitrary: s) case Nil hence "\,\,p\ \[], s\ \ s" by(simp add: skip) thus ?case by blast next case(Cons r rs s) from Cons.prems obtain t where t: "\,\,p\ \[r], s\ \ t" by simp blast with Cons show ?case proof(cases t) case Decision with t show ?thesis by (meson decision seq'_cons) next case Undecided from Cons obtain t' where t': "\,\,p\ \rs, s\ \ t'" by simp blast with Undecided t show ?thesis apply(rule_tac x=t' in exI) apply(rule seq'_cons) apply(simp) using iptables_bigstep_to_undecided by fastforce qed qed text\well founded relation.\ definition calls_chain :: "'a ruleset \ (string \ string) set" where "calls_chain \ = {(r, s). case \ r of Some rs \ \m. Rule m (Call s) \ set rs | None \ False}" lemma calls_chain_def2: "calls_chain \ = {(caller, callee). \rs m. \ caller = Some rs \ Rule m (Call callee) \ set rs}" unfolding calls_chain_def apply(safe) apply(simp split: option.split_asm) apply(simp) by blast text\example\ private lemma "calls_chain [ ''FORWARD'' \ [(Rule m1 Log), (Rule m2 (Call ''foo'')), (Rule m3 Accept), (Rule m' (Call ''baz''))], ''foo'' \ [(Rule m4 Log), (Rule m5 Return), (Rule m6 (Call ''bar''))], ''bar'' \ [], ''baz'' \ []] = {(''FORWARD'', ''foo''), (''FORWARD'', ''baz''), (''foo'', ''bar'')}" unfolding calls_chain_def by(auto split: option.split_asm if_split_asm) private lemma "wf (calls_chain [ ''FORWARD'' \ [(Rule m1 Log), (Rule m2 (Call ''foo'')), (Rule m3 Accept), (Rule m' (Call ''baz''))], ''foo'' \ [(Rule m4 Log), (Rule m5 Return), (Rule m6 (Call ''bar''))], ''bar'' \ [], ''baz'' \ []])" proof - have g: "calls_chain [''FORWARD'' \ [(Rule m1 Log), (Rule m2 (Call ''foo'')), (Rule m3 Accept), (Rule m' (Call ''baz''))], ''foo'' \ [(Rule m4 Log), (Rule m5 Return), (Rule m6 (Call ''bar''))], ''bar'' \ [], ''baz'' \ []] = {(''FORWARD'', ''foo''), (''FORWARD'', ''baz''), (''foo'', ''bar'')}" by(auto simp add: calls_chain_def split: option.split_asm if_split_asm) show ?thesis unfolding g apply(simp) apply safe apply(erule rtranclE, simp_all) apply(erule rtranclE, simp_all) done qed text\In our proof, we will need the reverse.\ private definition called_by_chain :: "'a ruleset \ (string \ string) set" where "called_by_chain \ = {(callee, caller). case \ caller of Some rs \ \m. Rule m (Call callee) \ set rs | None \ False}" private lemma called_by_chain_converse: "calls_chain \ = converse (called_by_chain \)" apply(simp add: calls_chain_def called_by_chain_def) by blast private lemma wf_called_by_chain: "finite (calls_chain \) \ wf (calls_chain \) \ wf (called_by_chain \)" apply(frule Wellfounded.wf_acyclic) apply(drule(1) Wellfounded.finite_acyclic_wf_converse) apply(simp add: called_by_chain_converse) done private lemma helper_cases_call_subchain_defined_or_return: "(\x\ran \. wf_chain \ x) \ \rsg\ran \. \r\set rsg. (\chain. get_action r \ Goto chain) \ get_action r \ Unknown \ \y m. \r\set rs_called. r = Rule m (Call y) \ (\t. \,\,p\ \[Rule m (Call y)], Undecided\ \ t) \ wf_chain \ rs_called \ \r\set rs_called. (\chain. get_action r \ Goto chain) \ get_action r \ Unknown \ (\t. \,\,p\ \rs_called, Undecided\ \ t) \ (\rs_called1 rs_called2 m'. rs_called = (rs_called1 @ [Rule m' Return] @ rs_called2) \ matches \ m' p \ \,\,p\ \rs_called1, Undecided\ \ Undecided)" proof(induction rs_called arbitrary:) case Nil hence "\t. \,\,p\ \[], Undecided\ \ t" apply(rule_tac x=Undecided in exI) by(simp add: skip) thus ?case by simp next case (Cons r rs) from Cons.prems have "wf_chain \ [r]" by(simp add: wf_chain_def) from Cons.prems have IH:"(\t'. \,\,p\ \rs, Undecided\ \ t') \ (\rs_called1 rs_called2 m'. rs = (rs_called1 @ [Rule m' Return] @ rs_called2) \ matches \ m' p \ \,\,p\ \rs_called1, Undecided\ \ Undecided)" apply - apply(rule Cons.IH) apply(auto dest: wf_chain_fst) done from Cons.prems have case_call: "r = Rule m (Call y) \ (\t. \,\,p\ \[Rule m (Call y)], Undecided\ \ t)" for y m by(simp) obtain m a where r: "r = Rule m a" by(cases r) simp from Cons.prems have a_not: "(\chain. a \ Goto chain) \ a \ Unknown" by(simp add: r) have ex_neq_ret: "a \ Return \ \t. \,\,p\ \[Rule m a], Undecided\ \ t" proof(cases "matches \ m p") case False thus ?thesis by(rule_tac x=Undecided in exI)(simp add: nomatch; fail) next case True assume "a \ Return" show ?thesis proof(cases a) case Accept with True show ?thesis by(rule_tac x="Decision FinalAllow" in exI) (simp add: accept; fail) next case Drop with True show ?thesis by(rule_tac x="Decision FinalDeny" in exI) (simp add: drop; fail) next case Log with True show ?thesis by(rule_tac x="Undecided" in exI)(simp add: log; fail) next case Reject with True show ?thesis by(rule_tac x="Decision FinalDeny" in exI) (simp add: reject; fail) next case Call with True show ?thesis apply(simp) apply(rule case_call) apply(simp add: r; fail) done next case Empty with True show ?thesis by(rule_tac x="Undecided" in exI) (simp add: empty; fail) next case Return with \a \ Return\ show ?thesis by simp qed(simp_all add: a_not) qed have *: "?case" if pre: "rs = rs_called1 @ Rule m' Return # rs_called2 \ matches \ m' p \ \,\,p\ \rs_called1, Undecided\ \ Undecided" for rs_called1 m' rs_called2 proof(cases "matches \ m p") case False thus ?thesis apply - apply(rule disjI2) apply(rule_tac x="r#rs_called1" in exI) apply(rule_tac x=rs_called2 in exI) apply(rule_tac x=m' in exI) apply(simp add: r pre) apply(rule_tac t=Undecided in seq_cons) apply(simp add: r nomatch; fail) apply(simp add: pre; fail) done next case True from pre have rule_case_dijs1: "\X. \,\,p\ \[Rule m a], Undecided\ \ Decision X \ ?thesis" apply - apply(rule disjI1) apply(elim exE conjE, rename_tac X) apply(simp) apply(rule_tac x="Decision X" in exI) apply(rule_tac t="Decision X" in seq_cons) apply(simp add: r; fail) apply(simp add: decision; fail) done from pre have rule_case_dijs2: "\,\,p\ \[Rule m a], Undecided\ \ Undecided \ ?thesis" apply - apply(rule disjI2) apply(rule_tac x="r#rs_called1" in exI) apply(rule_tac x=rs_called2 in exI) apply(rule_tac x=m' in exI) apply(simp add: r) apply(rule_tac t=Undecided in seq_cons) apply(simp; fail) apply(simp;fail) done show ?thesis proof(cases a) case Accept show ?thesis apply(rule rule_case_dijs1) apply(rule_tac x="FinalAllow" in exI) using True pre Accept by(simp add: accept) next case Drop show ?thesis apply(rule rule_case_dijs1) apply(rule_tac x="FinalDeny" in exI) using True Drop by(simp add: deny) next case Log show ?thesis apply(rule rule_case_dijs2) using Log True by(simp add: log) next case Reject show ?thesis apply(rule rule_case_dijs1) apply(rule_tac x="FinalDeny" in exI) using Reject True by(simp add: reject) next case (Call x5) have "\t. \,\,p\ \[Rule m (Call x5)], Undecided\ \ t" by(rule case_call) (simp add: r Call) with Call pre True show ?thesis apply(simp) apply(elim exE, rename_tac t_called) apply(case_tac t_called) apply(simp) apply(rule disjI2) apply(rule_tac x="r#rs_called1" in exI) apply(rule_tac x=rs_called2 in exI) apply(rule_tac x=m' in exI) apply(simp add: r) apply(rule_tac t=Undecided in seq_cons) apply(simp add: r; fail) apply(simp; fail) apply(rule disjI1) apply(rule_tac x=t_called in exI) apply(rule_tac t=t_called in seq_cons) apply(simp add: r; fail) apply(simp add: decision; fail) done next case Empty show ?thesis apply(rule rule_case_dijs2) using Empty True by(simp add: pre empty) next case Return show ?thesis apply(rule disjI2) apply(rule_tac x="[]" in exI) apply(rule_tac x="rs_called1 @ Rule m' Return # rs_called2" in exI) apply(rule_tac x=m in exI) using Return True pre by(simp add: skip r) qed(simp_all add: a_not) qed from IH have **: "a \ Return \ (\t. \,\,p\ \[Rule m a], Undecided\ \ t) \ ?case" proof(elim disjE, goal_cases) case 2 from this obtain rs_called1 m' rs_called2 where a1: "rs = rs_called1 @ [Rule m' Return] @ rs_called2" and a2: "matches \ m' p" and a3: "\,\,p\ \rs_called1, Undecided\ \ Undecided" by blast show ?case apply(rule *) using a1 a2 a3 by simp next case 1 thus ?case proof(cases "a \ Return") case True with 1 obtain t1 t2 where t1: "\,\,p\ \[Rule m a], Undecided\ \ t1" and t2: "\,\,p\ \rs, Undecided\ \ t2" by blast from t1 t2 show ?thesis apply - apply(rule disjI1) apply(simp add: r) apply(cases t1) apply(simp_all) apply(rule_tac x=t2 in exI) apply(rule_tac seq'_cons) apply(simp_all) apply (meson decision seq_cons) done next case False show ?thesis proof(cases "matches \ m p") assume "\ matches \ m p" with 1 show ?thesis apply - apply(rule disjI1) apply(elim exE) apply(rename_tac t') apply(rule_tac x=t' in exI) apply(rule_tac t=Undecided in seq_cons) apply(simp add: r nomatch; fail) by(simp) next assume "matches \ m p" with False show ?thesis apply - apply(rule disjI2) apply(rule_tac x="[]" in exI) apply(rule_tac x=rs in exI) apply(rule_tac x=m in exI) apply(simp add: r skip; fail) done qed qed qed thus ?case using ex_neq_ret by blast qed lemma helper_defined_single: assumes "wf (called_by_chain \)" and "\rsg \ ran \ \ {[Rule m a]}. wf_chain \ rsg" and "\rsg \ ran \ \ {[Rule m a]}. \ r \ set rsg. (\(\chain. get_action r = Goto chain)) \ get_action r \ Unknown" and "a \ Return" (*no toplevel Return*) shows "\t. \,\,p\ \[Rule m a], s\ \ t" proof(cases s) case (Decision decision) thus ?thesis apply(rule_tac x="Decision decision" in exI) apply(simp) using iptables_bigstep.decision by fast next case Undecided have "\t. \,\,p\ \[Rule m a], Undecided\ \ t" proof(cases "matches \ m p") case False with assms show ?thesis apply(rule_tac x=Undecided in exI) apply(rule_tac t=Undecided in seq'_cons) apply (metis empty_iff empty_set insert_iff list.simps(15) nomatch' rule.sel(1)) apply(simp add: skip; fail) done next case True show ?thesis proof(cases a) case Unknown with assms(3) show ?thesis by simp next case Goto with assms(3) show ?thesis by auto next case Accept with True show ?thesis by(auto intro: iptables_bigstep.intros) next case Drop with True show ?thesis by(auto intro: iptables_bigstep.intros) next case Reject with True show ?thesis by(auto intro: iptables_bigstep.intros) next case Log with True show ?thesis by(auto intro: iptables_bigstep.intros) next case Empty with True show ?thesis by(auto intro: iptables_bigstep.intros) next case Return with assms show ?thesis by simp next case (Call chain_name) thm wf_induct_rule[where r="(calls_chain \)" and P="\x. \t. \,\,p\ \[Rule m (Call x)], Undecided\ \ t"] \ \Only the assumptions we will need\ from assms have "wf (called_by_chain \)" "\rsg\ran \. wf_chain \ rsg" "\rsg\ran \. \r\set rsg. (\chain. get_action r \ Goto chain) \ get_action r \ Unknown" by auto \ \strengthening the IH to do a well-founded induction\ hence "matches \ m p \ wf_chain \ [Rule m (Call chain_name)] \ (\t. \,\,p\ \[Rule m (Call chain_name)], Undecided\ \ t)" proof(induction arbitrary: m rule: wf_induct_rule[where r="called_by_chain \"]) case (less chain_name_neu) from less.prems have "\ chain_name_neu \ None" by(simp add: wf_chain_def) from this obtain rs_called where rs_called: "\ chain_name_neu = Some rs_called" by blast from less rs_called have "wf_chain \ rs_called" by (simp add: ranI) from less rs_called have "rs_called \ ran \" by (simp add: ranI) (*get good IH*) from less.prems rs_called have "\y m. \r \ set rs_called. r = Rule m (Call y) \ (y, chain_name_neu) \ called_by_chain \ \ wf_chain \ [Rule m (Call y)]" apply(simp) apply(intro impI allI conjI) apply(simp add: called_by_chain_def) apply blast apply(simp add: wf_chain_def) apply (meson ranI rule.sel(2)) done with less have "\y m. \r\set rs_called. r = Rule m (Call y) \ (\t. \,\,p\ \[Rule m (Call y)], Undecided\ \ t)" apply(intro allI, rename_tac y my) apply(case_tac "matches \ my p") apply blast apply(intro ballI impI) apply(rule_tac x=Undecided in exI) apply(simp add: nomatch; fail) done from less.prems(4) rs_called \rs_called \ ran \\ helper_cases_call_subchain_defined_or_return[OF less.prems(3) less.prems(4) this \wf_chain \ rs_called\] have "(\t. \,\,p\ \rs_called, Undecided\ \ t) \ (\rs_called1 rs_called2 m'. \ chain_name_neu = Some (rs_called1@[Rule m' Return]@rs_called2) \ matches \ m' p \ \,\,p\ \rs_called1, Undecided\ \ Undecided)" by simp thus ?case proof(elim disjE exE conjE) fix t assume a: "\,\,p\ \rs_called, Undecided\ \ t" show ?case using call_result[OF less.prems(1) rs_called a] by(blast) next fix m' rs_called1 rs_called2 assume a1: "\ chain_name_neu = Some (rs_called1 @ [Rule m' Return] @ rs_called2)" and a2: "matches \ m' p" and a3: "\,\,p\ \rs_called1, Undecided\ \ Undecided" show ?case using call_return[OF less.prems(1) a1 a2 a3 ] by(blast) qed qed with True assms Call show ?thesis by simp qed qed with Undecided show ?thesis by simp qed private lemma helper_defined_ruleset_calledby: "wf (called_by_chain \) \ \rsg \ ran \ \ {rs}. wf_chain \ rsg \ \rsg \ ran \ \ {rs}. \ r \ set rsg. (\(\chain. get_action r = Goto chain)) \ get_action r \ Unknown \ \ r \ set rs. get_action r \ Return \ \t. \,\,p\ \rs, s\ \ t" apply(rule iptables_bigstep_defined_if_singleton_rules) apply(intro ballI, rename_tac r, case_tac r, rename_tac m a, simp) apply(rule helper_defined_single) apply(simp; fail) apply(simp add: wf_chain_def; fail) apply fastforce apply fastforce done corollary semantics_bigstep_defined: "finite (calls_chain \) \ wf (calls_chain \) \ \ \call relation finite and terminating\ \rsg \ ran \ \ {rs}. wf_chain \ rsg \ \ \All calls to defined chains\ \rsg \ ran \ \ {rs}. \ r \ set rsg. (\x. get_action r \ Goto x) \ get_action r \ Unknown \ \ \no bad actions\ \ r \ set rs. get_action r \ Return \ \no toplevel return\ \ \t. \,\,p\ \rs, s\ \ t" apply(drule(1) wf_called_by_chain) apply(thin_tac "wf (calls_chain \)") apply(rule helper_defined_ruleset_calledby) apply(simp_all) done end text\Common Algorithms\ lemma iptables_bigstep_rm_LogEmpty: "\,\,p\ \rm_LogEmpty rs, s\ \ t \ \,\,p\ \rs, s\ \ t" proof(induction rs arbitrary: s) case Nil thus ?case by(simp) next case (Cons r rs) have step_IH: "(\s. \,\,p\ \rs1, s\ \ t = \,\,p\ \rs2, s\ \ t) \ \,\,p\ \r#rs1, s\ \ t = \,\,p\ \r#rs2, s\ \ t" for rs1 rs2 r by (meson seq'_cons seqE_cons) have case_log: "\,\,p\ \Rule m Log # rs, s\ \ t \ \,\,p\ \rs, s\ \ t" for m apply(rule iffI) apply(erule seqE_cons) apply (metis append_Nil log_remove seq') apply(rule_tac t=s in seq'_cons) apply(cases s) apply(cases "matches \ m p") apply(simp add: log; fail) apply(simp add: nomatch; fail) apply(simp add: decision; fail) apply simp done have case_empty: "\,\,p\ \Rule m Empty # rs, s\ \ t \ \,\,p\ \rs, s\ \ t" for m apply(rule iffI) apply(erule seqE_cons) apply (metis append_Nil empty_empty seq') apply(rule_tac t=s in seq'_cons) apply(cases s) apply(cases "matches \ m p") apply(simp add: empty; fail) apply(simp add: nomatch; fail) apply(simp add: decision; fail) apply simp done from Cons show ?case apply(cases r, rename_tac m a) apply(case_tac a) apply(simp_all) apply(simp_all cong: step_IH) apply(simp_all add: case_log case_empty) done qed lemma iptables_bigstep_rw_Reject: "\,\,p\ \rw_Reject rs, s\ \ t \ \,\,p\ \rs, s\ \ t" proof(induction rs arbitrary: s) case Nil thus ?case by(simp) next case (Cons r rs) have step_IH: "(\s. \,\,p\ \rs1, s\ \ t = \,\,p\ \rs2, s\ \ t) \ \,\,p\ \r#rs1, s\ \ t = \,\,p\ \r#rs2, s\ \ t" for rs1 rs2 r by (meson seq'_cons seqE_cons) have fst_rule: "(\t. \,\,p\ \[r1], s\ \ t \ \,\,p\ \[r2], s\ \ t) \ \,\,p\ \r1 # rs, s\ \ t \ \,\,p\ \r2 # rs, s\ \ t" for r1 r2 rs s t by (meson seq'_cons seqE_cons) have dropreject: "\,\,p\ \[Rule m Drop], s\ \ t = \,\,p\ \[Rule m Reject], s\ \ t" for m t apply(cases s) apply(cases "matches \ m p") using drop reject dropD rejectD apply fast using nomatch nomatchD apply fast using decision decisionD apply fast done from Cons show ?case apply(cases r, rename_tac m a) apply simp apply(case_tac a) apply(simp_all) apply(simp_all cong: step_IH) apply(rule fst_rule) apply(simp add: dropreject) done qed end diff --git a/thys/Jacobson_Basic_Algebra/Group_Theory.thy b/thys/Jacobson_Basic_Algebra/Group_Theory.thy --- a/thys/Jacobson_Basic_Algebra/Group_Theory.thy +++ b/thys/Jacobson_Basic_Algebra/Group_Theory.thy @@ -1,1700 +1,1700 @@ (* Copyright (c) 2014-2019 by Clemens Ballarin This file is licensed under the 3-clause BSD license. *) theory Group_Theory imports Set_Theory begin hide_const monoid hide_const group hide_const inverse no_notation quotient (infixl "'/'/" 90) section \Monoids and Groups\ subsection \Monoids of Transformations and Abstract Monoids\ text \Def 1.1\ text \p 28, ll 28--30\ locale monoid = fixes M and composition (infixl "\" 70) and unit ("\") assumes composition_closed [intro, simp]: "\ a \ M; b \ M \ \ a \ b \ M" and unit_closed [intro, simp]: "\ \ M" and associative [intro]: "\ a \ M; b \ M; c \ M \ \ (a \ b) \ c = a \ (b \ c)" and left_unit [intro, simp]: "a \ M \ \ \ a = a" and right_unit [intro, simp]: "a \ M \ a \ \ = a" text \p 29, ll 27--28\ locale submonoid = monoid M "(\)" \ for N and M and composition (infixl "\" 70) and unit ("\") + assumes subset [intro, simp]: "N \ M" and sub_composition_closed: "\ a \ N; b \ N \ \ a \ b \ N" and sub_unit_closed: "\ \ N" begin text \p 29, ll 27--28\ lemma sub [intro, simp]: "a \ N \ a \ M" using subset by blast text \p 29, ll 32--33\ sublocale sub: monoid N "(\)" \ by unfold_locales (auto simp: sub_composition_closed sub_unit_closed) end (* submonoid *) text \p 29, ll 33--34\ theorem submonoid_transitive: assumes "submonoid K N composition unit" and "submonoid N M composition unit" shows "submonoid K M composition unit" proof - interpret K: submonoid K N composition unit by fact interpret M: submonoid N M composition unit by fact show ?thesis by unfold_locales auto qed text \p 28, l 23\ locale transformations = fixes S :: "'a set" (* assumes non_vacuous: "S \ {}" *) (* Jacobson requires this but we don't need it, strange. *) text \Monoid of all transformations\ text \p 28, ll 23--24\ sublocale transformations \ monoid "S \\<^sub>E S" "compose S" "identity S" by unfold_locales (auto simp: PiE_def compose_eq compose_assoc Id_compose compose_Id) text \@{term N} is a monoid of transformations of the set @{term S}.\ text \p 29, ll 34--36\ locale transformation_monoid = transformations S + submonoid M "S \\<^sub>E S" "compose S" "identity S" for M and S begin text \p 29, ll 34--36\ lemma transformation_closed [intro, simp]: "\ \ \ M; x \ S \ \ \ x \ S" by (metis PiE_iff sub) text \p 29, ll 34--36\ lemma transformation_undefined [intro, simp]: "\ \ \ M; x \ S \ \ \ x = undefined" by (metis PiE_arb sub) end (* transformation_monoid *) subsection \Groups of Transformations and Abstract Groups\ context monoid begin text \Invertible elements\ text \p 31, ll 3--5\ definition invertible where "u \ M \ invertible u \ (\v \ M. u \ v = \ \ v \ u = \)" text \p 31, ll 3--5\ lemma invertibleI [intro]: "\ u \ v = \; v \ u = \; u \ M; v \ M \ \ invertible u" unfolding invertible_def by fast text \p 31, ll 3--5\ lemma invertibleE [elim]: "\ invertible u; \v. \ u \ v = \ \ v \ u = \; v \ M \ \ P; u \ M \ \ P" unfolding invertible_def by fast text \p 31, ll 6--7\ theorem inverse_unique: "\ u \ v' = \; v \ u = \; u \ M; v \ M; v' \ M \ \ v = v'" by (metis associative left_unit right_unit) text \p 31, l 7\ definition inverse where "inverse = (\u \ M. THE v. v \ M \ u \ v = \ \ v \ u = \)" text \p 31, l 7\ theorem inverse_equality: "\ u \ v = \; v \ u = \; u \ M; v \ M \ \ inverse u = v" unfolding inverse_def using inverse_unique by simp blast text \p 31, l 7\ lemma invertible_inverse_closed [intro, simp]: "\ invertible u; u \ M \ \ inverse u \ M" using inverse_equality by auto text \p 31, l 7\ lemma inverse_undefined [intro, simp]: "u \ M \ inverse u = undefined" by (simp add: inverse_def) text \p 31, l 7\ lemma invertible_left_inverse [simp]: "\ invertible u; u \ M \ \ inverse u \ u = \" using inverse_equality by auto text \p 31, l 7\ lemma invertible_right_inverse [simp]: "\ invertible u; u \ M \ \ u \ inverse u = \" using inverse_equality by auto text \p 31, l 7\ lemma invertible_left_cancel [simp]: "\ invertible x; x \ M; y \ M; z \ M \ \ x \ y = x \ z \ y = z" by (metis associative invertible_def left_unit) text \p 31, l 7\ lemma invertible_right_cancel [simp]: "\ invertible x; x \ M; y \ M; z \ M \ \ y \ x = z \ x \ y = z" by (metis associative invertible_def right_unit) text \p 31, l 7\ lemma inverse_unit [simp]: "inverse \ = \" using inverse_equality by blast text \p 31, ll 7--8\ theorem invertible_inverse_invertible [intro, simp]: "\ invertible u; u \ M \ \ invertible (inverse u)" using invertible_left_inverse invertible_right_inverse by blast text \p 31, l 8\ theorem invertible_inverse_inverse [simp]: "\ invertible u; u \ M \ \ inverse (inverse u) = u" by (simp add: inverse_equality) end (* monoid *) context submonoid begin text \Reasoning about @{term invertible} and @{term inverse} in submonoids.\ text \p 31, l 7\ lemma submonoid_invertible [intro, simp]: "\ sub.invertible u; u \ N \ \ invertible u" using invertibleI by blast text \p 31, l 7\ lemma submonoid_inverse_closed [intro, simp]: "\ sub.invertible u; u \ N \ \ inverse u \ N" using inverse_equality by auto end (* submonoid *) text \Def 1.2\ text \p 31, ll 9--10\ locale group = monoid G "(\)" \ for G and composition (infixl "\" 70) and unit ("\") + assumes invertible [simp, intro]: "u \ G \ invertible u" text \p 31, ll 11--12\ locale subgroup = submonoid G M "(\)" \ + sub: group G "(\)" \ for G and M and composition (infixl "\" 70) and unit ("\") begin text \Reasoning about @{term invertible} and @{term inverse} in subgroups.\ text \p 31, ll 11--12\ lemma subgroup_inverse_equality [simp]: "u \ G \ inverse u = sub.inverse u" by (simp add: inverse_equality) text \p 31, ll 11--12\ lemma subgroup_inverse_iff [simp]: "\ invertible x; x \ M \ \ inverse x \ G \ x \ G" using invertible_inverse_inverse sub.invertible_inverse_closed by fastforce end (* subgroup *) text \p 31, ll 11--12\ lemma subgroup_transitive [trans]: assumes "subgroup K H composition unit" and "subgroup H G composition unit" shows "subgroup K G composition unit" proof - interpret K: subgroup K H composition unit by fact interpret H: subgroup H G composition unit by fact show ?thesis by unfold_locales auto qed context monoid begin text \Jacobson states both directions, but the other one is trivial.\ text \p 31, ll 12--15\ theorem subgroupI: fixes G assumes subset [THEN subsetD, intro]: "G \ M" and [intro]: "\ \ G" and [intro]: "\g h. \ g \ G; h \ G \ \ g \ h \ G" and [intro]: "\g. g \ G \ invertible g" and [intro]: "\g. g \ G \ inverse g \ G" shows "subgroup G M (\) \" proof - interpret sub: monoid G "(\)" \ by unfold_locales auto show ?thesis proof unfold_locales fix u assume [intro]: "u \ G" show "sub.invertible u" using invertible_left_inverse invertible_right_inverse by blast qed auto qed text \p 31, l 16\ definition "Units = {u \ M. invertible u}" text \p 31, l 16\ lemma mem_UnitsI: "\ invertible u; u \ M \ \ u \ Units" unfolding Units_def by clarify text \p 31, l 16\ lemma mem_UnitsD: "\ u \ Units \ \ invertible u \ u \ M" unfolding Units_def by clarify text \p 31, ll 16--21\ interpretation units: subgroup Units M proof (rule subgroupI) fix u1 u2 assume Units [THEN mem_UnitsD, simp]: "u1 \ Units" "u2 \ Units" have "(u1 \ u2) \ (inverse u2 \ inverse u1) = (u1 \ (u2 \ inverse u2)) \ inverse u1" by (simp add: associative del: invertible_left_inverse invertible_right_inverse) also have "\ = \" by simp finally have inv1: "(u1 \ u2) \ (inverse u2 \ inverse u1) = \" by simp \ \ll 16--18\ have "(inverse u2 \ inverse u1) \ (u1 \ u2) = (inverse u2 \ (inverse u1 \ u1)) \ u2" by (simp add: associative del: invertible_left_inverse invertible_right_inverse) also have "\ = \" by simp finally have inv2: "(inverse u2 \ inverse u1) \ (u1 \ u2) = \" by simp \ \l 9, “and similarly”\ show "u1 \ u2 \ Units" using inv1 inv2 invertibleI mem_UnitsI by auto qed (auto simp: Units_def) text \p 31, ll 21--22\ theorem group_of_Units [intro, simp]: "group Units (\) \" .. text \p 31, l 19\ lemma composition_invertible [simp, intro]: "\ invertible x; invertible y; x \ M; y \ M \ \ invertible (x \ y)" using mem_UnitsD mem_UnitsI by blast text \p 31, l 20\ lemma unit_invertible: "invertible \" by fast text \Useful simplification rules\ text \p 31, l 22\ lemma invertible_right_inverse2: "\ invertible u; u \ M; v \ M \ \ u \ (inverse u \ v) = v" by (simp add: associative [THEN sym]) text \p 31, l 22\ lemma invertible_left_inverse2: "\ invertible u; u \ M; v \ M \ \ inverse u \ (u \ v) = v" by (simp add: associative [THEN sym]) text \p 31, l 22\ lemma inverse_composition_commute: assumes [simp]: "invertible x" "invertible y" "x \ M" "y \ M" shows "inverse (x \ y) = inverse y \ inverse x" proof - have "inverse (x \ y) \ (x \ y) = (inverse y \ inverse x) \ (x \ y)" by (simp add: invertible_left_inverse2 associative) then show ?thesis by (simp del: invertible_left_inverse) qed end (* monoid *) text \p 31, l 24\ context transformations begin text \p 31, ll 25--26\ theorem invertible_is_bijective: assumes dom: "\ \ S \\<^sub>E S" shows "invertible \ \ bij_betw \ S S" proof - from dom interpret map \ S S by unfold_locales show ?thesis by (auto simp add: bij_betw_iff_has_inverse invertible_def) qed text \p 31, ll 26--27\ theorem Units_bijective: "Units = {\ \ S \\<^sub>E S. bij_betw \ S S}" unfolding Units_def by (auto simp add: invertible_is_bijective) text \p 31, ll 26--27\ lemma Units_bij_betwI [intro, simp]: "\ \ Units \ bij_betw \ S S" by (simp add: Units_bijective) text \p 31, ll 26--27\ lemma Units_bij_betwD [dest, simp]: "\ \ \ S \\<^sub>E S; bij_betw \ S S \ \ \ \ Units" unfolding Units_bijective by simp text \p 31, ll 28--29\ abbreviation "Sym \ Units" text \p 31, ll 26--28\ sublocale symmetric: group "Sym" "compose S" "identity S" by (fact group_of_Units) end (* transformations *) text \p 32, ll 18--19\ locale transformation_group = transformations S + symmetric: subgroup G Sym "compose S" "identity S" for G and S begin text \p 32, ll 18--19\ lemma transformation_group_closed [intro, simp]: "\ \ \ G; x \ S \ \ \ x \ S" using bij_betwE by blast text \p 32, ll 18--19\ lemma transformation_group_undefined [intro, simp]: "\ \ \ G; x \ S \ \ \ x = undefined" by (metis compose_def symmetric.sub.right_unit restrict_apply) end (* transformation_group *) subsection \Isomorphisms. Cayley's Theorem\ text \Def 1.3\ text \p 37, ll 7--11\ locale monoid_isomorphism = bijective_map \ M M' + source: monoid M "(\)" \ + target: monoid M' "(\')" "\'" for \ and M and composition (infixl "\" 70) and unit ("\") and M' and composition' (infixl "\''" 70) and unit' ("\''") + assumes commutes_with_composition: "\ x \ M; y \ M \ \ \ x \' \ y = \ (x \ y)" and commutes_with_unit: "\ \ = \'" text \p 37, l 10\ definition isomorphic_as_monoids (infixl "\\<^sub>M" 50) where "\ \\<^sub>M \' \ (let (M, composition, unit) = \; (M', composition', unit') = \' in (\\. monoid_isomorphism \ M composition unit M' composition' unit'))" text \p 37, ll 11--12\ locale monoid_isomorphism' = bijective_map \ M M' + source: monoid M "(\)" \ + target: monoid M' "(\')" "\'" for \ and M and composition (infixl "\" 70) and unit ("\") and M' and composition' (infixl "\''" 70) and unit' ("\''") + assumes commutes_with_composition: "\ x \ M; y \ M \ \ \ x \' \ y = \ (x \ y)" text \p 37, ll 11--12\ sublocale monoid_isomorphism \ monoid_isomorphism' by unfold_locales (simp add: commutes_with_composition) text \Both definitions are equivalent.\ text \p 37, ll 12--15\ sublocale monoid_isomorphism' \ monoid_isomorphism proof unfold_locales { fix y assume "y \ M'" then obtain x where "\ x = y" "x \ M" by (metis image_iff surjective) then have "y \' \ \ = y" using commutes_with_composition by auto } then show "\ \ = \'" by fastforce qed (simp add: commutes_with_composition) context monoid_isomorphism begin text \p 37, ll 30--33\ theorem inverse_monoid_isomorphism: "monoid_isomorphism (restrict (inv_into M \) M') M' (\') \' M (\) \" using commutes_with_composition commutes_with_unit surjective by unfold_locales auto end (* monoid_isomorphism *) text \We only need that @{term \} is symmetric.\ text \p 37, ll 28--29\ theorem isomorphic_as_monoids_symmetric: "(M, composition, unit) \\<^sub>M (M', composition', unit') \ (M', composition', unit') \\<^sub>M (M, composition, unit)" by (simp add: isomorphic_as_monoids_def) (meson monoid_isomorphism.inverse_monoid_isomorphism) text \p 38, l 4\ locale left_translations_of_monoid = monoid begin (* We take the liberty of omitting "left_" from the name of the translation operation. The derived transformation monoid and group won't be qualified with "left" either. This avoids qualifications such as "left.left_...". In contexts where left and right translations are used simultaneously, notably subgroup_of_group, qualifiers are needed. *) text \p 38, ll 5--7\ definition translation ("'(_')\<^sub>L") where "translation = (\a \ M. \x \ M. a \ x)" text \p 38, ll 5--7\ lemma translation_map [intro, simp]: "a \ M \ (a)\<^sub>L \ M \\<^sub>E M" unfolding translation_def by simp text \p 38, ll 5--7\ lemma Translations_maps [intro, simp]: "translation ` M \ M \\<^sub>E M" by (simp add: image_subsetI) text \p 38, ll 5--7\ lemma translation_apply: "\ a \ M; b \ M \ \ (a)\<^sub>L b = a \ b" unfolding translation_def by auto text \p 38, ll 5--7\ lemma translation_exist: "f \ translation ` M \ \a \ M. f = (a)\<^sub>L" by auto text \p 38, ll 5--7\ lemmas Translations_E [elim] = translation_exist [THEN bexE] text \p 38, l 10\ theorem translation_unit_eq [simp]: "identity M = (\)\<^sub>L" unfolding translation_def by auto text \p 38, ll 10--11\ theorem translation_composition_eq [simp]: assumes [simp]: "a \ M" "b \ M" shows "compose M (a)\<^sub>L (b)\<^sub>L = (a \ b)\<^sub>L" unfolding translation_def by rule (simp add: associative compose_def) (* Activate @{locale monoid} to simplify subsequent proof. *) text \p 38, ll 7--9\ sublocale transformation: transformations M . text \p 38, ll 7--9\ theorem Translations_transformation_monoid: "transformation_monoid (translation ` M) M" by unfold_locales auto text \p 38, ll 7--9\ sublocale transformation: transformation_monoid "translation ` M" M by (fact Translations_transformation_monoid) text \p 38, l 12\ sublocale map translation M "translation ` M" by unfold_locales (simp add: translation_def) text \p 38, ll 12--16\ theorem translation_isomorphism [intro]: "monoid_isomorphism translation M (\) \ (translation ` M) (compose M) (identity M)" proof unfold_locales have "inj_on translation M" proof (rule inj_onI) fix a b assume [simp]: "a \ M" "b \ M" "(a)\<^sub>L = (b)\<^sub>L" have "(a)\<^sub>L \ = (b)\<^sub>L \" by simp then show "a = b" by (simp add: translation_def) qed then show "bij_betw translation M (translation ` M)" by (simp add: inj_on_imp_bij_betw) qed simp_all text \p 38, ll 12--16\ sublocale monoid_isomorphism translation M "(\)" \ "translation ` M" "compose M" "identity M" .. end (* left_translations_of_monoid *) context monoid begin text \p 38, ll 1--2\ interpretation left_translations_of_monoid .. text \p 38, ll 1--2\ theorem cayley_monoid: "\M' composition' unit'. transformation_monoid M' M \ (M, (\), \) \\<^sub>M (M', composition', unit')" by (simp add: isomorphic_as_monoids_def) (fast intro: Translations_transformation_monoid) end (* monoid *) text \p 38, l 17\ locale left_translations_of_group = group begin text \p 38, ll 17--18\ sublocale left_translations_of_monoid where M = G .. text \p 38, ll 17--18\ notation translation ("'(_')\<^sub>L") text \ The group of left translations is a subgroup of the symmetric group, hence @{term transformation.sub.invertible}. \ text \p 38, ll 20--22\ theorem translation_invertible [intro, simp]: assumes [simp]: "a \ G" shows "transformation.sub.invertible (a)\<^sub>L" proof show "compose G (a)\<^sub>L (inverse a)\<^sub>L = identity G" by simp next show "compose G (inverse a)\<^sub>L (a)\<^sub>L = identity G" by simp qed auto text \p 38, ll 19--20\ theorem translation_bijective [intro, simp]: "a \ G \ bij_betw (a)\<^sub>L G G" by (blast intro: transformation.invertible_is_bijective [THEN iffD1]) text \p 38, ll 18--20\ theorem Translations_transformation_group: "transformation_group (translation ` G) G" proof unfold_locales show "(translation ` G) \ transformation.Sym" unfolding transformation.Units_bijective by auto next fix \ assume \: "\ \ translation ` G" then obtain a where a: "a \ G" and eq: "\ = (a)\<^sub>L" .. with translation_invertible show "transformation.sub.invertible \" by simp qed auto text \p 38, ll 18--20\ sublocale transformation: transformation_group "translation ` G" G by (fact Translations_transformation_group) end (* left_translations_of_group *) context group begin text \p 38, ll 2--3\ interpretation left_translations_of_group .. text \p 38, ll 2--3\ theorem cayley_group: "\G' composition' unit'. transformation_group G' G \ (G, (\), \) \\<^sub>M (G', composition', unit')" by (simp add: isomorphic_as_monoids_def) (fast intro: Translations_transformation_group) end (* group *) text \Exercise 3\ text \p 39, ll 9--10\ locale right_translations_of_group = group begin text \p 39, ll 9--10\ definition translation ("'(_')\<^sub>R") where "translation = (\a \ G. \x \ G. x \ a)" text \p 39, ll 9--10\ abbreviation "Translations \ translation ` G" text \The isomorphism that will be established is a map different from @{term translation}.\ text \p 39, ll 9--10\ interpretation aux: map translation G Translations by unfold_locales (simp add: translation_def) text \p 39, ll 9--10\ lemma translation_map [intro, simp]: "a \ G \ (a)\<^sub>R \ G \\<^sub>E G" unfolding translation_def by simp text \p 39, ll 9--10\ lemma Translation_maps [intro, simp]: "Translations \ G \\<^sub>E G" by (simp add: image_subsetI) text \p 39, ll 9--10\ lemma translation_apply: "\ a \ G; b \ G \ \ (a)\<^sub>R b = b \ a" unfolding translation_def by auto text \p 39, ll 9--10\ lemma translation_exist: "f \ Translations \ \a \ G. f = (a)\<^sub>R" by auto text \p 39, ll 9--10\ lemmas Translations_E [elim] = translation_exist [THEN bexE] text \p 39, ll 9--10\ lemma translation_unit_eq [simp]: "identity G = (\)\<^sub>R" unfolding translation_def by auto text \p 39, ll 10--11\ lemma translation_composition_eq [simp]: assumes [simp]: "a \ G" "b \ G" shows "compose G (a)\<^sub>R (b)\<^sub>R = (b \ a)\<^sub>R" unfolding translation_def by rule (simp add: associative compose_def) text \p 39, ll 10--11\ sublocale transformation: transformations G . text \p 39, ll 10--11\ lemma Translations_transformation_monoid: "transformation_monoid Translations G" by unfold_locales auto text \p 39, ll 10--11\ sublocale transformation: transformation_monoid Translations G by (fact Translations_transformation_monoid) text \p 39, ll 10--11\ lemma translation_invertible [intro, simp]: assumes [simp]: "a \ G" shows "transformation.sub.invertible (a)\<^sub>R" proof show "compose G (a)\<^sub>R (inverse a)\<^sub>R = identity G" by simp next show "compose G (inverse a)\<^sub>R (a)\<^sub>R = identity G" by simp qed auto text \p 39, ll 10--11\ lemma translation_bijective [intro, simp]: "a \ G \ bij_betw (a)\<^sub>R G G" by (blast intro: transformation.invertible_is_bijective [THEN iffD1]) text \p 39, ll 10--11\ theorem Translations_transformation_group: "transformation_group Translations G" proof unfold_locales show "Translations \ transformation.Sym" unfolding transformation.Units_bijective by auto next fix \ assume \: "\ \ Translations" then obtain a where a: "a \ G" and eq: "\ = (a)\<^sub>R" .. with translation_invertible show "transformation.sub.invertible \" by simp qed auto text \p 39, ll 10--11\ sublocale transformation: transformation_group Translations G by (rule Translations_transformation_group) text \p 39, ll 10--11\ lemma translation_inverse_eq [simp]: assumes [simp]: "a \ G" shows "transformation.sub.inverse (a)\<^sub>R = (inverse a)\<^sub>R" proof (rule transformation.sub.inverse_equality) show "compose G (a)\<^sub>R (inverse a)\<^sub>R = identity G" by simp next show "compose G (inverse a)\<^sub>R (a)\<^sub>R = identity G" by simp qed auto text \p 39, ll 10--11\ theorem translation_inverse_monoid_isomorphism [intro]: "monoid_isomorphism (\a\G. transformation.symmetric.inverse (a)\<^sub>R) G (\) \ Translations (compose G) (identity G)" (is "monoid_isomorphism ?inv _ _ _ _ _ _") proof unfold_locales show "?inv \ G \\<^sub>E Translations" by (simp del: translation_unit_eq) next note bij_betw_compose [trans] have "bij_betw inverse G G" by (rule bij_betwI [where g = inverse]) auto also have "bij_betw translation G Translations" by (rule bij_betwI [where g = "\\\Translations. \ \"]) (auto simp: translation_apply) finally show "bij_betw ?inv G Translations" by (simp cong: bij_betw_cong add: compose_eq del: translation_unit_eq) next fix x and y assume [simp]: "x \ G" "y \ G" show "compose G (?inv x) (?inv y) = (?inv (x \ y))" by (simp add: inverse_composition_commute del: translation_unit_eq) next show "?inv \ = identity G" by (simp del: translation_unit_eq) simp qed text \p 39, ll 10--11\ sublocale monoid_isomorphism "\a\G. transformation.symmetric.inverse (a)\<^sub>R" G "(\)" \ Translations "compose G" "identity G" .. end (* right_translations_of_group *) subsection \Generalized Associativity. Commutativity\ text \p 40, l 27; p 41, ll 1--2\ locale commutative_monoid = monoid + assumes commutative: "\ x \ M; y \ M \ \ x \ y = y \ x" text \p 41, l 2\ locale abelian_group = group + commutative_monoid G "(\)" \ subsection \Orbits. Cosets of a Subgroup\ context transformation_group begin text \p 51, ll 18--20\ definition Orbit_Relation where "Orbit_Relation = {(x, y). x \ S \ y \ S \ (\\ \ G. y = \ x)}" text \p 51, ll 18--20\ lemma Orbit_Relation_memI [intro]: "\ \\ \ G. y = \ x; x \ S \ \ (x, y) \ Orbit_Relation" unfolding Orbit_Relation_def by auto text \p 51, ll 18--20\ lemma Orbit_Relation_memE [elim]: "\ (x, y) \ Orbit_Relation; \\. \ \ \ G; x \ S; y = \ x \ \ Q \ \ Q" unfolding Orbit_Relation_def by auto text \p 51, ll 20--23, 26--27\ sublocale orbit: equivalence S Orbit_Relation proof (unfold_locales, auto simp: Orbit_Relation_def) fix x assume x: "x \ S" then have id: "x = identity S x" by simp with x show "\\ \ G. x = \ x" by fast fix \ assume \: "\ \ G" with x id have y: "x = compose S (symmetric.inverse \) \ x" by auto with x \ show "\\' \ G. x = \' (\ x)" by (metis compose_eq symmetric.sub.invertible symmetric.submonoid_inverse_closed) fix \ assume \: "\ \ G" with x have "\ (\ x) = compose S \ \ x" by (simp add: compose_eq) with \ \ show "\\ \ G. \ (\ x) = \ x" by fast qed text \p 51, ll 23--24\ theorem orbit_equality: "x \ S \ orbit.Class x = {\ x | \. \ \ G}" by (simp add: orbit.Class_def) (blast intro: orbit.symmetric dest: orbit.symmetric) end (* transformation_group *) context monoid_isomorphism begin text \p 52, ll 16--17\ theorem image_subgroup: assumes "subgroup G M (\) \" shows "subgroup (\ ` G) M' (\') \'" proof - interpret subgroup G M "(\)" \ by fact interpret image: monoid "\ ` G" "(\')" "\'" by unfold_locales (auto simp add: commutes_with_composition commutes_with_unit [symmetric]) show ?thesis proof (unfold_locales, auto) fix x assume x: "x \ G" show "image.invertible (\ x)" proof show "\ (sub.inverse x) \ \ ` G" using x by simp qed (auto simp: x commutes_with_composition commutes_with_unit) qed qed end (* monoid_isomorphism *) text \ Technical device to achieve Jacobson's notation for @{text Right_Coset} and @{text Left_Coset}. The definitions are pulled out of @{text subgroup_of_group} to a context where @{text H} is not a parameter. \ text \p 52, l 20\ locale coset_notation = fixes composition (infixl "\" 70) begin text \Equation 23\ text \p 52, l 20\ definition Right_Coset (infixl "|\" 70) where "H |\ x = {h \ x | h. h \ H}" text \p 53, ll 8--9\ definition Left_Coset (infixl "\|" 70) where "x \| H = {x \ h | h. h \ H}" text \p 52, l 20\ lemma Right_Coset_memI [intro]: "h \ H \ h \ x \ H |\ x" unfolding Right_Coset_def by blast text \p 52, l 20\ lemma Right_Coset_memE [elim]: "\ a \ H |\ x; \h. \ h \ H; a = h \ x \ \ P \ \ P" unfolding Right_Coset_def by blast text \p 53, ll 8--9\ lemma Left_Coset_memI [intro]: "h \ H \ x \ h \ x \| H" unfolding Left_Coset_def by blast text \p 53, ll 8--9\ lemma Left_Coset_memE [elim]: "\ a \ x \| H; \h. \ h \ H; a = x \ h \ \ P \ \ P" unfolding Left_Coset_def by blast end (* coset_notation *) text \p 52, l 12\ locale subgroup_of_group = subgroup H G "(\)" \ + coset_notation "(\)" + group G "(\)" \ for H and G and composition (infixl "\" 70) and unit ("\") begin text \p 52, ll 12--14\ interpretation left: left_translations_of_group .. interpretation right: right_translations_of_group .. text \ @{term "left.translation ` H"} denotes Jacobson's @{text "H\<^sub>L(G)"} and @{term "left.translation ` G"} denotes Jacobson's @{text "G\<^sub>L"}. \ text \p 52, ll 16--18\ theorem left_translations_of_subgroup_are_transformation_group [intro]: "transformation_group (left.translation ` H) G" proof - have "subgroup (left.translation ` H) (left.translation ` G) (compose G) (identity G)" by (rule left.image_subgroup) unfold_locales also have "subgroup (left.translation ` G) left.transformation.Sym (compose G) (identity G)" .. finally interpret right_coset: subgroup "left.translation ` H" left.transformation.Sym "compose G" "identity G" . show ?thesis .. qed text \p 52, l 18\ interpretation transformation_group "left.translation ` H" G .. text \p 52, ll 19--20\ theorem Right_Coset_is_orbit: "x \ G \ H |\ x = orbit.Class x" using left.translation_apply by (auto simp: orbit_equality Right_Coset_def) (metis imageI sub) text \p 52, ll 24--25\ theorem Right_Coset_Union: "(\x\G. H |\ x) = G" by (simp add: Right_Coset_is_orbit) text \p 52, l 26\ theorem Right_Coset_bij: assumes G [simp]: "x \ G" "y \ G" shows "bij_betw (inverse x \ y)\<^sub>R (H |\ x) (H |\ y)" proof (rule bij_betw_imageI) show "inj_on (inverse x \ y)\<^sub>R (H |\ x)" by (fastforce intro: inj_onI simp add: Right_Coset_is_orbit right.translation_apply orbit.block_closed) next show "(inverse x \ y)\<^sub>R ` (H |\ x) = H |\ y" by (force simp add: right.translation_apply associative invertible_right_inverse2) qed text \p 52, ll 25--26\ theorem Right_Cosets_cardinality: "\ x \ G; y \ G \ \ card (H |\ x) = card (H |\ y)" by (fast intro: bij_betw_same_card Right_Coset_bij) text \p 52, l 27\ theorem Right_Coset_unit: "H |\ \ = H" by (force simp add: Right_Coset_def) text \p 52, l 27\ theorem Right_Coset_cardinality: "x \ G \ card (H |\ x) = card H" using Right_Coset_unit Right_Cosets_cardinality unit_closed by presburger text \p 52, ll 31--32\ definition "index = card orbit.Partition" text \Theorem 1.5\ text \p 52, ll 33--35; p 53, ll 1--2\ theorem lagrange: "finite G \ card G = card H * index" unfolding index_def apply (subst card_partition) apply (auto simp: finite_UnionD orbit.complete orbit.disjoint) apply (metis Right_Coset_cardinality Right_Coset_is_orbit orbit.Block_self orbit.element_exists) done end (* subgroup_of_group *) text \Left cosets\ context subgroup begin text \p 31, ll 11--12\ lemma image_of_inverse [intro, simp]: "x \ G \ x \ inverse ` G" by (metis image_eqI sub.invertible sub.invertible_inverse_closed sub.invertible_inverse_inverse subgroup_inverse_equality) end (* subgroup *) context group begin (* Does Jacobson show this somewhere? *) text \p 53, ll 6--7\ lemma inverse_subgroupI: assumes sub: "subgroup H G (\) \" shows "subgroup (inverse ` H) G (\) \" proof - from sub interpret subgroup H G "(\)" \ . interpret inv: monoid "inverse ` H" "(\)" \ by unfold_locales (auto simp del: subgroup_inverse_equality) interpret inv: group "inverse ` H" "(\)" \ by unfold_locales (force simp del: subgroup_inverse_equality) show ?thesis by unfold_locales (auto simp del: subgroup_inverse_equality) qed text \p 53, ll 6--7\ lemma inverse_subgroupD: assumes sub: "subgroup (inverse ` H) G (\) \" and inv: "H \ Units" shows "subgroup H G (\) \" proof - from sub have "subgroup (inverse ` inverse ` H) G (\) \" by (rule inverse_subgroupI) moreover from inv [THEN subsetD, simplified Units_def] have "inverse ` inverse ` H = H" by (simp cong: image_cong add: image_comp) ultimately show ?thesis by simp qed end (* group *) context subgroup_of_group begin text \p 53, l 6\ interpretation right_translations_of_group .. text \ @{term "translation ` H"} denotes Jacobson's @{text "H\<^sub>R(G)"} and @{term "Translations"} denotes Jacobson's @{text "G\<^sub>R"}. \ text \p 53, ll 6--7\ theorem right_translations_of_subgroup_are_transformation_group [intro]: "transformation_group (translation ` H) G" proof - have "subgroup ((\a\G. transformation.symmetric.inverse (a)\<^sub>R) ` H) Translations (compose G) (identity G)" by (rule image_subgroup) unfold_locales also have "subgroup Translations transformation.Sym (compose G) (identity G)" .. finally interpret left_coset: subgroup "translation ` H" transformation.Sym "compose G" "identity G" by (auto intro: transformation.symmetric.inverse_subgroupD cong: image_cong simp: image_image transformation.symmetric.Units_def simp del: translation_unit_eq) show ?thesis .. qed text \p 53, ll 6--7\ interpretation transformation_group "translation ` H" G .. text \Equation 23 for left cosets\ text \p 53, ll 7--8\ theorem Left_Coset_is_orbit: "x \ G \ x \| H = orbit.Class x" using translation_apply by (auto simp: orbit_equality Left_Coset_def) (metis imageI sub) end (* subgroup_of_group *) subsection \Congruences. Quotient Monoids and Groups\ text \Def 1.4\ text \p 54, ll 19--22\ locale monoid_congruence = monoid + equivalence where S = M + assumes cong: "\ (a, a') \ E; (b, b') \ E \ \ (a \ b, a' \ b') \ E" begin text \p 54, ll 26--28\ theorem Class_cong: "\ Class a = Class a'; Class b = Class b'; a \ M; a' \ M; b \ M; b' \ M \ \ Class (a \ b) = Class (a' \ b')" by (simp add: Class_equivalence cong) text \p 54, ll 28--30\ definition quotient_composition (infixl "[\]" 70) where "quotient_composition = (\A \ M / E. \B \ M / E. THE C. \a \ A. \b \ B. C = Class (a \ b))" text \p 54, ll 28--30\ theorem Class_commutes_with_composition: "\ a \ M; b \ M \ \ Class a [\] Class b = Class (a \ b)" by (auto simp: quotient_composition_def intro: Class_cong [OF Class_eq Class_eq] del: equalityI) text \p 54, ll 30--31\ theorem quotient_composition_closed [intro, simp]: "\ A \ M / E; B \ M / E \ \ A [\] B \ M / E" by (erule quotient_ClassE)+ (simp add: Class_commutes_with_composition) text \p 54, l 32; p 55, ll 1--3\ sublocale quotient: monoid "M / E" "([\])" "Class \" by unfold_locales (auto simp: Class_commutes_with_composition associative elim!: quotient_ClassE) end (* monoid_congruence *) text \p 55, ll 16--17\ locale group_congruence = group + monoid_congruence where M = G begin text \p 55, ll 16--17\ notation quotient_composition (infixl "[\]" 70) text \p 55, l 18\ theorem Class_right_inverse: "a \ G \ Class a [\] Class (inverse a) = Class \" by (simp add: Class_commutes_with_composition) text \p 55, l 18\ theorem Class_left_inverse: "a \ G \ Class (inverse a) [\] Class a = Class \" by (simp add: Class_commutes_with_composition) text \p 55, l 18\ theorem Class_invertible: "a \ G \ quotient.invertible (Class a)" by (blast intro!: Class_right_inverse Class_left_inverse)+ text \p 55, l 18\ theorem Class_commutes_with_inverse: "a \ G \ quotient.inverse (Class a) = Class (inverse a)" by (rule quotient.inverse_equality) (auto simp: Class_right_inverse Class_left_inverse) text \p 55, l 17\ sublocale quotient: group "G / E" "([\])" "Class \" by unfold_locales (metis Block_self Class_invertible element_exists) end (* group_congruence *) text \Def 1.5\ text \p 55, ll 22--25\ locale normal_subgroup = subgroup_of_group K G "(\)" \ for K and G and composition (infixl "\" 70) and unit ("\") + assumes normal: "\ g \ G; k \ K \ \ inverse g \ k \ g \ K" text \Lemmas from the proof of Thm 1.6\ context subgroup_of_group begin text \We use @{term H} for @{term K}.\ text \p 56, ll 14--16\ theorem Left_equals_Right_coset_implies_normality: assumes [simp]: "\g. g \ G \ g \| H = H |\ g" shows "normal_subgroup H G (\) \" proof fix g k assume [simp]: "g \ G" "k \ H" have "k \ g \ g \| H" by auto then obtain k' where "k \ g = g \ k'" and "k' \ H" by blast then show "inverse g \ k \ g \ H" by (simp add: associative invertible_left_inverse2) qed end (* subgroup_of_group *) text \Thm 1.6, first part\ context group_congruence begin text \Jacobson's $K$\ text \p 56, l 29\ definition "Normal = Class \" text \p 56, ll 3--6\ interpretation subgroup "Normal" G "(\)" \ unfolding Normal_def proof (rule subgroupI) fix k1 and k2 assume K: "k1 \ Class \" "k2 \ Class \" then have "k1 \ k2 \ Class (k1 \ k2)" by blast also have "\ = Class k1 [\] Class k2" using K by (auto simp add: Class_commutes_with_composition Class_closed) also have "\ = Class \ [\] Class \" using K by (metis ClassD Class_eq unit_closed) also have "\ = Class \" by simp finally show "k1 \ k2 \ Class \" . next fix k assume K: "k \ Class \" then have "inverse k \ Class (inverse k)" by blast also have "\ = quotient.inverse (Class k)" using Class_commutes_with_inverse K by blast also have "\ = quotient.inverse (Class \)" using Block_self K by auto also have "\ = Class \" using quotient.inverse_unit by blast finally show "inverse k \ Class \" . qed auto text \Coset notation\ text \p 56, ll 5--6\ interpretation subgroup_of_group "Normal" G "(\)" \ .. text \Equation 25 for right cosets\ text \p 55, ll 29--30; p 56, ll 6--11\ theorem Right_Coset_Class_unit: assumes g: "g \ G" shows "Normal |\ g = Class g" unfolding Normal_def proof auto fix a \ \ll 6--8\ assume a: "a \ Class g" from a g have "a \ inverse g \ Class (a \ inverse g)" by blast also from a g have "\ = Class a [\] Class (inverse g)" by (simp add: Class_commutes_with_composition block_closed) also from a g have "\ = Class g [\] quotient.inverse (Class g)" using Block_self Class_commutes_with_inverse by auto also from g have "\ = Class \" by simp finally show "a \ Class \ |\ g" unfolding Right_Coset_def by simp (metis Class_closed a associative g inverse_equality invertible invertible_def right_unit) next fix a \ \ll 8--9\ assume a: "a \ Class \ |\ g" then obtain k where eq: "a = k \ g" and k: "k \ Class \" by blast with g have "Class a = Class k [\] Class g" using Class_commutes_with_composition by auto also from k have "\ = Class \ [\] Class g" using Block_self by auto also from g have "\ = Class g" by simp finally show "a \ Class g" using g eq k composition_closed quotient.unit_closed by blast qed text \Equation 25 for left cosets\ text \p 55, ll 29--30; p 56, ll 6--11\ theorem Left_Coset_Class_unit: assumes g: "g \ G" shows "g \| Normal = Class g" unfolding Normal_def proof auto fix a \ \ll 6--8\ assume a: "a \ Class g" from a g have "inverse g \ a \ Class (inverse g \ a)" by blast also from a g have "\ = Class (inverse g) [\] Class a" by (simp add: Class_commutes_with_composition block_closed) also from a g have "\ = quotient.inverse (Class g) [\] Class g" using Block_self Class_commutes_with_inverse by auto also from g have "\ = Class \" by simp finally show "a \ g \| Class \" unfolding Left_Coset_def by simp (metis Class_closed a associative g inverse_equality invertible invertible_def right_unit) next fix a \ \ll 8--9, ``the same thing holds''\ assume a: "a \ g \| Class \" then obtain k where eq: "a = g \ k" and k: "k \ Class \" by blast with g have "Class a = Class g [\] Class k" using Class_commutes_with_composition by auto also from k have "\ = Class g [\] Class \" using Block_self by auto also from g have "\ = Class g" by simp finally show "a \ Class g" using g eq k composition_closed quotient.unit_closed by blast qed text \Thm 1.6, statement of first part\ text \p 55, ll 28--29; p 56, ll 12--16\ theorem Class_unit_is_normal: "normal_subgroup Normal G (\) \" proof - { fix g assume "g \ G" then have "g \| Normal = Normal |\ g" by (simp add: Right_Coset_Class_unit Left_Coset_Class_unit) } then show ?thesis by (rule Left_equals_Right_coset_implies_normality) qed sublocale normal: normal_subgroup Normal G "(\)" \ by (fact Class_unit_is_normal) end (* group_congruence *) context normal_subgroup begin text \p 56, ll 16--19\ theorem Left_equals_Right_coset: "g \ G \ g \| K = K |\ g" proof assume [simp]: "g \ G" show "K |\ g \ g \| K" proof fix x assume x: "x \ K |\ g" then obtain k where "x = k \ g" and [simp]: "k \ K" by (auto simp add: Right_Coset_def) then have "x = g \ (inverse g \ k \ g)" by (simp add: associative invertible_right_inverse2) also from normal have "\ \ g \| K" by (auto simp add: Left_Coset_def) finally show "x \ g \| K" . qed next assume [simp]: "g \ G" show "g \| K \ K |\ g" proof fix x assume x: "x \ g \| K" then obtain k where "x = g \ k" and [simp]: "k \ K" by (auto simp add: Left_Coset_def) then have "x = (inverse (inverse g) \ k \ inverse g) \ g" by (simp add: associative del: invertible_right_inverse) also from normal [where g = "inverse g"] have "\ \ K |\ g" by (auto simp add: Right_Coset_def) finally show "x \ K |\ g" . qed qed text \Thm 1.6, second part\ text \p 55, ll 31--32; p 56, ll 20--21\ definition "Congruence = {(a, b). a \ G \ b \ G \ inverse a \ b \ K}" text \p 56, ll 21--22\ interpretation right_translations_of_group .. text \p 56, ll 21--22\ interpretation transformation_group "translation ` K" G rewrites "Orbit_Relation = Congruence" proof - interpret transformation_group "translation ` K" G .. show "Orbit_Relation = Congruence" unfolding Orbit_Relation_def Congruence_def by (force simp: invertible_left_inverse2 invertible_right_inverse2 translation_apply simp del: restrict_apply) qed rule text \p 56, ll 20--21\ lemma CongruenceI: "\ a = b \ k; a \ G; b \ G; k \ K \ \ (a, b) \ Congruence" by (clarsimp simp: Congruence_def associative inverse_composition_commute) text \p 56, ll 20--21\ lemma CongruenceD: "(a, b) \ Congruence \ \k\K. a = b \ k" by (drule orbit.symmetric) (force simp: Congruence_def invertible_right_inverse2) text \ ``We showed in the last section that the relation we are considering is an equivalence relation in @{term G} for any subgroup @{term K} of @{term G}. We now proceed to show that normality of @{term K} ensures that [...] $a \equiv b \pmod{K}$ is a congruence.'' \ text \p 55, ll 30--32; p 56, ll 1, 22--28\ sublocale group_congruence where E = Congruence rewrites "Normal = K" proof - show "group_congruence G (\) \ Congruence" proof unfold_locales note CongruenceI [intro] CongruenceD [dest] fix a g b h assume 1: "(a, g) \ Congruence" and 2: "(b, h) \ Congruence" then have G: "a \ G" "g \ G" "b \ G" "h \ G" unfolding Congruence_def by clarify+ from 1 obtain k1 where a: "a = g \ k1" and k1: "k1 \ K" by blast from 2 obtain k2 where b: "b = h \ k2" and k2: "k2 \ K" by blast from G Left_equals_Right_coset have "K |\ h = h \| K" by blast with k1 obtain k3 where c: "k1 \ h = h \ k3" and k3: "k3 \ K" unfolding Left_Coset_def Right_Coset_def by blast from G k1 k2 a b have "a \ b = g \ k1 \ h \ k2" by (simp add: associative) also from G k1 k3 c have "\ = g \ h \ k3 \ k2" by (simp add: associative) also have "\ = (g \ h) \ (k3 \ k2)" using G k2 k3 by (simp add: associative) finally show "(a \ b, g \ h) \ Congruence" using G k2 k3 by blast qed then interpret group_congruence where E = Congruence . show "Normal = K" unfolding Normal_def orbit.Class_def unfolding Congruence_def using invertible_inverse_inverse submonoid_inverse_closed by fastforce qed end (* normal_subgroup *) (* deletes translations and orbits, recovers Class for congruence class *) context group begin text \Pulled out of @{locale normal_subgroup} to achieve standard notation.\ text \p 56, ll 31--32\ abbreviation Factor_Group (infixl "'/'/" 75) where "S // K \ S / (normal_subgroup.Congruence K G (\) \)" end (* group *) context normal_subgroup begin text \p 56, ll 28--29\ theorem Class_unit_normal_subgroup: "Class \ = K" unfolding Class_def unfolding Congruence_def using invertible_inverse_inverse submonoid_inverse_closed by fastforce text \p 56, ll 1--2; p 56, l 29\ theorem Class_is_Left_Coset: "g \ G \ Class g = g \| K" using Left_Coset_Class_unit Class_unit_normal_subgroup by simp text \p 56, l 29\ lemma Left_CosetE: "\ A \ G // K; \a. a \ G \ P (a \| K) \ \ P A" by (metis Class_is_Left_Coset quotient_ClassE) text \Equation 26\ text \p 56, ll 32--34\ theorem factor_composition [simp]: "\ g \ G; h \ G \ \ (g \| K) [\] (h \| K) = g \ h \| K" using Class_commutes_with_composition Class_is_Left_Coset by auto text \p 56, l 35\ theorem factor_unit: "K = \ \| K" using Class_is_Left_Coset Class_unit_normal_subgroup by blast text \p 56, l 35\ theorem factor_inverse [simp]: "g \ G \ quotient.inverse (g \| K) = (inverse g \| K)" using Class_commutes_with_inverse Class_is_Left_Coset by auto end (* normal_subgroup *) text \p 57, ll 4--5\ locale subgroup_of_abelian_group = subgroup_of_group H G "(\)" \ + abelian_group G "(\)" \ for H and G and composition (infixl "\" 70) and unit ("\") text \p 57, ll 4--5\ sublocale subgroup_of_abelian_group \ normal_subgroup H G "(\)" \ using commutative invertible_right_inverse2 by unfold_locales auto subsection \Homomorphims\ text \Def 1.6\ text \p 58, l 33; p 59, ll 1--2\ locale monoid_homomorphism = map \ M M'+ source: monoid M "(\)" \ + target: monoid M' "(\')" "\'" for \ and M and composition (infixl "\" 70) and unit ("\") and M' and composition' (infixl "\''" 70) and unit' ("\''") + assumes commutes_with_composition: "\ x \ M; y \ M \ \ \ (x \ y) = \ x \' \ y" and commutes_with_unit: "\ \ = \'" begin text \Jacobson notes that @{thm [source] commutes_with_unit} is not necessary for groups, but doesn't make use of that later.\ text \p 58, l 33; p 59, ll 1--2\ notation source.invertible ("invertible _" [100] 100) notation source.inverse ("inverse _" [100] 100) -notation target.invertible ("invertible' _" [100] 100) -notation target.inverse ("inverse' _" [100] 100) +notation target.invertible ("invertible'' _" [100] 100) +notation target.inverse ("inverse'' _" [100] 100) end (* monoid_homomorphism *) text \p 59, ll 29--30\ locale monoid_epimorphism = monoid_homomorphism + surjective_map \ M M' text \p 59, l 30\ locale monoid_monomorphism = monoid_homomorphism + injective_map \ M M' text \p 59, ll 30--31\ sublocale monoid_isomorphism \ monoid_epimorphism by unfold_locales (auto simp: commutes_with_composition commutes_with_unit) text \p 59, ll 30--31\ sublocale monoid_isomorphism \ monoid_monomorphism by unfold_locales (auto simp: commutes_with_composition commutes_with_unit) context monoid_homomorphism begin text \p 59, ll 33--34\ theorem invertible_image_lemma: assumes "invertible a" "a \ M" shows "\ a \' \ (inverse a) = \'" and "\ (inverse a) \' \ a = \'" using assms commutes_with_composition commutes_with_unit source.inverse_equality by auto (metis source.invertible_inverse_closed source.invertible_left_inverse) text \p 59, l 34; p 60, l 1\ theorem invertible_target_invertible [intro, simp]: "\ invertible a; a \ M \ \ invertible' (\ a)" using invertible_image_lemma by blast text \p 60, l 1\ theorem invertible_commutes_with_inverse: "\ invertible a; a \ M \ \ \ (inverse a) = inverse' (\ a)" using invertible_image_lemma target.inverse_equality by fastforce end (* monoid_homomorphism *) text \p 60, ll 32--34; p 61, l 1\ sublocale monoid_congruence \ natural: monoid_homomorphism Class M "(\)" \ "M / E" "([\])" "Class \" by unfold_locales (auto simp: PiE_I Class_commutes_with_composition) text \Fundamental Theorem of Homomorphisms of Monoids\ text \p 61, ll 5, 14--16\ sublocale monoid_homomorphism \ image: submonoid "\ ` M" M' "(\')" "\'" by unfold_locales (auto simp: commutes_with_composition [symmetric] commutes_with_unit [symmetric]) text \p 61, l 4\ locale monoid_homomorphism_fundamental = monoid_homomorphism begin text \p 61, ll 17--18\ sublocale fiber_relation \ M M' .. notation Fiber_Relation ("E'(_')") text \p 61, ll 6--7, 18--20\ sublocale monoid_congruence where E = "E(\)" using Class_eq by unfold_locales (rule Class_equivalence [THEN iffD1], auto simp: left_closed right_closed commutes_with_composition Fiber_equality) text \p 61, ll 7--9\ text \ @{term induced} denotes Jacobson's $\bar{\eta}$. We have the commutativity of the diagram, where @{term induced} is unique: @{thm [display] factorization} @{thm [display] uniqueness}. \ text \p 61, l 20\ notation quotient_composition (infixl "[\]" 70) text \p 61, ll 7--8, 22--25\ sublocale induced: monoid_homomorphism induced "M / E(\)" "([\])" "Class \" "M'" "(\')" "\'" apply unfold_locales apply (auto simp: commutes_with_unit) apply (fastforce simp: commutes_with_composition commutes_with_unit Class_commutes_with_composition) done text \p 61, ll 9, 26\ sublocale natural: monoid_epimorphism Class M "(\)" \ "M / E(\)" "([\])" "Class \" .. text \p 61, ll 9, 26--27\ sublocale induced: monoid_monomorphism induced "M / E(\)" "([\])" "Class \" "M'" "(\')" "\'" .. end (* monoid_homomorphism_fundamental *) text \p 62, ll 12--13\ locale group_homomorphism = monoid_homomorphism \ G "(\)" \ G' "(\')" "\'" + source: group G "(\)" \ + target: group G' "(\')" "\'" for \ and G and composition (infixl "\" 70) and unit ("\") and G' and composition' (infixl "\''" 70) and unit' ("\''") begin text \p 62, l 13\ sublocale image: subgroup "\ ` G" G' "(\')" "\'" using invertible_image_lemma by unfold_locales auto text \p 62, ll 13--14\ definition "Ker = \ -` {\'} \ G" text \p 62, ll 13--14\ lemma Ker_equality: "Ker = {a | a. a \ G \ \ a = \'}" unfolding Ker_def by auto text \p 62, ll 13--14\ lemma Ker_closed [intro, simp]: "a \ Ker \ a \ G" unfolding Ker_def by simp text \p 62, ll 13--14\ lemma Ker_image [intro]: (* loops as a simprule *) "a \ Ker \ \ a = \'" unfolding Ker_def by simp text \p 62, ll 13--14\ lemma Ker_memI [intro]: (* loops as a simprule *) "\ \ a = \'; a \ G \ \ a \ Ker" unfolding Ker_def by simp text \p 62, ll 15--16\ sublocale kernel: normal_subgroup Ker G proof - interpret kernel: submonoid Ker G unfolding Ker_def by unfold_locales (auto simp: commutes_with_composition commutes_with_unit) interpret kernel: subgroup Ker G by unfold_locales (force intro: source.invertible_right_inverse simp: Ker_image invertible_commutes_with_inverse) show "normal_subgroup Ker G (\) \" apply unfold_locales unfolding Ker_def by (auto simp: commutes_with_composition invertible_image_lemma(2)) qed text \p 62, ll 17--20\ theorem injective_iff_kernel_unit: "inj_on \ G \ Ker = {\}" proof (rule Not_eq_iff [THEN iffD1, OF iffI]) assume "Ker \ {\}" then obtain b where b: "b \ Ker" "b \ \" by blast then have "\ b = \ \" by (simp add: Ker_image) with b show "\ inj_on \ G" by (meson inj_onD kernel.sub source.unit_closed) next assume "\ inj_on \ G" then obtain a b where "a \ b" and ab: "a \ G" "b \ G" "\ a = \ b" by (meson inj_onI) then have "inverse a \ b \ \" "\ (inverse a \ b) = \'" using ab source.invertible_right_inverse2 by force (metis ab commutes_with_composition invertible_image_lemma(2) source.invertible source.invertible_inverse_closed) then have "inverse a \ b \ Ker" using Ker_memI ab by blast then show "Ker \ {\}" using \inverse a \ b \ \\ by blast qed end (* group_homomorphism *) text \p 62, l 24\ locale group_epimorphism = group_homomorphism + monoid_epimorphism \ G "(\)" \ G' "(\')" "\'" text \p 62, l 21\ locale normal_subgroup_in_kernel = group_homomorphism + contained: normal_subgroup L G "(\)" \ for L + assumes subset: "L \ Ker" begin text \p 62, l 21\ notation contained.quotient_composition (infixl "[\]" 70) text \"homomorphism onto @{term "G // L"}"\ text \p 62, ll 23--24\ sublocale natural: group_epimorphism contained.Class G "(\)" \ "G // L" "([\])" "contained.Class \" .. text \p 62, ll 25--26\ theorem left_coset_equality: assumes eq: "a \| L = b \| L" and [simp]: "a \ G" and b: "b \ G" shows "\ a = \ b" proof - obtain l where l: "b = a \ l" "l \ L" by (metis b contained.Class_is_Left_Coset contained.Class_self eq kernel.Left_Coset_memE) then have "\ a = \ a \' \ l" using Ker_image monoid_homomorphism.commutes_with_composition subset by auto also have "\ = \ b" by (simp add: commutes_with_composition l) finally show ?thesis . qed text \$\bar{\eta}$\ text \p 62, ll 26--27\ definition "induced = (\A \ G // L. THE b. \a \ G. a \| L = A \ b = \ a)" text \p 62, ll 26--27\ lemma induced_closed [intro, simp]: assumes [simp]: "A \ G // L" shows "induced A \ G'" proof - obtain a where a: "a \ G" "a \| L = A" using contained.Class_is_Left_Coset contained.Partition_def assms by auto have "(THE b. \a \ G. a \| L = A \ b = \ a) \ G'" apply (rule theI2) using a by (auto intro: left_coset_equality) then show ?thesis unfolding induced_def by simp qed text \p 62, ll 26--27\ lemma induced_undefined [intro, simp]: "A \ G // L \ induced A = undefined" unfolding induced_def by simp text \p 62, ll 26--27\ theorem induced_left_coset_closed [intro, simp]: "a \ G \ induced (a \| L) \ G'" using contained.Class_is_Left_Coset contained.Class_in_Partition by auto text \p 62, ll 26--27\ theorem induced_left_coset_equality [simp]: assumes [simp]: "a \ G" shows "induced (a \| L) = \ a" proof - have "(THE b. \a' \ G. a' \| L = a \| L \ b = \ a') = \ a" by (rule the_equality) (auto intro: left_coset_equality) then show ?thesis unfolding induced_def using contained.Class_is_Left_Coset contained.Class_in_Partition by auto qed text \p 62, l 27\ theorem induced_Left_Coset_commutes_with_composition [simp]: "\ a \ G; b \ G \ \ induced ((a \| L) [\] (b \| L)) = induced (a \| L) \' induced (b \| L)" by (simp add: commutes_with_composition) text \p 62, ll 27--28\ theorem induced_group_homomorphism: "group_homomorphism induced (G // L) ([\]) (contained.Class \) G' (\') \'" apply unfold_locales apply (auto elim!: contained.Left_CosetE simp: commutes_with_composition commutes_with_unit) using contained.factor_unit induced_left_coset_equality apply (fastforce simp: contained.Class_unit_normal_subgroup) done text \p 62, l 28\ sublocale induced: group_homomorphism induced "G // L" "([\])" "contained.Class \" G' "(\')" "\'" by (fact induced_group_homomorphism) text \p 62, ll 28--29\ theorem factorization_lemma: "a \ G \ compose G induced contained.Class a = \ a" unfolding compose_def by (simp add: contained.Class_is_Left_Coset) text \p 62, ll 29--30\ theorem factorization [simp]: "compose G induced contained.Class = \" by rule (simp add: compose_def contained.Class_is_Left_Coset) text \ Jacobson does not state the uniqueness of @{term induced} explicitly but he uses it later, for rings, on p 107. \ text \p 62, l 30\ theorem uniqueness: assumes map: "\ \ G // L \\<^sub>E G'" and factorization: "compose G \ contained.Class = \" shows "\ = induced" proof fix A show "\ A = induced A" proof (cases "A \ G // L") case True then obtain a where [simp]: "A = contained.Class a" "a \ G" by fast then have "\ (contained.Class a) = \ a" by (metis compose_eq factorization) also have "\ = induced (contained.Class a)" by (simp add: contained.Class_is_Left_Coset) finally show ?thesis by simp qed (simp add: induced_def PiE_arb [OF map]) qed text \p 62, l 31\ theorem induced_image: "induced ` (G // L) = \ ` G" by (metis factorization contained.natural.surjective surj_compose) text \p 62, l 33\ interpretation L: normal_subgroup L Ker by unfold_locales (auto simp: subset, metis kernel.sub kernel.subgroup_inverse_equality contained.normal) text \p 62, ll 31--33\ theorem induced_kernel: "induced.Ker = Ker / L.Congruence" (* Ker // L is apparently not the right thing *) proof - have "induced.Ker = { a \| L | a. a \ G \ a \ Ker }" unfolding induced.Ker_equality by simp (metis (hide_lams) contained.Class_is_Left_Coset Ker_image Ker_memI induced_left_coset_equality contained.Class_in_Partition contained.representant_exists) also have "\ = Ker / L.Congruence" using L.Class_is_Left_Coset L.Class_in_Partition by auto (metis L.Class_is_Left_Coset L.representant_exists kernel.sub) finally show ?thesis . qed text \p 62, ll 34--35\ theorem induced_inj_on: "inj_on induced (G // L) \ L = Ker" apply (simp add: induced.injective_iff_kernel_unit induced_kernel contained.Class_unit_normal_subgroup) apply rule using L.block_exists apply auto [1] using L.Block_self L.Class_unit_normal_subgroup L.quotient.unit_closed L.representant_exists apply auto done end (* normal_subgroup_in_kernel *) text \Fundamental Theorem of Homomorphisms of Groups\ text \p 63, l 1\ locale group_homomorphism_fundamental = group_homomorphism begin text \p 63, l 1\ notation kernel.quotient_composition (infixl "[\]" 70) text \p 63, l 1\ sublocale normal_subgroup_in_kernel where L = Ker by unfold_locales rule text \p 62, ll 36--37; p 63, l 1\ text \ @{term induced} denotes Jacobson's $\bar{\eta}$. We have the commutativity of the diagram, where @{term induced} is unique: @{thm [display] factorization} @{thm [display] uniqueness} \ end (* group_homomorphism_fundamental *) text \p 63, l 5\ locale group_isomorphism = group_homomorphism + bijective_map \ G G' begin text \p 63, l 5\ sublocale monoid_isomorphism \ G "(\)" \ G' "(\')" "\'" by unfold_locales (auto simp: commutes_with_composition) text \p 63, l 6\ lemma inverse_group_isomorphism: "group_isomorphism (restrict (inv_into G \) G') G' (\') \' G (\) \" using commutes_with_composition commutes_with_unit surjective by unfold_locales auto end (* group_isomorphism *) text \p 63, l 6\ definition isomorphic_as_groups (infixl "\\<^sub>G" 50) where "\ \\<^sub>G \' \ (let (G, composition, unit) = \; (G', composition', unit') = \' in (\\. group_isomorphism \ G composition unit G' composition' unit'))" text \p 63, l 6\ lemma isomorphic_as_groups_symmetric: "(G, composition, unit) \\<^sub>G (G', composition', unit') \ (G', composition', unit') \\<^sub>G (G, composition, unit)" by (simp add: isomorphic_as_groups_def) (meson group_isomorphism.inverse_group_isomorphism) text \p 63, l 1\ sublocale group_isomorphism \ group_epimorphism .. text \p 63, l 1\ locale group_epimorphism_fundamental = group_homomorphism_fundamental + group_epimorphism begin text \p 63, ll 1--2\ interpretation image: group_homomorphism induced "G // Ker" "([\])" "kernel.Class \" "(\ ` G)" "(\')" "\'" by (simp add: surjective group_homomorphism_fundamental.intro induced_group_homomorphism) text \p 63, ll 1--2\ sublocale image: group_isomorphism induced "G // Ker" "([\])" "kernel.Class \" "(\ ` G)" "(\')" "\'" using induced_group_homomorphism by unfold_locales (auto simp: bij_betw_def induced_image induced_inj_on induced.commutes_with_composition) end (* group_epimorphism_fundamental *) context group_homomorphism begin text \p 63, ll 5--7\ theorem image_isomorphic_to_factor_group: "\K composition unit. normal_subgroup K G (\) \ \ (\ ` G, (\'), \') \\<^sub>G (G // K, composition, unit)" proof - interpret image: group_epimorphism_fundamental where G' = "\ ` G" by unfold_locales (auto simp: commutes_with_composition) have "group_isomorphism image.induced (G // Ker) ([\]) (kernel.Class \) (\ ` G) (\') \'" .. then have "(\ ` G, (\'), \') \\<^sub>G (G // Ker, ([\]), kernel.Class \)" by (simp add: isomorphic_as_groups_def) (meson group_isomorphism.inverse_group_isomorphism) moreover have "normal_subgroup Ker G (\) \" .. ultimately show ?thesis by blast qed end (* group_homomorphism *) end diff --git a/thys/Jinja/BV/JVM_SemiType.thy b/thys/Jinja/BV/JVM_SemiType.thy --- a/thys/Jinja/BV/JVM_SemiType.thy +++ b/thys/Jinja/BV/JVM_SemiType.thy @@ -1,256 +1,256 @@ (* Title: HOL/MicroJava/BV/JVM.thy Author: Gerwin Klein Copyright 2000 TUM *) section \The JVM Type System as Semilattice\ theory JVM_SemiType imports SemiType begin type_synonym ty\<^sub>l = "ty err list" type_synonym ty\<^sub>s = "ty list" type_synonym ty\<^sub>i = "ty\<^sub>s \ ty\<^sub>l" type_synonym ty\<^sub>i' = "ty\<^sub>i option" type_synonym ty\<^sub>m = "ty\<^sub>i' list" type_synonym ty\<^sub>P = "mname \ cname \ ty\<^sub>m" definition stk_esl :: "'c prog \ nat \ ty\<^sub>s esl" where "stk_esl P mxs \ upto_esl mxs (SemiType.esl P)" definition loc_sl :: "'c prog \ nat \ ty\<^sub>l sl" where "loc_sl P mxl \ Listn.sl mxl (Err.sl (SemiType.esl P))" definition sl :: "'c prog \ nat \ nat \ ty\<^sub>i' err sl" where "sl P mxs mxl \ Err.sl(Opt.esl(Product.esl (stk_esl P mxs) (Err.esl(loc_sl P mxl))))" definition states :: "'c prog \ nat \ nat \ ty\<^sub>i' err set" where "states P mxs mxl \ fst(sl P mxs mxl)" definition le :: "'c prog \ nat \ nat \ ty\<^sub>i' err ord" where "le P mxs mxl \ fst(snd(sl P mxs mxl))" definition sup :: "'c prog \ nat \ nat \ ty\<^sub>i' err binop" where "sup P mxs mxl \ snd(snd(sl P mxs mxl))" definition sup_ty_opt :: "['c prog,ty err,ty err] \ bool" ("_ \ _ \\<^sub>\ _" [71,71,71] 70) where "sup_ty_opt P \ Err.le (subtype P)" definition sup_state :: "['c prog,ty\<^sub>i,ty\<^sub>i] \ bool" ("_ \ _ \\<^sub>i _" [71,71,71] 70) where "sup_state P \ Product.le (Listn.le (subtype P)) (Listn.le (sup_ty_opt P))" definition sup_state_opt :: "['c prog,ty\<^sub>i',ty\<^sub>i'] \ bool" - ("_ \ _ \' _" [71,71,71] 70) + ("_ \ _ \'' _" [71,71,71] 70) where "sup_state_opt P \ Opt.le (sup_state P)" abbreviation sup_loc :: "['c prog,ty\<^sub>l,ty\<^sub>l] \ bool" ("_ \ _ [\\<^sub>\] _" [71,71,71] 70) where "P \ LT [\\<^sub>\] LT' \ list_all2 (sup_ty_opt P) LT LT'" notation (ASCII) sup_ty_opt ("_ |- _ <=T _" [71,71,71] 70) and sup_state ("_ |- _ <=i _" [71,71,71] 70) and sup_state_opt ("_ |- _ <=' _" [71,71,71] 70) and sup_loc ("_ |- _ [<=T] _" [71,71,71] 70) subsection "Unfolding" lemma JVM_states_unfold: "states P mxs mxl \ err(opt((Union {list n (types P) |n. n <= mxs}) \ list mxl (err(types P))))" (*<*) apply (unfold states_def sl_def Opt.esl_def Err.sl_def stk_esl_def loc_sl_def Product.esl_def Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def) apply simp done (*>*) lemma JVM_le_unfold: "le P m n \ Err.le(Opt.le(Product.le(Listn.le(subtype P))(Listn.le(Err.le(subtype P)))))" (*<*) apply (unfold le_def sl_def Opt.esl_def Err.sl_def stk_esl_def loc_sl_def Product.esl_def Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def) apply simp done (*>*) lemma sl_def2: "JVM_SemiType.sl P mxs mxl \ (states P mxs mxl, JVM_SemiType.le P mxs mxl, JVM_SemiType.sup P mxs mxl)" (*<*) by (unfold JVM_SemiType.sup_def states_def JVM_SemiType.le_def) simp (*>*) lemma JVM_le_conv: "le P m n (OK t1) (OK t2) = P \ t1 \' t2" (*<*) by (simp add: JVM_le_unfold Err.le_def lesub_def sup_state_opt_def sup_state_def sup_ty_opt_def) (*>*) lemma JVM_le_Err_conv: "le P m n = Err.le (sup_state_opt P)" (*<*) by (unfold sup_state_opt_def sup_state_def sup_ty_opt_def JVM_le_unfold) simp (*>*) lemma err_le_unfold [iff]: "Err.le r (OK a) (OK b) = r a b" (*<*) by (simp add: Err.le_def lesub_def) (*>*) subsection \Semilattice\ lemma order_sup_state_opt [intro, simp]: "wf_prog wf_mb P \ order (sup_state_opt P)" (*<*) by (unfold sup_state_opt_def sup_state_def sup_ty_opt_def) blast (*>*) lemma semilat_JVM [intro?]: "wf_prog wf_mb P \ semilat (JVM_SemiType.sl P mxs mxl)" (*<*) apply (unfold JVM_SemiType.sl_def stk_esl_def loc_sl_def) apply (blast intro: err_semilat_Product_esl err_semilat_upto_esl Listn_sl err_semilat_JType_esl) done (*>*) lemma acc_JVM [intro]: "wf_prog wf_mb P \ acc (JVM_SemiType.le P mxs mxl)" (*<*) by (unfold JVM_le_unfold) blast (*>*) subsection \Widening with \\\\ lemma subtype_refl[iff]: "subtype P t t" (*<*) by (simp add: fun_of_def) (*>*) lemma sup_ty_opt_refl [iff]: "P \ T \\<^sub>\ T" (*<*) apply (unfold sup_ty_opt_def) apply (fold lesub_def) apply (rule le_err_refl) apply (simp add: lesub_def) done (*>*) lemma Err_any_conv [iff]: "P \ Err \\<^sub>\ T = (T = Err)" (*<*) by (unfold sup_ty_opt_def) (rule Err_le_conv [simplified lesub_def]) (*>*) lemma any_Err [iff]: "P \ T \\<^sub>\ Err" (*<*) by (unfold sup_ty_opt_def) (rule le_Err [simplified lesub_def]) (*>*) lemma OK_OK_conv [iff]: "P \ OK T \\<^sub>\ OK T' = P \ T \ T'" (*<*) by (simp add: sup_ty_opt_def fun_of_def) (*>*) lemma any_OK_conv [iff]: "P \ X \\<^sub>\ OK T' = (\T. X = OK T \ P \ T \ T')" (*<*) apply (unfold sup_ty_opt_def) apply (rule le_OK_conv [simplified lesub_def]) done (*>*) lemma OK_any_conv: "P \ OK T \\<^sub>\ X = (X = Err \ (\T'. X = OK T' \ P \ T \ T'))" (*<*) apply (unfold sup_ty_opt_def) apply (rule OK_le_conv [simplified lesub_def]) done (*>*) lemma sup_ty_opt_trans [intro?, trans]: "\P \ a \\<^sub>\ b; P \ b \\<^sub>\ c\ \ P \ a \\<^sub>\ c" (*<*) by (auto intro: widen_trans simp add: sup_ty_opt_def Err.le_def lesub_def fun_of_def split: err.splits) (*>*) subsection "Stack and Registers" lemma stk_convert: "P \ ST [\] ST' = Listn.le (subtype P) ST ST'" (*<*) by (simp add: Listn.le_def lesub_def) (*>*) lemma sup_loc_refl [iff]: "P \ LT [\\<^sub>\] LT" (*<*) by (rule list_all2_refl) simp (*>*) lemmas sup_loc_Cons1 [iff] = list_all2_Cons1 [of "sup_ty_opt P"] for P lemma sup_loc_def: "P \ LT [\\<^sub>\] LT' \ Listn.le (sup_ty_opt P) LT LT'" (*<*) by (simp add: Listn.le_def lesub_def) (*>*) lemma sup_loc_widens_conv [iff]: "P \ map OK Ts [\\<^sub>\] map OK Ts' = P \ Ts [\] Ts'" (*<*) by (simp add: list_all2_map1 list_all2_map2) (*>*) lemma sup_loc_trans [intro?, trans]: "\P \ a [\\<^sub>\] b; P \ b [\\<^sub>\] c\ \ P \ a [\\<^sub>\] c" (*<*) by (rule list_all2_trans, rule sup_ty_opt_trans) (*>*) subsection "State Type" lemma sup_state_conv [iff]: "P \ (ST,LT) \\<^sub>i (ST',LT') = (P \ ST [\] ST' \ P \ LT [\\<^sub>\] LT')" (*<*) by (auto simp add: sup_state_def stk_convert lesub_def Product.le_def sup_loc_def) (*>*) lemma sup_state_conv2: "P \ s1 \\<^sub>i s2 = (P \ fst s1 [\] fst s2 \ P \ snd s1 [\\<^sub>\] snd s2)" (*<*) by (cases s1, cases s2) simp (*>*) lemma sup_state_refl [iff]: "P \ s \\<^sub>i s" (*<*) by (auto simp add: sup_state_conv2) (*>*) lemma sup_state_trans [intro?, trans]: "\P \ a \\<^sub>i b; P \ b \\<^sub>i c\ \ P \ a \\<^sub>i c" (*<*) by (auto intro: sup_loc_trans widens_trans simp add: sup_state_conv2) (*>*) lemma sup_state_opt_None_any [iff]: "P \ None \' s" (*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*) lemma sup_state_opt_any_None [iff]: "P \ s \' None = (s = None)" (*<*) by (simp add: sup_state_opt_def Opt.le_def) (*>*) lemma sup_state_opt_Some_Some [iff]: "P \ Some a \' Some b = P \ a \\<^sub>i b" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_any_Some: "P \ (Some s) \' X = (\s'. X = Some s' \ P \ s \\<^sub>i s')" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_refl [iff]: "P \ s \' s" (*<*) by (simp add: sup_state_opt_def Opt.le_def lesub_def) (*>*) lemma sup_state_opt_trans [intro?, trans]: "\P \ a \' b; P \ b \' c\ \ P \ a \' c" (*<*) apply (unfold sup_state_opt_def Opt.le_def lesub_def) apply (simp del: split_paired_All) apply (rule sup_state_trans, assumption+) done (*>*) end diff --git a/thys/Jinja/J/Progress.thy b/thys/Jinja/J/Progress.thy --- a/thys/Jinja/J/Progress.thy +++ b/thys/Jinja/J/Progress.thy @@ -1,626 +1,626 @@ (* Title: Jinja/J/SmallProgress.thy Author: Tobias Nipkow Copyright 2003 Technische Universitaet Muenchen *) section \Progress of Small Step Semantics\ theory Progress imports Equivalence WellTypeRT DefAss "../Common/Conform" begin lemma final_addrE: "\ P,E,h \ e : Class C; final e; \a. e = addr a \ R; \a. e = Throw a \ R \ \ R" (*<*)by(auto simp:final_def)(*>*) lemma finalRefE: "\ P,E,h \ e : T; is_refT T; final e; e = null \ R; \a C. \ e = addr a; T = Class C \ \ R; \a. e = Throw a \ R \ \ R" (*<*)by(auto simp:final_def is_refT_def)(*>*) text\Derivation of new induction scheme for well typing:\ inductive WTrt' :: "[J_prog,heap,env,expr,ty] \ bool" and WTrts' :: "[J_prog,heap,env,expr list, ty list] \ bool" and WTrt2' :: "[J_prog,env,heap,expr,ty] \ bool" - ("_,_,_ \ _ :' _" [51,51,51]50) + ("_,_,_ \ _ :'' _" [51,51,51]50) and WTrts2' :: "[J_prog,env,heap,expr list, ty list] \ bool" ("_,_,_ \ _ [:''] _" [51,51,51]50) for P :: J_prog and h :: heap where "P,E,h \ e :' T \ WTrt' P h E e T" | "P,E,h \ es [:'] Ts \ WTrts' P h E es Ts" | "is_class P C \ P,E,h \ new C :' Class C" | "\ P,E,h \ e :' T; is_refT T; is_class P C \ \ P,E,h \ Cast C e :' Class C" | "typeof\<^bsub>h\<^esub> v = Some T \ P,E,h \ Val v :' T" | "E v = Some T \ P,E,h \ Var v :' T" | "\ P,E,h \ e\<^sub>1 :' T\<^sub>1; P,E,h \ e\<^sub>2 :' T\<^sub>2 \ \ P,E,h \ e\<^sub>1 \Eq\ e\<^sub>2 :' Boolean" | "\ P,E,h \ e\<^sub>1 :' Integer; P,E,h \ e\<^sub>2 :' Integer \ \ P,E,h \ e\<^sub>1 \Add\ e\<^sub>2 :' Integer" | "\ P,E,h \ Var V :' T; P,E,h \ e :' T'; P \ T' \ T \<^cancel>\V \ This\ \ \ P,E,h \ V:=e :' Void" | "\ P,E,h \ e :' Class C; P \ C has F:T in D \ \ P,E,h \ e\F{D} :' T" | "P,E,h \ e :' NT \ P,E,h \ e\F{D} :' T" | "\ P,E,h \ e\<^sub>1 :' Class C; P \ C has F:T in D; P,E,h \ e\<^sub>2 :' T\<^sub>2; P \ T\<^sub>2 \ T \ \ P,E,h \ e\<^sub>1\F{D}:=e\<^sub>2 :' Void" | "\ P,E,h \ e\<^sub>1:'NT; P,E,h \ e\<^sub>2 :' T\<^sub>2 \ \ P,E,h \ e\<^sub>1\F{D}:=e\<^sub>2 :' Void" | "\ P,E,h \ e :' Class C; P \ C sees M:Ts \ T = (pns,body) in D; P,E,h \ es [:'] Ts'; P \ Ts' [\] Ts \ \ P,E,h \ e\M(es) :' T" | "\ P,E,h \ e :' NT; P,E,h \ es [:'] Ts \ \ P,E,h \ e\M(es) :' T" | "P,E,h \ [] [:'] []" | "\ P,E,h \ e :' T; P,E,h \ es [:'] Ts \ \ P,E,h \ e#es [:'] T#Ts" | "\ typeof\<^bsub>h\<^esub> v = Some T\<^sub>1; P \ T\<^sub>1 \ T; P,E(V\T),h \ e\<^sub>2 :' T\<^sub>2 \ \ P,E,h \ {V:T := Val v; e\<^sub>2} :' T\<^sub>2" | "\ P,E(V\T),h \ e :' T'; \ assigned V e \ \ P,E,h \ {V:T; e} :' T'" | "\ P,E,h \ e\<^sub>1:' T\<^sub>1; P,E,h \ e\<^sub>2:'T\<^sub>2 \ \ P,E,h \ e\<^sub>1;;e\<^sub>2 :' T\<^sub>2" | "\ P,E,h \ e :' Boolean; P,E,h \ e\<^sub>1:' T\<^sub>1; P,E,h \ e\<^sub>2:' T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ P \ T\<^sub>2 \ T\<^sub>1; P \ T\<^sub>1 \ T\<^sub>2 \ T = T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ T = T\<^sub>1 \ \ P,E,h \ if (e) e\<^sub>1 else e\<^sub>2 :' T" (* "\ P,E,h \ e :' Boolean; P,E,h \ e\<^sub>1:' T\<^sub>1; P,E,h \ e\<^sub>2:' T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ \ P,E,h \ if (e) e\<^sub>1 else e\<^sub>2 :' T\<^sub>2" "\ P,E,h \ e :' Boolean; P,E,h \ e\<^sub>1:' T\<^sub>1; P,E,h \ e\<^sub>2:' T\<^sub>2; P \ T\<^sub>2 \ T\<^sub>1 \ \ P,E,h \ if (e) e\<^sub>1 else e\<^sub>2 :' T\<^sub>1" *) | "\ P,E,h \ e :' Boolean; P,E,h \ c:' T \ \ P,E,h \ while(e) c :' Void" | "\ P,E,h \ e :' T\<^sub>r; is_refT T\<^sub>r \ \ P,E,h \ throw e :' T" | "\ P,E,h \ e\<^sub>1 :' T\<^sub>1; P,E(V \ Class C),h \ e\<^sub>2 :' T\<^sub>2; P \ T\<^sub>1 \ T\<^sub>2 \ \ P,E,h \ try e\<^sub>1 catch(C V) e\<^sub>2 :' T\<^sub>2" (*<*) lemmas WTrt'_induct = WTrt'_WTrts'.induct [split_format (complete)] and WTrt'_inducts = WTrt'_WTrts'.inducts [split_format (complete)] inductive_cases WTrt'_elim_cases[elim!]: "P,E,h \ V :=e :' T" (*>*) lemma [iff]: "P,E,h \ e\<^sub>1;;e\<^sub>2 :' T\<^sub>2 = (\T\<^sub>1. P,E,h \ e\<^sub>1:' T\<^sub>1 \ P,E,h \ e\<^sub>2:' T\<^sub>2)" (*<*) apply(rule iffI) apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) done (*>*) lemma [iff]: "P,E,h \ Val v :' T = (typeof\<^bsub>h\<^esub> v = Some T)" (*<*) apply(rule iffI) apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) done (*>*) lemma [iff]: "P,E,h \ Var v :' T = (E v = Some T)" (*<*) apply(rule iffI) apply (auto elim: WTrt'.cases intro!:WTrt'_WTrts'.intros) done (*>*) lemma wt_wt': "P,E,h \ e : T \ P,E,h \ e :' T" and wts_wts': "P,E,h \ es [:] Ts \ P,E,h \ es [:'] Ts" (*<*) apply (induct rule:WTrt_inducts) prefer 14 apply(case_tac "assigned V e") apply(clarsimp simp add:fun_upd_same assigned_def simp del:fun_upd_apply) apply(erule (2) WTrt'_WTrts'.intros) apply(erule (1) WTrt'_WTrts'.intros) apply(blast intro:WTrt'_WTrts'.intros)+ done (*>*) lemma wt'_wt: "P,E,h \ e :' T \ P,E,h \ e : T" and wts'_wts: "P,E,h \ es [:'] Ts \ P,E,h \ es [:] Ts" (*<*) apply (induct rule:WTrt'_inducts) prefer 16 apply(rule WTrt_WTrts.intros) apply(rule WTrt_WTrts.intros) apply(rule WTrt_WTrts.intros) apply simp apply(erule (2) WTrt_WTrts.intros) apply(blast intro:WTrt_WTrts.intros)+ done (*>*) corollary wt'_iff_wt: "(P,E,h \ e :' T) = (P,E,h \ e : T)" (*<*)by(blast intro:wt_wt' wt'_wt)(*>*) corollary wts'_iff_wts: "(P,E,h \ es [:'] Ts) = (P,E,h \ es [:] Ts)" (*<*)by(blast intro:wts_wts' wts'_wts)(*>*) (*<*) lemmas WTrt_inducts2 = WTrt'_inducts [unfolded wt'_iff_wt wts'_iff_wts, case_names WTrtNew WTrtCast WTrtVal WTrtVar WTrtBinOpEq WTrtBinOpAdd WTrtLAss WTrtFAcc WTrtFAccNT WTrtFAss WTrtFAssNT WTrtCall WTrtCallNT WTrtNil WTrtCons WTrtInitBlock WTrtBlock WTrtSeq WTrtCond WTrtWhile WTrtThrow WTrtTry, consumes 1] (*>*) theorem assumes wf: "wwf_J_prog P" and hconf: "P \ h \" shows progress: "P,E,h \ e : T \ (\l. \ \ e \dom l\; \ final e \ \ \e' s'. P \ \e,(h,l)\ \ \e',s'\)" and "P,E,h \ es [:] Ts \ (\l. \ \s es \dom l\; \ finals es \ \ \es' s'. P \ \es,(h,l)\ [\] \es',s'\)" (*<*) proof (induct rule:WTrt_inducts2) case WTrtNew show ?case proof cases assume "\a. h a = None" with assms WTrtNew show ?thesis by (fastforce del:exE intro!:RedNew simp add:new_Addr_def elim!:wf_Fields_Ex[THEN exE]) next assume "\(\a. h a = None)" with assms WTrtNew show ?thesis by(fastforce intro:RedNewFail simp add:new_Addr_def) qed next case (WTrtCast E e T C) have wte: "P,E,h \ e : T" and ref: "is_refT T" and IH: "\l. \\ e \dom l\; \ final e\ \ \e' s'. P \ \e,(h,l)\ \ \e',s'\" and D: "\ (Cast C e) \dom l\" by fact+ from D have De: "\ e \dom l\" by auto show ?case proof cases assume "final e" with wte ref show ?thesis proof (rule finalRefE) assume "e = null" thus ?case by(fastforce intro:RedCastNull) next fix D a assume A: "T = Class D" "e = addr a" show ?thesis proof cases assume "P \ D \\<^sup>* C" thus ?thesis using A wte by(fastforce intro:RedCast) next assume "\ P \ D \\<^sup>* C" thus ?thesis using A wte by(force intro!:RedCastFail) qed next fix a assume "e = Throw a" thus ?thesis by(blast intro!:red_reds.CastThrow) qed next assume nf: "\ final e" from IH[OF De nf] show ?thesis by (blast intro:CastRed) qed next case WTrtVal thus ?case by(simp add:final_def) next case WTrtVar thus ?case by(fastforce intro:RedVar simp:hyper_isin_def) next case (WTrtBinOpEq E e1 T1 e2 T2) show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v1 assume [simp]: "e1 = Val v1" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v2 assume "e2 = Val v2" thus ?thesis using WTrtBinOpEq by(fastforce intro:RedBinOp) next fix a assume "e2 = Throw a" thus ?thesis by(auto intro:red_reds.BinOpThrow2) qed next assume "\ final e2" with WTrtBinOpEq show ?thesis by simp (fast intro!:BinOpRed2) qed next fix a assume "e1 = Throw a" thus ?thesis by simp (fast intro:red_reds.BinOpThrow1) qed next assume "\ final e1" with WTrtBinOpEq show ?thesis by simp (fast intro:BinOpRed1) qed next case (WTrtBinOpAdd E e1 e2) show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v1 assume [simp]: "e1 = Val v1" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v2 assume "e2 = Val v2" thus ?thesis using WTrtBinOpAdd by(fastforce intro:RedBinOp) next fix a assume "e2 = Throw a" thus ?thesis by(auto intro:red_reds.BinOpThrow2) qed next assume "\ final e2" with WTrtBinOpAdd show ?thesis by simp (fast intro!:BinOpRed2) qed next fix a assume "e1 = Throw a" thus ?thesis by simp (fast intro:red_reds.BinOpThrow1) qed next assume "\ final e1" with WTrtBinOpAdd show ?thesis by simp (fast intro:BinOpRed1) qed next case (WTrtLAss E V T e T') show ?case proof cases assume "final e" with WTrtLAss show ?thesis by(auto simp:final_def intro!:RedLAss red_reds.LAssThrow) next assume "\ final e" with WTrtLAss show ?thesis by simp (fast intro:LAssRed) qed next case (WTrtFAcc E e C F T D) have wte: "P,E,h \ e : Class C" and field: "P \ C has F:T in D" by fact+ show ?case proof cases assume "final e" with wte show ?thesis proof (rule final_addrE) fix a assume e: "e = addr a" with wte obtain fs where hp: "h a = Some(C,fs)" by auto with hconf have "P,h \ (C,fs) \" using hconf_def by fastforce then obtain v where "fs(F,D) = Some v" using field by(fastforce dest:has_fields_fun simp:oconf_def has_field_def) with hp e show ?thesis by(fastforce intro:RedFAcc) next fix a assume "e = Throw a" thus ?thesis by(fastforce intro:red_reds.FAccThrow) qed next assume "\ final e" with WTrtFAcc show ?thesis by(fastforce intro!:FAccRed) qed next case (WTrtFAccNT E e F D T) show ?case proof cases assume "final e" \ \@{term e} is @{term null} or @{term throw}\ with WTrtFAccNT show ?thesis by(fastforce simp:final_def intro: RedFAccNull red_reds.FAccThrow) next assume "\ final e" \ \@{term e} reduces by IH\ with WTrtFAccNT show ?thesis by simp (fast intro:FAccRed) qed next case (WTrtFAss E e1 C F T D e2 T2) have wte1: "P,E,h \ e1 : Class C" by fact show ?case proof cases assume "final e1" with wte1 show ?thesis proof (rule final_addrE) fix a assume e1: "e1 = addr a" show ?thesis proof cases assume "final e2" thus ?thesis proof (rule finalE) fix v assume "e2 = Val v" thus ?thesis using e1 wte1 by(fastforce intro:RedFAss) next fix a assume "e2 = Throw a" thus ?thesis using e1 by(fastforce intro:red_reds.FAssThrow2) qed next assume "\ final e2" with WTrtFAss e1 show ?thesis by simp (fast intro!:FAssRed2) qed next fix a assume "e1 = Throw a" thus ?thesis by(fastforce intro:red_reds.FAssThrow1) qed next assume "\ final e1" with WTrtFAss show ?thesis by simp (blast intro!:FAssRed1) qed next case (WTrtFAssNT E e\<^sub>1 e\<^sub>2 T\<^sub>2 F D) show ?case proof cases assume e1: "final e\<^sub>1" \ \@{term e\<^sub>1} is @{term null} or @{term throw}\ show ?thesis proof cases assume "final e\<^sub>2" \ \@{term e\<^sub>2} is @{term Val} or @{term throw}\ with WTrtFAssNT e1 show ?thesis by(fastforce simp:final_def intro: RedFAssNull red_reds.FAssThrow1 red_reds.FAssThrow2) next assume "\ final e\<^sub>2" \ \@{term e\<^sub>2} reduces by IH\ with WTrtFAssNT e1 show ?thesis by (fastforce simp:final_def intro!:red_reds.FAssRed2 red_reds.FAssThrow1) qed next assume "\ final e\<^sub>1" \ \@{term e\<^sub>1} reduces by IH\ with WTrtFAssNT show ?thesis by (fastforce intro:FAssRed1) qed next case (WTrtCall E e C M Ts T pns body D es Ts') have wte: "P,E,h \ e : Class C" and "method": "P \ C sees M:Ts\T = (pns,body) in D" and wtes: "P,E,h \ es [:] Ts'"and sub: "P \ Ts' [\] Ts" and IHes: "\l. \\s es \dom l\; \ finals es\ \ \es' s'. P \ \es,(h,l)\ [\] \es',s'\" and D: "\ (e\M(es)) \dom l\" by fact+ show ?case proof cases assume "final e" with wte show ?thesis proof (rule final_addrE) fix a assume e_addr: "e = addr a" show ?thesis proof cases assume es: "\vs. es = map Val vs" from wte e_addr obtain fs where ha: "h a = Some(C,fs)" by auto show ?thesis using e_addr ha "method" WTrts_same_length[OF wtes] sub es sees_wf_mdecl[OF wf "method"] by (fastforce intro!: RedCall simp:list_all2_iff wf_mdecl_def) next assume "\(\vs. es = map Val vs)" hence not_all_Val: "\(\e \ set es. \v. e = Val v)" by(simp add:ex_map_conv) let ?ves = "takeWhile (\e. \v. e = Val v) es" let ?rest = "dropWhile (\e. \v. e = Val v) es" let ?ex = "hd ?rest" let ?rst = "tl ?rest" from not_all_Val have nonempty: "?rest \ []" by auto hence es: "es = ?ves @ ?ex # ?rst" by simp have "\e \ set ?ves. \v. e = Val v" by(fastforce dest:set_takeWhileD) then obtain vs where ves: "?ves = map Val vs" using ex_map_conv by blast show ?thesis proof cases assume "final ?ex" moreover from nonempty have "\(\v. ?ex = Val v)" by(auto simp:neq_Nil_conv simp del:dropWhile_eq_Nil_conv) (simp add:dropWhile_eq_Cons_conv) ultimately obtain b where ex_Throw: "?ex = Throw b" by(fast elim!:finalE) show ?thesis using e_addr es ex_Throw ves by(fastforce intro:CallThrowParams) next assume not_fin: "\ final ?ex" have "finals es = finals(?ves @ ?ex # ?rst)" using es by(rule arg_cong) also have "\ = finals(?ex # ?rst)" using ves by simp finally have "finals es = finals(?ex # ?rst)" . hence "\ finals es" using not_finals_ConsI[OF not_fin] by blast thus ?thesis using e_addr D IHes by(fastforce intro!:CallParams) qed qed next fix a assume "e = Throw a" with WTrtCall.prems show ?thesis by(fast intro!:CallThrowObj) qed next assume "\ final e" with WTrtCall show ?thesis by simp (blast intro!:CallObj) qed next case (WTrtCallNT E e es Ts M T) show ?case proof cases assume "final e" moreover { fix v assume e: "e = Val v" hence "e = null" using WTrtCallNT by simp have ?case proof cases assume "finals es" moreover { fix vs assume "es = map Val vs" with WTrtCallNT e have ?thesis by(fastforce intro: RedCallNull) } moreover { fix vs a es' assume "es = map Val vs @ Throw a # es'" with WTrtCallNT e have ?thesis by(fastforce intro: CallThrowParams) } ultimately show ?thesis by(fastforce simp:finals_def) next assume "\ finals es" \ \@{term es} reduces by IH\ with WTrtCallNT e show ?thesis by(fastforce intro: CallParams) qed } moreover { fix a assume "e = Throw a" with WTrtCallNT have ?case by(fastforce intro: CallThrowObj) } ultimately show ?thesis by(fastforce simp:final_def) next assume "\ final e" \ \@{term e} reduces by IH\ with WTrtCallNT show ?thesis by (fastforce intro:CallObj) qed next case WTrtNil thus ?case by simp next case (WTrtCons E e T es Ts) have IHe: "\l. \\ e \dom l\; \ final e\ \ \e' s'. P \ \e,(h,l)\ \ \e',s'\" and IHes: "\l. \\s es \dom l\; \ finals es\ \ \es' s'. P \ \es,(h,l)\ [\] \es',s'\" and D: "\s (e#es) \dom l\" and not_fins: "\ finals(e # es)" by fact+ have De: "\ e \dom l\" and Des: "\s es (\dom l\ \ \ e)" using D by auto show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume e: "e = Val v" hence Des': "\s es \dom l\" using De Des by auto have not_fins_tl: "\ finals es" using not_fins e by simp show ?thesis using e IHes[OF Des' not_fins_tl] by (blast intro!:ListRed2) next fix a assume "e = Throw a" hence False using not_fins by simp thus ?thesis .. qed next assume "\ final e" with IHe[OF De] show ?thesis by(fast intro!:ListRed1) qed next case (WTrtInitBlock v T\<^sub>1 T E V e\<^sub>2 T\<^sub>2) have IH2: "\l. \\ e\<^sub>2 \dom l\; \ final e\<^sub>2\ \ \e' s'. P \ \e\<^sub>2,(h,l)\ \ \e',s'\" and D: "\ {V:T := Val v; e\<^sub>2} \dom l\" by fact+ show ?case proof cases assume "final e\<^sub>2" then show ?thesis proof (rule finalE) fix v\<^sub>2 assume "e\<^sub>2 = Val v\<^sub>2" thus ?thesis by(fast intro:RedInitBlock) next fix a assume "e\<^sub>2 = Throw a" thus ?thesis by(fast intro:red_reds.InitBlockThrow) qed next assume not_fin2: "\ final e\<^sub>2" from D have D2: "\ e\<^sub>2 \dom(l(V\v))\" by (auto simp:hyperset_defs) from IH2[OF D2 not_fin2] obtain h' l' e' where red2: "P \ \e\<^sub>2,(h, l(V\v))\ \ \e',(h', l')\" by auto from red_lcl_incr[OF red2] have "V \ dom l'" by auto with red2 show ?thesis by(fastforce intro:InitBlockRed) qed next case (WTrtBlock E V T e T') have IH: "\l. \\ e \dom l\; \ final e\ \ \e' s'. P \ \e,(h,l)\ \ \e',s'\" and unass: "\ assigned V e" and D: "\ {V:T; e} \dom l\" by fact+ show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume "e = Val v" thus ?thesis by(fast intro:RedBlock) next fix a assume "e = Throw a" thus ?thesis by(fast intro:red_reds.BlockThrow) qed next assume not_fin: "\ final e" from D have De: "\ e \dom(l(V:=None))\" by(simp add:hyperset_defs) from IH[OF De not_fin] obtain h' l' e' where red: "P \ \e,(h,l(V:=None))\ \ \e',(h',l')\" by auto show ?thesis proof (cases "l' V") assume "l' V = None" with red unass show ?thesis by(blast intro: BlockRedNone) next fix v assume "l' V = Some v" with red unass show ?thesis by(blast intro: BlockRedSome) qed qed next case (WTrtSeq E e1 T1 e2 T2) show ?case proof cases assume "final e1" thus ?thesis by(fast elim:finalE intro:RedSeq red_reds.SeqThrow) next assume "\ final e1" with WTrtSeq show ?thesis by simp (blast intro:SeqRed) qed next case (WTrtCond E e e\<^sub>1 T\<^sub>1 e\<^sub>2 T\<^sub>2 T) have wt: "P,E,h \ e : Boolean" by fact show ?case proof cases assume "final e" thus ?thesis proof (rule finalE) fix v assume val: "e = Val v" then obtain b where v: "v = Bool b" using wt by auto show ?thesis proof (cases b) case True with val v show ?thesis by(auto intro:RedCondT) next case False with val v show ?thesis by(auto intro:RedCondF) qed next fix a assume "e = Throw a" thus ?thesis by(fast intro:red_reds.CondThrow) qed next assume "\ final e" with WTrtCond show ?thesis by simp (fast intro:CondRed) qed next case WTrtWhile show ?case by(fast intro:RedWhile) next case (WTrtThrow E e T\<^sub>r T) show ?case proof cases assume "final e" \ \Then @{term e} must be @{term throw} or @{term null}\ with WTrtThrow show ?thesis by(fastforce simp:final_def is_refT_def intro:red_reds.ThrowThrow red_reds.RedThrowNull) next assume "\ final e" \ \Then @{term e} must reduce\ with WTrtThrow show ?thesis by simp (blast intro:ThrowRed) qed next case (WTrtTry E e1 T1 V C e2 T2) have wt1: "P,E,h \ e1 : T1" by fact show ?case proof cases assume "final e1" thus ?thesis proof (rule finalE) fix v assume "e1 = Val v" thus ?thesis by(fast intro:RedTry) next fix a assume e1_Throw: "e1 = Throw a" with wt1 obtain D fs where ha: "h a = Some(D,fs)" by fastforce show ?thesis proof cases assume "P \ D \\<^sup>* C" with e1_Throw ha show ?thesis by(fastforce intro!:RedTryCatch) next assume "\ P \ D \\<^sup>* C" with e1_Throw ha show ?thesis by(force intro!:RedTryFail) qed qed next assume "\ final e1" with WTrtTry show ?thesis by simp (fast intro:TryRed) qed qed (*>*) end diff --git a/thys/JinjaThreads/BV/JVM_SemiType.thy b/thys/JinjaThreads/BV/JVM_SemiType.thy --- a/thys/JinjaThreads/BV/JVM_SemiType.thy +++ b/thys/JinjaThreads/BV/JVM_SemiType.thy @@ -1,235 +1,235 @@ (* Title: JinjaThreads/BV/JVM_SemiType.thy Author: Gerwin Klein, Andreas Lochbihler Based on the theory Jinja/BV/JVM_SemiType *) chapter \Bytecode verifier\ section \The JVM Type System as Semilattice\ theory JVM_SemiType imports "../Common/SemiType" begin type_synonym ty\<^sub>l = "ty err list" type_synonym ty\<^sub>s = "ty list" type_synonym ty\<^sub>i = "ty\<^sub>s \ ty\<^sub>l" type_synonym ty\<^sub>i' = "ty\<^sub>i option" type_synonym ty\<^sub>m = "ty\<^sub>i' list" type_synonym ty\<^sub>P = "mname \ cname \ ty\<^sub>m" definition stk_esl :: "'c prog \ nat \ ty\<^sub>s esl" where "stk_esl P mxs \ upto_esl mxs (SemiType.esl P)" definition loc_sl :: "'c prog \ nat \ ty\<^sub>l sl" where "loc_sl P mxl \ Listn.sl mxl (Err.sl (SemiType.esl P))" definition sl :: "'c prog \ nat \ nat \ ty\<^sub>i' err sl" where "sl P mxs mxl \ Err.sl(Opt.esl(Product.esl (stk_esl P mxs) (Err.esl(loc_sl P mxl))))" definition "states" :: "'c prog \ nat \ nat \ ty\<^sub>i' err set" where "states P mxs mxl \ fst(sl P mxs mxl)" definition le :: "'c prog \ nat \ nat \ ty\<^sub>i' err ord" where "le P mxs mxl \ fst(snd(sl P mxs mxl))" definition sup :: "'c prog \ nat \ nat \ ty\<^sub>i' err binop" where "sup P mxs mxl \ snd(snd(sl P mxs mxl))" definition sup_ty_opt :: "['c prog,ty err,ty err] \ bool" ("_ \ _ \\<^sub>\ _" [71,71,71] 70) where "sup_ty_opt P \ Err.le (widen P)" definition sup_state :: "['c prog,ty\<^sub>i,ty\<^sub>i] \ bool" ("_ \ _ \\<^sub>i _" [71,71,71] 70) where "sup_state P \ Product.le (Listn.le (widen P)) (Listn.le (sup_ty_opt P))" definition sup_state_opt :: "['c prog,ty\<^sub>i',ty\<^sub>i'] \ bool" - ("_ \ _ \' _" [71,71,71] 70) + ("_ \ _ \'' _" [71,71,71] 70) where "sup_state_opt P \ Opt.le (sup_state P)" abbreviation sup_loc :: "['c prog,ty\<^sub>l,ty\<^sub>l] \ bool" ("_ \ _ [\\<^sub>\] _" [71,71,71] 70) where "P \ LT [\\<^sub>\] LT' \ list_all2 (sup_ty_opt P) LT LT'" notation (ASCII) sup_ty_opt ("_ |- _ <=T _" [71,71,71] 70) and sup_state ("_ |- _ <=i _" [71,71,71] 70) and sup_state_opt ("_ |- _ <=' _" [71,71,71] 70) and sup_loc ("_ |- _ [<=T] _" [71,71,71] 70) subsection "Unfolding" lemma JVM_states_unfold: "states P mxs mxl \ err(opt((Union {list n (types P) |n. n <= mxs}) \ list mxl (err(types P))))" apply (unfold states_def sl_def Opt.esl_def Err.sl_def stk_esl_def loc_sl_def Product.esl_def Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def) apply simp done lemma JVM_le_unfold: "le P m n \ Err.le(Opt.le(Product.le(Listn.le(widen P))(Listn.le(Err.le(widen P)))))" apply (unfold le_def sl_def Opt.esl_def Err.sl_def stk_esl_def loc_sl_def Product.esl_def Listn.sl_def upto_esl_def SemiType.esl_def Err.esl_def) apply simp done lemma sl_def2: "JVM_SemiType.sl P mxs mxl \ (states P mxs mxl, JVM_SemiType.le P mxs mxl, JVM_SemiType.sup P mxs mxl)" by (unfold JVM_SemiType.sup_def states_def JVM_SemiType.le_def) simp lemma JVM_le_conv: "le P m n (OK t1) (OK t2) = P \ t1 \' t2" by (simp add: JVM_le_unfold Err.le_def lesub_def sup_state_opt_def sup_state_def sup_ty_opt_def) lemma JVM_le_Err_conv: "le P m n = Err.le (sup_state_opt P)" by (unfold sup_state_opt_def sup_state_def sup_ty_opt_def JVM_le_unfold) simp lemma err_le_unfold [iff]: "Err.le r (OK a) (OK b) = r a b" by (simp add: Err.le_def lesub_def) subsection \Semilattice\ lemma order_sup_state_opt [intro, simp]: "wf_prog wf_mb P \ order (sup_state_opt P)" by (unfold sup_state_opt_def sup_state_def sup_ty_opt_def) blast lemma semilat_JVM [intro?]: "wf_prog wf_mb P \ semilat (JVM_SemiType.sl P mxs mxl)" apply (unfold JVM_SemiType.sl_def stk_esl_def loc_sl_def) apply (blast intro: err_semilat_Product_esl err_semilat_upto_esl Listn_sl err_semilat_JType_esl) done lemma acc_JVM [intro]: "wf_prog wf_mb P \ acc (JVM_SemiType.states P mxs mxl) (JVM_SemiType.le P mxs mxl)" by(unfold JVM_le_unfold JVM_states_unfold) blast subsection \Widening with \\\\ lemma widen_refl[iff]: "widen P t t" by (simp add: fun_of_def) lemma sup_ty_opt_refl [iff]: "P \ T \\<^sub>\ T" apply (unfold sup_ty_opt_def) apply (fold lesub_def) apply (rule le_err_refl) apply (simp add: lesub_def) done lemma Err_any_conv [iff]: "P \ Err \\<^sub>\ T = (T = Err)" by (unfold sup_ty_opt_def) (rule Err_le_conv [simplified lesub_def]) lemma any_Err [iff]: "P \ T \\<^sub>\ Err" by (unfold sup_ty_opt_def) (rule le_Err [simplified lesub_def]) lemma OK_OK_conv [iff]: "P \ OK T \\<^sub>\ OK T' = P \ T \ T'" by (simp add: sup_ty_opt_def fun_of_def) lemma any_OK_conv [iff]: "P \ X \\<^sub>\ OK T' = (\T. X = OK T \ P \ T \ T')" apply (unfold sup_ty_opt_def) apply (rule le_OK_conv [simplified lesub_def]) done lemma OK_any_conv: "P \ OK T \\<^sub>\ X = (X = Err \ (\T'. X = OK T' \ P \ T \ T'))" apply (unfold sup_ty_opt_def) apply (rule OK_le_conv [simplified lesub_def]) done lemma sup_ty_opt_trans [intro?, trans]: "\P \ a \\<^sub>\ b; P \ b \\<^sub>\ c\ \ P \ a \\<^sub>\ c" by (auto intro: widen_trans simp add: sup_ty_opt_def Err.le_def lesub_def fun_of_def split: err.splits) subsection "Stack and Registers" lemma stk_convert: "P \ ST [\] ST' = Listn.le (widen P) ST ST'" by (simp add: Listn.le_def lesub_def) lemma sup_loc_refl [iff]: "P \ LT [\\<^sub>\] LT" by (rule list_all2_refl) simp lemmas sup_loc_Cons1 [iff] = list_all2_Cons1 [of "sup_ty_opt P"] for P lemma sup_loc_def: "P \ LT [\\<^sub>\] LT' \ Listn.le (sup_ty_opt P) LT LT'" by (simp add: Listn.le_def lesub_def) lemma sup_loc_widens_conv [iff]: "P \ map OK Ts [\\<^sub>\] map OK Ts' = P \ Ts [\] Ts'" by (simp add: list_all2_map1 list_all2_map2) lemma sup_loc_trans [intro?, trans]: "\P \ a [\\<^sub>\] b; P \ b [\\<^sub>\] c\ \ P \ a [\\<^sub>\] c" by (rule list_all2_trans, rule sup_ty_opt_trans) subsection "State Type" lemma sup_state_conv [iff]: "P \ (ST,LT) \\<^sub>i (ST',LT') = (P \ ST [\] ST' \ P \ LT [\\<^sub>\] LT')" by (auto simp add: sup_state_def stk_convert lesub_def Product.le_def sup_loc_def) lemma sup_state_conv2: "P \ s1 \\<^sub>i s2 = (P \ fst s1 [\] fst s2 \ P \ snd s1 [\\<^sub>\] snd s2)" by (cases s1, cases s2) simp lemma sup_state_refl [iff]: "P \ s \\<^sub>i s" by (auto simp add: sup_state_conv2 intro: list_all2_refl) lemma sup_state_trans [intro?, trans]: "\P \ a \\<^sub>i b; P \ b \\<^sub>i c\ \ P \ a \\<^sub>i c" by (auto intro: sup_loc_trans widens_trans simp add: sup_state_conv2) lemma sup_state_opt_None_any [iff]: "P \ None \' s" by (simp add: sup_state_opt_def Opt.le_def) lemma sup_state_opt_any_None [iff]: "P \ s \' None = (s = None)" by (simp add: sup_state_opt_def Opt.le_def) lemma sup_state_opt_Some_Some [iff]: "P \ Some a \' Some b = P \ a \\<^sub>i b" by (simp add: sup_state_opt_def Opt.le_def lesub_def) lemma sup_state_opt_any_Some: "P \ (Some s) \' X = (\s'. X = Some s' \ P \ s \\<^sub>i s')" by (simp add: sup_state_opt_def Opt.le_def lesub_def) lemma sup_state_opt_refl [iff]: "P \ s \' s" by (simp add: sup_state_opt_def Opt.le_def lesub_def) lemma sup_state_opt_trans [intro?, trans]: "\P \ a \' b; P \ b \' c\ \ P \ a \' c" apply (unfold sup_state_opt_def Opt.le_def lesub_def) apply (simp del: split_paired_All) apply (rule sup_state_trans, assumption+) done end diff --git a/thys/JinjaThreads/Compiler/J1JVMBisim.thy b/thys/JinjaThreads/Compiler/J1JVMBisim.thy --- a/thys/JinjaThreads/Compiler/J1JVMBisim.thy +++ b/thys/JinjaThreads/Compiler/J1JVMBisim.thy @@ -1,5582 +1,5582 @@ (* Title: JinjaThreads/Compiler/J1JVMBisim.thy Author: Andreas Lochbihler *) section \The delay bisimulation between intermediate language and JVM\ theory J1JVMBisim imports Execs "../BV/BVNoTypeError" J1 begin declare Listn.lesub_list_impl_same_size[simp del] lemma (in JVM_heap_conf_base') \exec_1_\exec_1_d: "\ wf_jvm_prog\<^bsub>\\<^esub> P; \exec_1 P t \ \'; \ |- t:\ [ok] \ \ \exec_1_d P t \ \'" by(auto simp add: \exec_1_def \exec_1_d_def welltyped_commute[symmetric] elim: jvmd_NormalE) context JVM_conf_read begin lemma \Exec_1r_preserves_correct_state: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and exec: "\Exec_1r P t \ \'" shows "\ |- t:\ [ok] \ \ |- t:\' [ok]" using exec by(induct)(blast intro: BV_correct_1[OF wf])+ lemma \Exec_1t_preserves_correct_state: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and exec: "\Exec_1t P t \ \'" shows "\ |- t:\ [ok] \ \ |- t:\' [ok]" using exec by(induct)(blast intro: BV_correct_1[OF wf])+ lemma \Exec_1r_\Exec_1_dr: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "\ \Exec_1r P t \ \'; \ |- t:\ [ok] \ \ \Exec_1_dr P t \ \'" apply(induct rule: rtranclp_induct) apply(blast intro: rtranclp.rtrancl_into_rtrancl \exec_1_\exec_1_d[OF wf] \Exec_1r_preserves_correct_state[OF wf])+ done lemma \Exec_1t_\Exec_1_dt: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" shows "\ \Exec_1t P t \ \'; \ |- t:\ [ok] \ \ \Exec_1_dt P t \ \'" apply(induct rule: tranclp_induct) apply(blast intro: tranclp.trancl_into_trancl \exec_1_\exec_1_d[OF wf] \Exec_1t_preserves_correct_state[OF wf])+ done lemma \Exec_1_dr_preserves_correct_state: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and exec: "\Exec_1_dr P t \ \'" shows "\ |- t: \ [ok] \ \ |- t: \' [ok]" using exec by(induct)(blast intro: BV_correct_1[OF wf])+ lemma \Exec_1_dt_preserves_correct_state: assumes wf: "wf_jvm_prog\<^bsub>\\<^esub> P" and exec: "\Exec_1_dt P t \ \'" shows "\ |- t:\ [ok] \ \ |- t:\' [ok]" using exec by(induct)(blast intro: BV_correct_1[OF wf])+ end locale J1_JVM_heap_base = J1_heap_base + JVM_heap_base + constrains addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" begin inductive bisim1 :: "'m prog \ 'heap \ 'addr expr1 \ ('addr expr1 \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" and bisims1 :: "'m prog \ 'heap \ 'addr expr1 list \ ('addr expr1 list \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" and bisim1_syntax :: "'m prog \ 'addr expr1 \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" ("_,_,_ \ _ \ _" [50, 0, 0, 0, 50] 100) and bisims1_syntax :: "'m prog \ 'addr expr1 list \ 'heap \ ('addr expr1 list \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" ("_,_,_ \ _ [\] _" [50, 0, 0, 0, 50] 100) for P :: "'m prog" and h :: 'heap where "P, e, h \ exs \ s \ bisim1 P h e exs s" | "P, es, h \ esxs [\] s \ bisims1 P h es esxs s" | bisim1Val2: "pc = length (compE2 e) \ P, e, h \ (Val v, xs) \ (v # [], xs, pc, None)" | bisim1New: "P, new C, h \ (new C, xs) \ ([], xs, 0, None)" | bisim1NewThrow: "P, new C, h \ (THROW OutOfMemory, xs) \ ([], xs, 0, \addr_of_sys_xcpt OutOfMemory\)" | bisim1NewArray: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, newA T\e\, h \ (newA T\e'\, xs) \ (stk, loc, pc, xcp)" | bisim1NewArrayThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, newA T\e\, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1NewArrayFail: "P, newA T\e\, h \ (Throw a, xs) \ ([v], xs, length (compE2 e), \a\)" | bisim1Cast: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, Cast T e, h \ (Cast T e', xs) \ (stk, loc, pc, xcp)" | bisim1CastThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, Cast T e, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1CastFail: "P, Cast T e, h \ (THROW ClassCast, xs) \ ([v], xs, length (compE2 e), \addr_of_sys_xcpt ClassCast\)" | bisim1InstanceOf: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e instanceof T, h \ (e' instanceof T, xs) \ (stk, loc, pc, xcp)" | bisim1InstanceOfThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, e instanceof T, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Val: "P, Val v, h \ (Val v, xs) \ ([], xs, 0, None)" | bisim1Var: "P, Var V, h \ (Var V, xs) \ ([], xs, 0, None)" | bisim1BinOp1: "P, e1, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e1\bop\e2, h \ (e'\bop\e2, xs) \ (stk, loc, pc, xcp)" | bisim1BinOp2: "P, e2, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e1\bop\e2, h \ (Val v1 \bop\ e', xs) \ (stk @ [v1], loc, length (compE2 e1) + pc, xcp)" | bisim1BinOpThrow1: "P, e1, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, e1\bop\e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1BinOpThrow2: "P, e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, e1\bop\e2, h \ (Throw a, xs) \ (stk @ [v1], loc, length (compE2 e1) + pc, \a\)" | bisim1BinOpThrow: "P, e1\bop\e2, h \ (Throw a, xs) \ ([v1, v2], xs, length (compE2 e1) + length (compE2 e2), \a\)" | bisim1LAss1: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, V:=e, h \ (V:=e', xs) \ (stk, loc, pc, xcp)" | bisim1LAss2: "P, V:=e, h \ (unit, xs) \ ([], xs, Suc (length (compE2 e)), None)" | bisim1LAssThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, V:=e, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1AAcc1: "P, a, h \ (a', xs) \ (stk, loc, pc, xcp) \ P, a\i\, h \ (a'\i\, xs) \ (stk, loc, pc, xcp)" | bisim1AAcc2: "P, i, h \ (i', xs) \ (stk, loc, pc, xcp) \ P, a\i\, h \ (Val v\i'\, xs) \ (stk @ [v], loc, length (compE2 a) + pc, xcp)" | bisim1AAccThrow1: "P, a, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\i\, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1AAccThrow2: "P, i, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\i\, h \ (Throw ad, xs) \ (stk @ [v], loc, length (compE2 a) + pc, \ad\)" | bisim1AAccFail: "P, a\i\, h \ (Throw ad, xs) \ ([v, v'], xs, length (compE2 a) + length (compE2 i), \ad\)" | bisim1AAss1: "P, a, h \ (a', xs) \ (stk, loc, pc, xcp) \ P, a\i\ := e, h \ (a'\i\ := e, xs) \ (stk, loc, pc, xcp)" | bisim1AAss2: "P, i, h \ (i', xs) \ (stk, loc, pc, xcp) \ P, a\i\ := e, h \ (Val v\i'\ := e, xs) \ (stk @ [v], loc, length (compE2 a) + pc, xcp)" | bisim1AAss3: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, a\i\ := e, h \ (Val v\Val v'\ := e', xs) \ (stk @ [v', v], loc, length (compE2 a) + length (compE2 i) + pc, xcp)" | bisim1AAssThrow1: "P, a, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\i\ := e, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1AAssThrow2: "P, i, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\i\ := e, h \ (Throw ad, xs) \ (stk @ [v], loc, length (compE2 a) + pc, \ad\)" | bisim1AAssThrow3: "P, e, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\i\ := e, h \ (Throw ad, xs) \ (stk @ [v', v], loc, length (compE2 a) + length (compE2 i) + pc, \ad\)" | bisim1AAssFail: "P, a\i\ := e, h \ (Throw ad, xs) \ ([v', v, v''], xs, length (compE2 a) + length (compE2 i) + length (compE2 e), \ad\)" | bisim1AAss4: "P, a\i\ := e, h \ (unit, xs) \ ([], xs, Suc (length (compE2 a) + length (compE2 i) + length (compE2 e)), None)" | bisim1ALength: "P, a, h \ (a', xs) \ (stk, loc, pc, xcp) \ P, a\length, h \ (a'\length, xs) \ (stk, loc, pc, xcp)" | bisim1ALengthThrow: "P, a, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\length, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1ALengthNull: "P, a\length, h \ (THROW NullPointer, xs) \ ([Null], xs, length (compE2 a), \addr_of_sys_xcpt NullPointer\)" | bisim1FAcc: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e\F{D}, h \ (e'\F{D}, xs) \ (stk, loc, pc, xcp)" | bisim1FAccThrow: "P, e, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e\F{D}, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1FAccNull: "P, e\F{D}, h \ (THROW NullPointer, xs) \ ([Null], xs, length (compE2 e), \addr_of_sys_xcpt NullPointer\)" | bisim1FAss1: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e\F{D} := e2, h \ (e'\F{D} := e2, xs) \ (stk, loc, pc, xcp)" | bisim1FAss2: "P, e2, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e\F{D} := e2, h \ (Val v\F{D} := e', xs) \ (stk @ [v], loc, length (compE2 e) + pc, xcp)" | bisim1FAssThrow1: "P, e, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e\F{D} := e2, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1FAssThrow2: "P, e2, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e\F{D} := e2, h \ (Throw ad, xs) \ (stk @ [v], loc, length (compE2 e) + pc, \ad\)" | bisim1FAssNull: "P, e\F{D} := e2, h \ (THROW NullPointer, xs) \ ([v, Null], xs, length (compE2 e) + length (compE2 e2), \addr_of_sys_xcpt NullPointer\)" | bisim1FAss3: "P, e\F{D} := e2, h \ (unit, xs) \ ([], xs, Suc (length (compE2 e) + length (compE2 e2)), None)" | bisim1CAS1: "P, e1, h \ (e1', xs) \ (stk, loc, pc, xcp) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (e1'\compareAndSwap(D\F, e2, e3), xs) \ (stk, loc, pc, xcp)" | bisim1CAS2: "P, e2, h \ (e2', xs) \ (stk, loc, pc, xcp) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (Val v\compareAndSwap(D\F, e2', e3), xs) \ (stk @ [v], loc, length (compE2 e1) + pc, xcp)" | bisim1CAS3: "P, e3, h \ (e3', xs) \ (stk, loc, pc, xcp) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (Val v\compareAndSwap(D\F, Val v', e3'), xs) \ (stk @ [v', v], loc, length (compE2 e1) + length (compE2 e2) + pc, xcp)" | bisim1CASThrow1: "P, e1, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1CASThrow2: "P, e2, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (Throw ad, xs) \ (stk @ [v], loc, length (compE2 e1) + pc, \ad\)" | bisim1CASThrow3: "P, e3, h \ (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e1\compareAndSwap(D\F, e2, e3), h \ (Throw ad, xs) \ (stk @ [v', v], loc, length (compE2 e1) + length (compE2 e2) + pc, \ad\)" | bisim1CASFail: "P, e1\compareAndSwap(D\F, e2, e3), h \ (Throw ad, xs) \ ([v', v, v''], xs, length (compE2 e1) + length (compE2 e2) + length (compE2 e3), \ad\)" | bisim1Call1: "P, obj, h \ (obj', xs) \ (stk, loc, pc, xcp) \ P, obj\M(ps), h \ (obj'\M(ps), xs) \ (stk, loc, pc, xcp)" | bisim1CallParams: "P, ps, h \ (ps', xs) [\] (stk, loc, pc, xcp) \ P, obj\M(ps), h \ (Val v\M(ps'), xs) \ (stk @ [v], loc, length (compE2 obj) + pc, xcp)" | bisim1CallThrowObj: "P, obj, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, obj\M(ps), h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1CallThrowParams: "P, ps, h \ (map Val vs @ Throw a # ps', xs) [\] (stk, loc, pc, \a\) \ P, obj\M(ps), h \ (Throw a, xs) \ (stk @ [v], loc, length (compE2 obj) + pc, \a\)" | bisim1CallThrow: "length ps = length vs \ P, obj\M(ps), h \ (Throw a, xs) \ (vs @ [v], xs, length (compE2 obj) + length (compEs2 ps), \a\)" | bisim1BlockSome1: "P, {V:T=\v\; e}, h \ ({V:T=\v\; e}, xs) \ ([], xs, 0, None)" | bisim1BlockSome2: "P, {V:T=\v\; e}, h \ ({V:T=\v\; e}, xs) \ ([v], xs, Suc 0, None)" | bisim1BlockSome4: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, {V:T=\v\; e}, h \ ({V:T=None; e'}, xs) \ (stk, loc, Suc (Suc pc), xcp)" | bisim1BlockThrowSome: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, {V:T=\v\; e}, h \ (Throw a, xs) \ (stk, loc, Suc (Suc pc), \a\)" | bisim1BlockNone: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, {V:T=None; e}, h \ ({V:T=None; e'}, xs) \ (stk, loc, pc, xcp)" | bisim1BlockThrowNone: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, {V:T=None; e}, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Sync1: "P, e1, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, sync\<^bsub>V\<^esub> (e1) e2, h \ (sync\<^bsub>V\<^esub> (e') e2, xs) \ (stk, loc, pc, xcp)" | bisim1Sync2: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (sync\<^bsub>V\<^esub> (Val v) e2, xs) \ ([v, v], xs, Suc (length (compE2 e1)), None)" | bisim1Sync3: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (sync\<^bsub>V\<^esub> (Val v) e2, xs) \ ([v], xs[V := v], Suc (Suc (length (compE2 e1))), None)" | bisim1Sync4: "P, e2, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, sync\<^bsub>V\<^esub> (e1) e2, h \ (insync\<^bsub>V\<^esub> (a) e', xs) \ (stk, loc, Suc (Suc (Suc (length (compE2 e1) + pc))), xcp)" | bisim1Sync5: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (insync\<^bsub>V\<^esub> (a) Val v, xs) \ ([xs ! V, v], xs, 4 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync6: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Val v, xs) \ ([v], xs, 5 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync7: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (insync\<^bsub>V\<^esub> (a) Throw a', xs) \ ([Addr a'], xs, 6 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync8: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (insync\<^bsub>V\<^esub> (a) Throw a', xs) \ ([xs ! V, Addr a'], xs, 7 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync9: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Throw a, xs) \ ([Addr a], xs, 8 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync10: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Throw a, xs) \ ([Addr a], xs, 8 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1Sync11: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (THROW NullPointer, xs) \ ([Null], xs, Suc (Suc (length (compE2 e1))), \addr_of_sys_xcpt NullPointer\)" | bisim1Sync12: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Throw a, xs) \ ([v, v'], xs, 4 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1Sync14: "P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Throw a, xs) \ ([v, Addr a'], xs, 7 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1SyncThrow: "P, e1, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, sync\<^bsub>V\<^esub> (e1) e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1InSync: \ \This rule only exists such that @{text "P,e,h \ (e, xs) \ ([], xs, 0, None)"} holds for all @{text "e"}\ "P, insync\<^bsub>V\<^esub> (a) e, h \ (insync\<^bsub>V\<^esub> (a) e, xs) \ ([], xs, 0, None)" | bisim1Seq1: "P, e1, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e1;;e2, h \ (e';;e2, xs) \ (stk, loc, pc, xcp)" | bisim1SeqThrow1: "P, e1, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, e1;;e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Seq2: "P, e2, h \ exs \ (stk, loc, pc, xcp) \ P, e1;;e2, h \ exs \ (stk, loc, Suc (length (compE2 e1) + pc), xcp)" | bisim1Cond1: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, if (e) e1 else e2, h \ (if (e') e1 else e2, xs) \ (stk, loc, pc, xcp)" | bisim1CondThen: "P, e1, h \ exs \ (stk, loc, pc, xcp) \ P, if (e) e1 else e2, h \ exs \ (stk, loc, Suc (length (compE2 e) + pc), xcp)" | bisim1CondElse: "P, e2, h \ exs \ (stk, loc, pc, xcp) \ P, if (e) e1 else e2, h \ exs \ (stk, loc, Suc (Suc (length (compE2 e) + length (compE2 e1) + pc)), xcp)" | bisim1CondThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, if (e) e1 else e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1While1: "P, while (c) e, h \ (while (c) e, xs) \ ([], xs, 0, None)" | bisim1While3: "P, c, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, while (c) e, h \ (if (e') (e;; while (c) e) else unit, xs) \ (stk, loc, pc, xcp)" | bisim1While4: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, while (c) e, h \ (e';; while (c) e, xs) \ (stk, loc, Suc (length (compE2 c) + pc), xcp)" | bisim1While6: "P, while (c) e, h \ (while (c) e, xs) \ ([], xs, Suc (Suc (length (compE2 c) + length (compE2 e))), None)" | bisim1While7: "P, while (c) e, h \ (unit, xs) \ ([], xs, Suc (Suc (Suc (length (compE2 c) + length (compE2 e)))), None)" | bisim1WhileThrow1: "P, c, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, while (c) e, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1WhileThrow2: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, while (c) e, h \ (Throw a, xs) \ (stk, loc, Suc (length (compE2 c) + pc), \a\)" | bisim1Throw1: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, throw e, h \ (throw e', xs) \ (stk, loc, pc, xcp)" | bisim1Throw2: "P, throw e, h \ (Throw a, xs) \ ([Addr a], xs, length (compE2 e), \a\)" | bisim1ThrowNull: "P, throw e, h \ (THROW NullPointer, xs) \ ([Null], xs, length (compE2 e), \addr_of_sys_xcpt NullPointer\)" | bisim1ThrowThrow: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, throw e, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Try: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, try e catch(C V) e2, h \ (try e' catch(C V) e2, xs) \ (stk, loc, pc, xcp)" | bisim1TryCatch1: "\ P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\); typeof_addr h a = \Class_type C'\; P \ C' \\<^sup>* C \ \ P, try e catch(C V) e2, h \ ({V:Class C=None; e2}, xs[V := Addr a]) \ ([Addr a], loc, Suc (length (compE2 e)), None)" | bisim1TryCatch2: "P, e2, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, try e catch(C V) e2, h \ ({V:Class C=None; e'}, xs) \ (stk, loc, Suc (Suc (length (compE2 e) + pc)), xcp)" | bisim1TryFail: "\ P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\); typeof_addr h a = \Class_type C'\; \ P \ C' \\<^sup>* C \ \ P, try e catch(C V) e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1TryCatchThrow: "P, e2, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ P, try e catch(C V) e2, h \ (Throw a, xs) \ (stk, loc, Suc (Suc (length (compE2 e) + pc)), \a\)" | bisims1Nil: "P, [], h \ ([], xs) [\] ([], xs, 0, None)" | bisims1List1: "P, e, h \ (e', xs) \ (stk, loc, pc, xcp) \ P, e#es, h \ (e'#es, xs) [\] (stk, loc, pc, xcp)" | bisims1List2: "P, es, h \ (es', xs) [\] (stk, loc, pc, xcp) \ P, e#es, h \ (Val v # es', xs) [\] (stk @ [v], loc, length (compE2 e) + pc, xcp)" inductive_cases bisim1_cases: "P,e,h \ (Val v, xs) \ (stk, loc, pc, xcp)" lemma bisim1_refl: "P,e,h \ (e, xs) \ ([], xs, 0, None)" and bisims1_refl: "P,es,h \ (es, xs) [\] ([], xs, 0, None)" apply(induct e and es rule: call.induct calls.induct) apply(auto intro: bisim1_bisims1.intros simp add: nat_fun_sum_eq_conv) apply(rename_tac option a) apply(case_tac option) apply(auto intro: bisim1_bisims1.intros split: if_split_asm) done lemma bisims1_lengthD: "P, es, h \ (es', xs) [\] s \ length es = length es'" apply(induct es arbitrary: es' s) apply(auto elim: bisims1.cases) done text \ Derive an alternative induction rule for @{term bisim1} such that (i) induction hypothesis are generated for all subexpressions and (ii) the number of surrounding blocks is passed through. \ inductive bisim1' :: "'m prog \ 'heap \ 'addr expr1 \ nat \ ('addr expr1 \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" and bisims1' :: "'m prog \ 'heap \ 'addr expr1 list \ nat \ ('addr expr1 list \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" and bisim1'_syntax :: "'m prog \ 'addr expr1 \ nat \ 'heap \ ('addr expr1 \ 'addr locals1) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" - ("_,_,_,_ \' _ \ _" [50, 0, 0, 0, 0, 50] 100) + ("_,_,_,_ \'' _ \ _" [50, 0, 0, 0, 0, 50] 100) and bisims1'_syntax :: "'m prog \ 'addr expr1 list \ nat \ 'heap \ ('addr expr1 list \ 'addr val list) \ ('addr val list \ 'addr val list \ pc \ 'addr option) \ bool" - ("_,_,_,_ \' _ [\] _" [50, 0, 0, 0, 0, 50] 100) + ("_,_,_,_ \'' _ [\] _" [50, 0, 0, 0, 0, 50] 100) for P :: "'m prog" and h :: 'heap where "P, e, n, h \' exs \ s \ bisim1' P h e n exs s" | "P, es, n, h \' esxs [\] s \ bisims1' P h es n esxs s" | bisim1Val2': "P, e, n, h \' (Val v, xs) \ (v # [], xs, length (compE2 e), None)" | bisim1New': "P, new C, n, h \' (new C, xs) \ ([], xs, 0, None)" | bisim1NewThrow': "P, new C, n, h \' (THROW OutOfMemory, xs) \ ([], xs, 0, \addr_of_sys_xcpt OutOfMemory\)" | bisim1NewArray': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, newA T\e\, n, h \' (newA T\e'\, xs) \ (stk, loc, pc, xcp)" | bisim1NewArrayThrow': "P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, newA T\e\, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1NewArrayFail': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, newA T\e\, n, h \' (Throw a, xs) \ ([v], xs, length (compE2 e), \a\)" | bisim1Cast': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, Cast T e, n, h \' (Cast T e', xs) \ (stk, loc, pc, xcp)" | bisim1CastThrow': "P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, Cast T e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1CastFail': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, Cast T e, n, h \' (THROW ClassCast, xs) \ ([v], xs, length (compE2 e), \addr_of_sys_xcpt ClassCast\)" | bisim1InstanceOf': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, e instanceof T, n, h \' (e' instanceof T, xs) \ (stk, loc, pc, xcp)" | bisim1InstanceOfThrow': "P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, e instanceof T, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Val': "P, Val v, n, h \' (Val v, xs) \ ([], xs, 0, None)" | bisim1Var': "P, Var V, n, h \' (Var V, xs) \ ([], xs, 0, None)" | bisim1BinOp1': "\ P, e1, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1\bop\e2, n, h \' (e'\bop\e2, xs) \ (stk, loc, pc, xcp)" | bisim1BinOp2': "\ P, e2, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None) \ \ P, e1\bop\e2, n, h \' (Val v1 \bop\ e', xs) \ (stk @ [v1], loc, length (compE2 e1) + pc, xcp)" | bisim1BinOpThrow1': "\ P, e1, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1\bop\e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1BinOpThrow2': "\ P, e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None) \ \ P, e1\bop\e2, n, h \' (Throw a, xs) \ (stk @ [v1], loc, length (compE2 e1) + pc, \a\)" | bisim1BinOpThrow': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1\bop\e2, n, h \' (Throw a, xs) \ ([v1, v2], xs, length (compE2 e1) + length (compE2 e2), \a\)" | bisim1LAss1': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, V:=e, n, h \' (V:=e', xs) \ (stk, loc, pc, xcp)" | bisim1LAss2': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, V:=e, n, h \' (unit, xs) \ ([], xs, Suc (length (compE2 e)), None)" | bisim1LAssThrow': "P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, V:=e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1AAcc1': "\ P, a, n, h \' (a', xs) \ (stk, loc, pc, xcp); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None) \ \ P, a\i\, n, h \' (a'\i\, xs) \ (stk, loc, pc, xcp)" | bisim1AAcc2': "\ P, i, n, h \' (i', xs) \ (stk, loc, pc, xcp); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None) \ \ P, a\i\, n, h \' (Val v\i'\, xs) \ (stk @ [v], loc, length (compE2 a) + pc, xcp)" | bisim1AAccThrow1': "\ P, a, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None) \ \ P, a\i\, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1AAccThrow2': "\ P, i, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None) \ \ P, a\i\, n, h \' (Throw ad, xs) \ (stk @ [v], loc, length (compE2 a) + pc, \ad\)" | bisim1AAccFail': "\ \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None) \ \ P, a\i\, n, h \' (Throw ad, xs) \ ([v, v'], xs, length (compE2 a) + length (compE2 i), \ad\)" | bisim1AAss1': "\ P, a, n, h \' (a', xs) \ (stk, loc, pc, xcp); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (a'\i\ := e, xs) \ (stk, loc, pc, xcp)" | bisim1AAss2': "\ P, i, n, h \' (i', xs) \ (stk, loc, pc, xcp); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Val v\i'\ := e, xs) \ (stk @ [v], loc, length (compE2 a) + pc, xcp)" | bisim1AAss3': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Val v\Val v'\ := e', xs) \ (stk @ [v', v], loc, length (compE2 a) + length (compE2 i) + pc, xcp)" | bisim1AAssThrow1': "\ P, a, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1AAssThrow2': "\ P, i, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Throw ad, xs) \ (stk @ [v], loc, length (compE2 a) + pc, \ad\)" | bisim1AAssThrow3': "\ P, e, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Throw ad, xs) \ (stk @ [v', v], loc, length (compE2 a) + length (compE2 i) + pc, \ad\)" | bisim1AAssFail': "\ \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (Throw ad, xs) \ ([v', v, v''], xs, length (compE2 a) + length (compE2 i) + length (compE2 e), \ad\)" | bisim1AAss4': "\ \xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None); \xs. P, i, n, h \' (i, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, a\i\ := e, n, h \' (unit, xs) \ ([], xs, Suc (length (compE2 a) + length (compE2 i) + length (compE2 e)), None)" | bisim1ALength': "P, a, n, h \' (a', xs) \ (stk, loc, pc, xcp) \ P, a\length, n, h \' (a'\length, xs) \ (stk, loc, pc, xcp)" | bisim1ALengthThrow': "P, a, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, a\length, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1ALengthNull': "(\xs. P, a, n, h \' (a, xs) \ ([], xs, 0, None)) \ P, a\length, n, h \' (THROW NullPointer, xs) \ ([Null], xs, length (compE2 a), \addr_of_sys_xcpt NullPointer\)" | bisim1FAcc': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, e\F{D}, n, h \' (e'\F{D}, xs) \ (stk, loc, pc, xcp)" | bisim1FAccThrow': "P, e, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\) \ P, e\F{D}, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1FAccNull': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, e\F{D}, n, h \' (THROW NullPointer, xs) \ ([Null], xs, length (compE2 e), \addr_of_sys_xcpt NullPointer\)" | bisim1FAss1': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (e'\F{D} := e2, xs) \ (stk, loc, pc, xcp)" | bisim1FAss2': "\ P, e2, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (Val v\F{D} := e', xs) \ (stk @ [v], loc, length (compE2 e) + pc, xcp)" | bisim1FAssThrow1': "\ P, e, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1FAssThrow2': "\ P, e2, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (Throw ad, xs) \ (stk @ [v], loc, length (compE2 e) + pc, \ad\)" | bisim1FAssNull': "\ \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (THROW NullPointer, xs) \ ([v, Null], xs, length (compE2 e) + length (compE2 e2), \addr_of_sys_xcpt NullPointer\)" | bisim1FAss3': "\ \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e\F{D} := e2, n, h \' (unit, xs) \ ([], xs, Suc (length (compE2 e) + length (compE2 e2)), None)" | bisim1CAS1': "\ P, e1, n, h \' (e1', xs) \ (stk, loc, pc, xcp); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None); \xs. P, e3, n, h \' (e3, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (e1'\compareAndSwap(D\F, e2, e3), xs) \ (stk, loc, pc, xcp)" | bisim1CAS2': "\ P, e2, n, h \' (e2', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e3, n, h \' (e3, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Val v\compareAndSwap(D\F, e2', e3), xs) \ (stk @ [v], loc, length (compE2 e1) + pc, xcp)" | bisim1CAS3': "\ P, e3, n, h \' (e3', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Val v\compareAndSwap(D\F, Val v', e3'), xs) \ (stk @ [v', v], loc, length (compE2 e1) + length (compE2 e2) + pc, xcp)" | bisim1CASThrow1': "\ P, e1, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None); \xs. P, e3, n, h \' (e3, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\)" | bisim1CASThrow2': "\ P, e2, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e3, n, h \' (e3, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Throw ad, xs) \ (stk @ [v], loc, length (compE2 e1) + pc, \ad\)" | bisim1CASThrow3': "\ P, e3, n, h \' (Throw ad, xs) \ (stk, loc, pc, \ad\); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Throw ad, xs) \ (stk @ [v', v], loc, length (compE2 e1) + length (compE2 e2) + pc, \ad\)" | bisim1CASFail': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None); \xs. P, e3, n, h \' (e3, xs) \ ([], xs, 0, None) \ \ P, e1\compareAndSwap(D\F, e2, e3), n, h \' (Throw ad, xs) \ ([v', v, v''], xs, length (compE2 e1) + length (compE2 e2) + length (compE2 e3), \ad\)" | bisim1Call1': "\ P, obj, n, h \' (obj', xs) \ (stk, loc, pc, xcp); \xs. P, ps, n, h \' (ps, xs) [\] ([], xs, 0, None) \ \ P, obj\M(ps), n, h \' (obj'\M(ps), xs) \ (stk, loc, pc, xcp)" | bisim1CallParams': "\ P, ps, n, h \' (ps', xs) [\] (stk, loc, pc, xcp); ps \ []; \xs. P, obj, n, h \' (obj, xs) \ ([], xs, 0, None) \ \ P, obj\M(ps), n, h \' (Val v\M(ps'), xs) \ (stk @ [v], loc, length (compE2 obj) + pc, xcp)" | bisim1CallThrowObj': "\ P, obj, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, ps, n, h \' (ps, xs) [\] ([], xs, 0, None)\ \ P, obj\M(ps), n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1CallThrowParams': "\ P, ps, n, h \' (map Val vs @ Throw a # ps', xs) [\] (stk, loc, pc, \a\); \xs. P, obj, n, h \' (obj, xs) \ ([], xs, 0, None) \ \ P, obj\M(ps), n, h \' (Throw a, xs) \ (stk @ [v], loc, length (compE2 obj) + pc, \a\)" | bisim1CallThrow': "\ length ps = length vs; \xs. P, obj, n, h \' (obj, xs) \ ([], xs, 0, None); \xs. P, ps, n, h \' (ps, xs) [\] ([], xs, 0, None) \ \ P, obj\M(ps), n, h \' (Throw a, xs) \ (vs @ [v], xs, length (compE2 obj) + length (compEs2 ps), \a\)" | bisim1BlockSome1': "(\xs. P, e, Suc n, h \' (e, xs) \ ([], xs, 0, None)) \ P, {V:T=\v\; e}, n, h \' ({V:T=\v\; e}, xs) \ ([], xs, 0, None)" | bisim1BlockSome2': "(\xs. P, e, Suc n, h \' (e, xs) \ ([], xs, 0, None)) \ P, {V:T=\v\; e}, n, h \' ({V:T=\v\; e}, xs) \ ([v], xs, Suc 0, None)" | bisim1BlockSome4': "P, e, Suc n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, {V:T=\v\; e}, n, h \' ({V:T=None; e'}, xs) \ (stk, loc, Suc (Suc pc), xcp)" | bisim1BlockThrowSome': "P, e, Suc n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, {V:T=\v\; e}, n, h \' (Throw a, xs) \ (stk, loc, Suc (Suc pc), \a\)" | bisim1BlockNone': "P, e, Suc n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, {V:T=None; e}, n, h \' ({V:T=None; e'}, xs) \ (stk, loc, pc, xcp)" | bisim1BlockThrowNone': "P, e, Suc n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, {V:T=None; e}, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Sync1': "\ P, e1, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (sync\<^bsub>V\<^esub> (e') e2, xs) \ (stk, loc, pc, xcp)" | bisim1Sync2': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (sync\<^bsub>V\<^esub> (Val v) e2, xs) \ ([v, v], xs, Suc (length (compE2 e1)), None)" | bisim1Sync3': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (sync\<^bsub>V\<^esub> (Val v) e2, xs) \ ([v], xs[V := v], Suc (Suc (length (compE2 e1))), None)" | bisim1Sync4': "\ P, e2, Suc n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (insync\<^bsub>V\<^esub> (a) e', xs) \ (stk, loc, Suc (Suc (Suc (length (compE2 e1) + pc))), xcp)" | bisim1Sync5': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (insync\<^bsub>V\<^esub> (a) Val v, xs) \ ([xs ! V, v], xs, 4 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync6': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Val v, xs) \ ([v], xs, 5 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync7': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (insync\<^bsub>V\<^esub> (a) Throw a', xs) \ ([Addr a'], xs, 6 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync8': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (insync\<^bsub>V\<^esub> (a) Throw a', xs) \ ([xs ! V, Addr a'], xs, 7 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync9': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Throw a, xs) \ ([Addr a], xs, 8 + length (compE2 e1) + length (compE2 e2), None)" | bisim1Sync10': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Throw a, xs) \ ([Addr a], xs, 8 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1Sync11': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (THROW NullPointer, xs) \ ([Null], xs, Suc (Suc (length (compE2 e1))), \addr_of_sys_xcpt NullPointer\)" | bisim1Sync12': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Throw a, xs) \ ([v, v'], xs, 4 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1Sync14': "\ \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Throw a, xs) \ ([v, Addr a'], xs, 7 + length (compE2 e1) + length (compE2 e2), \a\)" | bisim1SyncThrow': "\ P, e1, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, sync\<^bsub>V\<^esub> (e1) e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1InSync': "P, insync\<^bsub>V\<^esub> (a) e, n, h \' (insync\<^bsub>V\<^esub> (a) e, xs) \ ([], xs, 0, None)" | bisim1Seq1': "\ P, e1, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1;;e2, n, h \' (e';;e2, xs) \ (stk, loc, pc, xcp)" | bisim1SeqThrow1': "\ P, e1, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, e1;;e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Seq2': "\ P, e2, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None) \ \ P, e1;;e2, n, h \' (e', xs) \ (stk, loc, Suc (length (compE2 e1) + pc), xcp)" | bisim1Cond1': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, if (e) e1 else e2, n, h \' (if (e') e1 else e2, xs) \ (stk, loc, pc, xcp)" | bisim1CondThen': "\ P, e1, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, if (e) e1 else e2, n, h \' (e', xs) \ (stk, loc, Suc (length (compE2 e) + pc), xcp)" | bisim1CondElse': "\ P, e2, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None) \ \ P, if (e) e1 else e2, n, h \' (e', xs) \ (stk, loc, Suc (Suc (length (compE2 e) + length (compE2 e1) + pc)), xcp)" | bisim1CondThrow': "\ P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e1, n, h \' (e1, xs) \ ([], xs, 0, None); \xs. P, e2, n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, if (e) e1 else e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1While1': "\ \xs. P, c, n, h \' (c, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (while (c) e, xs) \ ([], xs, 0, None)" | bisim1While3': "\ P, c, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (if (e') (e;; while (c) e) else unit, xs) \ (stk, loc, pc, xcp)" | bisim1While4': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, c, n, h \' (c, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (e';; while (c) e, xs) \ (stk, loc, Suc (length (compE2 c) + pc), xcp)" | bisim1While6': "\ \xs. P, c, n, h \' (c, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (while (c) e, xs) \ ([], xs, Suc (Suc (length (compE2 c) + length (compE2 e))), None)" | bisim1While7': "\ \xs. P, c, n, h \' (c, xs) \ ([], xs, 0, None); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (unit, xs) \ ([], xs, Suc (Suc (Suc (length (compE2 c) + length (compE2 e)))), None)" | bisim1WhileThrow1': "\ P, c, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1WhileThrow2': "\ P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, c, n, h \' (c, xs) \ ([], xs, 0, None) \ \ P, while (c) e, n, h \' (Throw a, xs) \ (stk, loc, Suc (length (compE2 c) + pc), \a\)" | bisim1Throw1': "P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp) \ P, throw e, n, h \' (throw e', xs) \ (stk, loc, pc, xcp)" | bisim1Throw2': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, throw e, n, h \' (Throw a, xs) \ ([Addr a], xs, length (compE2 e), \a\)" | bisim1ThrowNull': "(\xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None)) \ P, throw e, n, h \' (THROW NullPointer, xs) \ ([Null], xs, length (compE2 e), \addr_of_sys_xcpt NullPointer\)" | bisim1ThrowThrow': "P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\) \ P, throw e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1Try': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, try e catch(C V) e2, n, h \' (try e' catch(C V) e2, xs) \ (stk, loc, pc, xcp)" | bisim1TryCatch1': "\ P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); typeof_addr h a = \Class_type C'\; P \ C' \\<^sup>* C; \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, try e catch(C V) e2, n, h \' ({V:Class C=None; e2}, xs[V := Addr a]) \ ([Addr a], loc, Suc (length (compE2 e)), None)" | bisim1TryCatch2': "\ P, e2, Suc n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, try e catch(C V) e2, n, h \' ({V:Class C=None; e'}, xs) \ (stk, loc, Suc (Suc (length (compE2 e) + pc)), xcp)" | bisim1TryFail': "\ P, e, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); typeof_addr h a = \Class_type C'\; \ P \ C' \\<^sup>* C; \xs. P, e2, Suc n, h \' (e2, xs) \ ([], xs, 0, None) \ \ P, try e catch(C V) e2, n, h \' (Throw a, xs) \ (stk, loc, pc, \a\)" | bisim1TryCatchThrow': "\ P, e2, Suc n, h \' (Throw a, xs) \ (stk, loc, pc, \a\); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, try e catch(C V) e2, n, h \' (Throw a, xs) \ (stk, loc, Suc (Suc (length (compE2 e) + pc)), \a\)" | bisims1Nil': "P, [], n, h \' ([], xs) [\] ([], xs, 0, None)" | bisims1List1': "\ P, e, n, h \' (e', xs) \ (stk, loc, pc, xcp); \xs. P, es, n, h \' (es, xs) [\] ([], xs, 0, None) \ \ P, e#es, n, h \' (e'#es, xs) [\] (stk, loc, pc, xcp)" | bisims1List2': "\ P, es, n, h \' (es', xs) [\] (stk, loc, pc, xcp); \xs. P, e, n, h \' (e, xs) \ ([], xs, 0, None) \ \ P, e#es, n, h \' (Val v # es', xs) [\] (stk @ [v], loc, length (compE2 e) + pc, xcp)" lemma bisim1'_refl: "P,e,n,h \' (e,xs) \ ([],xs,0,None)" and bisims1'_refl: "P,es,n,h \' (es,xs) [\] ([],xs,0,None)" apply(induct e and es arbitrary: n xs and n xs rule: call.induct calls.induct) apply(auto intro: bisim1'_bisims1'.intros simp add: nat_fun_sum_eq_conv) apply(rename_tac option a b c) apply(case_tac option) apply(auto intro: bisim1'_bisims1'.intros simp add: fun_eq_iff split: if_split_asm) done lemma bisim1_imp_bisim1': "P, e, h \ exs \ s \ P, e, n, h \' exs \ s" and bisims1_imp_bisims1': "P, es, h \ esxs [\] s \ P, es, n, h \' esxs [\] s" proof(induct arbitrary: n and n rule: bisim1_bisims1.inducts) case (bisim1CallParams ps ps' xs stk loc pc xcp obj M v) show ?case proof(cases "ps = []") case True with \P,ps,h \ (ps', xs) [\] (stk, loc, pc, xcp)\ have "ps' = []" "pc = 0" "stk = []" "loc = xs" "xcp = None" by(auto elim: bisims1.cases) moreover have "P,obj,n,h \' (Val v,xs) \ ([v],xs,length (compE2 obj),None)" by(blast intro: bisim1Val2' bisim1'_refl) hence "P,obj\M([]),n,h \' (Val v\M([]),xs) \ ([v],xs,length (compE2 obj),None)" by-(rule bisim1Call1', auto intro!: bisims1Nil' simp add: bsoks_def) ultimately show ?thesis using True by simp next case False with bisim1CallParams show ?thesis by(auto intro: bisim1CallParams' bisims1'_refl bisim1'_refl) qed qed(auto intro: bisim1'_bisims1'.intros bisim1'_refl bisims1'_refl) lemma bisim1'_imp_bisim1: "P, e, n, h \' exs \ s \ P, e, h \ exs \ s" and bisims1'_imp_bisims1: "P, es, n, h \' esxs [\] s \ P, es, h \ esxs [\] s" apply(induct rule: bisim1'_bisims1'.inducts) apply(blast intro: bisim1_bisims1.intros)+ done lemma bisim1'_eq_bisim1: "bisim1' P h e n = bisim1 P h e" and bisims1'_eq_bisims1: "bisims1' P h es n = bisims1 P h es" by(blast intro!: ext bisim1_imp_bisim1' bisims1_imp_bisims1' bisim1'_imp_bisim1 bisims1'_imp_bisims1)+ end (* FIXME: Take lemmas out of locale to speed up opening the context *) lemmas bisim1_bisims1_inducts = J1_JVM_heap_base.bisim1'_bisims1'.inducts [simplified J1_JVM_heap_base.bisim1'_eq_bisim1 J1_JVM_heap_base.bisims1'_eq_bisims1, consumes 1, case_names bisim1Val2 bisim1New bisim1NewThrow bisim1NewArray bisim1NewArrayThrow bisim1NewArrayFail bisim1Cast bisim1CastThrow bisim1CastFail bisim1InstanceOf bisim1InstanceOfThrow bisim1Val bisim1Var bisim1BinOp1 bisim1BinOp2 bisim1BinOpThrow1 bisim1BinOpThrow2 bisim1BinOpThrow bisim1LAss1 bisim1LAss2 bisim1LAssThrow bisim1AAcc1 bisim1AAcc2 bisim1AAccThrow1 bisim1AAccThrow2 bisim1AAccFail bisim1AAss1 bisim1AAss2 bisim1AAss3 bisim1AAssThrow1 bisim1AAssThrow2 bisim1AAssThrow3 bisim1AAssFail bisim1AAss4 bisim1ALength bisim1ALengthThrow bisim1ALengthNull bisim1FAcc bisim1FAccThrow bisim1FAccNull bisim1FAss1 bisim1FAss2 bisim1FAssThrow1 bisim1FAssThrow2 bisim1FAssNull bisim1FAss3 bisim1CAS1 bisim1CAS2 bisim1CAS3 bisim1CASThrow1 bisim1CASThrow2 bisim1CASThrow3 bisim1CASFail bisim1Call1 bisim1CallParams bisim1CallThrowObj bisim1CallThrowParams bisim1CallThrow bisim1BlockSome1 bisim1BlockSome2 bisim1BlockSome4 bisim1BlockThrowSome bisim1BlockNone bisim1BlockThrowNone bisim1Sync1 bisim1Sync2 bisim1Sync3 bisim1Sync4 bisim1Sync5 bisim1Sync6 bisim1Sync7 bisim1Sync8 bisim1Sync9 bisim1Sync10 bisim1Sync11 bisim1Sync12 bisim1Sync14 bisim1SyncThrow bisim1InSync bisim1Seq1 bisim1SeqThrow1 bisim1Seq2 bisim1Cond1 bisim1CondThen bisim1CondElse bisim1CondThrow bisim1While1 bisim1While3 bisim1While4 bisim1While6 bisim1While7 bisim1WhileThrow1 bisim1WhileThrow2 bisim1Throw1 bisim1Throw2 bisim1ThrowNull bisim1ThrowThrow bisim1Try bisim1TryCatch1 bisim1TryCatch2 bisim1TryFail bisim1TryCatchThrow bisims1Nil bisims1List1 bisims1List2] lemmas bisim1_bisims1_inducts_split = bisim1_bisims1_inducts[split_format (complete)] context J1_JVM_heap_base begin lemma bisim1_pc_length_compE2: "P,E,h \ (e, xs) \ (stk, loc, pc, xcp) \ pc \ length (compE2 E)" and bisims1_pc_length_compEs2: "P,Es,h \ (es, xs) [\] (stk, loc, pc, xcp) \ pc \ length (compEs2 Es)" apply(induct "(stk, loc, pc, xcp)" and "(stk, loc, pc, xcp)" arbitrary: stk loc pc xcp and stk loc pc xcp rule: bisim1_bisims1.inducts) apply(auto) done lemma bisim1_pc_length_compE2D: "P,e,h \ (e', xs) \ (stk,loc,length (compE2 e),xcp) \ xcp = None \ call1 e' = None \ (\v. stk = [v] \ (is_val e' \ e' = Val v \ xs = loc))" and bisims1_pc_length_compEs2D: "P,es,h \ (es', xs) [\] (stk,loc,length (compEs2 es),xcp) \ xcp = None \ calls1 es' = None \ (\vs. stk = rev vs \ length vs = length es \ (is_vals es' \ es' = map Val vs \ xs = loc))" proof(induct "(e', xs)" "(stk, loc, length (compE2 e), xcp)" and "(es', xs)" "(stk, loc, length (compEs2 es), xcp)" arbitrary: e' xs stk loc xcp and es' xs stk loc xcp rule: bisim1_bisims1.inducts) case (bisims1List2 es es' xs stk loc pc xcp e v) then obtain vs where "xcp = None" "calls1 es' = None" "stk = rev vs" "length vs = length es" "is_vals es' \ es' = map Val vs \ xs = loc" by auto thus ?case by(clarsimp)(rule_tac x="v#vs" in exI, auto) qed(simp_all (no_asm_use), (fastforce dest: bisim1_pc_length_compE2 bisims1_pc_length_compEs2 split: bop.split_asm if_split_asm)+) corollary bisim1_call_pcD: "\ P,e,h \ (e', xs) \ (stk, loc, pc, xcp); call1 e' = \aMvs\ \ \ pc < length (compE2 e)" and bisims1_calls_pcD: "\ P,es,h \ (es', xs) [\] (stk, loc, pc, xcp); calls1 es' = \aMvs\ \ \ pc < length (compEs2 es)" proof - assume bisim: "P,e,h \ (e', xs) \ (stk, loc, pc, xcp)" and call: "call1 e' = \aMvs\" { assume "pc = length (compE2 e)" with bisim call have False by(auto dest: bisim1_pc_length_compE2D) } moreover from bisim have "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) ultimately show "pc < length (compE2 e)" by(cases "pc < length (compE2 e)")(auto) next assume bisim: "P,es,h \ (es', xs) [\] (stk, loc, pc, xcp)" and call: "calls1 es' = \aMvs\" { assume "pc = length (compEs2 es)" with bisim call have False by(auto dest: bisims1_pc_length_compEs2D) } moreover from bisim have "pc \ length (compEs2 es)" by(rule bisims1_pc_length_compEs2) ultimately show "pc < length (compEs2 es)" by(cases "pc < length (compEs2 es)")(auto) qed lemma bisim1_length_xs: "P,e,h \ (e',xs) \ (stk, loc, pc, xcp) \ length xs = length loc" and bisims1_length_xs: "P,es,h \ (es',xs) [\] (stk, loc, pc, xcp) \ length xs = length loc" by(induct "(e',xs)" "(stk, loc, pc, xcp)" and "(es',xs)" "(stk, loc, pc, xcp)" arbitrary: e' xs stk loc pc xcp and es' xs stk loc pc xcp rule: bisim1_bisims1.inducts) auto lemma bisim1_Val_length_compE2D: "P,e,h \ (Val v,xs) \ (stk, loc, length (compE2 e), xcp) \ stk = [v] \ xs = loc \ xcp = None" and bisims1_Val_length_compEs2D: "P,es,h \ (map Val vs,xs) [\] (stk, loc, length (compEs2 es), xcp) \ stk = rev vs \ xs = loc \ xcp = None" by(auto dest: bisim1_pc_length_compE2D bisims1_pc_length_compEs2D) lemma bisim_Val_loc_eq_xcp_None: "P, e, h \ (Val v, xs) \ (stk, loc, pc, xcp) \ xs = loc \ xcp = None" and bisims_Val_loc_eq_xcp_None: "P, es, h \ (map Val vs, xs) [\] (stk, loc, pc, xcp) \ xs = loc \ xcp = None" apply(induct "(Val v :: 'addr expr1, xs)" "(stk, loc, pc, xcp)" and "(map Val vs :: 'addr expr1 list, xs)" "(stk, loc, pc, xcp)" arbitrary: v xs stk loc pc xcp and vs xs stk loc pc xcp rule: bisim1_bisims1.inducts) apply(auto) done lemma bisim_Val_pc_not_Invoke: "\ P,e,h \ (Val v,xs) \ (stk,loc,pc,xcp); pc < length (compE2 e) \ \ compE2 e ! pc \ Invoke M n'" and bisims_Val_pc_not_Invoke: "\ P,es,h \ (map Val vs,xs) [\] (stk,loc,pc,xcp); pc < length (compEs2 es) \ \ compEs2 es ! pc \ Invoke M n'" apply(induct "(Val v :: 'addr expr1, xs)" "(stk, loc, pc, xcp)" and "(map Val vs :: 'addr expr1 list, xs)" "(stk, loc, pc, xcp)" arbitrary: v xs stk loc pc xcp and vs xs stk loc pc xcp rule: bisim1_bisims1.inducts) apply(auto simp add: nth_append compEs2_map_Val dest: bisim1_pc_length_compE2) done lemma bisim1_VarD: "P, E, h \ (Var V,xs) \ (stk,loc,pc,xcp) \ xs = loc" and "P, es, h \ (es', xs) [\] (stk, loc, pc, xcp) \ True" by(induct "(Var V :: 'addr expr1, xs)" "(stk, loc, pc, xcp)" and arbitrary: V xs stk loc pc xcp and rule: bisim1_bisims1.inducts) auto lemma bisim1_ThrowD: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, xcp) \ pc < length (compE2 e) \ (xcp = \a\ \ xcp = None) \ xs = loc" and bisims1_ThrowD: "P, es, h \ (map Val vs @ Throw a # es', xs) [\] (stk, loc, pc, xcp) \ pc < length (compEs2 es) \ (xcp = \a\ \ xcp = None) \ xs = loc" apply(induct "(Throw a :: 'addr expr1, xs)" "(stk, loc, pc, xcp)" and "(map Val vs @ Throw a # es', xs)" "(stk, loc, pc, xcp)" arbitrary: xs stk loc pc xcp and vs es' xs stk loc pc xcp rule: bisim1_bisims1.inducts) apply(fastforce dest: bisim1_pc_length_compE2 bisim_Val_loc_eq_xcp_None simp add: Cons_eq_append_conv)+ done lemma fixes P :: "'addr J1_prog" shows bisim1_Invoke_stkD: "\ P,e,h \ exs \ (stk,loc,pc,None); pc < length (compE2 e); compE2 e ! pc = Invoke M n' \ \ \vs v stk'. stk = vs @ v # stk' \ length vs = n'" and bisims1_Invoke_stkD: "\ P,es,h \ esxs [\] (stk,loc,pc,None); pc < length (compEs2 es); compEs2 es ! pc = Invoke M n' \ \ \vs v stk'. stk = vs @ v # stk' \ length vs = n'" proof(induct "(stk, loc, pc, None :: 'addr option)" and "(stk, loc, pc, None :: 'addr option)" arbitrary: stk loc pc and stk loc pc rule: bisim1_bisims1.inducts) case bisim1Call1 thus ?case apply(clarsimp simp add: nth_append append_eq_append_conv2 neq_Nil_conv split: if_split_asm) apply(drule bisim1_pc_length_compE2, clarsimp simp add: neq_Nil_conv nth_append) apply(frule bisim1_pc_length_compE2, clarsimp) apply(drule bisim1_pc_length_compE2D, fastforce) done next case bisim1CallParams thus ?case apply(clarsimp simp add: nth_append append_eq_Cons_conv split: if_split_asm) apply(fastforce simp add: append_eq_append_conv2 Cons_eq_append_conv) apply(frule bisims1_pc_length_compEs2, clarsimp) apply(drule bisims1_pc_length_compEs2D, fastforce simp add: append_eq_append_conv2) done qed(fastforce simp add: nth_append append_eq_append_conv2 neq_Nil_conv split: if_split_asm bop.split_asm dest: bisim1_pc_length_compE2 bisims1_pc_length_compEs2)+ lemma fixes P :: "'addr J1_prog" shows bisim1_call_xcpNone: "P,e,h \ (e',xs) \ (stk,loc,pc,\a\) \ call1 e' = None" and bisims1_calls_xcpNone: "P,es,h \ (es',xs) [\] (stk,loc,pc,\a\) \ calls1 es' = None" apply(induct "(e', xs)" "(stk, loc, pc, \a :: 'addr\)" and "(es',xs)" "(stk, loc, pc, \a :: 'addr\)" arbitrary: e' xs stk loc pc and es' xs stk loc pc rule: bisim1_bisims1.inducts) apply(auto dest: bisim_Val_loc_eq_xcp_None bisims_Val_loc_eq_xcp_None simp add: is_vals_conv) done lemma bisims1_map_Val_append: assumes bisim: "P, es', h \ (es'', xs) [\] (stk, loc, pc, xcp)" shows "length es = length vs \ P, es @ es', h \ (map Val vs @ es'', xs) [\] (stk @ rev vs, loc, length (compEs2 es) + pc, xcp)" proof(induction vs arbitrary: es) case Nil thus ?case using bisim by simp next case (Cons v vs) from \length es = length (v # vs)\ obtain e es''' where [simp]: "es = e # es'''" by(cases es, auto) with \length es = length (v # vs)\ have len: "length es''' = length vs" by simp from Cons.IH[OF len] show ?case by(simp add: add.assoc append_assoc[symmetric] del: append_assoc)(rule bisims1List2, auto) qed lemma bisim1_hext_mono: "\ P,e,h \ exs \ s; hext h h' \ \ P,e,h' \ exs \ s" (is "PROP ?thesis1") and bisims1_hext_mono: "\ P,es,h \ esxs [\] s; hext h h' \ \ P,es,h' \ esxs [\] s" (is "PROP ?thesis2") proof - assume hext: "hext h h'" have "P,e,h \ exs \ s \ P,e,h' \ exs \ s" and "P,es,h \ esxs [\] s \ P,es,h' \ esxs [\] s" apply(induct rule: bisim1_bisims1.inducts) apply(insert hext) apply(auto intro: bisim1_bisims1.intros dest: hext_objD) done thus "PROP ?thesis1" and "PROP ?thesis2" by auto qed declare match_ex_table_append_not_pcs [simp] match_ex_table_eq_NoneI[simp] outside_pcs_compxE2_not_matches_entry[simp] outside_pcs_compxEs2_not_matches_entry[simp] lemma bisim1_xcp_Some_not_caught: "P, e, h \ (Throw a, xs) \ (stk, loc, pc, \a\) \ match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e pc' d) = None" and bisims1_xcp_Some_not_caught: "P, es, h \ (map Val vs @ Throw a # es', xs) [\] (stk, loc, pc, \a\) \ match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxEs2 es pc' d) = None" proof(induct "(Throw a :: 'addr expr1, xs)" "(stk, loc, pc, \a :: 'addr\)" and "(map Val vs @ Throw a # es' :: 'addr expr1 list, xs)" "(stk, loc, pc, \a :: 'addr\)" arbitrary: xs stk loc pc pc' d and xs stk loc pc vs es' pc' d rule: bisim1_bisims1.inducts) case bisim1Sync10 thus ?case by(simp add: matches_ex_entry_def) next case bisim1Sync11 thus ?case by(simp add: matches_ex_entry_def) next case (bisim1SyncThrow e1 xs stk loc pc e2) note IH = \match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e1 pc' d) = None\ from \P,e1,h \ (Throw a,xs) \ (stk,loc,pc,\a\)\ have "pc < length (compE2 e1)" by(auto dest: bisim1_ThrowD) with IH show ?case by(auto simp add: match_ex_table_append matches_ex_entry_def dest: match_ex_table_pc_length_compE2 intro: match_ex_table_not_pcs_None) next case bisims1List1 thus ?case by(auto simp add: Cons_eq_append_conv dest: bisim1_ThrowD bisim_Val_loc_eq_xcp_None) next case (bisims1List2 es es'' xs stk loc pc e v) hence "\pc' d. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxEs2 es pc' d) = None" by(auto simp add: Cons_eq_append_conv) from this[of "pc' + length (compE2 e)" "Suc d"] show ?case by(auto simp add: add.assoc) next case (bisim1BlockThrowSome e xs stk loc pc T v) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e pc' d) = None" by auto from this[of "2+pc'"] show ?case by(auto) next case (bisim1Seq2 e2 stk loc pc e1 xs) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e2 pc' d) = None" by auto from this[of "Suc (pc' + length (compE2 e1))"] show ?case by(simp add: add.assoc) next case (bisim1CondThen e1 stk loc pc e e2 xs) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e1 pc' d) = None" by auto note this[of "Suc (pc' + length (compE2 e))"] moreover from \P,e1,h \ (Throw a,xs) \ (stk,loc,pc,\a\)\ have "pc < length (compE2 e1)" by(auto dest: bisim1_ThrowD) ultimately show ?case by(simp add: add.assoc match_ex_table_eq_NoneI outside_pcs_compxE2_not_matches_entry) next case (bisim1CondElse e2 stk loc pc e e1 xs) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e2 pc' d) = None" by auto note this[of "Suc (Suc (pc' + (length (compE2 e) + length (compE2 e1))))"] thus ?case by(simp add: add.assoc) next case (bisim1WhileThrow2 e xs stk loc pc c) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e pc' d) = None" by auto from this[of "Suc (pc' + (length (compE2 c)))"] show ?case by(simp add: add.assoc) next case (bisim1Throw1 e xs stk loc pc) thus ?case by(auto dest: bisim_Val_loc_eq_xcp_None) next case (bisim1TryFail e xs stk loc pc C' C e2) hence "match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e pc' d) = None" by auto moreover from \P,e,h \ (Throw a,xs) \ (stk,loc,pc,\a\)\ have "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) ultimately show ?case using \typeof_addr h a = \Class_type C'\\ \\ P \ C' \\<^sup>* C\ by(simp add: matches_ex_entry_def cname_of_def) next case (bisim1TryCatchThrow e2 xs stk loc pc e C) hence "\pc'. match_ex_table (compP f P) (cname_of h a) (pc' + pc) (compxE2 e2 pc' d) = None" by auto from this[of "Suc (Suc (pc' + (length (compE2 e))))"] show ?case by(simp add: add.assoc matches_ex_entry_def) next case bisim1Sync12 thus ?case by(auto dest: bisim1_ThrowD simp add: match_ex_table_append eval_nat_numeral, simp add: matches_ex_entry_def) next case bisim1Sync14 thus ?case by(auto dest: bisim1_ThrowD simp add: match_ex_table_append eval_nat_numeral, simp add: matches_ex_entry_def) qed(fastforce dest: bisim1_ThrowD simp add: add.assoc[symmetric])+ declare match_ex_table_append_not_pcs [simp del] match_ex_table_eq_NoneI[simp del] outside_pcs_compxE2_not_matches_entry[simp del] outside_pcs_compxEs2_not_matches_entry[simp del] lemma bisim1_xcp_pcD: "P,e,h \ (e', xs) \ (stk, loc, pc, \a\) \ pc < length (compE2 e)" and bisims1_xcp_pcD: "P,es,h \ (es', xs) [\] (stk, loc, pc, \a\) \ pc < length (compEs2 es)" by(induct "(e', xs)" "(stk, loc, pc, \a :: 'addr\)" and "(es', xs)" "(stk, loc, pc, \a :: 'addr\)" arbitrary: e' xs stk loc pc and es' xs stk loc pc rule: bisim1_bisims1.inducts) auto declare nth_Cons_subtract[simp] declare nth_append [simp] lemma bisim1_Val_\Exec_move: "\ P, E, h \ (Val v, xs) \ (stk, loc, pc, xcp); pc < length (compE2 E) \ \ xs = loc \ xcp = None \ \Exec_mover_a P t E h (stk, xs, pc, None) ([v], xs, length (compE2 E), None)" and bisims1_Val_\Exec_moves: "\ P, Es, h \ (map Val vs, xs) [\] (stk, loc, pc, xcp); pc < length (compEs2 Es) \ \ xs = loc \ xcp = None \ \Exec_movesr_a P t Es h (stk, xs, pc, None) (rev vs, xs, length (compEs2 Es), None)" proof(induct "(Val v :: 'addr expr1, xs)" "(stk, loc, pc, xcp)" and "(map Val vs :: 'addr expr1 list, xs)" "(stk, loc, pc, xcp)" arbitrary: v xs stk loc pc xcp and vs xs stk loc pc xcp rule: bisim1_bisims1.inducts) case bisim1Val thus ?case by(auto intro!: \Execr1step exec_instr \move2Val simp add: exec_move_def) next case (bisim1LAss2 V e xs) have "\Exec_mover_a P t (V:=e) h ([], xs, Suc (length (compE2 e)), None) ([Unit], xs, Suc (Suc (length (compE2 e))), None)" by(auto intro!: \Execr1step exec_instr \move2LAssRed2 simp add: nth_append exec_move_def) with bisim1LAss2 show ?case by simp next case (bisim1AAss4 a i e xs) have "\Exec_mover_a P t (a\i\ := e) h ([], xs, Suc (length (compE2 a) + length (compE2 i) + length (compE2 e)), None) ([Unit], xs, Suc (Suc (length (compE2 a) + length (compE2 i) + length (compE2 e))), None)" by(auto intro!: \Execr1step exec_instr \move2AAssRed simp add: nth_append exec_move_def) with bisim1AAss4 show ?case by(simp add: ac_simps) next case (bisim1FAss3 e F D e2 xs) have "\Exec_mover_a P t (e\F{D} := e2) h ([], xs, Suc (length (compE2 e) + length (compE2 e2)), None) ([Unit], xs, Suc (Suc (length (compE2 e) + length (compE2 e2))), None)" by(auto intro!: \Execr1step exec_instr \move2FAssRed simp add: nth_append exec_move_def) with bisim1FAss3 show ?case by simp next case (bisim1Sync6 V e1 e2 v xs) have "\Exec_mover_a P t (sync\<^bsub>V\<^esub> (e1) e2) h ([v], xs, 5 + length (compE2 e1) + length (compE2 e2), None) ([v], xs, 9 + length (compE2 e1) + length (compE2 e2), None)" by(rule \Execr1step)(auto intro: exec_instr \move2Sync6 simp add: exec_move_def) with bisim1Sync6 show ?case by(auto simp add: eval_nat_numeral) next case (bisim1Seq2 e2 stk loc pc xcp e1 v xs) from \Suc (length (compE2 e1) + pc) < length (compE2 (e1;; e2))\ have pc: "pc < length (compE2 e2)" by simp with \pc < length (compE2 e2) \ xs = loc \ xcp = None \ \Exec_mover_a P t e2 h (stk, xs, pc, None) ([v], xs, length (compE2 e2), None)\ have "xs = loc" "xcp = None" "\Exec_mover_a P t e2 h (stk, xs, pc, None) ([v], xs, length (compE2 e2), None)" by auto moreover hence "\Exec_mover_a P t (e1;;e2) h (stk, xs, Suc (length (compE2 e1) + pc), None) ([v], xs, Suc (length (compE2 e1) + length (compE2 e2)), None)" by -(rule Seq_\ExecrI2) ultimately show ?case by(simp) next case (bisim1CondThen e1 stk loc pc xcp e e2 v xs) from \P, e1, h \ (Val v,xs) \ (stk,loc,pc,xcp)\ have "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) have e: "\Exec_mover_a P t (if (e) e1 else e2) h ([v], xs, Suc (length (compE2 e) + (length (compE2 e1))), None) ([v], xs, length (compE2 (if (e) e1 else e2)), None)" by(rule \Execr1step)(auto simp add: nth_append exec_move_def intro!: exec_instr \move2CondThenExit) show ?case proof(cases "pc < length (compE2 e1)") case True with \pc < length (compE2 e1) \ xs = loc \ xcp = None \ \Exec_mover_a P t e1 h (stk, xs, pc, None) ([v], xs, length (compE2 e1), None)\ have s: "xs = loc" "xcp = None" and "\Exec_mover_a P t e1 h (stk, xs, pc, None) ([v], xs, length (compE2 e1), None)" by auto hence "\Exec_mover_a P t (if (e) e1 else e2) h (stk, xs, Suc (length (compE2 e) + pc), None) ([v], xs, Suc (length (compE2 e) + length (compE2 e1)), None)" by -(rule Cond_\ExecrI2) with e True s show ?thesis by(simp) next case False with \pc \ length (compE2 e1)\ have pc: "pc = length (compE2 e1)" by auto with \P, e1, h \ (Val v,xs) \ (stk,loc,pc,xcp)\ have "stk = [v]" "xs = loc" "xcp = None" by(auto dest: bisim1_Val_length_compE2D) with pc e show ?thesis by(simp) qed next case (bisim1CondElse e2 stk loc pc xcp e e1 v xs) from \P, e2, h \ (Val v,xs) \ (stk,loc,pc,xcp)\ have "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e2)") case True with \pc < length (compE2 e2) \ xs = loc \ xcp = None \ \Exec_mover_a P t e2 h (stk, xs, pc, None) ([v], xs, length (compE2 e2), None)\ have s: "xs = loc" "xcp = None" and e: "\Exec_mover_a P t e2 h (stk, xs, pc, None) ([v], xs, length (compE2 e2), None)" by auto from e have "\Exec_mover_a P t (if (e) e1 else e2) h (stk, xs, Suc (Suc (length (compE2 e) + length (compE2 e1) + pc)), None) ([v], xs, Suc (Suc (length (compE2 e) + length (compE2 e1) + length (compE2 e2))), None)" by(rule Cond_\ExecrI3) with True s show ?thesis by(simp add: add.assoc) next case False with \pc \ length (compE2 e2)\ have pc: "pc = length (compE2 e2)" by auto with \P, e2, h \ (Val v,xs) \ (stk,loc,pc,xcp)\ have "stk = [v]" "xs = loc" "xcp = None" by(auto dest: bisim1_Val_length_compE2D) with pc show ?thesis by(simp add: add.assoc) qed next case (bisim1While7 c e xs) have "\Exec_mover_a P t (while (c) e) h ([], xs, Suc (Suc (Suc (length (compE2 c) + length (compE2 e)))), None) ([Unit], xs, length (compE2 (while (c) e)), None)" by(auto intro!: \Execr1step exec_instr \move2While4 simp add: nth_append exec_move_def) thus ?case by(simp) next case (bisims1List1 e e' xs stk loc pc xcp es) from \e' # es = map Val vs\ obtain v vs' where [simp]: "vs = v # vs'" "e' = Val v" "es = map Val vs'" by auto from \P,e,h \ (e',xs) \ (stk,loc,pc,xcp)\ have length: "pc \ length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) hence "xs = loc \ xcp = None \ \Exec_mover_a P t e h (stk, xs, pc, None) ([v], xs, length (compE2 e), None)" proof(cases "pc < length (compE2 e)") case True with \\e' = Val v; pc < length (compE2 e)\ \ xs = loc \ xcp = None \ \Exec_mover_a P t e h (stk, xs, pc, None) ([v], xs, length (compE2 e), None)\ show ?thesis by auto next case False with length have pc: "pc = length (compE2 e)" by auto with \P,e,h \ (e',xs) \ (stk,loc,pc,xcp)\ have "stk = [v]" "xs = loc" "xcp = None" by(auto dest: bisim1_Val_length_compE2D) with pc show ?thesis by(auto) qed hence s: "xs = loc" "xcp = None" and exec1: "\Exec_mover_a P t e h (stk, xs, pc, None) ([v], xs, length (compE2 e), None)" by auto from exec1 have "\Exec_movesr_a P t (e # es) h (stk, xs, pc, None) ([v], xs, length (compE2 e), None)" by(auto intro: \Exec_mover_\Exec_movesr) moreover have "\Exec_movesr_a P t (map Val vs') h ([], xs, 0, None) (rev vs', xs, length (compEs2 (map Val vs')), None)" by(rule \Exec_movesr_map_Val) hence "\Exec_movesr_a P t ([e] @ map Val vs') h ([] @ [v], xs, length (compEs2 [e]) + 0, None) (rev vs' @ [v], xs, length (compEs2 [e]) + length (compEs2 (map Val vs')), None)" by -(rule append_\Exec_movesr, auto) ultimately show ?case using s by(auto) next case (bisims1List2 es es' xs stk loc pc xcp e v) from \Val v # es' = map Val vs\ obtain vs' where [simp]: "vs = v # vs'" "es' = map Val vs'" by auto from \P,es,h \ (es',xs) [\] (stk,loc,pc,xcp)\ have length: "pc \ length (compEs2 es)" by(auto dest: bisims1_pc_length_compEs2) hence "xs = loc \ xcp = None \ \Exec_movesr_a P t es h (stk, xs, pc, None) (rev vs', xs, length (compEs2 es), None)" proof(cases "pc < length (compEs2 es)") case True with \\es' = map Val vs'; pc < length (compEs2 es)\ \ xs = loc \ xcp = None \ \Exec_movesr_a P t es h (stk, xs, pc, None) (rev vs', xs, length (compEs2 es), None)\ show ?thesis by auto next case False with length have pc: "pc = length (compEs2 es)" by auto with \P,es,h \ (es',xs) [\] (stk,loc,pc,xcp)\ have "stk = rev vs'" "xs = loc" "xcp = None" by(auto dest: bisims1_Val_length_compEs2D) with pc show ?thesis by(auto) qed hence s: "xs = loc" "xcp = None" and exec1: "\Exec_movesr_a P t es h (stk, xs, pc, None) (rev vs', xs, length (compEs2 es), None)" by auto from exec1 have "\Exec_movesr_a P t ([e] @ es) h (stk @ [v], xs, length (compEs2 [e]) + pc, None) (rev vs' @ [v], xs, length (compEs2 [e]) + length (compEs2 es), None)" by -(rule append_\Exec_movesr, auto) thus ?case using s by(auto) qed(auto) lemma bisim1Val2D1: assumes bisim: "P, e, h \ (Val v,xs) \ (stk,loc,pc,xcp)" shows "xcp = None \ xs = loc \ \Exec_mover_a P t e h (stk, loc, pc, xcp) ([v], loc, length (compE2 e), None)" proof - from bisim have "xcp = None" "xs = loc" by(auto dest: bisim_Val_loc_eq_xcp_None) moreover have "\Exec_mover_a P t e h (stk, loc, pc, xcp) ([v], loc, length (compE2 e), None)" proof(cases "pc < length (compE2 e)") case True from bisim1_Val_\Exec_move[OF bisim True] show ?thesis by auto next case False from bisim have "pc \ length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) with False have "pc = length (compE2 e)" by auto with bisim have "stk = [v]" "loc = xs" "xcp=None" by(auto dest: bisim1_Val_length_compE2D) with \pc = length (compE2 e)\ show ?thesis by(auto) qed ultimately show ?thesis by simp qed lemma bisim1_Throw_\Exec_movet: "\ P, e, h \ (Throw a,xs) \ (stk,loc,pc,None) \ \ \pc'. \Exec_movet_a P t e h (stk, loc, pc, None) ([Addr a], loc, pc', \a\) \ P, e, h \ (Throw a,xs) \ ([Addr a], loc, pc', \a\) \ xs = loc" and bisims1_Throw_\Exec_movest: "\ P, es, h \ (map Val vs @ Throw a # es',xs) [\] (stk,loc,pc,None) \ \ \pc'. \Exec_movest_a P t es h (stk, loc, pc, None) (Addr a # rev vs, loc, pc', \a\) \ P, es, h \ (map Val vs @ Throw a # es',xs) [\] (Addr a # rev vs, loc, pc', \a\) \ xs = loc" proof(induct e "n :: nat" "Throw a :: 'addr expr1" xs stk loc pc "None :: 'addr option" and es "n :: nat" "map Val vs @ Throw a # es' :: 'addr expr1 list" xs stk loc pc "None :: 'addr option" arbitrary: and vs rule: bisim1_bisims1_inducts_split) case (bisim1Sync9 e1 n e2 V xs) let ?pc = "8 + length (compE2 e1) + length (compE2 e2)" have "\Exec_movet_a P t (sync\<^bsub>V\<^esub> (e1) e2) h ([Addr a], xs, ?pc, None) ([Addr a], xs, ?pc, \a\)" by(rule \Exect1step)(auto intro: exec_instr \move2_\moves2.intros simp add: is_Ref_def exec_move_def) moreover have "P,sync\<^bsub>V\<^esub> (e1) e2,h \ (Throw a,xs) \ ([Addr a],xs,?pc,\a\)" by(rule bisim1Sync10) ultimately show ?case by auto next case (bisim1Seq2 e2 n xs stk loc pc e1) then obtain pc' where "\Exec_movet_a P t e2 h (stk, loc, pc, None) ([Addr a], loc, pc', \a\)" "P, e2, h \ (Throw a,xs) \ ([Addr a],loc,pc',\a\)" "xs = loc" by auto thus ?case by(auto intro: Seq_\ExectI2 bisim1_bisims1.bisim1Seq2) next case (bisim1CondThen e1 n xs stk loc pc e e2) then obtain pc' where exec: "\Exec_movet_a P t e1 h (stk, loc, pc, None) ([Addr a], loc, pc', \a\)" and bisim: "P, e1, h \ (Throw a,xs) \ ([Addr a],loc,pc',\a\)" and s: "xs = loc" by auto from exec have "\Exec_movet_a P t (if (e) e1 else e2) h (stk, loc, Suc (length (compE2 e) + pc), None) ([Addr a], loc, Suc (length (compE2 e) + pc'), \a\)" by(rule Cond_\ExectI2) moreover from bisim have "P, if (e) e1 else e2, h \ (Throw a, xs) \ ([Addr a], loc, Suc (length (compE2 e) + pc'), \a\)" by(rule bisim1_bisims1.bisim1CondThen) ultimately show ?case using s by(auto) next case (bisim1CondElse e2 n xs stk loc pc e e1) then obtain pc' where exec: "\Exec_movet_a P t e2 h (stk, loc, pc, None) ([Addr a], loc, pc', \a\)" and bisim: "P, e2, h \ (Throw a, xs ) \ ([Addr a], loc, pc', \a\)" and s: "xs = loc" by auto let "?pc pc" = "Suc (Suc (length (compE2 e) + length (compE2 e1) + pc))" from exec have "\Exec_movet_a P t (if (e) e1 else e2) h (stk, loc, (?pc pc), None) ([Addr a], loc, ?pc pc', \a\)" by(rule Cond_\ExectI3) moreover from bisim have "P, if (e) e1 else e2, h \ (Throw a, xs ) \ ([Addr a], loc, ?pc pc', \a\)" by(rule bisim1_bisims1.bisim1CondElse) ultimately show ?case using s by auto next case (bisim1Throw1 e n xs stk loc pc) note bisim = \P, e, h \ (addr a, xs) \ (stk, loc, pc, None)\ hence s: "xs = loc" and exec: "\Exec_mover_a P t e h (stk, loc, pc, None) ([Addr a], loc, length (compE2 e), None)" by(auto dest: bisim1Val2D1) from exec have "\Exec_mover_a P t (throw e) h (stk, loc, pc, None) ([Addr a], loc, length (compE2 e), None)" by(rule Throw_\ExecrI) also have "\Exec_movet_a P t (throw e) h ([Addr a], loc, length (compE2 e), None) ([Addr a], loc, length (compE2 e), \a\)" by(rule \Exect1step, auto intro: exec_instr \move2Throw2 simp add: is_Ref_def exec_move_def) also have "P, throw e, h \ (Throw a, loc ) \ ([Addr a], loc, length (compE2 e), \a\)" by(rule bisim1Throw2) ultimately show ?case using s by auto next case (bisims1List1 e n e' xs stk loc pc es vs) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ show ?case proof(cases "is_val e'") case True with \e' # es = map Val vs @ Throw a # es'\ obtain v vs' where "vs = v # vs'" "e' = Val v" and es: "es = map Val vs' @ Throw a # es'" by(auto simp add: Cons_eq_append_conv) with bisim have "P,e,h \ (Val v, xs) \ (stk, loc, pc, None)" by simp from bisim1Val2D1[OF this] have [simp]: "xs = loc" and exec: "\Exec_mover_a P t e h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" by auto from exec have "\Exec_movesr_a P t (e # es) h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" by(rule \Exec_mover_\Exec_movesr) also from es \es = map Val vs' @ Throw a # es' \ \pc'. \Exec_movest_a P t es h ([], loc, 0, None) (Addr a # rev vs', loc, pc', \a\) \ P,es,h \ (map Val vs' @ Throw a # es', loc) [\] (Addr a # rev vs', loc, pc', \a\) \ loc = loc\ obtain pc' where execes: "\Exec_movest_a P t es h ([], loc, 0, None) (Addr a # rev vs', loc, pc', \a\)" and bisim': "P,es,h \ (map Val vs' @ Throw a # es', loc) [\] (Addr a # rev vs', loc, pc', \a\)" by auto from append_\Exec_movest[OF _ execes, of "[v]" "[e]"] have "\Exec_movest_a P t (e # es) h ([v], loc, length (compE2 e), None) (Addr a # rev vs' @ [v], loc, length (compE2 e) + pc', \a\)" by simp also from bisims1List2[OF bisim', of e v] es \e' = Val v\ \vs = v # vs'\ have "P,e # es,h \ (e' # es, xs) [\] ((Addr a # rev vs), loc, length (compE2 e) + pc', \a\)" by simp ultimately show ?thesis using \vs = v # vs'\ es \e' = Val v\ by auto next case False with \e' # es = map Val vs @ Throw a # es'\ have [simp]: "e' = Throw a" "es = es'" "vs = []" by(auto simp add: Cons_eq_append_conv) from \e' = Throw a \ \pc'. \Exec_movet_a P t e h (stk, loc, pc, None) ([Addr a], loc, pc', \a\) \ P,e,h \ (Throw a, xs ) \ ([Addr a], loc, pc', \a\) \ xs = loc\ obtain pc' where "\Exec_movet_a P t e h (stk, loc, pc, None) ([Addr a], loc, pc', \a\)" and bisim: "P,e,h \ (Throw a, xs ) \ ([Addr a], loc, pc', \a\)" and s: "xs = loc" by auto hence "\Exec_movest_a P t (e # es) h (stk, loc, pc, None) ([Addr a], loc, pc', \a\)" by-(rule \Exec_movet_\Exec_movest) moreover from bisim have "P,e#es,h \ (Throw a#es,xs) [\] ([Addr a],loc,pc',\a\)" by(rule bisim1_bisims1.bisims1List1) ultimately show ?thesis using s by auto qed next case (bisims1List2 es n es'' xs stk loc pc e v) note IH = \\vs. es'' = map Val vs @ Throw a # es' \ \pc'. \Exec_movest_a P t es h (stk, loc, pc, None) (Addr a # rev vs, loc, pc', \a\) \ P,es,h \ (map Val vs @ Throw a # es',xs) [\] (Addr a # rev vs,loc,pc',\a\) \ xs = loc\ from \Val v # es'' = map Val vs @ Throw a # es'\ obtain vs' where [simp]: "vs = v # vs'" "es'' = map Val vs' @ Throw a # es'" by(auto simp add: Cons_eq_append_conv) from IH[OF \es'' = map Val vs' @ Throw a # es'\] obtain pc' where exec: "\Exec_movest_a P t es h (stk, loc, pc, None) (Addr a # rev vs', loc, pc', \a\)" and bisim: "P,es,h \ (map Val vs' @ Throw a # es',xs) [\] (Addr a # rev vs',loc,pc',\a\)" and [simp]: "xs = loc" by auto from append_\Exec_movest[OF _ exec, of "[v]" "[e]"] have "\Exec_movest_a P t (e # es) h (stk @ [v], loc, length (compE2 e) + pc, None) (Addr a # rev vs, loc, length (compE2 e) + pc', \a\)" by simp moreover from bisim have "P,e#es,h \ (Val v # map Val vs' @ Throw a # es',xs) [\] ((Addr a # rev vs')@[v],loc,length (compE2 e) + pc',\a\)" by(rule bisim1_bisims1.bisims1List2) ultimately show ?case by(auto) qed(auto) lemma bisim1_Throw_\Exec_mover: "\ P, e, h \ (Throw a,xs) \ (stk,loc,pc,None) \ \ \pc'. \Exec_mover_a P t e h (stk, loc, pc, None) ([Addr a], loc, pc', \a\) \ P, e, h \ (Throw a,xs) \ ([Addr a], loc, pc', \a\) \ xs = loc" by(drule bisim1_Throw_\Exec_movet)(blast intro: tranclp_into_rtranclp) lemma bisims1_Throw_\Exec_movesr: "\ P, es, h \ (map Val vs @ Throw a # es',xs) [\] (stk,loc,pc,None) \ \ \pc'. \Exec_movesr_a P t es h (stk, loc, pc, None) (Addr a # rev vs, loc, pc', \a\) \ P, es, h \ (map Val vs @ Throw a # es',xs) [\] (Addr a # rev vs, loc, pc', \a\) \ xs = loc" by(drule bisims1_Throw_\Exec_movest)(blast intro: tranclp_into_rtranclp) declare split_beta [simp] lemma bisim1_inline_call_Throw: "\ P,e,h \ (e', xs) \ (stk, loc, pc, None); call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0; pc < length (compE2 e) \ \ n0 = length vs \ P,e,h \ (inline_call (Throw A) e', xs) \ (stk, loc, pc, \A\)" (is "\ _; _; _; _ \ \ ?concl e n e' xs pc stk loc") and bisims1_inline_calls_Throw: "\ P,es,h \ (es', xs) [\] (stk, loc, pc, None); calls1 es' = \(a, M, vs)\; compEs2 es ! pc = Invoke M n0; pc < length (compEs2 es) \ \ n0 = length vs \ P,es,h \ (inline_calls (Throw A) es', xs) [\] (stk, loc, pc, \A\)" (is "\ _; _; _; _ \ \ ?concls es n es' xs pc stk loc") proof(induct e "n :: nat" e' xs stk loc pc "None :: 'addr option" and es "n :: nat" es' xs stk loc pc "None :: 'addr option" rule: bisim1_bisims1_inducts_split) case (bisim1BinOp1 e1 n e' xs stk loc pc e2 bop) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e1 ! pc = Invoke M n0; pc < length (compE2 e1) \ \ ?concl e1 n e' xs pc stk loc\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note ins = \compE2 (e1 \bop\ e2) ! pc = Invoke M n0\ note call = \call1 (e' \bop\ e2) = \(a, M, vs)\\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e1)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis by(auto simp add: nth_append intro: bisim1_bisims1.bisim1BinOp1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e1)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e1)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 e1)" by(cases "pc < length (compE2 e1)") auto from call ins show ?thesis by simp qed next case bisim1BinOp2 thus ?case by(auto split: if_split_asm bop.split_asm dest: bisim1_bisims1.bisim1BinOp2) next case (bisim1AAcc1 A n a' xs stk loc pc i) note IH1 = \\call1 a' = \(a, M, vs)\; compE2 A ! pc = Invoke M n0; pc < length (compE2 A) \ \ ?concl A n a' xs pc stk loc\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note ins = \compE2 (A\i\) ! pc = Invoke M n0\ note call = \call1 (a'\i\) = \(a, M, vs)\\ show ?case proof(cases "is_val a'") case False with bisim1 call have "pc < length (compE2 A)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis by(auto simp add: nth_append intro: bisim1_bisims1.bisim1AAcc1) next case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "pc \ length (compE2 A)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 A)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 A)" by(cases "pc < length (compE2 A)") auto from call ins show ?thesis by simp qed next case bisim1AAcc2 thus ?case by(auto split: if_split_asm dest: bisim1_bisims1.bisim1AAcc2) next case (bisim1AAss1 A n a' xs stk loc pc i e) note IH1 = \\call1 a' = \(a, M, vs)\; compE2 A ! pc = Invoke M n0; pc < length (compE2 A) \ \ ?concl A n a' xs pc stk loc\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note ins = \compE2 (A\i\ := e) ! pc = Invoke M n0\ note call = \call1 (a'\i\ := e) = \(a, M, vs)\\ show ?case proof(cases "is_val a'") case False with bisim1 call have "pc < length (compE2 A)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis by(auto simp add: nth_append intro: bisim1_bisims1.bisim1AAss1) next case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "pc \ length (compE2 A)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 A)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 A)" by(cases "pc < length (compE2 A)") auto from call ins show ?thesis by simp qed next case (bisim1AAss2 i n i' xs stk loc pc A e v) note IH1 = \\call1 i' = \(a, M, vs)\; compE2 i ! pc = Invoke M n0; pc < length (compE2 i) \ \ ?concl i n i' xs pc stk loc\ note bisim1 = \P,i,h \ (i', xs) \ (stk, loc, pc, None)\ note ins = \compE2 (A\i\ := e) ! (length (compE2 A) + pc) = Invoke M n0\ note call = \call1 (Val v\i'\ := e) = \(a, M, vs)\\ show ?case proof(cases "is_val i'") case False with bisim1 call have "pc < length (compE2 i)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis by(auto simp add: nth_append intro: bisim1_bisims1.bisim1AAss2) next case True then obtain v where [simp]: "i' = Val v" by auto from bisim1 have "pc \ length (compE2 i)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 i)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 i)" by(cases "pc < length (compE2 i)") auto from call ins show ?thesis by simp qed next case bisim1AAss3 thus ?case by(auto split: if_split_asm nat.split_asm simp add: nth_Cons dest: bisim1_bisims1.bisim1AAss3) next case (bisim1FAss1 e n e' xs stk loc pc e2 F D) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0; pc < length (compE2 e) \ \ ?concl e n e' xs pc stk loc\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note ins = \compE2 (e\F{D} := e2) ! pc = Invoke M n0\ note call = \call1 (e'\F{D} := e2) = \(a, M, vs)\\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis by(auto simp add: nth_append intro: bisim1_bisims1.bisim1FAss1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 e)" by(cases "pc < length (compE2 e)") auto from call ins show ?thesis by simp qed next case bisim1FAss2 thus ?case by(auto split: if_split_asm nat.split_asm simp add: nth_Cons dest: bisim1_bisims1.bisim1FAss2) next case (bisim1CAS1 E n e' xs stk loc pc e2 e3 D F) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 E ! pc = Invoke M n0; pc < length (compE2 E) \ \ ?concl E n e' xs pc stk loc\ note bisim1 = \P,E,h \ (e', xs) \ (stk, loc, pc, None)\ note ins = \compE2 _ ! pc = Invoke M n0\ note call = \call1 _ = \(a, M, vs)\\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 E)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis by(auto simp add: nth_append intro: bisim1_bisims1.bisim1CAS1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 E)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 E)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 E)" by(cases "pc < length (compE2 E)") auto from call ins show ?thesis by simp qed next case (bisim1CAS2 e2 n e2' xs stk loc pc e1 e3 D F v) note IH1 = \\call1 e2' = \(a, M, vs)\; compE2 e2 ! pc = Invoke M n0; pc < length (compE2 e2) \ \ ?concl e2 n e2' xs pc stk loc\ note bisim1 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, None)\ note ins = \compE2 _ ! (length (compE2 e1) + pc) = Invoke M n0\ note call = \call1 _ = \(a, M, vs)\\ show ?case proof(cases "is_val e2'") case False with bisim1 call have "pc < length (compE2 e2)" by(auto intro: bisim1_call_pcD) with call ins IH1 False show ?thesis by(auto simp add: nth_append intro: bisim1_bisims1.bisim1CAS2) next case True then obtain v where [simp]: "e2' = Val v" by auto from bisim1 have "pc \ length (compE2 e2)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e2)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 e2)" by(cases "pc < length (compE2 e2)") auto from call ins show ?thesis by simp qed next case (bisim1Call1 obj n obj' xs stk loc pc ps M') note IH1 = \\call1 obj' = \(a, M, vs)\; compE2 obj ! pc = Invoke M n0; pc < length (compE2 obj) \ \ ?concl obj n obj' xs pc stk loc\ note IH2 = \\xs. \calls1 ps = \(a, M, vs)\; compEs2 ps ! 0 = Invoke M n0; 0 < length (compEs2 ps) \ \ ?concls ps n ps xs 0 [] xs\ note ins = \compE2 (obj\M'(ps)) ! pc = Invoke M n0\ note bisim1 = \P,obj,h \ (obj', xs) \ (stk, loc, pc, None)\ note call = \call1 (obj'\M'(ps)) = \(a, M, vs)\\ thus ?case proof(cases rule: call1_callE) case CallObj with bisim1 call have "pc < length (compE2 obj)" by(auto intro: bisim1_call_pcD) with call ins CallObj IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1Call1) next case (CallParams v) from bisim1 have "pc \ length (compE2 obj)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 obj)" with bisim1 ins CallParams have False by(auto dest: bisim_Val_pc_not_Invoke) } ultimately have [simp]: "pc = length (compE2 obj)" by(cases "pc < length (compE2 obj)") auto with bisim1 CallParams have [simp]: "stk = [v]" "loc = xs" by(auto dest: bisim1_Val_length_compE2D) from IH2[of loc] CallParams ins show ?thesis apply(clarsimp simp add: compEs2_map_Val is_vals_conv split: if_split_asm) apply(drule bisim1_bisims1.bisim1CallParams) apply(auto simp add: neq_Nil_conv) done next case [simp]: Call from bisim1 have "pc \ length (compE2 obj)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 obj)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 obj)" by(cases "pc < length (compE2 obj)") auto with ins have [simp]: "vs = []" by(auto simp add: nth_append compEs2_map_Val split: if_split_asm) from bisim1 have [simp]: "stk = [Addr a]" "xs = loc" by(auto dest: bisim1_Val_length_compE2D) from ins show ?thesis by(auto intro: bisim1CallThrow[of "[]" "[]", simplified]) qed next case (bisim1CallParams ps n ps' xs stk loc pc obj M' v) note IH2 = \\calls1 ps' = \(a, M, vs)\; compEs2 ps ! pc = Invoke M n0; pc < length (compEs2 ps) \ \ ?concls ps n ps' xs pc stk loc\ note ins = \compE2 (obj\M'(ps)) ! (length (compE2 obj) + pc) = Invoke M n0\ note bisim2 = \P,ps,h \ (ps', xs) [\] (stk, loc, pc, None)\ note call = \call1 (Val v\M'(ps')) = \(a, M, vs)\\ thus ?case proof(cases rule: call1_callE) case CallObj thus ?thesis by simp next case (CallParams v') hence [simp]: "v' = v" and call': "calls1 ps' = \(a, M, vs)\" by auto from bisim2 call' have "pc < length (compEs2 ps)" by(auto intro: bisims1_calls_pcD) with IH2 CallParams ins show ?thesis by(auto simp add: is_vals_conv split: if_split_asm intro: bisim1_bisims1.bisim1CallParams) next case Call hence [simp]: "v = Addr a" "M' = M" "ps' = map Val vs" by auto from bisim2 have "pc \ length (compEs2 ps)" by(auto dest: bisims1_pc_length_compEs2) moreover { assume pc: "pc < length (compEs2 ps)" with bisim2 ins have False by(auto dest: bisims_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compEs2 ps)" by(cases "pc < length (compEs2 ps)") auto from bisim2 have [simp]: "stk = rev vs" "xs = loc" by(auto dest: bisims1_Val_length_compEs2D) from bisim2 have "length ps = length vs" by(auto dest: bisims1_lengthD) with ins show ?thesis by(auto intro: bisim1CallThrow) qed next case (bisims1List1 e n e' xs stk loc pc es) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0; pc < length (compE2 e) \ \ ?concl e n e' xs pc stk loc\ note IH2 = \\xs. \calls1 es = \(a, M, vs)\; compEs2 es ! 0 = Invoke M n0; 0 < length (compEs2 es) \ \ ?concls es n es xs 0 [] xs\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \calls1 (e' # es) = \(a, M, vs)\\ note ins = \compEs2 (e # es) ! pc = Invoke M n0\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto simp add: nth_append split_beta intro: bisim1_bisims1.bisims1List1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 e)" by(cases "pc < length (compE2 e)") auto with bisim1 have [simp]: "stk = [v]" "loc = xs" by(auto dest: bisim1_Val_length_compE2D) from call have "es \ []" by(cases es) simp_all with IH2[of loc] call ins show ?thesis by(auto split: if_split_asm dest: bisims1List2) qed qed(auto split: if_split_asm bop.split_asm intro: bisim1_bisims1.intros dest: bisim1_pc_length_compE2) lemma bisim1_max_stack: "P,e,h \ (e', xs) \ (stk, loc, pc, xcp) \ length stk \ max_stack e" and bisims1_max_stacks: "P,es,h \ (es', xs) [\] (stk, loc, pc, xcp) \ length stk \ max_stacks es" apply(induct "(e', xs)" "(stk, loc, pc, xcp)" and "(es', xs)" "(stk, loc, pc, xcp)" arbitrary: e' xs stk loc pc xcp and es' xs stk loc pc xcp rule: bisim1_bisims1.inducts) apply(auto simp add: max_stack1[simplified] max_def max_stacks_ge_length) apply(drule sym, simp add: max_stacks_ge_length, drule sym, simp, rule le_trans[OF max_stacks_ge_length], simp) done inductive bisim1_fr :: "'addr J1_prog \ 'heap \ 'addr expr1 \ 'addr locals1 \ 'addr frame \ bool" for P :: "'addr J1_prog" and h :: 'heap where "\ P \ C sees M:Ts\T = \body\ in D; P,blocks1 0 (Class D#Ts) body, h \ (e, xs) \ (stk, loc, pc, None); call1 e = \(a, M', vs)\; max_vars e \ length xs \ \ bisim1_fr P h (e, xs) (stk, loc, C, M, pc)" declare bisim1_fr.intros [intro] declare bisim1_fr.cases [elim] lemma bisim1_fr_hext_mono: "\ bisim1_fr P h exs fr; hext h h' \ \ bisim1_fr P h' exs fr" by(auto intro: bisim1_hext_mono) lemma bisim1_max_vars: "P,E,h \ (e, xs) \ (stk, loc, pc, xcp) \ max_vars E \ max_vars e" and bisims1_max_varss: "P,Es,h \ (es,xs) [\] (stk,loc,pc,xcp) \ max_varss Es \ max_varss es" apply(induct E "(e, xs)" "(stk, loc, pc, xcp)" and Es "(es, xs)" "(stk, loc, pc, xcp)" arbitrary: e xs stk loc pc xcp and es xs stk loc pc xcp rule: bisim1_bisims1.inducts) apply(auto) done lemma bisim1_call_\Exec_move: "\ P,e,h \ (e', xs) \ (stk, loc, pc, None); call1 e' = \(a, M', vs)\; n + max_vars e' \ length xs; \ contains_insync e \ \ \pc' loc' stk'. \Exec_mover_a P t e h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None) \ pc' < length (compE2 e) \ compE2 e ! pc' = Invoke M' (length vs) \ P,e,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" (is "\ _; _; _; _ \ \ ?concl e n e' xs pc stk loc") and bisims1_calls_\Exec_moves: "\ P,es,h \ (es',xs) [\] (stk, loc, pc, None); calls1 es' = \(a, M', vs)\; n + max_varss es' \ length xs; \ contains_insyncs es \ \ \pc' stk' loc'. \Exec_movesr_a P t es h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None) \ pc' < length (compEs2 es) \ compEs2 es ! pc' = Invoke M' (length vs) \ P,es,h \ (es', xs) [\] (rev vs @ Addr a # stk', loc', pc', None)" (is "\_; _; _; _ \ \ ?concls es n es' xs pc stk loc") proof(induct e "n :: nat" e' xs stk loc pc xcp\"None :: 'addr option" and es "n :: nat" es' xs stk loc pc xcp\"None :: 'addr option" rule: bisim1_bisims1_inducts_split) case bisim1Val2 thus ?case by auto next case bisim1New thus ?case by auto next case bisim1NewArray thus ?case by auto (fastforce intro: bisim1_bisims1.bisim1NewArray elim!: NewArray_\ExecrI intro!: exI) next case bisim1Cast thus ?case by(auto)(fastforce intro: bisim1_bisims1.bisim1Cast elim!: Cast_\ExecrI intro!: exI)+ next case bisim1InstanceOf thus ?case by(auto)(fastforce intro: bisim1_bisims1.bisim1InstanceOf elim!: InstanceOf_\ExecrI intro!: exI)+ next case bisim1Val thus ?case by auto next case bisim1Var thus ?case by auto next case (bisim1BinOp1 e1 n e' xs stk loc pc e2 bop) note IH1 = \\call1 e' = \(a, M', vs)\; n + max_vars e' \ length xs; \ contains_insync e1 \ \ ?concl e1 n e' xs pc stk loc\ note IH2 = \\xs. \call1 e2 = \(a, M', vs)\; n + max_vars e2 \ length xs; \ contains_insync e2 \ \ ?concl e2 n e2 xs 0 [] xs\ note call = \call1 (e' \bop\ e2) = \(a, M', vs)\\ note len = \n + max_vars (e' \bop\ e2) \ length xs\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note cs = \\ contains_insync (e1 \bop\ e2)\ show ?case proof(cases "is_val e'") case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "\Exec_mover_a P t e1 h (stk, loc, pc, None) ([v], loc, length (compE2 e1), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_mover_a P t (e1\bop\e2) h (stk, loc, pc, None) ([v], loc, length (compE2 e1), None)" by-(rule BinOp_\ExecrI1) also from call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e2 h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e2 ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e2)" and bisim': "P,e2,h \ (e2, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from BinOp_\ExecrI2[OF exec, of e1 bop v] have "\Exec_mover_a P t (e1\bop\e2) h ([v], loc, length (compE2 e1), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 e1) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,e1\bop\e2,h \ (Val v \bop\ e2, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e1) + pc', None)" by(rule bisim1BinOp2) ultimately show ?thesis using ins by fastforce next case False with IH1 len False call cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1BinOp1 elim!: BinOp_\ExecrI1 intro!: exI) qed next case (bisim1BinOp2 e2 n e' xs stk loc pc e1 bop v1) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e2,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from exec have "\Exec_mover_a P t (e1 \bop\ e2) h (stk @ [v1], loc, length (compE2 e1) + pc, None) ((rev vs @ Addr a # stk') @ [v1], loc', length (compE2 e1) + pc', None)" by(rule BinOp_\ExecrI2) moreover from bisim' have "P,e1 \bop\ e2,h \ (Val v1 \bop\ e', xs) \ ((rev vs @ Addr a # stk') @ [v1], loc', length (compE2 e1) + pc', None)" by(rule bisim1_bisims1.bisim1BinOp2) ultimately show ?case using pc' by(fastforce) next case bisim1LAss1 thus ?case by(auto)(fastforce intro: bisim1_bisims1.bisim1LAss1 elim!: LAss_\ExecrI intro!: exI) next case bisim1LAss2 thus ?case by simp next case (bisim1AAcc1 A n a' xs stk loc pc i) note IH1 = \\call1 a' = \(a, M', vs)\; n + max_vars a' \ length xs; \ contains_insync A\ \ ?concl A n a' xs pc stk loc\ note IH2 = \\xs. \call1 i = \(a, M', vs)\; n + max_vars i \ length xs; \ contains_insync i\ \ ?concl i n i xs 0 [] xs\ note call = \call1 (a'\i\) = \(a, M', vs)\\ note len = \n + max_vars (a'\i\) \ length xs\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note cs = \\ contains_insync (A\i\)\ show ?case proof(cases "is_val a'") case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "\Exec_mover_a P t A h (stk, loc, pc, None) ([v], loc, length (compE2 A), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_mover_a P t (A\i\) h (stk, loc, pc, None) ([v], loc, length (compE2 A), None)" by-(rule AAcc_\ExecrI1) also from call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t i h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 i ! pc' = Invoke M' (length vs)" "pc' < length (compE2 i)" and bisim': "P,i,h \ (i, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from AAcc_\ExecrI2[OF exec, of A v] have "\Exec_mover_a P t (A\i\) h ([v], loc, length (compE2 A), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 A) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,A\i\,h \ (Val v\i\, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 A) + pc', None)" by(rule bisim1AAcc2) ultimately show ?thesis using ins by fastforce next case False with IH1 len False call cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1AAcc1 elim!: AAcc_\ExecrI1 intro!: exI) qed next case (bisim1AAcc2 i n i' xs stk loc pc A v) then obtain pc' loc' stk' where pc': "pc' < length (compE2 i)" "compE2 i ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t i h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,i,h \ (i', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from exec have "\Exec_mover_a P t (A\i\) h (stk @ [v], loc, length (compE2 A) + pc, None) ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 A) + pc', None)" by(rule AAcc_\ExecrI2) moreover from bisim' have "P,A\i\,h \ (Val v\i'\, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 A) + pc', None)" by(rule bisim1_bisims1.bisim1AAcc2) ultimately show ?case using pc' by(fastforce) next case (bisim1AAss1 A n a' xs stk loc pc i e) note IH1 = \\call1 a' = \(a, M', vs)\; n + max_vars a' \ length xs; \ contains_insync A \ \ ?concl A n a' xs pc stk loc\ note IH2 = \\xs. \call1 i = \(a, M', vs)\; n + max_vars i \ length xs; \ contains_insync i\ \ ?concl i n i xs 0 [] xs\ note IH3 = \\xs. \call1 e = \(a, M', vs)\; n + max_vars e \ length xs; \ contains_insync e\ \ ?concl e n e xs 0 [] xs\ note call = \call1 (a'\i\ := e) = \(a, M', vs)\\ note len = \n + max_vars (a'\i\ := e) \ length xs\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note bisim2 = \P,i,h \ (i, loc) \ ([], loc, 0, None)\ note cs = \\ contains_insync (A\i\ := e)\ show ?case proof(cases "is_val a'") case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "\Exec_mover_a P t A h (stk, loc, pc, None) ([v], loc, length (compE2 A), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence exec: "\Exec_mover_a P t (A\i\ := e) h (stk, loc, pc, None) ([v], loc, length (compE2 A), None)" by-(rule AAss_\ExecrI1) show ?thesis proof(cases "is_val i") case True then obtain v' where [simp]: "i = Val v'" by auto note exec also from bisim2 have "\Exec_mover_a P t i h ([], loc, 0, None) ([v'], loc, length (compE2 i), None)" by(auto dest!: bisim1Val2D1) from AAss_\ExecrI2[OF this, of A e v] have "\Exec_mover_a P t (A\i\ := e) h ([v], loc, length (compE2 A), None) ([v', v], loc, length (compE2 A) + length (compE2 i), None)" by simp also (rtranclp_trans) from call IH3[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e h ([], loc, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e)" and bisim': "P,e,h \ (e, loc) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from AAss_\ExecrI3[OF exec, of A i v' v] have "\Exec_mover_a P t (A\i\ := e) h ([v', v], loc, length (compE2 A) + length (compE2 i), None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,A\i\ := e,h \ (Val v\Val v'\ := e, xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by - (rule bisim1AAss3, simp) ultimately show ?thesis using ins by fastforce next case False note exec also from False call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t i h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 i ! pc' = Invoke M' (length vs)" "pc' < length (compE2 i)" and bisim': "P,i,h \ (i, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from AAss_\ExecrI2[OF exec, of A e v] have "\Exec_mover_a P t (A\i\ := e) h ([v], loc, length (compE2 A), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 A) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,A\i\ := e,h \ (Val v\i\ := e, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 A) + pc', None)" by(rule bisim1AAss2) ultimately show ?thesis using ins False by(fastforce intro!: exI) qed next case False with IH1 len False call cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1AAss1 elim!: AAss_\ExecrI1 intro!: exI) qed next case (bisim1AAss2 i n i' xs stk loc pc A e v) note IH2 = \\call1 i' = \(a, M', vs)\; n + max_vars i' \ length xs; \ contains_insync i\ \ ?concl i n i' xs pc stk loc\ note IH3 = \\xs. \call1 e = \(a, M', vs)\; n + max_vars e \ length xs; \ contains_insync e\ \ ?concl e n e xs 0 [] xs\ note call = \call1 (Val v\i'\ := e) = \(a, M', vs)\\ note len = \n + max_vars (Val v\i'\ := e) \ length xs\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, None)\ note cs = \\ contains_insync (A\i\ := e)\ show ?case proof(cases "is_val i'") case True then obtain v' where [simp]: "i' = Val v'" by auto from bisim2 have exec: "\Exec_mover_a P t i h (stk, loc, pc, None) ([v'], loc, length (compE2 i), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) from AAss_\ExecrI2[OF exec, of A e v] have "\Exec_mover_a P t (A\i\ := e) h (stk @ [v], loc, length (compE2 A) + pc, None) ([v', v], loc, length (compE2 A) + length (compE2 i), None)" by simp also from call IH3[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e)" and bisim': "P,e,h \ (e, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from AAss_\ExecrI3[OF exec, of A i v' v] have "\Exec_mover_a P t (A\i\ := e) h ([v', v], loc, length (compE2 A) + length (compE2 i), None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,A\i\ := e,h \ (Val v\Val v'\ := e, xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by(rule bisim1AAss3) ultimately show ?thesis using ins by(fastforce intro!: exI) next case False with IH2 len call cs obtain pc' loc' stk' where ins: "pc' < length (compE2 i)" "compE2 i ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t i h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,i,h \ (i', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from bisim' have "P,A\i\ := e,h \ (Val v\i'\ := e, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 A) + pc', None)" by(rule bisim1_bisims1.bisim1AAss2) with AAss_\ExecrI2[OF exec, of A e v] ins False show ?thesis by(auto intro!: exI) qed next case (bisim1AAss3 e n e' xs stk loc pc A i v v') then obtain pc' loc' stk' where pc': "pc' < length (compE2 e)" "compE2 e ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from exec have "\Exec_mover_a P t (A\i\:=e) h (stk @ [v', v], loc, length (compE2 A) + length (compE2 i) + pc, None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by(rule AAss_\ExecrI3) moreover from bisim' have "P,A\i\ := e,h \ (Val v\Val v'\ := e', xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 A) + length (compE2 i) + pc', None)" by(rule bisim1_bisims1.bisim1AAss3) ultimately show ?case using pc' by(fastforce intro!: exI) next case bisim1AAss4 thus ?case by simp next case bisim1ALength thus ?case by(auto)(fastforce intro: bisim1_bisims1.bisim1ALength elim!: ALength_\ExecrI intro!: exI) next case bisim1FAcc thus ?case by(auto)(fastforce intro: bisim1_bisims1.bisim1FAcc elim!: FAcc_\ExecrI intro!: exI) next case (bisim1FAss1 e n e' xs stk loc pc e2 F D) note IH1 = \\call1 e' = \(a, M', vs)\; n + max_vars e' \ length xs; \ contains_insync e\ \ ?concl e n e' xs pc stk loc\ note IH2 = \\xs. \call1 e2 = \(a, M', vs)\; n + max_vars e2 \ length xs; \ contains_insync e2\ \ ?concl e2 n e2 xs 0 [] xs\ note call = \call1 (e'\F{D} := e2) = \(a, M', vs)\\ note len = \n + max_vars (e'\F{D} := e2) \ length xs\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note cs = \\ contains_insync (e\F{D} := e2)\ show ?case proof(cases "is_val e'") case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "\Exec_mover_a P t e h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_mover_a P t (e\F{D} := e2) h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" by-(rule FAss_\ExecrI1) also from call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e2 h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e2 ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e2)" and bisim': "P,e2,h \ (e2, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from FAss_\ExecrI2[OF exec, of e F D v] have "\Exec_mover_a P t (e\F{D} := e2) h ([v], loc, length (compE2 e), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 e) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,e\F{D} := e2,h \ (Val v\F{D} := e2, xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e) + pc', None)" by(rule bisim1FAss2) ultimately show ?thesis using ins by fastforce next case False with IH1 len False call cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1FAss1 elim!: FAss_\ExecrI1 intro!: exI) qed next case (bisim1FAss2 e2 n e' xs stk loc pc e F D v) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e2,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from exec have "\Exec_mover_a P t (e\F{D} := e2) h (stk @ [v], loc, length (compE2 e) + pc, None) ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e) + pc', None)" by(rule FAss_\ExecrI2) moreover from bisim' have "P,e\F{D} := e2,h \ (Val v\F{D} := e', xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e) + pc', None)" by(rule bisim1_bisims1.bisim1FAss2) ultimately show ?case using pc' by(fastforce) next case bisim1FAss3 thus ?case by simp next case (bisim1CAS1 e1 n e' xs stk loc pc e2 e3 D F) note IH1 = \\call1 e' = \(a, M', vs)\; n + max_vars e' \ length xs; \ contains_insync e1 \ \ ?concl e1 n e' xs pc stk loc\ note IH2 = \\xs. \call1 e2 = \(a, M', vs)\; n + max_vars e2 \ length xs; \ contains_insync e2\ \ ?concl e2 n e2 xs 0 [] xs\ note IH3 = \\xs. \call1 e3 = \(a, M', vs)\; n + max_vars e3 \ length xs; \ contains_insync e3\ \ ?concl e3 n e3 xs 0 [] xs\ note call = \call1 _ = \(a, M', vs)\\ note len = \n + max_vars _ \ length xs\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note bisim2 = \P,e2,h \ (e2, loc) \ ([], loc, 0, None)\ note cs = \\ contains_insync _\ show ?case proof(cases "is_val e'") case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "\Exec_mover_a P t e1 h (stk, loc, pc, None) ([v], loc, length (compE2 e1), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence exec: "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h (stk, loc, pc, None) ([v], loc, length (compE2 e1), None)" by-(rule CAS_\ExecrI1) show ?thesis proof(cases "is_val e2") case True then obtain v' where [simp]: "e2 = Val v'" by auto note exec also from bisim2 have "\Exec_mover_a P t e2 h ([], loc, 0, None) ([v'], loc, length (compE2 e2), None)" by(auto dest!: bisim1Val2D1) from CAS_\ExecrI2[OF this, of e1 D F e3] have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h ([v], loc, length (compE2 e1), None) ([v', v], loc, length (compE2 e1) + length (compE2 e2), None)" by simp also (rtranclp_trans) from call IH3[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e3 h ([], loc, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e3 ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e3)" and bisim': "P,e3,h \ (e3, loc) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from CAS_\ExecrI3[OF exec, of e1 D F e2 v' v] have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h ([v', v], loc, length (compE2 e1) + length (compE2 e2), None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,e1\compareAndSwap(D\F, e2, e3),h \ (Val v\compareAndSwap(D\F, Val v', e3), xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by - (rule bisim1CAS3, simp) ultimately show ?thesis using ins by fastforce next case False note exec also from False call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e2 h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e2 ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e2)" and bisim': "P,e2,h \ (e2, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from CAS_\ExecrI2[OF exec, of e1 D F e3 v] have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h ([v], loc, length (compE2 e1), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 e1) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,e1\compareAndSwap(D\F, e2, e3),h \ (Val v\compareAndSwap(D\F, e2, e3), xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e1) + pc', None)" by(rule bisim1CAS2) ultimately show ?thesis using ins False by(fastforce intro!: exI) qed next case False with IH1 len False call cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1CAS1 elim!: CAS_\ExecrI1 intro!: exI) qed next case (bisim1CAS2 e2 n e2' xs stk loc pc e1 e3 D F v) note IH2 = \\call1 e2' = \(a, M', vs)\; n + max_vars e2' \ length xs; \ contains_insync e2\ \ ?concl e2 n e2' xs pc stk loc\ note IH3 = \\xs. \call1 e3 = \(a, M', vs)\; n + max_vars e3 \ length xs; \ contains_insync e3\ \ ?concl e3 n e3 xs 0 [] xs\ note call = \call1 _ = \(a, M', vs)\\ note len = \n + max_vars _ \ length xs\ note bisim2 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, None)\ note cs = \\ contains_insync _\ show ?case proof(cases "is_val e2'") case True then obtain v' where [simp]: "e2' = Val v'" by auto from bisim2 have exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) ([v'], loc, length (compE2 e2), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) from CAS_\ExecrI2[OF exec, of e1 D F e3 v] have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h (stk @ [v], loc, length (compE2 e1) + pc, None) ([v', v], loc, length (compE2 e1) + length (compE2 e2), None)" by simp also from call IH3[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_mover_a P t e3 h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compE2 e3 ! pc' = Invoke M' (length vs)" "pc' < length (compE2 e3)" and bisim': "P,e3,h \ (e3, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from CAS_\ExecrI3[OF exec, of e1 D F e2 v' v] have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h ([v', v], loc, length (compE2 e1) + length (compE2 e2), None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,e1\compareAndSwap(D\F, e2, e3),h \ (Val v\compareAndSwap(D\F, Val v', e3), xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by(rule bisim1CAS3) ultimately show ?thesis using ins by(fastforce intro!: exI) next case False with IH2 len call cs obtain pc' loc' stk' where ins: "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e2,h \ (e2', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from bisim' have "P,e1\compareAndSwap(D\F, e2, e3),h \ (Val v\compareAndSwap(D\F, e2', e3), xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 e1) + pc', None)" by(rule bisim1_bisims1.bisim1CAS2) with CAS_\ExecrI2[OF exec, of e1 D F e3 v] ins False show ?thesis by(auto intro!: exI) qed next case (bisim1CAS3 e3 n e3' xs stk loc pc e1 e2 D F v v') then obtain pc' loc' stk' where pc': "pc' < length (compE2 e3)" "compE2 e3 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e3 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e3,h \ (e3', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by fastforce from exec have "\Exec_mover_a P t (e1\compareAndSwap(D\F, e2, e3)) h (stk @ [v', v], loc, length (compE2 e1) + length (compE2 e2) + pc, None) ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by(rule CAS_\ExecrI3) moreover from bisim' have "P,e1\compareAndSwap(D\F, e2, e3),h \ (Val v\compareAndSwap(D\F, Val v', e3'), xs) \ ((rev vs @ Addr a # stk') @ [v', v], loc', length (compE2 e1) + length (compE2 e2) + pc', None)" by(rule bisim1_bisims1.bisim1CAS3) ultimately show ?case using pc' by(fastforce intro!: exI) next case (bisim1Call1 obj n obj' xs stk loc pc ps M) note IH1 = \\call1 obj' = \(a, M', vs)\; n + max_vars obj' \ length xs; \ contains_insync obj\ \ ?concl obj n obj' xs pc stk loc\ note IH2 = \\xs. \calls1 ps = \(a, M', vs)\; n + max_varss ps \ length xs; \ contains_insyncs ps\ \ ?concls ps n ps xs 0 [] xs\ note len = \n + max_vars (obj'\M(ps)) \ length xs\ note bisim1 = \P,obj,h \ (obj', xs) \ (stk, loc, pc, None)\ note call = \call1 (obj'\M(ps)) = \(a, M', vs)\\ note cs = \\ contains_insync (obj\M(ps))\ from call show ?case proof(cases rule: call1_callE) case CallObj hence "\ is_val obj'" by auto with CallObj IH1 len cs show ?thesis by(clarsimp)(fastforce intro: bisim1_bisims1.bisim1Call1 elim!: Call_\ExecrI1 intro!: exI) next case (CallParams v) with bisim1 have "\Exec_mover_a P t obj h (stk, loc, pc, None) ([v], loc, length (compE2 obj), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_mover_a P t (obj\M(ps)) h (stk, loc, pc, None) ([v], loc, length (compE2 obj), None)" by-(rule Call_\ExecrI1) also from IH2[of loc] CallParams len cs obtain pc' stk' loc' where exec: "\Exec_movesr_a P t ps h ([], loc, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compEs2 ps ! pc' = Invoke M' (length vs)" "pc' < length (compEs2 ps)" and bisim': "P,ps,h \ (ps, xs) [\] (rev vs @ Addr a # stk',loc',pc',None)" by auto from Call_\ExecrI2[OF exec, of obj M v] have "\Exec_mover_a P t (obj\M(ps)) h ([v], loc, length (compE2 obj), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 obj) + pc', None)" by simp also (rtranclp_trans) have "P,obj\M(ps),h \ (Val v\M(ps), xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 obj) + pc', None)" using bisim' by(rule bisim1CallParams) ultimately show ?thesis using ins CallParams by fastforce next case [simp]: Call from bisim1 have "\Exec_mover_a P t obj h (stk, loc, pc, None) ([Addr a], loc, length (compE2 obj), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_mover_a P t (obj\M(ps)) h (stk, loc, pc, None) ([Addr a], loc, length (compE2 obj), None)" by-(rule Call_\ExecrI1) also have "\Exec_movesr_a P t ps h ([], xs, 0, None) (rev vs, xs, length (compEs2 ps), None)" proof(cases vs) case Nil with Call show ?thesis by(auto) next case Cons with Call bisims1_Val_\Exec_moves[OF bisims1_refl[of P h "map Val vs" loc]] show ?thesis by(auto simp add: bsoks_def) qed from Call_\ExecrI2[OF this, of obj M "Addr a"] have "\Exec_mover_a P t (obj\M(ps)) h ([Addr a], loc, length (compE2 obj), None) (rev vs @ [Addr a], xs, length (compE2 obj) + length (compEs2 ps), None)" by simp also (rtranclp_trans) have "P,ps,h \ (map Val vs,xs) [\] (rev vs,xs,length (compEs2 ps),None)" by(rule bisims1_map_Val_append[OF bisims1Nil, simplified])(simp_all add: bsoks_def) hence "P,obj\M(ps),h \ (addr a\M(map Val vs), xs) \ (rev vs @ [Addr a], xs, length (compE2 obj) + length (compEs2 ps), None)" by(rule bisim1CallParams) ultimately show ?thesis by fastforce qed next case (bisim1CallParams ps n ps' xs stk loc pc obj M v) note IH2 = \\calls1 ps' = \(a, M', vs)\; n + max_varss ps' \ length xs; \ contains_insyncs ps\ \ ?concls ps n ps' xs pc stk loc\ note bisim2 = \P,ps,h \ (ps', xs) [\] (stk, loc, pc, None)\ note call = \call1 (Val v\M(ps')) = \(a, M', vs)\\ note len = \n + max_vars (Val v\M(ps')) \ length xs\ note cs = \\ contains_insync (obj\M(ps))\ from call show ?case proof(cases rule: call1_callE) case CallObj thus ?thesis by simp next case (CallParams v') with IH2 len cs obtain pc' stk' loc' where exec: "\Exec_movesr_a P t ps h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "pc' < length (compEs2 ps)" "compEs2 ps ! pc' = Invoke M' (length vs)" and bisim': "P,ps,h \ (ps', xs) [\] (rev vs @ Addr a # stk',loc',pc',None)" by auto from exec have "\Exec_mover_a P t (obj\M(ps)) h (stk @ [v], loc, length (compE2 obj) + pc, None) ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 obj) + pc', None)" by(rule Call_\ExecrI2) moreover have "P,obj\M(ps),h \ (Val v\M(ps'), xs) \ ((rev vs @ Addr a # stk') @ [v], loc', length (compE2 obj) + pc', None)" using bisim' by(rule bisim1_bisims1.bisim1CallParams) ultimately show ?thesis using ins by fastforce next case Call hence [simp]: "v = Addr a" "ps' = map Val vs" "M' = M" by simp_all have "xs = loc \ \Exec_movesr_a P t ps h (stk, loc, pc, None) (rev vs, loc, length (compEs2 ps), None)" proof(cases "pc < length (compEs2 ps)") case True with bisim2 show ?thesis by(auto dest: bisims1_Val_\Exec_moves) next case False from bisim2 have "pc \ length (compEs2 ps)" by(rule bisims1_pc_length_compEs2) with False have "pc = length (compEs2 ps)" by simp with bisim2 show ?thesis by(auto dest: bisims1_Val_length_compEs2D) qed then obtain [simp]: "xs = loc" and exec: "\Exec_movesr_a P t ps h (stk, loc, pc, None) (rev vs, loc, length (compEs2 ps), None)" .. from exec have "\Exec_mover_a P t (obj\M(ps)) h (stk @ [v], loc, length (compE2 obj) + pc, None) (rev vs @ [v], loc, length (compE2 obj) + length (compEs2 ps), None)" by(rule Call_\ExecrI2) moreover from bisim2 have len: "length ps = length ps'" by(auto dest: bisims1_lengthD) moreover have "P,ps,h \ (map Val vs,xs) [\] (rev vs,xs,length (compEs2 ps),None)" using len by-(rule bisims1_map_Val_append[OF bisims1Nil, simplified], simp_all) hence "P,obj\M(ps),h \ (addr a\M(map Val vs), xs) \ (rev vs @ [Addr a], xs, length (compE2 obj) + length (compEs2 ps), None)" by(rule bisim1_bisims1.bisim1CallParams) ultimately show ?thesis by fastforce qed next case bisim1BlockSome1 thus ?case by simp next case bisim1BlockSome2 thus ?case by simp next case (bisim1BlockSome4 e n e' xs stk loc pc V T v) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e)" "compE2 e ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto note Block_\ExecrI_Some[OF exec, of V T v] moreover from bisim' have "P,{V:T=\v\; e},h \ ({V:T=None; e'}, xs) \ (rev vs @ Addr a # stk', loc', Suc (Suc pc'), None)" by(rule bisim1_bisims1.bisim1BlockSome4) ultimately show ?case using pc' by fastforce next case (bisim1BlockNone e n e' xs stk loc pc V T) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e)" "compE2 e ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto note Block_\ExecrI_None[OF exec, of V T] moreover from bisim' have "P,{V:T=None; e},h \ ({V:T=None; e'}, xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by(rule bisim1_bisims1.bisim1BlockNone) ultimately show ?case using pc' by fastforce next case bisim1Sync1 thus ?case by (auto)(fastforce intro: bisim1_bisims1.bisim1Sync1 elim!: Sync_\ExecrI intro!: exI) next case bisim1Sync2 thus ?case by simp next case bisim1Sync3 thus ?case by simp next case bisim1Sync4 thus ?case by (auto)(fastforce intro: bisim1_bisims1.bisim1Sync4 elim!: Insync_\ExecrI intro!: exI) next case bisim1Sync5 thus ?case by simp next case bisim1Sync6 thus ?case by simp next case bisim1Sync7 thus ?case by simp next case bisim1Sync8 thus ?case by simp next case bisim1Sync9 thus ?case by simp next case bisim1InSync thus ?case by simp next case bisim1Seq1 thus ?case by (auto)(fastforce intro: bisim1_bisims1.bisim1Seq1 elim!: Seq_\ExecrI1 intro!: exI) next case (bisim1Seq2 e2 n e' xs stk loc pc e1) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e2,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from Seq_\ExecrI2[OF exec, of e1] pc' bisim' show ?case by(fastforce intro: bisim1_bisims1.bisim1Seq2 intro!: exI) next case bisim1Cond1 thus ?case by (auto)(fastforce intro: bisim1_bisims1.bisim1Cond1 elim!: Cond_\ExecrI1 intro!: exI)+ next case (bisim1CondThen e1 n e' xs stk loc pc e e2) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e1)" "compE2 e1 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e1 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e1,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from Cond_\ExecrI2[OF exec] pc' bisim' show ?case by(fastforce intro: bisim1_bisims1.bisim1CondThen intro!: exI) next case (bisim1CondElse e2 n e' xs stk loc pc e e1) then obtain pc' loc' stk' where pc': "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and exec: "\Exec_mover_a P t e2 h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,e2,h \ (e', xs) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from Cond_\ExecrI3[OF exec] pc' bisim' show ?case by (fastforce intro: bisim1_bisims1.bisim1CondElse intro!: exI) next case bisim1While1 thus ?case by simp next case bisim1While3 thus ?case by (auto)(fastforce intro: bisim1_bisims1.bisim1While3 elim!: While_\ExecrI1 intro!: exI)+ next case bisim1While4 thus ?case by (auto)(fastforce intro!: While_\ExecrI2 bisim1_bisims1.bisim1While4 exI)+ next case bisim1While6 thus ?case by simp next case bisim1While7 thus ?case by simp next case bisim1Throw1 thus ?case by (auto)(fastforce intro!: exI bisim1_bisims1.bisim1Throw1 elim!: Throw_\ExecrI)+ next case bisim1Try thus ?case by (auto)(fastforce intro: bisim1_bisims1.bisim1Try elim!: Try_\ExecrI1 intro!: exI)+ next case (bisim1TryCatch1 e n a' xs stk loc pc C' C e2 V) note IH2 = \\xs. \call1 e2 = \(a, M', vs)\; Suc n + max_vars e2 \ length xs; \ contains_insync e2 \ \ ?concl e2 (Suc V) e2 xs 0 [] xs\ note bisim1 = \P,e,h \ (Throw a', xs) \ (stk, loc, pc, \a'\)\ note bisim2 = \\xs. P,e2,h \ (e2, xs) \ ([], xs, 0, None)\ note len = \n + max_vars {V:Class C=None; e2} \ length (xs[V := Addr a'])\ note cs = \\ contains_insync (try e catch(C V) e2)\ from bisim1 have [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from len have "\Exec_mover_a P t (try e catch(C V) e2) h ([Addr a'], loc, Suc (length (compE2 e)), None) ([], loc[V := Addr a'], Suc (Suc (length (compE2 e))), None)" by -(rule \Execr1step,auto simp add: exec_move_def intro: \move2_\moves2.intros exec_instr) also from IH2[of "loc[V := Addr a']"] len \call1 {V:Class C=None; e2} = \(a, M', vs)\\ cs obtain pc' loc' stk' where exec: "\Exec_mover_a P t e2 h ([], loc[V := Addr a'], 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "pc' < length (compE2 e2)" "compE2 e2 ! pc' = Invoke M' (length vs)" and bisim': "P,e2,h \ (e2, loc[V := Addr a']) \ (rev vs @ Addr a # stk', loc', pc', None)" by auto from Try_\ExecrI2[OF exec, of e C V] have "\Exec_mover_a P t (try e catch(C V) e2) h ([], loc[V := Addr a'], Suc (Suc (length (compE2 e))), None) (rev vs @ Addr a # stk', loc', Suc (Suc (length (compE2 e) + pc')), None)" by simp also from bisim' have "P,try e catch(C V) e2,h \ ({V:Class C=None; e2}, loc[V := Addr a']) \ (rev vs @ Addr a # stk', loc', (Suc (Suc (length (compE2 e) + pc'))), None)" by(rule bisim1TryCatch2) ultimately show ?case using ins by fastforce next case bisim1TryCatch2 thus ?case by (auto)(fastforce intro!: Try_\ExecrI2 bisim1_bisims1.bisim1TryCatch2 exI)+ next case bisims1Nil thus ?case by simp next case (bisims1List1 e n e' xs stk loc pc es) note IH1 = \\call1 e' = \(a, M', vs)\; n + max_vars e' \ length xs; \ contains_insync e\ \ ?concl e n e' xs pc stk loc\ note IH2 = \\xs. \calls1 es = \(a, M', vs)\; n + max_varss es \ length xs; \ contains_insyncs es\ \ ?concls es n es xs 0 [] xs\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \calls1 (e' # es) = \(a, M', vs)\\ note len = \n + max_varss (e' # es) \ length xs\ note cs = \\ contains_insyncs (e # es)\ show ?case proof(cases "is_val e'") case True then obtain v where [simp]: "e' = Val v" by auto with bisim1 have "\Exec_mover_a P t e h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" and [simp]: "xs = loc" by(auto dest!: bisim1Val2D1) hence "\Exec_movesr_a P t (e # es) h (stk, loc, pc, None) ([v], loc, length (compE2 e), None)" by-(rule \Exec_mover_\Exec_movesr) also from call IH2[of loc] len cs obtain pc' stk' loc' where exec: "\Exec_movesr_a P t es h ([], xs, 0, None) (rev vs @ Addr a # stk', loc', pc', None)" and ins: "compEs2 es ! pc' = Invoke M' (length vs)" "pc' < length (compEs2 es)" and bisim': "P,es,h \ (es, xs) [\] (rev vs @ Addr a # stk',loc',pc',None)" by auto from append_\Exec_movesr[OF _ exec, of "[v]" "[e]"] have "\Exec_movesr_a P t (e # es) h ([v], loc, length (compE2 e), None) (rev vs @ Addr a # (stk' @ [v]), loc', length (compE2 e) + pc', None)" by simp also (rtranclp_trans) from bisim' have "P,e # es,h \ (Val v # es, xs) [\] ((rev vs @ Addr a # stk') @ [v],loc',length (compE2 e) + pc',None)" by(rule bisim1_bisims1.bisims1List2) ultimately show ?thesis using ins by fastforce next case False with call IH1 len cs show ?thesis by (auto)(fastforce intro!: \Exec_mover_\Exec_movesr bisim1_bisims1.bisims1List1 exI)+ qed next case (bisims1List2 es n es' xs stk loc pc e v) then obtain pc' stk' loc' where pc': "pc' < length (compEs2 es)" "compEs2 es ! pc' = Invoke M' (length vs)" and exec: "\Exec_movesr_a P t es h (stk, loc, pc, None) (rev vs @ Addr a # stk', loc', pc', None)" and bisim': "P,es,h \ (es', xs) [\] (rev vs @ Addr a # stk', loc', pc', None)" by auto note append_\Exec_movesr[OF _ exec, of "[v]" "[e]"] moreover from bisim' have "P,e#es,h \ (Val v# es', xs) [\] ((rev vs @ Addr a # stk') @ [v],loc',length (compE2 e) + pc',None)" by(rule bisim1_bisims1.bisims1List2) ultimately show ?case using pc' by fastforce qed lemma fixes P :: "'addr J1_prog" shows bisim1_inline_call_Val: "\ P,e,h \ (e', xs) \ (stk, loc, pc, None); call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0 \ \ length stk \ Suc (length vs) \ n0 = length vs \ P,e,h \ (inline_call (Val v) e', xs) \ (v # drop (Suc (length vs)) stk, loc, Suc pc, None)" (is "\ _; _; _ \ \ ?concl e n e' xs pc stk loc") and bisims1_inline_calls_Val: "\ P,es,h \ (es',xs) [\] (stk,loc,pc,None); calls1 es' = \(a, M, vs)\; compEs2 es ! pc = Invoke M n0 \ \ length stk \ Suc (length vs) \ n0 = length vs \ P,es,h \ (inline_calls (Val v) es', xs) [\] (v # drop (Suc (length vs)) stk,loc,Suc pc,None)" (is "\ _; _; _ \ \ ?concls es n es' xs pc stk loc") proof(induct "(e', xs)" "(stk, loc, pc, None :: 'addr option)" and "(es', xs)" "(stk, loc, pc, None :: 'addr option)" arbitrary: e' xs stk loc pc and es' xs stk loc pc rule: bisim1_bisims1.inducts) case bisim1Val2 thus ?case by simp next case bisim1New thus ?case by simp next case bisim1NewArray thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1NewArray) next case bisim1Cast thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Cast) next case bisim1InstanceOf thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1InstanceOf) next case bisim1Val thus ?case by simp next case bisim1Var thus ?case by simp next case (bisim1BinOp1 e1 e' xs stk loc pc bop e2) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e1 ! pc = Invoke M n0 \ \ ?concl e1 n e' xs pc stk loc\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 (e' \bop\ e2) = \(a, M, vs)\\ note ins = \compE2 (e1 \bop\ e2) ! pc = Invoke M n0\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e1)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1BinOp1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e1)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e1)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 e1)" by(cases "pc < length (compE2 e1)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1BinOp2 e2 e' xs stk loc pc e1 bop v1) note IH2 = \\call1 e' = \(a, M, vs)\; compE2 e2 ! pc = Invoke M n0\ \ ?concl e2 n e' xs pc stk loc\ note bisim2 = \P,e2,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 (Val v1 \bop\ e') = \(a, M, vs)\\ note ins = \compE2 (e1 \bop\ e2) ! (length (compE2 e1) + pc) = Invoke M n0\ from call bisim2 have pc: "pc < length (compE2 e2)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 e2 ! pc = Invoke M n0" by(simp) from IH2 ins' pc call show ?case by(auto dest: bisim1_bisims1.bisim1BinOp2) next case bisim1LAss1 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1LAss1) next case bisim1LAss2 thus ?case by simp next case (bisim1AAcc1 A a' xs stk loc pc i) note IH1 = \\call1 a' = \(a, M, vs)\; compE2 A ! pc = Invoke M n0\ \ ?concl A n a' xs pc stk loc\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note call = \call1 (a'\i\) = \(a, M, vs)\\ note ins = \compE2 (A\i\) ! pc = Invoke M n0\ show ?case proof(cases "is_val a'") case False with bisim1 call have "pc < length (compE2 A)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1AAcc1) next case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "pc \ length (compE2 A)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 A)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 A)" by(cases "pc < length (compE2 A)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1AAcc2 i i' xs stk loc pc A v) note IH2 = \\call1 i' = \(a, M, vs)\; compE2 i ! pc = Invoke M n0\ \ ?concl i n i' xs pc stk loc\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, None)\ note call = \call1 (Val v\i'\) = \(a, M, vs)\\ note ins = \compE2 (A\i\) ! (length (compE2 A) + pc) = Invoke M n0\ from call bisim2 have pc: "pc < length (compE2 i)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 i ! pc = Invoke M n0" by(simp) from IH2 ins' pc call show ?case by(auto dest: bisim1_bisims1.bisim1AAcc2) next case (bisim1AAss1 A a' xs stk loc pc i e) note IH1 = \\call1 a' = \(a, M, vs)\; compE2 A ! pc = Invoke M n0\ \ ?concl A n a' xs pc stk loc\ note bisim1 = \P,A,h \ (a', xs) \ (stk, loc, pc, None)\ note call = \call1 (a'\i\ := e) = \(a, M, vs)\\ note ins = \compE2 (A\i\ := e) ! pc = Invoke M n0\ show ?case proof(cases "is_val a'") case False with bisim1 call have "pc < length (compE2 A)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1AAss1) next case True then obtain v where [simp]: "a' = Val v" by auto from bisim1 have "pc \ length (compE2 A)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 A)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 A)" by(cases "pc < length (compE2 A)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1AAss2 i i' xs stk loc pc A e v) note IH2 = \\call1 i' = \(a, M, vs)\; compE2 i ! pc = Invoke M n0\ \ ?concl i n i' xs pc stk loc\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, None)\ note call = \call1 (Val v\i'\ := e) = \(a, M, vs)\\ note ins = \compE2 (A\i\ := e) ! (length (compE2 A) + pc) = Invoke M n0\ show ?case proof(cases "is_val i'") case False with bisim2 call have pc: "pc < length (compE2 i)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 i ! pc = Invoke M n0" by(simp) from IH2 ins' pc False call show ?thesis by(auto dest: bisim1_bisims1.bisim1AAss2) next case True then obtain v where [simp]: "i' = Val v" by auto from bisim2 have "pc \ length (compE2 i)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 i)" with bisim2 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 i)" by(cases "pc < length (compE2 i)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1AAss3 e e' xs stk loc pc i A v v') note IH2 = \\call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0\ \ ?concl e n e' xs pc stk loc\ note bisim3 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 (Val v\Val v'\ := e') = \(a, M, vs)\\ note ins = \compE2 (i\A\ := e) ! (length (compE2 i) + length (compE2 A) + pc) = Invoke M n0\ from call bisim3 have pc: "pc < length (compE2 e)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 e ! pc = Invoke M n0" by(simp) from IH2 ins' pc call show ?case by(auto dest: bisim1_bisims1.bisim1AAss3) next case bisim1AAss4 thus ?case by simp next case bisim1ALength thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1ALength) next case bisim1FAcc thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1FAcc) next case (bisim1FAss1 e1 e' xs stk loc pc F D e2) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e1 ! pc = Invoke M n0\ \ ?concl e1 n e' xs pc stk loc\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 (e'\F{D} := e2) = \(a, M, vs)\\ note ins = \compE2 (e1\F{D} := e2) ! pc = Invoke M n0\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e1)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1FAss1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e1)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e1)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 e1)" by(cases "pc < length (compE2 e1)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1FAss2 e2 e' xs stk loc pc e1 F D v1) note IH2 = \\call1 e' = \(a, M, vs)\; compE2 e2 ! pc = Invoke M n0\ \ ?concl e2 n e' xs pc stk loc\ note bisim2 = \P,e2,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 (Val v1\F{D} := e') = \(a, M, vs)\\ note ins = \compE2 (e1\F{D} := e2) ! (length (compE2 e1) + pc) = Invoke M n0\ from call bisim2 have pc: "pc < length (compE2 e2)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 e2 ! pc = Invoke M n0" by(simp) from IH2 ins' pc call show ?case by(auto dest: bisim1_bisims1.bisim1FAss2) next case bisim1FAss3 thus ?case by simp next case (bisim1CAS1 e1 e' xs stk loc pc D F e2 E3) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e1 ! pc = Invoke M n0\ \ ?concl e1 n e' xs pc stk loc\ note bisim1 = \P,e1,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \call1 _ = \(a, M, vs)\\ note ins = \compE2 _ ! pc = Invoke M n0\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e1)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisim1CAS1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e1)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e1)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 e1)" by(cases "pc < length (compE2 e1)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1CAS2 e2 e2' xs stk loc pc e1 D F e3 v) note IH2 = \\call1 e2' = \(a, M, vs)\; compE2 e2 ! pc = Invoke M n0\ \ ?concl e2 n e2' xs pc stk loc\ note bisim2 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, None)\ note call = \call1 _ = \(a, M, vs)\\ note ins = \compE2 _ ! (length (compE2 e1) + pc) = Invoke M n0\ show ?case proof(cases "is_val e2'") case False with bisim2 call have pc: "pc < length (compE2 e2)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 e2 ! pc = Invoke M n0" by(simp) from IH2 ins' pc False call show ?thesis by(auto dest: bisim1_bisims1.bisim1CAS2) next case True then obtain v where [simp]: "e2' = Val v" by auto from bisim2 have "pc \ length (compE2 e2)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e2)" with bisim2 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 e2)" by(cases "pc < length (compE2 e2)") auto with ins have False by(simp) thus ?thesis .. qed next case (bisim1CAS3 e3 e3' xs stk loc pc e1 D F e2 v v') note IH2 = \\call1 e3' = \(a, M, vs)\; compE2 e3 ! pc = Invoke M n0\ \ ?concl e3 n e3' xs pc stk loc\ note bisim3 = \P,e3,h \ (e3', xs) \ (stk, loc, pc, None)\ note call = \call1 _ = \(a, M, vs)\\ note ins = \compE2 _ ! (length (compE2 e1) + length (compE2 e2) + pc) = Invoke M n0\ from call bisim3 have pc: "pc < length (compE2 e3)" by(auto intro: bisim1_call_pcD) with ins have ins': "compE2 e3 ! pc = Invoke M n0" by(simp) from IH2 ins' pc call show ?case by(auto dest: bisim1_bisims1.bisim1CAS3) next case (bisim1Call1 obj obj' xs stk loc pc M' ps) note IH1 = \\call1 obj' = \(a, M, vs)\; compE2 obj ! pc = Invoke M n0\ \ ?concl obj n obj' xs pc stk loc\ note bisim1 = \P,obj,h \ (obj', xs) \ (stk, loc, pc, None)\ note call = \call1 (obj'\M'(ps)) = \(a, M, vs)\\ note ins = \compE2 (obj\M'(ps)) ! pc = Invoke M n0\ show ?case proof(cases "is_val obj'") case False with call bisim1 have "pc < length (compE2 obj)" by(auto intro: bisim1_call_pcD) with call False ins IH1 False show ?thesis by(auto intro: bisim1_bisims1.bisim1Call1) next case True then obtain v' where [simp]: "obj' = Val v'" by auto from bisim1 have "pc \ length (compE2 obj)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 obj)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 obj)" by(cases "pc < length (compE2 obj)") auto with ins have [simp]: "ps = []" "M' = M" by(auto simp add: nth_append split: if_split_asm)(auto simp add: neq_Nil_conv nth_append) from ins call have [simp]: "vs = []" by(auto split: if_split_asm) with bisim1 have [simp]: "stk = [v']" "xs = loc" by(auto dest: bisim1_pc_length_compE2D) from bisim1Val2[of "length (compE2 (obj\M([])))" "obj\M([])" P h v loc] call ins show ?thesis by(auto simp add: is_val_iff) qed next case (bisim1CallParams ps ps' xs stk loc pc obj M' v') note IH2 = \\calls1 ps' = \(a, M, vs)\; compEs2 ps ! pc = Invoke M n0\ \ ?concls ps n ps' xs pc stk loc\ note bisim = \P,ps,h \ (ps', xs) [\] (stk, loc, pc, None)\ note call = \call1 (Val v'\M'(ps')) = \(a, M, vs)\\ note ins = \compE2 (obj\M'(ps)) ! (length (compE2 obj) + pc) = Invoke M n0\ from call show ?case proof(cases rule: call1_callE) case CallObj thus ?thesis by simp next case (CallParams v'') hence [simp]: "v'' = v'" and call': "calls1 ps' = \(a, M, vs)\" by simp_all from bisim call' have pc: "pc < length (compEs2 ps)" by(rule bisims1_calls_pcD) with ins have ins': "compEs2 ps ! pc = Invoke M n0" by(simp add: nth_append) with IH2 call' ins pc have "P,ps,h \ (inline_calls (Val v) ps', xs) [\] (v # drop (Suc (length vs)) stk, loc, Suc pc, None)" and len: "Suc (length vs) \ length stk" and n0: "n0 = length vs" by auto hence "P,obj\M'(ps),h \ (Val v'\M'(inline_calls (Val v) ps'), xs) \ ((v # drop (Suc (length vs)) stk) @ [v'], loc, length (compE2 obj) + Suc pc, None)" by-(rule bisim1_bisims1.bisim1CallParams) thus ?thesis using call' len n0 by(auto simp add: is_vals_conv) next case Call hence [simp]: "v' = Addr a" "M' = M" "ps' = map Val vs" by auto from bisim have "pc \ length (compEs2 ps)" by(auto dest: bisims1_pc_length_compEs2) moreover { assume pc: "pc < length (compEs2 ps)" with bisim ins have False by(auto dest: bisims_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compEs2 ps)" by(cases "pc < length (compEs2 ps)") auto from bisim have [simp]: "stk = rev vs" "xs = loc" by(auto dest: bisims1_Val_length_compEs2D) hence "P,obj\M(ps),h \ (Val v, loc) \ ([v], loc, length (compE2 (obj\M(ps))), None)" by-(rule bisim1Val2, simp) moreover from bisim have "length ps = length ps'" by(rule bisims1_lengthD) ultimately show ?thesis using ins by(auto) qed next case bisim1BlockSome1 thus ?case by simp next case bisim1BlockSome2 thus ?case by simp next case bisim1BlockSome4 thus ?case by(auto intro: bisim1_bisims1.bisim1BlockSome4) next case bisim1BlockNone thus ?case by(auto intro: bisim1_bisims1.bisim1BlockNone) next case bisim1Sync1 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Sync1) next case bisim1Sync2 thus ?case by simp next case bisim1Sync3 thus ?case by simp next case bisim1Sync4 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 bisim1_bisims1.bisim1Sync4) next case bisim1Sync5 thus ?case by simp next case bisim1Sync6 thus ?case by simp next case bisim1Sync7 thus ?case by simp next case bisim1Sync8 thus ?case by simp next case bisim1Sync9 thus ?case by simp next case bisim1InSync thus ?case by(simp) next case bisim1Seq1 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Seq1) next case bisim1Seq2 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2)(fastforce dest: bisim1_bisims1.bisim1Seq2) next case bisim1Cond1 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Cond1) next case (bisim1CondThen e1 stk loc pc e e2 e' xs) thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2) (fastforce dest: bisim1_bisims1.bisim1CondThen[where e=e and ?e2.0=e2]) next case (bisim1CondElse e2 stk loc pc e e1 e' xs) thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2) (fastforce dest: bisim1_bisims1.bisim1CondElse[where e=e and ?e1.0=e1]) next case bisim1While1 thus ?case by simp next case bisim1While3 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1While3) next case bisim1While4 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2)(fastforce dest: bisim1_bisims1.bisim1While4) next case bisim1While6 thus ?case by simp next case bisim1While7 thus ?case by simp next case bisim1Throw1 thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Throw1) next case bisim1Try thus ?case by(auto split: if_split_asm dest: bisim1_pc_length_compE2 intro: bisim1_bisims1.bisim1Try) next case bisim1TryCatch1 thus ?case by simp next case bisim1TryCatch2 thus ?case by(fastforce dest: bisim1_bisims1.bisim1TryCatch2) next case bisims1Nil thus ?case by simp next case (bisims1List1 e e' xs stk loc pc es) note IH1 = \\call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M n0\ \ ?concl e n e' xs pc stk loc\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, None)\ note call = \calls1 (e' # es) = \(a, M, vs)\\ note ins = \compEs2 (e # es) ! pc = Invoke M n0\ show ?case proof(cases "is_val e'") case False with bisim1 call have "pc < length (compE2 e)" by(auto intro: bisim1_call_pcD) with call ins False IH1 show ?thesis by(auto intro: bisim1_bisims1.bisims1List1) next case True then obtain v where [simp]: "e' = Val v" by auto from bisim1 have "pc \ length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) moreover { assume pc: "pc < length (compE2 e)" with bisim1 ins have False by(auto dest: bisim_Val_pc_not_Invoke simp add: nth_append) } ultimately have [simp]: "pc = length (compE2 e)" by(cases "pc < length (compE2 e)") auto with ins call have False by(cases es)(auto) thus ?thesis .. qed next case (bisims1List2 es es' xs stk loc pc e v') note IH = \\calls1 es' = \(a, M, vs)\; compEs2 es ! pc = Invoke M n0\ \ ?concls es n es' xs pc stk loc\ note call = \calls1 (Val v' # es') = \(a, M, vs)\\ note bisim = \P,es,h \ (es', xs) [\] (stk, loc, pc, None)\ note ins = \compEs2 (e # es) ! (length (compE2 e) + pc) = Invoke M n0\ from call have call': "calls1 es' = \(a, M, vs)\" by simp with bisim have pc: "pc < length (compEs2 es)" by(rule bisims1_calls_pcD) with ins have ins': "compEs2 es ! pc = Invoke M n0" by(simp add: nth_append) from IH call ins pc show ?case by(auto split: if_split_asm dest: bisim1_bisims1.bisims1List2) qed lemma bisim1_fv: "P,e,h \ (e', xs) \ s \ fv e' \ fv e" and bisims1_fvs: "P,es,h \ (es', xs) [\] s \ fvs es' \ fvs es" apply(induct "(e', xs)" s and "(es', xs)" s arbitrary: e' xs and es' xs rule: bisim1_bisims1.inducts) apply(auto) done lemma bisim1_syncvars: "\ P,e,h \ (e', xs) \ s; syncvars e \ \ syncvars e'" and bisims1_syncvarss: "\ P,es,h \ (es', xs) [\] s; syncvarss es \ \ syncvarss es'" apply(induct "(e', xs)" s and "(es', xs)" s arbitrary: e' xs and es' xs rule: bisim1_bisims1.inducts) apply(auto dest: bisim1_fv) done declare pcs_stack_xlift [simp] lemma bisim1_Val_\red1r: "\ P, E, h \ (e, xs) \ ([v], loc, length (compE2 E), None); n + max_vars e \ length xs; \ E n \ \ \red1r P t h (e, xs) (Val v, loc)" and bisims1_Val_\Reds1r: "\ P, Es, h \ (es, xs) [\] (rev vs, loc, length (compEs2 Es), None); n + max_varss es \ length xs; \s Es n \ \ \reds1r P t h (es, xs) (map Val vs, loc)" proof(induct E n e xs stk\"[v]" loc pc\"length (compE2 E)" xcp\"None::'addr option" and Es n es xs stk\"rev vs" loc pc\"length (compEs2 Es)" xcp\"None::'addr option" arbitrary: v and vs rule: bisim1_bisims1_inducts_split) case bisim1BlockSome2 thus ?case by(simp (no_asm_use)) next case (bisim1BlockSome4 e n e' xs loc pc V T val) from \\ {V:T=\val\; e} n\ have [simp]: "n = V" and "\ e (Suc n)" by auto note len = \n + max_vars {V:T=None; e'} \ length xs\ hence V: "V < length xs" by simp from \P,e,h \ (e', xs) \ ([v], loc, pc, None)\ have lenxs: "length xs = length loc" by(auto dest: bisim1_length_xs) note IH = \\pc = length (compE2 e); Suc n + max_vars e' \ length xs; \ e (Suc n)\ \ \red1r P t h (e', xs) (Val v, loc)\ with len \Suc (Suc pc) = length (compE2 {V:T=\val\; e})\ \\ e (Suc n)\ have "\red1r P t h (e', xs) (Val v, loc)" by(simp) hence "\red1r P t h ({V:T=None; e'}, xs) ({V:T=None; Val v}, loc)" by(rule Block_None_\red1r_xt) thus ?case using V lenxs by(auto elim!: rtranclp.rtrancl_into_rtrancl intro: Red1Block \move1BlockRed) next case (bisim1BlockNone e n e' xs loc V T) from \\ {V:T=None; e} n\ have [simp]: "n = V" and "\ e (Suc n)" by auto note len = \n + max_vars {V:T=None; e'} \ length xs\ hence V: "V < length xs" by simp from \P,e,h \ (e', xs) \ ([v], loc, length (compE2 {V:T=None; e}), None)\ have lenxs: "length xs = length loc" by(auto dest: bisim1_length_xs) note IH = \\length (compE2 {V:T=None; e}) = length (compE2 e); Suc n + max_vars e' \ length xs; \ e (Suc n) \ \ \red1r P t h (e', xs) (Val v, loc)\ with len \\ e (Suc n)\ have "\red1r P t h (e', xs) (Val v, loc)" by(simp) hence "\red1r P t h ({V:T=None; e'}, xs) ({V:T=None; Val v}, loc)" by(rule Block_None_\red1r_xt) thus ?case using V lenxs by(auto elim!: rtranclp.rtrancl_into_rtrancl intro: Red1Block \move1BlockRed) next case (bisim1TryCatch2 e2 n e' xs loc pc e C V) from \\ (try e catch(C V) e2) n\ have [simp]: "n = V" and "\ e2 (Suc n)" by auto note len = \n + max_vars {V:Class C=None; e'} \ length xs\ hence V: "V < length xs" by simp from \P,e2,h \ (e', xs) \ ([v], loc, pc, None)\ have lenxs: "length xs = length loc" by(auto dest: bisim1_length_xs) note IH = \\pc = length (compE2 e2); Suc n + max_vars e' \ length xs; \ e2 (Suc n)\ \ \red1r P t h (e', xs) (Val v, loc)\ with len \Suc (Suc (length (compE2 e) + pc)) = length (compE2 (try e catch(C V) e2))\ \\ e2 (Suc n)\ have "\red1r P t h (e', xs) (Val v, loc)" by(simp) hence "\red1r P t h ({V:Class C=None; e'}, xs) ({V:Class C=None; Val v}, loc)" by(rule Block_None_\red1r_xt) thus ?case using V lenxs by(auto elim!: rtranclp.rtrancl_into_rtrancl intro: Red1Block \move1BlockRed) next case (bisims1List1 e n e' xs loc es) note bisim = \P,e,h \ (e', xs) \ (rev vs, loc, length (compEs2 (e # es)), None)\ then have es: "es = []" and pc: "length (compEs2 (e # es)) = length (compE2 e)" by(auto dest: bisim1_pc_length_compE2) with bisim obtain val where stk: "rev vs = [val]" and e': "is_val e' \ e' = Val val" by(auto dest: bisim1_pc_length_compE2D) with es pc bisims1List1 have "\red1r P t h (e', xs) (Val val, loc)" by simp with stk es show ?case by(auto intro: \red1r_inj_\reds1r) next case (bisims1List2 es n es' xs stk loc pc e v) from \stk @ [v] = rev vs\ obtain vs' where vs: "vs = v # vs'" by(cases vs) auto with bisims1List2 show ?case by(auto intro: \reds1r_cons_\reds1r) qed(fastforce dest: bisim1_pc_length_compE2 bisims1_pc_length_compEs2)+ lemma exec_meth_stk_split: "\ P,E,h \ (e, xs) \ (stk, loc, pc, xcp); exec_meth_d (compP2 P) (compE2 E) (stack_xlift (length STK) (compxE2 E 0 0)) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp') \ \ \stk''. stk' = stk'' @ STK \ exec_meth_d (compP2 P) (compE2 E) (compxE2 E 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc', xcp')" (is "\ _; ?exec E stk STK loc pc xcp stk' loc' pc' xcp' \ \ ?concl E stk STK loc pc xcp stk' loc' pc' xcp'") and exec_meth_stk_splits: "\ P,Es,h \ (es,xs) [\] (stk,loc,pc,xcp); exec_meth_d (compP2 P) (compEs2 Es) (stack_xlift (length STK) (compxEs2 Es 0 0)) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp') \ \ \stk''. stk' = stk'' @ STK \ exec_meth_d (compP2 P) (compEs2 Es) (compxEs2 Es 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc', xcp')" (is "\ _; ?execs Es stk STK loc pc xcp stk' loc' pc' xcp' \ \ ?concls Es stk STK loc pc xcp stk' loc' pc' xcp'") proof(induct E "n :: nat" e xs stk loc pc xcp and Es "n :: nat" es xs stk loc pc xcp arbitrary: stk' loc' pc' xcp' STK and stk' loc' pc' xcp' STK rule: bisim1_bisims1_inducts_split) case bisim1InSync thus ?case by(auto elim!: exec_meth.cases intro!: exec_meth.intros) next case bisim1Val2 thus ?case by(auto dest: exec_meth_length_compE2_stack_xliftD) next case bisim1New thus ?case by (fastforce elim: exec_meth.cases intro: exec_meth.intros split: if_split_asm cong del: image_cong_simp) next case bisim1NewThrow thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1NewArray e n e' xs stk loc pc xcp T) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (newA T\e\) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply simp apply (erule exec_meth.cases) apply (auto 4 4 intro: exec_meth.intros split: if_split_asm cong del: image_cong_simp) done qed next case (bisim1NewArrayThrow e n a xs stk loc pc T) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (newA T\e\) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case bisim1NewArrayFail thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1Cast e n e' xs stk loc pc xcp T) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (Cast T e) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply(simp) by(erule exec_meth.cases)(auto intro!: exec_meth.intros split: if_split_asm) qed next case (bisim1CastThrow e n a xs stk loc pc T) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (Cast T e) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case bisim1CastFail thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1InstanceOf e n e' xs stk loc pc xcp T) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (e instanceof T) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply(simp) by(erule exec_meth.cases)(auto intro!: exec_meth.intros split: if_split_asm) qed next case (bisim1InstanceOfThrow e n a xs stk loc pc T) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e instanceof T) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case bisim1Val thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case bisim1Var thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1BinOp1 e1 n e1' xs stk loc pc xcp e2 bop) note IH1 = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?exec e2 [] STK xs 0 None stk' loc' pc' xcp' \ ?concl e2 [] STK xs 0 None stk' loc' pc' xcp'\ note bisim1 = \P,e1,h \ (e1', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,e2,h \ (e2, loc) \ ([], loc, 0, None)\ note exec = \?exec (e1 \bop\ e2) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim1 have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e1)") case True with exec have "?exec e1 stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 e1)" by simp with exec have "pc' \ length (compE2 e1)" by(simp add: compxE2_size_convs stack_xlift_compxE2)(auto split: bop.splits elim!: exec_meth_drop_xt_pc) then obtain PC where PC: "pc' = PC + length (compE2 e1)" by -(rule_tac PC34="pc' - length (compE2 e1)" in that, simp) from pc bisim1 obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec pc have "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0))) t h (stk @ STK, loc, length (compE2 e1) + 0, xcp) ta h' (stk', loc', pc', xcp')" by-(rule exec_meth_take, auto) hence "?exec e2 [] (v # STK) loc 0 None stk' loc' (pc' - length (compE2 e1)) xcp'" using \stk = [v]\ \xcp = None\ by -(rule exec_meth_drop_xt, auto simp add: stack_xlift_compxE2 shift_compxE2) from IH2[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ [BinOpInstr bop]) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v]) (compxE2 e2 0 0))) t h ([] @ [v], loc, length (compE2 e1) + 0, None) ta h' (stk'' @ [v], loc', length (compE2 e1) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using \stk = [v]\ \xcp = None\ stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1BinOp2 e2 n e2' xs stk loc pc xcp e1 bop v1) note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim1 = \P,e1,h \ (e1, xs) \ ([], xs, 0, None)\ note bisim2 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec (e1 \bop\ e2) (stk @ [v1]) STK loc (length (compE2 e1) + pc) xcp stk' loc' pc' xcp'\ from bisim2 have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e2)") case True from exec have "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ [BinOpInstr bop]) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e2 stk (v1 # STK) loc pc xcp stk' loc' (pc' - length (compE2 e1)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e1), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v1]) (compxE2 e2 0 0)) t h (stk @ [v1], loc, pc, xcp) ta h' (stk'' @ [v1], loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e1) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ [BinOpInstr bop]) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e1) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 e1)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(simp add: stack_xlift_compxE2 shift_compxE2) next case False with pc have pc: "pc = length (compE2 e2)" by simp with bisim2 obtain v2 where [simp]: "stk = [v2]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec pc show ?thesis by(fastforce elim: exec_meth.cases split: sum.split_asm intro!: exec_meth.intros) qed next case (bisim1BinOpThrow1 e1 n a xs stk loc pc e2 bop) note bisim1 = \P,e1,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e1 \bop\ e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 e1)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e1 @ (compE2 e2 @ [BinOpInstr bop])) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH1[OF this] show ?case by(auto) next case (bisim1BinOpThrow2 e2 n a xs stk loc pc e1 bop v1) note bisim2 = \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e1 \bop\ e2) (stk @ [v1]) STK loc (length (compE2 e1) + pc) \a\ stk' loc' pc' xcp'\ from bisim2 have pc: "pc < length (compE2 e2)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ [BinOpInstr bop]) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, \a\) ta h' (stk', loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e2 stk (v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 e1)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 e1), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v1]) (compxE2 e2 0 0)) t h (stk @ [v1], loc, pc, \a\) ta h' (stk'' @ [v1], loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e1) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule append_exec_meth_xt)(auto) hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ [BinOpInstr bop]) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e1) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule exec_meth_append) moreover from exec' have pc': "pc' \ length (compE2 e1)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case bisim1BinOpThrow thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1LAss1 e n e' xs stk loc pc xcp V) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (V := e) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply(simp) by(erule exec_meth.cases)(auto intro!: exec_meth.intros) qed next case (bisim1LAss2 e n xs V) thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1LAssThrow e n a xs stk loc pc V) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (V := e) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case (bisim1AAcc1 a n a' xs stk loc pc xcp i) note IH1 = \\stk' loc' pc' xcp' STK. ?exec a stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl a stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?exec i [] STK xs 0 None stk' loc' pc' xcp' \ ?concl i [] STK xs 0 None stk' loc' pc' xcp'\ note bisim1 = \P,a,h \ (a', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,i,h \ (i, loc) \ ([], loc, 0, None)\ note exec = \?exec (a\i\) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim1 have pc: "pc \ length (compE2 a)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 a)") case True with exec have "?exec a stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 a)" by simp with exec have "pc' \ length (compE2 a)" by(simp add: compxE2_size_convs stack_xlift_compxE2)(auto elim!: exec_meth_drop_xt_pc) then obtain PC where PC: "pc' = PC + length (compE2 a)" by -(rule_tac PC34="pc' - length (compE2 a)" in that, simp) from pc bisim1 obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec pc have "exec_meth_d (compP2 P) (compE2 a @ compE2 i) (stack_xlift (length STK) (compxE2 a 0 0 @ compxE2 i (length (compE2 a)) (Suc 0))) t h (stk @ STK, loc, length (compE2 a) + 0, xcp) ta h' (stk', loc', pc', xcp')" by-(rule exec_meth_take, auto) hence "?exec i [] (v # STK) loc 0 None stk' loc' (pc' - length (compE2 a)) xcp'" using \stk = [v]\ \xcp = None\ by -(rule exec_meth_drop_xt, auto simp add: stack_xlift_compxE2 shift_compxE2) from IH2[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ [ALoad]) (compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v]) (compxE2 i 0 0))) t h ([] @ [v], loc, length (compE2 a) + 0, None) ta h' (stk'' @ [v], loc', length (compE2 a) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using \stk = [v]\ \xcp = None\ stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1AAcc2 i n i' xs stk loc pc xcp a v1) note IH2 = \\stk' loc' pc' xcp' STK. ?exec i stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl i stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec (a\i\) (stk @ [v1]) STK loc (length (compE2 a) + pc) xcp stk' loc' pc' xcp'\ from bisim2 have pc: "pc \ length (compE2 i)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 i)") case True from exec have "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ [ALoad]) (stack_xlift (length STK) (compxE2 a 0 0) @ shift (length (compE2 a)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 a) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 a @ compE2 i) (stack_xlift (length STK) (compxE2 a 0 0) @ shift (length (compE2 a)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 a) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length STK) (compxE2 i 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec i stk (v1 # STK) loc pc xcp stk' loc' (pc' - length (compE2 a)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 a), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v1]) (compxE2 i 0 0)) t h (stk @ [v1], loc, pc, xcp) ta h' (stk'' @ [v1], loc', pc' - length (compE2 a), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 a @ compE2 i) (compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v1]) (compxE2 i 0 0))) t h (stk @ [v1], loc, length (compE2 a) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ [ALoad]) (compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v1]) (compxE2 i 0 0))) t h (stk @ [v1], loc, length (compE2 a) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 a)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(simp add: stack_xlift_compxE2 shift_compxE2) next case False with pc have pc: "pc = length (compE2 i)" by simp with bisim2 obtain v2 where [simp]: "stk = [v2]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec pc show ?thesis by(clarsimp)(erule exec_meth.cases, auto intro!: exec_meth.intros split: if_split_asm) qed next case (bisim1AAccThrow1 A n a xs stk loc pc i) note bisim1 = \P,A,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec A stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl A stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (A\i\) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 A)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 A @ (compE2 i @ [ALoad])) (stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec A stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH1[OF this] show ?case by(auto) next case (bisim1AAccThrow2 i n a xs stk loc pc A v1) note bisim2 = \P,i,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH2 = \\stk' loc' pc' xcp' STK. ?exec i stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl i stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (A\i\) (stk @ [v1]) STK loc (length (compE2 A) + pc) \a\ stk' loc' pc' xcp'\ from bisim2 have pc: "pc < length (compE2 i)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 A @ compE2 i) @ [ALoad]) (stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 A) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 A @ compE2 i) (stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 A) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length STK) (compxE2 i 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, \a\) ta h' (stk', loc', pc' - length (compE2 A), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec i stk (v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 A)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 A), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v1]) (compxE2 i 0 0)) t h (stk @ [v1], loc, pc, \a\) ta h' (stk'' @ [v1], loc', pc' - length (compE2 A), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 A @ compE2 i) (compxE2 A 0 0 @ shift (length (compE2 A)) (stack_xlift (length [v1]) (compxE2 i 0 0))) t h (stk @ [v1], loc, length (compE2 A) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 A) + (pc' - length (compE2 A)), xcp')" by(rule append_exec_meth_xt)(auto) hence "exec_meth_d (compP2 P) ((compE2 A @ compE2 i) @ [ALoad]) (compxE2 A 0 0 @ shift (length (compE2 A)) (stack_xlift (length [v1]) (compxE2 i 0 0))) t h (stk @ [v1], loc, length (compE2 A) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 A) + (pc' - length (compE2 A)), xcp')" by(rule exec_meth_append) moreover from exec' have pc': "pc' \ length (compE2 A)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case bisim1AAccFail thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1AAss1 a n a' xs stk loc pc xcp i e) note IH1 = \\stk' loc' pc' xcp' STK. ?exec a stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl a stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?exec i [] STK xs 0 None stk' loc' pc' xcp' \ ?concl i [] STK xs 0 None stk' loc' pc' xcp'\ note bisim1 = \P,a,h \ (a', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,i,h \ (i, loc) \ ([], loc, 0, None)\ note exec = \?exec (a\i\ := e) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim1 have pc: "pc \ length (compE2 a)" by(rule bisim1_pc_length_compE2) from exec have exec': "exec_meth_d (compP2 P) (compE2 a @ compE2 i @ compE2 e @ [AStore, Push Unit]) (stack_xlift (length STK) (compxE2 a 0 0) @ shift (length (compE2 a)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0) @ compxE2 e (length (compE2 i)) (Suc (Suc 0))))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) show ?case proof(cases "pc < length (compE2 a)") case True with exec' have "?exec a stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 a)" by simp with exec' have "pc' \ length (compE2 a)" by -(erule exec_meth_drop_xt_pc, auto) then obtain PC where PC: "pc' = PC + length (compE2 a)" by -(rule_tac PC34="pc' - length (compE2 a)" in that, simp) from pc bisim1 obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec PC pc have "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e @ [AStore, Push Unit]) (stack_xlift (length STK) (compxE2 a 0 0 @ shift (length (compE2 a)) (compxE2 i 0 (Suc 0))) @ shift (length (compE2 a @ compE2 i)) (compxE2 e 0 (length STK + Suc (Suc 0)))) t h (v # STK, loc, length (compE2 a) + 0, None) ta h' (stk', loc', length (compE2 a) + PC, xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence "exec_meth_d (compP2 P) (compE2 a @ compE2 i) (stack_xlift (length STK) (compxE2 a 0 0 @ shift (length (compE2 a)) (compxE2 i 0 (Suc 0)))) t h (v # STK, loc, length (compE2 a) + 0, None) ta h' (stk', loc', length (compE2 a) + PC, xcp')" by(rule exec_meth_take_xt) simp hence "?exec i [] (v # STK) loc 0 None stk' loc' ((length (compE2 a) + PC) - length (compE2 a)) xcp'" by -(rule exec_meth_drop_xt, auto simp add: stack_xlift_compxE2 shift_compxE2) from IH2[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ (compE2 e @ [AStore, Push Unit])) ((compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v]) (compxE2 i 0 0))) @ shift (length (compE2 a @ compE2 i)) (compxE2 e 0 (Suc (Suc 0)))) t h ([] @ [v], loc, length (compE2 a) + 0, None) ta h' (stk'' @ [v], loc', length (compE2 a) + PC, xcp')" apply - apply(rule exec_meth_append_xt) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using \stk = [v]\ \xcp = None\ stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1AAss2 i n i' xs stk loc pc xcp a e v) note IH2 = \\stk' loc' pc' xcp' STK. ?exec i stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl i stk STK loc pc xcp stk' loc' pc' xcp'\ note IH3 = \\xs stk' loc' pc' xcp' STK. ?exec e [] STK xs 0 None stk' loc' pc' xcp' \ ?concl e [] STK xs 0 None stk' loc' pc' xcp'\ note bisim2 = \P,i,h \ (i', xs) \ (stk, loc, pc, xcp)\ note bisim3 = \P,e,h \ (e, loc) \ ([], loc, 0, None)\ note exec = \?exec (a\i\ := e) (stk @ [v]) STK loc (length (compE2 a) + pc) xcp stk' loc' pc' xcp'\ from bisim2 have pc: "pc \ length (compE2 i)" by(rule bisim1_pc_length_compE2) from exec have exec': "\v'. exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e @ [AStore, Push Unit]) ((compxE2 a 0 (length STK) @ shift (length (compE2 a)) (stack_xlift (length (v # STK)) (compxE2 i 0 0))) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length (v' # v # STK)) (compxE2 e 0 0))) t h (stk @ v # STK, loc, length (compE2 a) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) show ?case proof(cases "pc < length (compE2 i)") case True with exec'[of arbitrary] have exec'': "exec_meth_d (compP2 P) (compE2 a @ compE2 i) (compxE2 a 0 (length STK) @ shift (length (compE2 a)) (stack_xlift (length (v # STK)) (compxE2 i 0 0))) t h (stk @ v # STK, loc, length (compE2 a) + pc, xcp) ta h' (stk', loc', pc', xcp')" by-(erule exec_meth_take_xt, simp) hence "?exec i stk (v # STK) loc pc xcp stk' loc' (pc' - length (compE2 a)) xcp'" by(rule exec_meth_drop_xt) auto from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec''': "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 a), xcp')" by blast from exec''' have "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e @ [AStore, Push Unit]) ((compxE2 a 0 0 @ shift (length (compE2 a)) (stack_xlift (length [v]) (compxE2 i 0 0))) @ shift (length (compE2 a @ compE2 i)) (compxE2 e 0 (Suc (Suc 0)))) t h (stk @ [v], loc, length (compE2 a) + pc, xcp) ta h' (stk'' @ [v], loc', length (compE2 a) + (pc' - length (compE2 a)), xcp')" apply - apply(rule exec_meth_append_xt) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by auto moreover from exec'' have "pc' \ length (compE2 a)" by(rule exec_meth_drop_xt_pc) auto ultimately show ?thesis using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2) next case False with pc have pc: "pc = length (compE2 i)" by simp with exec'[of arbitrary] have "pc' \ length (compE2 a @ compE2 i)" by-(erule exec_meth_drop_xt_pc, auto simp add: shift_compxE2 stack_xlift_compxE2) then obtain PC where PC: "pc' = PC + length (compE2 a) + length (compE2 i)" by -(rule_tac PC34="pc' - length (compE2 a @ compE2 i)" in that, simp) from pc bisim2 obtain v' where stk: "stk = [v']" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) from exec'[of v'] have "exec_meth_d (compP2 P) (compE2 e @ [AStore, Push Unit]) (stack_xlift (length (v' # v # STK)) (compxE2 e 0 0)) t h (v' # v # STK, loc, 0, xcp) ta h' (stk', loc', pc' - length (compE2 a @ compE2 i), xcp')" unfolding stk pc append_Cons append_Nil by -(rule exec_meth_drop_xt, simp only: add_0_right length_append, auto simp add: shift_compxE2 stack_xlift_compxE2) with PC xcp have "?exec e [] (v' # v # STK) loc 0 None stk' loc' PC xcp'" by -(rule exec_meth_take,auto) from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v' # v # STK" and "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "exec_meth_d (compP2 P) (((compE2 a @ compE2 i) @ compE2 e) @ [AStore, Push Unit]) ((compxE2 a 0 0 @ compxE2 i (length (compE2 a)) (Suc 0)) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length [v', v]) (compxE2 e 0 0))) t h ([] @ [v', v], loc, length (compE2 a @ compE2 i) + 0, None) ta h' (stk'' @ [v', v], loc', length (compE2 a @ compE2 i) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by auto thus ?thesis using stk xcp stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1AAss3 e n e' xs stk loc pc xcp a i v1 v2) note IH3 = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim3 = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec (a\i\ := e) (stk @ [v2, v1]) STK loc (length (compE2 a) + length (compE2 i) + pc) xcp stk' loc' pc' xcp'\ from bisim3 have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True from exec have "exec_meth_d (compP2 P) (((compE2 a @ compE2 i) @ compE2 e) @ [AStore, Push Unit]) ((compxE2 a 0 (length STK) @ compxE2 i (length (compE2 a)) (Suc (length STK))) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 a @ compE2 i) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) hence exec': "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e) ((compxE2 a 0 (length STK) @ compxE2 i (length (compE2 a)) (Suc (length STK))) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 a @ compE2 i) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "?exec e stk (v2 # v1 # STK) loc pc xcp stk' loc' (pc' - length (compE2 a @ compE2 i)) xcp'" by(rule exec_meth_drop_xt) auto from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v2 # v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 a @ compE2 i), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e) (stack_xlift (length [v2, v1]) (compxE2 e 0 0)) t h (stk @ [v2, v1], loc, pc, xcp) ta h' (stk'' @ [v2, v1], loc', pc' - length (compE2 a @ compE2 i), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) ((compE2 a @ compE2 i) @ compE2 e) ((compxE2 a 0 0 @ compxE2 i (length (compE2 a)) (Suc 0)) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length [v2, v1]) (compxE2 e 0 0))) t h (stk @ [v2, v1], loc, length (compE2 a @ compE2 i) + pc, xcp) ta h' (stk'' @ [v2, v1], loc', length (compE2 a @ compE2 i) + (pc' - length (compE2 a @ compE2 i)), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) (((compE2 a @ compE2 i) @ compE2 e) @ [AStore, Push Unit]) ((compxE2 a 0 0 @ compxE2 i (length (compE2 a)) (Suc 0)) @ shift (length (compE2 a @ compE2 i)) (stack_xlift (length [v2, v1]) (compxE2 e 0 0))) t h (stk @ [v2, v1], loc, length (compE2 a @ compE2 i) + pc, xcp) ta h' (stk'' @ [v2, v1], loc', length (compE2 a @ compE2 i) + (pc' - length (compE2 a @ compE2 i)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 a @ compE2 i)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(simp add: stack_xlift_compxE2 shift_compxE2) next case False with pc have pc: "pc = length (compE2 e)" by simp with bisim3 obtain v3 where [simp]: "stk = [v3]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec pc show ?thesis apply(simp) by(erule exec_meth.cases)(auto intro!: exec_meth.intros split: if_split_asm) qed next case (bisim1AAssThrow1 A n a xs stk loc pc i e) note bisim1 = \P,A,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec A stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl A stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (A\i\ := e) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 A)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 A @ (compE2 i @ compE2 e @ [AStore, Push Unit])) (stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0) @ compxE2 e (length (compE2 i)) (Suc (Suc 0))))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec A stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH1[OF this] show ?case by(auto) next case (bisim1AAssThrow2 i n a xs stk loc pc A e v1) note bisim2 = \P,i,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH2 = \\stk' loc' pc' xcp' STK. ?exec i stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl i stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (A\i\ := e) (stk @ [v1]) STK loc (length (compE2 A) + pc) \a\ stk' loc' pc' xcp'\ from bisim2 have pc: "pc < length (compE2 i)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 A @ compE2 i) @ compE2 e @ [AStore, Push Unit]) ((stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) @ (shift (length (compE2 A @ compE2 i)) (compxE2 e 0 (Suc (Suc (length STK)))))) t h (stk @ v1 # STK, loc, length (compE2 A) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence exec': "exec_meth_d (compP2 P) (compE2 A @ compE2 i) (stack_xlift (length STK) (compxE2 A 0 0) @ shift (length (compE2 A)) (stack_xlift (length STK) (compxE2 i 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 A) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take_xt)(simp add: pc) hence "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length STK) (compxE2 i 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, \a\) ta h' (stk', loc', pc' - length (compE2 A), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec i stk (v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 A)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 i) (compxE2 i 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 A), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 i) (stack_xlift (length [v1]) (compxE2 i 0 0)) t h (stk @ [v1], loc, pc, \a\) ta h' (stk'' @ [v1], loc', pc' - length (compE2 A), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 A @ compE2 i) (compxE2 A 0 0 @ shift (length (compE2 A)) (stack_xlift (length [v1]) (compxE2 i 0 0))) t h (stk @ [v1], loc, length (compE2 A) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 A) + (pc' - length (compE2 A)), xcp')" by(rule append_exec_meth_xt)(auto) hence "exec_meth_d (compP2 P) ((compE2 A @ compE2 i) @ compE2 e @ [AStore, Push Unit]) ((compxE2 A 0 0 @ shift (length (compE2 A)) (stack_xlift (length [v1]) (compxE2 i 0 0))) @ (shift (length (compE2 A @ compE2 i)) (compxE2 e 0 (Suc (Suc 0))))) t h (stk @ [v1], loc, length (compE2 A) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 A) + (pc' - length (compE2 A)), xcp')" by(rule exec_meth_append_xt) moreover from exec' have pc': "pc' \ length (compE2 A)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case (bisim1AAssThrow3 e n a xs stk loc pc A i v2 v1) note bisim3 = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH3 = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (A\i\ := e) (stk @ [v2, v1]) STK loc (length (compE2 A) + length (compE2 i) + pc) \a\ stk' loc' pc' xcp'\ from bisim3 have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (((compE2 A @ compE2 i) @ compE2 e) @ [AStore, Push Unit]) ((stack_xlift (length STK) (compxE2 A 0 0 @ compxE2 i (length (compE2 A)) (Suc 0))) @ shift (length (compE2 A @ compE2 i)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 A @ compE2 i) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence exec': "exec_meth_d (compP2 P) ((compE2 A @ compE2 i) @ compE2 e) ((stack_xlift (length STK) (compxE2 A 0 0 @ compxE2 i (length (compE2 A)) (Suc 0))) @ shift (length (compE2 A @ compE2 i)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 A @ compE2 i) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "?exec e stk (v2 # v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 A @ compE2 i)) xcp'" by(rule exec_meth_drop_xt) auto from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v2 # v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 A @ compE2 i), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e) (stack_xlift (length [v2, v1]) (compxE2 e 0 0)) t h (stk @ [v2, v1], loc, pc, \a\) ta h' (stk'' @ [v2, v1], loc', pc' - length (compE2 A @ compE2 i), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) ((compE2 A @ compE2 i) @ compE2 e) ((compxE2 A 0 0 @ compxE2 i (length (compE2 A)) (Suc 0)) @ shift (length (compE2 A @ compE2 i)) (stack_xlift (length [v2, v1]) (compxE2 e 0 0))) t h (stk @ [v2, v1], loc, length (compE2 A @ compE2 i) + pc, \a\) ta h' (stk'' @ [v2, v1], loc', length (compE2 A @ compE2 i) + (pc' - length (compE2 A @ compE2 i)), xcp')" by(rule append_exec_meth_xt)(auto) hence "exec_meth_d (compP2 P) (((compE2 A @ compE2 i) @ compE2 e) @ [AStore, Push Unit]) ((compxE2 A 0 0 @ compxE2 i (length (compE2 A)) (Suc 0)) @ shift (length (compE2 A @ compE2 i)) (stack_xlift (length [v2, v1]) (compxE2 e 0 0))) t h (stk @ [v2, v1], loc, length (compE2 A @ compE2 i) + pc, \a\) ta h' (stk'' @ [v2, v1], loc', length (compE2 A @ compE2 i) + (pc' - length (compE2 A @ compE2 i)), xcp')" by(rule exec_meth_append) moreover from exec' have pc': "pc' \ length (compE2 A @ compE2 i)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case bisim1AAssFail thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case bisim1AAss4 thus ?case by -(erule exec_meth.cases, auto intro!: exec_meth.exec_instr) next case (bisim1ALength a n a' xs stk loc pc xcp) note bisim = \P,a,h \ (a', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec a stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl a stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (a\length) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 a)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 a)") case True with exec have "?exec a stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 a)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply(simp) by(erule exec_meth.cases)(auto intro!: exec_meth.intros split: if_split_asm) qed next case (bisim1ALengthThrow e n a xs stk loc pc) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e\length) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case bisim1ALengthNull thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1FAcc e n e' xs stk loc pc xcp F D) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (e\F{D}) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp with bisim obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec show ?thesis apply(simp) by(erule exec_meth.cases)(fastforce intro!: exec_meth.intros simp add: is_Ref_def split: if_split_asm)+ qed next case (bisim1FAccThrow e n a xs stk loc pc F D) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e\F{D}) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp)(erule exec_meth_take[OF _ pc]) from IH[OF this] show ?case by(auto) next case bisim1FAccNull thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1FAss1 e n e' xs stk loc pc xcp e2 F D) note IH1 = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?exec e2 [] STK xs 0 None stk' loc' pc' xcp' \ ?concl e2 [] STK xs 0 None stk' loc' pc' xcp'\ note bisim1 = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,e2,h \ (e2, loc) \ ([], loc, 0, None)\ note exec = \?exec (e\F{D} := e2) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim1 have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxE2_size_convs)(erule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 e)" by simp with exec have "pc' \ length (compE2 e)" by(simp add: compxE2_size_convs stack_xlift_compxE2)(auto elim!: exec_meth_drop_xt_pc) then obtain PC where PC: "pc' = PC + length (compE2 e)" by -(rule_tac PC34="pc' - length (compE2 e)" in that, simp) from pc bisim1 obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec pc have "exec_meth_d (compP2 P) (compE2 e @ compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0 @ compxE2 e2 (length (compE2 e)) (Suc 0))) t h (stk @ STK, loc, length (compE2 e) + 0, xcp) ta h' (stk', loc', pc', xcp')" by-(rule exec_meth_take, auto) hence "?exec e2 [] (v # STK) loc 0 None stk' loc' (pc' - length (compE2 e)) xcp'" using \stk = [v]\ \xcp = None\ by -(rule exec_meth_drop_xt, auto simp add: stack_xlift_compxE2 shift_compxE2) from IH2[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "exec_meth_d (compP2 P) ((compE2 e @ compE2 e2) @ [Putfield F D, Push Unit]) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v]) (compxE2 e2 0 0))) t h ([] @ [v], loc, length (compE2 e) + 0, None) ta h' (stk'' @ [v], loc', length (compE2 e) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using \stk = [v]\ \xcp = None\ stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1FAss2 e2 n e' xs stk loc pc xcp e F D v1) note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim2 = \P,e2,h \ (e', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec (e\F{D} := e2) (stk @ [v1]) STK loc (length (compE2 e) + pc) xcp stk' loc' pc' xcp'\ from bisim2 have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e2)") case True from exec have "exec_meth_d (compP2 P) ((compE2 e @ compE2 e2) @ [Putfield F D, Push Unit]) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 e @ compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, xcp) ta h' (stk', loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e2 stk (v1 # STK) loc pc xcp stk' loc' (pc' - length (compE2 e)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v1]) (compxE2 e2 0 0)) t h (stk @ [v1], loc, pc, xcp) ta h' (stk'' @ [v1], loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e @ compE2 e2) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) ((compE2 e @ compE2 e2) @ [Putfield F D, Push Unit]) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e) + pc, xcp) ta h' (stk'' @ [v1], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(simp add: stack_xlift_compxE2 shift_compxE2) next case False with pc have pc: "pc = length (compE2 e2)" by simp with bisim2 obtain v2 where [simp]: "stk = [v2]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec pc show ?thesis apply(simp) by(erule exec_meth.cases)(fastforce intro!: exec_meth.intros split: if_split_asm)+ qed next case (bisim1FAssThrow1 e n a xs stk loc pc e2 F D) note bisim1 = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e\F{D} := e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 e)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e @ (compE2 e2 @ [Putfield F D, Push Unit])) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH1[OF this] show ?case by(auto) next case (bisim1FAssThrow2 e2 n a xs stk loc pc e F D v1) note bisim2 = \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e\F{D} := e2) (stk @ [v1]) STK loc (length (compE2 e) + pc) \a\ stk' loc' pc' xcp'\ from bisim2 have pc: "pc < length (compE2 e2)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 e @ compE2 e2) @ [Putfield F D, Push Unit]) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (compE2 e @ compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, \a\) ta h' (stk', loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e2 stk (v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 e)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v1]) (compxE2 e2 0 0)) t h (stk @ [v1], loc, pc, \a\) ta h' (stk'' @ [v1], loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e @ compE2 e2) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule append_exec_meth_xt)(auto) hence "exec_meth_d (compP2 P) ((compE2 e @ compE2 e2) @ [Putfield F D, Push Unit]) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule exec_meth_append) moreover from exec' have pc': "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case bisim1FAssNull thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case bisim1FAss3 thus ?case by -(erule exec_meth.cases, auto intro!: exec_meth.exec_instr) next case (bisim1CAS1 e1 n e1' xs stk loc pc xcp e2 e3 D F) note IH1 = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?exec e2 [] STK xs 0 None stk' loc' pc' xcp' \ ?concl e2 [] STK xs 0 None stk' loc' pc' xcp'\ note bisim1 = \P,e1,h \ (e1', xs) \ (stk, loc, pc, xcp)\ note bisim2 = \P,e2,h \ (e2, loc) \ ([], loc, 0, None)\ note exec = \?exec _ stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim1 have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) from exec have exec': "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2 @ compE2 e3 @ [CAS F D]) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0) @ compxE2 e3 (length (compE2 e2)) (Suc (Suc 0))))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) show ?case proof(cases "pc < length (compE2 e1)") case True with exec' have "?exec e1 stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 e1)" by simp with exec' have "pc' \ length (compE2 e1)" by -(erule exec_meth_drop_xt_pc, auto) then obtain PC where PC: "pc' = PC + length (compE2 e1)" by -(rule_tac PC34="pc' - length (compE2 e1)" in that, simp) from pc bisim1 obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec PC pc have "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3 @ [CAS F D]) (stack_xlift (length STK) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (compxE2 e2 0 (Suc 0))) @ shift (length (compE2 e1 @ compE2 e2)) (compxE2 e3 0 (length STK + Suc (Suc 0)))) t h (v # STK, loc, length (compE2 e1) + 0, None) ta h' (stk', loc', length (compE2 e1) + PC, xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (compxE2 e2 0 (Suc 0)))) t h (v # STK, loc, length (compE2 e1) + 0, None) ta h' (stk', loc', length (compE2 e1) + PC, xcp')" by(rule exec_meth_take_xt) simp hence "?exec e2 [] (v # STK) loc 0 None stk' loc' ((length (compE2 e1) + PC) - length (compE2 e1)) xcp'" by -(rule exec_meth_drop_xt, auto simp add: stack_xlift_compxE2 shift_compxE2) from IH2[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ (compE2 e3 @ [CAS F D])) ((compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v]) (compxE2 e2 0 0))) @ shift (length (compE2 e1 @ compE2 e2)) (compxE2 e3 0 (Suc (Suc 0)))) t h ([] @ [v], loc, length (compE2 e1) + 0, None) ta h' (stk'' @ [v], loc', length (compE2 e1) + PC, xcp')" apply - apply(rule exec_meth_append_xt) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using \stk = [v]\ \xcp = None\ stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1CAS2 e2 n e2' xs stk loc pc xcp e1 e3 D F v) note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note IH3 = \\xs stk' loc' pc' xcp' STK. ?exec e3 [] STK xs 0 None stk' loc' pc' xcp' \ ?concl e3 [] STK xs 0 None stk' loc' pc' xcp'\ note bisim2 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, xcp)\ note bisim3 = \P,e3,h \ (e3, loc) \ ([], loc, 0, None)\ note exec = \?exec _ (stk @ [v]) STK loc (length (compE2 e1) + pc) xcp stk' loc' pc' xcp'\ from bisim2 have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) from exec have exec': "\v'. exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3 @ [CAS F D]) ((compxE2 e1 0 (length STK) @ shift (length (compE2 e1)) (stack_xlift (length (v # STK)) (compxE2 e2 0 0))) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length (v' # v # STK)) (compxE2 e3 0 0))) t h (stk @ v # STK, loc, length (compE2 e1) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) show ?case proof(cases "pc < length (compE2 e2)") case True with exec'[of undefined] have exec'': "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (compxE2 e1 0 (length STK) @ shift (length (compE2 e1)) (stack_xlift (length (v # STK)) (compxE2 e2 0 0))) t h (stk @ v # STK, loc, length (compE2 e1) + pc, xcp) ta h' (stk', loc', pc', xcp')" by-(erule exec_meth_take_xt, simp) hence "?exec e2 stk (v # STK) loc pc xcp stk' loc' (pc' - length (compE2 e1)) xcp'" by(rule exec_meth_drop_xt) auto from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec''': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e1), xcp')" by blast from exec''' have "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3 @ [CAS F D]) ((compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v]) (compxE2 e2 0 0))) @ shift (length (compE2 e1 @ compE2 e2)) (compxE2 e3 0 (Suc (Suc 0)))) t h (stk @ [v], loc, length (compE2 e1) + pc, xcp) ta h' (stk'' @ [v], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" apply - apply(rule exec_meth_append_xt) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by auto moreover from exec'' have "pc' \ length (compE2 e1)" by(rule exec_meth_drop_xt_pc) auto ultimately show ?thesis using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2) next case False with pc have pc: "pc = length (compE2 e2)" by simp with exec'[of undefined] have "pc' \ length (compE2 e1 @ compE2 e2)" by-(erule exec_meth_drop_xt_pc, auto simp add: shift_compxE2 stack_xlift_compxE2) then obtain PC where PC: "pc' = PC + length (compE2 e1) + length (compE2 e2)" by -(rule_tac PC34="pc' - length (compE2 e1 @ compE2 e2)" in that, simp) from pc bisim2 obtain v' where stk: "stk = [v']" and xcp: "xcp = None" by(auto dest: bisim1_pc_length_compE2D) from exec'[of v'] have "exec_meth_d (compP2 P) (compE2 e3 @ [CAS F D]) (stack_xlift (length (v' # v # STK)) (compxE2 e3 0 0)) t h (v' # v # STK, loc, 0, xcp) ta h' (stk', loc', pc' - length (compE2 e1 @ compE2 e2), xcp')" unfolding stk pc append_Cons append_Nil by -(rule exec_meth_drop_xt, simp only: add_0_right length_append, auto simp add: shift_compxE2 stack_xlift_compxE2) with PC xcp have "?exec e3 [] (v' # v # STK) loc 0 None stk' loc' PC xcp'" by -(rule exec_meth_take,auto) from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v' # v # STK" and "exec_meth_d (compP2 P) (compE2 e3) (compxE2 e3 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto hence "exec_meth_d (compP2 P) (((compE2 e1 @ compE2 e2) @ compE2 e3) @ [CAS F D]) ((compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0)) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length [v', v]) (compxE2 e3 0 0))) t h ([] @ [v', v], loc, length (compE2 e1 @ compE2 e2) + 0, None) ta h' (stk'' @ [v', v], loc', length (compE2 e1 @ compE2 e2) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by auto thus ?thesis using stk xcp stk' pc PC by(clarsimp simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) qed next case (bisim1CAS3 e3 n e3' xs stk loc pc xcp e1 e2 D F v1 v2) note IH3 = \\stk' loc' pc' xcp' STK. ?exec e3 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e3 stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim3 = \P,e3,h \ (e3', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec _ (stk @ [v2, v1]) STK loc (length (compE2 e1) + length (compE2 e2) + pc) xcp stk' loc' pc' xcp'\ from bisim3 have pc: "pc \ length (compE2 e3)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e3)") case True from exec have "exec_meth_d (compP2 P) (((compE2 e1 @ compE2 e2) @ compE2 e3) @ [CAS F D]) ((compxE2 e1 0 (length STK) @ compxE2 e2 (length (compE2 e1)) (Suc (length STK))) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e3 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 e1 @ compE2 e2) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) hence exec': "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3) ((compxE2 e1 0 (length STK) @ compxE2 e2 (length (compE2 e1)) (Suc (length STK))) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e3 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 e1 @ compE2 e2) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "?exec e3 stk (v2 # v1 # STK) loc pc xcp stk' loc' (pc' - length (compE2 e1 @ compE2 e2)) xcp'" by(rule exec_meth_drop_xt) auto from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v2 # v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e3) (compxE2 e3 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e1 @ compE2 e2), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e3) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0)) t h (stk @ [v2, v1], loc, pc, xcp) ta h' (stk'' @ [v2, v1], loc', pc' - length (compE2 e1 @ compE2 e2), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3) ((compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0)) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0))) t h (stk @ [v2, v1], loc, length (compE2 e1 @ compE2 e2) + pc, xcp) ta h' (stk'' @ [v2, v1], loc', length (compE2 e1 @ compE2 e2) + (pc' - length (compE2 e1 @ compE2 e2)), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) (((compE2 e1 @ compE2 e2) @ compE2 e3) @ [CAS F D]) ((compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0)) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0))) t h (stk @ [v2, v1], loc, length (compE2 e1 @ compE2 e2) + pc, xcp) ta h' (stk'' @ [v2, v1], loc', length (compE2 e1 @ compE2 e2) + (pc' - length (compE2 e1 @ compE2 e2)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 e1 @ compE2 e2)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(simp add: stack_xlift_compxE2 shift_compxE2) next case False with pc have pc: "pc = length (compE2 e3)" by simp with bisim3 obtain v3 where [simp]: "stk = [v3]" "xcp = None" by(auto dest: dest: bisim1_pc_length_compE2D) with exec pc show ?thesis apply(simp) by(erule exec_meth.cases)(fastforce intro!: exec_meth.intros split: if_split_asm)+ qed next case (bisim1CASThrow1 e1 n a xs stk loc pc e2 e3 D F) note bisim1 = \P,e1,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec _ stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 e1)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e1 @ (compE2 e2 @ compE2 e3 @ [CAS F D])) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0) @ compxE2 e3 (length (compE2 e2)) (Suc (Suc 0))))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH1[OF this] show ?case by(auto) next case (bisim1CASThrow2 e2 n a xs stk loc pc e1 e3 D F v1) note bisim2 = \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH2 = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec _ (stk @ [v1]) STK loc (length (compE2 e1) + pc) \a\ stk' loc' pc' xcp'\ from bisim2 have pc: "pc < length (compE2 e2)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3 @ [CAS F D]) ((stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) @ (shift (length (compE2 e1 @ compE2 e2)) (compxE2 e3 0 (Suc (Suc (length STK)))))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence exec': "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0)))) t h (stk @ v1 # STK, loc, length (compE2 e1) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take_xt)(simp add: pc) hence "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length STK) (compxE2 e2 0 (Suc 0))) t h (stk @ v1 # STK, loc, pc, \a\) ta h' (stk', loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e2 stk (v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 e1)) xcp'" by(simp add: compxE2_stack_xlift_convs) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 e1), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e2) (stack_xlift (length [v1]) (compxE2 e2 0 0)) t h (stk @ [v1], loc, pc, \a\) ta h' (stk'' @ [v1], loc', pc' - length (compE2 e1), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e1 @ compE2 e2) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) t h (stk @ [v1], loc, length (compE2 e1) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule append_exec_meth_xt)(auto) hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3 @ [CAS F D]) ((compxE2 e1 0 0 @ shift (length (compE2 e1)) (stack_xlift (length [v1]) (compxE2 e2 0 0))) @ (shift (length (compE2 e1 @ compE2 e2)) (compxE2 e3 0 (Suc (Suc 0))))) t h (stk @ [v1], loc, length (compE2 e1) + pc, \a\) ta h' (stk'' @ [v1], loc', length (compE2 e1) + (pc' - length (compE2 e1)), xcp')" by(rule exec_meth_append_xt) moreover from exec' have pc': "pc' \ length (compE2 e1)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case (bisim1CASThrow3 e3 n a xs stk loc pc e1 e2 D F v2 v1) note bisim3 = \P,e3,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH3 = \\stk' loc' pc' xcp' STK. ?exec e3 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e3 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec _ (stk @ [v2, v1]) STK loc (length (compE2 e1) + length (compE2 e2) + pc) \a\ stk' loc' pc' xcp'\ from bisim3 have pc: "pc < length (compE2 e3)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (((compE2 e1 @ compE2 e2) @ compE2 e3) @ [CAS F D]) ((stack_xlift (length STK) (compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0))) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e3 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 e1 @ compE2 e2) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence exec': "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3) ((stack_xlift (length STK) (compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0))) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length (v2 # v1 # STK)) (compxE2 e3 0 0))) t h (stk @ v2 # v1 # STK, loc, length (compE2 e1 @ compE2 e2) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "?exec e3 stk (v2 # v1 # STK) loc pc \a\ stk' loc' (pc' - length (compE2 e1 @ compE2 e2)) xcp'" by(rule exec_meth_drop_xt) auto from IH3[OF this] obtain stk'' where stk': "stk' = stk'' @ v2 # v1 # STK" and exec'': "exec_meth_d (compP2 P) (compE2 e3) (compxE2 e3 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 e1 @ compE2 e2), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e3) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0)) t h (stk @ [v2, v1], loc, pc, \a\) ta h' (stk'' @ [v2, v1], loc', pc' - length (compE2 e1 @ compE2 e2), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) ((compE2 e1 @ compE2 e2) @ compE2 e3) ((compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0)) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0))) t h (stk @ [v2, v1], loc, length (compE2 e1 @ compE2 e2) + pc, \a\) ta h' (stk'' @ [v2, v1], loc', length (compE2 e1 @ compE2 e2) + (pc' - length (compE2 e1 @ compE2 e2)), xcp')" by(rule append_exec_meth_xt)(auto) hence "exec_meth_d (compP2 P) (((compE2 e1 @ compE2 e2) @ compE2 e3) @ [CAS F D]) ((compxE2 e1 0 0 @ compxE2 e2 (length (compE2 e1)) (Suc 0)) @ shift (length (compE2 e1 @ compE2 e2)) (stack_xlift (length [v2, v1]) (compxE2 e3 0 0))) t h (stk @ [v2, v1], loc, length (compE2 e1 @ compE2 e2) + pc, \a\) ta h' (stk'' @ [v2, v1], loc', length (compE2 e1 @ compE2 e2) + (pc' - length (compE2 e1 @ compE2 e2)), xcp')" by(rule exec_meth_append) moreover from exec' have pc': "pc' \ length (compE2 e1 @ compE2 e2)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: stack_xlift_compxE2 shift_compxE2) next case bisim1CASFail thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case (bisim1Call1 obj n obj' xs stk loc pc xcp ps M') note bisimObj = \P,obj,h \ (obj', xs) \ (stk, loc, pc, xcp)\ note IHobj = \\stk' loc' pc' xcp' STK. ?exec obj stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl obj stk STK loc pc xcp stk' loc' pc' xcp'\ note IHparams = \\xs stk' loc' pc' xcp' STK. ?execs ps [] STK xs 0 None stk' loc' pc' xcp' \ ?concls ps [] STK xs 0 None stk' loc' pc' xcp'\ note exec = \?exec (obj\M'(ps)) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisimObj have pc: "pc \ length (compE2 obj)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 obj)") case True from exec have "?exec obj stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxEs2_size_convs)(erule exec_meth_take_xt[OF _ True]) from IHobj[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 obj)" by simp with exec have "pc' \ length (compE2 obj)" by(simp add: compxEs2_size_convs stack_xlift_compxE2)(auto elim!: exec_meth_drop_xt_pc) then obtain PC where PC: "pc' = PC + length (compE2 obj)" by -(rule_tac PC34="pc' - length (compE2 obj)" in that, simp) from pc bisimObj obtain v where [simp]: "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) show ?thesis proof(cases ps) case Cons with exec pc have "exec_meth_d (compP2 P) (compE2 obj @ compEs2 ps) (stack_xlift (length STK) (compxE2 obj 0 0 @ compxEs2 ps (length (compE2 obj)) (Suc 0))) t h (stk @ STK, loc, length (compE2 obj) + 0, xcp) ta h' (stk', loc', pc', xcp')" by -(rule exec_meth_take, auto) hence "?execs ps [] (v # STK) loc 0 None stk' loc' (pc' - length (compE2 obj)) xcp'" apply - apply(rule exec_meth_drop_xt) apply(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) apply(auto simp add: stack_xlift_compxE2) done from IHparams[OF this] PC obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec': "exec_meth_d (compP2 P) (compEs2 ps) (compxEs2 ps 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', PC, xcp')" by auto from exec' have "exec_meth_d (compP2 P) ((compE2 obj @ compEs2 ps) @ [Invoke M' (length ps)]) (compxE2 obj 0 0 @ shift (length (compE2 obj)) (stack_xlift (length [v]) (compxEs2 ps 0 0))) t h ([] @ [v], loc, length (compE2 obj) + 0, None) ta h' (stk'' @ [v], loc', length (compE2 obj) + PC, xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) by(auto) thus ?thesis using stk' PC by(clarsimp simp add: shift_compxEs2 stack_xlift_compxEs2 ac_simps) next case Nil with exec pc show ?thesis apply(auto elim!: exec_meth.cases intro!: exec_meth.intros simp add: split_beta split: if_split_asm) apply(auto split: extCallRet.split_asm intro!: exec_meth.intros) apply(force intro!: exI) apply(force intro!: exI) apply(force intro!: exI) done qed qed next case (bisim1CallParams ps n ps' xs stk loc pc xcp obj M' v) note bisimParam = \P,ps,h \ (ps',xs) [\] (stk,loc,pc,xcp)\ note IHparam = \\stk' loc' pc' xcp' STK. ?execs ps stk STK loc pc xcp stk' loc' pc' xcp' \ ?concls ps stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (obj\M'(ps)) (stk @ [v]) STK loc (length (compE2 obj) + pc) xcp stk' loc' pc' xcp'\ show ?case proof(cases ps) case Nil with bisimParam have "pc = 0" "xcp = None" by(auto elim: bisims1.cases) with exec Nil show ?thesis apply(auto elim!: exec_meth.cases intro!: exec_meth.intros simp add: split_beta extRet2JVM_def split: if_split_asm) apply(auto split: extCallRet.split_asm simp add: neq_Nil_conv) apply(force intro!: exec_meth.intros)+ done next case Cons from bisimParam have pc: "pc \ length (compEs2 ps)" by(rule bisims1_pc_length_compEs2) show ?thesis proof(cases "pc < length (compEs2 ps)") case True from exec have "exec_meth_d (compP2 P) ((compE2 obj @ compEs2 ps) @ [Invoke M' (length ps)]) (stack_xlift (length STK) (compxE2 obj 0 0) @ shift (length (compE2 obj)) (stack_xlift (length (v # STK)) (compxEs2 ps 0 0))) t h (stk @ v # STK, loc, length (compE2 obj) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) hence exec': "exec_meth_d (compP2 P) (compE2 obj @ compEs2 ps) (stack_xlift (length STK) (compxE2 obj 0 0) @ shift (length (compE2 obj)) (stack_xlift (length (v # STK)) (compxEs2 ps 0 0))) t h (stk @ v # STK, loc, length (compE2 obj) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: True) hence "?execs ps stk (v # STK) loc pc xcp stk' loc' (pc' - length (compE2 obj)) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IHparam[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec'': "exec_meth_d (compP2 P) (compEs2 ps) (compxEs2 ps 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 obj), xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compEs2 ps) (stack_xlift (length [v]) (compxEs2 ps 0 0)) t h (stk @ [v], loc, pc, xcp) ta h' (stk'' @ [v], loc', pc' - length (compE2 obj), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 obj @ compEs2 ps) (compxE2 obj 0 0 @ shift (length (compE2 obj)) (stack_xlift (length [v]) (compxEs2 ps 0 0))) t h (stk @ [v], loc, length (compE2 obj) + pc, xcp) ta h' (stk'' @ [v], loc', length (compE2 obj) + (pc' - length (compE2 obj)), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) ((compE2 obj @ compEs2 ps) @ [Invoke M' (length ps)]) (compxE2 obj 0 0 @ shift (length (compE2 obj)) (stack_xlift (length [v]) (compxEs2 ps 0 0))) t h (stk @ [v], loc, length (compE2 obj) + pc, xcp) ta h' (stk'' @ [v], loc', length (compE2 obj) + (pc' - length (compE2 obj)), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length (compE2 obj)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(auto simp add: shift_compxEs2 stack_xlift_compxEs2) next case False with pc have pc: "pc = length (compEs2 ps)" by simp with bisimParam obtain vs where "stk = vs" "length vs = length ps" "xcp = None" by(auto dest: bisims1_pc_length_compEs2D) with exec pc Cons show ?thesis apply(auto elim!: exec_meth.cases intro!: exec_meth.intros simp add: split_beta extRet2JVM_def split: if_split_asm) apply(auto simp add: neq_Nil_conv split: extCallRet.split_asm) apply(force intro!: exec_meth.intros)+ done qed qed next case (bisim1CallThrowObj obj n a xs stk loc pc ps M') note bisim = \P,obj,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec obj stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl obj stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (obj\M'(ps)) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have "pc < length (compE2 obj)" and [simp]: "xs = loc" by(auto dest: bisim1_ThrowD) with exec have "?exec obj stk STK loc pc \a\ stk' loc' pc' xcp'" by(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs)(erule exec_meth_take_xt) from IH[OF this] show ?case by auto next case (bisim1CallThrowParams ps n vs a ps' xs stk loc pc obj M' v) note bisim = \P,ps,h \ (map Val vs @ Throw a # ps',xs) [\] (stk,loc,pc,\a\)\ note IH = \\stk' loc' pc' xcp' STK. ?execs ps stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concls ps stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (obj\M'(ps)) (stk @ [v]) STK loc (length (compE2 obj) + pc) \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compEs2 ps)" "loc = xs" by(auto dest: bisims1_ThrowD) from exec have "exec_meth_d (compP2 P) ((compE2 obj @ compEs2 ps) @ [Invoke M' (length ps)]) (stack_xlift (length STK) (compxE2 obj 0 0) @ shift (length (compE2 obj)) (stack_xlift (length (v # STK)) (compxEs2 ps 0 0))) t h (stk @ v # STK, loc, length (compE2 obj) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) hence exec': "exec_meth_d (compP2 P) (compE2 obj @ compEs2 ps) (stack_xlift (length STK) (compxE2 obj 0 0) @ shift (length (compE2 obj)) (stack_xlift (length (v # STK)) (compxEs2 ps 0 0))) t h (stk @ v # STK, loc, length (compE2 obj) + pc, \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(simp add: pc) hence "?execs ps stk (v # STK) loc pc \a\ stk' loc' (pc' - length (compE2 obj)) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec'': "exec_meth_d (compP2 P) (compEs2 ps) (compxEs2 ps 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length (compE2 obj), xcp')" by auto from exec'' have "exec_meth_d (compP2 P) ((compE2 obj @ compEs2 ps) @ [Invoke M' (length ps)]) (compxE2 obj 0 0 @ shift (length (compE2 obj)) (stack_xlift (length [v]) (compxEs2 ps 0 0))) t h (stk @ [v], loc, length (compE2 obj) + pc, \a\) ta h' (stk'' @ [v], loc', length (compE2 obj) + (pc' - length (compE2 obj)), xcp')" apply - apply(rule exec_meth_append) apply(rule append_exec_meth_xt) apply(erule exec_meth_stk_offer) apply auto done moreover from exec' have "pc' \ length (compE2 obj)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) next case bisim1CallThrow thus ?case by(auto elim!: exec_meth.cases dest: match_ex_table_pcsD simp add: stack_xlift_compxEs2 stack_xlift_compxE2) next case bisim1BlockSome1 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case bisim1BlockSome2 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1BlockSome4 e n e' xs stk loc pc xcp V T v) note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec {V:T=\v\; e} stk STK loc (Suc (Suc pc)) xcp stk' loc' pc' xcp'\ hence exec': "exec_meth_d (compP2 P) ([Push v, Store V] @ compE2 e) (shift (length [Push v, Store V]) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, length [Push v, Store V] + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e stk STK loc pc xcp stk' loc' (pc' - length [Push v, Store V]) xcp'" by(rule exec_meth_drop) auto from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length [Push v, Store V], xcp')" by auto from exec'' have "exec_meth_d (compP2 P) ([Push v, Store V] @ compE2 e) (shift (length [Push v, Store V]) (compxE2 e 0 0)) t h (stk, loc, length [Push v, Store V] + pc, xcp) ta h' (stk'', loc', length [Push v, Store V] + (pc' - length [Push v, Store V]), xcp')" by(rule append_exec_meth) auto moreover from exec' have "pc' \ length [Push v, Store V]" by(rule exec_meth_drop_pc)(auto simp add: stack_xlift_compxE2) hence "Suc (Suc (pc' - Suc (Suc 0))) = pc'" by(simp) ultimately show ?case using stk' by(auto simp add: compxE2_size_convs) next case (bisim1BlockThrowSome e n a xs stk loc pc V T v) note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec {V:T=\v\; e} stk STK loc (Suc (Suc pc)) \a\ stk' loc' pc' xcp'\ hence exec': "exec_meth_d (compP2 P) ([Push v, Store V] @ compE2 e) (shift (length [Push v, Store V]) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, length [Push v, Store V] + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e stk STK loc pc \a\ stk' loc' (pc' - length [Push v, Store V]) xcp'" by(rule exec_meth_drop) auto from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length [Push v, Store V], xcp')" by auto from exec'' have "exec_meth_d (compP2 P) ([Push v, Store V] @ compE2 e) (shift (length [Push v, Store V]) (compxE2 e 0 0)) t h (stk, loc, length [Push v, Store V] + pc, \a\) ta h' (stk'', loc', length [Push v, Store V] + (pc' - length [Push v, Store V]), xcp')" by(rule append_exec_meth) auto moreover from exec' have "pc' \ length [Push v, Store V]" by(rule exec_meth_drop_pc)(auto simp add: stack_xlift_compxE2) hence "Suc (Suc (pc' - Suc (Suc 0))) = pc'" by(simp) ultimately show ?case using stk' by(auto simp add: compxE2_size_convs) next case bisim1BlockNone thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case bisim1BlockThrowNone thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1Sync1 e1 n e1' xs stk loc pc xcp e2 V) note bisim = \P,e1,h \ (e1', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e1)") case True from exec have "exec_meth_d (compP2 P) (compE2 e1 @ (Dup # Store V # MEnter # compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc])) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 3 0) @ stack_xlift (length STK) [(3, 3 + length (compE2 e2), None, 6 + length (compE2 e2), 0)])) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence "?exec e1 stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule True) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec': "exec_meth_d (compP2 P) (compE2 e1) (compxE2 e1 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc', xcp')" by blast from exec' have "exec_meth_d (compP2 P) (compE2 e1 @ (Dup # Store V # MEnter # compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc])) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (compxE2 e2 3 0 @ [(3, 3 + length (compE2 e2), None, 6 + length (compE2 e2), 0)])) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc', xcp')" by(rule exec_meth_append_xt) thus ?thesis using stk' by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) next case False with pc have [simp]: "pc = length (compE2 e1)" by simp with bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) thus ?thesis using exec by(auto elim!: exec_meth.cases intro!: exec_meth.intros) qed next case bisim1Sync2 thus ?case by(fastforce elim!: exec_meth.cases intro!: exec_meth.intros) next case bisim1Sync3 thus ?case by(fastforce elim!: exec_meth.cases intro!: exec_meth.intros split: if_split_asm) next case (bisim1Sync4 e2 n e2' xs stk loc pc xcp e1 V a) note IH = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note bisim = \P,e2,h \ (e2', xs) \ (stk, loc, pc, xcp)\ note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) stk STK loc (Suc (Suc (Suc (length (compE2 e1) + pc)))) xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e2)") case True let ?pre = "compE2 e1 @ [Dup, Store V, MEnter]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e2 0 0) @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), length STK)])) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: stack_xlift_compxE2 shift_compxE2 eval_nat_numeral ac_simps) hence exec'': "exec_meth_d (compP2 P) (compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (stack_xlift (length STK) (compxE2 e2 0 0) @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), length STK)]) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc' - length ?pre, xcp')" by(rule exec_meth_drop_xt[where n=1])(auto simp add: stack_xlift_compxE2) from exec' have pc': "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc[where n'=1])(auto simp add: stack_xlift_compxE2) hence pc'': "(Suc (Suc (Suc (pc' - Suc (Suc (Suc 0)))))) = pc'" by simp show ?thesis proof(cases xcp) case None from exec'' None True have "?exec e2 stk STK loc pc xcp stk' loc' (pc' - length ?pre) xcp'" apply - apply (erule exec_meth.cases) apply (cases "compE2 e2 ! pc") apply (fastforce simp add: is_Ref_def intro: exec_meth.intros split: if_split_asm cong del: image_cong_simp)+ done from IH[OF this] obtain stk'' where stk: "stk' = stk'' @ STK" and exec''': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by blast from exec''' have "exec_meth_d (compP2 P) (compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (compxE2 e2 0 0 @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), 0)]) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by(rule exec_meth_append_xt) hence "exec_meth_d (compP2 P) (?pre @ compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (compxE2 e1 0 0 @ shift (length ?pre) (compxE2 e2 0 0 @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), 0)])) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt[where n=1]) auto thus ?thesis using stk pc' pc'' by(simp add: eval_nat_numeral shift_compxE2 ac_simps) next case (Some a) with exec'' have [simp]: "h' = h" "xcp' = None" "loc' = loc" "ta = \" by(auto elim!: exec_meth.cases simp add: match_ex_table_append split: if_split_asm dest!: match_ex_table_stack_xliftD) show ?thesis proof(cases "match_ex_table (compP2 P) (cname_of h a) pc (compxE2 e2 0 0)") case None with Some exec'' True have [simp]: "stk' = Addr a # STK" and pc': "pc' = length (compE2 e1) + length (compE2 e2) + 6" by(auto elim!: exec_meth.cases simp add: match_ex_table_append split: if_split_asm dest!: match_ex_table_stack_xliftD) with exec'' Some None have "exec_meth_d (compP2 P) (compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (compxE2 e2 0 0 @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), 0)]) t h (stk, loc, pc, \a\) \ h (Addr a # drop (length stk - 0) stk, loc, pc' - length ?pre, None)" by -(rule exec_catch, auto elim!: exec_meth.cases simp add: match_ex_table_append matches_ex_entry_def split: if_split_asm dest!: match_ex_table_stack_xliftD) hence "exec_meth_d (compP2 P) (?pre @ compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (compxE2 e1 0 0 @ shift (length ?pre) (compxE2 e2 0 0 @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), 0)])) t h (stk, loc, length ?pre + pc, \a\) \ h (Addr a # drop (length stk - 0) stk, loc, length ?pre + (pc' - length ?pre), None)" by(rule append_exec_meth_xt[where n=1]) auto with pc' Some show ?thesis by(simp add: eval_nat_numeral shift_compxE2 ac_simps) next case (Some pcd) with \xcp = \a\\ exec'' True have "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, \a\) \ h (Addr a # drop (length stk - snd pcd) stk, loc, pc' - length ?pre, None)" apply - apply(rule exec_catch) apply(auto elim!: exec_meth.cases simp add: match_ex_table_append split: if_split_asm dest!: match_ex_table_stack_xliftD) done hence "exec_meth_d (compP2 P) (compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (compxE2 e2 0 0 @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), 0)]) t h (stk, loc, pc, \a\) \ h (Addr a # drop (length stk - snd pcd) stk, loc, pc' - length ?pre, None)" by(rule exec_meth_append_xt) hence "exec_meth_d (compP2 P) (?pre @ compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (compxE2 e1 0 0 @ shift (length ?pre) (compxE2 e2 0 0 @ [(0, length (compE2 e2), None, 3 + length (compE2 e2), 0)])) t h (stk, loc, length ?pre + pc, \a\) \ h (Addr a # drop (length stk - snd pcd) stk, loc, length ?pre + (pc' - length ?pre), None)" by(rule append_exec_meth_xt[where n=1])(auto) moreover from Some \xcp = \a\\ exec'' True pc' have "pc' = length (compE2 e1) + 3 + fst pcd" "stk' = Addr a # drop (length stk - snd pcd) stk @ STK" by(auto elim!: exec_meth.cases dest!: match_ex_table_stack_xliftD simp: match_ex_table_append split: if_split_asm) ultimately show ?thesis using \xcp = \a\\ by(auto simp add: eval_nat_numeral shift_compxE2 ac_simps) qed qed next case False with pc have [simp]: "pc = length (compE2 e2)" by simp with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros split: if_split_asm simp add: match_ex_table_append_not_pcs eval_nat_numeral)(simp_all add: matches_ex_entry_def) qed next case bisim1Sync5 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros split: if_split_asm) next case bisim1Sync6 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros split: if_split_asm) next case bisim1Sync7 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros split: if_split_asm) next case bisim1Sync8 thus ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros split: if_split_asm) next case (bisim1Sync9 e1 n e2 V a xs) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [Addr a] STK xs (8 + length (compE2 e1) + length (compE2 e2)) None stk' loc' pc' xcp'\ let ?pre = "compE2 e1 @ Dup # Store V # MEnter # compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ [ThrowExc]) (stack_xlift (length STK) (compxE2 (sync\<^bsub>V\<^esub> (e1) e2) 0 0) @ shift (length ?pre) []) t h (Addr a # STK, xs, length ?pre + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: eval_nat_numeral) hence "exec_meth_d (compP2 P) [ThrowExc] [] t h (Addr a # STK, xs, 0, None) ta h' (stk', loc', pc' - length ?pre, xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) moreover from exec' have "pc' = 8 + length (compE2 e1) + length (compE2 e2)" "stk' = Addr a # STK" by(auto elim!: exec_meth.cases) ultimately show ?case by(fastforce elim!: exec_meth.cases intro: exec_meth.intros) next case (bisim1Sync10 e1 n e2 V a xs) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [Addr a] STK xs (8 + length (compE2 e1) + length (compE2 e2)) \a\ stk' loc' pc' xcp'\ hence "match_ex_table (compP2 P) (cname_of h a) (8 + length (compE2 e1) + length (compE2 e2)) (stack_xlift (length STK) (compxE2 (sync\<^bsub>V\<^esub> (e1) e2) 0 0)) \ None" by(rule exec_meth.cases) auto hence False by(auto split: if_split_asm simp add: match_ex_table_append_not_pcs)(simp add: matches_ex_entry_def) thus ?case .. next case (bisim1Sync11 e1 n e2 V xs) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [Null] STK xs (Suc (Suc (length (compE2 e1)))) \addr_of_sys_xcpt NullPointer\ stk' loc' pc' xcp'\ hence "match_ex_table (compP2 P) (cname_of h (addr_of_sys_xcpt NullPointer)) (2 + length (compE2 e1)) (stack_xlift (length STK) (compxE2 (sync\<^bsub>V\<^esub> (e1) e2) 0 0)) \ None" by(rule exec_meth.cases)(auto split: if_split_asm) hence False by(auto split: if_split_asm simp add: match_ex_table_append_not_pcs)(simp add: matches_ex_entry_def) thus ?case .. next case (bisim1SyncThrow e1 n a xs stk loc pc e2 V) note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ note IH = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc \a\ stk' loc' pc' xcp'\ note bisim = \P,e1,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ from bisim have pc: "pc < length (compE2 e1)" and [simp]: "loc = xs" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e1 @ Dup # Store V # MEnter # compE2 e2 @ [Load V, MExit, Goto 4, Load V, MExit, ThrowExc]) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 3 0) @ [(3, 3 + length (compE2 e2), None, 6 + length (compE2 e2), length STK)])) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) hence "?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH[OF this] show ?case by auto next case (bisim1Seq1 e1 n e1' xs stk loc pc xcp e2) note bisim1 = \P,e1,h \ (e1', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (e1;;e2) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim1 have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e1)") case True from exec have "exec_meth_d (compP2 P) (compE2 e1 @ Pop # compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 (Suc 0) 0))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) hence "?exec e1 stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule True) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e1)" by simp with bisim1 obtain v where "xcp = None" "stk = [v]" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(fastforce elim: exec_meth.cases intro: exec_meth.intros) qed next case (bisim1SeqThrow1 e1 n a xs stk loc pc e2) note bisim1 = \P,e1,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (e1;;e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim1 have pc: "pc < length (compE2 e1)" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e1 @ Pop # compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 (Suc 0) 0))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: shift_compxE2 stack_xlift_compxE2) hence "?exec e1 stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH[OF this] show ?case by(fastforce elim: exec_meth.cases intro: exec_meth.intros) next case (bisim1Seq2 e2 n e2' xs stk loc pc xcp e1) note bisim2 = \P,e2,h \ (e2', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (e1;;e2) stk STK loc (Suc (length (compE2 e1) + pc)) xcp stk' loc' pc' xcp'\ from bisim2 have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e2)") case False with pc have [simp]: "pc = length (compE2 e2)" by simp from bisim2 have "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec have False by(auto elim: exec_meth.cases) thus ?thesis .. next case True from exec have exec': "exec_meth_d (compP2 P) ((compE2 e1 @ [Pop]) @ compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1 @ [Pop])) (stack_xlift (length STK) (compxE2 e2 0 0))) t h (stk @ STK, loc, length ((compE2 e1) @ [Pop]) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec e2 stk STK loc pc xcp stk' loc' (pc' - length ((compE2 e1) @ [Pop])) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e1 @ [Pop]), xcp')" by auto from exec'' have "exec_meth_d (compP2 P) ((compE2 e1 @ [Pop]) @ compE2 e2) (compxE2 e1 0 0 @ shift (length (compE2 e1 @ [Pop])) (compxE2 e2 0 0)) t h (stk, loc, length ((compE2 e1) @ [Pop]) + pc, xcp) ta h' (stk'', loc', length ((compE2 e1) @ [Pop]) + (pc' - length ((compE2 e1) @ [Pop])), xcp')" by(rule append_exec_meth_xt) auto moreover from exec' have "pc' \ length ((compE2 e1) @ [Pop])" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2) qed next case (bisim1Cond1 e n e' xs stk loc pc xcp e1 e2) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (if (e) e1 else e2) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True from exec have "exec_meth_d (compP2 P) (compE2 e @ IfFalse (2 + int (length (compE2 e1))) # compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e1 (Suc 0) 0) @ stack_xlift (length STK) (compxE2 e2 (Suc (Suc (length (compE2 e1)))) 0))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: stack_xlift_compxE2 shift_compxE2 ac_simps) hence "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule True) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros) qed next case (bisim1CondThen e1 n e1' xs stk loc pc xcp e e2) note bisim = \P,e1,h \ (e1', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e1 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e1 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (if (e) e1 else e2) stk STK loc (Suc (length (compE2 e) + pc)) xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e1)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e1)") case True let ?pre = "compE2 e @ [IfFalse (2 + int (length (compE2 e1)))]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 (Suc 0) 0)))) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: stack_xlift_compxE2 shift_compxE2 ac_simps) hence "exec_meth_d (compP2 P) (compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (stack_xlift (length STK) (compxE2 e1 0 0) @ shift (length (compE2 e1)) (stack_xlift (length STK) (compxE2 e2 (Suc 0) 0))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc' - length ?pre, xcp')" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) hence "?exec e1 stk STK loc pc xcp stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_take_xt)(rule True) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e1) (compxE2 e1 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (compxE2 e2 (Suc 0) 0)) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by(rule exec_meth_append_xt) hence "exec_meth_d (compP2 P) (?pre @ compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (compxE2 e 0 0 @ shift (length ?pre) (compxE2 e1 0 0 @ shift (length (compE2 e1)) (compxE2 e2 (Suc 0) 0))) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt)(auto) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2 ac_simps) next case False with pc have [simp]: "pc = length (compE2 e1)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros) qed next case (bisim1CondElse e2 n e2' xs stk loc pc xcp e e1) note IH = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (if (e) e1 else e2) stk STK loc (Suc (Suc (length (compE2 e) + length (compE2 e1) + pc))) xcp stk' loc' pc' xcp'\ note bisim = \P,e2,h \ (e2', xs) \ (stk, loc, pc, xcp)\ from bisim have pc: "pc \ length (compE2 e2)" by(rule bisim1_pc_length_compE2) let ?pre = "compE2 e @ IfFalse (2 + int (length (compE2 e1))) # compE2 e1 @ [Goto (1 + int (length (compE2 e2)))]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0 @ compxE2 e1 (Suc (length (compE2 e))) 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e2 0 0))) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: stack_xlift_compxE2 shift_compxE2 ac_simps) hence "?exec e2 stk STK loc pc xcp stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2 shift_compxEs2 ac_simps) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by blast from exec'' have "exec_meth_d (compP2 P) (?pre @ compE2 e2) ((compxE2 e 0 0 @ compxE2 e1 (Suc (length (compE2 e))) 0) @ shift (length ?pre) (compxE2 e2 0 0)) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt)(auto) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) moreover hence "(Suc (Suc (pc' - Suc (Suc 0)))) = pc'" by simp ultimately show ?case using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2 ac_simps eval_nat_numeral) next case (bisim1CondThrow e n a xs stk loc pc e1 e2) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (if (e) e1 else e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 e @ IfFalse (2 + int (length (compE2 e1))) # compE2 e1 @ Goto (1 + int (length (compE2 e2))) # compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e1 (Suc 0) 0) @ stack_xlift (length STK) (compxE2 e2 (Suc (Suc (length (compE2 e1)))) 0))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: stack_xlift_compxE2 shift_compxE2 ac_simps) hence "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH[OF this] show ?case by auto next case (bisim1While1 c n e xs) note IH = \\stk' loc' pc' xcp' STK. ?exec c [] STK xs 0 None stk' loc' pc' xcp' \ ?concl c [] STK xs 0 None stk' loc' pc' xcp'\ note exec = \?exec (while (c) e) [] STK xs 0 None stk' loc' pc' xcp'\ hence "exec_meth_d (compP2 P) (compE2 c @ IfFalse (int (length (compE2 e)) + 3) # compE2 e @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length (compE2 c)) (stack_xlift (length STK) (compxE2 e (Suc 0) 0))) t h ([] @ STK, xs, 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec c [] STK xs 0 None stk' loc' pc' xcp'" by(rule exec_meth_take_xt) simp from IH[OF this] show ?case by auto next case (bisim1While3 c n c' xs stk loc pc xcp e) note bisim = \P,c,h \ (c', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec c stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl c stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (while (c) e) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 c)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 c)") case True from exec have "exec_meth_d (compP2 P) (compE2 c @ IfFalse (int (length (compE2 e)) + 3) # compE2 e @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length (compE2 c)) (stack_xlift (length STK) (compxE2 e (Suc 0) 0))) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec c stk STK loc pc xcp stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule True) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 c)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros) qed next case (bisim1While4 e n e' xs stk loc pc xcp c) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (while (c) e) stk STK loc (Suc (length (compE2 c) + pc)) xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True let ?pre = "compE2 c @ [IfFalse (int (length (compE2 e)) + 3)]" from exec have "exec_meth_d (compP2 P) ((?pre @ compE2 e) @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (?pre @ compE2 e) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(auto intro: True) hence "?exec e stk STK loc pc xcp stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by auto from exec'' have "exec_meth_d (compP2 P) (?pre @ compE2 e) (compxE2 c 0 0 @ shift (length ?pre) (compxE2 e 0 0)) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) ((?pre @ compE2 e) @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (compxE2 c 0 0 @ shift (length ?pre) (compxE2 e 0 0)) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) moreover have "-2 + (- int (length (compE2 e)) - int (length (compE2 c))) = - int (length (compE2 c)) + (-2 - int (length (compE2 e)))" by simp ultimately show ?thesis using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2 algebra_simps uminus_minus_left_commute) next case False with pc have [simp]: "pc = length (compE2 e)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros) qed next case (bisim1While6 c n e xs) note exec = \?exec (while (c) e) [] STK xs (Suc (Suc (length (compE2 c) + length (compE2 e)))) None stk' loc' pc' xcp'\ thus ?case by(rule exec_meth.cases)(simp_all, auto intro!: exec_meth.intros) next case (bisim1While7 c n e xs) note exec = \?exec (while (c) e) [] STK xs (Suc (Suc (Suc (length (compE2 c) + length (compE2 e))))) None stk' loc' pc' xcp'\ thus ?case by(rule exec_meth.cases)(simp_all, auto intro!: exec_meth.intros) next case (bisim1WhileThrow1 c n a xs stk loc pc e) note bisim = \P,c,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec c stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl c stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (while (c) e) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 c)" by(auto dest: bisim1_ThrowD) from exec have "exec_meth_d (compP2 P) (compE2 c @ IfFalse (int (length (compE2 e)) + 3) # compE2 e @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length (compE2 c)) (stack_xlift (length STK) (compxE2 e (Suc 0) 0))) t h (stk @ STK, loc, pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence "?exec c stk STK loc pc \a\ stk' loc' pc' xcp'" by(rule exec_meth_take_xt)(rule pc) from IH[OF this] show ?case by auto next case (bisim1WhileThrow2 e n a xs stk loc pc c) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (while (c) e) stk STK loc (Suc (length (compE2 c) + pc)) \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) let ?pre = "compE2 c @ [IfFalse (int (length (compE2 e)) + 3)]" from exec have "exec_meth_d (compP2 P) ((?pre @ compE2 e) @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, length ?pre + pc, \a\) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) hence exec': "exec_meth_d (compP2 P) (?pre @ compE2 e) (stack_xlift (length STK) (compxE2 c 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e 0 0))) t h (stk @ STK, loc, (length ?pre + pc), \a\) ta h' (stk', loc', pc', xcp')" by(rule exec_meth_take)(auto intro: pc) hence "?exec e stk STK loc pc \a\ stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length ?pre, xcp')" by auto from exec'' have "exec_meth_d (compP2 P) (?pre @ compE2 e) (compxE2 c 0 0 @ shift (length ?pre) (compxE2 e 0 0)) t h (stk, loc, length ?pre + pc, \a\) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) ((?pre @ compE2 e) @ [Pop, Goto (-2 + (- int (length (compE2 e)) - int (length (compE2 c)))), Push Unit]) (compxE2 c 0 0 @ shift (length ?pre) (compxE2 e 0 0)) t h (stk, loc, length ?pre + pc, \a\) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule exec_meth_append) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) moreover have "-2 + (- int (length (compE2 e)) - int (length (compE2 c))) = - int (length (compE2 c)) + (-2 - int (length (compE2 e)))" by simp ultimately show ?case using stk' by(auto simp add: shift_compxE2 stack_xlift_compxE2 algebra_simps uminus_minus_left_commute) next case (bisim1Throw1 e n e' xs stk loc pc xcp) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (throw e) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" by(auto elim: exec_meth_take) from IH[OF this] show ?thesis by auto next case False with pc have [simp]: "pc = length (compE2 e)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros split: if_split_asm) qed next case bisim1Throw2 thus ?case apply(auto elim!:exec_meth.cases intro: exec_meth.intros dest!: match_ex_table_stack_xliftD) apply(auto intro: exec_meth.intros dest!: match_ex_table_stack_xliftD intro!: exI) apply(auto simp add: le_Suc_eq) done next case bisim1ThrowNull thus ?case apply(auto elim!:exec_meth.cases intro: exec_meth.intros dest!: match_ex_table_stack_xliftD) apply(auto intro: exec_meth.intros dest!: match_ex_table_stack_xliftD intro!: exI) apply(auto simp add: le_Suc_eq) done next case (bisim1ThrowThrow e n a xs stk loc pc) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (throw e) stk STK loc pc \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) with exec have "?exec e stk STK loc pc \a\ stk' loc' pc' xcp'" by(auto elim: exec_meth_take simp add: compxE2_size_convs) from IH[OF this] show ?case by auto next case (bisim1Try e n e' xs stk loc pc xcp e2 C' V) note bisim = \P,e,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (try e catch(C' V) e2) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True from exec have exec': "exec_meth_d (compP2 P) (compE2 e @ Goto (int (length (compE2 e2)) + 2) # Store V # compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length STK) (compxE2 e2 (Suc (Suc 0)) 0)) @ [(0, length (compE2 e), \C'\, Suc (length (compE2 e)), length STK)]) t h (stk @ STK, loc, pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxE2_size_convs) show ?thesis proof(cases xcp) case None with exec' True have "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" apply - apply (erule exec_meth.cases) apply (cases "compE2 e ! pc") apply (fastforce simp add: is_Ref_def intro: exec_meth.intros split: if_split_asm cong del: image_cong_simp)+ done from IH[OF this] show ?thesis by auto next case (Some a) with exec' have [simp]: "h' = h" "loc' = loc" "xcp' = None" "ta = \" by(auto elim: exec_meth.cases) show ?thesis proof(cases "match_ex_table (compP2 P) (cname_of h a) pc (compxE2 e 0 0)") case (Some pcd) from exec \xcp = \a\\ Some pc have stk': "stk' = Addr a # (drop (length stk - snd pcd) stk) @ STK" by(auto elim!: exec_meth.cases simp add: match_ex_table_append split: if_split_asm dest!: match_ex_table_stack_xliftD) from exec' \xcp = \a\\ Some pc have "exec_meth_d (compP2 P) (compE2 e) (stack_xlift (length STK) (compxE2 e 0 0)) t h (stk @ STK, loc, pc, \a\) \ h (Addr a # (drop (length (stk @ STK) - (snd pcd + length STK)) (stk @ STK)), loc, pc', None)" apply - apply(rule exec_meth.intros) apply(auto elim!: exec_meth.cases simp add: match_ex_table_append split: if_split_asm dest!: match_ex_table_shift_pcD match_ex_table_stack_xliftD) done from IH[unfolded \ta = \\ \xcp = \a\\ \h' = h\, OF this] have stk: "Addr a # drop (length stk - snd pcd) (stk @ STK) = Addr a # drop (length stk - snd pcd) stk @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e) (compxE2 e 0 0) t h (stk, loc, pc, \a\) \ h (Addr a # drop (length stk - snd pcd) stk, loc, pc', None)" by auto thus ?thesis using Some stk' \xcp = \a\\ by(auto) next case None with Some exec pc have stk': "stk' = Addr a # STK" and pc': "pc' = Suc (length (compE2 e))" and subcls: "compP2 P \ cname_of h a \\<^sup>* C'" by(auto elim!: exec_meth.cases split: if_split_asm simp add: match_ex_table_append_not_pcs)(simp add: matches_ex_entry_def) moreover from Some True None pc' subcls have "exec_meth_d (compP2 P) (compE2 (try e catch(C' V) e2)) (compxE2 (try e catch(C' V) e2) 0 0) t h (stk, loc, pc, \a\) \ h (Addr a # drop (length stk - 0) stk, loc, pc', None)" by -(rule exec_catch,auto simp add: match_ex_table_append_not_pcs matches_ex_entry_def) ultimately show ?thesis using Some by auto qed qed next case False with pc have [simp]: "pc = length (compE2 e)" by simp from bisim obtain v where "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec show ?thesis by(auto elim!: exec_meth.cases intro!: exec_meth.intros split: if_split_asm) qed next case bisim1TryCatch1 thus ?case by(auto elim!: exec_meth.cases intro!: exec_meth.intros split: if_split_asm) next case (bisim1TryCatch2 e2 n e' xs stk loc pc xcp e C' V) note bisim = \P,e2,h \ (e', xs) \ (stk, loc, pc, xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?exec (try e catch(C' V) e2) stk STK loc (Suc (Suc (length (compE2 e) + pc))) xcp stk' loc' pc' xcp'\ let ?pre = "compE2 e @ [Goto (int (length (compE2 e2)) + 2), Store V]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e2 0 0))) t h (stk @ STK, loc, length ?pre + pc, xcp) ta h' (stk', loc', pc', xcp')" proof(cases) case (exec_catch xcp'' d) let ?stk = "stk @ STK" and ?PC = "Suc (Suc (length (compE2 e) + pc))" note s = \stk' = Addr xcp'' # drop (length ?stk - d) ?stk\ \ta = \\ \h' = h\ \xcp' = None\ \xcp = \xcp''\\ \loc' = loc\ from \match_ex_table (compP2 P) (cname_of h xcp'') ?PC (stack_xlift (length STK) (compxE2 (try e catch(C' V) e2) 0 0)) = \(pc', d)\\ \d \ length ?stk\ show ?thesis unfolding s by -(rule exec_meth.exec_catch, simp_all add: shift_compxE2 stack_xlift_compxE2, simp add: match_ex_table_append add: matches_ex_entry_def) qed(auto intro: exec_meth.intros simp add: shift_compxE2 stack_xlift_compxE2) hence "?exec e2 stk STK loc pc xcp stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length ?pre, xcp')" by auto from exec'' have "exec_meth_d (compP2 P) (?pre @ compE2 e2) (compxE2 e 0 0 @ shift (length ?pre) (compxE2 e2 0 0)) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) (?pre @ compE2 e2) (compxE2 e 0 0 @ shift (length ?pre) (compxE2 e2 0 0) @ [(0, length (compE2 e), \C'\, Suc (length (compE2 e)), 0)]) t h (stk, loc, length ?pre + pc, xcp) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule exec_meth.cases)(auto intro: exec_meth.intros simp add: match_ex_table_append_not_pcs) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) moreover hence "(Suc (Suc (pc' - Suc (Suc 0)))) = pc'" by simp ultimately show ?case using stk' by(auto simp add: shift_compxE2 eval_nat_numeral) next case (bisim1TryFail e n a xs stk loc pc C' C'' e2 V) note bisim = \P,e,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note exec = \?exec (try e catch(C'' V) e2) stk STK loc pc \a\ stk' loc' pc' xcp'\ note a = \typeof_addr h a = \Class_type C'\\ \\ P \ C' \\<^sup>* C''\ from bisim have "match_ex_table (compP2 P) (cname_of h a) (0 + pc) (compxE2 e 0 0) = None" unfolding compP2_def by(rule bisim1_xcp_Some_not_caught) moreover from bisim have "pc < length (compE2 e)" by(auto dest: bisim1_ThrowD) ultimately have False using exec a apply(auto elim!: exec_meth.cases simp add: outside_pcs_compxE2_not_matches_entry outside_pcs_not_matches_entry split: if_split_asm) apply(auto simp add: compP2_def match_ex_entry match_ex_table_append_not_pcs cname_of_def split: if_split_asm) done thus ?case .. next case (bisim1TryCatchThrow e2 n a xs stk loc pc e C' V) note bisim = \P,e2,h \ (Throw a, xs) \ (stk, loc, pc, \a\)\ note IH = \\stk' loc' pc' xcp' STK. ?exec e2 stk STK loc pc \a\ stk' loc' pc' xcp' \ ?concl e2 stk STK loc pc \a\ stk' loc' pc' xcp'\ note exec = \?exec (try e catch(C' V) e2) stk STK loc (Suc (Suc (length (compE2 e) + pc))) \a\ stk' loc' pc' xcp'\ from bisim have pc: "pc < length (compE2 e2)" by(auto dest: bisim1_ThrowD) let ?pre = "compE2 e @ [Goto (int (length (compE2 e2)) + 2), Store V]" from exec have exec': "exec_meth_d (compP2 P) (?pre @ compE2 e2) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length ?pre) (stack_xlift (length STK) (compxE2 e2 0 0))) t h (stk @ STK, loc, length ?pre + pc, \a\) ta h' (stk', loc', pc', xcp')" proof(cases) case (exec_catch d) let ?stk = "stk @ STK" and ?PC = "Suc (Suc (length (compE2 e) + pc))" note s = \stk' = Addr a # drop (length ?stk - d) ?stk\ \loc' = loc\ \ta = \\ \h' = h\ \xcp' = None\ from \match_ex_table (compP2 P) (cname_of h a) ?PC (stack_xlift (length STK) (compxE2 (try e catch(C' V) e2) 0 0)) = \(pc', d)\\ \d \ length ?stk\ show ?thesis unfolding s by -(rule exec_meth.exec_catch, simp_all add: shift_compxE2 stack_xlift_compxE2, simp add: match_ex_table_append add: matches_ex_entry_def) qed hence "?exec e2 stk STK loc pc \a\ stk' loc' (pc' - length ?pre) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ STK" and exec'': "exec_meth_d (compP2 P) (compE2 e2) (compxE2 e2 0 0) t h (stk, loc, pc, \a\) ta h' (stk'', loc', pc' - length ?pre, xcp')" by auto from exec'' have "exec_meth_d (compP2 P) (?pre @ compE2 e2) (compxE2 e 0 0 @ shift (length ?pre) (compxE2 e2 0 0)) t h (stk, loc, length ?pre + pc, \a\) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule append_exec_meth_xt) auto hence "exec_meth_d (compP2 P) (?pre @ compE2 e2) (compxE2 e 0 0 @ shift (length ?pre) (compxE2 e2 0 0) @ [(0, length (compE2 e), \C'\, Suc (length (compE2 e)), 0)]) t h (stk, loc, length ?pre + pc, \a\) ta h' (stk'', loc', length ?pre + (pc' - length ?pre), xcp')" by(rule exec_meth.cases)(auto intro!: exec_meth.intros simp add: match_ex_table_append_not_pcs) moreover from exec' have "pc' \ length ?pre" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) moreover hence "(Suc (Suc (pc' - Suc (Suc 0)))) = pc'" by simp ultimately show ?case using stk' by(auto simp add: shift_compxE2 eval_nat_numeral) next case bisims1Nil thus ?case by(auto elim: exec_meth.cases) next case (bisims1List1 e n e' xs stk loc pc xcp es) note bisim1 = \P,e,h \ (e', xs ) \ (stk, loc, pc, xcp)\ note IH1 = \\stk' loc' pc' xcp' STK. ?exec e stk STK loc pc xcp stk' loc' pc' xcp' \ ?concl e stk STK loc pc xcp stk' loc' pc' xcp'\ note IH2 = \\xs stk' loc' pc' xcp' STK. ?execs es [] STK xs 0 None stk' loc' pc' xcp' \ ?concls es [] STK xs 0 None stk' loc' pc' xcp'\ note exec = \?execs (e # es) stk STK loc pc xcp stk' loc' pc' xcp'\ from bisim1 have pc: "pc \ length (compE2 e)" by(rule bisim1_pc_length_compE2) show ?case proof(cases "pc < length (compE2 e)") case True with exec have "?exec e stk STK loc pc xcp stk' loc' pc' xcp'" by(simp add: compxEs2_size_convs)(erule exec_meth_take_xt) from IH1[OF this] show ?thesis by auto next case False with pc have pc: "pc = length (compE2 e)" by simp with bisim1 obtain v where s: "stk = [v]" "xcp = None" by(auto dest: bisim1_pc_length_compE2D) with exec pc have exec': "exec_meth_d (compP2 P) (compE2 e @ compEs2 es) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length (v # STK)) (compxEs2 es 0 0))) t h ([] @ v # STK, loc, length (compE2 e) + 0, None) ta h' (stk', loc', pc', xcp')" by(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) hence "?execs es [] (v # STK) loc 0 None stk' loc' (pc' - length (compE2 e)) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH2[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec'': "exec_meth_d (compP2 P) (compEs2 es) (compxEs2 es 0 0) t h ([], loc, 0, None) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" by auto from exec'' have "exec_meth_d (compP2 P) (compEs2 es) (stack_xlift (length [v]) (compxEs2 es 0 0)) t h ([] @ [v], loc, 0, None) ta h' (stk'' @ [v], loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e @ compEs2 es) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v]) (compxEs2 es 0 0))) t h ([] @ [v], loc, length (compE2 e) + 0, None) ta h' (stk'' @ [v], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule append_exec_meth_xt) auto moreover from exec' have "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?thesis using s pc stk' by(auto simp add: shift_compxEs2 stack_xlift_compxEs2) qed next case (bisims1List2 es n es' xs stk loc pc xcp e v) note bisim = \P,es,h \ (es',xs) [\] (stk,loc,pc,xcp)\ note IH = \\stk' loc' pc' xcp' STK. ?execs es stk STK loc pc xcp stk' loc' pc' xcp' \ ?concls es stk STK loc pc xcp stk' loc' pc' xcp'\ note exec = \?execs (e # es) (stk @ [v]) STK loc (length (compE2 e) + pc) xcp stk' loc' pc' xcp'\ from exec have exec': "exec_meth_d (compP2 P) (compE2 e @ compEs2 es) (stack_xlift (length STK) (compxE2 e 0 0) @ shift (length (compE2 e)) (stack_xlift (length (v # STK)) (compxEs2 es 0 0))) t h (stk @ v # STK, loc, length (compE2 e) + pc, xcp) ta h' (stk', loc', pc', xcp')" by(simp add: compxEs2_size_convs compxEs2_stack_xlift_convs) hence "?execs es stk (v # STK) loc pc xcp stk' loc' (pc' - length (compE2 e)) xcp'" by(rule exec_meth_drop_xt)(auto simp add: stack_xlift_compxE2) from IH[OF this] obtain stk'' where stk': "stk' = stk'' @ v # STK" and exec'': "exec_meth_d (compP2 P) (compEs2 es) (compxEs2 es 0 0) t h (stk, loc, pc, xcp) ta h' (stk'', loc', pc' - length (compE2 e), xcp')" by auto from exec'' have "exec_meth_d (compP2 P) (compEs2 es) (stack_xlift (length [v]) (compxEs2 es 0 0)) t h (stk @ [v], loc, pc, xcp) ta h' (stk'' @ [v], loc', pc' - length (compE2 e), xcp')" by(rule exec_meth_stk_offer) hence "exec_meth_d (compP2 P) (compE2 e @ compEs2 es) (compxE2 e 0 0 @ shift (length (compE2 e)) (stack_xlift (length [v]) (compxEs2 es 0 0))) t h (stk @ [v], loc, length (compE2 e) + pc, xcp) ta h' (stk'' @ [v], loc', length (compE2 e) + (pc' - length (compE2 e)), xcp')" by(rule append_exec_meth_xt) auto moreover from exec' have "pc' \ length (compE2 e)" by(rule exec_meth_drop_xt_pc)(auto simp add: stack_xlift_compxE2) ultimately show ?case using stk' by(auto simp add: shift_compxEs2 stack_xlift_compxEs2) next case (bisim1Sync12 e1 n e2 V a xs v v') note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [v, v'] STK xs (4 + length (compE2 e1) + length (compE2 e2)) \a\ stk' loc' pc' xcp'\ thus ?case by(auto elim!: exec_meth.cases split: if_split_asm simp add: match_ex_table_append_not_pcs)(simp add: matches_ex_entry_def) next case (bisim1Sync14 e1 n e2 V a xs v a') note exec = \?exec (sync\<^bsub>V\<^esub> (e1) e2) [v, Addr a'] STK xs (7 + length (compE2 e1) + length (compE2 e2)) \a\ stk' loc' pc' xcp'\ thus ?case by(auto elim!: exec_meth.cases split: if_split_asm simp add: match_ex_table_append_not_pcs)(simp add: matches_ex_entry_def) qed lemma shows bisim1_callD: "\ P,e,h \ (e', xs) \ (stk, loc, pc, xcp); call1 e' = \(a, M, vs)\; compE2 e ! pc = Invoke M' n0 \ \ M = M'" and bisims1_callD: "\ P,es,h \ (es',xs) [\] (stk,loc,pc, xcp); calls1 es' = \(a, M, vs)\; compEs2 es ! pc = Invoke M' n0 \ \ M = M'" proof(induct e "n :: nat" e' xs stk loc pc xcp and es "n :: nat" es' xs stk loc pc xcp rule: bisim1_bisims1_inducts_split) case bisim1AAss1 thus ?case apply(simp (no_asm_use) split: if_split_asm add: is_val_iff) apply(fastforce dest: bisim_Val_pc_not_Invoke) apply(fastforce dest: bisim_Val_pc_not_Invoke) apply(fastforce dest: bisim_Val_pc_not_Invoke bisim1_pc_length_compE2)+ done next case bisim1Call1 thus ?case apply(clarsimp split: if_split_asm simp add: is_vals_conv) apply(drule bisim_Val_pc_not_Invoke, simp, fastforce) apply(drule bisim_Val_pc_not_Invoke, simp, fastforce) apply(drule bisim1_pc_length_compE2, clarsimp simp add: neq_Nil_conv) apply(drule bisim1_pc_length_compE2, simp) apply(drule bisim1_pc_length_compE2, simp) apply(drule bisim1_pc_length_compE2, simp) apply(drule bisim1_call_pcD, simp, simp) apply(drule bisim1_call_pcD, simp, simp) done next case bisim1CallParams thus ?case apply(clarsimp split: if_split_asm simp add: is_vals_conv) apply(drule bisims_Val_pc_not_Invoke, simp, fastforce) apply(drule bisims1_pc_length_compEs2, simp) apply(drule bisims1_calls_pcD, simp, simp) done qed(fastforce split: if_split_asm dest: bisim1_pc_length_compE2 bisims1_pc_length_compEs2 bisim1_call_pcD bisims1_calls_pcD bisim1_call_xcpNone bisims1_calls_xcpNone bisim_Val_pc_not_Invoke bisims_Val_pc_not_Invoke)+ lemma bisim1_xcpD: "P,e,h \ (e', xs) \ (stk, loc, pc, \a\) \ pc < length (compE2 e)" and bisims1_xcpD: "P,es,h \ (es', xs) [\] (stk, loc, pc, \a\) \ pc < length (compEs2 es)" by(induct "(e', xs)" "(stk, loc, pc, \a :: 'addr\)" and "(es', xs)" "(stk, loc, pc, \a :: 'addr\)" arbitrary: e' xs stk loc pc and es' xs stk loc pc rule: bisim1_bisims1.inducts) simp_all lemma bisim1_match_Some_stk_length: "\ P,E,h \ (e, xs) \ (stk, loc, pc, \a\); match_ex_table (compP2 P) (cname_of h a) pc (compxE2 E 0 0) = \(pc', d)\ \ \ d \ length stk" and bisims1_match_Some_stk_length: "\ P,Es,h \ (es, xs) [\] (stk, loc, pc, \a\); match_ex_table (compP2 P) (cname_of h a) pc (compxEs2 Es 0 0) = \(pc', d)\ \ \ d \ length stk" proof(induct "(e, xs)" "(stk, loc, pc, \a :: 'addr\)" and "(es, xs)" "(stk, loc, pc, \a :: 'addr\)" arbitrary: pc' d e xs stk loc pc and pc' d es xs stk loc pc rule: bisim1_bisims1.inducts) case bisim1Call1 thus ?case by(fastforce dest: bisim1_xcpD simp add: match_ex_table_append match_ex_table_not_pcs_None) next case bisim1CallThrowObj thus ?case by(fastforce dest: bisim1_xcpD simp add: match_ex_table_append match_ex_table_not_pcs_None) next case bisim1Sync4 thus ?case apply(clarsimp simp add: match_ex_table_not_pcs_None match_ex_table_append matches_ex_entry_def split: if_split_asm) apply(fastforce simp add: match_ex_table_compxE2_shift_conv dest: bisim1_xcpD) done next case bisim1Try thus ?case by(fastforce simp add: match_ex_table_append matches_ex_entry_def match_ex_table_not_pcs_None dest: bisim1_xcpD split: if_split_asm) next case bisim1TryCatch2 thus ?case apply(clarsimp simp add: match_ex_table_not_pcs_None match_ex_table_append matches_ex_entry_def split: if_split_asm) apply(fastforce simp add: match_ex_table_compxE2_shift_conv dest: bisim1_xcpD) done next case bisim1TryFail thus ?case by(fastforce simp add: match_ex_table_append matches_ex_entry_def match_ex_table_not_pcs_None dest: bisim1_xcpD split: if_split_asm) next case bisim1TryCatchThrow thus ?case apply(clarsimp simp add: match_ex_table_not_pcs_None match_ex_table_append matches_ex_entry_def split: if_split_asm) apply(fastforce simp add: match_ex_table_compxE2_shift_conv dest: bisim1_xcpD) done next case bisims1List1 thus ?case by(fastforce simp add: match_ex_table_append split: if_split_asm dest: bisim1_xcpD match_ex_table_pcsD) qed(fastforce simp add: match_ex_table_not_pcs_None match_ex_table_append match_ex_table_compxE2_shift_conv match_ex_table_compxEs2_shift_conv match_ex_table_compxE2_stack_conv match_ex_table_compxEs2_stack_conv matches_ex_entry_def dest: bisim1_xcpD)+ end locale J1_JVM_heap_conf_base = J1_JVM_heap_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write + J1_heap_conf_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf P + JVM_heap_conf_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf "compP2 P" for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and hconf :: "'heap \ bool" and P :: "'addr J1_prog" begin inductive bisim1_list1 :: "'thread_id \ 'heap \ 'addr expr1 \ 'addr locals1 \ ('addr expr1 \ 'addr locals1) list \ 'addr option \ 'addr frame list \ bool" for t :: 'thread_id and h :: 'heap where bl1_Normal: "\ compTP P \ t:(xcp, h, (stk, loc, C, M, pc) # frs) \; P \ C sees M : Ts\T = \body\ in D; P,blocks1 0 (Class D#Ts) body, h \ (e, xs) \ (stk, loc, pc, xcp); max_vars e \ length xs; list_all2 (bisim1_fr P h) exs frs \ \ bisim1_list1 t h (e, xs) exs xcp ((stk, loc, C, M, pc) # frs)" | bl1_finalVal: "\ hconf h; preallocated h \ \ bisim1_list1 t h (Val v, xs) [] None []" | bl1_finalThrow: "\ hconf h; preallocated h \ \ bisim1_list1 t h (Throw a, xs) [] \a\ []" fun wbisim1 :: "'thread_id \ ((('addr expr1 \ 'addr locals1) \ ('addr expr1 \ 'addr locals1) list) \ 'heap, ('addr option \ 'addr frame list) \ 'heap) bisim" where "wbisim1 t ((ex, exs), h) ((xcp, frs), h') \ h = h' \ bisim1_list1 t h ex exs xcp frs" lemma new_thread_conf_compTP: assumes hconf: "hconf h" "preallocated h" and ha: "typeof_addr h a = \Class_type C\" and sub: "typeof_addr h (thread_id2addr t) = \Class_type C'\" "P \ C' \\<^sup>* Thread" and sees: "P \ C sees M: []\T = \meth\ in D" shows "compTP P \ t:(None, h, [([], Addr a # replicate (max_vars meth) undefined_value, D, M, 0)]) \" proof - from ha sees_method_decl_above[OF sees] have "P,h \ Addr a :\ Class D" by(simp add: conf_def) moreover hence "compP2 P,h \ Addr a :\ Class D" by(simp add: compP2_def) hence "compP2 P,h \ Addr a # replicate (max_vars meth) undefined_value [:\\<^sub>\] map (\i. if i = 0 then OK ([Class D] ! i) else Err) [0..i. if i = 0 then OK ([Class D] ! i) else Err) [0..l (Suc (max_vars meth)) [Class D] {0}) (compE2 meth @ [Return]) ([], Addr a # replicate (max_vars meth) undefined_value, D, M, 0)" by(simp add: TC0.ty\<^sub>l_def conf_f_def2 compP2_def) with hconf ha sub sees_method_compP[OF sees, where f="\C M Ts T. compMb2"] sees_method_idemp[OF sees] show ?thesis by(auto simp add: TC0.ty\<^sub>i'_def correct_state_def compTP_def tconf_def)(fastforce simp add: compP2_def compMb2_def tconf_def intro: sees_method_idemp)+ qed lemma ta_bisim12_extTA2J1_extTA2JVM: assumes nt: "\n T C M a h. \ n < length \ta\\<^bsub>t\<^esub>; \ta\\<^bsub>t\<^esub> ! n = NewThread T (C, M, a) h \ \ typeof_addr h a = \Class_type C\ \ (\C'. typeof_addr h (thread_id2addr T) = \Class_type C'\ \ P \ C' \\<^sup>* Thread) \ (\T meth D. P \ C sees M:[]\T =\meth\ in D) \ hconf h \ preallocated h" shows "ta_bisim wbisim1 (extTA2J1 P ta) (extTA2JVM (compP2 P) ta)" proof - { fix n t C M a m assume "n < length \ta\\<^bsub>t\<^esub>" and "\ta\\<^bsub>t\<^esub> ! n = NewThread t (C, M, a) m" from nt[OF this] obtain T meth D C' where ma: "typeof_addr m a = \Class_type C\" and sees: "P \ C sees M: []\T = \meth\ in D" and sub: "typeof_addr m (thread_id2addr t) = \Class_type C'\" "P \ C' \\<^sup>* Thread" and mconf: "hconf m" "preallocated m" by fastforce from sees_method_compP[OF sees, where f="\C M Ts T. compMb2"] have sees': "compP2 P \ C sees M: []\T = \(max_stack meth, max_vars meth, compE2 meth @ [Return], compxE2 meth 0 0)\ in D" by(simp add: compMb2_def compP2_def) have "bisim1_list1 t m ({0:Class D=None; meth}, Addr a # replicate (max_vars meth) undefined_value) ([]) None [([], Addr a # replicate (max_vars meth) undefined_value, D, M, 0)]" proof from mconf ma sub sees show "compTP P \ t:(None, m, [([], Addr a # replicate (max_vars meth) undefined_value, D, M, 0)]) \" by(rule new_thread_conf_compTP) from sees show "P \ D sees M: []\T = \meth\ in D" by(rule sees_method_idemp) show "list_all2 (bisim1_fr P m) [] []" by simp show "P,blocks1 0 [Class D] meth,m \ ({0:Class D=None; meth}, Addr a # replicate (max_vars meth) undefined_value) \ ([], Addr a # replicate (max_vars meth) undefined_value, 0, None)" by simp(rule bisim1_refl) qed simp with sees sees' have "bisim1_list1 t m ({0:Class (fst (method P C M))=None; the (snd (snd (snd (method P C M))))}, Addr a # replicate (max_vars (the (snd (snd (snd (method P C M)))))) undefined_value) [] None [([], Addr a # replicate (fst (snd (the (snd (snd (snd (method (compP2 P) C M))))))) undefined_value, fst (method (compP2 P) C M), M, 0)]" by simp } thus ?thesis apply(auto simp add: ta_bisim_def intro!: list_all2_all_nthI) apply(case_tac "\ta\\<^bsub>t\<^esub> ! n", auto simp add: extNTA2JVM_def) done qed end definition no_call2 :: "'addr expr1 \ pc \ bool" where "no_call2 e pc \ (pc \ length (compE2 e)) \ (pc < length (compE2 e) \ (\M n. compE2 e ! pc \ Invoke M n))" definition no_calls2 :: "'addr expr1 list \ pc \ bool" where "no_calls2 es pc \ (pc \ length (compEs2 es)) \ (pc < length (compEs2 es) \ (\M n. compEs2 es ! pc \ Invoke M n))" locale J1_JVM_conf_read = J1_JVM_heap_conf_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf P + JVM_conf_read addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf "compP2 P" for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and hconf :: "'heap \ bool" and P :: "'addr J1_prog" locale J1_JVM_heap_conf = J1_JVM_heap_conf_base addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf P + JVM_heap_conf addr2thread_id thread_id2addr spurious_wakeups empty_heap allocate typeof_addr heap_read heap_write hconf "compP2 P" for addr2thread_id :: "('addr :: addr) \ 'thread_id" and thread_id2addr :: "'thread_id \ 'addr" and spurious_wakeups :: bool and empty_heap :: "'heap" and allocate :: "'heap \ htype \ ('heap \ 'addr) set" and typeof_addr :: "'heap \ 'addr \ htype" and heap_read :: "'heap \ 'addr \ addr_loc \ 'addr val \ bool" and heap_write :: "'heap \ 'addr \ addr_loc \ 'addr val \ 'heap \ bool" and hconf :: "'heap \ bool" and P :: "'addr J1_prog" begin lemma red_external_ta_bisim21: "\ wf_prog wf_md P; P,t \ \a\M(vs), h\ -ta\ext \va, h'\; hconf h'; preallocated h' \ \ ta_bisim wbisim1 (extTA2J1 P ta) (extTA2JVM (compP2 P) ta)" apply(rule ta_bisim12_extTA2J1_extTA2JVM) apply(frule (1) red_external_new_thread_sees) apply(fastforce simp add: in_set_conv_nth) apply(frule red_ext_new_thread_heap) apply(fastforce simp add: in_set_conv_nth) apply(frule red_external_new_thread_exists_thread_object[unfolded compP2_def, simplified]) apply(fastforce simp add: in_set_conv_nth) apply simp done lemma ta_bisim_red_extTA2J1_extTA2JVM: assumes wf: "wf_prog wf_md P" and red: "uf,P,t' \1 \e, s\ -ta\ \e', s'\" and hconf: "hconf (hp s')" "preallocated (hp s')" shows "ta_bisim wbisim1 (extTA2J1 P ta) (extTA2JVM (compP2 P) ta)" proof - { fix n t C M a H assume len: "n < length \ta\\<^bsub>t\<^esub>" and tan: "\ta\\<^bsub>t\<^esub> ! n = NewThread t (C, M, a) H" hence nt: "NewThread t (C, M, a) H \ set \ta\\<^bsub>t\<^esub>" unfolding set_conv_nth by(auto intro!: exI) from red1_new_threadD[OF red nt] obtain ad M' vs va T C' Ts' Tr' D' where rede: "P,t' \ \ad\M'(vs),hp s\ -ta\ext \va,hp s'\" and ad: "typeof_addr (hp s) ad = \T\" by blast from red_ext_new_thread_heap[OF rede nt] have [simp]: "hp s' = H" by simp from red_external_new_thread_sees[OF wf rede nt] obtain T body D where Ha: "typeof_addr H a = \Class_type C\" and sees: "P \ C sees M:[]\T=\body\ in D" by auto have sees': "compP2 P \ C sees M:[]\T=\(max_stack body, max_vars body, compE2 body @ [Return], compxE2 body 0 0)\ in D" using sees unfolding compP2_def compMb2_def Let_def by(auto dest: sees_method_compP) from red_external_new_thread_exists_thread_object[unfolded compP2_def, simplified, OF rede nt] hconf Ha sees have "compTP P \ t:(None, H, [([], Addr a # replicate (max_vars body) undefined_value, D, M, 0)]) \" by(auto intro: new_thread_conf_compTP) hence "bisim1_list1 t H ({0:Class D=None; body}, Addr a # replicate (max_vars body) undefined_value) [] None [([], Addr a # replicate (max_vars body) undefined_value, D, M, 0)]" proof from sees show "P \ D sees M:[]\T=\body\ in D" by(rule sees_method_idemp) show "P,blocks1 0 [Class D] body,H \ ({0:Class D=None; body}, Addr a # replicate (max_vars body) undefined_value) \ ([], Addr a # replicate (max_vars body) undefined_value, 0, None)" by(auto intro: bisim1_refl) qed simp_all hence "bisim1_list1 t H ({0:Class (fst (method P C M))=None; the (snd (snd (snd (method P C M))))}, Addr a # replicate (max_vars (the (snd (snd (snd (method P C M)))))) undefined_value) [] None [([], Addr a # replicate (fst (snd (the (snd (snd (snd (method (compP2 P) C M))))))) undefined_value, fst (method (compP2 P) C M), M, 0)]" using sees sees' by simp } thus ?thesis apply(auto simp add: ta_bisim_def intro!: list_all2_all_nthI) apply(case_tac "\ta\\<^bsub>t\<^esub> ! n") apply(auto simp add: extNTA2JVM_def extNTA2J1_def) done qed end sublocale J1_JVM_conf_read < heap_conf?: J1_JVM_heap_conf by(unfold_locales) sublocale J1_JVM_conf_read < heap?: J1_heap apply(rule J1_heap.intro) apply(subst compP_heap[symmetric, where f="\_ _ _ _. compMb2", folded compP2_def]) apply(unfold_locales) done end diff --git a/thys/JinjaThreads/J/Annotate.thy b/thys/JinjaThreads/J/Annotate.thy --- a/thys/JinjaThreads/J/Annotate.thy +++ b/thys/JinjaThreads/J/Annotate.thy @@ -1,358 +1,358 @@ (* Title: Jinja/J/Annotate.thy Author: Tobias Nipkow, Andreas Lochbihler *) section \Program annotation\ theory Annotate imports WellType begin abbreviation (output) unanFAcc :: "'addr expr \ vname \ 'addr expr" ("(_\_)" [10,10] 90) where "unanFAcc e F \ FAcc e F (STR '''')" abbreviation (output) unanFAss :: "'addr expr \ vname \ 'addr expr \ 'addr expr" ("(_\_ := _)" [10,0,90] 90) where "unanFAss e F e' \ FAss e F (STR '''') e'" definition array_length_field_name :: vname where "array_length_field_name = STR ''length''" notation (output) array_length_field_name ("length") definition super :: vname where "super = STR ''super''" lemma super_neq_this [simp]: "super \ this" "this \ super" by(simp_all add: this_def super_def) inductive Anno :: "(ty \ ty \ ty \ bool) \ 'addr J_prog \ env \ 'addr expr \ 'addr expr \ bool" ("_,_,_ \ _ \ _" [51,51,0,0,51]50) and Annos :: "(ty \ ty \ ty \ bool) \ 'addr J_prog \ env \ 'addr expr list \ 'addr expr list \ bool" ("_,_,_ \ _ [\] _" [51,51,0,0,51]50) for is_lub :: "ty \ ty \ ty \ bool" and P :: "'addr J_prog" where AnnoNew: "is_lub,P,E \ new C \ new C" | AnnoNewArray: "is_lub,P,E \ i \ i' \ is_lub,P,E \ newA T\i\ \ newA T\i'\" | AnnoCast: "is_lub,P,E \ e \ e' \ is_lub,P,E \ Cast C e \ Cast C e'" | AnnoInstanceOf: "is_lub,P,E \ e \ e' \ is_lub,P,E \ e instanceof T \ e' instanceof T" | AnnoVal: "is_lub,P,E \ Val v \ Val v" | AnnoVarVar: "\ E V = \T\; V \ super \ \ is_lub,P,E \ Var V \ Var V" | AnnoVarField: \ \There is no need to handle access of array fields explicitly, because arrays do not implement methods, i.e. @{term "this"} is always of a @{term "Class"} type.\ "\ E V = None; V \ super; E this = \Class C\; P \ C sees V:T (fm) in D \ \ is_lub,P,E \ Var V \ Var this\V{D}" | AnnoBinOp: "\ is_lub,P,E \ e1 \ e1'; is_lub,P,E \ e2 \ e2' \ \ is_lub,P,E \ e1 \bop\ e2 \ e1' \bop\ e2'" | AnnoLAssVar: "\ E V = \T\; V \ super; is_lub,P,E \ e \ e' \ \ is_lub,P,E \ V:=e \ V:=e'" | AnnoLAssField: "\ E V = None; V \ super; E this = \Class C\; P \ C sees V:T (fm) in D; is_lub,P,E \ e \ e' \ \ is_lub,P,E \ V:=e \ Var this\V{D} := e'" | AnnoAAcc: "\ is_lub,P,E \ a \ a'; is_lub,P,E \ i \ i' \ \ is_lub,P,E \ a\i\ \ a'\i'\" | AnnoAAss: "\ is_lub,P,E \ a \ a'; is_lub,P,E \ i \ i'; is_lub,P,E \ e \ e' \ \ is_lub,P,E \ a\i\ := e \ a'\i'\ := e'" | AnnoALength: "is_lub,P,E \ a \ a' \ is_lub,P,E \ a\length \ a'\length" | \ \All arrays implicitly declare a final field called @{term "array_length_field_name"} to store the array length, which hides a potential field of the same name in @{term "Object"} (cf. JLS 6.4.5). The last premise implements the hiding because field lookup does does not model the implicit declaration.\ AnnoFAcc: "\ is_lub,P,E \ e \ e'; is_lub,P,E \ e' :: U; class_type_of' U = \C\; P \ C sees F:T (fm) in D; is_Array U \ F \ array_length_field_name \ \ is_lub,P,E \ e\F{STR ''''} \ e'\F{D}" | AnnoFAccALength: "\ is_lub,P,E \ e \ e'; is_lub,P,E \ e' :: T\\ \ \ is_lub,P,E \ e\array_length_field_name{STR ''''} \ e'\length" | AnnoFAccSuper: \ \In class C with super class D, "super" is syntactic sugar for "((D) this)" (cf. JLS, 15.11.2)\ "\ E this = \Class C\; C \ Object; class P C = \(D, fs, ms)\; P \ D sees F:T (fm) in D' \ \ is_lub,P,E \ Var super\F{STR ''''} \ (Cast (Class D) (Var this))\F{D'}" | AnnoFAss: "\ is_lub,P,E \ e1 \ e1'; is_lub,P,E \ e2 \ e2'; is_lub,P,E \ e1' :: U; class_type_of' U = \C\; P \ C sees F:T (fm) in D; is_Array U \ F \ array_length_field_name \ \ is_lub,P,E \ e1\F{STR ''''} := e2 \ e1'\F{D} := e2'" | AnnoFAssSuper: "\ E this = \Class C\; C \ Object; class P C = \(D, fs, ms)\; P \ D sees F:T (fm) in D'; is_lub,P,E \ e \ e' \ \ is_lub,P,E \ Var super\F{STR ''''} := e \ (Cast (Class D) (Var this))\F{D'} := e'" | AnnoCAS: "\ is_lub,P,E \ e1 \ e1'; is_lub,P,E \ e2 \ e2'; is_lub,P,E \ e3 \ e3' \ \ is_lub,P,E \ e1\compareAndSwap(D\F, e2, e3) \ e1'\compareAndSwap(D\F, e2', e3')" | AnnoCall: "\ is_lub,P,E \ e \ e'; is_lub,P,E \ es [\] es' \ \ is_lub,P,E \ Call e M es \ Call e' M es'" | AnnoBlock: "is_lub,P,E(V \ T) \ e \ e' \ is_lub,P,E \ {V:T=vo; e} \ {V:T=vo; e'}" | AnnoSync: "\ is_lub,P,E \ e1 \ e1'; is_lub,P,E \ e2 \ e2' \ \ is_lub,P,E \ sync(e1) e2 \ sync(e1') e2'" | AnnoComp: "\ is_lub,P,E \ e1 \ e1'; is_lub,P,E \ e2 \ e2' \ \ is_lub,P,E \ e1;;e2 \ e1';;e2'" | AnnoCond: "\ is_lub,P,E \ e \ e'; is_lub,P,E \ e1 \ e1'; is_lub,P,E \ e2 \ e2' \ \ is_lub,P,E \ if (e) e1 else e2 \ if (e') e1' else e2'" | AnnoLoop: "\ is_lub,P,E \ e \ e'; is_lub,P,E \ c \ c' \ \ is_lub,P,E \ while (e) c \ while (e') c'" | AnnoThrow: "is_lub,P,E \ e \ e' \ is_lub,P,E \ throw e \ throw e'" | AnnoTry: "\ is_lub,P,E \ e1 \ e1'; is_lub,P,E(V \ Class C) \ e2 \ e2' \ \ is_lub,P,E \ try e1 catch(C V) e2 \ try e1' catch(C V) e2'" | AnnoNil: "is_lub,P,E \ [] [\] []" | AnnoCons: "\ is_lub,P,E \ e \ e'; is_lub,P,E \ es [\] es' \ \ is_lub,P,E \ e#es [\] e'#es'" inductive_cases Anno_cases [elim!]: "is_lub',P,E \ new C \ e" "is_lub',P,E \ newA T\e\ \ e'" "is_lub',P,E \ Cast T e \ e'" "is_lub',P,E \ e instanceof T \ e'" "is_lub',P,E \ Val v \ e'" "is_lub',P,E \ Var V \ e'" "is_lub',P,E \ e1 \bop\ e2 \ e'" "is_lub',P,E \ V := e \ e'" "is_lub',P,E \ e1\e2\ \ e'" "is_lub',P,E \ e1\e2\ := e3 \ e'" "is_lub',P,E \ e\length \ e'" "is_lub',P,E \ e\F{D} \ e'" "is_lub',P,E \ e1\F{D} := e2 \ e'" "is_lub',P,E \ e1\compareAndSwap(D\F, e2, e3) \ e'" "is_lub',P,E \ e\M(es) \ e'" "is_lub',P,E \ {V:T=vo; e} \ e'" "is_lub',P,E \ sync(e1) e2 \ e'" "is_lub',P,E \ insync(a) e2 \ e'" "is_lub',P,E \ e1;; e2 \ e'" "is_lub',P,E \ if (e) e1 else e2 \ e'" "is_lub',P,E \ while(e1) e2 \ e'" "is_lub',P,E \ throw e \ e'" "is_lub',P,E \ try e1 catch(C V) e2 \ e'" inductive_cases Annos_cases [elim!]: "is_lub',P,E \ [] [\] es'" "is_lub',P,E \ e # es [\] es'" abbreviation Anno' :: "'addr J_prog \ env \ 'addr expr \ 'addr expr \ bool" ("_,_ \ _ \ _" [51,0,0,51]50) where "Anno' P \ Anno (TypeRel.is_lub P) P" abbreviation Annos' :: "'addr J_prog \ env \ 'addr expr list \ 'addr expr list \ bool" ("_,_ \ _ [\] _" [51,0,0,51]50) where "Annos' P \ Annos (TypeRel.is_lub P) P" definition annotate :: "'addr J_prog \ env \ 'addr expr \ 'addr expr" where "annotate P E e = THE_default e (\e'. P,E \ e \ e')" lemma fixes is_lub :: "ty \ ty \ ty \ bool" ("\ lub'((_,/ _)') = _" [51,51,51] 50) assumes is_lub_unique: "\T1 T2 T3 T4. \ \ lub(T1, T2) = T3; \ lub(T1, T2) = T4 \ \ T3 = T4" shows Anno_fun: "\ is_lub,P,E \ e \ e'; is_lub,P,E \ e \ e'' \ \ e' = e''" and Annos_fun: "\ is_lub,P,E \ es [\] es'; is_lub,P,E \ es [\] es'' \ \ es' = es''" proof(induct arbitrary: e'' and es'' rule: Anno_Annos.inducts) case (AnnoFAcc E e e' U C F T fm D) from \is_lub,P,E \ e\F{STR ''''} \ e''\ show ?case proof(rule Anno_cases) fix e''' U' C' T' fm' D' assume "is_lub,P,E \ e \ e'''" "is_lub,P,E \ e''' :: U'" and "class_type_of' U' = \C'\" and "P \ C' sees F:T' (fm') in D'" "e'' = e'''\F{D'}" from \is_lub,P,E \ e \ e'''\ have "e' = e'''" by(rule AnnoFAcc) with \is_lub,P,E \ e' :: U\ \is_lub,P,E \ e''' :: U'\ have "U = U'" by(auto intro: WT_unique is_lub_unique) with \class_type_of' U = \C\\ \class_type_of' U' = \C'\\ have "C = C'" by(auto) with \P \ C' sees F:T' (fm') in D'\ \P \ C sees F:T (fm) in D\ have "D' = D" by(auto dest: sees_field_fun) with \e'' = e'''\F{D'}\ \e' = e'''\ show ?thesis by simp next fix e''' T assume "e'' = e'''\length" and "is_lub,P,E \ e''' :: T\\" and "is_lub,P,E \ e \ e'''" and "F = array_length_field_name" from \is_lub,P,E \ e \ e'''\ have "e' = e'''" by(rule AnnoFAcc) with \is_lub,P,E \ e' :: U\ \is_lub,P,E \ e''' :: T\\\ have "U = T\\" by(auto intro: WT_unique is_lub_unique) with \class_type_of' U = \C\\ \is_Array U \ F \ array_length_field_name\ show ?thesis using \F = array_length_field_name\ by simp next fix C' D' fs ms T D'' assume "E this = \Class C'\" and "class P C' = \(D', fs, ms)\" and "e = Var super" and "e'' = Cast (Class D') (Var this)\F{D''}" with \is_lub,P,E \ e \ e'\ have False by(auto) thus ?thesis .. qed next case AnnoFAccALength thus ?case by(fastforce intro: WT_unique[OF is_lub_unique]) next case (AnnoFAss E e1 e1' e2 e2' U C F T fm D) from \is_lub,P,E \ e1\F{STR ''''} := e2 \ e''\ show ?case proof(rule Anno_cases) fix e1'' e2'' U' C' T' fm' D' assume "is_lub,P,E \ e1 \ e1''" "is_lub,P,E \ e2 \ e2''" and "is_lub,P,E \ e1'' :: U'" and "class_type_of' U' = \C'\" and "P \ C' sees F:T' (fm') in D'" and "e'' = e1''\F{D'} := e2''" from \is_lub,P,E \ e1 \ e1''\ have "e1' = e1''" by(rule AnnoFAss) moreover with \is_lub,P,E \ e1' :: U\ \is_lub,P,E \ e1'' :: U'\ have "U = U'" by(auto intro: WT_unique is_lub_unique) with \class_type_of' U = \C\\ \class_type_of' U' = \C'\\ have "C = C'" by(auto) with \P \ C' sees F:T' (fm') in D'\ \P \ C sees F:T (fm) in D\ have "D' = D" by(auto dest: sees_field_fun) moreover from \is_lub,P,E \ e2 \ e2''\ have "e2' = e2''" by(rule AnnoFAss) ultimately show ?thesis using \e'' = e1''\F{D'} := e2''\ by simp next fix C' D' fs ms T' fm' D'' e''' assume "e'' = Cast (Class D') (Var this)\F{D''} := e'''" and "E this = \Class C'\" and "class P C' = \(D', fs, ms)\" and "P \ D' sees F:T' (fm') in D''" and "is_lub,P,E \ e2 \ e'''" and "e1 = Var super" with \is_lub,P,E \ e1 \ e1'\ have False by(auto elim: Anno_cases) thus ?thesis .. qed qed(fastforce dest: sees_field_fun)+ subsection \Code generation\ -definition Anno_code :: "'addr J_prog \ env \ 'addr expr \ 'addr expr \ bool" ("_,_ \ _ \' _" [51,0,0,51]50) +definition Anno_code :: "'addr J_prog \ env \ 'addr expr \ 'addr expr \ bool" ("_,_ \ _ \'' _" [51,0,0,51]50) where "Anno_code P = Anno (is_lub_sup P) P" definition Annos_code :: "'addr J_prog \ env \ 'addr expr list \ 'addr expr list \ bool" ("_,_ \ _ [\''] _" [51,0,0,51]50) where "Annos_code P = Annos (is_lub_sup P) P" primrec block_types :: "('a, 'b, 'addr) exp \ ty list" and blocks_types :: "('a, 'b, 'addr) exp list \ ty list" where "block_types (new C) = []" | "block_types (newA T\e\) = block_types e" | "block_types (Cast U e) = block_types e" | "block_types (e instanceof U) = block_types e" | "block_types (e1\bop\e2) = block_types e1 @ block_types e2" | "block_types (Val v) = []" | "block_types (Var V) = []" | "block_types (V := e) = block_types e" | "block_types (a\i\) = block_types a @ block_types i" | "block_types (a\i\ := e) = block_types a @ block_types i @ block_types e" | "block_types (a\length) = block_types a" | "block_types (e\F{D}) = block_types e" | "block_types (e\F{D} := e') = block_types e @ block_types e'" | "block_types (e\compareAndSwap(D\F, e', e'')) = block_types e @ block_types e' @ block_types e''" | "block_types (e\M(es)) = block_types e @ blocks_types es" | "block_types {V:T=vo; e} = T # block_types e" | "block_types (sync\<^bsub>V\<^esub>(e) e') = block_types e @ block_types e'" | "block_types (insync\<^bsub>V\<^esub>(a) e) = block_types e" | "block_types (e;;e') = block_types e @ block_types e'" | "block_types (if (e) e1 else e2) = block_types e @ block_types e1 @ block_types e2" | "block_types (while (b) c) = block_types b @ block_types c" | "block_types (throw e) = block_types e" | "block_types (try e catch(C V) e') = block_types e @ Class C # block_types e'" | "blocks_types [] = []" | "blocks_types (e#es) = block_types e @ blocks_types es" lemma fixes is_lub1 :: "ty \ ty \ ty \ bool" ("\1 lub'((_,/ _)') = _" [51,51,51] 50) and is_lub2 :: "ty \ ty \ ty \ bool" ("\2 lub'((_,/ _)') = _" [51,51,51] 50) assumes wf: "wf_prog wf_md P" and is_lub1_into_is_lub2: "\T1 T2 T3. \ \1 lub(T1, T2) = T3; is_type P T1; is_type P T2 \ \ \2 lub(T1, T2) = T3" and is_lub2_is_type: "\T1 T2 T3. \ \2 lub(T1, T2) = T3; is_type P T1; is_type P T2 \ \ is_type P T3" shows Anno_change_is_lub: "\ is_lub1,P,E \ e \ e'; ran E \ set (block_types e) \ types P \ \ is_lub2,P,E \ e \ e'" and Annos_change_is_lub: "\ is_lub1,P,E \ es [\] es'; ran E \ set (blocks_types es) \ types P \ \ is_lub2,P,E \ es [\] es'" proof(induct rule: Anno_Annos.inducts) case (AnnoBlock E V T e e' vo) from \ran E \ set (block_types {V:T=vo; e}) \ types P\ have "ran (E(V \ T)) \ set (block_types e) \ types P" by(auto simp add: ran_def) thus ?case using AnnoBlock by(blast intro: Anno_Annos.intros) next case (AnnoTry E e1 e1' V C e2 e2') from \ran E \ set (block_types (try e1 catch(C V) e2)) \ types P\ have "ran (E(V \ Class C)) \ set (block_types e2) \ types P" by(auto simp add: ran_def) thus ?case using AnnoTry by(simp del: fun_upd_apply)(blast intro: Anno_Annos.intros) qed(simp_all del: is_Array.simps is_Array_conv, (blast intro: Anno_Annos.intros WT_change_is_lub[OF wf, where ?is_lub1.0=is_lub1 and ?is_lub2.0=is_lub2] is_lub1_into_is_lub2 is_lub2_is_type)+) lemma assumes wf: "wf_prog wf_md P" shows Anno_into_Anno_code: "\ P,E \ e \ e'; ran E \ set (block_types e) \ types P \ \ P,E \ e \' e'" and Annos_into_Annos_code: "\ P,E \ es [\] es'; ran E \ set (blocks_types es) \ types P \ \ P,E \ es [\'] es'" proof - assume anno: "P,E \ e \ e'" and ran: "ran E \ set (block_types e) \ types P" show "P,E \ e \' e'" unfolding Anno_code_def by(rule Anno_change_is_lub[OF wf _ _ anno ran])(blast intro!: is_lub_sup.intros intro: is_lub_subD[OF wf] sup_is_type[OF wf] elim!: is_lub_sup.cases)+ next assume annos: "P,E \ es [\] es'" and ran: "ran E \ set (blocks_types es) \ types P" show "P,E \ es [\'] es'" unfolding Annos_code_def by(rule Annos_change_is_lub[OF wf _ _ annos ran])(blast intro!: is_lub_sup.intros intro: is_lub_subD[OF wf] sup_is_type[OF wf] elim!: is_lub_sup.cases)+ qed lemma assumes wf: "wf_prog wf_md P" shows Anno_code_into_Anno: "\ P,E \ e \' e'; ran E \ set (block_types e) \ types P \ \ P,E \ e \ e'" and Annos_code_into_Annos: "\ P,E \ es [\'] es'; ran E \ set (blocks_types es) \ types P \ \ P,E \ es [\] es'" proof - assume anno: "P,E \ e \' e'" and ran: "ran E \ set (block_types e) \ types P" show "P,E \ e \ e'" by(rule Anno_change_is_lub[OF wf _ _ anno[unfolded Anno_code_def] ran])(blast elim!: is_lub_sup.cases intro: sup_is_lubI[OF wf] is_lub_is_type[OF wf])+ next assume annos: "P,E \ es [\'] es'" and ran: "ran E \ set (blocks_types es) \ types P" show "P,E \ es [\] es'" by(rule Annos_change_is_lub[OF wf _ _ annos[unfolded Annos_code_def] ran])(blast elim!: is_lub_sup.cases intro: sup_is_lubI[OF wf] is_lub_is_type[OF wf])+ qed lemma fixes is_lub assumes wf: "wf_prog wf_md P" shows WT_block_types_is_type: "is_lub,P,E \ e :: T \ set (block_types e) \ types P" and WTs_blocks_types_is_type: "is_lub,P,E \ es [::] Ts \ set (blocks_types es) \ types P" apply(induct rule: WT_WTs.inducts) apply(auto intro: is_class_sub_Throwable[OF wf]) done lemma fixes is_lub shows Anno_block_types: "is_lub,P,E \ e \ e' \ block_types e = block_types e'" and Annos_blocks_types: "is_lub,P,E \ es [\] es' \ blocks_types es = blocks_types es'" by(induct rule: Anno_Annos.inducts) auto code_pred (modes: (i \ i \ o \ bool) \ i \ i \ i \ o \ bool) [detect_switches, skip_proof] Anno . definition annotate_code :: "'addr J_prog \ env \ 'addr expr \ 'addr expr" where "annotate_code P E e = THE_default e (\e'. P,E \ e \' e')" code_pred (modes: i \ i \ i \ o \ bool) [inductify] Anno_code . lemma eval_Anno_i_i_i_o_conv: "Predicate.eval (Anno_code_i_i_i_o P E e) = (\e'. P,E \ e \' e')" by(auto intro!: ext intro: Anno_code_i_i_i_oI elim: Anno_code_i_i_i_oE) lemma annotate_code [code]: "annotate_code P E e = Predicate.singleton (\_. Code.abort (STR ''annotate'') (\_. e)) (Anno_code_i_i_i_o P E e)" by(simp add: THE_default_def Predicate.singleton_def annotate_code_def eval_Anno_i_i_i_o_conv) end diff --git a/thys/JinjaThreads/J/WellType.thy b/thys/JinjaThreads/J/WellType.thy --- a/thys/JinjaThreads/J/WellType.thy +++ b/thys/JinjaThreads/J/WellType.thy @@ -1,419 +1,419 @@ (* Title: JinjaThreads/J/WellType.thy Author: Tobias Nipkow, Andreas Lochbihler *) section \Well-typedness of Jinja expressions\ theory WellType imports Expr State "../Common/ExternalCallWF" "../Common/WellForm" "../Common/SemiType" begin declare Listn.lesub_list_impl_same_size[simp del] declare listE_length [simp del] type_synonym env = "vname \ ty" inductive WT :: "(ty \ ty \ ty \ bool) \ 'addr J_prog \ env \ 'addr expr \ ty \ bool" ("_,_,_ \ _ :: _" [51,51,51,51]50) and WTs :: "(ty \ ty \ ty \ bool) \ 'addr J_prog \ env \ 'addr expr list \ ty list \ bool" ("_,_,_ \ _ [::] _" [51,51,51,51]50) for is_lub :: "ty \ ty \ ty \ bool" ("\ lub'((_,/ _)') = _" [51,51,51] 50) and P :: "'addr J_prog" where WTNew: "is_class P C \ is_lub,P,E \ new C :: Class C" | WTNewArray: "\ is_lub,P,E \ e :: Integer; is_type P (T\\) \ \ is_lub,P,E \ newA T\e\ :: T\\" | WTCast: "\ is_lub,P,E \ e :: T; P \ U \ T \ P \ T \ U; is_type P U \ \ is_lub,P,E \ Cast U e :: U" | WTInstanceOf: "\ is_lub,P,E \ e :: T; P \ U \ T \ P \ T \ U; is_type P U; is_refT U \ \ is_lub,P,E \ e instanceof U :: Boolean" | WTVal: "typeof v = Some T \ is_lub,P,E \ Val v :: T" | WTVar: "E V = Some T \ is_lub,P,E \ Var V :: T" | WTBinOp: "\ is_lub,P,E \ e1 :: T1; is_lub,P,E \ e2 :: T2; P \ T1\bop\T2 :: T \ \ is_lub,P,E \ e1\bop\e2 :: T" | WTLAss: "\ E V = Some T; is_lub,P,E \ e :: T'; P \ T' \ T; V \ this \ \ is_lub,P,E \ V:=e :: Void" | WTAAcc: "\ is_lub,P,E \ a :: T\\; is_lub,P,E \ i :: Integer \ \ is_lub,P,E \ a\i\ :: T" | WTAAss: "\ is_lub,P,E \ a :: T\\; is_lub,P,E \ i :: Integer; is_lub,P,E \ e :: T'; P \ T' \ T \ \ is_lub,P,E \ a\i\ := e :: Void" | WTALength: "is_lub,P,E \ a :: T\\ \ is_lub,P,E \ a\length :: Integer" | WTFAcc: "\ is_lub,P,E \ e :: U; class_type_of' U = \C\; P \ C sees F:T (fm) in D \ \ is_lub,P,E \ e\F{D} :: T" | WTFAss: "\ is_lub,P,E \ e\<^sub>1 :: U; class_type_of' U = \C\; P \ C sees F:T (fm) in D; is_lub,P,E \ e\<^sub>2 :: T'; P \ T' \ T \ \ is_lub,P,E \ e\<^sub>1\F{D}:=e\<^sub>2 :: Void" | WTCAS: "\ is_lub,P,E \ e1 :: U; class_type_of' U = \C\; P \ C sees F:T (fm) in D; volatile fm; is_lub,P,E \ e2 :: T'; P \ T' \ T; is_lub,P,E \ e3 :: T''; P \ T'' \ T \ \ is_lub,P,E \ e1\compareAndSwap(D\F, e2, e3) :: Boolean" | WTCall: "\ is_lub,P,E \ e :: U; class_type_of' U = \C\; P \ C sees M:Ts \ T = meth in D; is_lub,P,E \ es [::] Ts'; P \ Ts' [\] Ts \ \ is_lub,P,E \ e\M(es) :: T" | WTBlock: "\ is_type P T; is_lub,P,E(V \ T) \ e :: T'; case vo of None \ True | \v\ \ \T'. typeof v = \T'\ \ P \ T' \ T \ \ is_lub,P,E \ {V:T=vo; e} :: T'" | WTSynchronized: "\ is_lub,P,E \ o' :: T; is_refT T; T \ NT; is_lub,P,E \ e :: T' \ \ is_lub,P,E \ sync(o') e :: T'" \ \Note that insync is not statically typable.\ | WTSeq: "\ is_lub,P,E \ e\<^sub>1::T\<^sub>1; is_lub,P,E \ e\<^sub>2::T\<^sub>2 \ \ is_lub,P,E \ e\<^sub>1;;e\<^sub>2 :: T\<^sub>2" | WTCond: "\ is_lub,P,E \ e :: Boolean; is_lub,P,E \ e\<^sub>1::T\<^sub>1; is_lub,P,E \ e\<^sub>2::T\<^sub>2; \ lub(T\<^sub>1, T\<^sub>2) = T \ \ is_lub,P,E \ if (e) e\<^sub>1 else e\<^sub>2 :: T" | WTWhile: "\ is_lub,P,E \ e :: Boolean; is_lub,P,E \ c::T \ \ is_lub,P,E \ while (e) c :: Void" | WTThrow: "\ is_lub,P,E \ e :: Class C; P \ C \\<^sup>* Throwable \ \ is_lub,P,E \ throw e :: Void" | WTTry: "\ is_lub,P,E \ e\<^sub>1 :: T; is_lub,P,E(V \ Class C) \ e\<^sub>2 :: T; P \ C \\<^sup>* Throwable \ \ is_lub,P,E \ try e\<^sub>1 catch(C V) e\<^sub>2 :: T" | WTNil: "is_lub,P,E \ [] [::] []" | WTCons: "\ is_lub,P,E \ e :: T; is_lub,P,E \ es [::] Ts \ \ is_lub,P,E \ e#es [::] T#Ts" abbreviation WT' :: "'addr J_prog \ env \ 'addr expr \ ty \ bool" ("_,_ \ _ :: _" [51,51,51] 50) where "WT' P \ WT (TypeRel.is_lub P) P" abbreviation WTs' :: "'addr J_prog \ env \ 'addr expr list \ ty list \ bool" ("_,_ \ _ [::] _" [51,51,51] 50) where "WTs' P \ WTs (TypeRel.is_lub P) P" declare WT_WTs.intros[intro!] inductive_simps WTs_iffs [iff]: "is_lub',P,E \ [] [::] Ts" "is_lub',P,E \ e#es [::] T#Ts" "is_lub',P,E \ e#es [::] Ts" lemma WTs_conv_list_all2: fixes is_lub shows "is_lub,P,E \ es [::] Ts = list_all2 (WT is_lub P E) es Ts" by(induct es arbitrary: Ts)(auto simp add: list_all2_Cons1 elim: WTs.cases) lemma WTs_append [iff]: "\is_lub Ts. (is_lub,P,E \ es\<^sub>1 @ es\<^sub>2 [::] Ts) = (\Ts\<^sub>1 Ts\<^sub>2. Ts = Ts\<^sub>1 @ Ts\<^sub>2 \ is_lub,P,E \ es\<^sub>1 [::] Ts\<^sub>1 \ is_lub,P,E \ es\<^sub>2[::]Ts\<^sub>2)" by(auto simp add: WTs_conv_list_all2 list_all2_append1 dest: list_all2_lengthD[symmetric]) inductive_simps WT_iffs [iff]: "is_lub',P,E \ Val v :: T" "is_lub',P,E \ Var V :: T" "is_lub',P,E \ e\<^sub>1;;e\<^sub>2 :: T\<^sub>2" "is_lub',P,E \ {V:T=vo; e} :: T'" inductive_cases WT_elim_cases[elim!]: "is_lub',P,E \ V :=e :: T" "is_lub',P,E \ sync(o') e :: T" "is_lub',P,E \ if (e) e\<^sub>1 else e\<^sub>2 :: T" "is_lub',P,E \ while (e) c :: T" "is_lub',P,E \ throw e :: T" "is_lub',P,E \ try e\<^sub>1 catch(C V) e\<^sub>2 :: T" "is_lub',P,E \ Cast D e :: T" "is_lub',P,E \ e instanceof U :: T" "is_lub',P,E \ a\F{D} :: T" "is_lub',P,E \ a\F{D} := v :: T" "is_lub',P,E \ e\compareAndSwap(D\F, e', e'') :: T" "is_lub',P,E \ e\<^sub>1 \bop\ e\<^sub>2 :: T" "is_lub',P,E \ new C :: T" "is_lub',P,E \ newA T\e\ :: T'" "is_lub',P,E \ a\i\ := e :: T" "is_lub',P,E \ a\i\ :: T" "is_lub',P,E \ a\length :: T" "is_lub',P,E \ e\M(ps) :: T" "is_lub',P,E \ sync(o') e :: T" "is_lub',P,E \ insync(a) e :: T" lemma fixes is_lub :: "ty \ ty \ ty \ bool" ("\ lub'((_,/ _)') = _" [51,51,51] 50) assumes is_lub_unique: "\T1 T2 T3 T4. \ \ lub(T1, T2) = T3; \ lub(T1, T2) = T4 \ \ T3 = T4" shows WT_unique: "\ is_lub,P,E \ e :: T; is_lub,P,E \ e :: T' \ \ T = T'" and WTs_unique: "\ is_lub,P,E \ es [::] Ts; is_lub,P,E \ es [::] Ts' \ \ Ts = Ts'" apply(induct arbitrary: T' and Ts' rule: WT_WTs.inducts) apply blast apply blast apply blast apply blast apply fastforce apply fastforce apply(fastforce dest: WT_binop_fun) apply fastforce apply fastforce apply fastforce apply fastforce apply(fastforce dest: sees_field_fun) apply(fastforce dest: sees_field_fun) apply blast apply(fastforce dest: sees_method_fun) apply fastforce apply fastforce apply fastforce apply(blast dest: is_lub_unique) apply fastforce apply fastforce apply blast apply fastforce apply fastforce done lemma fixes is_lub shows wt_env_mono: "is_lub,P,E \ e :: T \ (\E'. E \\<^sub>m E' \ is_lub,P,E' \ e :: T)" and wts_env_mono: "is_lub,P,E \ es [::] Ts \ (\E'. E \\<^sub>m E' \ is_lub,P,E' \ es [::] Ts)" apply(induct rule: WT_WTs.inducts) apply(simp add: WTNew) apply(simp add: WTNewArray) apply(fastforce simp: WTCast) apply(fastforce simp: WTInstanceOf) apply(fastforce simp: WTVal) apply(simp add: WTVar map_le_def dom_def) apply(fastforce simp: WTBinOp) apply(force simp:map_le_def) apply(simp add: WTAAcc) apply(simp add: WTAAss, fastforce) apply(simp add: WTALength, fastforce) apply(fastforce simp: WTFAcc) apply(fastforce simp: WTFAss del:WT_WTs.intros WT_elim_cases) apply blast apply(fastforce) apply(fastforce simp: map_le_def WTBlock) apply(fastforce simp: WTSynchronized) apply(fastforce simp: WTSeq) apply(fastforce simp: WTCond) apply(fastforce simp: WTWhile) apply(fastforce simp: WTThrow) apply(fastforce simp: WTTry map_le_def dom_def) apply(fastforce)+ done lemma fixes is_lub shows WT_fv: "is_lub,P,E \ e :: T \ fv e \ dom E" and WT_fvs: "is_lub,P,E \ es [::] Ts \ fvs es \ dom E" apply(induct rule:WT_WTs.inducts) apply(simp_all del: fun_upd_apply) apply fast+ done lemma fixes is_lub shows WT_expr_locks: "is_lub,P,E \ e :: T \ expr_locks e = (\ad. 0)" and WTs_expr_lockss: "is_lub,P,E \ es [::] Ts \ expr_lockss es = (\ad. 0)" by(induct rule: WT_WTs.inducts)(auto) lemma fixes is_lub :: "ty \ ty \ ty \ bool" ("\ lub'((_,/ _)') = _" [51,51,51] 50) assumes is_lub_is_type: "\T1 T2 T3. \ \ lub(T1, T2) = T3; is_type P T1; is_type P T2 \ \ is_type P T3" and wf: "wf_prog wf_md P" shows WT_is_type: "\ is_lub,P,E \ e :: T; ran E \ types P \ \ is_type P T" and WTs_is_type: "\ is_lub,P,E \ es [::] Ts; ran E \ types P \ \ set Ts \ types P" apply(induct rule: WT_WTs.inducts) apply simp apply simp apply simp apply simp apply (simp add:typeof_lit_is_type) apply (fastforce intro:nth_mem simp add: ran_def) apply(simp add: WT_binop_is_type) apply(simp) apply(simp del: is_type_array add: is_type_ArrayD) apply(simp) apply(simp) apply(simp add:sees_field_is_type[OF _ wf]) apply(simp) apply simp apply(fastforce dest: sees_wf_mdecl[OF wf] simp:wf_mdecl_def) apply(fastforce simp add: ran_def split: if_split_asm) apply(simp add: is_class_Object[OF wf]) apply(simp) apply(simp) apply(fastforce intro: is_lub_is_type) apply(simp) apply(simp) apply simp apply simp apply simp done lemma fixes is_lub1 :: "ty \ ty \ ty \ bool" ("\1 lub'((_,/ _)') = _" [51,51,51] 50) and is_lub2 :: "ty \ ty \ ty \ bool" ("\2 lub'((_,/ _)') = _" [51,51,51] 50) assumes wf: "wf_prog wf_md P" and is_lub1_into_is_lub2: "\T1 T2 T3. \ \1 lub(T1, T2) = T3; is_type P T1; is_type P T2 \ \ \2 lub(T1, T2) = T3" and is_lub2_is_type: "\T1 T2 T3. \ \2 lub(T1, T2) = T3; is_type P T1; is_type P T2 \ \ is_type P T3" shows WT_change_is_lub: "\ is_lub1,P,E \ e :: T; ran E \ types P \ \ is_lub2,P,E \ e :: T" and WTs_change_is_lub: "\ is_lub1,P,E \ es [::] Ts; ran E \ types P \ \ is_lub2,P,E \ es [::] Ts" proof(induct rule: WT_WTs.inducts) case (WTBlock U E V e' T vo) from \is_type P U\ \ran E \ types P\ have "ran (E(V \ U)) \ types P" by(auto simp add: ran_def) hence "is_lub2,P,E(V \ U) \ e' :: T" by(rule WTBlock) with \is_type P U\ show ?case using \case vo of None \ True | \v\ \ \T'. typeof v = \T'\ \ P \ T' \ U\ by auto next case (WTCond E e e1 T1 e2 T2 T) from \ran E \ types P\ have "is_lub2,P,E \ e :: Boolean" "is_lub2,P,E \ e1 :: T1" "is_lub2,P,E \ e2 :: T2" by(rule WTCond)+ moreover from is_lub2_is_type wf \is_lub2,P,E \ e1 :: T1\ \ran E \ types P\ have "is_type P T1" by(rule WT_is_type) from is_lub2_is_type wf \is_lub2,P,E \ e2 :: T2\ \ran E \ types P\ have "is_type P T2" by(rule WT_is_type) with \\1 lub(T1, T2) = T\ \is_type P T1\ have "\2 lub(T1, T2) = T" by(rule is_lub1_into_is_lub2) ultimately show ?case .. next case (WTTry E e1 T V C e2) from \ran E \ types P\ have "is_lub2,P,E \ e1 :: T" by(rule WTTry) moreover from \P \ C \\<^sup>* Throwable\ have "is_class P C" by(rule is_class_sub_Throwable[OF wf]) with \ran E \ types P\ have "ran (E(V \ Class C)) \ types P" by(auto simp add: ran_def) hence "is_lub2,P,E(V \ Class C) \ e2 :: T" by(rule WTTry) ultimately show ?case using \P \ C \\<^sup>* Throwable\ .. qed auto subsection \Code generator setup\ lemma WTBlock_code: "\is_lub. \ is_type P T; is_lub,P,E(V \ T) \ e :: T'; case vo of None \ True | \v\ \ case typeof v of None \ False | Some T' \ P \ T' \ T \ \ is_lub,P,E \ {V:T=vo; e} :: T'" by(auto) lemmas [code_pred_intro] = WTNew WTNewArray WTCast WTInstanceOf WTVal WTVar WTBinOp WTLAss WTAAcc WTAAss WTALength WTFAcc WTFAss WTCAS WTCall declare WTBlock_code [code_pred_intro WTBlock'] lemmas [code_pred_intro] = WTSynchronized WTSeq WTCond WTWhile WTThrow WTTry WTNil WTCons code_pred (modes: (i \ i \ o \ bool) \ i \ i \ i \ o \ bool, (i \ i \ o \ bool) \ i \ i \ i \ i \ bool) [detect_switches, skip_proof] WT proof - case WT from WT.prems show thesis proof cases case (WTBlock T V e vo) thus thesis using WT.WTBlock'[OF refl refl refl, of V T vo e] by(auto) qed(assumption|erule WT.that[OF refl refl refl]|rule refl)+ next case WTs from WTs.prems WTs.that show thesis by cases blast+ qed inductive is_lub_sup :: "'m prog \ ty \ ty \ ty \ bool" for P T1 T2 T3 where "sup P T1 T2 = OK T3 \ is_lub_sup P T1 T2 T3" code_pred (modes: i \ i \ i \ o \ bool, i \ i \ i \ i \ bool) is_lub_sup . -definition WT_code :: "'addr J_prog \ env \ 'addr expr \ ty \ bool" ("_,_ \ _ ::' _" [51,51,51] 50) +definition WT_code :: "'addr J_prog \ env \ 'addr expr \ ty \ bool" ("_,_ \ _ ::'' _" [51,51,51] 50) where "WT_code P \ WT (is_lub_sup P) P" definition WTs_code :: "'addr J_prog \ env \ 'addr expr list \ ty list \ bool" ("_,_ \ _ [::''] _" [51,51,51] 50) where "WTs_code P \ WTs (is_lub_sup P) P" lemma assumes wf: "wf_prog wf_md P" shows WT_code_into_WT: "\ P,E \ e ::' T; ran E \ types P \ \ P,E \ e :: T" and WTs_code_into_WTs: "\ P,E \ es [::'] Ts; ran E \ types P \ \ P,E \ es [::] Ts" proof - assume ran: "ran E \ types P" { assume wt: "P,E \ e ::' T" show "P,E \ e :: T" by(rule WT_change_is_lub[OF wf _ _ wt[unfolded WT_code_def] ran])(blast elim!: is_lub_sup.cases intro: sup_is_lubI[OF wf] is_lub_is_type[OF wf])+ } { assume wts: "P,E \ es [::'] Ts" show "P,E \ es [::] Ts" by(rule WTs_change_is_lub[OF wf _ _ wts[unfolded WTs_code_def] ran])(blast elim!: is_lub_sup.cases intro: sup_is_lubI[OF wf] is_lub_is_type[OF wf])+ } qed lemma assumes wf: "wf_prog wf_md P" shows WT_into_WT_code: "\ P,E \ e :: T; ran E \ types P \ \ P,E \ e ::' T" and WT_into_WTs_code_OK: "\ P,E \ es [::] Ts; ran E \ types P \ \ P,E \ es [::'] Ts" proof - assume ran: "ran E \ types P" { assume wt: "P,E \ e :: T" show "P,E \ e ::' T" unfolding WT_code_def by(rule WT_change_is_lub[OF wf _ _ wt ran])(blast intro!: is_lub_sup.intros intro: is_lub_subD[OF wf] sup_is_type[OF wf] elim!: is_lub_sup.cases)+ } { assume wts: "P,E \ es [::] Ts" show "P,E \ es [::'] Ts" unfolding WTs_code_def by(rule WTs_change_is_lub[OF wf _ _ wts ran])(blast intro!: is_lub_sup.intros intro: is_lub_subD[OF wf] sup_is_type[OF wf] elim!: is_lub_sup.cases)+ } qed theorem WT_eq_WT_code: assumes "wf_prog wf_md P" and "ran E \ types P" shows "P,E \ e :: T \ P,E \ e ::' T" using assms by(blast intro: WT_code_into_WT WT_into_WT_code) code_pred (modes: i \ i \ i \ i \ bool, i \ i \ i \ o \ bool) [inductify] WT_code . code_pred (modes: i \ i \ i \ i \ bool, i \ i \ i \ o \ bool) [inductify] WTs_code . end diff --git a/thys/JinjaThreads/MM/JMM_Type.thy b/thys/JinjaThreads/MM/JMM_Type.thy --- a/thys/JinjaThreads/MM/JMM_Type.thy +++ b/thys/JinjaThreads/MM/JMM_Type.thy @@ -1,352 +1,352 @@ (* Title: JinjaThreads/MM/JMM_Type.thy Author: Andreas Lochbihler *) section \JMM heap implementation 1\ theory JMM_Type imports "../Common/ExternalCallWF" "../Common/ConformThreaded" JMM_Heap begin subsection \Definitions\ text \ The JMM heap only stores type information. \ type_synonym 'addr JMM_heap = "'addr \ htype" translations (type) "'addr JMM_heap" <= (type) "'addr \ htype option" abbreviation jmm_empty :: "'addr JMM_heap" where "jmm_empty == Map.empty" definition jmm_allocate :: "'addr JMM_heap \ htype \ ('addr JMM_heap \ 'addr) set" where "jmm_allocate h hT = (\a. (h(a \ hT), a)) ` {a. h a = None}" definition jmm_typeof_addr :: "'addr JMM_heap \ 'addr \ htype" where "jmm_typeof_addr h = h" definition jmm_heap_read :: "'addr JMM_heap \ 'addr \ addr_loc \ 'addr val \ bool" where "jmm_heap_read h a ad v = True" context notes [[inductive_internals]] begin inductive jmm_heap_write :: "'addr JMM_heap \ 'addr \ addr_loc \ 'addr val \ 'addr JMM_heap \ bool" where "jmm_heap_write h a ad v h" end definition jmm_hconf :: "'m prog \ 'addr JMM_heap \ bool" ("_ \jmm _ \" [51,51] 50) where "P \jmm h \ \ ty_of_htype ` ran h \ {T. is_type P T}" definition jmm_allocated :: "'addr JMM_heap \ 'addr set" where "jmm_allocated h = dom (jmm_typeof_addr h)" definition jmm_spurious_wakeups :: bool where "jmm_spurious_wakeups = True" lemmas jmm_heap_ops_defs = jmm_allocate_def jmm_typeof_addr_def jmm_heap_read_def jmm_heap_write_def jmm_allocated_def jmm_spurious_wakeups_def type_synonym 'addr thread_id = "'addr" abbreviation (input) addr2thread_id :: "'addr \ 'addr thread_id" where "addr2thread_id \ \x. x" abbreviation (input) thread_id2addr :: "'addr thread_id \ 'addr" where "thread_id2addr \ \x. x" interpretation jmm: heap_base addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr jmm_heap_read jmm_heap_write . notation jmm.hext ("_ \jmm _" [51,51] 50) notation jmm.conf ("_,_ \jmm _ :\ _" [51,51,51,51] 50) notation jmm.addr_loc_type ("_,_ \jmm _@_ : _" [50, 50, 50, 50, 50] 51) notation jmm.confs ("_,_ \jmm _ [:\] _" [51,51,51,51] 50) notation jmm.tconf ("_,_ \jmm _ \t" [51,51,51] 50) text \Now a variation of the JMM with a different read operation that permits to read only type-conformant values\ interpretation jmm': heap_base addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr "jmm.heap_read_typed P" jmm_heap_write for P . -notation jmm'.hext ("_ \jmm' _" [51,51] 50) -notation jmm'.conf ("_,_ \jmm' _ :\ _" [51,51,51,51] 50) -notation jmm'.addr_loc_type ("_,_ \jmm' _@_ : _" [50, 50, 50, 50, 50] 51) -notation jmm'.confs ("_,_ \jmm' _ [:\] _" [51,51,51,51] 50) -notation jmm'.tconf ("_,_ \jmm' _ \t" [51,51,51] 50) +notation jmm'.hext ("_ \jmm'' _" [51,51] 50) +notation jmm'.conf ("_,_ \jmm'' _ :\ _" [51,51,51,51] 50) +notation jmm'.addr_loc_type ("_,_ \jmm'' _@_ : _" [50, 50, 50, 50, 50] 51) +notation jmm'.confs ("_,_ \jmm'' _ [:\] _" [51,51,51,51] 50) +notation jmm'.tconf ("_,_ \jmm'' _ \t" [51,51,51] 50) subsection \Heap locale interpretations\ subsection \Locale \heap\\ lemma jmm_heap: "heap addr2thread_id thread_id2addr jmm_allocate jmm_typeof_addr jmm_heap_write P" proof fix h' a h hT assume "(h', a) \ jmm_allocate h hT" thus "jmm_typeof_addr h' a = \hT\" by(auto simp add: jmm_heap_ops_defs) next fix h' :: "('addr :: addr) JMM_heap" and h hT a assume "(h', a) \ jmm_allocate h hT" thus "h \jmm h'" by(fastforce simp add: jmm_heap_ops_defs intro: jmm.hextI) next fix h a al v and h' :: "('addr :: addr) JMM_heap" assume "jmm_heap_write h a al v h'" thus "h \jmm h'" by cases auto qed simp interpretation jmm: heap addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr jmm_heap_read jmm_heap_write P for P by(rule jmm_heap) declare jmm.typeof_addr_thread_id2_addr_addr2thread_id [simp del] lemmas jmm'_heap = jmm_heap interpretation jmm': heap addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr "jmm.heap_read_typed P" jmm_heap_write P for P by(rule jmm'_heap) declare jmm'.typeof_addr_thread_id2_addr_addr2thread_id [simp del] subsection \Locale \heap_conf\\ interpretation jmm: heap_conf_base addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr jmm_heap_read jmm_heap_write "jmm_hconf P" P for P . -abbreviation (input) jmm'_hconf :: "'m prog \ 'addr JMM_heap \ bool" ("_ \jmm' _ \" [51,51] 50) +abbreviation (input) jmm'_hconf :: "'m prog \ 'addr JMM_heap \ bool" ("_ \jmm'' _ \" [51,51] 50) where "jmm'_hconf == jmm_hconf" interpretation jmm': heap_conf_base addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr "jmm.heap_read_typed P" jmm_heap_write "jmm'_hconf P" P for P . abbreviation jmm_heap_read_typeable :: "('addr :: addr) itself \ 'm prog \ bool" where "jmm_heap_read_typeable tytok P \ jmm.heap_read_typeable (jmm_hconf P :: 'addr JMM_heap \ bool) P" abbreviation jmm'_heap_read_typeable :: "('addr :: addr) itself \ 'm prog \ bool" where "jmm'_heap_read_typeable tytok P \ jmm'.heap_read_typeable TYPE('m) P (jmm_hconf P :: 'addr JMM_heap \ bool) P" lemma jmm_heap_read_typeable: "jmm_heap_read_typeable tytok P" by(rule jmm.heap_read_typeableI)(simp add: jmm_heap_read_def) lemma jmm'_heap_read_typeable: "jmm'_heap_read_typeable tytok P" by(rule jmm'.heap_read_typeableI)(auto simp add: jmm.heap_read_typed_def jmm_heap_read_def dest: jmm'.addr_loc_type_fun) lemma jmm_heap_conf: "heap_conf addr2thread_id thread_id2addr jmm_empty jmm_allocate jmm_typeof_addr jmm_heap_write (jmm_hconf P) P" proof show "P \jmm jmm_empty \" by(simp add: jmm_hconf_def) next fix h a hT assume "jmm_typeof_addr h a = \hT\" "P \jmm h \" thus "is_htype P hT" by(auto simp add: jmm_hconf_def jmm_heap_ops_defs intro: ranI) next fix h' h hT a assume "(h', a) \ jmm_allocate h hT" "P \jmm h \" "is_htype P hT" thus "P \jmm h' \" by(fastforce simp add: jmm_hconf_def jmm_heap_ops_defs ran_def split: if_split_asm) next fix h a al v h' T assume "jmm_heap_write h a al v h'" "P \jmm h \" and "jmm.addr_loc_type P h a al T" and "P,h \jmm v :\ T" thus "P \jmm h' \" by(cases) simp qed interpretation jmm: heap_conf addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr jmm_heap_read jmm_heap_write "jmm_hconf P" P for P by(rule jmm_heap_conf) lemmas jmm'_heap_conf = jmm_heap_conf interpretation jmm': heap_conf addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr "jmm.heap_read_typed P" jmm_heap_write "jmm'_hconf P" P for P by(rule jmm'_heap_conf) subsection \Locale \heap_progress\\ lemma jmm_heap_progress: "heap_progress addr2thread_id thread_id2addr jmm_empty jmm_allocate jmm_typeof_addr jmm_heap_read jmm_heap_write (jmm_hconf P) P" proof fix h a al T assume "P \jmm h \" and al: "jmm.addr_loc_type P h a al T" show "\v. jmm_heap_read h a al v \ P,h \jmm v :\ T" using jmm.defval_conf[of P h T] unfolding jmm_heap_ops_defs by blast next fix h a al T v assume "jmm.addr_loc_type P h a al T" show "\h'. jmm_heap_write h a al v h'" by(auto intro: jmm_heap_write.intros) qed interpretation jmm: heap_progress addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr jmm_heap_read jmm_heap_write "jmm_hconf P" P for P by(rule jmm_heap_progress) lemma jmm'_heap_progress: "heap_progress addr2thread_id thread_id2addr jmm_empty jmm_allocate jmm_typeof_addr (jmm.heap_read_typed P) jmm_heap_write (jmm'_hconf P) P" proof fix h a al T assume "P \jmm' h \" and al: "jmm'.addr_loc_type P h a al T" thus "\v. jmm.heap_read_typed P h a al v \ P,h \jmm' v :\ T" unfolding jmm.heap_read_typed_def jmm_heap_read_def by(auto dest: jmm'.addr_loc_type_fun intro: jmm'.defval_conf) next fix h a al T v assume "jmm'.addr_loc_type P h a al T" and "P,h \jmm' v :\ T" thus "\h'. jmm_heap_write h a al v h'" by(auto intro: jmm_heap_write.intros) qed interpretation jmm': heap_progress addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr "jmm.heap_read_typed P" jmm_heap_write "jmm'_hconf P" P for P by(rule jmm'_heap_progress) subsection \Locale \heap_conf_read\\ lemma jmm'_heap_conf_read: "heap_conf_read addr2thread_id thread_id2addr jmm_empty jmm_allocate jmm_typeof_addr (jmm.heap_read_typed P) jmm_heap_write (jmm'_hconf P) P" by(rule jmm.heap_conf_read_heap_read_typed) interpretation jmm': heap_conf_read addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr "jmm.heap_read_typed P" jmm_heap_write "jmm'_hconf P" P for P by(rule jmm'_heap_conf_read) interpretation jmm': heap_typesafe addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr "jmm.heap_read_typed P" jmm_heap_write "jmm'_hconf P" P for P .. subsection \Locale \allocated_heap\\ lemma jmm_allocated_heap: "allocated_heap addr2thread_id thread_id2addr jmm_empty jmm_allocate jmm_typeof_addr jmm_heap_write jmm_allocated P" proof show "jmm_allocated jmm_empty = {}" by(auto simp add: jmm_heap_ops_defs) next fix h' a h hT assume "(h', a) \ jmm_allocate h hT" thus "jmm_allocated h' = insert a (jmm_allocated h) \ a \ jmm_allocated h" by(auto simp add: jmm_heap_ops_defs split: if_split_asm) next fix h a al v h' assume "jmm_heap_write h a al v h'" thus "jmm_allocated h' = jmm_allocated h" by cases simp qed interpretation jmm: allocated_heap addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr jmm_heap_read jmm_heap_write jmm_allocated P for P by(rule jmm_allocated_heap) lemmas jmm'_allocated_heap = jmm_allocated_heap interpretation jmm': allocated_heap addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate jmm_typeof_addr "jmm.heap_read_typed P" jmm_heap_write jmm_allocated P for P by(rule jmm'_allocated_heap) subsection \Syntax translations\ -notation jmm'.external_WT' ("_,_ \jmm' (_\_'(_')) : _" [50,0,0,0,50] 60) +notation jmm'.external_WT' ("_,_ \jmm'' (_\_'(_')) : _" [50,0,0,0,50] 60) abbreviation jmm'_red_external :: "'m prog \ 'addr thread_id \ 'addr JMM_heap \ 'addr \ mname \ 'addr val list \ ('addr :: addr, 'addr thread_id, 'addr JMM_heap) external_thread_action \ 'addr extCallRet \ 'addr JMM_heap \ bool" where "jmm'_red_external P \ jmm'.red_external (TYPE('m)) P P" abbreviation jmm'_red_external_syntax :: "'m prog \ 'addr thread_id \ 'addr \ mname \ 'addr val list \ 'addr JMM_heap \ ('addr :: addr, 'addr thread_id, 'addr JMM_heap) external_thread_action \ 'addr extCallRet \ 'addr JMM_heap \ bool" - ("_,_ \jmm' (\(_\_'(_')),/_\) -_\ext (\(_),/(_)\)" [50, 0, 0, 0, 0, 0, 0, 0, 0] 51) + ("_,_ \jmm'' (\(_\_'(_')),/_\) -_\ext (\(_),/(_)\)" [50, 0, 0, 0, 0, 0, 0, 0, 0] 51) where "P,t \jmm' \a\M(vs), h\ -ta\ext \va, h'\ \ jmm'_red_external P t h a M vs ta va h'" abbreviation jmm'_red_external_aggr :: "'m prog \ 'addr thread_id \ 'addr \ mname \ 'addr val list \ 'addr JMM_heap \ (('addr :: addr, 'addr thread_id, 'addr JMM_heap) external_thread_action \ 'addr extCallRet \ 'addr JMM_heap) set" where "jmm'_red_external_aggr P \ jmm'.red_external_aggr TYPE('m) P P" abbreviation jmm'_heap_copy_loc :: "'m prog \ 'addr \ 'addr \ addr_loc \ 'addr JMM_heap \ ('addr :: addr, 'addr thread_id) obs_event list \ 'addr JMM_heap \ bool" where "jmm'_heap_copy_loc \ jmm'.heap_copy_loc TYPE('m)" abbreviation jmm'_heap_copies :: "'m prog \ 'addr \ 'addr \ addr_loc list \ 'addr JMM_heap \ ('addr :: addr, 'addr thread_id) obs_event list \ 'addr JMM_heap \ bool" where "jmm'_heap_copies \ jmm'.heap_copies TYPE('m)" abbreviation jmm'_heap_clone :: "'m prog \ 'addr JMM_heap \ 'addr \ 'addr JMM_heap \ (('addr :: addr, 'addr thread_id) obs_event list \ 'addr) option \ bool" where "jmm'_heap_clone P \ jmm'.heap_clone TYPE('m) P P" end diff --git a/thys/JinjaThreads/MM/JMM_Type2.thy b/thys/JinjaThreads/MM/JMM_Type2.thy --- a/thys/JinjaThreads/MM/JMM_Type2.thy +++ b/thys/JinjaThreads/MM/JMM_Type2.thy @@ -1,425 +1,425 @@ (* Title: JinjaThreads/MM/JMM_Type2.thy Author: Andreas Lochbihler *) section \JMM heap implementation 2\ theory JMM_Type2 imports "../Common/ExternalCallWF" "../Common/ConformThreaded" JMM_Heap begin subsection \Definitions\ datatype addr = Address htype nat \ \heap type and sequence number\ lemma rec_addr_conv_case_addr [simp]: "rec_addr = case_addr" by(auto intro!: ext split: addr.split) instantiation addr :: addr begin definition "hash_addr (a :: addr) = (case a of Address ht n \ int n)" definition "monitor_finfun_to_list (ls :: addr \f nat) = (SOME xs. set xs = {x. finfun_dom ls $ x })" instance proof fix ls :: "addr \f nat" show "set (monitor_finfun_to_list ls) = Collect (($) (finfun_dom ls))" unfolding monitor_finfun_to_list_addr_def using finite_list[OF finite_finfun_dom, where ?f.1 = "ls"] by(rule someI_ex) qed end primrec the_Address :: "addr \ htype \ nat" where "the_Address (Address hT n) = (hT, n)" text \ The JMM heap only stores which sequence numbers of a given @{typ "htype"} have already been allocated. \ type_synonym JMM_heap = "htype \ nat set" translations (type) "JMM_heap" <= (type) "htype \ nat set" definition jmm_allocate :: "JMM_heap \ htype \ (JMM_heap \ addr) set" where "jmm_allocate h hT = (let hhT = h hT in (\n. (h(hT := insert n hhT), Address hT n)) ` (- hhT))" abbreviation jmm_empty :: "JMM_heap" where "jmm_empty == (\_. {})" definition jmm_typeof_addr :: "'m prog \ JMM_heap \ addr \ htype" where "jmm_typeof_addr P h = (\hT. if is_htype P hT then Some hT else None) \ fst \ the_Address" definition jmm_typeof_addr' :: "'m prog \ addr \ htype" where "jmm_typeof_addr' P = (\hT. if is_htype P hT then Some hT else None) \ fst \ the_Address" lemma jmm_typeof_addr'_conv_jmm_type_addr: "jmm_typeof_addr' P = jmm_typeof_addr P h" by(simp add: jmm_typeof_addr_def jmm_typeof_addr'_def) lemma jmm_typeof_addr'_conv_jmm_typeof_addr: "(\_. jmm_typeof_addr' P) = jmm_typeof_addr P" by(simp add: jmm_typeof_addr_def jmm_typeof_addr'_def fun_eq_iff) lemma jmm_typeof_addr_conv_jmm_typeof_addr': "jmm_typeof_addr = (\P _. jmm_typeof_addr' P)" by(simp add: jmm_typeof_addr'_conv_jmm_typeof_addr) definition jmm_heap_read :: "JMM_heap \ addr \ addr_loc \ addr val \ bool" where "jmm_heap_read h a ad v = True" context notes [[inductive_internals]] begin inductive jmm_heap_write :: "JMM_heap \ addr \ addr_loc \ addr val \ JMM_heap \ bool" where "jmm_heap_write h a ad v h" end definition jmm_hconf :: "JMM_heap \ bool" where "jmm_hconf h \ True" definition jmm_allocated :: "JMM_heap \ addr set" where "jmm_allocated h = {Address CTn n|CTn n. n \ h CTn}" definition jmm_spurious_wakeups :: "bool" where "jmm_spurious_wakeups = True" lemmas jmm_heap_ops_defs = jmm_allocate_def jmm_typeof_addr_def jmm_heap_read_def jmm_heap_write_def jmm_allocated_def jmm_spurious_wakeups_def type_synonym thread_id = "addr" abbreviation (input) addr2thread_id :: "addr \ thread_id" where "addr2thread_id \ \x. x" abbreviation (input) thread_id2addr :: "thread_id \ addr" where "thread_id2addr \ \x. x" interpretation jmm: heap_base addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" jmm_heap_read jmm_heap_write for P . abbreviation jmm_hext :: "'m prog \ JMM_heap \ JMM_heap \ bool" ("_ \ _ \jmm _" [51,51,51] 50) where "jmm_hext \ jmm.hext TYPE('m)" abbreviation jmm_conf :: "'m prog \ JMM_heap \ addr val \ ty \ bool" ("_,_ \jmm _ :\ _" [51,51,51,51] 50) where "jmm_conf P \ jmm.conf TYPE('m) P P" abbreviation jmm_addr_loc_type :: "'m prog \ JMM_heap \ addr \ addr_loc \ ty \ bool" ("_,_ \jmm _@_ : _" [50, 50, 50, 50, 50] 51) where "jmm_addr_loc_type P \ jmm.addr_loc_type TYPE('m) P P" abbreviation jmm_confs :: "'m prog \ JMM_heap \ addr val list \ ty list \ bool" ("_,_ \jmm _ [:\] _" [51,51,51,51] 50) where "jmm_confs P \ jmm.confs TYPE('m) P P" abbreviation jmm_tconf :: "'m prog \ JMM_heap \ addr \ bool" ("_,_ \jmm _ \t" [51,51,51] 50) where "jmm_tconf P \ jmm.tconf TYPE('m) P P" interpretation jmm: allocated_heap_base addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" jmm_heap_read jmm_heap_write jmm_allocated for P . text \Now a variation of the JMM with a different read operation that permits to read only type-conformant values\ abbreviation jmm_heap_read_typed :: "'m prog \ JMM_heap \ addr \ addr_loc \ addr val \ bool" where "jmm_heap_read_typed P \ jmm.heap_read_typed TYPE('m) P P" interpretation jmm': heap_base addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" "jmm_heap_read_typed P" jmm_heap_write for P . -abbreviation jmm'_hext :: "'m prog \ JMM_heap \ JMM_heap \ bool" ("_ \ _ \jmm' _" [51,51,51] 50) +abbreviation jmm'_hext :: "'m prog \ JMM_heap \ JMM_heap \ bool" ("_ \ _ \jmm'' _" [51,51,51] 50) where "jmm'_hext \ jmm'.hext TYPE('m)" abbreviation jmm'_conf :: "'m prog \ JMM_heap \ addr val \ ty \ bool" - ("_,_ \jmm' _ :\ _" [51,51,51,51] 50) + ("_,_ \jmm'' _ :\ _" [51,51,51,51] 50) where "jmm'_conf P \ jmm'.conf TYPE('m) P P" abbreviation jmm'_addr_loc_type :: "'m prog \ JMM_heap \ addr \ addr_loc \ ty \ bool" - ("_,_ \jmm' _@_ : _" [50, 50, 50, 50, 50] 51) + ("_,_ \jmm'' _@_ : _" [50, 50, 50, 50, 50] 51) where "jmm'_addr_loc_type P \ jmm'.addr_loc_type TYPE('m) P P" abbreviation jmm'_confs :: "'m prog \ JMM_heap \ addr val list \ ty list \ bool" - ("_,_ \jmm' _ [:\] _" [51,51,51,51] 50) + ("_,_ \jmm'' _ [:\] _" [51,51,51,51] 50) where "jmm'_confs P \ jmm'.confs TYPE('m) P P" -abbreviation jmm'_tconf :: "'m prog \ JMM_heap \ addr \ bool" ("_,_ \jmm' _ \t" [51,51,51] 50) +abbreviation jmm'_tconf :: "'m prog \ JMM_heap \ addr \ bool" ("_,_ \jmm'' _ \t" [51,51,51] 50) where "jmm'_tconf P \ jmm'.tconf TYPE('m) P P" subsection \Heap locale interpretations\ subsection \Locale \heap\\ lemma jmm_heap: "heap addr2thread_id thread_id2addr jmm_allocate (jmm_typeof_addr P) jmm_heap_write P" proof fix h' a h hT assume "(h', a) \ jmm_allocate h hT" "is_htype P hT" thus "jmm_typeof_addr P h' a = \hT\" by(auto simp add: jmm_heap_ops_defs) next fix h hT h' a assume "(h', a) \ jmm_allocate h hT" thus "P \ h \jmm h'" by(auto simp add: jmm_heap_ops_defs intro: jmm.hextI) next fix h a al v h' assume "jmm_heap_write h a al v h'" thus "P \ h \jmm h'" by cases auto qed simp interpretation jmm: heap addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" jmm_heap_read jmm_heap_write P for P by(rule jmm_heap) declare jmm.typeof_addr_thread_id2_addr_addr2thread_id [simp del] lemmas jmm'_heap = jmm_heap interpretation jmm': heap addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" "jmm_heap_read_typed P" jmm_heap_write P for P by(rule jmm'_heap) declare jmm'.typeof_addr_thread_id2_addr_addr2thread_id [simp del] lemma jmm_heap_read_typed_default_val: "heap_base.heap_read_typed typeof_addr jmm_heap_read P h a al (default_val (THE T. heap_base.addr_loc_type typeof_addr P h a al T))" by(rule heap_base.heap_read_typedI)(simp_all add: heap_base.THE_addr_loc_type jmm_heap_read_def heap_base.defval_conf) lemma jmm_allocate_Eps: "(SOME ha. ha \ jmm_allocate h hT) = (h', a') \ jmm_allocate h hT \ {} \ (h', a') \ jmm_allocate h hT" by(auto dest: jmm.allocate_Eps) lemma jmm_allocate_eq_empty: "jmm_allocate h hT = {} \ h hT = UNIV" by(auto simp add: jmm_allocate_def) lemma jmm_allocate_otherD: "(h', a) \ jmm_allocate h hT \ \hT'. hT' \ hT \ h' hT' = h hT'" by(auto simp add: jmm_allocate_def) lemma jmm_start_heap_ok: "jmm.start_heap_ok" apply(simp add: jmm.start_heap_ok_def jmm.start_heap_data_def initialization_list_def sys_xcpts_list_def jmm.create_initial_object_simps) apply(split prod.split, clarify, clarsimp simp add: jmm.create_initial_object_simps jmm_allocate_eq_empty Thread_neq_sys_xcpts sys_xcpts_neqs dest!: jmm_allocate_Eps jmm_allocate_otherD)+ done subsection \Locale \heap_conf\\ interpretation jmm: heap_conf_base addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" jmm_heap_read jmm_heap_write jmm_hconf P for P . abbreviation (input) jmm'_hconf :: "JMM_heap \ bool" where "jmm'_hconf == jmm_hconf" interpretation jmm': heap_conf_base addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" "jmm_heap_read_typed P" jmm_heap_write jmm'_hconf P for P . abbreviation jmm_heap_read_typeable :: "'m prog \ bool" where "jmm_heap_read_typeable P \ jmm.heap_read_typeable TYPE('m) P jmm_hconf P" abbreviation jmm'_heap_read_typeable :: "'m prog \ bool" where "jmm'_heap_read_typeable P \ jmm'.heap_read_typeable TYPE('m) P jmm_hconf P" lemma jmm_heap_read_typeable: "jmm_heap_read_typeable P" by(rule jmm.heap_read_typeableI)(simp add: jmm_heap_read_def) lemma jmm'_heap_read_typeable: "jmm'_heap_read_typeable P" by(rule jmm'.heap_read_typeableI)(auto simp add: jmm_heap_read_def jmm.heap_read_typed_def dest: jmm'.addr_loc_type_fun) lemma jmm_heap_conf: "heap_conf addr2thread_id thread_id2addr jmm_empty jmm_allocate (jmm_typeof_addr P) jmm_heap_write jmm_hconf P" by(unfold_locales)(simp_all add: jmm_hconf_def jmm_heap_ops_defs split: if_split_asm) interpretation jmm: heap_conf addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" jmm_heap_read jmm_heap_write jmm_hconf P for P by(rule jmm_heap_conf) lemmas jmm'_heap_conf = jmm_heap_conf interpretation jmm': heap_conf addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" "jmm_heap_read_typed P" jmm_heap_write jmm'_hconf P for P by(rule jmm'_heap_conf) subsection \Locale \heap_progress\\ lemma jmm_heap_progress: "heap_progress addr2thread_id thread_id2addr jmm_empty jmm_allocate (jmm_typeof_addr P) jmm_heap_read jmm_heap_write jmm_hconf P" proof fix h a al T assume "jmm_hconf h" and al: "P,h \jmm a@al : T" show "\v. jmm_heap_read h a al v \ P,h \jmm v :\ T" using jmm.defval_conf[of P P h T] unfolding jmm_heap_ops_defs by blast next fix h a al T v assume "P,h \jmm a@al : T" show "\h'. jmm_heap_write h a al v h'" by(auto intro: jmm_heap_write.intros) qed interpretation jmm: heap_progress addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" jmm_heap_read jmm_heap_write jmm_hconf P for P by(rule jmm_heap_progress) lemma jmm'_heap_progress: "heap_progress addr2thread_id thread_id2addr jmm_empty jmm_allocate (jmm_typeof_addr P) (jmm_heap_read_typed P) jmm_heap_write jmm'_hconf P" proof fix h a al T assume "jmm'_hconf h" and al: "P,h \jmm' a@al : T" thus "\v. jmm_heap_read_typed P h a al v \ P,h \jmm' v :\ T" unfolding jmm_heap_read_def jmm.heap_read_typed_def by(blast dest: jmm'.addr_loc_type_fun intro: jmm'.defval_conf)+ next fix h a al T v assume "P,h \jmm' a@al : T" and "P,h \jmm' v :\ T" thus "\h'. jmm_heap_write h a al v h'" by(auto intro: jmm_heap_write.intros) qed interpretation jmm': heap_progress addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" "jmm_heap_read_typed P" jmm_heap_write jmm'_hconf P for P by(rule jmm'_heap_progress) subsection \Locale \heap_conf_read\\ lemma jmm'_heap_conf_read: "heap_conf_read addr2thread_id thread_id2addr jmm_empty jmm_allocate (jmm_typeof_addr P) (jmm_heap_read_typed P) jmm_heap_write jmm'_hconf P" by(rule jmm.heap_conf_read_heap_read_typed) interpretation jmm': heap_conf_read addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" "jmm_heap_read_typed P" jmm_heap_write jmm'_hconf P for P by(rule jmm'_heap_conf_read) interpretation jmm': heap_typesafe addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" "jmm_heap_read_typed P" jmm_heap_write jmm'_hconf P for P .. subsection \Locale \allocated_heap\\ lemma jmm_allocated_heap: "allocated_heap addr2thread_id thread_id2addr jmm_empty jmm_allocate (jmm_typeof_addr P) jmm_heap_write jmm_allocated P" proof show "jmm_allocated jmm_empty = {}" by(simp add: jmm_allocated_def) next fix h' a h hT assume "(h', a) \ jmm_allocate h hT" thus "jmm_allocated h' = insert a (jmm_allocated h) \ a \ jmm_allocated h" by(auto simp add: jmm_heap_ops_defs split: if_split_asm) next fix h a al v h' assume "jmm_heap_write h a al v h'" thus "jmm_allocated h' = jmm_allocated h" by cases simp qed interpretation jmm: allocated_heap addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" jmm_heap_read jmm_heap_write jmm_allocated P for P by(rule jmm_allocated_heap) lemmas jmm'_allocated_heap = jmm_allocated_heap interpretation jmm': allocated_heap addr2thread_id thread_id2addr jmm_spurious_wakeups jmm_empty jmm_allocate "jmm_typeof_addr P" "jmm_heap_read_typed P" jmm_heap_write jmm_allocated P for P by(rule jmm'_allocated_heap) subsection \Syntax translations\ -notation jmm'.external_WT' ("_,_ \jmm' (_\_'(_')) : _" [50,0,0,0,50] 60) +notation jmm'.external_WT' ("_,_ \jmm'' (_\_'(_')) : _" [50,0,0,0,50] 60) abbreviation jmm'_red_external :: "'m prog \ thread_id \ JMM_heap \ addr \ mname \ addr val list \ (addr, thread_id, JMM_heap) external_thread_action \ addr extCallRet \ JMM_heap \ bool" where "jmm'_red_external P \ jmm'.red_external (TYPE('m)) P P" abbreviation jmm'_red_external_syntax :: "'m prog \ thread_id \ addr \ mname \ addr val list \ JMM_heap \ (addr, thread_id, JMM_heap) external_thread_action \ addr extCallRet \ JMM_heap \ bool" - ("_,_ \jmm' (\(_\_'(_')),/_\) -_\ext (\(_),/(_)\)" [50, 0, 0, 0, 0, 0, 0, 0, 0] 51) + ("_,_ \jmm'' (\(_\_'(_')),/_\) -_\ext (\(_),/(_)\)" [50, 0, 0, 0, 0, 0, 0, 0, 0] 51) where "P,t \jmm' \a\M(vs), h\ -ta\ext \va, h'\ \ jmm'_red_external P t h a M vs ta va h'" abbreviation jmm'_red_external_aggr :: "'m prog \ thread_id \ addr \ mname \ addr val list \ JMM_heap \ ((addr, thread_id, JMM_heap) external_thread_action \ addr extCallRet \ JMM_heap) set" where "jmm'_red_external_aggr P \ jmm'.red_external_aggr TYPE('m) P P" abbreviation jmm'_heap_copy_loc :: "'m prog \ addr \ addr \ addr_loc \ JMM_heap \ (addr, thread_id) obs_event list \ JMM_heap \ bool" where "jmm'_heap_copy_loc \ jmm'.heap_copy_loc TYPE('m)" abbreviation jmm'_heap_copies :: "'m prog \ addr \ addr \ addr_loc list \ JMM_heap \ (addr, thread_id) obs_event list \ JMM_heap \ bool" where "jmm'_heap_copies \ jmm'.heap_copies TYPE('m)" abbreviation jmm'_heap_clone :: "'m prog \ JMM_heap \ addr \ JMM_heap \ ((addr, thread_id) obs_event list \ addr) option \ bool" where "jmm'_heap_clone P \ jmm'.heap_clone TYPE('m) P P" end diff --git a/thys/LTL/Rewriting.thy b/thys/LTL/Rewriting.thy --- a/thys/LTL/Rewriting.thy +++ b/thys/LTL/Rewriting.thy @@ -1,1052 +1,1052 @@ (* Author: Salomon Sickert Author: Benedikt Seidl Author: Alexander Schimpf (original entry: CAVA/LTL_Rewrite.thy) License: BSD *) section \Rewrite Rules for LTL Simplification\ theory Rewriting imports LTL "HOL-Library.Extended_Nat" begin text \This theory provides rewrite rules for the simplification of LTL formulas. It supports: \begin{itemize} \item Constants Removal \item @{const Next_ltln}-Normalisation \item Modal Simplification (based on pure eventual, pure universal, or suspendable formulas) \item Syntactic Implication Checking \end{itemize} It reuses parts of LTL\_Rewrite.thy (CAVA, LTL\_TO\_GBA). Furthermore, some rules were taken from \cite{DBLP:conf/cav/SomenziB00} and \cite{DBLP:conf/tacas/BabiakKRS12}. All functions are defined for @{type ltln}.\ subsection \Constant Eliminating Constructors\ definition mk_and where "mk_and x y \ case x of false\<^sub>n \ false\<^sub>n | true\<^sub>n \ y | _ \ (case y of false\<^sub>n \ false\<^sub>n | true\<^sub>n \ x | _ \ x and\<^sub>n y)" definition mk_or where "mk_or x y \ case x of false\<^sub>n \ y | true\<^sub>n \ true\<^sub>n | _ \ (case y of true\<^sub>n \ true\<^sub>n | false\<^sub>n \ x | _ \ x or\<^sub>n y)" fun remove_strong_ops where "remove_strong_ops (x U\<^sub>n y) = remove_strong_ops y" | "remove_strong_ops (x M\<^sub>n y) = x and\<^sub>n y" | "remove_strong_ops (x or\<^sub>n y) = remove_strong_ops x or\<^sub>n remove_strong_ops y" | "remove_strong_ops x = x" fun remove_weak_ops where "remove_weak_ops (x R\<^sub>n y) = remove_weak_ops y" | "remove_weak_ops (x W\<^sub>n y) = x or\<^sub>n y" | "remove_weak_ops (x and\<^sub>n y) = remove_weak_ops x and\<^sub>n remove_weak_ops y" | "remove_weak_ops x = x" definition mk_finally where "mk_finally x \ case x of true\<^sub>n \ true\<^sub>n | false\<^sub>n \ false\<^sub>n | _ \ F\<^sub>n (remove_strong_ops x)" definition mk_globally where "mk_globally x \ case x of true\<^sub>n \ true\<^sub>n | false\<^sub>n \ false\<^sub>n | _ \ G\<^sub>n (remove_weak_ops x)" definition mk_until where "mk_until x y \ case x of false\<^sub>n \ y | true\<^sub>n \ mk_finally y | _ \ (case y of true\<^sub>n \ true\<^sub>n | false\<^sub>n \ false\<^sub>n | _ \ x U\<^sub>n y)" definition mk_release where "mk_release x y \ case x of true\<^sub>n \ y | false\<^sub>n \ mk_globally y | _ \ (case y of true\<^sub>n \ true\<^sub>n | false\<^sub>n \ false\<^sub>n | _ \ x R\<^sub>n y)" definition mk_weak_until where "mk_weak_until x y \ case y of true\<^sub>n \ true\<^sub>n | false\<^sub>n \ mk_globally x | _ \ (case x of true\<^sub>n \ true\<^sub>n | false\<^sub>n \ y | _ \ x W\<^sub>n y)" definition mk_strong_release where "mk_strong_release x y \ case y of false\<^sub>n \ false\<^sub>n | true\<^sub>n \ mk_finally x | _ \ (case x of true\<^sub>n \ y | false\<^sub>n \ false\<^sub>n | _ \ x M\<^sub>n y)" definition mk_next where "mk_next x \ case x of true\<^sub>n \ true\<^sub>n | false\<^sub>n \ false\<^sub>n | _ \ X\<^sub>n x" -definition mk_next_pow ("X\<^sub>n'") +definition mk_next_pow ("X\<^sub>n''") where "mk_next_pow n x \ case x of true\<^sub>n \ true\<^sub>n | false\<^sub>n \ false\<^sub>n | _ \ (Next_ltln ^^ n) x" lemma mk_and_semantics [simp]: "w \\<^sub>n mk_and x y \ w \\<^sub>n x and\<^sub>n y" unfolding mk_and_def by (cases x; cases y; simp) lemma mk_or_semantics [simp]: "w \\<^sub>n mk_or x y \ w \\<^sub>n x or\<^sub>n y" unfolding mk_or_def by (cases x; cases y; simp) lemma remove_strong_ops_sound [simp]: "w \\<^sub>n F\<^sub>n (remove_strong_ops y) \ w \\<^sub>n F\<^sub>n y" by (induction y arbitrary: w) (auto; force)+ lemma remove_weak_ops_sound [simp]: "w \\<^sub>n G\<^sub>n (remove_weak_ops y) \ w \\<^sub>n G\<^sub>n y" by (induction y arbitrary: w) (auto; force)+ lemma mk_finally_semantics [simp]: "w \\<^sub>n mk_finally x \ w \\<^sub>n F\<^sub>n x" by (simp add: mk_finally_def del: semantics_ltln.simps(8,9) remove_strong_ops.simps split: ltln.splits) lemma mk_globally_semantics [simp]: "w \\<^sub>n mk_globally x \ w \\<^sub>n G\<^sub>n x" by (simp add: mk_globally_def del: semantics_ltln.simps(8,9) remove_weak_ops.simps split: ltln.splits) lemma mk_until_semantics [simp]: "w \\<^sub>n mk_until x y \ w \\<^sub>n x U\<^sub>n y" proof (cases x) case (True_ltln) show ?thesis unfolding True_ltln mk_until_def by (cases y) auto next case (False_ltln) thus ?thesis by (force simp: mk_until_def) qed (cases y; force simp: mk_until_def)+ lemma mk_release_semantics [simp]: "w \\<^sub>n mk_release x y \ w \\<^sub>n x R\<^sub>n y" proof (cases x) case (False_ltln) thus ?thesis unfolding False_ltln mk_release_def by (cases y) auto next case (True_ltln) thus ?thesis by (force simp: mk_release_def) qed (cases y; force simp: mk_release_def)+ lemma mk_weak_until_semantics [simp]: "w \\<^sub>n mk_weak_until x y \ w \\<^sub>n x W\<^sub>n y" proof (cases y) case (False_ltln) thus ?thesis unfolding False_ltln mk_weak_until_def by (cases x) auto next case (True_ltln) thus ?thesis by (force simp: mk_weak_until_def) qed (cases x; force simp: mk_weak_until_def)+ lemma mk_strong_release_semantics [simp]: "w \\<^sub>n mk_strong_release x y \ w \\<^sub>n x M\<^sub>n y" proof (cases y) case (True_ltln) show ?thesis unfolding True_ltln mk_strong_release_def by (cases x) auto next case (False_ltln) thus ?thesis by (force simp: mk_strong_release_def) qed (cases x; force simp: mk_strong_release_def)+ lemma mk_next_semantics [simp]: "w \\<^sub>n mk_next x \ w \\<^sub>n X\<^sub>n x" unfolding mk_next_def by (cases x; auto) lemma mk_next_pow_semantics [simp]: "w \\<^sub>n mk_next_pow i x \ suffix i w \\<^sub>n x" by (induction i arbitrary: w; cases x) (auto simp: mk_next_pow_def) lemma mk_next_pow_simp [simp, code_unfold]: "mk_next_pow 0 x = x" "mk_next_pow 1 x = mk_next x" by (cases x; simp add: mk_next_pow_def mk_next_def)+ subsection \Constant Propagation\ fun is_constant :: "'a ltln \ bool" where "is_constant true\<^sub>n = True" | "is_constant false\<^sub>n = True" | "is_constant _ = False" lemma is_constant_constructorsI: "is_constant x \ is_constant y \ is_constant (mk_and x y)" "\is_constant x \ \is_constant y \ \is_constant (mk_and x y)" "is_constant x \ is_constant y \ is_constant (mk_or x y)" "\is_constant x \ \is_constant y \ \is_constant (mk_or x y)" "is_constant x \ is_constant (mk_finally x)" "\is_constant x \ \is_constant (mk_finally x)" "is_constant x \ is_constant (mk_globally x)" "\is_constant x \ \is_constant (mk_globally x)" "is_constant x \ is_constant (mk_until y x)" "\is_constant x \ \is_constant (mk_until y x)" "is_constant x \ is_constant (mk_release y x)" "\is_constant x \ \is_constant (mk_release y x)" "is_constant x \ is_constant y \ is_constant (mk_weak_until x y)" "\is_constant x \ \is_constant y \ \is_constant (mk_weak_until x y)" "is_constant x \ is_constant y \ is_constant (mk_strong_release x y)" "\is_constant x \ \is_constant y \ \is_constant (mk_strong_release x y)" "is_constant x \ is_constant (mk_next x)" "\is_constant x \ \is_constant (mk_next x)" "is_constant x \ is_constant (mk_next_pow n x)" by (cases x; cases y; simp add: mk_and_def mk_or_def mk_finally_def mk_globally_def mk_until_def mk_release_def mk_weak_until_def mk_strong_release_def mk_next_def mk_next_pow_def)+ lemma is_constant_constructors_simps: "mk_next_pow n x = false\<^sub>n \ x = false\<^sub>n" "mk_next_pow n x = true\<^sub>n \ x = true\<^sub>n" "is_constant (mk_next_pow n x) \ is_constant x" by (induction n) (cases x; simp add: mk_next_pow_def)+ lemma is_constant_constructors_simps2: "is_constant (mk_and x y) \ (x = true\<^sub>n \ y = true\<^sub>n \ x = false\<^sub>n \ y = false\<^sub>n)" "is_constant (mk_or x y) \ (x = false\<^sub>n \ y = false\<^sub>n \ x = true\<^sub>n \ y = true\<^sub>n)" "is_constant (mk_finally x) \ is_constant x" "is_constant (mk_globally x) \ is_constant x" "is_constant (mk_until y x) \ is_constant x" "is_constant (mk_release y x) \ is_constant x" "is_constant (mk_next x) \ is_constant x" by ((cases x; cases y; simp add: mk_and_def), (cases x; cases y; simp add: mk_or_def), (meson is_constant_constructorsI)+) lemma is_constant_constructors_simps3: "is_constant (mk_weak_until x y) \ (x = false\<^sub>n \ y = false\<^sub>n \ x = true\<^sub>n \ y = true\<^sub>n)" "is_constant (mk_strong_release x y) \ (x = true\<^sub>n \ y = true\<^sub>n \ x = false\<^sub>n \ y = false\<^sub>n)" by (cases x; cases y; simp add: mk_weak_until_def mk_strong_release_def is_constant_constructors_simps2)+ lemma is_constant_semantics: "is_constant \ \ ((\w. w \\<^sub>n \) \ \(\w. w \\<^sub>n \))" by (cases \) auto lemma until_constant_simp: "is_constant \ \ w \\<^sub>n \ U\<^sub>n \ \ w \\<^sub>n \" by (cases \) auto lemma release_constant_simp: "is_constant \ \ w \\<^sub>n \ R\<^sub>n \ \ w \\<^sub>n \" by (cases \) auto lemma mk_next_pow_dist: "mk_next_pow (i + j) \ = mk_next_pow i (mk_next_pow j \)" by (cases j; simp) (cases \; simp add: mk_next_pow_def funpow_add; simp add: funpow_swap1) lemma mk_next_pow_until: "suffix (min i j) w \\<^sub>n (mk_next_pow (i - j) \) U\<^sub>n (mk_next_pow (j - i) \) \ w \\<^sub>n (mk_next_pow i \) U\<^sub>n (mk_next_pow j \)" by (simp add: mk_next_pow_dist min_def add.commute) lemma mk_next_pow_release: "suffix (min i j) w \\<^sub>n (mk_next_pow (i - j) \) R\<^sub>n (mk_next_pow (j - i) \) \ w \\<^sub>n (mk_next_pow i \) R\<^sub>n (mk_next_pow j \)" by (simp add: mk_next_pow_dist min_def add.commute) subsection \X-Normalisation\ text \The following rewrite functions pulls the X-operator up in the syntax tree. This preprocessing step enables the removal of X-operators in front of suspendable formulas. Furthermore constants are removed as far as possible.\ fun the_enat_0 :: "enat \ nat" where "the_enat_0 i = i" | "the_enat_0 \ = 0" lemma the_enat_0_simp [simp]: "the_enat_0 0 = 0" "the_enat_0 1 = 1" by (simp add: zero_enat_def one_enat_def)+ fun combine :: "('a ltln \ 'a ltln \ 'a ltln) \ ('a ltln * enat) \ ('a ltln * enat) \ ('a ltln * enat)" where "combine binop (\, i) (\, j) = ( let \ = binop (mk_next_pow (the_enat_0 (i - j)) \) (mk_next_pow (the_enat_0 (j - i)) \) in (\, if is_constant \ then \ else min i j))" lemma fst_combine: "fst (combine binop (\, i) (\, j)) = binop (mk_next_pow (the_enat_0 (i - j)) \) (mk_next_pow (the_enat_0 (j - i)) \)" unfolding combine.simps by (meson fstI) abbreviation to_ltln :: "('a ltln * enat) \ 'a ltln" where "to_ltln x \ mk_next_pow (the_enat_0 (snd x)) (fst x)" fun rewrite_X_enat :: "'a ltln \ ('a ltln * enat)" where "rewrite_X_enat true\<^sub>n = (true\<^sub>n, \)" | "rewrite_X_enat false\<^sub>n = (false\<^sub>n, \)" | "rewrite_X_enat prop\<^sub>n(a) = (prop\<^sub>n(a), 0)" | "rewrite_X_enat nprop\<^sub>n(a) = (nprop\<^sub>n(a), 0)" | "rewrite_X_enat (\ and\<^sub>n \) = combine mk_and (rewrite_X_enat \) (rewrite_X_enat \)" | "rewrite_X_enat (\ or\<^sub>n \) = combine mk_or (rewrite_X_enat \) (rewrite_X_enat \)" | "rewrite_X_enat (\ U\<^sub>n \) = combine mk_until (rewrite_X_enat \) (rewrite_X_enat \)" | "rewrite_X_enat (\ R\<^sub>n \) = combine mk_release (rewrite_X_enat \) (rewrite_X_enat \)" | "rewrite_X_enat (\ W\<^sub>n \) = combine mk_weak_until (rewrite_X_enat \) (rewrite_X_enat \)" | "rewrite_X_enat (\ M\<^sub>n \) = combine mk_strong_release (rewrite_X_enat \) (rewrite_X_enat \)" | "rewrite_X_enat (X\<^sub>n \) = (\(\, n). (\, eSuc n)) (rewrite_X_enat \)" definition "rewrite_X \ = to_ltln (rewrite_X_enat \)" lemma combine_infinity_invariant: assumes "i = \ \ is_constant x" assumes "j = \ \ is_constant y" shows "combine mk_and (x, i) (y, j) = (z, k) \ (k = \ \ is_constant z)" and "combine mk_or (x, i) (y, j) = (z, k) \ (k = \ \ is_constant z)" and "combine mk_until (x, i) (y, j) = (z, k) \ (k = \ \ is_constant z)" and "combine mk_release (x, i) (y, j) = (z, k) \ (k = \ \ is_constant z)" and "combine mk_weak_until (x, i) (y, j) = (z, k) \ (k = \ \ is_constant z)" and "combine mk_strong_release (x, i) (y, j) = (z, k) \ (k = \ \ is_constant z)" by (cases i; cases j; simp add: assms Let_def; force intro: is_constant_constructorsI)+ lemma combine_and_or_semantics: assumes "i = \ \ is_constant \" assumes "j = \ \ is_constant \" shows "w \\<^sub>n to_ltln (combine mk_and (\, i) (\, j)) \ w \\<^sub>n to_ltln (\, i) and\<^sub>n to_ltln (\, j)" and "w \\<^sub>n to_ltln (combine mk_or (\, i) (\, j)) \ w \\<^sub>n to_ltln (\, i) or\<^sub>n to_ltln (\, j)" by ((cases i; cases j; simp add: min_def is_constant_constructors_simps is_constant_constructors_simps2 assms), (cases \; insert assms; auto), (cases \; insert assms; auto), (blast elim!: is_constant.elims))+ lemma combine_until_release_semantics: assumes "i = \ \ is_constant \" assumes "j = \ \ is_constant \" shows "w \\<^sub>n to_ltln (combine mk_until (\, i) (\, j)) \ w \\<^sub>n to_ltln (\, i) U\<^sub>n to_ltln (\, j)" and "w \\<^sub>n to_ltln (combine mk_release (\, i) (\, j)) \ w \\<^sub>n to_ltln (\, i) R\<^sub>n to_ltln (\, j)" by ((cases i; cases j; simp add: is_constant_constructors_simps is_constant_constructors_simps2 until_constant_simp release_constant_simp mk_next_pow_until mk_next_pow_release del: semantics_ltln.simps), (blast dest: is_constant_semantics), (cases \; simp add: assms), (cases \; insert assms; auto simp: add.commute))+ (* TODO this proof is a bit slow and could be optimized *) lemma combine_weak_until_strong_release_semantics: assumes "i = \ \ is_constant \" assumes "j = \ \ is_constant \" shows "w \\<^sub>n to_ltln (combine mk_weak_until (\, i) (\, j)) \ w \\<^sub>n to_ltln (\, i) W\<^sub>n to_ltln (\, j)" and "w \\<^sub>n to_ltln (combine mk_strong_release (\, i) (\, j)) \ w \\<^sub>n to_ltln (\, i) M\<^sub>n to_ltln (\, j)" by ((cases i; cases j; simp add: min_def is_constant_constructors_simps is_constant_constructors_simps3 del: semantics_ltln.simps), (cases \; simp add: assms), (cases \; insert assms; auto simp: add.commute))+ lemma rewrite_X_enat_infinity_invariant: "snd (rewrite_X_enat \) = \ \ is_constant (fst (rewrite_X_enat \))" proof (induction \) case (And_ltln \ \) thus ?case by (simp add: combine_infinity_invariant[OF And_ltln(1,2), unfolded prod.collapse]) next case (Or_ltln \ \) thus ?case by (simp add: combine_infinity_invariant[OF Or_ltln(1,2), unfolded prod.collapse]) next case (Until_ltln \ \) thus ?case by (simp add: combine_infinity_invariant[OF Until_ltln(1,2), unfolded prod.collapse]) next case (Release_ltln \ \) thus ?case by (simp add: combine_infinity_invariant[OF Release_ltln(1,2), unfolded prod.collapse]) next case (WeakUntil_ltln \ \) thus ?case by (simp add: combine_infinity_invariant[OF WeakUntil_ltln(1,2), unfolded prod.collapse]) next case (StrongRelease_ltln \ \) thus ?case by (simp add: combine_infinity_invariant[OF StrongRelease_ltln(1,2), unfolded prod.collapse]) next case (Next_ltln \) thus ?case by (simp add: split_def) (metis eSuc_infinity eSuc_inject) qed auto lemma rewrite_X_enat_correct: "w \\<^sub>n \ \ w \\<^sub>n to_ltln (rewrite_X_enat \)" proof (induction \ arbitrary: w) case (And_ltln \ \) thus ?case using combine_and_or_semantics[OF rewrite_X_enat_infinity_invariant rewrite_X_enat_infinity_invariant] by fastforce next case (Or_ltln \ \) thus ?case using combine_and_or_semantics[OF rewrite_X_enat_infinity_invariant rewrite_X_enat_infinity_invariant] by fastforce next case (Until_ltln \ \) thus ?case unfolding rewrite_X_enat.simps combine_until_release_semantics[OF rewrite_X_enat_infinity_invariant rewrite_X_enat_infinity_invariant, unfolded prod.collapse] by fastforce next case (Release_ltln \ \) thus ?case unfolding rewrite_X_enat.simps combine_until_release_semantics[OF rewrite_X_enat_infinity_invariant rewrite_X_enat_infinity_invariant, unfolded prod.collapse] by fastforce next case (WeakUntil_ltln \ \) thus ?case unfolding rewrite_X_enat.simps combine_weak_until_strong_release_semantics[OF rewrite_X_enat_infinity_invariant rewrite_X_enat_infinity_invariant, unfolded prod.collapse] by fastforce next case (StrongRelease_ltln \ \) thus ?case unfolding rewrite_X_enat.simps combine_weak_until_strong_release_semantics[OF rewrite_X_enat_infinity_invariant rewrite_X_enat_infinity_invariant, unfolded prod.collapse] by fastforce next case (Next_ltln \) moreover have " w \\<^sub>n to_ltln (rewrite_X_enat (X\<^sub>n \)) \ suffix 1 w \\<^sub>n to_ltln (rewrite_X_enat \)" by (simp add: split_def; cases "snd (rewrite_X_enat \) \ \") (auto simp: eSuc_def, auto simp: rewrite_X_enat_infinity_invariant eSuc_def dest: is_constant_semantics) ultimately show ?case using semantics_ltln.simps(7) by blast qed auto lemma rewrite_X_sound [simp]: "w \\<^sub>n rewrite_X \ \ w \\<^sub>n \" using rewrite_X_enat_correct unfolding rewrite_X_def Let_def by auto subsection \Pure Eventual, Pure Universal, and Suspendable Formulas\ fun pure_eventual :: "'a ltln \ bool" where "pure_eventual true\<^sub>n = True" | "pure_eventual false\<^sub>n = True" | "pure_eventual (\ and\<^sub>n \') = (pure_eventual \ \ pure_eventual \')" | "pure_eventual (\ or\<^sub>n \') = (pure_eventual \ \ pure_eventual \')" | "pure_eventual (\ U\<^sub>n \') = (\ = true\<^sub>n \ pure_eventual \')" | "pure_eventual (\ R\<^sub>n \') = (pure_eventual \ \ pure_eventual \')" | "pure_eventual (\ W\<^sub>n \') = (pure_eventual \ \ pure_eventual \')" | "pure_eventual (\ M\<^sub>n \') = (pure_eventual \ \ pure_eventual \' \ \' = true\<^sub>n)" | "pure_eventual (X\<^sub>n \) = pure_eventual \" | "pure_eventual _ = False" fun pure_universal :: "'a ltln \ bool" where "pure_universal true\<^sub>n = True" | "pure_universal false\<^sub>n = True" | "pure_universal (\ and\<^sub>n \') = (pure_universal \ \ pure_universal \')" | "pure_universal (\ or\<^sub>n \') = (pure_universal \ \ pure_universal \')" | "pure_universal (\ U\<^sub>n \') = (pure_universal \ \ pure_universal \')" | "pure_universal (\ R\<^sub>n \') = (\ = false\<^sub>n \ pure_universal \')" | "pure_universal (\ W\<^sub>n \') = (pure_universal \ \ pure_universal \' \ \' = false\<^sub>n)" | "pure_universal (\ M\<^sub>n \') = (pure_universal \ \ pure_universal \')" | "pure_universal (X\<^sub>n \) = pure_universal \" | "pure_universal _ = False" fun suspendable :: "'a ltln \ bool" where "suspendable true\<^sub>n = True" | "suspendable false\<^sub>n = True" | "suspendable (\ and\<^sub>n \') = (suspendable \ \ suspendable \')" | "suspendable (\ or\<^sub>n \') = (suspendable \ \ suspendable \')" | "suspendable (\ U\<^sub>n \) = (\ = true\<^sub>n \ pure_universal \ \ suspendable \)" | "suspendable (\ R\<^sub>n \) = (\ = false\<^sub>n \ pure_eventual \ \ suspendable \)" | "suspendable (\ W\<^sub>n \) = (suspendable \ \ suspendable \ \ pure_eventual \ \ \ = false\<^sub>n)" | "suspendable (\ M\<^sub>n \) = (suspendable \ \ suspendable \ \ pure_universal \ \ \ = true\<^sub>n)" | "suspendable (X\<^sub>n \) = suspendable \" | "suspendable _ = False" lemma pure_eventual_left_append: "pure_eventual \ \ w \\<^sub>n \ \ (u \ w) \\<^sub>n \" proof (induction \ arbitrary: u w) case (Until_ltln \ \') moreover then obtain i where "suffix i w \\<^sub>n \'" by auto hence "\ = true\<^sub>n \ ?case" by simp (metis suffix_conc_length suffix_suffix) moreover have "pure_eventual \' \ (u \ w) \\<^sub>n \'" by (metis \suffix i w \\<^sub>n \'\ Until_ltln(2) prefix_suffix) hence "pure_eventual \' \ ?case" by force ultimately show ?case by fastforce next case (Release_ltln \ \') thus ?case by (cases "\i. suffix i w \\<^sub>n \'"; simp_all) (metis linear suffix_conc_snd gr0I not_less0 prefix_suffix suffix_0)+ next case (WeakUntil_ltln \ \') thus ?case by (cases "\i. suffix i w \\<^sub>n \'"; simp_all) (metis zero_le le0 nat_le_linear prefix_suffix suffix_0 suffix_conc_length suffix_conc_snd suffix_subseq_join)+ next case (StrongRelease_ltln \ \') moreover then obtain i where "suffix i w \\<^sub>n \ and\<^sub>n \'" by auto hence "\' = true\<^sub>n \ ?case" by simp (metis suffix_conc_length suffix_suffix) moreover have "pure_eventual \ \ pure_eventual \' \ (u \ w) \\<^sub>n \ and\<^sub>n \'" by (metis \suffix i w \\<^sub>n \ and\<^sub>n \'\ calculation(1) calculation(2) prefix_suffix semantics_ltln.simps(5)) hence "pure_eventual \ \ pure_eventual \' \ ?case" by force ultimately show ?case by fastforce qed (auto, metis diff_zero le_0_eq not_less_eq_eq suffix_conc_length suffix_conc_snd word_split) lemma pure_universal_suffix_closed: "pure_universal \ \ (u \ w) \\<^sub>n \ \ w \\<^sub>n \" proof (induction \ arbitrary: u w) case (Until_ltln \ \') hence "\i. suffix i (u \ w) \\<^sub>n \' \ (\j w) \\<^sub>n \)" using semantics_ltln.simps(8) by blast thus ?case by simp (metis Until_ltln(1-3) le_0_eq le_eq_less_or_eq le_less_linear prefix_suffix pure_universal.simps(5) suffix_conc_fst suffix_conc_snd) next case (Release_ltln \ \') moreover hence "\i. suffix i (u \ w) \\<^sub>n \' \ (\j w) \\<^sub>n \)" by simp ultimately show ?case by simp (metis semantics_ltln.simps(2) not_less0 prefix_suffix suffix_0 suffix_conc_length suffix_suffix) next case (WeakUntil_ltln \ \') moreover hence "\i. suffix i (u \ w) \\<^sub>n \ \ (\j\i. suffix j (u \ w) \\<^sub>n \')" by simp ultimately show ?case by simp (metis (full_types) le_antisym prefix_suffix semantics_ltln.simps(2) suffix_0 suffix_conc_length suffix_suffix zero_le) next case (StrongRelease_ltln \ \') hence "\i. suffix i (u \ w) \\<^sub>n \ \ (\j\i. suffix j (u \ w) \\<^sub>n \')" using semantics_ltln.simps(11) by blast thus ?case by simp (metis StrongRelease_ltln(1-3) diff_is_0_eq nat_le_linear prefix_conc_length prefix_suffix pure_universal.simps(8) subsequence_length suffix_conc_snd suffix_subseq_join) next case (Next_ltln \) thus ?case by (metis prefix_suffix pure_universal.simps(9) semantics_ltln.simps(7) semiring_normalization_rules(24) suffix_conc_length suffix_suffix) qed auto lemma suspendable_prefix_invariant: "suspendable \ \ (u \ w) \\<^sub>n \ \ w \\<^sub>n \" proof (induction \ arbitrary: u w) case (Until_ltln \ \') show ?case proof (cases "suspendable \'") case False hence "\ = true\<^sub>n" and "pure_universal \'" using Until_ltln by simp+ thus ?thesis by (simp; metis (no_types) linear pure_universal_suffix_closed suffix_conc_fst suffix_conc_length suffix_conc_snd suffix_suffix) qed (simp; metis Until_ltln(2) not_less0 prefix_suffix) next case (Release_ltln \ \') show ?case proof (cases "suspendable \'") case False hence "\ = false\<^sub>n" and "pure_eventual \'" using Release_ltln by simp+ thus ?thesis by (simp; metis (no_types) le_iff_add add_diff_cancel_left' linear pure_eventual_left_append suffix_0 suffix_conc_fst suffix_conc_snd) qed (simp; metis Release_ltln(2) not_less0 prefix_suffix) next case (WeakUntil_ltln \ \') show ?case proof (cases "suspendable \ \ suspendable \'") case False hence "\' = false\<^sub>n" and "pure_eventual \" using WeakUntil_ltln by simp+ thus ?thesis by (simp; metis (no_types) le_iff_add add_diff_cancel_left' linear pure_eventual_left_append suffix_0 suffix_conc_fst suffix_conc_snd) qed (simp; metis (full_types) WeakUntil_ltln.IH prefix_suffix) next case (StrongRelease_ltln \ \') show ?case proof (cases "suspendable \ \ suspendable \'") case False hence "\' = true\<^sub>n" and "pure_universal \" using StrongRelease_ltln by simp+ thus ?thesis by (simp; metis (no_types) linear pure_universal_suffix_closed suffix_conc_fst suffix_conc_length suffix_conc_snd suffix_suffix) qed (simp; metis (full_types) StrongRelease_ltln.IH(1) StrongRelease_ltln.IH(2) prefix_suffix) qed (simp_all, metis prefix_suffix) theorem pure_eventual_until_simp: assumes "pure_eventual \" shows "w \\<^sub>n \ U\<^sub>n \ \ w \\<^sub>n \" proof - have "\i. suffix i w \\<^sub>n \ \ w \\<^sub>n \" using pure_eventual_left_append[OF assms] prefix_suffix by metis thus ?thesis by force qed theorem pure_universal_release_simp: assumes "pure_universal \" shows "w \\<^sub>n \ R\<^sub>n \ \ w \\<^sub>n \" proof - have "\i. w \\<^sub>n \ \ suffix i w \\<^sub>n \" using pure_universal_suffix_closed[OF assms] prefix_suffix by metis thus ?thesis by force qed theorem pure_universal_weak_until_simp: assumes "pure_universal \" and "pure_universal \" shows "w \\<^sub>n \ W\<^sub>n \ \ w \\<^sub>n \ or\<^sub>n \" proof - have "\i. w \\<^sub>n \ \ suffix i w \\<^sub>n \" and "\i. w \\<^sub>n \ \ suffix i w \\<^sub>n \" using assms pure_universal_suffix_closed prefix_suffix by metis+ thus ?thesis by force qed theorem pure_eventual_strong_release_simp: assumes "pure_eventual \" and "pure_eventual \" shows "w \\<^sub>n \ M\<^sub>n \ \ w \\<^sub>n \ and\<^sub>n \" proof - have "\i. suffix i w \\<^sub>n \ \ w \\<^sub>n \" and "\i. suffix i w \\<^sub>n \ \ w \\<^sub>n \" using assms pure_eventual_left_append prefix_suffix by metis+ thus ?thesis by force qed theorem suspendable_formula_simp: assumes "suspendable \" shows "w \\<^sub>n X\<^sub>n \ \ w \\<^sub>n \" (is ?t1) and "w \\<^sub>n \ U\<^sub>n \ \ w \\<^sub>n \" (is ?t2) and "w \\<^sub>n \ R\<^sub>n \ \ w \\<^sub>n \" (is ?t3) proof - have "\i. suffix i w \\<^sub>n \ \ w \\<^sub>n \" using suspendable_prefix_invariant[OF assms] prefix_suffix by metis thus ?t1 ?t2 ?t3 by force+ qed theorem suspendable_formula_simp2: assumes "suspendable \" and "suspendable \" shows "w \\<^sub>n \ W\<^sub>n \ \ w \\<^sub>n \ or\<^sub>n \" (is ?t1) and "w \\<^sub>n \ M\<^sub>n \ \ w \\<^sub>n \ and\<^sub>n \" (is ?t2) proof - have "\i. suffix i w \\<^sub>n \ \ w \\<^sub>n \" and "\i. suffix i w \\<^sub>n \ \ w \\<^sub>n \" using assms suspendable_prefix_invariant prefix_suffix by metis+ thus ?t1 ?t2 by force+ qed fun rewrite_modal :: "'a ltln \ 'a ltln" where "rewrite_modal true\<^sub>n = true\<^sub>n" | "rewrite_modal false\<^sub>n = false\<^sub>n" | "rewrite_modal (\ and\<^sub>n \) = (rewrite_modal \ and\<^sub>n rewrite_modal \)" | "rewrite_modal (\ or\<^sub>n \) = (rewrite_modal \ or\<^sub>n rewrite_modal \)" | "rewrite_modal (\ U\<^sub>n \) = (if pure_eventual \ \ suspendable \ then rewrite_modal \ else (rewrite_modal \ U\<^sub>n rewrite_modal \))" | "rewrite_modal (\ R\<^sub>n \) = (if pure_universal \ \ suspendable \ then rewrite_modal \ else (rewrite_modal \ R\<^sub>n rewrite_modal \))" | "rewrite_modal (\ W\<^sub>n \) = (if pure_universal \ \ pure_universal \ \ suspendable \ \ suspendable \ then (rewrite_modal \ or\<^sub>n rewrite_modal \) else (rewrite_modal \ W\<^sub>n rewrite_modal \))" | "rewrite_modal (\ M\<^sub>n \) = (if pure_eventual \ \ pure_eventual \ \ suspendable \ \ suspendable \ then (rewrite_modal \ and\<^sub>n rewrite_modal \) else (rewrite_modal \ M\<^sub>n rewrite_modal \))" | "rewrite_modal (X\<^sub>n \) = (if suspendable \ then rewrite_modal \ else X\<^sub>n (rewrite_modal \))" | "rewrite_modal \ = \" lemma rewrite_modal_sound [simp]: "w \\<^sub>n rewrite_modal \ \ w \\<^sub>n \" proof (induction \ arbitrary: w) case (Until_ltln \ \) thus ?case apply (cases "pure_eventual \ \ suspendable \") apply (insert pure_eventual_until_simp[of \] suspendable_formula_simp[of \]) apply fastforce+ done next case (Release_ltln \ \) thus ?case apply (cases "pure_universal \ \ suspendable \") apply (insert pure_universal_release_simp[of \] suspendable_formula_simp[of \]) apply fastforce+ done next case (WeakUntil_ltln \ \) thus ?case apply (cases "pure_universal \ \ pure_universal \ \ suspendable \ \ suspendable \") apply (insert pure_universal_weak_until_simp[of \ \] suspendable_formula_simp2[of \ \]) apply fastforce+ done next case (StrongRelease_ltln \ \) thus ?case apply (cases "pure_eventual \ \ pure_eventual \ \ suspendable \ \ suspendable \") apply (insert pure_eventual_strong_release_simp[of \ \] suspendable_formula_simp2[of \ \]) apply fastforce+ done next case (Next_ltln \) thus ?case apply (cases "suspendable \") apply (insert suspendable_formula_simp[of \]) apply fastforce+ done qed auto lemma rewrite_modal_size: "size (rewrite_modal \) \ size \" by (induction \) auto subsection \Syntactical Implication Based Simplification\ inductive syntactical_implies :: "'a ltln \ 'a ltln \ bool" ("_ \\<^sub>s _" [80, 80]) where "_ \\<^sub>s true\<^sub>n" | "false\<^sub>n \\<^sub>s _" | "\ = \ \ \ \\<^sub>s \" | "\ \\<^sub>s \ \ (\ and\<^sub>n \) \\<^sub>s \" | "\ \\<^sub>s \ \ (\ and\<^sub>n \) \\<^sub>s \" | "\ \\<^sub>s \ \ \ \\<^sub>s \ \ \ \\<^sub>s (\ and\<^sub>n \)" | "\ \\<^sub>s \ \ \ \\<^sub>s (\ or\<^sub>n \)" | "\ \\<^sub>s \ \ \ \\<^sub>s (\ or\<^sub>n \)" | "\ \\<^sub>s \ \ \ \\<^sub>s \ \ (\ or\<^sub>n \) \\<^sub>s \" | "\ \\<^sub>s \ \ \ \\<^sub>s (\ U\<^sub>n \)" | "\ \\<^sub>s \ \ \ \\<^sub>s \ \ (\ U\<^sub>n \) \\<^sub>s \" | "\ \\<^sub>s \ \ \ \\<^sub>s \ \ (\ U\<^sub>n \) \\<^sub>s (\ U\<^sub>n \)" | "\ \\<^sub>s \ \ (\ R\<^sub>n \) \\<^sub>s \" | "\ \\<^sub>s \ \ \ \\<^sub>s \ \ \ \\<^sub>s (\ R\<^sub>n \)" | "\ \\<^sub>s \ \ \ \\<^sub>s \ \ (\ R\<^sub>n \) \\<^sub>s (\ R\<^sub>n \)" | "(false\<^sub>n R\<^sub>n \) \\<^sub>s \ \ (false\<^sub>n R\<^sub>n \) \\<^sub>s X\<^sub>n \" | "\ \\<^sub>s (true\<^sub>n U\<^sub>n \) \ (X\<^sub>n \) \\<^sub>s (true\<^sub>n U\<^sub>n \)" | "\ \\<^sub>s \ \ (X\<^sub>n \) \\<^sub>s (X\<^sub>n \)" lemma syntactical_implies_correct: "\ \\<^sub>s \ \ w \\<^sub>n \ \ w \\<^sub>n \" by (induction arbitrary: w rule: syntactical_implies.induct; auto; force) fun rewrite_syn_imp where "rewrite_syn_imp (\ and\<^sub>n \) = ( if \ \\<^sub>s \ then rewrite_syn_imp \ else if \ \\<^sub>s \ then rewrite_syn_imp \ else if \ \\<^sub>s (not\<^sub>n \) \ \ \\<^sub>s (not\<^sub>n \) then false\<^sub>n else mk_and (rewrite_syn_imp \) (rewrite_syn_imp \))" | "rewrite_syn_imp (\ or\<^sub>n \) = ( if \ \\<^sub>s \ then rewrite_syn_imp \ else if \ \\<^sub>s \ then rewrite_syn_imp \ else if (not\<^sub>n \) \\<^sub>s \ \ (not\<^sub>n \) \\<^sub>s \ then true\<^sub>n else mk_or (rewrite_syn_imp \) (rewrite_syn_imp \))" | "rewrite_syn_imp (\ U\<^sub>n \) = ( if \ \\<^sub>s \ then rewrite_syn_imp \ else if (not\<^sub>n \) \\<^sub>s \ then mk_finally (rewrite_syn_imp \) else mk_until (rewrite_syn_imp \) (rewrite_syn_imp \))" | "rewrite_syn_imp (\ R\<^sub>n \) = ( if \ \\<^sub>s \ then rewrite_syn_imp \ else if \ \\<^sub>s (not\<^sub>n \) then mk_globally (rewrite_syn_imp \) else mk_release (rewrite_syn_imp \) (rewrite_syn_imp \))" | "rewrite_syn_imp (X\<^sub>n \) = mk_next (rewrite_syn_imp \)" | "rewrite_syn_imp \ = \" lemma rewrite_syn_imp_sound: "w \\<^sub>n rewrite_syn_imp \ \ w \\<^sub>n \" proof (induction \ arbitrary: w) case And_ltln thus ?case by (simp add: Let_def; metis syntactical_implies_correct not\<^sub>n_semantics) next case (Or_ltln \ \) moreover have "(not\<^sub>n \) \\<^sub>s \ \ \w. w \\<^sub>n \ or\<^sub>n \" by (auto intro: syntactical_implies_correct[of "not\<^sub>n \"]) moreover have "(not\<^sub>n \) \\<^sub>s \ \ \w. w \\<^sub>n \ or\<^sub>n \" by (auto intro: syntactical_implies_correct[of "not\<^sub>n \"]) ultimately show ?case by (auto intro: syntactical_implies_correct) next case (Until_ltln \ \) moreover have "\ \\<^sub>s \ \ ?case" by (force simp add: Until_ltln dest: syntactical_implies_correct) moreover { assume A: "(not\<^sub>n \) \\<^sub>s \" and B: "\ \ \\<^sub>s \" hence [simp]: "rewrite_syn_imp (\ U\<^sub>n \) = mk_finally (rewrite_syn_imp \)" by simp { assume "\i. suffix i w \\<^sub>n \" moreover define i where "i \ LEAST i. suffix i w \\<^sub>n \" ultimately have "\j < i. \suffix j w \\<^sub>n \" and "suffix i w \\<^sub>n \" by (blast dest: not_less_Least , metis LeastI \\i. suffix i w \\<^sub>n \\ i_def) hence "\j < i. suffix j w \\<^sub>n \" and "suffix i w \\<^sub>n \" using syntactical_implies_correct[OF A] by auto } hence ?case by (simp del: rewrite_syn_imp.simps; unfold Until_ltln(2)) blast } ultimately show ?case by fastforce next case (Release_ltln \ \) moreover have "\ \\<^sub>s \ \ ?case" by (force simp add: Release_ltln dest: syntactical_implies_correct) moreover { assume A: "\ \\<^sub>s (not\<^sub>n \)" and B: "\ \ \\<^sub>s \" hence [simp]: "rewrite_syn_imp (\ R\<^sub>n \) = mk_globally (rewrite_syn_imp \)" by simp { assume "\i. \suffix i w \\<^sub>n \" moreover define i where "i \ LEAST i. \suffix i w \\<^sub>n \" ultimately have "\j < i. suffix j w \\<^sub>n \" and "\ suffix i w \\<^sub>n \" by (blast dest: not_less_Least , metis LeastI \\i. \suffix i w \\<^sub>n \\ i_def) hence "\j < i. \suffix j w \\<^sub>n \" and "\ suffix i w \\<^sub>n \" using syntactical_implies_correct[OF A] by auto } hence ?case by (simp del: rewrite_syn_imp.simps; unfold Release_ltln(2)) blast } ultimately show ?case by fastforce qed auto subsection \Iterated Rewriting\ fun iterate where "iterate f x 0 = x" | "iterate f x (Suc n) = (let x' = f x in if x = x' then x else iterate f x' n)" definition "rewrite_iter_fast \ \ iterate (rewrite_modal o rewrite_X) \ (size \)" definition "rewrite_iter_slow \ \ iterate (rewrite_syn_imp o rewrite_modal o rewrite_X) \ (size \)" text \The rewriting functions defined in the previous subsections can be used as-is. However, in the most cases one wants to iterate these rules until the formula cannot be simplified further. @{const rewrite_iter_fast} pulls X operators up in the syntax tree and the uses the "modal" simplification rules. @{const rewrite_iter_slow} additionally tries to simplify the formula using syntactic implication checking.\ lemma iterate_sound: assumes "\\. w \\<^sub>n f \ \ w \\<^sub>n \" shows "w \\<^sub>n iterate f \ n \ w \\<^sub>n \" by (induction n arbitrary: \; simp add: assms Let_def) theorem rewrite_iter_fast_sound [simp]: "w \\<^sub>n rewrite_iter_fast \ \ w \\<^sub>n \" using iterate_sound[of _ "rewrite_modal o rewrite_X"] unfolding comp_def rewrite_modal_sound rewrite_X_sound rewrite_iter_fast_def by blast theorem rewrite_iter_slow_sound [simp]: "w \\<^sub>n rewrite_iter_slow \ \ w \\<^sub>n \" using iterate_sound[of _ "rewrite_syn_imp o rewrite_modal o rewrite_X"] unfolding comp_def rewrite_modal_sound rewrite_X_sound rewrite_syn_imp_sound rewrite_iter_slow_def by blast subsection \Preservation of atoms\ lemma iterate_atoms: assumes "\\. atoms_ltln (f \) \ atoms_ltln \" shows "atoms_ltln (iterate f \ n) \ atoms_ltln \" by (induction n arbitrary: \) (auto, metis (mono_tags, lifting) assms in_mono) lemma rewrite_modal_atoms: "atoms_ltln (rewrite_modal \) \ atoms_ltln \" by (induction \) auto lemma mk_and_atoms: "atoms_ltln (mk_and \ \) \ atoms_ltln \ \ atoms_ltln \" by (auto simp: mk_and_def split: ltln.splits) lemma mk_or_atoms: "atoms_ltln (mk_or \ \) \ atoms_ltln \ \ atoms_ltln \" by (auto simp: mk_or_def split: ltln.splits) lemma remove_strong_ops_atoms: "atoms_ltln (remove_strong_ops \) \ atoms_ltln \" by (induction \) auto lemma remove_weak_ops_atoms: "atoms_ltln (remove_weak_ops \) \ atoms_ltln \" by (induction \) auto lemma mk_finally_atoms: "atoms_ltln (mk_finally \) \ atoms_ltln \" by (auto simp: mk_finally_def split: ltln.splits) (insert remove_strong_ops_atoms, fast+) lemma mk_globally_atoms: "atoms_ltln (mk_globally \) \ atoms_ltln \" by (auto simp: mk_globally_def split: ltln.splits) (insert remove_weak_ops_atoms, fast+) lemma mk_until_atoms: "atoms_ltln (mk_until \ \) \ atoms_ltln \ \ atoms_ltln \" by (auto simp: mk_until_def split: ltln.splits) (insert mk_finally_atoms, fastforce+) lemma mk_release_atoms: "atoms_ltln (mk_release \ \) \ atoms_ltln \ \ atoms_ltln \" by (auto simp: mk_release_def split: ltln.splits) (insert mk_globally_atoms, fastforce+) lemma mk_weak_until_atoms: "atoms_ltln (mk_weak_until \ \) \ atoms_ltln \ \ atoms_ltln \" by (auto simp: mk_weak_until_def split: ltln.splits) (insert mk_globally_atoms, fastforce+) lemma mk_strong_release_atoms: "atoms_ltln (mk_strong_release \ \) \ atoms_ltln \ \ atoms_ltln \" by (auto simp: mk_strong_release_def split: ltln.splits) (insert mk_finally_atoms, fastforce+) lemma mk_next_atoms: "atoms_ltln (mk_next \) = atoms_ltln \" by (auto simp: mk_next_def split: ltln.splits) lemma mk_next_pow_atoms: "atoms_ltln (mk_next_pow n \) = atoms_ltln \" by (induction n) (auto simp: mk_next_pow_def split: ltln.splits) lemma combine_atoms: assumes "\\ \. atoms_ltln (f \ \) \ atoms_ltln \ \ atoms_ltln \" shows "atoms_ltln (fst (combine f x y)) \ atoms_ltln (fst x) \ atoms_ltln (fst y)" by (metis assms fst_combine mk_next_pow_atoms prod.collapse) lemmas combine_mk_atoms = combine_atoms[OF mk_and_atoms] combine_atoms[OF mk_or_atoms] combine_atoms[OF mk_until_atoms] combine_atoms[OF mk_release_atoms] combine_atoms[OF mk_weak_until_atoms] combine_atoms[OF mk_strong_release_atoms] lemma rewrite_X_enat_atoms: "atoms_ltln (fst (rewrite_X_enat \)) \ atoms_ltln \" by (induction \) (simp_all add: case_prod_beta, insert combine_mk_atoms, fast+) lemma rewrite_X_atoms: "atoms_ltln (rewrite_X \) \ atoms_ltln \" by (induction \) (simp_all add: rewrite_X_def mk_next_pow_atoms case_prod_beta, insert combine_mk_atoms, fast+) lemma rewrite_syn_imp_atoms: "atoms_ltln (rewrite_syn_imp \) \ atoms_ltln \" proof (induction \) case (And_ltln \1 \2) then show ?case using mk_and_atoms by simp fast next case (Or_ltln \1 \2) then show ?case using mk_or_atoms by simp fast next case (Next_ltln \) then show ?case using mk_next_atoms by simp fast next case (Until_ltln \1 \2) then show ?case using mk_finally_atoms mk_until_atoms by simp fast next case (Release_ltln \1 \2) then show ?case using mk_globally_atoms mk_release_atoms by simp fast qed simp_all lemma rewrite_iter_fast_atoms: "atoms_ltln (rewrite_iter_fast \) \ atoms_ltln \" proof - have 1: "\\. atoms_ltln (rewrite_modal (rewrite_X \)) \ atoms_ltln \" using rewrite_modal_atoms rewrite_X_atoms by force show ?thesis by (simp add: rewrite_iter_fast_def 1 iterate_atoms) qed lemma rewrite_iter_slow_atoms: "atoms_ltln (rewrite_iter_slow \) \ atoms_ltln \" proof - have 1: "\\. atoms_ltln (rewrite_syn_imp (rewrite_modal (rewrite_X \))) \ atoms_ltln \" using rewrite_syn_imp_atoms rewrite_modal_atoms rewrite_X_atoms by force show ?thesis by (simp add: rewrite_iter_slow_def 1 iterate_atoms) qed subsection \Simplifier\ text \We now define a convenience wrapper for the rewriting engine\ datatype mode = Nop | Fast | Slow fun simplify :: "mode \ 'a ltln \ 'a ltln" where "simplify Nop = id" | "simplify Fast = rewrite_iter_fast" | "simplify Slow = rewrite_iter_slow" theorem simplify_correct: "w \\<^sub>n simplify m \ \ w \\<^sub>n \" by (cases m) simp+ lemma simplify_atoms: "atoms_ltln (simplify m \) \ atoms_ltln \" by (cases m) (insert rewrite_iter_fast_atoms rewrite_iter_slow_atoms, fastforce+) subsection \Code Generation\ code_pred syntactical_implies . export_code simplify checking lemma "rewrite_iter_fast (F\<^sub>n (G\<^sub>n (X\<^sub>n prop\<^sub>n(''a'')))) = (F\<^sub>n (G\<^sub>n prop\<^sub>n(''a'')))" by eval lemma "rewrite_iter_fast (X\<^sub>n prop\<^sub>n(''a'') U\<^sub>n (X\<^sub>n nprop\<^sub>n(''a''))) = X\<^sub>n (prop\<^sub>n(''a'') U\<^sub>n nprop\<^sub>n(''a''))" by eval lemma "rewrite_iter_slow (X\<^sub>n prop\<^sub>n(''a'') U\<^sub>n (X\<^sub>n nprop\<^sub>n(''a''))) = X\<^sub>n (F\<^sub>n nprop\<^sub>n(''a''))" by eval end diff --git a/thys/Psi_Calculi/Semantics.thy b/thys/Psi_Calculi/Semantics.thy --- a/thys/Psi_Calculi/Semantics.thy +++ b/thys/Psi_Calculi/Semantics.thy @@ -1,9323 +1,9323 @@ (* Title: Psi-calculi Author/Maintainer: Jesper Bengtson (jebe@itu.dk), 2012 *) theory Semantics imports Frame begin nominal_datatype ('a, 'b, 'c) boundOutput = - BOut "'a::fs_name" "('a, 'b::fs_name, 'c::fs_name) psi" ("_ \' _" [110, 110] 110) + BOut "'a::fs_name" "('a, 'b::fs_name, 'c::fs_name) psi" ("_ \'' _" [110, 110] 110) | BStep "\name\ ('a, 'b, 'c) boundOutput" ("\\_\_" [110, 110] 110) primrec BOresChain :: "name list \ ('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput \ ('a, 'b, 'c) boundOutput" where Base: "BOresChain [] B = B" | Step: "BOresChain (x#xs) B = \\x\(BOresChain xs B)" abbreviation BOresChainJudge ("\\*_\_" [80, 80] 80) where "\\*xvec\B \ BOresChain xvec B" lemma BOresChainEqvt[eqvt]: fixes perm :: "name prm" and lst :: "name list" and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" shows "perm \ (\\*xvec\B) = \\*(perm \ xvec)\(perm \ B)" by(induct_tac xvec, auto) lemma BOresChainSimps[simp]: fixes xvec :: "name list" and N :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and N' :: 'a and P' :: "('a, 'b, 'c) psi" and B :: "('a, 'b, 'c) boundOutput" and B' :: "('a, 'b, 'c) boundOutput" shows "(\\*xvec\N \' P = N' \' P') = (xvec = [] \ N = N' \ P = P')" and "(N' \' P' = \\*xvec\N \' P) = (xvec = [] \ N = N' \ P = P')" and "(N' \' P' = N \' P) = (N = N' \ P = P')" and "(\\*xvec\B = \\*xvec\B') = (B = B')" by(induct xvec) (auto simp add: boundOutput.inject alpha) lemma outputFresh[simp]: fixes Xs :: "name set" and xvec :: "name list" and N :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" shows "(Xs \* (N \' P)) = ((Xs \* N) \ (Xs \* P))" and "(xvec \* (N \' P)) = ((xvec \* N) \ (xvec \* P))" by(auto simp add: fresh_star_def) lemma boundOutputFresh: fixes x :: name and xvec :: "name list" and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" shows "(x \ (\\*xvec\B)) = (x \ set xvec \ x \ B)" by (induct xvec) (simp_all add: abs_fresh) lemma boundOutputFreshSet: fixes Xs :: "name set" and xvec :: "name list" and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" and yvec :: "name list" and x :: name shows "Xs \* (\\*xvec\B) = (\x\Xs. x \ set xvec \ x \ B)" and "yvec \* (\\*xvec\B) = (\x\(set yvec). x \ set xvec \ x \ B)" and "Xs \* (\\x\B) = Xs \* [x].B" and "xvec \* (\\x\B) = xvec \* [x].B" by(simp add: fresh_star_def boundOutputFresh)+ lemma BOresChainSupp: fixes xvec :: "name list" and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" shows "(supp(\\*xvec\B)::name set) = (supp B) - (supp xvec)" by(induct xvec) (auto simp add: boundOutput.supp supp_list_nil supp_list_cons abs_supp supp_atm) lemma boundOutputFreshSimps[simp]: fixes Xs :: "name set" and xvec :: "name list" and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" and yvec :: "name list" and x :: name shows "Xs \* xvec \ (Xs \* (\\*xvec\B)) = (Xs \* B)" and "yvec \* xvec \ yvec \* (\\*xvec\B) = yvec \* B" and "xvec \* (\\*xvec\B)" and "x \ xvec \ x \ \\*xvec\B = x \ B" apply(simp add: boundOutputFreshSet) apply(force simp add: fresh_star_def name_list_supp fresh_def) apply(simp add: boundOutputFreshSet) apply(force simp add: fresh_star_def name_list_supp fresh_def) apply(simp add: boundOutputFreshSet) by(simp add: BOresChainSupp fresh_def) lemma boundOutputChainAlpha: fixes p :: "name prm" and xvec :: "name list" and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" and yvec :: "name list" assumes xvecFreshB: "(p \ xvec) \* B" and S: "set p \ set xvec \ set (p \ xvec)" and "(set xvec) \ (set yvec)" shows "(\\*yvec\B) = (\\*(p \ yvec)\(p \ B))" proof - note pt_name_inst at_name_inst S moreover from \(set xvec) \ (set yvec)\ have "set xvec \* (\\*yvec\B)" by(force simp add: boundOutputFreshSet) moreover from xvecFreshB \(set xvec) \ (set yvec)\ have "set (p \ xvec) \* (\\*yvec\B)" by (simp add: boundOutputFreshSet) (simp add: fresh_star_def) ultimately have "(\\*yvec\B) = p \ (\\*yvec\B)" by (rule_tac pt_freshs_freshs [symmetric]) then show ?thesis by(simp add: eqvts) qed lemma boundOutputChainAlpha': fixes p :: "name prm" and xvec :: "name list" and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" and yvec :: "name list" and zvec :: "name list" assumes xvecFreshB: "xvec \* B" and S: "set p \ set xvec \ set yvec" and "yvec \* (\\*zvec\B)" shows "(\\*zvec\B) = (\\*(p \ zvec)\(p \ B))" proof - note pt_name_inst at_name_inst S \yvec \* (\\*zvec\B)\ moreover from xvecFreshB have "set (xvec) \* (\\*zvec\B)" by (simp add: boundOutputFreshSet) (simp add: fresh_star_def) ultimately have "(\\*zvec\B) = p \ (\\*zvec\B)" by (rule_tac pt_freshs_freshs [symmetric]) auto then show ?thesis by(simp add: eqvts) qed lemma boundOutputChainAlpha'': fixes p :: "name prm" and xvec :: "name list" and M :: "'a::fs_name" and P :: "('a::fs_name, 'b::fs_name, 'c::fs_name) psi" and yvec :: "name list" assumes "(p \ xvec) \* M" and "(p \ xvec) \* P" and "set p \ set xvec \ set (p \ xvec)" and "(set xvec) \ (set yvec)" shows "(\\*yvec\M \' P) = (\\*(p \ yvec)\(p \ M) \' (p \ P))" using assms by(subst boundOutputChainAlpha) auto lemma boundOutputChainSwap: fixes x :: name and y :: name and N :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and xvec :: "name list" assumes "y \ N" and "y \ P" and "x \ (set xvec)" shows "\\*xvec\N \' P = \\*([(x, y)] \ xvec)\([(x ,y)] \ N) \' ([(x, y)] \ P)" proof(case_tac "x=y") assume "x=y" thus ?thesis by simp next assume "x \ y" with assms show ?thesis by(rule_tac xvec="[x]" in boundOutputChainAlpha'') (auto simp add: calc_atm) qed lemma alphaBoundOutput: fixes x :: name and y :: name and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" assumes "y \ B" shows "\\x\B = \\y\([(x, y)] \ B)" using assms by(auto simp add: boundOutput.inject alpha fresh_left calc_atm) lemma boundOutputEqFresh: fixes B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" and C :: "('a, 'b, 'c) boundOutput" and x :: name and y :: name assumes "\\x\B = \\y\C" and "x \ B" shows "y \ C" using assms by(auto simp add: boundOutput.inject alpha fresh_left calc_atm) lemma boundOutputEqSupp: fixes B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" and C :: "('a, 'b, 'c) boundOutput" and x :: name and y :: name assumes "\\x\B = \\y\C" and "x \ supp B" shows "y \ supp C" using assms apply(auto simp add: boundOutput.inject alpha fresh_left calc_atm) apply(drule_tac pi="[(x, y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) by(simp add: eqvts calc_atm) lemma boundOutputChainEq: fixes xvec :: "name list" and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" and yvec :: "name list" and B' :: "('a, 'b, 'c) boundOutput" assumes "\\*xvec\B = \\*yvec\B'" and "xvec \* yvec" and "length xvec = length yvec" shows "\p. (set p) \ (set xvec) \ set (yvec) \ distinctPerm p \ B = p \ B' \ (set (map fst p)) \ (supp B) \ xvec \* B' \ yvec \* B" proof - obtain n where "n = length xvec" by auto with assms show ?thesis proof(induct n arbitrary: xvec yvec B B') case(0 xvec yvec B B') have Eq: "\\*xvec\B = \\*yvec\B'" by fact from \0 = length xvec\ have "xvec = []" by auto moreover with \length xvec = length yvec\ have "yvec = []" by(case_tac yvec) auto ultimately show ?case using Eq by(simp add: boundOutput.inject) next case(Suc n xvec yvec B B') from \Suc n = length xvec\ obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" by(case_tac xvec) auto from \\\*xvec\B = \\*yvec\B'\ \xvec = x # xvec'\ \length xvec = length yvec\ obtain y yvec' where "\\*(x#xvec')\B = \\*(y#yvec')\B'" and "yvec = y#yvec'" and "length xvec' = length yvec'" by(case_tac yvec) auto hence EQ: "\\x\(\\*xvec'\B) = \\y\(\\*yvec'\B')" by simp from \xvec = x#xvec'\ \yvec=y#yvec'\ \xvec \* yvec\ have "x \ y" and "xvec' \* yvec'" and "x \ yvec'" and "y \ xvec'" by auto have IH: "\xvec yvec B B'. \\\*xvec\(B::('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput) = \\*yvec\B'; xvec \* yvec; length xvec = length yvec; n = length xvec\ \ \p. (set p) \ (set xvec) \ (set yvec) \ distinctPerm p \ B = p \ B' \ set(map fst p) \ supp B \ xvec \* B' \ yvec \* B" by fact from EQ \x \ y\ have EQ': "\\*xvec'\B = ([(x, y)] \ (\\*yvec'\B'))" and xFreshB': "x \ (\\*yvec'\B')" and yFreshB: "y \ (\\*xvec'\B)" by(metis boundOutput.inject alpha)+ from xFreshB' \x \ yvec'\ have "x \ B'" by(auto simp add: boundOutputFresh) (simp add: fresh_def name_list_supp)+ from yFreshB \y \ xvec'\ have "y \ B" by(auto simp add: boundOutputFresh) (simp add: fresh_def name_list_supp)+ show ?case proof(case_tac "x \ \\*xvec'\B") assume xFreshB: "x \ \\*xvec'\B" with EQ have yFreshB': "y \ \\*yvec'\B'" by(rule boundOutputEqFresh) with xFreshB' EQ' have "\\*xvec'\B = \\*yvec'\B'" by(simp) with \xvec' \* yvec'\ \length xvec' = length yvec'\ \length xvec' = n\ IH obtain p where S: "(set p) \ (set xvec') \ (set yvec')" and "distinctPerm p" and "B = p \ B'" and "set(map fst p) \ supp B" and "xvec' \* B'" and "yvec' \* B" by blast from S have "(set p) \ set(x#xvec') \ set(y#yvec')" by auto moreover note \xvec = x#xvec'\ \yvec=y#yvec'\ \distinctPerm p\ \B = p \ B'\ \xvec' \* B'\ \x \ B'\ \x \ B'\ \yvec' \* B\ \y \ B\ \set(map fst p) \ supp B\ ultimately show ?case by auto next assume "\(x \ \\*xvec'\B)" hence xSuppB: "x \ supp(\\*xvec'\B)" by(simp add: fresh_def) with EQ have ySuppB': "y \ supp (\\*yvec'\B')" by(rule boundOutputEqSupp) hence "y \ yvec'" by(induct yvec') (auto simp add: boundOutput.supp abs_supp) with \x \ yvec'\ EQ' have "\\*xvec'\B = \\*yvec'\([(x, y)] \ B')" by(simp add: eqvts) with \xvec' \* yvec'\ \length xvec' = length yvec'\ \length xvec' = n\ IH obtain p where S: "(set p) \ (set xvec') \ (set yvec')" and "distinctPerm p" and "B = p \ [(x, y)] \ B'" and "set(map fst p) \ supp B" and "xvec' \* ([(x, y)] \ B')" and "yvec' \* B" by blast from xSuppB have "x \ xvec'" by(induct xvec') (auto simp add: boundOutput.supp abs_supp) with \x \ yvec'\ \y \ xvec'\ \y \ yvec'\ S have "x \ p" and "y \ p" apply(induct p) by(auto simp add: name_list_supp) (auto simp add: fresh_def) from S have "(set ((x, y)#p)) \ (set(x#xvec')) \ (set(y#yvec'))" by force moreover from \x \ y\ \x \ p\ \y \ p\ S \distinctPerm p\ have "distinctPerm((x,y)#p)" by simp moreover from \B = p \ [(x, y)] \ B'\ \x \ p\ \y \ p\ have "B = [(x, y)] \ p \ B'" by(subst perm_compose) simp hence "B = ((x, y)#p) \ B'" by simp moreover from \xvec' \* ([(x, y)] \ B')\ have "([(x, y)] \ xvec') \* ([(x, y)] \ [(x, y)] \ B')" by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \x \ xvec'\ \y \ xvec'\ \x \ B'\ have "(x#xvec') \* B'" by simp moreover from \y \ B\ \yvec' \* B\ have "(y#yvec') \* B" by simp moreover from \set(map fst p) \ supp B\ xSuppB \x \ xvec'\ have "set(map fst ((x, y)#p)) \ supp B" by(simp add: BOresChainSupp) ultimately show ?case using \xvec=x#xvec'\ \yvec=y#yvec'\ by metis qed qed qed lemma boundOutputChainEqLength: fixes xvec :: "name list" and M :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and yvec :: "name list" and N :: "'a::fs_name" and Q :: "('a, 'b::fs_name, 'c::fs_name) psi" assumes "\\*xvec\M \' P = \\*yvec\N \' Q" shows "length xvec = length yvec" proof - obtain n where "n = length xvec" by auto with assms show ?thesis proof(induct n arbitrary: xvec yvec M P N Q) case(0 xvec yvec M P N Q) from \0 = length xvec\ have "xvec = []" by auto moreover with \\\*xvec\M \' P = \\*yvec\N \' Q\ have "yvec = []" by(case_tac yvec) auto ultimately show ?case by simp next case(Suc n xvec yvec M P N Q) from \Suc n = length xvec\ obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" by(case_tac xvec) auto from \\\*xvec\M \' P = \\*yvec\N \' Q\ \xvec = x # xvec'\ obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' Q" and "yvec = y#yvec'" by(case_tac yvec) auto hence EQ: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' Q)" by simp have IH: "\xvec yvec M P N Q. \\\*xvec\M \' P = \\*yvec\N \' (Q::('a, 'b, 'c) psi); n = length xvec\ \ length xvec = length yvec" by fact show ?case proof(case_tac "x = y") assume "x = y" with EQ have "\\*xvec'\M \' P = \\*yvec'\N \' Q" by(simp add: alpha boundOutput.inject) with IH \length xvec' = n\ have "length xvec' = length yvec'" by blast with \xvec = x#xvec'\ \yvec=y#yvec'\ show ?case by simp next assume "x \ y" with EQ have "\\*xvec'\M \' P = [(x, y)] \ \\*yvec'\N \' Q" by(simp add: alpha boundOutput.inject) hence "\\*xvec'\M \' P = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ Q)" by(simp add: eqvts) with IH \length xvec' = n\ have "length xvec' = length ([(x, y)] \ yvec')" by blast hence "length xvec' = length yvec'" by simp with \xvec = x#xvec'\ \yvec=y#yvec'\ show ?case by simp qed qed qed lemma boundOutputChainEq': fixes xvec :: "name list" and M :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and yvec :: "name list" and N :: 'a and Q :: "('a::fs_name, 'b::fs_name, 'c::fs_name) psi" assumes "\\*xvec\M \' P = \\*yvec\N \' Q" and "xvec \* yvec" shows "\p. (set p) \ (set xvec) \ set (yvec) \ distinctPerm p \ M = p \ N \ P = p \ Q \ xvec \* N \ xvec \* Q \ yvec \* M \ yvec \* P" using assms apply(frule_tac boundOutputChainEqLength) apply(drule_tac boundOutputChainEq) by(auto simp add: boundOutput.inject) lemma boundOutputChainEq'': fixes xvec :: "name list" and M :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and yvec :: "name list" and N :: 'a and Q :: "('a::fs_name, 'b::fs_name, 'c::fs_name) psi" assumes "\\*xvec\M \' P = \\*yvec\N \' Q" and "xvec \* yvec" and "distinct xvec" and "distinct yvec" obtains p where "(set p) \ (set xvec) \ set (p \ xvec)" and "distinctPerm p" and "yvec = p \ xvec" and "N = p \ M" and "Q = p \ P" and "xvec \* N" and "xvec \* Q" and "(p \ xvec) \* M" and "(p \ xvec) \* P" proof - assume "\p. \set p \ set xvec \ set (p \ xvec); distinctPerm p; yvec = p \ xvec; N = p \ M; Q = p \ P; xvec \* N; xvec \* Q; (p \ xvec) \* M; (p \ xvec) \* P\ \ thesis" moreover obtain n where "n = length xvec" by auto with assms have "\p. (set p) \ (set xvec) \ set (yvec) \ distinctPerm p \ yvec = p \ xvec \ N = p \ M \ Q = p \ P \ xvec \* N \ xvec \* Q \ (p \ xvec) \* M \ (p \ xvec) \* P" proof(induct n arbitrary: xvec yvec M P N Q) case(0 xvec yvec M P N Q) have Eq: "\\*xvec\M \' P = \\*yvec\N \' Q" by fact from \0 = length xvec\ have "xvec = []" by auto moreover with Eq have "yvec = []" by(case_tac yvec) auto ultimately show ?case using Eq by(simp add: boundOutput.inject) next case(Suc n xvec yvec M P N Q) from \Suc n = length xvec\ obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" by(case_tac xvec) auto from \\\*xvec\M \' P = \\*yvec\N \' Q\ \xvec = x # xvec'\ obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' Q" and "yvec = y#yvec'" by(case_tac yvec) auto hence EQ: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' Q)" by simp from \xvec = x#xvec'\ \yvec=y#yvec'\ \xvec \* yvec\ have "x \ y" and "xvec' \* yvec'" and "x \ yvec'" and "y \ xvec'" by auto from \distinct xvec\ \distinct yvec\ \xvec=x#xvec'\ \yvec=y#yvec'\ have "x \ xvec'" and "y \ yvec'" and "distinct xvec'" and "distinct yvec'" by simp+ have IH: "\xvec yvec M P N Q. \\\*xvec\(M::'a) \' (P::('a, 'b, 'c) psi) = \\*yvec\N \' Q; xvec \* yvec; distinct xvec; distinct yvec; n = length xvec\ \ \p. (set p) \ (set xvec) \ (set yvec) \ distinctPerm p \ yvec = p \ xvec \ N = p \ M \ Q = p \ P \ xvec \* N \ xvec \* Q \ (p \ xvec) \* M \ (p \ xvec) \* P" by fact from EQ \x \ y\ \x \ yvec'\ \y \ yvec'\ \y \ xvec'\ \x \ xvec'\ have "\\*xvec'\M \' P = \\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ Q)" and "x \ N" and "x \ Q" and "y \ M" and "y \ P" apply - apply(simp add: boundOutput.inject alpha eqvts) apply(simp add: boundOutput.inject alpha eqvts) apply(simp add: boundOutput.inject alpha eqvts) by(simp add: boundOutput.inject alpha' eqvts)+ with \xvec' \* yvec'\ \distinct xvec'\ \distinct yvec'\ \length xvec' = n\ IH obtain p where S: "(set p) \ (set xvec') \ (set yvec')" and "distinctPerm p" and "yvec' = p \ xvec'" and "([(x, y)] \ N) = p \ M" and "([(x, y)] \ Q) = p \ P" and "xvec' \* ([(x, y)] \ N)" and "xvec' \* ([(x, y)] \ Q)" and "yvec' \* M" and "yvec' \* P" by metis from S have "set((x, y)#p) \ set(x#xvec') \ set(y#yvec')" by auto moreover from \x \ xvec'\ \x \ yvec'\ \y \ xvec'\ \y \ yvec'\ S have "x \ p" and "y \ p" apply(induct p) by(auto simp add: fresh_prod name_list_supp) (auto simp add: fresh_def) with S \distinctPerm p\ \x \ y\ have "distinctPerm((x, y)#p)" by auto moreover from \yvec' = p \ xvec'\ \x \ p\ \y \ p\ \x \ xvec'\ \y \ xvec'\ have "(y#yvec') = ((x, y)#p) \ (x#xvec')" by(simp add: eqvts calc_atm perm_compose freshChainSimps) moreover from \([(x, y)] \ N) = p \ M\ have "([(x, y)] \ [(x, y)] \ N) = [(x, y)] \ p \ M" by(simp add: pt_bij) hence "N = ((x, y)#p) \ M" by simp moreover from \([(x, y)] \ Q) = p \ P\ have "([(x, y)] \ [(x, y)] \ Q) = [(x, y)] \ p \ P" by(simp add: pt_bij) hence "Q = ((x, y)#p) \ P" by simp moreover from \xvec' \* ([(x, y)] \ N)\ have "([(x, y)] \ xvec') \* ([(x, y)] \ [(x, y)] \ N)" by(subst fresh_star_bij) with \x \ xvec'\ \y \ xvec'\ have "xvec' \* N" by simp with \x \ N\ have "(x#xvec') \* N" by simp moreover from \xvec' \* ([(x, y)] \ Q)\ have "([(x, y)] \ xvec') \* ([(x, y)] \ [(x, y)] \ Q)" by(subst fresh_star_bij) with \x \ xvec'\ \y \ xvec'\ have "xvec' \* Q" by simp with \x \ Q\ have "(x#xvec') \* Q" by simp moreover from \y \ M\ \yvec' \* M\ have "(y#yvec') \* M" by simp moreover from \y \ P\ \yvec' \* P\ have "(y#yvec') \* P" by simp ultimately show ?case using \xvec=x#xvec'\ \yvec=y#yvec'\ by metis qed ultimately show ?thesis by blast qed lemma boundOutputEqSupp': fixes x :: name and xvec :: "name list" and M :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and y :: name and yvec :: "name list" and N :: 'a and Q :: "('a, 'b, 'c) psi" assumes Eq: "\\x\(\\*xvec\M \' P) = \\y\(\\*yvec\N \' Q)" and "x \ y" and "x \ yvec" and "x \ xvec" and "y \ xvec" and "y \ yvec" and "xvec \* yvec" and "x \ supp M" shows "y \ supp N" proof - from Eq \x \ y\ \x \ yvec\ \y \ yvec\ have "\\*xvec\M \' P = \\*yvec\([(x, y)] \ N) \' ([(x, y)] \ Q)" by(simp add: boundOutput.inject alpha eqvts) then obtain p where S: "set p \ set xvec \ set yvec" and "M = p \ [(x, y)] \ N" and "distinctPerm p" using \xvec \* yvec\ by(blast dest: boundOutputChainEq') with \x \ supp M\ have "x \ supp(p \ [(x, y)] \ N)" by simp hence "(p \ x) \ p \ supp(p \ [(x, y)] \ N)" by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) with \x \ xvec\ \x \ yvec\ S \distinctPerm p\ have "x \ supp([(x, y)] \ N)" by(simp add: eqvts) hence "([(x, y)] \ x) \ ([(x, y)] \ (supp([(x, y)] \ N)))" by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) with \x \ y\ show ?thesis by(simp add: calc_atm eqvts) qed lemma boundOutputChainOpenIH: fixes xvec :: "name list" and x :: name and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" and yvec :: "name list" and y :: name and B' :: "('a, 'b, 'c) boundOutput" assumes Eq: "\\*xvec\(\\x\B) = \\*yvec\(\\y\B')" and L: "length xvec = length yvec" and xFreshB': "x \ B'" and xFreshxvec: "x \ xvec" and xFreshyvec: "x \ yvec" shows "\\*xvec\B = \\*yvec\([(x, y)] \ B')" using assms proof(induct n=="length xvec" arbitrary: xvec yvec y B' rule: nat.induct) case(zero xvec yvec y B') have "0 = length xvec" and "length xvec = length yvec" by fact+ moreover have "\\*xvec\\\x\B = \\*yvec\\\y\B'" by fact ultimately show ?case by(auto simp add: boundOutput.inject alpha) next case(Suc n xvec yvec y B') have L: "length xvec = length yvec" and "Suc n = length xvec" by fact+ then obtain x' xvec' y' yvec' where xEq: "xvec = x'#xvec'" and yEq: "yvec = y'#yvec'" and L': "length xvec' = length yvec'" by(cases xvec, auto, cases yvec, auto) have xFreshB': "x \ B'" by fact have "x \ xvec" and "x \ yvec" by fact+ with xEq yEq have xineqx': "x \ x'" and xFreshxvec': "x \ xvec'" and xineqy': "x \ y'" and xFreshyvec': "x \ yvec'" by simp+ have "\\*xvec\\\x\B = \\*yvec\\\y\B'" by fact with xEq yEq have Eq: "\\x'\(\\*xvec'\\\x\B) = \\y'\(\\*yvec'\\\y\B')" by simp have "Suc n = length xvec" by fact with xEq have L'': "n = length xvec'" by simp have "\\x'\(\\*xvec'\B) = \\y'\(\\*yvec'\([(x, y)] \ B'))" proof(case_tac "x'=y'") assume x'eqy': "x' = y'" with Eq have "\\*xvec'\\\x\B = \\*yvec'\\\y\B'" by(simp add: boundOutput.inject alpha) hence "\\*xvec'\B = \\*yvec'\([(x, y)] \ B')" using L' xFreshB' xFreshxvec' xFreshyvec' L'' by(rule_tac Suc) with x'eqy' show ?thesis by(simp add: boundOutput.inject alpha) next assume x'ineqy': "x' \ y'" with Eq have Eq': "\\*xvec'\\\x\B = \\*([(x', y')] \ yvec')\\\([(x', y')] \ y)\([(x', y')] \ B')" and x'FreshB': "x' \ \\*yvec'\\\y\B'" by(simp add: boundOutput.inject alpha eqvts)+ from L' have "length xvec' = length ([(x', y')] \ yvec')" by simp moreover from xineqx' xineqy' xFreshB' have "x \ [(x', y')] \ B'" by(simp add: fresh_left calc_atm) moreover from xineqx' xineqy' xFreshyvec' have "x \ [(x', y')] \ yvec'" by(simp add: fresh_left calc_atm) ultimately have "\\*xvec'\B = \\*([(x', y')] \ yvec')\([(x, ([(x', y')] \ y))] \ [(x', y')] \ B')" using Eq' xFreshxvec' L'' by(rule_tac Suc) moreover from x'FreshB' have "x' \ \\*yvec'\([(x, y)] \ B')" proof(case_tac "x' \ yvec'") assume "x' \ yvec'" with x'FreshB' have x'FreshB': "x' \ \\y\B'" by(simp add: fresh_def BOresChainSupp) show ?thesis proof(case_tac "x'=y") assume x'eqy: "x' = y" show ?thesis proof(case_tac "x=y") assume "x=y" with xFreshB' x'eqy show ?thesis by(simp add: BOresChainSupp fresh_def) next assume "x \ y" with \x \ B'\ have "y \ [(x, y)] \ B'" by(simp add: fresh_left calc_atm) with x'eqy show ?thesis by(simp add: BOresChainSupp fresh_def) qed next assume x'ineqy: "x' \ y" with x'FreshB' have "x' \ B'" by(simp add: abs_fresh) with xineqx' x'ineqy have "x' \ ([(x, y)] \ B')" by(simp add: fresh_left calc_atm) thus ?thesis by(simp add: BOresChainSupp fresh_def) qed next assume "\x' \ yvec'" thus ?thesis by(simp add: BOresChainSupp fresh_def) qed ultimately show ?thesis using x'ineqy' xineqx' xineqy' apply(simp add: boundOutput.inject alpha eqvts) apply(subst perm_compose[of "[(x', y')]"]) by(simp add: calc_atm) qed with xEq yEq show ?case by simp qed lemma boundOutputPar1Dest: fixes xvec :: "name list" and M :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and yvec :: "name list" and N :: 'a and Q :: "('a, 'b, 'c) psi" and R :: "('a, 'b, 'c) psi" assumes "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" and "xvec \* R" and "yvec \* R" obtains T where "P = T \ R" and "\\*xvec\M \' T = \\*yvec\N \' Q" proof - assume "\T. \P = T \ R; \\*xvec\M \' T = \\*yvec\N \' Q\ \ thesis" moreover obtain n where "n = length xvec" by auto with assms have "\T. P = T \ R \ \\*xvec\M \' T = \\*yvec\N \' Q" proof(induct n arbitrary: xvec yvec M N P Q R) case(0 xvec yvec M N P Q R) have Eq: "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" by fact from \0 = length xvec\ have "xvec = []" by auto moreover with Eq have "yvec = []" by(case_tac yvec) auto ultimately show ?case using Eq by(simp add: boundOutput.inject) next case(Suc n xvec yvec M N P Q R) from \Suc n = length xvec\ obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" by(case_tac xvec) auto from \\\*xvec\M \' P = \\*yvec\N \' (Q \ R)\ \xvec = x # xvec'\ obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' (Q \ R)" and "yvec = y#yvec'" by(case_tac yvec) auto hence EQ: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' (Q \ R))" by simp from \xvec \* R\ \yvec \* R\ \xvec = x#xvec'\ \yvec = y#yvec'\ have "x \ R" and "xvec' \* R" and "y \ R" and "yvec' \* R" by auto show ?case proof(case_tac "x = y") assume "x = y" with EQ have "\\*xvec'\M \' P = \\*yvec'\N \' (Q \ R)" by(simp add: boundOutput.inject alpha) with \xvec' \* R\ \yvec' \* R\ \length xvec' = n\ obtain T where "P = T \ R" and "\\*xvec'\M \' T = \\*yvec'\N \' Q" by(drule_tac Suc) auto with \xvec=x#xvec'\ \yvec=y#yvec'\ \x=y\ show ?case by(force simp add: boundOutput.inject alpha) next assume "x \ y" with EQ \x \ R\ \y \ R\ have "\\*xvec'\M \' P = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' (([(x, y)] \ Q) \ R)" and xFreshQR: "x \ \\*yvec'\N \' (Q \ R)" by(simp add: boundOutput.inject alpha eqvts)+ moreover from \yvec' \* R\ have "([(x, y)] \ yvec') \* ([(x, y)] \ R)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \x \ R\ \y \ R\ have "([(x, y)] \ yvec') \* R" by simp moreover note \xvec' \* R\ \length xvec' = n\ ultimately obtain T where "P = T \ R" and A: "\\*xvec'\M \' T = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ Q)" by(drule_tac Suc) auto from A have "\\x\(\\*xvec'\M \' T) = \\x\(\\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ Q))" by(simp add: boundOutput.inject alpha) moreover from xFreshQR have "x \ \\*yvec'\N \' Q" by(force simp add: boundOutputFresh) ultimately show ?thesis using \P = T \ R\ \xvec=x#xvec'\ \yvec=y#yvec'\ xFreshQR by(force simp add: alphaBoundOutput name_swap eqvts) qed qed ultimately show ?thesis by blast qed lemma boundOutputPar1Dest': fixes xvec :: "name list" and M :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and yvec :: "name list" and N :: 'a and Q :: "('a, 'b, 'c) psi" and R :: "('a, 'b, 'c) psi" assumes "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" and "xvec \* yvec" obtains T p where "set p \ set xvec \ set yvec" and "P = T \ (p \ R)" and "\\*xvec\M \' T = \\*yvec\N \' Q" proof - assume "\p T. \set p \ set xvec \ set yvec; P = T \ (p \ R); \\*xvec\M \' T = \\*yvec\N \' Q\ \ thesis" moreover obtain n where "n = length xvec" by auto with assms have "\p T. set p \ set xvec \ set yvec \ P = T \ (p \ R) \ \\*xvec\M \' T = \\*yvec\N \' Q" proof(induct n arbitrary: xvec yvec M N P Q R) case(0 xvec yvec M N P Q R) have Eq: "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" by fact from \0 = length xvec\ have "xvec = []" by auto moreover with Eq have "yvec = []" by(case_tac yvec) auto ultimately show ?case using Eq by(simp add: boundOutput.inject) next case(Suc n xvec yvec M N P Q R) from \Suc n = length xvec\ obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" by(case_tac xvec) auto from \\\*xvec\M \' P = \\*yvec\N \' (Q \ R)\ \xvec = x # xvec'\ obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' (Q \ R)" and "yvec = y#yvec'" by(case_tac yvec) auto hence Eq: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' (Q \ R))" by simp from \xvec = x#xvec'\ \yvec=y#yvec'\ \xvec \* yvec\ have "x \ y" and "x \ yvec'" and "y \ xvec'" and "xvec' \* yvec'" by auto from Eq \x \ y\ have Eq': "\\*xvec'\M \' P = [(x, y)] \ \\*yvec'\N \' (Q \ R)" and xFreshQR: "x \ \\*yvec'\N \' (Q \ R)" by(simp add: boundOutput.inject alpha)+ have IH: "\xvec yvec M N P Q R. \\\*xvec\M \' (P::('a, 'b, 'c) psi) = \\*yvec\N \' (Q \ R); xvec \* yvec; n = length xvec\ \ \p T. set p \ set xvec \ set yvec \ P = T \ (p \ R) \ \\*xvec\M \' T = \\*yvec\N \' Q" by fact show ?case proof(case_tac "x \ \\*xvec'\M \' P") assume "x \ \\*xvec'\M \' P" with Eq have yFreshQR: "y \ \\*yvec'\N \' (Q \ R)" by(rule boundOutputEqFresh) with Eq' xFreshQR have "\\*xvec'\M \' P = \\*yvec'\N \' (Q \ R)" by simp with \xvec' \* yvec'\ \length xvec' = n\ obtain p T where S: "set p \ set xvec' \ set yvec'" and "P = T \ (p \ R)" and A: "\\*xvec'\M \' T = \\*yvec'\N \' Q" by(drule_tac IH) auto from yFreshQR xFreshQR have yFreshQ: "y \ \\*yvec'\N \' Q" and xFreshQ: "x \ \\*yvec'\N \' Q" by(force simp add: BOresChainSupp fresh_def boundOutput.supp psi.supp)+ hence "\\x\(\\*yvec'\N \' Q) = \\y\(\\*yvec'\N \' Q)" by (subst alphaBoundOutput) simp+ with A have "\\x\(\\*xvec'\M \' T) = \\y\(\\*yvec'\N \' Q)" by simp with \xvec=x#xvec'\ \yvec=y#yvec'\ S \P = T \ (p \ R)\ show ?case by auto next assume "\(x \ \\*xvec'\M \' P)" hence "x \ supp(\\*xvec'\M \' P)" by(simp add: fresh_def) with Eq have "y \ supp(\\*yvec'\N \' (Q \ R))" by(rule boundOutputEqSupp) hence "y \ yvec'" by(simp add: BOresChainSupp fresh_def) with Eq' \x \ yvec'\ have "\\*xvec'\M \' P = \\*yvec'\([(x, y)] \ N) \' (([(x, y)] \ Q) \ ([(x, y)] \ R))" by(simp add: eqvts) moreover note \xvec' \* yvec'\ \length xvec' = n\ ultimately obtain p T where S: "set p \ set xvec' \ set yvec'" and "P = T \ (p \ [(x, y)] \ R)" and A: "\\*xvec'\M \' T = \\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ Q)" by(drule_tac IH) auto from S have "set(p@[(x, y)]) \ set(x#xvec') \ set(y#yvec')" by auto moreover from \P = T \ (p \ [(x, y)] \ R)\ have "P = T \ ((p @ [(x, y)]) \ R)" by(simp add: pt2[OF pt_name_inst]) moreover from xFreshQR have xFreshQ: "x \ \\*yvec'\N \' Q" by(force simp add: BOresChainSupp fresh_def boundOutput.supp psi.supp)+ with \x \ yvec'\ \y \ yvec'\ \x \ y\ have "y \ \\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ Q)" by(simp add: fresh_left calc_atm) with \x \ yvec'\ \y \ yvec'\ have "\\x\(\\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ Q)) = \\y\(\\*yvec'\N \' Q)" by(subst alphaBoundOutput) (assumption | simp add: eqvts)+ with A have "\\x\(\\*xvec'\M \' T) = \\y\(\\*yvec'\N \' Q)" by simp ultimately show ?thesis using \xvec=x#xvec'\ \yvec=y#yvec'\ by(rule_tac x="p@[(x, y)]" in exI) force qed qed ultimately show ?thesis by blast qed lemma boundOutputPar2Dest: fixes xvec :: "name list" and M :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and yvec :: "name list" and N :: 'a and Q :: "('a, 'b, 'c) psi" and R :: "('a, 'b, 'c) psi" assumes "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" and "xvec \* Q" and "yvec \* Q" obtains T where "P = Q \ T" and "\\*xvec\M \' T = \\*yvec\N \' R" proof - assume "\T. \P = Q \ T; \\*xvec\M \' T = \\*yvec\N \' R\ \ thesis" moreover obtain n where "n = length xvec" by auto with assms have "\T. P = Q \ T \ \\*xvec\M \' T = \\*yvec\N \' R" proof(induct n arbitrary: xvec yvec M N P Q R) case(0 xvec yvec M N P Q R) have Eq: "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" by fact from \0 = length xvec\ have "xvec = []" by auto moreover with Eq have "yvec = []" by(case_tac yvec) auto ultimately show ?case using Eq by(simp add: boundOutput.inject) next case(Suc n xvec yvec M N P Q R) from \Suc n = length xvec\ obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" by(case_tac xvec) auto from \\\*xvec\M \' P = \\*yvec\N \' (Q \ R)\ \xvec = x # xvec'\ obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' (Q \ R)" and "yvec = y#yvec'" by(case_tac yvec) auto hence EQ: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' (Q \ R))" by simp from \xvec \* Q\ \yvec \* Q\ \xvec = x#xvec'\ \yvec = y#yvec'\ have "x \ Q" and "xvec' \* Q" and "y \ Q" and "yvec' \* Q" by auto have IH: "\xvec yvec M N P Q R. \\\*xvec\M \' (P::('a, 'b, 'c) psi) = \\*yvec\N \' (Q \ R); xvec \* Q; yvec \* Q; n = length xvec\ \ \T. P = Q \ T \ \\*xvec\M \' T = \\*yvec\N \' R" by fact show ?case proof(case_tac "x = y") assume "x = y" with EQ have "\\*xvec'\M \' P = \\*yvec'\N \' (Q \ R)" by(simp add: boundOutput.inject alpha) with \xvec' \* Q\ \yvec' \* Q\ \length xvec' = n\ obtain T where "P = Q \ T" and "\\*xvec'\M \' T = \\*yvec'\N \' R" by(drule_tac IH) auto with \xvec=x#xvec'\ \yvec=y#yvec'\ \x=y\ show ?case by(force simp add: boundOutput.inject alpha) next assume "x \ y" with EQ \x \ Q\ \y \ Q\ have "\\*xvec'\M \' P = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' (Q \ ([(x, y)] \ R))" and xFreshQR: "x \ \\*yvec'\N \' (Q \ R)" by(simp add: boundOutput.inject alpha eqvts)+ moreover from \yvec' \* Q\ have "([(x, y)] \ yvec') \* ([(x, y)] \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \x \ Q\ \y \ Q\ have "([(x, y)] \ yvec') \* Q" by simp moreover note \xvec' \* Q\ \length xvec' = n\ ultimately obtain T where "P = Q \ T" and A: "\\*xvec'\M \' T = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ R)" by(drule_tac IH) auto from A have "\\x\(\\*xvec'\M \' T) = \\x\(\\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ R))" by(simp add: boundOutput.inject alpha) moreover from xFreshQR have "x \ \\*yvec'\N \' R" by(force simp add: boundOutputFresh) ultimately show ?thesis using \P = Q \ T\ \xvec=x#xvec'\ \yvec=y#yvec'\ xFreshQR by(force simp add: alphaBoundOutput name_swap eqvts) qed qed ultimately show ?thesis by blast qed lemma boundOutputPar2Dest': fixes xvec :: "name list" and M :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and yvec :: "name list" and N :: 'a and Q :: "('a, 'b, 'c) psi" and R :: "('a, 'b, 'c) psi" assumes "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" and "xvec \* yvec" obtains T p where "set p \ set xvec \ set yvec" and "P = (p \ Q) \ T" and "\\*xvec\M \' T = \\*yvec\N \' R" proof - assume "\p T. \set p \ set xvec \ set yvec; P = (p \ Q) \ T; \\*xvec\M \' T = \\*yvec\N \' R\ \ thesis" moreover obtain n where "n = length xvec" by auto with assms have "\p T. set p \ set xvec \ set yvec \ P = (p \ Q) \ T \ \\*xvec\M \' T = \\*yvec\N \' R" proof(induct n arbitrary: xvec yvec M N P Q R) case(0 xvec yvec M N P Q R) have Eq: "\\*xvec\M \' P = \\*yvec\N \' (Q \ R)" by fact from \0 = length xvec\ have "xvec = []" by auto moreover with Eq have "yvec = []" by(case_tac yvec) auto ultimately show ?case using Eq by(simp add: boundOutput.inject) next case(Suc n xvec yvec M N P Q R) from \Suc n = length xvec\ obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" by(case_tac xvec) auto from \\\*xvec\M \' P = \\*yvec\N \' (Q \ R)\ \xvec = x # xvec'\ obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' (Q \ R)" and "yvec = y#yvec'" by(case_tac yvec) auto hence Eq: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' (Q \ R))" by simp from \xvec = x#xvec'\ \yvec=y#yvec'\ \xvec \* yvec\ have "x \ y" and "x \ yvec'" and "y \ xvec'" and "xvec' \* yvec'" by auto from Eq \x \ y\ have Eq': "\\*xvec'\M \' P = [(x, y)] \ \\*yvec'\N \' (Q \ R)" and xFreshQR: "x \ \\*yvec'\N \' (Q \ R)" by(simp add: boundOutput.inject alpha)+ have IH: "\xvec yvec M N P Q R. \\\*xvec\M \' (P::('a, 'b, 'c) psi) = \\*yvec\N \' (Q \ R); xvec \* yvec; n = length xvec\ \ \p T. set p \ set xvec \ set yvec \ P = (p \ Q) \ T \ \\*xvec\M \' T = \\*yvec\N \' R" by fact show ?case proof(case_tac "x \ \\*xvec'\M \' P") assume "x \ \\*xvec'\M \' P" with Eq have yFreshQR: "y \ \\*yvec'\N \' (Q \ R)" by(rule boundOutputEqFresh) with Eq' xFreshQR have "\\*xvec'\M \' P = \\*yvec'\N \' (Q \ R)" by simp with \xvec' \* yvec'\ \length xvec' = n\ obtain p T where S: "set p \ set xvec' \ set yvec'" and "P = (p \ Q) \ T" and A: "\\*xvec'\M \' T = \\*yvec'\N \' R" by(drule_tac IH) auto from yFreshQR xFreshQR have yFreshR: "y \ \\*yvec'\N \' R" and xFreshQ: "x \ \\*yvec'\N \' R" by(force simp add: BOresChainSupp fresh_def boundOutput.supp psi.supp)+ hence "\\x\(\\*yvec'\N \' R) = \\y\(\\*yvec'\N \' R)" by (subst alphaBoundOutput) simp+ with A have "\\x\(\\*xvec'\M \' T) = \\y\(\\*yvec'\N \' R)" by simp with \xvec=x#xvec'\ \yvec=y#yvec'\ S \P = (p \ Q) \ T\ show ?case by auto next assume "\(x \ \\*xvec'\M \' P)" hence "x \ supp(\\*xvec'\M \' P)" by(simp add: fresh_def) with Eq have "y \ supp(\\*yvec'\N \' (Q \ R))" by(rule boundOutputEqSupp) hence "y \ yvec'" by(simp add: BOresChainSupp fresh_def) with Eq' \x \ yvec'\ have "\\*xvec'\M \' P = \\*yvec'\([(x, y)] \ N) \' (([(x, y)] \ Q) \ ([(x, y)] \ R))" by(simp add: eqvts) moreover note \xvec' \* yvec'\ \length xvec' = n\ ultimately obtain p T where S: "set p \ set xvec' \ set yvec'" and "P = (p \ [(x, y)] \ Q) \ T" and A: "\\*xvec'\M \' T = \\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ R)" by(drule_tac IH) auto from S have "set(p@[(x, y)]) \ set(x#xvec') \ set(y#yvec')" by auto moreover from \P = (p \ [(x, y)] \ Q) \ T\ have "P = ((p @ [(x, y)]) \ Q) \ T" by(simp add: pt2[OF pt_name_inst]) moreover from xFreshQR have xFreshR: "x \ \\*yvec'\N \' R" by(force simp add: BOresChainSupp fresh_def boundOutput.supp psi.supp)+ with \x \ yvec'\ \y \ yvec'\ \x \ y\ have "y \ \\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ R)" by(simp add: fresh_left calc_atm) with \x \ yvec'\ \y \ yvec'\ have "\\x\(\\*yvec'\([(x, y)] \ N) \' ([(x, y)] \ R)) = \\y\(\\*yvec'\N \' R)" by(subst alphaBoundOutput) (assumption | simp add: eqvts)+ with A have "\\x\(\\*xvec'\M \' T) = \\y\(\\*yvec'\N \' R)" by simp ultimately show ?thesis using \xvec=x#xvec'\ \yvec=y#yvec'\ by(rule_tac x="p@[(x, y)]" in exI) force qed qed ultimately show ?thesis by blast qed lemma boundOutputApp: fixes xvec :: "name list" and yvec :: "name list" and B :: "('a::fs_name, 'b::fs_name, 'c::fs_name) boundOutput" shows "\\*(xvec@yvec)\B = \\*xvec\(\\*yvec\B)" by(induct xvec) auto lemma openInjectAuxAuxAux: fixes x :: name and xvec :: "name list" shows "\y yvec. x # xvec = yvec @ [y] \ length xvec = length yvec" apply(induct xvec arbitrary: x) apply auto apply(subgoal_tac "\y yvec. a # xvec = yvec @ [y] \ length xvec = length yvec") apply(clarify) apply(rule_tac x=y in exI) by auto lemma openInjectAuxAux: fixes xvec1 :: "name list" and xvec2 :: "name list" and yvec :: "name list" assumes "length(xvec1@xvec2) = length yvec" shows "\yvec1 yvec2. yvec = yvec1@yvec2 \ length xvec1 = length yvec1 \ length xvec2 = length yvec2" using assms apply(induct yvec arbitrary: xvec1) apply simp apply simp apply(case_tac xvec1) apply simp apply simp apply(subgoal_tac "\yvec1 yvec2. yvec = yvec1 @ yvec2 \ length list = length yvec1 \ length xvec2 = length yvec2") apply(clarify) apply(rule_tac x="a#yvec1" in exI) apply(rule_tac x=yvec2 in exI) by auto lemma openInjectAux: fixes xvec1 :: "name list" and x :: name and xvec2 :: "name list" and yvec :: "name list" assumes "length(xvec1@x#xvec2) = length yvec" shows "\yvec1 y yvec2. yvec = yvec1@y#yvec2 \ length xvec1 = length yvec1 \ length xvec2 = length yvec2" using assms apply(case_tac yvec) apply simp apply simp apply(subgoal_tac "\(yvec1::name list) (yvec2::name list). yvec1@yvec2 = list \ length xvec1 = length yvec1 \ length xvec2 = length yvec2") apply(clarify) apply hypsubst_thin apply simp apply(subgoal_tac "\y (yvec::name list). a # yvec1 = yvec @ [y] \ length yvec1 = length yvec") apply(clarify) apply(rule_tac x=yvec in exI) apply(rule_tac x=y in exI) apply simp apply(rule_tac x=yvec2 in exI) apply simp apply(rule openInjectAuxAuxAux) apply(insert openInjectAuxAux) apply simp by blast lemma boundOutputOpenDest: fixes yvec :: "name list" and M :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and xvec1 :: "name list" and x :: name and xvec2 :: "name list" and N :: 'a and Q :: "('a, 'b, 'c) psi" assumes Eq: "\\*(xvec1@x#xvec2)\M \' P = \\*yvec\N \' Q" and "x \ xvec1" and "x \ yvec" and "x \ N" and "x \ Q" and "distinct yvec" obtains yvec1 y yvec2 where "yvec=yvec1@y#yvec2" and "length xvec1 = length yvec1" and "length xvec2 = length yvec2" and "\\*(xvec1@xvec2)\M \' P = \\*(yvec1@yvec2)\([(x, y)] \ N) \' ([(x, y)] \ Q)" proof - assume Ass: "\yvec1 y yvec2. \yvec = yvec1 @ y # yvec2; length xvec1 = length yvec1; length xvec2 = length yvec2; \\*(xvec1 @ xvec2)\M \' P = \\*(yvec1 @ yvec2)\([(x, y)] \ N) \' ([(x, y)] \ Q)\ \ thesis" from Eq have "length(xvec1@x#xvec2) = length yvec" by(rule boundOutputChainEqLength) then obtain yvec1 y yvec2 where A: "yvec = yvec1@y#yvec2" and "length xvec1 = length yvec1" and "length xvec2 = length yvec2" by(metis openInjectAux sym) from \distinct yvec\ A have "y \ yvec2" by simp from A \x \ yvec\ have "x \ yvec2" and "x \ yvec1" by simp+ with Eq \length xvec1 = length yvec1\ \x \ N\ \x \ Q\ \y \ yvec2\ \x \ xvec1\ A have "\\*(xvec1@xvec2)\M \' P = \\*(yvec1@yvec2)\([(x, y)] \ N) \' ([(x, y)] \ Q)" by(force dest: boundOutputChainOpenIH simp add: boundOutputApp BOresChainSupp fresh_def boundOutput.supp eqvts) with \length xvec1 = length yvec1\ \length xvec2 = length yvec2\ A Ass show ?thesis by blast qed lemma boundOutputOpenDest': fixes yvec :: "name list" and M :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and xvec1 :: "name list" and x :: name and xvec2 :: "name list" and N :: 'a and Q :: "('a, 'b, 'c) psi" assumes Eq: "\\*(xvec1@x#xvec2)\M \' P = \\*yvec\N \' Q" and "x \ xvec1" and "x \ yvec" and "x \ N" and "x \ Q" obtains yvec1 y yvec2 where "yvec=yvec1@y#yvec2" and "length xvec1 = length yvec1" and "length xvec2 = length yvec2" and "\\*(xvec1@xvec2)\M \' P = \\*(yvec1@[(x, y)] \ yvec2)\([(x, y)] \ N) \' ([(x, y)] \ Q)" proof - assume Ass: "\yvec1 y yvec2. \yvec = yvec1 @ y # yvec2; length xvec1 = length yvec1; length xvec2 = length yvec2; \\*(xvec1 @ xvec2)\M \' P = \\*(yvec1 @ ([(x, y)] \ yvec2))\([(x, y)] \ N) \' ([(x, y)] \ Q)\ \ thesis" from Eq have "length(xvec1@x#xvec2) = length yvec" by(rule boundOutputChainEqLength) then obtain yvec1 y yvec2 where A: "yvec = yvec1@y#yvec2" and "length xvec1 = length yvec1" and "length xvec2 = length yvec2" by(metis openInjectAux sym) from A \x \ yvec\ have "x \ yvec2" and "x \ yvec1" by simp+ with Eq \length xvec1 = length yvec1\ \x \ N\ \x \ Q\ \x \ xvec1\ A have "\\*(xvec1@xvec2)\M \' P = \\*(yvec1@([(x, y)] \ yvec2))\([(x, y)] \ N) \' ([(x, y)] \ Q)" by(force dest: boundOutputChainOpenIH simp add: boundOutputApp BOresChainSupp fresh_def boundOutput.supp eqvts) with \length xvec1 = length yvec1\ \length xvec2 = length yvec2\ A Ass show ?thesis by blast qed lemma boundOutputScopeDest: fixes xvec :: "name list" and M :: "'a::fs_name" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and yvec :: "name list" and N :: 'a and x :: name and Q :: "('a, 'b, 'c) psi" assumes "\\*xvec\M \' P = \\*yvec\N \' \\z\Q" and "z \ xvec" and "z \ yvec" obtains R where "P = \\z\R" and "\\*xvec\M \' R = \\*yvec\N \' Q" proof - assume "\R. \P = \\z\R; \\*xvec\M \' R = \\*yvec\N \' Q\ \ thesis" moreover obtain n where "n = length xvec" by auto with assms have "\R. P = \\z\R \ \\*xvec\M \' R = \\*yvec\N \' Q" proof(induct n arbitrary: xvec yvec M N P Q z) case(0 xvec yvec M N P Q z) have Eq: "\\*xvec\M \' P = \\*yvec\N \' \\z\Q" by fact from \0 = length xvec\ have "xvec = []" by auto moreover with Eq have "yvec = []" by(case_tac yvec) auto ultimately show ?case using Eq by(simp add: boundOutput.inject) next case(Suc n xvec yvec M N P Q z) from \Suc n = length xvec\ obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" by(case_tac xvec) auto from \\\*xvec\M \' P = \\*yvec\N \' (\\z\Q)\ \xvec = x # xvec'\ obtain y yvec' where "\\*(x#xvec')\M \' P = \\*(y#yvec')\N \' \\z\Q" and "yvec = y#yvec'" by(case_tac yvec) auto hence EQ: "\\x\(\\*xvec'\M \' P) = \\y\(\\*yvec'\N \' \\z\Q)" by simp from \z \ xvec\ \z \ yvec\ \xvec = x#xvec'\ \yvec = y#yvec'\ have "z \ x" and "z \ y" and "z \ xvec'" and "z \ yvec'" by simp+ have IH: "\xvec yvec M N P Q z. \\\*xvec\M \' (P::('a, 'b, 'c) psi) = \\*yvec\N \' \\z\Q; z \ xvec; z \ yvec; n = length xvec\ \ \R. P = \\z\R \ \\*xvec\M \' R = \\*yvec\N \' Q" by fact show ?case proof(case_tac "x = y") assume "x = y" with EQ have "\\*xvec'\M \' P = \\*yvec'\N \' \\z\Q" by(simp add: boundOutput.inject alpha) with \z \ xvec'\ \z \ yvec'\ \length xvec' = n\ obtain R where "P = \\z\R" and "\\*xvec'\M \' R = \\*yvec'\N \' Q" by(drule_tac IH) auto with \xvec=x#xvec'\ \yvec=y#yvec'\ \x=y\ show ?case by(force simp add: boundOutput.inject alpha) next assume "x \ y" with EQ \z \ x\ \z \ y\ have "\\*xvec'\M \' P = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' \\z\([(x, y)] \ Q)" and xFreshzQ: "x \ \\*yvec'\N \' \\z\Q" by(simp add: boundOutput.inject alpha eqvts)+ moreover from \z \ x\ \z \ y\ \z \ yvec'\ \x \ y\ have "z \ ([(x, y)] \ yvec')" by(simp add: fresh_left calc_atm) moreover note \z \ xvec'\ \length xvec' = n\ ultimately obtain R where "P = \\z\R" and A: "\\*xvec'\M \' R = \\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ Q)" by(drule_tac IH) auto from A have "\\x\(\\*xvec'\M \' R) = \\x\(\\*([(x, y)] \ yvec')\([(x, y)] \ N) \' ([(x, y)] \ Q))" by(simp add: boundOutput.inject alpha) moreover from xFreshzQ \z \ x\ have "x \ \\*yvec'\N \' Q" by(simp add: boundOutputFresh abs_fresh) ultimately show ?thesis using \P = \\z\R\ \xvec=x#xvec'\ \yvec=y#yvec'\ xFreshzQ by(force simp add: alphaBoundOutput name_swap eqvts) qed qed ultimately show ?thesis by blast qed nominal_datatype ('a, 'b, 'c) residual = RIn "'a::fs_name" 'a "('a, 'b::fs_name, 'c::fs_name) psi" | ROut 'a "('a, 'b, 'c) boundOutput" | RTau "('a, 'b, 'c) psi" nominal_datatype 'a action = In "'a::fs_name" 'a ("_\_\" [90, 90] 90) | Out "'a::fs_name" "name list" 'a ("_\\*_\\_\" [90, 90, 90] 90) | Tau ("\" 90) nominal_primrec bn :: "('a::fs_name) action \ name list" where "bn (M\N\) = []" | "bn (M\\*xvec\\N\) = xvec" | "bn (\) = []" by(rule TrueI)+ lemma bnEqvt[eqvt]: fixes p :: "name prm" and \ :: "('a::fs_name) action" shows "(p \ bn \) = bn(p \ \)" by(nominal_induct \ rule: action.strong_induct) auto nominal_primrec create_residual :: "('a::fs_name) action \ ('a, 'b::fs_name, 'c::fs_name) psi \ ('a, 'b, 'c) residual" ("_ \ _" [80, 80] 80) where "(M\N\) \ P = RIn M N P" | "M\\*xvec\\N\ \ P = ROut M (\\*xvec\(N \' P))" | "\ \ P = (RTau P)" by(rule TrueI)+ nominal_primrec subject :: "('a::fs_name) action \ 'a option" where "subject (M\N\) = Some M" | "subject (M\\*xvec\\N\) = Some M" | "subject (\) = None" by(rule TrueI)+ nominal_primrec object :: "('a::fs_name) action \ 'a option" where "object (M\N\) = Some N" | "object (M\\*xvec\\N\) = Some N" | "object (\) = None" by(rule TrueI)+ lemma optionFreshChain[simp]: fixes xvec :: "name list" and X :: "name set" shows "xvec \* (Some x) = xvec \* x" and "X \* (Some x) = X \* x" and "xvec \* None" and "X \* None" by(auto simp add: fresh_star_def fresh_some fresh_none) lemmas [simp] = fresh_some fresh_none lemma actionFresh[simp]: fixes x :: name and \ :: "('a::fs_name) action" shows "(x \ \) = (x \ (subject \) \ x \ (bn \) \ x \ (object \))" by(nominal_induct \ rule: action.strong_induct) auto lemma actionFreshChain[simp]: fixes X :: "name set" and \ :: "('a::fs_name) action" and xvec :: "name list" shows "(X \* \) = (X \* (subject \) \ X \* (bn \) \ X \* (object \))" and "(xvec \* \) = (xvec \* (subject \) \ xvec \* (bn \) \ xvec \* (object \))" by(auto simp add: fresh_star_def) lemma subjectEqvt[eqvt]: fixes p :: "name prm" and \ :: "('a::fs_name) action" shows "(p \ subject \) = subject(p \ \)" by(nominal_induct \ rule: action.strong_induct) auto lemma okjectEqvt[eqvt]: fixes p :: "name prm" and \ :: "('a::fs_name) action" shows "(p \ object \) = object(p \ \)" by(nominal_induct \ rule: action.strong_induct) auto lemma create_residualEqvt[eqvt]: fixes p :: "name prm" and \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" shows "(p \ (\ \ P)) = (p \ \) \ (p \ P)" by(nominal_induct \ rule: action.strong_induct) (auto simp add: eqvts) lemma residualFresh: fixes x :: name and \ :: "'a::fs_name action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" shows "(x \ (\ \ P)) = (x \ (subject \) \ (x \ (set(bn(\))) \ (x \ object(\) \ x \ P)))" by(nominal_induct \ rule: action.strong_induct) (auto simp add: fresh_some fresh_none boundOutputFresh) lemma residualFresh2[simp]: fixes x :: name and \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" assumes "x \ \" and "x \ P" shows "x \ \ \ P" using assms by(nominal_induct \ rule: action.strong_induct) auto lemma residualFreshChain2[simp]: fixes xvec :: "name list" and X :: "name set" and \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" shows "\xvec \* \; xvec \* P\ \ xvec \* (\ \ P)" and "\X \* \; X \* P\ \ X \* (\ \ P)" by(auto simp add: fresh_star_def) lemma residualFreshSimp[simp]: fixes x :: name and M :: "'a::fs_name" and N :: 'a and P :: "('a, 'b::fs_name, 'c::fs_name) psi" shows "x \ (M\N\ \ P) = (x \ M \ x \ N \ x \ P)" and "x \ (M\\*xvec\\N\ \ P) = (x \ M \ x \ (\\*xvec\(N \' P)))" and "x \ (\ \ P) = (x \ P)" by(auto simp add: residualFresh) lemma residualInject': shows "(\ \ P = RIn M N Q) = (P = Q \ \ = M\N\)" and "(\ \ P = ROut M B) = (\xvec N. \ = M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" and "(\ \ P = RTau Q) = (\ = \ \ P = Q)" and "(RIn M N Q = \ \ P) = (P = Q \ \ = M\N\)" and "(ROut M B = \ \ P) = (\xvec N. \ = M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" and "(RTau Q = \ \ P) = (\ = \ \ P = Q)" proof - show "(\ \ P = RIn M N Q) = (P = Q \ \ = M\N\)" by(nominal_induct \ rule: action.strong_induct) (auto simp add: residual.inject action.inject) next show "(\ \ P = ROut M B) = (\xvec N. \ = M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" by(nominal_induct \ rule: action.strong_induct) (auto simp add: residual.inject action.inject) next show "(\ \ P = RTau Q) = (\ = \ \ P = Q)" by(nominal_induct \ rule: action.strong_induct) (auto simp add: residual.inject action.inject) next show "(RIn M N Q = \ \ P) = (P = Q \ \ = M\N\)" by(nominal_induct \ rule: action.strong_induct) (auto simp add: residual.inject action.inject) next show "(ROut M B = \ \ P) = (\xvec N. \ = M\\*xvec\\N\ \ B = \\*xvec\(N \' P))" by(nominal_induct \ rule: action.strong_induct) (auto simp add: residual.inject action.inject) next show "(RTau Q = \ \ P) = (\ = \ \ P = Q)" by(nominal_induct \ rule: action.strong_induct) (auto simp add: residual.inject action.inject) qed lemma residualFreshChainSimp[simp]: fixes xvec :: "name list" and X :: "name set" and M :: "'a::fs_name" and N :: 'a and yvec :: "name list" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" shows "xvec \* (M\N\ \ P) = (xvec \* M \ xvec \* N \ xvec \* P)" and "xvec \* (M\\*yvec\\N\ \ P) = (xvec \* M \ xvec \* (\\*yvec\(N \' P)))" and "xvec \* (\ \ P) = (xvec \* P)" and "X \* (M\N\ \ P) = (X \* M \ X \* N \ X \* P)" and "X \* (M\\*yvec\\N\ \ P) = (X \* M \ X \* (\\*yvec\(N \' P)))" and "X \* (\ \ P) = (X \* P)" by(auto simp add: fresh_star_def) lemma residualFreshChainSimp2[simp]: fixes xvec :: "name list" and X :: "name set" and M :: "'a::fs_name" and N :: 'a and yvec :: "name list" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" shows "xvec \* (RIn M N P) = (xvec \* M \ xvec \* N \ xvec \* P)" and "xvec \* (ROut M B) = (xvec \* M \ xvec \* B)" and "xvec \* (RTau P) = (xvec \* P)" and "X \* (RIn M N P) = (X \* M \ X \* N \ X \* P)" and "X \* (ROut M B) = (X \* M \ X \* B)" and "X \* (RTau P) = (X \* P)" by(auto simp add: fresh_star_def) lemma freshResidual3[dest]: fixes x :: name and \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" assumes "x \ bn \" and "x \ \ \ P" shows "x \ \" and "x \ P" using assms by(nominal_induct rule: action.strong_induct) auto lemma freshResidualChain3[dest]: fixes xvec :: "name list" and \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" assumes "xvec \* (\ \ P)" and "xvec \* bn \" shows "xvec \* \" and "xvec \* P" using assms by(nominal_induct rule: action.strong_induct) auto lemma freshResidual4[dest]: fixes x :: name and \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" assumes "x \ \ \ P" shows "x \ subject \" using assms by(nominal_induct rule: action.strong_induct) auto lemma freshResidualChain4[dest]: fixes xvec :: "name list" and \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" assumes "xvec \* (\ \ P)" shows "xvec \* subject \" using assms by(nominal_induct rule: action.strong_induct) auto lemma alphaOutputResidual: fixes M :: "'a::fs_name" and xvec :: "name list" and N :: 'a and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and p :: "name prm" assumes "(p \ xvec) \* N" and "(p \ xvec) \* P" and "set p \ set xvec \ set(p \ xvec)" and "set xvec \ set yvec" shows "M\\*yvec\\N\ \ P = M\\*(p \ yvec)\\(p \ N)\ \ (p \ P)" using assms by(simp add: boundOutputChainAlpha'') lemmas[simp del] = create_residual.simps lemma residualInject'': assumes "bn \ = bn \" shows "(\ \ P = \ \ Q) = (\ = \ \ P = Q)" using assms apply(nominal_induct \ rule: action.strong_induct) apply(auto simp add: residual.inject create_residual.simps residualInject' action.inject boundOutput.inject) by(rule_tac x="bn \" in exI) auto lemmas residualInject = residual.inject create_residual.simps residualInject' residualInject'' lemma bnFreshResidual[simp]: fixes \ :: "('a::fs_name) action" shows "(bn \) \* (\ \ P) = bn \ \* (subject \)" by(nominal_induct \ rule: action.strong_induct) (auto simp add: residualFresh fresh_some fresh_star_def) lemma actionCases[case_names cInput cOutput cTau]: fixes \ :: "('a::fs_name) action" assumes "\M N. \ = M\N\ \ Prop" and "\M xvec N. \ = M\\*xvec\\N\ \ Prop" and "\ = \ \ Prop" shows Prop using assms by(nominal_induct \ rule: action.strong_induct) auto lemma actionPar1Dest: fixes \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and \ :: "('a::fs_name) action" and Q :: "('a, 'b, 'c) psi" and R :: "('a, 'b, 'c) psi" assumes "\ \ P = \ \ (Q \ R)" and "bn \ \* bn \" obtains T p where "set p \ set(bn \) \ set(bn \)" and "P = T \ (p \ R)" and "\ \ T = \ \ Q" using assms apply(cases rule: actionCases[where \=\]) apply(auto simp add: residualInject) by(drule_tac boundOutputPar1Dest') auto lemma actionPar2Dest: fixes \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and \ :: "('a::fs_name) action" and Q :: "('a, 'b, 'c) psi" and R :: "('a, 'b, 'c) psi" assumes "\ \ P = \ \ (Q \ R)" and "bn \ \* bn \" obtains T p where "set p \ set(bn \) \ set(bn \)" and "P = (p \ Q) \ T" and "\ \ T = \ \ R" using assms apply(cases rule: actionCases[where \=\]) apply(auto simp add: residualInject) by(drule_tac boundOutputPar2Dest') auto lemma actionScopeDest: fixes \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" fixes \ :: "('a::fs_name) action" and x :: name and Q :: "('a, 'b, 'c) psi" assumes "\ \ P = \ \ \\x\Q" and "x \ bn \" and "x \ bn \" obtains R where "P = \\x\R" and "\ \ R = \ \ Q" using assms apply(cases rule: actionCases[where \=\]) apply(auto simp add: residualInject) by(drule_tac boundOutputScopeDest) auto abbreviation outputJudge ("_\_\" [110, 110] 110) where "M\N\ \ M\\*([])\\N\" declare [[unify_trace_bound=100]] locale env = substPsi substTerm substAssert substCond + assertion SCompose' SImp' SBottom' SChanEq' for substTerm :: "('a::fs_name) \ name list \ 'a::fs_name list \ 'a" and substAssert :: "('b::fs_name) \ name list \ 'a::fs_name list \ 'b" and substCond :: "('c::fs_name) \ name list \ 'a::fs_name list \ 'c" and SCompose' :: "'b \ 'b \ 'b" and SImp' :: "'b \ 'c \ bool" and SBottom' :: 'b and SChanEq' :: "'a \ 'a \ 'c" begin notation SCompose' (infixr "\" 90) notation SImp' ("_ \ _" [85, 85] 85) notation FrameImp ("_ \\<^sub>F _" [85, 85] 85) abbreviation FBottomJudge ("\\<^sub>F" 90) where "\\<^sub>F \ (FAssert SBottom')" notation SChanEq' ("_ \ _" [90, 90] 90) notation substTerm ("_[_::=_]" [100, 100, 100] 100) notation subs ("_[_::=_]" [100, 100, 100] 100) notation AssertionStatEq ("_ \ _" [80, 80] 80) notation FrameStatEq ("_ \\<^sub>F _" [80, 80] 80) notation SBottom' ("\" 190) abbreviation insertAssertion' ("insertAssertion") where "insertAssertion' \ assertionAux.insertAssertion (\)" inductive semantics :: "'b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) residual \ bool" ("_ \ _ \ _" [50, 50, 50] 50) where cInput: "\\ \ M \ K; distinct xvec; set xvec \ supp N; xvec \* Tvec; length xvec = length Tvec; xvec \* \; xvec \* M; xvec \* K\ \ \ \ M\\*xvec N\.P \ K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]" | Output: "\\ \ M \ K\ \ \ \ M\N\.P \ K\N\ \ P" | Case: "\\ \ P \ Rs; (\, P) mem Cs; \ \ \; guarded P\ \ \ \ Cases Cs \ Rs" | cPar1: "\(\ \ \\<^sub>Q) \ P \\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* P'; distinct(bn \); bn \ \* \; bn \ \* \\<^sub>Q; bn \ \* Q; bn \ \* P; bn \ \* (subject \)\ \ \ \ P \ Q \\ \ (P' \ Q)" | cPar2: "\(\ \ \\<^sub>P) \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* Q'; distinct(bn \); bn \ \* \; bn \ \* \\<^sub>P; bn \ \* P; bn \ \* Q; bn \ \* (subject \)\ \ \ \ P \ Q \\ \ (P \ Q')" | cComm1: "\\ \ \\<^sub>Q \ P \ M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \ K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; distinct xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K\ \ \ \ P \ Q \ \ \ \\*xvec\(P' \ Q')" | cComm2: "\\ \ \\<^sub>Q \ P \ M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \ K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; distinct xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K\ \ \ \ P \ Q \ \ \ \\*xvec\(P' \ Q')" | cOpen: "\\ \ P \ M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; x \ xvec; x \ yvec; x \ M; x \ \; distinct xvec; distinct yvec; xvec \* \; xvec \* P; xvec \* M; xvec \* yvec; yvec \* \; yvec \* P; yvec \* M\ \ \ \ \\x\P \ M\\*(xvec@x#yvec)\\N\ \ P'" | cScope: "\\ \ P \\ \ P'; x \ \; x \ \; bn \ \* \; bn \ \* P; bn \ \* (subject \); distinct(bn \)\ \ \ \ \\x\P \\ \ (\\x\P')" | Bang: "\\ \ P \ !P \ Rs; guarded P\ \ \ \ !P \ Rs" abbreviation semanticsBottomJudge ("_ \ _" [50, 50] 50) where "P \ Rs \ \ \ P \ Rs" equivariance env.semantics nominal_inductive2 env.semantics avoids cInput: "set xvec" | cPar1: "set A\<^sub>Q \ set(bn \)" | cPar2: "set A\<^sub>P \ set(bn \)" | cComm1: "set A\<^sub>P \ set A\<^sub>Q \ set xvec" | cComm2: "set A\<^sub>P \ set A\<^sub>Q \ set xvec" | cOpen: "{x} \ set xvec \ set yvec" | cScope: "{x} \ set(bn \)" apply(auto intro: substTerm.subst4Chain subst4Chain simp add: abs_fresh residualFresh) apply(force simp add: fresh_star_def abs_fresh) apply(simp add: boundOutputFresh) apply(simp add: boundOutputFreshSet) apply(simp add: boundOutputFreshSet) by(simp add: fresh_star_def abs_fresh) lemma nilTrans[dest]: fixes \ :: 'b and Rs :: "('a, 'b, 'c) residual" and M :: 'a and xvec :: "name list" and N :: 'a and P :: "('a, 'b, 'c) psi" and K :: 'a and yvec :: "name list" and N' :: 'a and P' :: "('a, 'b, 'c) psi" and CsP :: "('c \ ('a, 'b, 'c) psi) list" and \' :: 'b shows "\ \ \ \ Rs \ False" and "\ \ M\\*xvec N\.P \K\\*yvec\\N'\ \ P' \ False" and "\ \ M\\*xvec N\.P \\ \ P' \ False" and "\ \ M\N\.P \K\N'\ \ P' \ False" and "\ \ M\N\.P \\ \ P' \ False" and "\ \ \\'\ \ Rs \ False" apply(cases rule: semantics.cases) apply auto apply(cases rule: semantics.cases) apply(auto simp add: residualInject) apply(cases rule: semantics.cases) apply(auto simp add: residualInject) apply(cases rule: semantics.cases) apply(auto simp add: residualInject) apply(cases rule: semantics.cases) apply(auto simp add: residualInject) by(cases rule: semantics.cases) (auto simp add: residualInject) lemma residualEq: fixes \ :: "'a action" and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and Q :: "('a, 'b, 'c) psi" assumes "\ \ P = \ \ Q" and "bn \ \* (bn \)" and "distinct(bn \)" and "distinct(bn \)" and "bn \ \* (\ \ P)" and "bn \ \* (\ \ Q)" obtains p where "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "\ = p \ \" and "Q = p \ P" and "bn \ \* \" and "bn \ \* Q" and "bn(p \ \) \* \" and "bn(p \ \) \* P" using assms proof(nominal_induct \ rule: action.strong_induct) case(In M N) thus ?case by(simp add: residualInject) next case(Out M xvec N) thus ?case by(auto simp add: residualInject) (drule_tac boundOutputChainEq'', auto) next case Tau thus ?case by(simp add: residualInject) qed lemma semanticsInduct[consumes 3, case_names cAlpha cInput cOutput cCase cPar1 cPar2 cComm1 cComm2 cOpen cScope cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a action \ ('a, 'b, 'c) psi \ bool" and C :: "'d::fs_name" assumes "\ \ P \\ \ P'" and "bn \ \* (subject \)" and "distinct(bn \)" and rAlpha: "\\ P \ P' p C. \bn \ \* \; bn \ \* P; bn \ \* (subject \); bn \ \* C; bn \ \* (bn(p \ \)); set p \ set(bn \) \ set(bn(p \ \)); distinctPerm p; (bn(p \ \)) \* \; (bn(p \ \)) \* P'; Prop C \ P \ P'\ \ Prop C \ P (p \ \) (p \ P')" and rInput: "\\ M K xvec N Tvec P C. \\ \ M \ K; distinct xvec; set xvec \ supp N; length xvec = length Tvec; xvec \* \; xvec \* M; xvec \* K; xvec \* C\ \ Prop C \ (M\\*xvec N\.P) (K\(N[xvec::=Tvec])\) (P[xvec::=Tvec])" and rOutput: "\\ M K N P C. \\ \ M \ K\ \ Prop C \ (M\N\.P) (K\N\) P" and rCase: "\\ P \ P' \ Cs C. \\ \ P \\ \ P'; \C. Prop C \ P \ P'; (\, P) mem Cs; \ \ \; guarded P\ \ Prop C \ (Cases Cs) \ P'" and rPar1: "\\ \\<^sub>Q P \ P' A\<^sub>Q Q C. \\ \ \\<^sub>Q \ P \\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>Q) P \ P'; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* P'; A\<^sub>Q \* C; distinct(bn \); bn \ \* Q; bn \ \* \; bn \ \* \\<^sub>Q; bn \ \* P; bn \ \* subject \; bn \ \* C\ \ Prop C \ (P \ Q) \ (P' \ Q)" and rPar2: "\\ \\<^sub>P Q \ Q' A\<^sub>P P C. \\ \ \\<^sub>P \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C (\ \ \\<^sub>P) Q \ Q'; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* Q'; A\<^sub>P \* C; distinct(bn \); bn \ \* Q; bn \ \* \; bn \ \* \\<^sub>P; bn \ \* P; bn \ \* subject \; bn \ \* C\ \ Prop C \ (P \ Q) \ (P \ Q')" and rComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C. \\ \ \\<^sub>Q \ P \M\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P (M\N\) P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q (K\\*xvec\\N\) Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; distinct xvec; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q'))" and rComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; \C. Prop C (\ \ \\<^sub>Q) P (M\\*xvec\\N\) P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \K\N\ \ Q'; \C. Prop C (\ \ \\<^sub>P) Q (K\N\) Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; distinct xvec; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q'))" and rOpen: "\\ P M xvec yvec N P' x C. \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; \C. Prop C \ P (M\\*(xvec@yvec)\\N\) P'; x \ \; x \ M; x \ xvec; x \ yvec; xvec \* \; xvec \* P; xvec \* M; distinct xvec; distinct yvec; yvec \* \; yvec \* P; yvec \* M; yvec \* C; x \ C; xvec \* C\ \ Prop C \ (\\x\P) (M\\*(xvec@x#yvec)\\N\) P'" and rScope: "\\ P \ P' x C. \\ \ P \\ \ P'; \C. Prop C \ P \ P'; x \ \; x \ \; bn \ \* \; bn \ \* P; bn \ \* (subject \); x \ C; bn \ \* C; distinct(bn \)\ \ Prop C \ (\\x\P) \ (\\x\P')" and rBang: "\\ P \ P' C. \\ \ P \ !P \\ \ P'; guarded P; \C. Prop C \ (P \ !P) \ P'\ \ Prop C \ (!P) \ P'" shows "Prop C \ P \ P'" using \\ \ P \\ \ P'\ \bn \ \* (subject \)\ \distinct(bn \)\ proof(nominal_induct x3=="\ \ P'" avoiding: \ C arbitrary: P' rule: semantics.strong_induct) case(cInput \ M K xvec N Tvec P \ C P') thus ?case by(force intro: rInput simp add: residualInject) next case(Output \ M K N P \ C P') thus ?case by(force intro: rOutput simp add: residualInject) next case(Case \ P Rs \ Cs \ C) thus ?case by(auto intro: rCase) next case(cPar1 \ \\<^sub>Q P \ P' Q A\<^sub>Q \' C P'') note \\ \ (P' \ Q) = \' \ P''\ moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto moreover note \distinct (bn \)\ \distinct(bn \')\ moreover from \bn \ \* subject \\ \bn \' \* subject \'\ have "bn \ \* (\ \ P' \ Q)" and "bn \' \* (\' \ P'')" by simp+ ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" and \Eq: "\' = p \ \" and P'eq: "P'' = p \ (P' \ Q)" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* (P' \ Q)" by(rule residualEq) note \\ \ \\<^sub>Q \ P \\ \ P'\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ moreover from \bn \ \* subject \\ \distinct(bn \)\ have "\C. Prop C (\ \ \\<^sub>Q) P \ P'" by(rule_tac cPar1) auto moreover note \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* C\ \bn \ \* Q\ \distinct(bn \)\ \bn \ \* \\ \bn \ \* \\<^sub>Q\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ ultimately have "Prop C \ (P \ Q) \ (P' \ Q)" by(rule_tac rPar1) with \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* bn \'\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (P' \ Q)\ \A\<^sub>Q \* C\ have "Prop C \ (P \ Q) (p \ \) (p \ (P' \ Q))" by(rule_tac rAlpha) auto with \Eq P'eq \distinctPerm p\ show ?case by simp next case(cPar2 \ \\<^sub>P Q \ Q' P A\<^sub>P \' C Q'') note \\ \ (P \ Q') = \' \ Q''\ moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto moreover note \distinct (bn \)\ \distinct(bn \')\ moreover from \bn \ \* subject \\ \bn \' \* subject \'\ have "bn \ \* (\ \ P \ Q')" and "bn \' \* (\' \ Q'')" by simp+ ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" and \Eq: "\' = p \ \" and Q'eq: "Q'' = p \ (P \ Q')" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* (P \ Q')" by(rule residualEq) note \\ \ \\<^sub>P \ Q \\ \ Q'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ moreover from \bn \ \* subject \\ \distinct(bn \)\ have "\C. Prop C (\ \ \\<^sub>P) Q \ Q'" by(rule_tac cPar2) auto moreover note \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* Q'\ \A\<^sub>P \* C\ \bn \ \* Q\ \distinct(bn \)\ \bn \ \* \\ \bn \ \* \\<^sub>P\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ ultimately have "Prop C \ (P \ Q) \ (P \ Q')" by(rule_tac rPar2) with \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* (bn \')\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (P \ Q')\ have "Prop C \ (P \ Q) (p \ \) (p \ (P \ Q'))" by(rule_tac rAlpha) auto with \Eq Q'eq \distinctPerm p\ show ?case by simp next case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q \ C P'') hence "Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q'))" by(rule_tac rComm1) (assumption | simp)+ thus ?case using \\ \ \\*xvec\(P' \ Q') = \ \ P''\ by(simp add: residualInject) next case(cComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q \ C P'') hence "Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q'))" by(rule_tac rComm2) (assumption | simp)+ thus ?case using \\ \ \\*xvec\(P' \ Q') = \ \ P''\ by(simp add: residualInject) next case(cOpen \ P M xvec yvec N P' x \ C P'') note \M\\*(xvec@x#yvec)\\N\ \ P' = \ \ P''\ moreover from \xvec \* \\ \x \ \\ \yvec \* \\ have "(xvec@x#yvec) \* (bn \)" by auto moreover from \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \distinct xvec\ \distinct yvec\ have "distinct(xvec@x#yvec)" by(auto simp add: fresh_star_def) (simp add: fresh_def name_list_supp) moreover note \distinct(bn \)\ moreover from \xvec \* M\ \x \ M\ \yvec \* M\ have "(xvec@x#yvec) \* M" by auto hence "(xvec@x#yvec) \* (M\\*(xvec@x#yvec)\\N\ \ P')" by auto moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P'')" by simp ultimately obtain p where S: "(set p) \ (set(xvec@x#yvec)) \ (set(p \ (xvec@x#yvec)))" and "distinctPerm p" and \eq: "\ = (p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\" and P'eq: "P'' = (p \ P')" and A: "(xvec@x#yvec) \* ((p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\)" and B: "(p \ (xvec@x#yvec)) \* (M\\*(xvec@x#yvec)\\N\)" and C: "(p \ (xvec@x#yvec)) \* P'" by(rule_tac residualEq) (assumption | simp)+ note \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ \x \ (supp N)\ moreover { fix C from \xvec \* M\ \yvec \* M\ have "(xvec@yvec) \* M" by simp moreover from \distinct xvec\ \distinct yvec\ \xvec \* yvec\ have "distinct(xvec@yvec)" by auto (simp add: fresh_star_def name_list_supp fresh_def) ultimately have "Prop C \ P (M\\*(xvec@yvec)\\N\) P'" by(rule_tac cOpen) auto } moreover note \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ \xvec \* \\ \xvec \* P\ \xvec \* M\ \yvec \* \\ \yvec \* P\ \yvec \* M\ \yvec \* C\ \x \ C\ \xvec \* C\ \distinct xvec\ \distinct yvec\ ultimately have "Prop C \ (\\x\P) (M\\*(xvec@x#yvec)\\N\) P'" by(rule_tac rOpen) with \xvec \* \\ \yvec \* \\ \xvec \* P\ \yvec \* P\ \xvec \* M\ \yvec \* M\ \yvec \* C\ S \distinctPerm p\ \x \ C\ \xvec \* C\ \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ A B C have "Prop C \ (\\x\P) (p \ (M\\*(xvec@x#yvec)\\N\)) (p \ P')" apply(rule_tac \="M\\*(xvec@x#yvec)\\N\" in rAlpha) apply(assumption | simp)+ apply(fastforce simp add: fresh_star_def abs_fresh) by(assumption | simp)+ with \eq P'eq show ?case by simp next case(cScope \ P \ P' x \' C P'') note \\ \ (\\x\P') = \' \ P''\ moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto moreover note \distinct (bn \)\ \distinct(bn \')\ moreover from \bn \ \* subject \\ \bn \' \* subject \'\ have "bn \ \* (\ \ \\x\P')" and "bn \' \* (\' \ P'')" by simp+ ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" and \Eq: "\' = p \ \" and P'eq: "P'' = p \ (\\x\P')" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* (\\x\P')" by(rule residualEq) note \\ \ P \\ \ P'\ moreover from \bn \ \* subject \\ \distinct(bn \)\ have "\C. Prop C \ P \ P'" by(rule_tac cScope) auto moreover note \x \ \\ \x \ \\ \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \x \ C\ \bn \ \* C\ \distinct(bn \)\ ultimately have "Prop C \ (\\x\P) \ (\\x\P')" by(rule rScope) with \bn \ \* \\ \bn \ \* P\ \x \ \\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* (bn \')\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (\\x\P')\ have "Prop C \ (\\x\P) (p \ \) (p \ (\\x\P'))" by(rule_tac rAlpha) simp+ with \Eq P'eq \distinctPerm p\ show ?case by simp next case(Bang \ P Rs \ C) thus ?case by(rule_tac rBang) auto qed lemma outputInduct[consumes 1, case_names cOutput cCase cPar1 cPar2 cOpen cScope cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and B :: "('a, 'b, 'c) boundOutput" and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a \ ('a, 'b, 'c) boundOutput \ bool" and C :: "'d::fs_name" assumes "\ \ P \ROut M B" and rOutput: "\\ M K N P C. \\ \ M \ K\ \ Prop C \ (M\N\.P) K (N \' P)" and rCase: "\\ P M B \ Cs C. \\ \ P \(ROut M B); \C. Prop C \ P M B; (\, P) mem Cs; \ \ \; guarded P\ \ Prop C \ (Cases Cs) M B" and rPar1: "\\ \\<^sub>Q P M xvec N P' A\<^sub>Q Q C. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>Q) P M (\\*xvec\N \' P'); A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* C; xvec \* Q; xvec \* \; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* C\ \ Prop C \ (P \ Q) M (\\*xvec\N \' (P' \ Q))" and rPar2: "\\ \\<^sub>P Q M xvec N Q' A\<^sub>P P C. \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C (\ \ \\<^sub>P) Q M (\\*xvec\N \' Q'); A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* C; xvec \* P; xvec \* \; xvec \* \\<^sub>P; xvec \* Q; xvec \* M; xvec \* C\ \ Prop C \ (P \ Q) M (\\*xvec\N \' (P \ Q'))" and rOpen: "\\ P M xvec yvec N P' x C. \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; \C. Prop C \ P M (\\*(xvec@yvec)\N \' P'); x \ \; x \ M; x \ xvec; x \ yvec; xvec \* \; xvec \* P; xvec \* M; xvec \* yvec; yvec \* \; yvec \* P; yvec \* M; yvec \* C; x \ C; xvec \* C\ \ Prop C \ (\\x\P) M (\\*(xvec@x#yvec)\N \' P')" and rScope: "\\ P M xvec N P' x C. \\ \ P \M\\*xvec\\N\ \ P'; \C. Prop C \ P M (\\*xvec\N \' P'); x \ \; x \ M; x \ xvec; x \ N; xvec \* \; xvec \* P; xvec \* M; x \ C; xvec \* C\ \ Prop C \ (\\x\P) M (\\*xvec\N \' \\x\P')" and rBang: "\\ P M B C. \\ \ P \ !P \(ROut M B); guarded P; \C. Prop C \ (P \ !P) M B\ \ Prop C \ (!P) M B" shows "Prop C \ P M B" using \\ \ P \(ROut M B)\ proof(nominal_induct \ P Rs=="(ROut M B)" avoiding: C arbitrary: B rule: semantics.strong_induct) case(cInput \ M K xvec N Tvec P C) thus ?case by(simp add: residualInject) next case(Output \ M K N P C) thus ?case by(force simp add: residualInject intro: rOutput) next case(Case \ P Rs \ Cs C) thus ?case by(force intro: rCase) next case(cPar1 \ \\<^sub>Q P \ P' Q A\<^sub>Q C) thus ?case by(force intro: rPar1 simp add: residualInject) next case(cPar2 \ \\<^sub>P Q \ Q' P A\<^sub>P C) thus ?case by(force intro: rPar2 simp add: residualInject) next case cComm1 thus ?case by(simp add: residualInject) next case cComm2 thus ?case by(simp add: residualInject) next case(cOpen \ P M xvec yvec N P' x C B) thus ?case by(force intro: rOpen simp add: residualInject) next case(cScope \ P M \ P' x C) thus ?case by(force intro: rScope simp add: residualInject) next case(Bang \ P Rs C) thus ?case by(force intro: rBang) qed lemma boundOutputBindObject: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and yvec :: "name list" and N :: 'a and P' :: "('a, 'b, 'c) psi" and y :: name assumes "\ \ P \\ \ P'" and "bn \ \* subject \" and "distinct(bn \)" and "y \ set(bn \)" shows "y \ supp(object \)" using assms proof(nominal_induct avoiding: P' arbitrary: y rule: semanticsInduct) case(cAlpha \ P \ P' p P'' y) from \y \ set(bn(p \ \))\ have "(p \ y) \ (p \ set(bn(p \ \)))" by(rule pt_set_bij2[OF pt_name_inst, OF at_name_inst]) hence "(p \ y) \ set(bn \)" using \distinctPerm p\ by(simp add: eqvts) hence "(p \ y) \ supp(object \)" by(rule cAlpha) hence "(p \ p \ y) \ (p \ supp(object \))" by(rule pt_set_bij2[OF pt_name_inst, OF at_name_inst]) thus ?case using \distinctPerm p\ by(simp add: eqvts) next case cInput thus ?case by(simp add: supp_list_nil) next case cOutput thus ?case by(simp add: supp_list_nil) next case cCase thus ?case by simp next case cPar1 thus ?case by simp next case cPar2 thus ?case by simp next case cComm1 thus ?case by(simp add: supp_list_nil) next case cComm2 thus ?case by(simp add: supp_list_nil) next case cOpen thus ?case by(auto simp add: supp_list_cons supp_list_append supp_atm supp_some) next case cScope thus ?case by simp next case cBang thus ?case by simp qed lemma alphaBoundOutputChain': fixes yvec :: "name list" and xvec :: "name list" and B :: "('a, 'b, 'c) boundOutput" assumes "length xvec = length yvec" and "yvec \* B" and "yvec \* xvec" and "distinct yvec" shows "\\*xvec\B = \\*yvec\([xvec yvec] \\<^sub>v B)" using assms proof(induct rule: composePermInduct) case cBase show ?case by simp next case(cStep x xvec y yvec) thus ?case apply auto by(subst alphaBoundOutput[of y]) (auto simp add: eqvts) qed lemma alphaBoundOutputChain'': fixes yvec :: "name list" and xvec :: "name list" and N :: 'a and P :: "('a, 'b, 'c) psi" assumes "length xvec = length yvec" and "yvec \* N" and "yvec \* P" and "yvec \* xvec" and "distinct yvec" shows "\\*xvec\(N \' P) = \\*yvec\(([xvec yvec] \\<^sub>v N) \' ([xvec yvec] \\<^sub>v P))" proof - from assms have "\\*xvec\(N \' P) = \\*yvec\([xvec yvec] \\<^sub>v (N \' P))" by(simp add: alphaBoundOutputChain') thus ?thesis by simp qed lemma alphaDistinct: fixes xvec :: "name list" and N :: 'a and P :: "('a, 'b, 'c) psi" and yvec :: "name list" and M :: 'a and Q :: "('a, 'b, 'c) psi" assumes "\ \ P = \ \ Q" and "distinct(bn \)" and "\x. x \ set(bn \) \ x \ supp(object \)" and "bn \ \* bn \" and "bn \ \* (object \)" and "bn \ \* Q" shows "distinct(bn \)" using assms proof(rule_tac actionCases[where \=\], auto simp add: residualInject supp_some) fix xvec M yvec N assume Eq: "\\*xvec\N \' P = \\*yvec\M \' Q" assume "distinct xvec" and "xvec \* M" and "xvec \* yvec" and "xvec \* Q" assume Mem: "\x. x \ set xvec \ x \ (supp N)" show "distinct yvec" proof - from Eq have "length xvec = length yvec" by(rule boundOutputChainEqLength) with Eq \distinct xvec\ \xvec \* yvec\ \xvec \* M\ \xvec \* Q\ Mem show ?thesis proof(induct n=="length xvec" arbitrary: xvec yvec M Q rule: nat.induct) case(zero xvec yvec M Q) thus ?case by simp next case(Suc n xvec yvec M Q) have L: "length xvec = length yvec" and "Suc n = length xvec" by fact+ then obtain x xvec' y yvec' where xEq: "xvec = x#xvec'" and yEq: "yvec = y#yvec'" and L': "length xvec' = length yvec'" by(cases xvec, auto, cases yvec, auto) have xvecFreshyvec: "xvec \* yvec" and xvecDist: "distinct xvec" by fact+ with xEq yEq have xineqy: "x \ y" and xvec'Freshyvec': "xvec' \* yvec'" and xvec'Dist: "distinct xvec'" and xFreshxvec': "x \ xvec'" and xFreshyvec': "x \ yvec'" and yFreshxvec': "y \ xvec'" by auto have Eq: "\\*xvec\N \' P = \\*yvec\M \' Q" by fact with xEq yEq xineqy have Eq': "\\*xvec'\N \' P = \\*([(x, y)] \ yvec')\([(x, y)] \ M) \' ([(x, y)] \ Q)" by(simp add: boundOutput.inject alpha eqvts) moreover have Mem:"\x. x \ set xvec \ x \ supp N" by fact with xEq have "\x. x \ set xvec' \ x \ supp N" by simp moreover have "xvec \* M" by fact with xEq xFreshxvec' yFreshxvec' have "xvec' \* ([(x, y)] \ M)" by simp moreover have xvecFreshQ: "xvec \* Q" by fact with xEq xFreshxvec' yFreshxvec' have "xvec' \* ([(x, y)] \ Q)" by simp moreover have "Suc n = length xvec" by fact with xEq have "n = length xvec'" by simp moreover from xvec'Freshyvec' xFreshxvec' yFreshxvec' have "xvec' \* ([(x, y)] \ yvec')" by simp moreover from L' have "length xvec' = length([(x, y)] \ yvec')" by simp ultimately have "distinct([(x, y)] \ yvec')" using xvec'Dist by(rule_tac Suc) (assumption | simp)+ hence "distinct yvec'" by simp from Mem xEq have xSuppN: "x \ supp N" by simp from L \distinct xvec\ \xvec \* yvec\ \xvec \* M\ \xvec \* Q\ have "\\*yvec\M \' Q = \\*xvec\([yvec xvec] \\<^sub>v M) \' ([yvec xvec] \\<^sub>v Q)" by(simp add: alphaBoundOutputChain'') with Eq have "N = [yvec xvec] \\<^sub>v M" by simp with xEq yEq have "N = [(y, x)] \ [yvec' xvec'] \\<^sub>v M" by simp with xSuppN have ySuppM: "y \ supp([yvec' xvec'] \\<^sub>v M)" by(drule_tac pi="[(x, y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) (simp add: calc_atm eqvts name_swap) have "y \ yvec'" proof(simp add: fresh_def, rule notI) assume "y \ supp yvec'" hence "y mem yvec'" by(induct yvec') (auto simp add: supp_list_nil supp_list_cons supp_atm) moreover from \xvec \* M\ xEq xFreshxvec' have "xvec' \* M" by simp ultimately have "y \ [yvec' xvec'] \\<^sub>v M" using L' xvec'Freshyvec' xvec'Dist by(force intro: freshChainPerm) with ySuppM show "False" by(simp add: fresh_def) qed with \distinct yvec'\ yEq show ?case by simp qed qed qed lemma boundOutputDistinct: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" assumes "\ \ P \\ \ P'" shows "distinct(bn \)" using assms proof(nominal_induct \ P x3=="\ \ P'" avoiding: \ P' rule: semantics.strong_induct) case cInput thus ?case by(simp add: residualInject) next case Output thus ?case by(simp add: residualInject) next case Case thus ?case by(simp add: residualInject) next case cPar1 thus ?case by(force intro: alphaDistinct boundOutputBindObject) next case cPar2 thus ?case by(force intro: alphaDistinct boundOutputBindObject) next case cComm1 thus ?case by(simp add: residualInject) next case cComm2 thus ?case by(simp add: residualInject) next case(cOpen \ P M xvec yvec N P' x \ P'') note \M\\*(xvec@x#yvec)\\N\ \ P' = \ \ P''\ moreover from \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \distinct xvec\ \distinct yvec\ have "distinct(bn(M\\*(xvec@x#yvec)\\N\))" by auto (simp add: fresh_star_def fresh_def name_list_supp) moreover { fix y from \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ \x \ supp N\ \x \ xvec\ \x \ yvec\ \x \ M\ \x \ \\ \distinct xvec\ \distinct yvec\ \xvec \* \\ \xvec \* P\ \xvec \* M\ \xvec \* yvec\ \yvec \* \\ \yvec \* P\ \yvec \* M\ have "\ \ \\x\P \M\\*(xvec@x#yvec)\\N\ \ P'" by(rule semantics.cOpen) moreover moreover from \xvec \* M\ \x \ M\ \yvec \* M\ have "bn(M\\*(xvec@x#yvec)\\N\) \* (subject(M\\*(xvec@x#yvec)\\N\))" by simp moreover note \distinct(bn(M\\*(xvec@x#yvec)\\N\))\ moreover assume "y \ set(bn(M\\*(xvec@x#yvec)\\N\))" ultimately have "y \ supp(object(M\\*(xvec@x#yvec)\\N\))" by(rule_tac boundOutputBindObject) } moreover from \xvec \* \\ \x \ \\ \yvec \* \\ have "bn(M\\*(xvec@x#yvec)\\N\) \* bn \" and "bn(M\\*(xvec@x#yvec)\\N\) \* object \" by simp+ moreover from \xvec \* P''\ \x \ P''\ \yvec \* P''\ have "bn(M\\*(xvec@x#yvec)\\N\) \* P''" by simp ultimately show ?case by(rule alphaDistinct) next case cScope thus ?case by(rule_tac alphaDistinct, auto) (rule_tac boundOutputBindObject, auto) next case Bang thus ?case by simp qed lemma inputDistinct: fixes \ :: 'b and M :: 'a and xvec :: "name list" and N :: 'a and P :: "('a, 'b, 'c) psi" and Rs :: "('a, 'b, 'c) residual" assumes "\ \ M\\*xvec N\.P \ Rs" shows "distinct xvec" using assms by(nominal_induct \ P=="M\\*xvec N\.P" Rs avoiding: xvec N P rule: semantics.strong_induct) (auto simp add: psi.inject intro: alphaInputDistinct) lemma outputInduct'[consumes 2, case_names cAlpha cOutput cCase cPar1 cPar2 cOpen cScope cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and yvec :: "name list" and N :: 'a and P' :: "('a, 'b, 'c) psi" and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a \ name list \ 'a \ ('a, 'b, 'c) psi \ bool" and C :: "'d::fs_name" assumes "\ \ P \M\\*xvec\\N\ \ P'" and "xvec \* M" and rAlpha: "\\ P M xvec N P' p C. \xvec \* \; xvec \* P; xvec \* M; xvec \* C; xvec \* (p \ xvec); set p \ set xvec \ set(p \ xvec); distinctPerm p; (p \ xvec) \* N; (p \ xvec) \* P'; Prop C \ P M xvec N P'\ \ Prop C \ P M (p \ xvec) (p \ N) (p \ P')" and rOutput: "\\ M K N P C. \\ \ M \ K\ \ Prop C \ (M\N\.P) K ([]) N P" and rCase: "\\ P M xvec N P' \ Cs C. \\ \ P \M\\*xvec\\N\ \ P'; \C. Prop C \ P M xvec N P'; (\, P) mem Cs; \ \ \; guarded P\ \ Prop C \ (Cases Cs) M xvec N P'" and rPar1: "\\ \\<^sub>Q P M xvec N P' A\<^sub>Q Q C. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>Q) P M xvec N P'; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* C; xvec \* Q; xvec \* \; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* C\ \ Prop C \ (P \ Q) M xvec N (P' \ Q)" and rPar2: "\\ \\<^sub>P Q M xvec N Q' A\<^sub>P P C. \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C (\ \ \\<^sub>P) Q M xvec N Q'; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* C; xvec \* Q; xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* M; xvec \* C\ \ Prop C \ (P \ Q) M xvec N (P \ Q')" and rOpen: "\\ P M xvec yvec N P' x C. \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; \C. Prop C \ P M (xvec@yvec) N P'; x \ \; x \ M; x \ xvec; x \ yvec; xvec \* \; xvec \* P; xvec \* M; yvec \* \; yvec \* P; yvec \* M; yvec \* C; x \ C; xvec \* C\ \ Prop C \ (\\x\P) M (xvec@x#yvec) N P'" and rScope: "\\ P M xvec N P' x C. \\ \ P \M\\*xvec\\N\ \ P'; \C. Prop C \ P M xvec N P'; x \ \; x \ M; x \ xvec; x \ N; xvec \* \; xvec \* P; xvec \* M; x \ C; xvec \* C\ \ Prop C \ (\\x\P) M xvec N (\\x\P')" and rBang: "\\ P M xvec N P' C. \\ \ P \ !P \M\\*xvec\\N\ \ P'; guarded P; \C. Prop C \ (P \ !P) M xvec N P'\ \ Prop C \ (!P) M xvec N P'" shows "Prop C \ P M xvec N P'" proof - note \\ \ P \M\\*xvec\\N\ \ P'\ moreover from \xvec \* M\ have "bn(M\\*xvec\\N\) \* subject(M\\*xvec\\N\)" by simp moreover from \\ \ P \M\\*xvec\\N\ \ P'\ have "distinct(bn(M\\*xvec\\N\))" by(rule boundOutputDistinct) ultimately show ?thesis proof(nominal_induct \ P \=="M\\*xvec\\N\" P' avoiding: C arbitrary: M xvec N rule: semanticsInduct) case(cAlpha \ P \ P' p C M xvec N) from \(p \ \) = M\\*xvec\\N\\ have "(p \ p \ \) = p \ (M\\*xvec\\N\)" by(simp add: fresh_bij) with \distinctPerm p\ have A: "\ = (p \ M)\\*(p \ xvec)\\(p \ N)\" by(simp add: eqvts) with \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \ \ \bn \ \* C\ \bn \ \* bn(p \ \)\ \distinctPerm p\ have "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* (p \ M)" and "(p \ xvec) \* C" and "(p \ xvec) \* (p \ p \ xvec)" by auto moreover from A \set p \ set(bn \) \ set(bn(p \ \))\ \distinctPerm p\ have S: "set p \ set(p \ xvec) \ set(p \ p \ xvec)" by simp moreover note \distinctPerm p\ moreover from A \bn(p \ \) \* \\ \bn(p \ \) \* P'\ have "(p \ p \ xvec) \* (p \ N)" and "(p \ p \ xvec) \* P'" by simp+ moreover from A have "Prop C \ P (p \ M) (p \ xvec) (p \ N) P'" by(rule cAlpha) ultimately have "Prop C \ P (p \ M) (p \ p \ xvec) (p \ p \ N) (p \ P')" by(rule rAlpha) moreover from A \bn \ \* subject \\ have "(p \ xvec) \* (p \ M)" by simp hence "xvec \* M" by(simp add: fresh_star_bij) from A \bn(p \ \) \* \\ \distinctPerm p\ have "xvec \* (p \ M)" by simp hence "(p \ xvec) \* (p \ p \ M)" by(simp add: fresh_star_bij) with \distinctPerm p\ have "(p \ xvec) \* M" by simp with \xvec \* M\ S \distinctPerm p\ have "(p \ M) = M" by simp ultimately show ?case using S \distinctPerm p\ by simp next case cInput thus ?case by(simp add: residualInject) next case cOutput thus ?case by(force dest: rOutput simp add: action.inject) next case cCase thus ?case by(force intro: rCase) next case cPar1 thus ?case by(force intro: rPar1) next case cPar2 thus ?case by(force intro: rPar2) next case cComm1 thus ?case by(simp add: action.inject) next case cComm2 thus ?case by(simp add: action.inject) next case cOpen thus ?case by(fastforce intro: rOpen simp add: action.inject) next case cScope thus ?case by(fastforce intro: rScope) next case cBang thus ?case by(fastforce intro: rBang) qed qed lemma inputInduct[consumes 1, case_names cInput cCase cPar1 cPar2 cScope cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a \ 'a \ ('a, 'b, 'c) psi \ bool" and C :: "'d::fs_name" assumes Trans: "\ \ P \M\N\ \ P'" and rInput: "\\ M K xvec N Tvec P C. \\ \ M \ K; distinct xvec; set xvec \ supp N; length xvec = length Tvec; xvec \* \; xvec \* M; xvec \* K; xvec \* C\ \ Prop C \ (M\\*xvec N\.P) K (N[xvec::=Tvec]) (P[xvec::=Tvec])" and rCase: "\\ P M N P' \ Cs C. \\ \ P \M\N\ \ P'; \C. Prop C \ P M N P'; (\, P) mem Cs; \ \ \; guarded P\ \ Prop C \ (Cases Cs) M N P'" and rPar1: "\\ \\<^sub>Q P M N P' A\<^sub>Q Q C. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>Q) P M N P'; distinct A\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* C\ \ Prop C \ (P \ Q) M N (P' \ Q)" and rPar2: "\\ \\<^sub>P Q M N Q' A\<^sub>P P C. \\ \ \\<^sub>P \ Q \M\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C (\ \ \\<^sub>P) Q M N Q'; distinct A\<^sub>P; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* C\ \ Prop C \ (P \ Q) M N (P \ Q')" and rScope: "\\ P M N P' x C. \\ \ P \M\N\ \ P'; \C. Prop C \ P M N P'; x \ \; x \ M; x \ N; x \ C\ \ Prop C \ (\\x\P) M N (\\x\P')" and rBang: "\\ P M N P' C. \\ \ P \ !P \M\N\ \ P'; guarded P; \C. Prop C \ (P \ !P) M N P'\ \ Prop C \ (!P) M N P'" shows "Prop C \ P M N P'" using Trans proof(nominal_induct \ P Rs=="M\N\ \ P'" avoiding: C arbitrary: P' rule: semantics.strong_induct) case(cInput \ M K xvec N Tvec P C) thus ?case by(force intro: rInput simp add: residualInject action.inject) next case(Output \ M K N P C) thus ?case by(simp add: residualInject) next case(Case \ P Rs \ CS C) thus ?case by(force intro: rCase) next case(cPar1 \ \\<^sub>Q P \ P' Q A\<^sub>Q C P'') thus ?case by(force intro: rPar1 simp add: residualInject) next case(cPar2 \ \P Q \ Q' xvec P C Q'') thus ?case by(force intro: rPar2 simp add: residualInject) next case(cComm1 \ \Q P M N P' xvec \P Q K zvec Q' yvec C PQ) thus ?case by(simp add: residualInject) next case(cComm2 \ \Q P M zvec N P' xvec \P Q K yvec Q' C PQ) thus ?case by(simp add: residualInject) next case(cOpen \ P M xvec N P' x yvec C P'') thus ?case by(simp add: residualInject) next case(cScope \ P \ P' x C P'') thus ?case by(force intro: rScope simp add: residualInject) next case(Bang \ P Rs C) thus ?case by(force intro: rBang) qed lemma tauInduct[consumes 1, case_names cCase cPar1 cPar2 cComm1 cComm2 cScope cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Rs :: "('a, 'b, 'c) residual" and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi \ bool" and C :: "'d::fs_name" assumes Trans: "\ \ P \\ \ P'" and rCase: "\\ P P' \ Cs C. \\ \ P \\ \ P'; \C. Prop C \ P P'; (\, P) mem Cs; \ \ \; guarded P\ \ Prop C \ (Cases Cs) P'" and rPar1: "\\ \\<^sub>Q P P' A\<^sub>Q Q C. \\ \ \\<^sub>Q \ P \\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>Q) P P'; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* P'; A\<^sub>Q \* C\ \ Prop C \ (P \ Q) (P' \ Q)" and rPar2: "\\ \\<^sub>P Q Q' A\<^sub>P P C. \\ \ \\<^sub>P \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C (\ \ \\<^sub>P) Q Q'; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* Q'; A\<^sub>P \* C\ \ Prop C \ (P \ Q) (P \ Q')" and rComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) (\\*xvec\(P' \ Q'))" and rComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) (\\*xvec\(P' \ Q'))" and rScope: "\\ P P' x C. \\ \ P \\ \ P'; \C. Prop C \ P P'; x \ \; x \ C\ \ Prop C \ (\\x\P) (\\x\P')" and rBang: "\\ P P' C. \\ \ P \ !P \\ \ P'; guarded P; \C. Prop C \ (P \ !P) P'\ \ Prop C \ (!P) P'" shows "Prop C \ P P'" using Trans proof(nominal_induct \ P Rs=="\ \ P'" avoiding: C arbitrary: P' rule: semantics.strong_induct) case(cInput M K xvec N Tvec P C) thus ?case by(simp add: residualInject) next case(Output \ M K N P C) thus ?case by(simp add: residualInject) next case(Case \ P Rs \ Cs C) thus ?case by(force intro: rCase simp add: residualInject) next case(cPar1 \ \\<^sub>Q P \ P' A\<^sub>Q Q C P'') thus ?case by(force intro: rPar1 simp add: residualInject) next case(cPar2 \ \\<^sub>P Q \ Q' A\<^sub>P P C Q'') thus ?case by(force intro: rPar2 simp add: residualInject) next case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C PQ) thus ?case by(force intro: rComm1 simp add: residualInject) next case(cComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \P Q' A\<^sub>Q C PQ) thus ?case by(force intro: rComm2 simp add: residualInject) next case(cOpen \ P M xvec N P' x yvec C P'') thus ?case by(simp add: residualInject) next case(cScope \ P \ P' x C P'') thus ?case by(force intro: rScope simp add: residualInject) next case(Bang \ P Rs C ) thus ?case by(force intro: rBang simp add: residualInject) qed lemma semanticsFrameInduct[consumes 3, case_names cAlpha cInput cOutput cCase cPar1 cPar2 cComm1 cComm2 cOpen cScope cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Rs :: "('a, 'b, 'c) residual" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) residual \ name list \ 'b \ bool" and C :: "'d::fs_name" assumes Trans: "\ \ P \ Rs" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and rAlpha: "\\ P A\<^sub>P \\<^sub>P p Rs C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* Rs; A\<^sub>P \* C; set p \ set A\<^sub>P \ set(p \ A\<^sub>P); distinctPerm p; Prop C \ P Rs A\<^sub>P \\<^sub>P\ \ Prop C \ P Rs (p \ A\<^sub>P) (p \ \\<^sub>P)" and rInput: "\\ M K xvec N Tvec P C. \\ \ M \ K; distinct xvec; set xvec \ supp N; length xvec = length Tvec; xvec \* \; xvec \* M; xvec \* K; xvec \* C\ \ Prop C \ (M\\*xvec N\.P) (K\(N[xvec::=Tvec])\ \ (P[xvec::=Tvec])) ([]) (\)" and rOutput: "\\ M K N P C. \ \ M \ K \ Prop C \ (M\N\.P) (K\N\ \ P) ([]) (\)" and rCase: "\\ P Rs \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \ Rs; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P Rs A\<^sub>P \\<^sub>P; (\, P) mem Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Rs; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) Rs ([]) (\)" and rPar1: "\\ \\<^sub>Q P \ P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. \\ \ \\<^sub>Q \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>Q) P (\ \ P') A\<^sub>P \\<^sub>P; distinct(bn \); A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; bn \ \* \; bn \ \* P; bn \ \* Q; bn \ \* subject \; bn \ \* \\<^sub>P; bn \ \* \\<^sub>Q; A\<^sub>P \* C; A\<^sub>Q \* C; bn \ \* C\ \ Prop C \ (P \ Q) (\ \ (P' \ Q)) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rPar2: "\\ \\<^sub>P Q \ Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. \\ \ \\<^sub>P \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>P) Q (\ \ Q') A\<^sub>Q \\<^sub>Q; distinct(bn \); A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; bn \ \* \; bn \ \* P; bn \ \* Q; bn \ \* subject \; bn \ \* \\<^sub>P; bn \ \* \\<^sub>Q; A\<^sub>P \* C; A\<^sub>Q \* C; bn \ \* C\ \ Prop C \ (P \ Q) (\ \ (P \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C (\ \ \\<^sub>Q) P ((M\N\) \ P') A\<^sub>P \\<^sub>P; \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; \C. Prop C (\ \ \\<^sub>P) Q (K\\*xvec\\N\ \ Q') A\<^sub>Q \\<^sub>Q; distinct xvec; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C (\ \ \\<^sub>Q) P (M\\*xvec\\N\ \ P') A\<^sub>P \\<^sub>P; \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>P) Q (K\N\ \ Q') A\<^sub>Q \\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rOpen: "\\ P M xvec yvec N P' x A\<^sub>P \\<^sub>P C. \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P (M\\*(xvec@yvec)\\N\ \ P') A\<^sub>P \\<^sub>P; x \ supp N; x \ \; x \ M; x \ A\<^sub>P; x \ xvec; x \ yvec; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* xvec; A\<^sub>P \* yvec; xvec \* yvec; distinct xvec; distinct yvec; xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; yvec \* \\<^sub>P; yvec \* \; yvec \* P; yvec \* M; A\<^sub>P \* C; x \ C; xvec \* C; yvec \* C\ \ Prop C \ (\\x\P) (M\\*(xvec@x#yvec)\\N\ \ P') (x#A\<^sub>P) \\<^sub>P" and rScope: "\\ P \ P' x A\<^sub>P \\<^sub>P C. \\ \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P (\ \ P') A\<^sub>P \\<^sub>P; x \ \; x \ \; x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* \; A\<^sub>P \* P'; distinct(bn \); bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* \\<^sub>P; A\<^sub>P \* C; x \ C; bn \ \* C\ \ Prop C \ (\\x\P) (\ \ (\\x\P')) (x#A\<^sub>P) \\<^sub>P" and rBang: "\\ P Rs A\<^sub>P \\<^sub>P C. \\ \ P \ !P \ Rs; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ (P \ !P) Rs A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; supp \\<^sub>P = ({}::name set); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Rs; A\<^sub>P \* C\ \ Prop C \ (!P) Rs ([]) (\)" shows "Prop C \ P Rs A\<^sub>P \\<^sub>P" using Trans FrP \distinct A\<^sub>P\ proof(nominal_induct avoiding: A\<^sub>P \\<^sub>P C rule: semantics.strong_induct) case(cInput \ M K xvec N Tvec P A\<^sub>P \\<^sub>P C) from \extractFrame (M\\*xvec N\.P) = \A\<^sub>P, \\<^sub>P\\ have "A\<^sub>P = []" and "\\<^sub>P = \" by auto with \\ \ M \ K\ \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ \xvec \* \\ \xvec \* M\ \xvec \* K\ \xvec \* C\ show ?case by(blast intro: rInput) next case(Output \ M K N P A\<^sub>P \\<^sub>P) from \extractFrame (M\N\.P) = \A\<^sub>P, \\<^sub>P\\ have "A\<^sub>P = []" and "\\<^sub>P = \" by auto with \\ \ M \ K\ show ?case by(blast intro: rOutput) next case(Case \ P Rs \ Cs A\<^sub>c\<^sub>P \\<^sub>c\<^sub>P C) obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "A\<^sub>P \* (\, P, Rs, C)" by(rule freshFrame) hence "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* Rs" and "A\<^sub>P \* C" by simp+ note \\ \ P \ Rs\ FrP \distinct A\<^sub>P\ moreover from FrP \distinct A\<^sub>P\ \\A\<^sub>P \\<^sub>P C. \extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P\ \ Prop C \ P Rs A\<^sub>P \\<^sub>P\ have "\C. Prop C \ P Rs A\<^sub>P \\<^sub>P" by simp moreover note \(\, P) mem Cs\ \\ \ \\ \guarded P\ moreover from \guarded P\ FrP have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" by(metis guardedStatEq)+ moreover note \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Rs\ \A\<^sub>P \* C\ ultimately have "Prop C \ (Cases Cs) Rs ([]) (\)" by(rule rCase) thus ?case using \extractFrame(Cases Cs) = \A\<^sub>c\<^sub>P, \\<^sub>c\<^sub>P\\ by simp next case(cPar1 \ \\<^sub>Q P \ P' Q A\<^sub>Q A\<^sub>P\<^sub>Q \\<^sub>P\<^sub>Q C) obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" "A\<^sub>P \* (P, Q, \, \, P', A\<^sub>Q, A\<^sub>P\<^sub>Q, C, \\<^sub>Q)" by(rule freshFrame) hence "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* \" and "A\<^sub>P \* \" and "A\<^sub>P \* P'" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* C" and "A\<^sub>P \* \\<^sub>Q" by simp+ have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact from \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>Q\ FrP have "A\<^sub>Q \* \\<^sub>P" by(force dest: extractFrameFreshChain) from \bn \ \* P\ \A\<^sub>P \* \\ FrP have "bn \ \* \\<^sub>P" by(force dest: extractFrameFreshChain) from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" by(auto simp add: fresh_star_def fresh_def name_list_supp) ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" and \eq: "\\<^sub>P\<^sub>Q = p \ (\\<^sub>P \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ by(rule_tac frameChainEq') (assumption | simp add: eqvts)+ note \\ \ \\<^sub>Q \ P \\ \ P'\ FrP \distinct A\<^sub>P\ FrQ \distinct A\<^sub>Q\ moreover from FrP \distinct A\<^sub>P\ \\A\<^sub>P \\<^sub>P C. \extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P\ \ Prop C (\ \ \\<^sub>Q) P (\ \ P') A\<^sub>P \\<^sub>P\ have "\C. Prop C (\ \ \\<^sub>Q) P (\ \ P') A\<^sub>P \\<^sub>P" by simp moreover note \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* \\<^sub>P\ \distinct(bn \)\ \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* \\<^sub>P\ \bn \ \* \\<^sub>Q\ \A\<^sub>P \* C\ \A\<^sub>Q \* C\ \bn \ \* C\ ultimately have "Prop C \ (P \ Q) (\ \ (P' \ Q)) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" by(rule_tac rPar1) with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>P \* C\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* C\ S \distinctPerm p\ Aeq have "Prop C \ (P \ Q) (\ \ (P' \ Q)) (p \ (A\<^sub>P@A\<^sub>Q)) (p \ (\\<^sub>P \ \\<^sub>Q))" by(rule_tac rAlpha) (assumption | simp add: eqvts)+ with \eq Aeq show ?case by(simp add: eqvts) next case(cPar2 \ \\<^sub>P Q \ Q' P A\<^sub>P A\<^sub>P\<^sub>Q \\<^sub>P\<^sub>Q C) obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>Q" "A\<^sub>Q \* (P, Q, \, \, Q', A\<^sub>P, A\<^sub>P\<^sub>Q, C, \\<^sub>P)" by(rule freshFrame) hence "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* \" and "A\<^sub>Q \* Q'" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* C" and "A\<^sub>Q \* \\<^sub>P" by simp+ from \A\<^sub>Q \* A\<^sub>P\ have "A\<^sub>P \* A\<^sub>Q" by simp have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact from \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ FrQ have "A\<^sub>P \* \\<^sub>Q" by(force dest: extractFrameFreshChain) from \bn \ \* Q\ \A\<^sub>Q \* \\ FrQ have "bn \ \* \\<^sub>Q" by(force dest: extractFrameFreshChain) from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" by(auto simp add: fresh_star_def fresh_def name_list_supp) ultimately obtain p where S: "(set p \ (set(A\<^sub>P@A\<^sub>Q)) \ (set A\<^sub>P\<^sub>Q))" and "distinctPerm p" and \eq: "\\<^sub>P\<^sub>Q = p \ (\\<^sub>P \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = ((p \ A\<^sub>P)@(p \ A\<^sub>Q))" using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ by(rule_tac frameChainEq') (assumption | simp add: eqvts)+ note \\ \ \\<^sub>P \ Q \\ \ Q'\ FrP \distinct A\<^sub>P\ FrQ \distinct A\<^sub>Q\ moreover from FrQ \distinct A\<^sub>Q\ \\A\<^sub>Q \\<^sub>Q C. \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q\ \ Prop C (\ \ \\<^sub>P) Q (\ \ Q') A\<^sub>Q \\<^sub>Q\ have "\C. Prop C (\ \ \\<^sub>P) Q (\ \ Q') A\<^sub>Q \\<^sub>Q" by simp moreover note \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* Q'\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* \\<^sub>P\ \distinct(bn \)\ \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* \\<^sub>P\ \bn \ \* \\<^sub>Q\ \A\<^sub>P \* C\ \A\<^sub>Q \* C\ \bn \ \* C\ ultimately have "Prop C \ (P \ Q) (\ \ (P \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" by(rule_tac rPar2) with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* Q'\ \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>P \* C\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* C\ S \distinctPerm p\ Aeq have "Prop C \ (P \ Q) (\ \ (P \ Q')) (p \ (A\<^sub>P@A\<^sub>Q)) (p \ (\\<^sub>P \ \\<^sub>Q))" by(rule_tac rAlpha) (assumption | simp add: eqvts)+ with \eq Aeq show ?case by(simp add: eqvts) next case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q A\<^sub>P\<^sub>Q \\<^sub>P\<^sub>Q C) from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" by(auto simp add: fresh_star_def fresh_def name_list_supp) from cComm1 have "Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" by(rule_tac rComm1) moreover from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ have "\(A\<^sub>P@A\<^sub>Q), (\\<^sub>P \ \\<^sub>Q)\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp with \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct(A\<^sub>P@A\<^sub>Q)\ \distinct A\<^sub>P\<^sub>Q\ obtain p where S: "(set p \ (set(A\<^sub>P@A\<^sub>Q)) \ (set A\<^sub>P\<^sub>Q))" and "distinctPerm p" and \eq: "\\<^sub>P\<^sub>Q = p \ (\\<^sub>P \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = p \ (A\<^sub>P@A\<^sub>Q)" by(rule_tac frameChainEq') (assumption | simp)+ moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* xvec\ \A\<^sub>Q \* xvec\ \A\<^sub>P \* P'\ \A\<^sub>Q \* P'\ \A\<^sub>P \* Q'\ \A\<^sub>Q \* Q'\ \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \A\<^sub>P \* C\ \A\<^sub>Q \* C\ ultimately have "Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (p \ (A\<^sub>P@A\<^sub>Q)) (p \ (\\<^sub>P \ \\<^sub>Q))" by(rule_tac rAlpha) auto with \eq Aeq show ?case by simp next case(cComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q A\<^sub>P\<^sub>Q \\<^sub>P\<^sub>Q C) from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" by(auto simp add: fresh_star_def fresh_def name_list_supp) from cComm2 have "Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" by(rule_tac rComm2) moreover from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ have "\(A\<^sub>P@A\<^sub>Q), (\\<^sub>P \ \\<^sub>Q)\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp with \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct(A\<^sub>P@A\<^sub>Q)\ \distinct A\<^sub>P\<^sub>Q\ obtain p where S: "(set p \ (set(A\<^sub>P@A\<^sub>Q)) \ (set A\<^sub>P\<^sub>Q))" and "distinctPerm p" and \eq: "\\<^sub>P\<^sub>Q = p \ (\\<^sub>P \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = p \ (A\<^sub>P@A\<^sub>Q)" by(rule_tac frameChainEq') (assumption | simp)+ moreover note \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* xvec\ \A\<^sub>Q \* xvec\ \A\<^sub>P \* P'\ \A\<^sub>Q \* P'\ \A\<^sub>P \* Q'\ \A\<^sub>Q \* Q'\ \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \A\<^sub>P \* C\ \A\<^sub>Q \* C\ ultimately have "Prop C \ (P \ Q) (\ \ \\*xvec\(P' \ Q')) (p \ (A\<^sub>P@A\<^sub>Q)) (p \ (\\<^sub>P \ \\<^sub>Q))" by(rule_tac rAlpha) auto with \eq Aeq show ?case by simp next case(cOpen \ P M xvec yvec N P' x A\<^sub>x\<^sub>P \\<^sub>x\<^sub>P C) obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "A\<^sub>P \* (\, P, M, xvec, yvec, N, P', A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P, C, x)" by(rule freshFrame) hence "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* M" and "A\<^sub>P \* xvec"and "A\<^sub>P \* yvec" and "A\<^sub>P \* N" and "A\<^sub>P \* P'" and "A\<^sub>P \* A\<^sub>x\<^sub>P" and "A\<^sub>P \* \\<^sub>x\<^sub>P" and "A\<^sub>P \* C" and "x \ A\<^sub>P" by simp+ from \xvec \* P\ \A\<^sub>P \* xvec\ FrP have "xvec \* \\<^sub>P" by(force dest: extractFrameFreshChain) from \yvec \* P\ \A\<^sub>P \* yvec\ FrP have "yvec \* \\<^sub>P" by(force dest: extractFrameFreshChain) from \extractFrame(\\x\P) = \A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P\\ FrP have "\(x#A\<^sub>P), \\<^sub>P\ = \A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P\" by simp moreover from \x \ A\<^sub>P\ \distinct A\<^sub>P\ have "distinct(x#A\<^sub>P)" by simp ultimately obtain p where S: "set p \ set (x#A\<^sub>P) \ set (p \ (x#A\<^sub>P))" and "distinctPerm p" and \eq: "\\<^sub>x\<^sub>P = p \ \\<^sub>P" and Aeq: "A\<^sub>x\<^sub>P = (p \ x)#(p \ A\<^sub>P)" using \A\<^sub>P \* A\<^sub>x\<^sub>P\\x \ A\<^sub>x\<^sub>P\ \distinct A\<^sub>x\<^sub>P\ by(rule_tac frameChainEq') (assumption | simp add: eqvts)+ note \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ FrP \distinct A\<^sub>P\ moreover from FrP \distinct A\<^sub>P\ \\A\<^sub>P \\<^sub>P C. \extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P\ \ Prop C \ P (M\\*(xvec@yvec)\\N\ \ P') A\<^sub>P \\<^sub>P\ have "\C. Prop C \ P (M\\*(xvec@yvec)\\N\ \ P') A\<^sub>P \\<^sub>P" by simp moreover note \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ \x \ supp N\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \xvec \* \\ \xvec \* P\ \xvec \* M\ \xvec \* \\<^sub>P\ \yvec \* \\ \yvec \* P\ \yvec \* M\ \yvec \* \\<^sub>P\ \A\<^sub>P \* C\ \x \ C\ \xvec \* C\ \yvec \* C\ \xvec \* yvec\ \distinct xvec\ \distinct yvec\ ultimately have "Prop C \ (\\x\P) (M\\*(xvec@x#yvec)\\N\ \ P') (x#A\<^sub>P) \\<^sub>P" by(rule_tac rOpen) with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* A\<^sub>x\<^sub>P\ \A\<^sub>P \* C\ \x \ A\<^sub>x\<^sub>P\ \A\<^sub>P \* A\<^sub>x\<^sub>P\ \x \ A\<^sub>P\ \x \ \\ \x \ M\ \x \ C\ \x \ xvec\ \x \ yvec\ Aeq S \distinctPerm p\ have "Prop C \ (\\x\P) (M\\*(xvec@x#yvec)\\N\ \ P') (p \ (x#A\<^sub>P)) (p \ \\<^sub>P)" by(rule_tac A\<^sub>P="x#A\<^sub>P" in rAlpha) (assumption | simp add: abs_fresh fresh_star_def boundOutputFresh)+ with \eq Aeq show ?case by(simp add: eqvts) next case(cScope \ P \ P' x A\<^sub>x\<^sub>P \\<^sub>x\<^sub>P C) obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "A\<^sub>P \* (\, P, \, P', A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P, C, x)" by(rule freshFrame) hence "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* \" and "A\<^sub>P \* P'" and "A\<^sub>P \* A\<^sub>x\<^sub>P" and "A\<^sub>P \* \\<^sub>x\<^sub>P" and "A\<^sub>P \* C" and "x \ A\<^sub>P" by simp+ from \bn \ \* P\ \A\<^sub>P \* \\ FrP have "bn \ \* \\<^sub>P" by(force dest: extractFrameFreshChain) from \extractFrame(\\x\P) = \A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P\\ FrP have "\(x#A\<^sub>P), \\<^sub>P\ = \A\<^sub>x\<^sub>P, \\<^sub>x\<^sub>P\" by simp moreover from \x \ A\<^sub>P\ \distinct A\<^sub>P\ have "distinct(x#A\<^sub>P)" by simp ultimately obtain p where S: "set p \ set (x#A\<^sub>P) \ set (p \ (x#A\<^sub>P))" and "distinctPerm p" and \eq: "\\<^sub>x\<^sub>P = p \ \\<^sub>P" and Aeq: "A\<^sub>x\<^sub>P = (p \ x)#(p \ A\<^sub>P)" using \A\<^sub>P \* A\<^sub>x\<^sub>P\\x \ A\<^sub>x\<^sub>P\ \distinct A\<^sub>x\<^sub>P\ by(rule_tac frameChainEq') (assumption | simp add: eqvts)+ note \\ \ P \\ \ P'\ FrP \distinct A\<^sub>P\ moreover from FrP \distinct A\<^sub>P\ \\A\<^sub>P \\<^sub>P C. \extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P\ \ Prop C \ P (\ \ P') A\<^sub>P \\<^sub>P\ have "\C. Prop C \ P (\ \ P') A\<^sub>P \\<^sub>P" by simp moreover note \x \ \\ \x \ \\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \distinct(bn \)\ \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* \\<^sub>P\ \A\<^sub>P \* C\ \x \ C\ \bn \ \* C\ ultimately have "Prop C \ (\\x\P) (\ \ (\\x\P')) (x#A\<^sub>P) \\<^sub>P" by(rule_tac rScope) with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \A\<^sub>P \* A\<^sub>x\<^sub>P\ \A\<^sub>P \* C\ \x \ A\<^sub>x\<^sub>P\ \A\<^sub>P \* A\<^sub>x\<^sub>P\ \x \ A\<^sub>P\ \x \ \\ \x \ \\ \x \ C\ Aeq S \distinctPerm p\ have "Prop C \ (\\x\P) (\ \ (\\x\P')) (p \ (x#A\<^sub>P)) (p \ \\<^sub>P)" by(rule_tac A\<^sub>P="x#A\<^sub>P" in rAlpha) (assumption | simp add: abs_fresh fresh_star_def)+ with \eq Aeq show ?case by(simp add: eqvts) next case(Bang \ P Rs A\<^sub>b\<^sub>P \\<^sub>b\<^sub>P C) obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "A\<^sub>P \* (\, P, Rs, C)" by(rule freshFrame) hence "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* Rs" and "A\<^sub>P \* C" by simp+ note \\ \ P \ !P \ Rs\ \guarded P\ FrP \distinct A\<^sub>P\ moreover from FrP have "extractFrame (P \ !P) = \A\<^sub>P, \\<^sub>P \ \\" by simp with \distinct A\<^sub>P\ \\A\<^sub>P \\<^sub>P C. \extractFrame (P \ !P) = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P\ \ Prop C \ (P \ !P) Rs A\<^sub>P \\<^sub>P\ have "\C. Prop C \ (P \ !P) Rs A\<^sub>P (\\<^sub>P \ \)" by simp moreover from \guarded P\ FrP have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" by(metis guardedStatEq)+ moreover note \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Rs\ \A\<^sub>P \* C\ ultimately have "Prop C \ (!P) Rs ([]) (\)" by(rule rBang) thus ?case using \extractFrame(!P) = \A\<^sub>b\<^sub>P, \\<^sub>b\<^sub>P\\ by simp qed lemma semanticsFrameInduct'[consumes 5, case_names cAlpha cFrameAlpha cInput cOutput cCase cPar1 cPar2 cComm1 cComm2 cOpen cScope cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Rs :: "('a, 'b, 'c) residual" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a action \ ('a, 'b, 'c) psi \ name list \ 'b \ bool" and C :: "'d::fs_name" assumes Trans: "\ \ P \\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "bn \ \* subject \" and "distinct(bn \)" and rAlpha: "\\ P \ P' p A\<^sub>P \\<^sub>P C. \bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* \\<^sub>P; bn \ \* C; bn \ \* (p \ \); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* C; set p \ set(bn \) \ set(bn(p \ \)); distinctPerm p; bn(p \ \) \* \; (bn(p \ \)) \* P'; Prop C \ P \ P' A\<^sub>P \\<^sub>P\ \ Prop C \ P (p \ \) (p \ P') A\<^sub>P \\<^sub>P" and rFrameAlpha: "\\ P A\<^sub>P \\<^sub>P p \ P' C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* C; set p \ set A\<^sub>P \ set(p \ A\<^sub>P); distinctPerm p; A\<^sub>P \* subject \; Prop C \ P \ P' A\<^sub>P \\<^sub>P\ \ Prop C \ P \ P' (p \ A\<^sub>P) (p \ \\<^sub>P)" and rInput: "\\ M K xvec N Tvec P C. \\ \ M \ K; distinct xvec; set xvec \ supp N; length xvec = length Tvec; xvec \* \; xvec \* M; xvec \* K; xvec \* C\ \ Prop C \ (M\\*xvec N\.P) (K\(N[xvec::=Tvec])\) (P[xvec::=Tvec]) ([]) (\)" and rOutput: "\\ M K N P C. \ \ M \ K \ Prop C \ (M\N\.P) (K\N\) P ([]) (\)" and rCase: "\\ P \ P' \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P \ P' A\<^sub>P \\<^sub>P; (\, P) mem Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) \ P' ([]) (\)" and rPar1: "\\ \\<^sub>Q P \ P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. \\ \ \\<^sub>Q \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>Q) P \ P' A\<^sub>P \\<^sub>P; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; bn \ \* \; bn \ \* P; bn \ \* Q; bn \ \* subject \; bn \ \* \\<^sub>P; bn \ \* \\<^sub>Q; A\<^sub>P \* C; A\<^sub>Q \* C; bn \ \* C\ \ Prop C \ (P \ Q) \ (P' \ Q) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rPar2: "\\ \\<^sub>P Q \ Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. \\ \ \\<^sub>P \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>P) Q \ Q' A\<^sub>Q \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; bn \ \* \; bn \ \* P; bn \ \* Q; bn \ \* subject \; bn \ \* \\<^sub>P; bn \ \* \\<^sub>Q; A\<^sub>P \* C; A\<^sub>Q \* C; bn \ \* C\ \ Prop C \ (P \ Q) \ (P \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C (\ \ \\<^sub>Q) P (M\N\) P' A\<^sub>P \\<^sub>P; \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; \C. Prop C (\ \ \\<^sub>P) Q (K\\*xvec\\N\) Q' A\<^sub>Q \\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C (\ \ \\<^sub>Q) P (M\\*xvec\\N\) P' A\<^sub>P \\<^sub>P; \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>P) Q (K\N\) Q' A\<^sub>Q \\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) (\) (\\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rOpen: "\\ P M xvec yvec N P' x A\<^sub>P \\<^sub>P y C. \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P (M\\*(xvec@yvec)\\N\) P' A\<^sub>P \\<^sub>P; x \ supp N; x \ \; x \ M; x \ A\<^sub>P; x \ xvec; x \ yvec; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* xvec; A\<^sub>P \* yvec; xvec \* yvec; distinct xvec; distinct yvec; xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; yvec \* \; yvec \* P; yvec \* M; A\<^sub>P \* C; x \ C; xvec \* C; yvec \* C; y \ x; y \ \; y \ P; y \ M; y \ xvec; y \ yvec; y \ N; y \ P'; y \ A\<^sub>P; y \ \\<^sub>P; y \ C\ \ Prop C \ (\\x\P) (M\\*(xvec@y#yvec)\\([(x, y)] \ N)\) ([(x, y)] \ P') (x#A\<^sub>P) \\<^sub>P" and rScope: "\\ P \ P' x A\<^sub>P \\<^sub>P C. \\ \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P \ P' A\<^sub>P \\<^sub>P; x \ \; x \ \; x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* \; A\<^sub>P \* P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* \\<^sub>P; A\<^sub>P \* C; x \ C; bn \ \* C\ \ Prop C \ (\\x\P) \ (\\x\P') (x#A\<^sub>P) \\<^sub>P" and rBang: "\\ P \ P' A\<^sub>P \\<^sub>P C. \\ \ P \ !P \\ \ P'; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ (P \ !P) \ P' A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; supp \\<^sub>P = ({}::name set); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (!P) \ P' ([]) (\)" shows "Prop C \ P \ P' A\<^sub>P \\<^sub>P" using Trans FrP \distinct A\<^sub>P\ \bn \ \* subject \\ \distinct(bn \)\ proof(nominal_induct \ P Rs=="\ \ P'" A\<^sub>P \\<^sub>P avoiding: C \ P' rule: semanticsFrameInduct) case cAlpha thus ?case using rFrameAlpha by auto next case cInput thus ?case using rInput by(auto simp add: residualInject) next case cOutput thus ?case using rOutput by(auto simp add: residualInject) next case cCase thus ?case using rCase by(auto simp add: residualInject) next case(cPar1 \ \\<^sub>Q P \ P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C \' P'') note \\ \ (P' \ Q) = \' \ P''\ moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto moreover note \distinct (bn \)\ \distinct(bn \')\ moreover from \bn \ \* subject \\ \bn \' \* subject \'\ have "bn \ \* (\ \ P' \ Q)" and "bn \' \* (\' \ P'')" by simp+ ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" and \Eq: "\' = p \ \" and P'eq: "P'' = p \ (P' \ Q)" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* (P' \ Q)" by(rule residualEq) note \\ \ \\<^sub>Q \ P \\ \ P'\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ moreover from \bn \ \* subject \\ \distinct(bn \)\ \A\<^sub>P \* \\ have "\C. Prop C (\ \ \\<^sub>Q) P \ P' A\<^sub>P \\<^sub>P" by(rule_tac cPar1) auto moreover note \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* C\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \A\<^sub>P \* C\ \bn \ \* Q\ \distinct(bn \)\ \bn \ \* \\ \bn \ \* \\<^sub>Q\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ \bn \ \* \\<^sub>P\ ultimately have "Prop C \ (P \ Q) \ (P' \ Q) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" by(rule_tac rPar1) with \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* (bn \')\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (P' \ Q)\ \bn \ \* \\<^sub>P\ \bn \ \* \\<^sub>Q\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* \'\ \A\<^sub>Q \* \'\ \Eq \bn \ \* \\<^sub>P\ \bn \ \* \'\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* P'\ \A\<^sub>Q \* P'\ \A\<^sub>P \* C\ \A\<^sub>Q \* C\ have "Prop C \ (P \ Q) (p \ \) (p \ (P' \ Q)) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" by(rule_tac rAlpha) auto with \Eq P'eq \distinctPerm p\ show ?case by simp next case(cPar2 \ \\<^sub>P Q \ Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C \' Q'') note \\ \ (P \ Q') = \' \ Q''\ moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto moreover note \distinct (bn \)\ \distinct(bn \')\ moreover from \bn \ \* subject \\ \bn \' \* subject \'\ have "bn \ \* (\ \ P \ Q')" and "bn \' \* (\' \ Q'')" by simp+ ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" and \Eq: "\' = p \ \" and Q'eq: "Q'' = p \ (P \ Q')" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* (P \ Q')" by(rule residualEq) note \\ \ \\<^sub>P \ Q \\ \ Q'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ moreover from \bn \ \* subject \\ \distinct(bn \)\ \A\<^sub>Q \* \\ have "\C. Prop C (\ \ \\<^sub>P) Q \ Q' A\<^sub>Q \\<^sub>Q" by(rule_tac cPar2) auto moreover note \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* C\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ \A\<^sub>P \* Q'\ \A\<^sub>P \* C\ \bn \ \* Q\ \distinct(bn \)\ \bn \ \* \\ \bn \ \* \\<^sub>Q\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ \bn \ \* \\<^sub>P\ ultimately have "Prop C \ (P \ Q) \ (P \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" by(rule_tac rPar2) auto with \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* (bn \')\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (P \ Q')\ \bn \ \* \\<^sub>P\ \bn \ \* \\<^sub>Q\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* \'\ \A\<^sub>Q \* \'\ \Eq \bn \ \* \'\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P \* Q'\ \A\<^sub>Q \* Q'\ \A\<^sub>P \* C\ \A\<^sub>Q \* C\ have "Prop C \ (P \ Q) (p \ \) (p \ (P \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" by(rule_tac rAlpha) auto with \Eq Q'eq \distinctPerm p\ show ?case by simp next case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C \ P'') thus ?case using rComm1 apply(auto) apply(drule_tac x="M\N\" in meta_spec) apply(drule_tac x="K\\*xvec\\N\" in meta_spec) apply(drule_tac x=P' in meta_spec) apply(drule_tac x=Q' in meta_spec) apply auto apply(drule_tac x=\ in meta_spec) apply(drule_tac x=\\<^sub>Q in meta_spec) apply(drule_tac x=P in meta_spec) apply(drule_tac x=M in meta_spec) apply(drule_tac x=N in meta_spec) apply(drule_tac x=P' in meta_spec) apply(drule_tac x=A\<^sub>P in meta_spec) apply(drule_tac x=\\<^sub>P in meta_spec) apply(drule_tac x=Q in meta_spec) apply(drule_tac x=K in meta_spec) apply(drule_tac x=xvec in meta_spec) apply(drule_tac x=Q' in meta_spec) apply(drule_tac x=A\<^sub>Q in meta_spec) apply auto apply(subgoal_tac "\C. Prop C (\ \ \\<^sub>P) Q (K\\*xvec\\N\) Q' A\<^sub>Q \\<^sub>Q") apply(subgoal_tac "\C. Prop C (\ \ \\<^sub>Q) P (M\N\) P' A\<^sub>P \\<^sub>P") by(auto simp add: residualInject) next case(cComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C \ Q'') thus ?case using rComm2 apply(drule_tac x="M\\*xvec\\N\" in meta_spec) apply(drule_tac x="K\N\" in meta_spec) apply(drule_tac x=P' in meta_spec) apply(drule_tac x=Q' in meta_spec) apply auto apply(drule_tac x=\ in meta_spec) apply(drule_tac x=\\<^sub>Q in meta_spec) apply(drule_tac x=P in meta_spec) apply(drule_tac x=M in meta_spec) apply(drule_tac x=xvec in meta_spec) apply(drule_tac x=N in meta_spec) apply(drule_tac x=P' in meta_spec) apply(drule_tac x=A\<^sub>P in meta_spec) apply(drule_tac x=\\<^sub>P in meta_spec) apply(drule_tac x=Q in meta_spec) apply(drule_tac x=K in meta_spec) apply(drule_tac x=Q' in meta_spec) apply(drule_tac x=A\<^sub>Q in meta_spec) apply auto apply(subgoal_tac "\C. Prop C (\ \ \\<^sub>Q) P (M\\*xvec\\N\) P' A\<^sub>P \\<^sub>P") apply(subgoal_tac "\C. Prop C (\ \ \\<^sub>P) Q (K\N\) Q' A\<^sub>Q \\<^sub>Q") by(auto simp add: residualInject) next case(cOpen \ P M xvec yvec N P' x A\<^sub>P \\<^sub>P C \ P'') note \M\\*(xvec@x#yvec)\\N\ \ P' = \ \ P''\ moreover from \xvec \* \\ \x \ \\ \yvec \* \\ have "(xvec@x#yvec) \* (bn \)" by auto moreover from \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \distinct xvec\ \distinct yvec\ have "distinct(xvec@x#yvec)" by(auto simp add: fresh_star_def) (simp add: fresh_def name_list_supp) moreover note \distinct(bn \)\ moreover from \xvec \* M\ \x \ M\ \yvec \* M\ have "(xvec@x#yvec) \* M" by auto hence "(xvec@x#yvec) \* (M\\*(xvec@x#yvec)\\N\ \ P')" by auto moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P'')" by simp ultimately obtain p where S: "(set p) \ (set(xvec@x#yvec)) \ (set(p \ (xvec@x#yvec)))" and "distinctPerm p" and \eq: "\ = (p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\" and P'eq: "P'' = (p \ P')" and A: "(xvec@x#yvec) \* ((p \ M)\\*(p \ (xvec@x#yvec))\\(p \ N)\)" and B: "(p \ (xvec@x#yvec)) \* (M\\*(xvec@x#yvec)\\N\)" and C: "(p \ (xvec@x#yvec)) \* P'" by(rule_tac residualEq) (assumption | simp)+ note \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ \x \ (supp N)\ moreover { fix C from \xvec \* M\ \yvec \* M\ have "(xvec@yvec) \* M" by simp moreover from \distinct xvec\ \distinct yvec\ \xvec \* yvec\ have "distinct(xvec@yvec)" by auto (simp add: fresh_star_def name_list_supp fresh_def) ultimately have "Prop C \ P (M\\*(xvec@yvec)\\N\) P' A\<^sub>P \\<^sub>P" using \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* M\ \A\<^sub>P \* N\ by(rule_tac cOpen) auto } moreover obtain y::name where "y \ \" and "y \ x" and "y \ P" and "y \ xvec" and "y \ yvec" and "y \ \" and "y \ P'" and "y \ A\<^sub>P" and "y \ \\<^sub>P" and "y \ M" and "y \ N" and "y \ C" and "y \ p" by(generate_fresh "name") auto moreover note \x \ \\ \x \ M\ \x \ xvec\ \x \ yvec\ \xvec \* \\ \xvec \* P\ \xvec \* M\ \yvec \* \\ \yvec \* P\ \yvec \* M\ \yvec \* C\ \x \ C\ \xvec \* C\ \distinct xvec\ \distinct yvec\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \x \ A\<^sub>P\ \xvec \* yvec\ \xvec \* \\<^sub>P\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* xvec\ \A\<^sub>P \* yvec\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* C\ ultimately have "Prop C \ (\\x\P) (M\\*(xvec@y#yvec)\\([(x, y)] \ N)\) ([(x, y)] \ P') (x#A\<^sub>P) \\<^sub>P" by(rule_tac rOpen) moreover have "(([(x, y)] \ p) \ [(x, y)] \ M) = [(x, y)] \ p \ M" by(subst perm_compose[symmetric]) simp with \y \ M\ \x \ \\ \eq \y \ p\ \x \ M\ have D: "(([(x, y)] \ p) \ M) = p \ M" by(auto simp add: eqvts freshChainSimps) moreover have "(([(x, y)] \ p) \ [(x, y)] \ xvec) = [(x, y)] \ p \ xvec" by(subst perm_compose[symmetric]) simp with \y \ xvec\ \x \ \\ \eq \y \ p\ \x \ xvec\ have E: "(([(x, y)] \ p) \ xvec) = p \ xvec" by(auto simp add: eqvts freshChainSimps) moreover have "(([(x, y)] \ p) \ [(x, y)] \ yvec) = [(x, y)] \ p \ yvec" by(subst perm_compose[symmetric]) simp with \y \ yvec\ \x \ \\ \eq \y \ p\ \x \ yvec\ have F: "(([(x, y)] \ p) \ yvec) = p \ yvec" by(auto simp add: eqvts freshChainSimps) moreover have "(([(x, y)] \ p) \ [(x, y)] \ x) = [(x, y)] \ p \ x" by(subst perm_compose[symmetric]) simp with \y \ x\ \y \ p\ have G: "(([(x, y)] \ p) \ y) = p \ x" apply(simp add: freshChainSimps calc_atm) apply(subgoal_tac "y \ p \ x") apply(clarsimp) using A \eq apply(simp add: eqvts) apply(subst fresh_atm[symmetric]) apply(simp only: freshChainSimps) by simp moreover have "(([(x, y)] \ p) \ [(x, y)] \ N) = [(x, y)] \ p \ N" by(subst perm_compose[symmetric]) simp with \y \ N\ \x \ \\ \y \ p\ \eq have H: "(([(x, y)] \ p) \ [(x, y)] \ N) = p \ N" by(auto simp add: eqvts freshChainSimps) moreover have "(([(x, y)] \ p) \ [(x, y)] \ P') = [(x, y)] \ p \ P'" by(subst perm_compose[symmetric]) simp with \y \ P'\ \x \ P''\ \y \ p\ P'eq have I: "(([(x, y)] \ p) \ [(x, y)] \ P') = p \ P'" by(auto simp add: eqvts freshChainSimps) from \y \ p\ \y \ x\ have "y \ p \ x" apply(subst fresh_atm[symmetric]) apply(simp only: freshChainSimps) by simp moreover from S have "([(x, y)] \ set p) \ [(x, y)] \ (set(xvec@x#yvec) \ set(p \ (xvec@x#yvec)))" by(simp) with \y \ p \ x\ \(([(x, y)] \ p) \ y) = p \ x\ \x \ xvec\ \y \ xvec\ \x \ yvec\ \y \ yvec\ \y \ p\ \x \ \\ \eq have "set([(x, y)] \ p) \ set(xvec@y#yvec) \ set(([(x, y)] \ p) \ (xvec@y#yvec))" by(simp add: eqvts calc_atm perm_compose) moreover note \xvec \* \\ \yvec \* \\ \xvec \* P\ \yvec \* P\ \xvec \* M\ \yvec \* M\ \yvec \* C\ S \distinctPerm p\ \x \ C\ \xvec \* C\ \xvec \* \\<^sub>P\ \yvec \* \\<^sub>P\ \x \ \\ \A\<^sub>P \* xvec\ \x \ A\<^sub>P\ \A\<^sub>P \* yvec\ \A\<^sub>P \* M\ \x \ xvec\ \x \ yvec\ \x \ M\ \x \ A\<^sub>P\ \A\<^sub>P \* N\ A B C \eq \A\<^sub>P \* \\ \y \ \\ \y \ x\ \y \ P\ \y \ M\ \y \ \\<^sub>P\ \y \ C\ \xvec \* \\ \x \ \\ \yvec \* \\ \y \ \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\ \y \ A\<^sub>P\ \y \ N\ \A\<^sub>P \* P'\ \y \ P'\ \A\<^sub>P \* C\ P'eq ultimately have "Prop C \ (\\x\P) (([(x, y)] \ p) \ (M\\*(xvec@y#yvec)\\([(x, y)] \ N)\)) (([(x, y)] \ p) \ [(x, y)] \ P') (x#A\<^sub>P) \\<^sub>P" apply(rule_tac \="M\\*(xvec@y#yvec)\\([(x, y)] \ N)\" in rAlpha) apply(assumption | simp)+ apply(simp add: eqvts) apply(assumption | simp add: abs_fresh)+ apply(simp add: fresh_left calc_atm) apply(assumption | simp)+ apply(simp add: fresh_left calc_atm) apply(assumption | simp)+ by(simp add: eqvts fresh_left)+ with \eq P'eq D E F G H I show ?case by(simp add: eqvts) next case(cScope \ P \ P' x A\<^sub>P \\<^sub>P C \' P'') note \\ \ (\\x\P') = \' \ P''\ moreover from \bn \ \* \'\ have "bn \ \* (bn \')" by auto moreover note \distinct (bn \)\ \distinct(bn \')\ moreover from \bn \ \* subject \\ \bn \' \* subject \'\ have "bn \ \* (\ \ \\x\P')" and "bn \' \* (\' \ P'')" by simp+ ultimately obtain p where S: "(set p) \ (set(bn \)) \ (set(bn(p \ \)))" and "distinctPerm p" and \Eq: "\' = p \ \" and P'eq: "P'' = p \ (\\x\P')" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* (\\x\P')" by(rule residualEq) note \\ \ P \\ \ P'\ moreover from \bn \ \* subject \\ \distinct(bn \)\ have "\C. Prop C \ P \ P' A\<^sub>P \\<^sub>P" by(rule_tac cScope) auto moreover note \x \ \\ \x \ \\ \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* \\<^sub>P\ \x \ C\ \bn \ \* C\ \distinct(bn \)\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* P'\ \A\<^sub>P \* C\ ultimately have "Prop C \ (\\x\P) \ (\\x\P') (x#A\<^sub>P) \\<^sub>P" by(rule_tac rScope) with \bn \ \* \\ \bn \ \* P\ \x \ \\ \bn \ \* subject \\ \bn \ \* C\ \bn \ \* (bn \')\ S \distinctPerm p\ \bn(p \ \) \* \\ \bn(p \ \) \* (\\x\P')\ \A\<^sub>P \* \\ \A\<^sub>P \* \'\ \Eq \x \ \'\ \bn \ \* \\<^sub>P\ \bn \ \* \'\ \x \ \\ \A\<^sub>P \* \\ \x \ A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>P \* P'\ \x \ C\ \A\<^sub>P \* C\ have "Prop C \ (\\x\P) (p \ \) (p \ (\\x\P')) (x#A\<^sub>P) \\<^sub>P" by(rule_tac rAlpha) (simp add: abs_fresh)+ with \Eq P'eq \distinctPerm p\ show ?case by simp next case(cBang \ P Rs A\<^sub>P \\<^sub>P C \) thus ?case by(rule_tac rBang) auto qed lemma inputFrameInduct[consumes 3, case_names cAlpha cInput cCase cPar1 cPar2 cScope cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a \ 'a \ ('a, 'b, 'c) psi \ name list \ 'b \ bool" and C :: "'d::fs_name" assumes Trans: "\ \ P \M\N\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and rAlpha: "\\ P M N P' A\<^sub>P \\<^sub>P p C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* C; set p \ set A\<^sub>P \ set(p \ A\<^sub>P); distinctPerm p; Prop C \ P M N P' A\<^sub>P \\<^sub>P\ \ Prop C \ P M N P' (p \ A\<^sub>P) (p \ \\<^sub>P)" and rInput: "\\ M K xvec N Tvec P C. \\ \ M \ K; distinct xvec; set xvec \ supp N; length xvec = length Tvec; xvec \* \; xvec \* M; xvec \* K; xvec \* C\ \ Prop C \ (M\\*xvec N\.P) K (N[xvec::=Tvec]) (P[xvec::=Tvec]) ([]) (\)" and rCase: "\\ P M N P' \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P M N P' A\<^sub>P \\<^sub>P; (\, P) mem Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) M N P' ([]) (\)" and rPar1: "\\ \\<^sub>Q P M N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>Q) P M N P' A\<^sub>P \\<^sub>P; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* C; A\<^sub>Q \* C\ \ Prop C \ (P \ Q) M N (P' \ Q) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rPar2: "\\ \\<^sub>P Q M N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. \\ \ \\<^sub>P \ Q \M\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>P) Q M N Q' A\<^sub>Q \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* N; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* C; A\<^sub>Q \* C\ \ Prop C \ (P \ Q) M N (P \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rScope: "\\ P M N P' x A\<^sub>P \\<^sub>P C. \\ \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P M N P' A\<^sub>P \\<^sub>P; x \ \; x \ M; x \ N; x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* C; x \ C\ \ Prop C \ (\\x\P) M N (\\x\P') (x#A\<^sub>P) \\<^sub>P" and rBang: "\\ P M N P' A\<^sub>P \\<^sub>P C. \\ \ P \ !P \M\N\ \ P'; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ (P \ !P) M N P' A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (!P) M N P' ([]) (\)" shows "Prop C \ P M N P' A\<^sub>P \\<^sub>P" using assms by(nominal_induct \ P Rs=="M\N\ \ P'" A\<^sub>P \\<^sub>P avoiding: C arbitrary: P' rule: semanticsFrameInduct) (auto simp add: residualInject) lemma outputFrameInduct[consumes 3, case_names cAlpha cOutput cCase cPar1 cPar2 cOpen cScope cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and B :: "('a, 'b, 'c) boundOutput" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a \ ('a, 'b, 'c) boundOutput \ name list \ 'b \ bool" and C :: "'d::fs_name" assumes Trans: "\ \ P \ROut M B" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and rAlpha: "\\ P M A\<^sub>P \\<^sub>P p B C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* B; A\<^sub>P \* C; set p \ set A\<^sub>P \ set(p \ A\<^sub>P); distinctPerm p; Prop C \ P M B A\<^sub>P \\<^sub>P\ \ Prop C \ P M B (p \ A\<^sub>P) (p \ \\<^sub>P)" and rOutput: "\\ M K N P C. \ \ M \ K \ Prop C \ (M\N\.P) K (N \' P) ([]) (\)" and rCase: "\\ P M B \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \(ROut M B); extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P M B A\<^sub>P \\<^sub>P; (\, P) mem Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* B; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) M B ([]) (\)" and rPar1: "\\ \\<^sub>Q P M xvec N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>Q) P M (\\*xvec\N \' P') A\<^sub>P \\<^sub>P; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; xvec \* \; xvec \* P; xvec \* Q; xvec \* M; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) M (\\*xvec\N \' (P' \ Q)) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rPar2: "\\ \\<^sub>P Q M xvec N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>P) Q M (\\*xvec\N \' Q') A\<^sub>Q \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* M; A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* M; A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; xvec \* \; xvec \* P; xvec \* Q; xvec \* M; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) M (\\*xvec\N \' (P \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rOpen: "\\ P M xvec yvec N P' x A\<^sub>P \\<^sub>P C. \\ \ P \M\\*(xvec@yvec)\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P M (\\*(xvec@yvec)\N \' P') A\<^sub>P \\<^sub>P; x \ supp N; x \ \; x \ M; x \ A\<^sub>P; x \ xvec; x \ yvec; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* xvec; A\<^sub>P \* yvec; xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; yvec \* \; yvec \* P; yvec \* M; A\<^sub>P \* C; x \ C; xvec \* C; yvec \* C\ \ Prop C \ (\\x\P) M (\\*(xvec@x#yvec)\N \' P') (x#A\<^sub>P) \\<^sub>P" and rScope: "\\ P M xvec N P' x A\<^sub>P \\<^sub>P C. \\ \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P M (\\*xvec\N \' P') A\<^sub>P \\<^sub>P; x \ \; x \ M; x \ xvec; x \ N; x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* xvec; xvec \* \; xvec \* P; xvec \* M; xvec \* \\<^sub>P; A\<^sub>P \* C; x \ C; xvec \* C\ \ Prop C \ (\\x\P) M (\\*xvec\N \' (\\x\P')) (x#A\<^sub>P) \\<^sub>P" and rBang: "\\ P M B A\<^sub>P \\<^sub>P C. \\ \ P \ !P \ROut M B; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ (P \ !P) M B A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; supp \\<^sub>P = ({}::name set); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* C\ \ Prop C \ (!P) M B ([]) (\)" shows "Prop C \ P M B A\<^sub>P \\<^sub>P" proof - { fix B assume "\ \ P \ROut M B" hence "Prop C \ P M B A\<^sub>P \\<^sub>P" using FrP \distinct A\<^sub>P\ proof(nominal_induct \ P Rs=="ROut M B" A\<^sub>P \\<^sub>P avoiding: C arbitrary: B rule: semanticsFrameInduct) case cAlpha thus ?case by(fastforce intro: rAlpha) next case cInput thus ?case by(simp add: residualInject) next case cOutput thus ?case by(force intro: rOutput simp add: residualInject) next case cCase thus ?case by(force intro: rCase simp add: residualInject) next case cPar1 thus ?case by(fastforce intro: rPar1 simp add: residualInject) next case cPar2 thus ?case by(fastforce intro: rPar2 simp add: residualInject) next case cComm1 thus ?case by(simp add: residualInject) next case cComm2 thus ?case by(simp add: residualInject) next case cOpen thus ?case by(fastforce intro: rOpen simp add: residualInject) next case cScope thus ?case by(force intro: rScope simp add: residualInject) next case cBang thus ?case by(force intro: rBang simp add: residualInject) qed } with Trans show ?thesis by(simp add: residualInject) qed lemma tauFrameInduct[consumes 3, case_names cAlpha cCase cPar1 cPar2 cComm1 cComm2 cScope cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and P' :: "('a, 'b, 'c) psi" and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi \ name list \ 'b \ bool" and C :: "'d::fs_name" assumes Trans: "\ \ P \\ \ P'" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and rAlpha: "\\ P P' A\<^sub>P \\<^sub>P p C. \A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* P'; A\<^sub>P \* (p \ A\<^sub>P); A\<^sub>P \* C; set p \ set A\<^sub>P \ set (p \ A\<^sub>P); distinctPerm p; Prop C \ P P' A\<^sub>P \\<^sub>P\ \ Prop C \ P P' (p \ A\<^sub>P) (p \ \\<^sub>P)" and rCase: "\\ P P' \ Cs A\<^sub>P \\<^sub>P C. \\ \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P P' A\<^sub>P \\<^sub>P; (\, P) mem Cs; \ \ \; guarded P; \\<^sub>P \ \; (supp \\<^sub>P) = ({}::name set); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (Cases Cs) P' ([]) (\)" and rPar1: "\\ \\<^sub>Q P P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C. \\ \ \\<^sub>Q \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>Q) P P' A\<^sub>P \\<^sub>P; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* P'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* P'; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* C; A\<^sub>Q \* C\ \ Prop C \ (P \ Q) (P' \ Q) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rPar2: "\\ \\<^sub>P Q Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C. \\ \ \\<^sub>P \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \C. Prop C (\ \ \\<^sub>P) Q Q' A\<^sub>Q \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* Q'; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* C; A\<^sub>Q \* C\ \ Prop C \ (P \ Q) (P \ Q') (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rComm1: "\\ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) (\\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rComm2: "\\ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* K; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C\ \ Prop C \ (P \ Q) (\\*xvec\(P' \ Q')) (A\<^sub>P@A\<^sub>Q) (\\<^sub>P \ \\<^sub>Q)" and rScope: "\\ P P' x A\<^sub>P \\<^sub>P C. \\ \ P \\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ P P' A\<^sub>P \\<^sub>P; x \ \; x \ A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* P'; A\<^sub>P \* C; x \ C\ \ Prop C \ (\\x\P) (\\x\P') (x#A\<^sub>P) \\<^sub>P" and rBang: "\\ P P' A\<^sub>P \\<^sub>P C. \\ \ P \ !P \\ \ P'; guarded P; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \C. Prop C \ (P \ !P) P' A\<^sub>P (\\<^sub>P \ \); \\<^sub>P \ \; supp \\<^sub>P = ({}::name set); A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* P'; A\<^sub>P \* C\ \ Prop C \ (!P) P' ([]) (\)" shows "Prop C \ P P' A\<^sub>P \\<^sub>P" using Trans FrP \distinct A\<^sub>P\ proof(nominal_induct \ P Rs=="\ \ P'" A\<^sub>P \\<^sub>P avoiding: C arbitrary: P' rule: semanticsFrameInduct) case cAlpha thus ?case by(force intro: rAlpha simp add: residualInject) next case cInput thus ?case by(simp add: residualInject) next case cOutput thus ?case by(simp add: residualInject) next case cCase thus ?case by(force intro: rCase simp add: residualInject) next case cPar1 thus ?case by(force intro: rPar1 simp add: residualInject) next case cPar2 thus ?case by(force intro: rPar2 simp add: residualInject) next case cComm1 thus ?case by(force intro: rComm1 simp add: residualInject) next case cComm2 thus ?case by(force intro: rComm2 simp add: residualInject) next case cOpen thus ?case by(simp add: residualInject) next case cScope thus ?case by(force intro: rScope simp add: residualInject) next case cBang thus ?case by(force intro: rBang simp add: residualInject) qed lemma inputFreshDerivative: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and x :: name assumes "\ \ P \M\N\ \ P'" and "x \ P" and "x \ N" shows "x \ P'" proof - have "bn(M\N\) \* subject(M\N\)" and "distinct(bn(M\N\))" by simp+ with \\ \ P \M\N\ \ P'\ show ?thesis using \x \ P\ \x \ N\ proof(nominal_induct \ P \=="M\N\" P' avoiding: x rule: semanticsInduct) case(cAlpha \ P \ P' p x) thus ?case by simp next case(cInput \ M' K xvec N' Tvec P x) from \K\(N'[xvec::=Tvec])\ = M\N\\ have "M = K" and NeqN': "N = N'[xvec::=Tvec]" by(simp add: action.inject)+ note \length xvec = length Tvec\ \distinct xvec\ then moreover have "x \ Tvec" using \set xvec \ supp N'\ \x \ N\ NeqN' by(blast intro: substTerm.subst3) moreover from \xvec \* x\ \x \ M'\\*xvec N'\.P\ have "x \ P" by(simp add: inputChainFresh) (simp add: name_list_supp fresh_def) ultimately show ?case using \xvec \* x\ by auto next case(cOutput \ M K N P x) thus ?case by simp next case(cCase \ P P' \ Cs x) thus ?case by(induct Cs, auto) next case(cPar1 \ \\<^sub>Q P P' xvec Q x) thus ?case by simp next case(cPar2 \ \\<^sub>P Q Q' xvec P x) thus ?case by simp next case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q x) thus ?case by simp next case(cComm2 \ \\<^sub>Q P M xwec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q x) thus ?case by simp next case(cOpen \ P M xvec yvec N P' x y) thus ?case by simp next case(cScope \ P P' x y) thus ?case by(simp add: abs_fresh) next case(cBang \ P P' x) thus ?case by simp qed qed lemma inputFreshChainDerivative: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and xvec :: "name list" assumes "\ \ P \M\N\ \ P'" and "xvec \* P" and "xvec \* N" shows "xvec \* P'" using assms by(induct xvec) (auto intro: inputFreshDerivative) lemma outputFreshDerivative: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and N :: 'a and P' :: "('a, 'b, 'c) psi" and x :: name assumes "\ \ P \M\\*xvec\\N\ \ P'" and "xvec \* M" and "distinct xvec" and "x \ P" and "x \ xvec" shows "x \ N" and "x \ P'" proof - note \\ \ P \M\\*xvec\\N\ \ P'\ moreover from \xvec \* M\ have "bn(M\\*xvec\\N\) \* subject(M\\*xvec\\N\)" by simp moreover from \distinct xvec\ have "distinct(bn(M\\*xvec\\N\))" by simp ultimately show "x \ N" using \x \ P\ \x \ xvec\ proof(nominal_induct \ P \=="M\\*xvec\\N\" P' avoiding: x arbitrary: M xvec N rule: semanticsInduct) case(cAlpha \ P \ P' p x M xvec N) have S: "set p \ set(bn \) \ set(bn(p \ \))" by fact from \(p \ \) = M\\*xvec\\N\\ have "(p \ p \ \) = p \ (M\\*xvec\\N\)" by(simp add: fresh_star_bij) with \distinctPerm p\ have "\ = (p \ M)\\*(p \ xvec)\\(p \ N)\" by simp moreover from \(p \ \) = M\\*xvec\\N\\ \x \ xvec\ have "x \ (bn(p \ \))" by simp with \(bn \) \* x\ \x \ xvec\ S have "x \ (p \ xvec)" by(drule_tac pt_fresh_bij1[OF pt_name_inst, OF at_name_inst, where pi=p and x=xvec]) simp ultimately have "x \ (p \ N)" using \x \ P\ by(rule_tac cAlpha) hence "(p \ x) \ (p \ p \ N)" by(simp add: pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) with \distinctPerm p\ \bn(\) \* x\ \x \ (bn(p \ \))\S show ?case by simp next case cInput thus ?case by simp next case cOutput thus ?case by(simp add: action.inject) next case cCase thus ?case by(rule_tac cCase) (auto dest: memFresh) next case cPar1 thus ?case by simp next case cPar2 thus ?case by simp next case cComm1 thus ?case by simp next case cComm2 thus ?case by simp next case(cOpen \ P M xvec yvec N P' x y M' zvec N') from \M\\*(xvec@x#yvec)\\N\ = M'\\*zvec\\N'\\ have "zvec = xvec@x#yvec" and "N = N'" by(simp add: action.inject)+ from \y \ \\x\P\ \x \ y\ have "y \ P" by(simp add: abs_fresh) moreover from \y \ zvec\ \zvec = xvec@x#yvec\have "y \ (xvec@yvec)" by simp ultimately have "y \ N" by(rule_tac cOpen) auto with \N = N'\ show ?case by simp next case cScope thus ?case by(auto simp add: abs_fresh) next case cBang thus ?case by simp qed next note \\ \ P \M\\*xvec\\N\ \ P'\ moreover from \xvec \* M\ have "bn(M\\*xvec\\N\) \* subject(M\\*xvec\\N\)" by simp moreover from \distinct xvec\ have "distinct(bn(M\\*xvec\\N\))" by simp ultimately show "x \ P'" using \x \ P\ \x \ xvec\ proof(nominal_induct \ P \=="M\\*xvec\\N\" P' avoiding: x arbitrary: M xvec N rule: semanticsInduct) case(cAlpha \ P \ P' p x M xvec N) have S: "set p \ set(bn \) \ set(bn(p \ \))" by fact from \(p \ \) = M\\*xvec\\N\\ have "(p \ p \ \) = p \ (M\\*xvec\\N\)" by(simp add: fresh_star_bij) with \distinctPerm p\ have "\ = (p \ M)\\*(p \ xvec)\\(p \ N)\" by simp moreover from \(p \ \) = M\\*xvec\\N\\ \x \ xvec\ have "x \ (bn(p \ \))" by simp with \(bn \) \* x\ \x \ xvec\ S have "x \ (p \ xvec)" by(drule_tac pt_fresh_bij1[OF pt_name_inst, OF at_name_inst, where pi=p and x=xvec]) simp ultimately have "x \ P'" using \x \ P\ by(rule_tac cAlpha) hence "(p \ x) \ (p \ P')" by(simp add: pt_fresh_bij1[OF pt_name_inst, OF at_name_inst]) with \distinctPerm p\ \bn(\) \* x\ \x \ (bn(p \ \))\S show ?case by simp next case cInput thus ?case by simp next case cOutput thus ?case by(simp add: action.inject) next case cCase thus ?case by(fastforce simp add: action.inject dest: memFresh) next case cPar1 thus ?case by simp next case cPar2 thus ?case by simp next case cComm1 thus ?case by simp next case cComm2 thus ?case by simp next case(cOpen \ P M xvec yvec N P' x y M' zvec N') from \M\\*(xvec@x#yvec)\\N\ = M'\\*zvec\\N'\\ have "zvec = xvec@x#yvec" by(simp add: action.inject) from \y \ \\x\P\ \x \ y\ have "y \ P" by(simp add: abs_fresh) moreover from \y \ zvec\ \zvec = xvec@x#yvec\have "y \ (xvec@yvec)" by simp ultimately show "y \ P'" by(rule_tac cOpen) auto next case cScope thus ?case by(auto simp add: abs_fresh) next case cBang thus ?case by simp qed qed lemma outputFreshChainDerivative: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and N :: 'a and P' :: "('a, 'b, 'c) psi" and yvec :: "name list" assumes "\ \ P \M\\*xvec\\N\ \ P'" and "xvec \* M" and "distinct xvec" and "yvec \* P" and "yvec \* xvec" shows "yvec \* N" and "yvec \* P'" using assms by(induct yvec) (auto intro: outputFreshDerivative) lemma tauFreshDerivative: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and P' :: "('a, 'b, 'c) psi" and x :: name assumes "\ \ P \\ \ P'" and "x \ P" shows "x \ P'" proof - have "bn(\) \* subject(\)" and "distinct(bn(\))" by simp+ with \\ \ P \\ \ P'\ show ?thesis using \x \ P\ proof(nominal_induct \ P \=="(\::('a action))" P' avoiding: x rule: semanticsInduct) case cAlpha thus ?case by simp next case cInput thus ?case by simp next case cOutput thus ?case by simp next case cCase thus ?case by(auto dest: memFresh) next case cPar1 thus ?case by simp next case cPar2 thus ?case by simp next case cComm1 thus ?case by(fastforce dest: inputFreshDerivative outputFreshDerivative simp add: resChainFresh) next case cComm2 thus ?case by(fastforce dest: inputFreshDerivative outputFreshDerivative simp add: resChainFresh) next case cOpen thus ?case by simp next case cScope thus ?case by(simp add: abs_fresh) next case cBang thus ?case by simp qed qed lemma tauFreshChainDerivative: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and xvec :: "name list" assumes "\ \ P \\ \ P'" and "xvec \* P" shows "xvec \* P'" using assms by(induct xvec) (auto intro: tauFreshDerivative) lemma freeFreshDerivative: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and x :: name assumes "\ \ P \\ \ P'" and "bn \ \* subject \" and "distinct(bn \)" and "x \ \" and "x \ P" shows "x \ P'" using assms by(rule_tac actionCases[where \=\]) (auto intro: inputFreshDerivative tauFreshDerivative outputFreshDerivative) lemma freeFreshChainDerivative: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and xvec :: "name list" assumes "\ \ P \\ \ P'" and "bn \ \* subject \" and "distinct(bn \)" and "xvec \* P" and "xvec \* \" shows "xvec \* P'" using assms by(auto intro: freeFreshDerivative simp add: fresh_star_def) lemma Input: fixes \ :: 'b and M :: 'a and K :: 'a and xvec :: "name list" and N :: 'a and Tvec :: "'a list" assumes "\ \ M \ K" and "distinct xvec" and "set xvec \ supp N" and "length xvec = length Tvec" shows "\ \ M\\*xvec N\.P \K\N[xvec::=Tvec]\ \ P[xvec::=Tvec]" proof - obtain p where xvecFreshPsi: "((p::name prm) \ (xvec::name list)) \* \" and xvecFreshM: "(p \ xvec) \* M" and xvecFreshN: "(p \ xvec) \* N" and xvecFreshK: "(p \ xvec) \* K" and xvecFreshTvec: "(p \ xvec) \* Tvec" and xvecFreshP: "(p \ xvec) \* P" and S: "(set p) \ (set xvec) \ (set(p \ xvec))" and dp: "distinctPerm p" by(rule_tac xvec=xvec and c="(\, M, K, N, P, Tvec)" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) note \\ \ M \ K\ moreover from \distinct xvec\ have "distinct(p \ xvec)" by simp moreover from \(set xvec) \ (supp N)\ have "(p \ (set xvec)) \ (p \ (supp N))" by simp hence "set(p \ xvec) \ supp(p \ N)" by(simp add: eqvts) moreover from \length xvec = length Tvec\ have "length(p \ xvec) = length Tvec" by simp ultimately have "\ \ M\\*(p \ xvec) (p \ N)\.(p \ P) \K\(p \ N)[(p \ xvec)::=Tvec]\ \ (p \ P)[(p \ xvec)::=Tvec]" using xvecFreshPsi xvecFreshM xvecFreshK xvecFreshTvec by(rule_tac cInput) thus ?thesis using xvecFreshN xvecFreshP S \length xvec = length Tvec\ dp by(auto simp add: inputChainAlpha' substTerm.renaming renaming) qed lemma residualAlpha: fixes p :: "name prm" and \ :: "'a action" and P :: "('a, 'b, 'c) psi" assumes "bn(p \ \) \* object \" and "bn(p \ \) \* P" and "bn \ \* subject \" and "bn(p \ \) \* subject \" and "set p \ set(bn \) \ set(bn(p \ \))" shows "\ \ P = (p \ \) \ (p \ P)" using assms apply(rule_tac \=\ in actionCases) apply(simp only: eqvts bn.simps) apply simp apply(simp add: boundOutputChainAlpha'' residualInject) by simp lemma Par1: fixes \ :: 'b and \\<^sub>Q :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and A\<^sub>Q :: "name list" and Q :: "('a, 'b, 'c) psi" assumes Trans: "\ \ \\<^sub>Q \ P \\ \ P'" and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "bn \ \* Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* P" and "A\<^sub>Q \* \" shows "\ \ P \ Q \\ \ (P' \ Q)" proof - { fix \ :: 'b and \\<^sub>Q :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and A\<^sub>Q :: "name list" and Q :: "('a, 'b, 'c) psi" assume "\ \ \\<^sub>Q \ P \\ \ P'" and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "bn \ \* Q" and "bn \ \* subject \" and "A\<^sub>Q \* \" and "A\<^sub>Q \* P" and "A\<^sub>Q \* \" and "distinct A\<^sub>Q" have "\ \ P \ Q \\ \ (P' \ Q)" proof - from \\ \ \\<^sub>Q \ P \\ \ P'\ have "distinct(bn \)" by(rule boundOutputDistinct) obtain q::"name prm" where "bn(q \ \) \* \" and "bn(q \ \) \* P" and "bn(q \ \) \* Q" and "bn(q \ \) \* \" and "bn(q \ \) \* A\<^sub>Q" and "bn(q \ \) \* P'" and "bn(q \ \) \* \\<^sub>Q" and Sq: "(set q) \ (set (bn \)) \ (set(bn(q \ \)))" by(rule_tac xvec="bn \" and c="(\, P, Q, \, A\<^sub>Q, \\<^sub>Q, P')" in name_list_avoiding) (auto simp add: eqvts) obtain p::"name prm" where "(p \ A\<^sub>Q) \* \" and "(p \ A\<^sub>Q) \* P" and "(p \ A\<^sub>Q) \* Q" and "(p \ A\<^sub>Q) \* \" and "(p \ A\<^sub>Q) \* \" and "(p \ A\<^sub>Q) \* (q \ \)" and "(p \ A\<^sub>Q) \* P'" and "(p \ A\<^sub>Q) \* (q \ P')" and "(p \ A\<^sub>Q) \* \\<^sub>Q" and Sp: "(set p) \ (set A\<^sub>Q) \ (set(p \ A\<^sub>Q))" by(rule_tac xvec=A\<^sub>Q and c="(\, P, Q, \, bn \, q \ \, P', (q \ P'), \\<^sub>Q)" in name_list_avoiding) auto from \distinct(bn \)\ have "distinct(bn(q \ \))" by(rule_tac \=\ in actionCases) (auto simp add: eqvts) from \A\<^sub>Q \* \\ \bn(q \ \) \* A\<^sub>Q\ Sq have "A\<^sub>Q \* (q \ \)" apply(rule_tac \=\ in actionCases) apply(simp only: bn.simps eqvts, simp) apply(simp add: freshChainSimps) by simp from \bn \ \* subject \\ have "(q \ (bn \)) \* (q \ (subject \))" by(simp add: fresh_star_bij) hence "bn(q \ \) \* subject(q \ \)" by(simp add: eqvts) from \\ \ \\<^sub>Q \ P \\ \ P'\ \bn(q \ \) \* \\ \bn(q \ \) \* P'\ \bn \ \* (subject \)\ Sq have Trans: "\ \ \\<^sub>Q \ P \(q \ \) \ (q \ P')" by(force simp add: residualAlpha) hence "A\<^sub>Q \* (q \ P')" using \bn(q \ \) \* subject(q \ \)\ \distinct(bn(q \ \))\ \A\<^sub>Q \* P\ \A\<^sub>Q \* (q \ \)\ by(auto intro: freeFreshChainDerivative) from Trans have "(p \ (\ \ \\<^sub>Q)) \ (p \ P) \p \ ((q \ \) \ (q \ P'))" by(rule semantics.eqvt) with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* (q \ \)\ \(p \ A\<^sub>Q) \* (q \ \)\ \A\<^sub>Q \* (q \ P')\ \(p \ A\<^sub>Q) \* \\ \(p \ A\<^sub>Q) \* P\ \(p \ A\<^sub>Q) \* (q \ P')\ Sp have "\ \ (p \ \\<^sub>Q) \ P \(q \ \) \ (q \ P')" by(simp add: eqvts) moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \(p \ A\<^sub>Q) \* \\<^sub>Q\ Sp have "extractFrame Q = \(p \ A\<^sub>Q), (p \ \\<^sub>Q)\" by(simp add: frameChainAlpha' eqvts) moreover from \(bn(q \ \)) \* \\<^sub>Q\ \(bn(q \ \)) \* A\<^sub>Q\ \(p \ A\<^sub>Q) \* (q \ \)\ Sp have "(bn(q \ \)) \* (p \ \\<^sub>Q)" by(simp add: freshAlphaPerm) moreover from \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>Q)" by simp ultimately have "\ \ P \ Q \(q \ \) \ ((q \ P') \ Q)" using \(p \ A\<^sub>Q) \* P\ \(p \ A\<^sub>Q) \* Q\ \(p \ A\<^sub>Q) \* \\ \(p \ A\<^sub>Q) \* (q \ \)\ \(p \ A\<^sub>Q) \* (q \ P')\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q\ \(bn(q \ \)) \* P\ \(bn(q \ \)) \* (subject (q \ \))\ \distinct(bn(q \ \))\ by(rule_tac cPar1) thus ?thesis using \bn(q \ \) \* \\ \bn(q \ \) \* P'\ \bn \ \* subject \\ \bn(q \ \) \* Q\ \bn \ \* Q\ Sq by(force simp add: residualAlpha) qed } note Goal = this from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* \\ obtain A\<^sub>Q' where FrQ: "extractFrame Q = \A\<^sub>Q', \\<^sub>Q\" and "distinct A\<^sub>Q'" and "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* \" by(rule_tac C="(\, P, \)" in distinctFrame) auto show ?thesis proof(induct rule: actionCases[where \=\]) case(cInput M N) from Trans FrQ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \A\<^sub>Q' \* \\ \distinct A\<^sub>Q'\ \bn \ \* Q\ show ?case using \\ = M\N\\ by(force intro: Goal) next case cTau from Trans FrQ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \A\<^sub>Q' \* \\ \distinct A\<^sub>Q'\ \bn \ \* Q\ show ?case using \\ = \\ by(force intro: Goal) next case(cOutput M xvec N) from \\ = M\\*xvec\\N\\ \A\<^sub>Q' \* \\ \bn \ \* Q\ have "xvec \* A\<^sub>Q'" and "xvec \* Q" by simp+ obtain p where "(p \ xvec) \* N" and "(p \ xvec) \* P'" and "(p \ xvec) \* Q" and "(p \ xvec) \* M" and "(p \ xvec) \* A\<^sub>Q'" and S: "set p \ set xvec \ set(p \ xvec)" by(rule_tac xvec=xvec and c="(N, P', Q, M, A\<^sub>Q')" in name_list_avoiding) auto from Trans \\=M\\*xvec\\N\\ have "\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" by simp with \(p \ xvec) \* N\ \(p \ xvec) \* P'\ S have "\ \ \\<^sub>Q \ P \M\\*(p \ xvec)\\(p \ N)\ \ (p \ P')" by(simp add: boundOutputChainAlpha'' create_residual.simps) moreover from \xvec \* A\<^sub>Q'\ \(p \ xvec) \* A\<^sub>Q'\ \A\<^sub>Q' \* \\ S have "A\<^sub>Q' \* (p \ \)" by(simp add: freshChainSimps del: actionFreshChain) ultimately have "\ \ P \ Q \M\\*(p \ xvec)\\(p \ N)\ \ (p \ P') \ Q" using FrQ \A\<^sub>Q' \* \\ \A\<^sub>Q' \* P\ \distinct A\<^sub>Q'\ \(p \ xvec) \* Q\ \A\<^sub>Q' \* \\ \(p \ xvec) \* M\ \\ = M\\*xvec\\N\\ by(force intro: Goal) with \(p \ xvec) \* N\ \(p \ xvec) \* P'\ \(p \ xvec) \* Q\ \xvec \* Q\ S \\ = M\\*xvec\\N\\ show ?case by(simp add: boundOutputChainAlpha'' eqvts create_residual.simps) qed qed lemma Par2: fixes \ :: 'b and \\<^sub>P :: 'b and Q :: "('a, 'b, 'c) psi" and \ :: "'a action" and Q' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and P :: "('a, 'b, 'c) psi" assumes Trans: "\ \ \\<^sub>P \ Q \\ \ Q'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "bn \ \* P" and "A\<^sub>P \* \" and "A\<^sub>P \* Q" and "A\<^sub>P \* \" shows "\ \ P \ Q \\ \ (P \ Q')" proof - { fix \ :: 'b and \\<^sub>P :: 'b and Q :: "('a, 'b, 'c) psi" and \ :: "'a action" and Q' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and P :: "('a, 'b, 'c) psi" assume "\ \ \\<^sub>P \ Q \\ \ Q'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "bn \ \* P" and "bn \ \* subject \" and "A\<^sub>P \* \" and "A\<^sub>P \* Q" and "A\<^sub>P \* \" and "distinct A\<^sub>P" have "\ \ P \ Q \\ \ (P \ Q')" proof - from \\ \ \\<^sub>P \ Q \\ \ Q'\ have "distinct(bn \)" by(rule boundOutputDistinct) obtain q::"name prm" where "bn(q \ \) \* \" and "bn(q \ \) \* P" and "bn(q \ \) \* Q" and "bn(q \ \) \* \" and "bn(q \ \) \* A\<^sub>P" and "bn(q \ \) \* Q'" and "bn(q \ \) \* \\<^sub>P" and Sq: "(set q) \ (set (bn \)) \ (set(bn(q \ \)))" by(rule_tac xvec="bn \" and c="(\, P, Q, \, A\<^sub>P, \\<^sub>P, Q')" in name_list_avoiding) (auto simp add: eqvts) obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* (q \ \)" and "(p \ A\<^sub>P) \* Q'" and "(p \ A\<^sub>P) \* (q \ Q')" and "(p \ A\<^sub>P) \* \\<^sub>P" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" by(rule_tac xvec=A\<^sub>P and c="(\, P, Q, \, q \ \, Q', (q \ Q'), \\<^sub>P)" in name_list_avoiding) auto from \distinct(bn \)\ have "distinct(bn(q \ \))" by(rule_tac \=\ in actionCases) (auto simp add: eqvts) from \A\<^sub>P \* \\ \bn(q \ \) \* A\<^sub>P\ Sq have "A\<^sub>P \* (q \ \)" apply(rule_tac \=\ in actionCases) apply(simp only: bn.simps eqvts, simp) apply(simp add: freshChainSimps) by simp from \bn \ \* subject \\ have "(q \ (bn \)) \* (q \ (subject \))" by(simp add: fresh_star_bij) hence "bn(q \ \) \* subject(q \ \)" by(simp add: eqvts) from \\ \ \\<^sub>P \ Q \\ \ Q'\ \bn(q \ \) \* \\ \bn(q \ \) \* Q'\ \bn \ \* (subject \)\ Sq have Trans: "\ \ \\<^sub>P \ Q \(q \ \) \ (q \ Q')" by(force simp add: residualAlpha) hence "A\<^sub>P \* (q \ Q')" using \bn(q \ \) \* subject(q \ \)\ \distinct(bn(q \ \))\ \A\<^sub>P \* Q\ \A\<^sub>P \* (q \ \)\ by(auto intro: freeFreshChainDerivative) from Trans have "(p \ (\ \ \\<^sub>P)) \ (p \ Q) \p \ ((q \ \) \ (q \ Q'))" by(rule semantics.eqvt) with \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>P \* (q \ \)\ \(p \ A\<^sub>P) \* (q \ \)\ \A\<^sub>P \* (q \ Q')\ \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* (q \ Q')\ Sp have "\ \ (p \ \\<^sub>P) \ Q \(q \ \) \ (q \ Q')" by(simp add: eqvts) moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \(p \ A\<^sub>P) \* \\<^sub>P\ Sp have "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" by(simp add: frameChainAlpha' eqvts) moreover from \(bn(q \ \)) \* \\<^sub>P\ \(bn(q \ \)) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ \)\ Sp have "(bn(q \ \)) \* (p \ \\<^sub>P)" by(simp add: freshAlphaPerm) moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp ultimately have "\ \ P \ Q \(q \ \) \ (P \ (q \ Q'))" using \(p \ A\<^sub>P) \* P\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* (q \ \)\ \(p \ A\<^sub>P) \* (q \ Q')\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q\ \(bn(q \ \)) \* P\ \(bn(q \ \)) \* (subject (q \ \))\ \distinct(bn(q \ \))\ by(rule_tac cPar2) thus ?thesis using \bn(q \ \) \* \\ \bn(q \ \) \* Q'\ \bn \ \* subject \\ \bn(q \ \) \* P\ \bn \ \* P\ Sq by(force simp add: residualAlpha) qed } note Goal = this from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ obtain A\<^sub>P' where FrP: "extractFrame P = \A\<^sub>P', \\<^sub>P\" and "distinct A\<^sub>P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* Q" and "A\<^sub>P' \* \" by(rule_tac C="(\, Q, \)" in distinctFrame) auto show ?thesis proof(induct rule: actionCases[where \=\]) case(cInput M N) from Trans FrP \A\<^sub>P' \* \\ \A\<^sub>P' \* Q\ \A\<^sub>P' \* \\ \distinct A\<^sub>P'\ \bn \ \* P\ show ?case using \\ = M\N\\ by(force intro: Goal) next case cTau from Trans FrP \A\<^sub>P' \* \\ \A\<^sub>P' \* Q\ \A\<^sub>P' \* \\ \distinct A\<^sub>P'\ \bn \ \* P\ show ?case using \\ = \\ by(force intro: Goal) next case(cOutput M xvec N) from \\ = M\\*xvec\\N\\ \A\<^sub>P' \* \\ \bn \ \* P\ have "xvec \* A\<^sub>P'" and "xvec \* P" by simp+ obtain p where "(p \ xvec) \* N" and "(p \ xvec) \* Q'" and "(p \ xvec) \* P" and "(p \ xvec) \* M" and "(p \ xvec) \* A\<^sub>P'" and S: "set p \ set xvec \ set(p \ xvec)" by(rule_tac xvec=xvec and c="(N, Q', P, M, A\<^sub>P')" in name_list_avoiding) auto from Trans \\=M\\*xvec\\N\\ have "\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'" by simp with \(p \ xvec) \* N\ \(p \ xvec) \* Q'\ S have "\ \ \\<^sub>P \ Q \M\\*(p \ xvec)\\(p \ N)\ \ (p \ Q')" by(simp add: boundOutputChainAlpha'' create_residual.simps) moreover from \xvec \* A\<^sub>P'\ \(p \ xvec) \* A\<^sub>P'\ \A\<^sub>P' \* \\ S have "A\<^sub>P' \* (p \ \)" by(simp add: freshChainSimps del: actionFreshChain) ultimately have "\ \ P \ Q \M\\*(p \ xvec)\\(p \ N)\ \ P \ (p \ Q')" using FrP \A\<^sub>P' \* \\ \A\<^sub>P' \* Q\ \distinct A\<^sub>P'\ \(p \ xvec) \* P\ \A\<^sub>P' \* \\ \(p \ xvec) \* M\ \\ = M\\*xvec\\N\\ by(force intro: Goal) with \(p \ xvec) \* N\ \(p \ xvec) \* Q'\ \(p \ xvec) \* P\ \xvec \* P\ S \\ = M\\*xvec\\N\\ show ?case by(simp add: boundOutputChainAlpha'' eqvts create_residual.simps) qed qed lemma Open: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and yvec :: "name list" and N :: 'a and P' :: "('a, 'b, 'c) psi" and x :: name assumes Trans: "\ \ P \M\\*(xvec@yvec)\\N\ \ P'" and "x \ supp N" and "x \ \" and "x \ M" and "x \ xvec" and "x \ yvec" shows "\ \ \\x\P \M\\*(xvec@x#yvec)\\N\ \ P'" proof - from Trans have "distinct(xvec@yvec)" by(force dest: boundOutputDistinct) hence "xvec \* yvec" by(induct xvec) auto obtain p where "(p \ yvec) \* \" and "(p \ yvec) \* P" and "(p \ yvec) \* M" and "(p \ yvec) \* yvec" and "(p \ yvec) \* N" and "(p \ yvec) \* P'" and "x \ (p \ yvec)" and "(p \ yvec) \* xvec" and Sp: "(set p) \ (set yvec) \ (set(p \ yvec))" by(rule_tac xvec=yvec and c="(\, P, M, xvec, yvec, N, P', x)" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) obtain q where "(q \ xvec) \* \" and "(q \ xvec) \* P" and "(q \ xvec) \* M" and "(q \ xvec) \* xvec" and "(q \ xvec) \* N" and "(q \ xvec) \* P'" and "x \ (q \ xvec)" and "(q \ xvec) \* yvec" and "(q \ xvec) \* p" and "(q \ xvec) \* (p \ yvec)" and Sq: "(set q) \ (set xvec) \ (set(q \ xvec))" by(rule_tac xvec=xvec and c="(\, P, M, xvec, yvec, p \ yvec, N, P', x, p)" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) note \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ moreover from \(p \ yvec) \* N\ \(q \ xvec) \* N\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq have "((p@q) \ (xvec @ yvec)) \* N" apply(simp only: eqvts) apply(simp only: pt2[OF pt_name_inst]) by simp moreover from \(p \ yvec) \* P'\ \(q \ xvec) \* P'\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq have "((p@q) \ (xvec @ yvec)) \* P'" by(simp del: freshAlphaPerm add: eqvts pt2[OF pt_name_inst]) moreover from Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ have Spq: "set(p@q) \ set(xvec@yvec) \ set((p@q) \ (xvec@yvec))" by(simp add: pt2[OF pt_name_inst] eqvts) blast ultimately have "\ \ P \M\\*((p@q) \ (xvec@yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" apply(simp add: create_residual.simps) by(erule_tac rev_mp) (subst boundOutputChainAlpha, auto) with Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ have "\ \ P \M\\*((q \ xvec)@(p \ yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" by(simp add: eqvts pt2[OF pt_name_inst] del: freshAlphaPerm) moreover from \x \ supp N\ have "((p@q) \ x) \ (p@q) \ (supp N)" by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) with \x \ xvec\ \x \ yvec\ \x \ (q \ xvec)\ \x \ (p \ yvec)\ Sp Sq have "x \ supp((p@q)\ N)" by(simp add: eqvts pt2[OF pt_name_inst]) moreover from \distinct(xvec@yvec)\ have "distinct(q \ xvec)" and "distinct(p \ yvec)" by auto moreover note \x \ (q \ xvec)\ \x \ (p \ yvec)\ \x \ M\ \x \ \\ \(q \ xvec) \* \\ \(q \ xvec) \* P\ \(q \ xvec) \* M\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* \\ \(p \ yvec) \* P\ \(p \ yvec) \* M\ \distinct(q \ xvec)\ ultimately have "\ \ \\x\P \M\\*((q \ xvec)@x#(p \ yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" by(rule_tac cOpen) with \x \ xvec\ \x \ yvec\ \x \ (q \ xvec)\ \x \ (p \ yvec)\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq have "\ \ \\x\P \M\\*((p@q) \ (xvec@x#yvec))\\((p@q) \ N)\ \ ((p@q) \ P')" by(simp add: eqvts pt2[OF pt_name_inst] del: freshAlphaPerm) thus ?thesis using \((p@q) \ (xvec @ yvec)) \* N\ \((p@q) \ (xvec @ yvec)) \* P'\ Spq apply(simp add: create_residual.simps) by(erule_tac rev_mp) (subst boundOutputChainAlpha, auto) qed lemma Scope: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and x :: name assumes "\ \ P \\ \ P'" and "x \ \" and "x \ \" shows "\ \ \\x\P \\ \ \\x\P'" proof - { fix \ P M xvec N P' x assume "\ \ P \M\\*xvec\\N\ \ P'" and "(x::name) \ \" and "x \ M" and "x \ xvec" and "x \ N" obtain p::"name prm" where "(p \ xvec) \* \" and "(p \ xvec) \* P" and "(p \ xvec) \* M" and "(p \ xvec) \* xvec" and "(p \ xvec) \* N" and "(p \ xvec) \* P'" and "x \ (p \ xvec)" and S: "(set p) \ (set xvec) \ (set(p \ xvec))" by(rule_tac xvec=xvec and c="(\, P, M, xvec, N, P', x)" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) from \\ \ P \M\\*xvec\\N\ \ P'\ \(p \ xvec) \* N\ \(p \ xvec) \* P'\ S have "\ \ P \M\\*(p \ xvec)\\(p \ N)\ \ (p \ P')" by(simp add: boundOutputChainAlpha'' create_residual.simps) moreover hence "distinct(p \ xvec)" by(force dest: boundOutputDistinct) moreover note \x \ \\ \x \ M\ \x \ (p \ xvec)\ moreover from \x \ xvec\ \x \ p \ xvec\ \x \ N\ S have "x \ (p \ N)" by(simp add: fresh_left del: freshAlphaSwap) ultimately have "\ \ \\x\P \M\\*(p \ xvec)\\(p \ N)\ \ \\x\(p \ P')" using \(p \ xvec) \* \\ \(p \ xvec) \* P\ \(p \ xvec) \* M\ by(rule_tac cScope) auto moreover from \x \ xvec\ \x \ p \ xvec\ S have "p \ x = x" by simp ultimately have "\ \ \\x\P \M\\*(p \ xvec)\\(p \ N)\ \ (p \ (\\x\P'))" by simp moreover from \(p \ xvec) \* P'\ \x \ xvec\ \x \ (p \ xvec)\ have "(p \ xvec) \* \\x\P'" by(simp add: abs_fresh_star) ultimately have"\ \ \\x\P \M\\*xvec\\N\ \ \\x\P'" using \(p \ xvec) \* N\ S by(simp add: boundOutputChainAlpha'' create_residual.simps) } note Goal = this show ?thesis proof(induct rule: actionCases[where \=\]) case(cInput M N) with assms show ?case by(force intro: cScope) next case(cOutput M xvec N) with assms show ?case by(force intro: Goal) next case cTau with assms show ?case by(force intro: cScope) qed qed lemma inputSwapFrameSubject: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and x :: name and y :: name assumes "\ \ P \M\N\ \ P'" and "x \ P" and "y \ P" shows "([(x, y)] \ \) \ P \ ([(x, y)] \ M)\N\ \ P'" using assms proof(nominal_induct avoiding: x y rule: inputInduct) case(cInput \ M K xvec N Tvec P x y) from \x \ M\\*xvec N\.P\ have "x \ M" by simp from \y \ M\\*xvec N\.P\ have "y \ M" by simp from \\ \ M \ K\ have "([(x, y)] \ \) \ ([(x, y)] \ M) \ ([(x, y)] \ K)" by(rule chanEqClosed) with \x \ M\ \y \ M\ have "([(x, y)] \ \) \ M \ ([(x, y)] \ K)" by(simp) thus ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ by(rule Input) next case(cCase \ P M N P' \ Cs x y) from \x \ Cases Cs\ \y \ Cases Cs\ \(\, P) mem Cs\ have "x \ \" and "x \ P" and "y \ \" and "y \ P" by(auto dest: memFresh) from \x \ P\ \y \ P\ have "([(x ,y)] \ \) \ P \ ([(x, y)] \ M)\N\ \ P'" by(rule cCase) moreover note \(\, P) mem Cs\ moreover from \\ \ \\ have "([(x, y)] \ \) \ ([(x, y)] \ \)" by(rule statClosed) with \x \ \\ \y \ \\ have "([(x, y)] \ \) \ \" by simp ultimately show ?case using \guarded P\ by(rule Case) next case(cPar1 \ \\<^sub>Q P M N P' A\<^sub>Q Q x y) from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ (\ \ \\<^sub>Q)) \ P \([(x, y)] \ M)\N\ \ P'\ have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>Q) \ P \([(x, y)] \ M)\N\ \ P'" by(simp add: eqvts) moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "([(x, y)] \ (extractFrame Q)) = ([(x, y)] \ \A\<^sub>Q, \\<^sub>Q\)" by simp with \A\<^sub>Q \* x\ \x \ Q\ \A\<^sub>Q \* y\ \y \ Q\ have "\A\<^sub>Q, ([(x, y)] \ \\<^sub>Q)\ = extractFrame Q" by(simp add: eqvts) moreover from \A\<^sub>Q \* \\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \)" by simp moreover from \A\<^sub>Q \* M\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ M)" by simp ultimately show ?case using \A\<^sub>Q \* P\ \A\<^sub>Q \* N\ by(rule_tac Par1) auto next case(cPar2 \ \\<^sub>P Q M N Q' A\<^sub>P P x y) from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ from \x \ Q\ \y \ Q\ \\x y. \x \ Q; y \ Q\ \ ([(x, y)] \ (\ \ \\<^sub>P)) \ Q \([(x, y)] \ M)\N\ \ Q'\ have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>P) \ Q \([(x, y)] \ M)\N\ \ Q'" by(simp add: eqvts) moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ have "([(x, y)] \ (extractFrame P)) = ([(x, y)] \ \A\<^sub>P, \\<^sub>P\)" by simp with \A\<^sub>P \* x\ \x \ P\ \A\<^sub>P \* y\ \y \ P\ have "\A\<^sub>P, ([(x, y)] \ \\<^sub>P)\ = extractFrame P" by(simp add: eqvts) moreover from \A\<^sub>P \* \\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \)" by simp moreover from \A\<^sub>P \* M\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ M)" by simp ultimately show ?case using \A\<^sub>P \* Q\ \A\<^sub>P \* N\ by(rule_tac Par2) auto next case(cScope \ P M N P' z x y) from \x \ \\z\P\ \z \ x\ have "x \ P" by(simp add: abs_fresh) from \y \ \\z\P\ \z \ y\ have "y \ P" by(simp add: abs_fresh) from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ \) \ P \([(x, y)] \ M)\N\ \ P'\ have "([(x, y)] \ \) \ P \([(x, y)] \ M)\N\ \ P'" by simp moreover with \z \ \\ have "([(x, y)] \ z) \ [(x, y)] \ \" by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) with \z \ x\ \z \ y\ have "z \ [(x, y)] \ \" by simp moreover with \z \ M\ have "([(x, y)] \ z) \ [(x, y)] \ M" by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) with \z \ x\ \z \ y\ have "z \ [(x, y)] \ M" by simp ultimately show ?case using \z \ N\ by(rule_tac Scope) (assumption | simp)+ next case(cBang \ P M N P' x y) thus ?case by(force intro: Bang) qed lemma inputPermFrameSubject: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and p :: "name prm" and Xs :: "name set" and Ys :: "name set" assumes "\ \ P \M\N\ \ P'" and S: "set p \ Xs \ Ys" and "Xs \* P" and "Ys \* P" shows "(p \ \) \ P \ (p \ M)\N\ \ P'" using S proof(induct p) case Nil from \\ \ P \M\N\ \ P'\ show ?case by simp next case(Cons a p) from \set(a#p) \ Xs \ Ys\ have "set p \ Xs \ Ys" by auto with \set p \ Xs \ Ys \ (p \ \) \ P \ (p \ M)\N\ \ P'\ have Trans: "(p \ \) \ P \ (p \ M)\N\ \ P'" by simp from \set(a#p) \ Xs \ Ys\ show ?case proof(cases a, clarsimp) fix a b assume "a \ Xs" and "b \ Ys" with \Xs \* P\ \Ys \* P\ have "a \ P" and "b \ P" by(auto simp add: fresh_star_def) with Trans show "([(a, b)] \ p \ \) \ P \ ([(a, b)] \ p \ M)\N\ \ P'" by(rule inputSwapFrameSubject) qed qed lemma inputSwapSubject: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and x :: name and y :: name assumes "\ \ P \M\N\ \ P'" and "x \ P" and "y \ P" and "x \ \" and "y \ \" shows "\ \ P \ ([(x, y)] \ M)\N\ \ P'" proof - from \\ \ P \M\N\ \ P'\ \x \ P\ \y \ P\ have "([(x, y)] \ \) \ P \([(x, y)] \ M)\N\ \ P'" by(rule inputSwapFrameSubject) with \x \ \\ \y \ \\ show ?thesis by simp qed lemma inputPermSubject: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and p :: "name prm" and Xs :: "name set" and Ys :: "name set" assumes "\ \ P \M\N\ \ P'" and S: "set p \ Xs \ Ys" and "Xs \* P" and "Ys \* P" and "Xs \* \" and "Ys \* \" shows "\ \ P \ (p \ M)\N\ \ P'" proof - from \\ \ P \M\N\ \ P'\ S \Xs \* P\ \Ys \* P\ have "(p \ \) \ P \(p \ M)\N\ \ P'" by(rule inputPermFrameSubject) with \Xs \* \\ \Ys \* \\ S show ?thesis by simp qed lemma inputSwapFrame: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and x :: name and y :: name assumes "\ \ P \M\N\ \ P'" and "x \ P" and "y \ P" and "x \ M" and "y \ M" shows "([(x, y)] \ \) \ P \ M\N\ \ P'" proof - from \\ \ P \M\N\ \ P'\ \x \ P\ \y \ P\ have "([(x, y)] \ \) \ P \([(x, y)] \ M)\N\ \ P'" by(rule inputSwapFrameSubject) with \x \ M\ \y \ M\ show ?thesis by simp qed lemma inputPermFrame: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and p :: "name prm" and Xs :: "name set" and Ys :: "name set" assumes "\ \ P \M\N\ \ P'" and S: "set p \ Xs \ Ys" and "Xs \* P" and "Ys \* P" and "Xs \* M" and "Ys \* M" shows "(p \ \) \ P \ M\N\ \ P'" proof - from \\ \ P \M\N\ \ P'\ S \Xs \* P\ \Ys \* P\ have "(p \ \) \ P \(p \ M)\N\ \ P'" by(rule inputPermFrameSubject) with \Xs \* M\ \Ys \* M\ S show ?thesis by simp qed lemma inputAlpha: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and p :: "name prm" and xvec :: "name list" assumes "\ \ P \M\N\ \ P'" and "set p \ (set xvec) \ (set (p \ xvec))" and "distinctPerm p" and "xvec \* P" and "(p \ xvec) \* P" shows "\ \ P \M\(p \ N)\ \ (p \ P')" proof - from \\ \ P \M\N\ \ P'\ \set p \ (set xvec) \ (set (p \ xvec))\ \xvec \* P\ \(p \ xvec) \* P\ have "(p \ \) \ P \(p \ M)\N\ \ P'" by(rule_tac inputPermFrameSubject) auto hence "(p \ p \ \) \ (p \ P) \(p \ ((p \ M)\N\ \ P'))" by(rule eqvts) with \distinctPerm p\ \xvec \* P\ \(p \ xvec) \* P\ \set p \ (set xvec) \ (set (p \ xvec))\ show ?thesis by(simp add: eqvts) qed lemma frameFresh[dest]: fixes x :: name and A\<^sub>F :: "name list" and \\<^sub>F :: 'b assumes "x \ A\<^sub>F" and "x \ \A\<^sub>F, \\<^sub>F\" shows "x \ \\<^sub>F" using assms by(simp add: frameResChainFresh) (simp add: fresh_def name_list_supp) lemma outputSwapFrameSubject: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and N :: 'a and x :: name and y :: name assumes "\ \ P \M\\*xvec\\N\ \ P'" and "xvec \* M" and "x \ P" and "y \ P" shows "([(x, y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" using assms proof(nominal_induct avoiding: x y rule: outputInduct') case cAlpha thus ?case by(simp add: create_residual.simps boundOutputChainAlpha'') next case(cOutput \ M K N P x y) from \x \ M\N\.P\ have "x \ M" by simp from \y \ M\N\.P\ have "y \ M" by simp from \\ \ M \ K\ have "([(x, y)] \ \) \ ([(x, y)] \ M) \ ([(x, y)] \ K)" by(rule chanEqClosed) with \x \ M\ \y \ M\ have "([(x, y)] \ \) \ M \ ([(x, y)] \ K)" by(simp) thus ?case by(rule Output) next case(cCase \ P M xvec N P' \ Cs x y) from \x \ Cases Cs\ \y \ Cases Cs\ \(\, P) mem Cs\ have "x \ \" and "x \ P" and "y \ \" and "y \ P" by(auto dest: memFresh) from \x \ P\ \y \ P\ have "([(x ,y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" by(rule cCase) moreover note \(\, P) mem Cs\ moreover from \\ \ \\ have "([(x, y)] \ \) \ ([(x, y)] \ \)" by(rule statClosed) with \x \ \\ \y \ \\ have "([(x, y)] \ \) \ \" by simp ultimately show ?case using \guarded P\ by(rule Case) next case(cPar1 \ \\<^sub>Q P M xvec N P' A\<^sub>Q Q x y) from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ (\ \ \\<^sub>Q)) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'\ have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>Q) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" by(simp add: eqvts) moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ have "([(x, y)] \ \A\<^sub>Q, \\<^sub>Q\) = ([(x, y)] \ (extractFrame Q))" by simp with \A\<^sub>Q \* x\ \x \ Q\ \A\<^sub>Q \* y\ \y \ Q\ have "\A\<^sub>Q, ([(x, y)] \ \\<^sub>Q)\ = extractFrame Q" by(simp add: eqvts) moreover from \A\<^sub>Q \* \\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ \)" by simp moreover from \A\<^sub>Q \* M\ have "([(x, y)] \ A\<^sub>Q) \* ([(x, y)] \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>Q \* x\ \A\<^sub>Q \* y\ have "A\<^sub>Q \* ([(x, y)] \ M)" by simp ultimately show ?case using \A\<^sub>Q \* P\ \A\<^sub>Q \* N\ \xvec \* Q\ \A\<^sub>Q \* xvec\ by(rule_tac Par1) auto next case(cPar2 \ \\<^sub>P Q M xvec N Q' A\<^sub>P P x y) from \x \ P \ Q\ have "x \ P" and "x \ Q" by simp+ from \y \ P \ Q\ have "y \ P" and "y \ Q" by simp+ from \x \ Q\ \y \ Q\ \\x y. \x \ Q; y \ Q\ \ ([(x, y)] \ (\ \ \\<^sub>P)) \ Q \([(x, y)] \ M)\\*xvec\\N\ \ Q'\ have "([(x, y)] \ \) \ ([(x, y)] \ \\<^sub>P) \ Q \([(x, y)] \ M)\\*xvec\\N\ \ Q'" by(simp add: eqvts) moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ have "([(x, y)] \ \A\<^sub>P, \\<^sub>P\) = ([(x, y)] \ (extractFrame P))" by simp with \A\<^sub>P \* x\ \x \ P\ \A\<^sub>P \* y\ \y \ P\ have "\A\<^sub>P, ([(x, y)] \ \\<^sub>P)\ = extractFrame P" by(simp add: eqvts) moreover from \A\<^sub>P \* \\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ \)" by simp moreover from \A\<^sub>P \* M\ have "([(x, y)] \ A\<^sub>P) \* ([(x, y)] \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* x\ \A\<^sub>P \* y\ have "A\<^sub>P \* ([(x, y)] \ M)" by simp ultimately show ?case using \A\<^sub>P \* Q\ \A\<^sub>P \* N\ \xvec \* P\ \A\<^sub>P \* xvec\ by(rule_tac Par2) auto next case(cOpen \ P M xvec yvec N P' z x y) from \x \ \\z\P\ \z \ x\ have "x \ P" by(simp add: abs_fresh) from \y \ \\z\P\ \z \ y\ have "y \ P" by(simp add: abs_fresh) from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ \) \ P \([(x, y)] \ M)\\*(xvec@yvec)\\N\ \ P'\ have "([(x, y)] \ \) \ P \([(x, y)] \ M)\\*(xvec@yvec)\\N\ \ P'" by simp moreover with \z \ \\ have "([(x, y)] \ z) \ [(x, y)] \ \" by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) with \z \ x\ \z \ y\ have "z \ [(x, y)] \ \" by simp moreover with \z \ M\ have "([(x, y)] \ z) \ [(x, y)] \ M" by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) with \z \ x\ \z \ y\ have "z \ [(x, y)] \ M" by simp ultimately show ?case using \z \ supp N\ \z \ xvec\ \z \ yvec\ by(rule_tac Open) (assumption | simp)+ next case(cScope \ P M xvec N P' z x y) from \x \ \\z\P\ \z \ x\ have "x \ P" by(simp add: abs_fresh) from \y \ \\z\P\ \z \ y\ have "y \ P" by(simp add: abs_fresh) from \x \ P\ \y \ P\ \\x y. \x \ P; y \ P\ \ ([(x, y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'\ have "([(x, y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" by simp moreover with \z \ \\ have "([(x, y)] \ z) \ [(x, y)] \ \" by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) with \z \ x\ \z \ y\ have "z \ [(x, y)] \ \" by simp moreover with \z \ M\ have "([(x, y)] \ z) \ [(x, y)] \ M" by(simp add: pt_fresh_bij[OF pt_name_inst, OF at_name_inst]) with \z \ x\ \z \ y\ have "z \ [(x, y)] \ M" by simp ultimately show ?case using \z \ N\ \z \ xvec\ by(rule_tac Scope) (assumption | simp)+ next case(cBang \ P M B x y) thus ?case by(force intro: Bang) qed lemma outputPermFrameSubject: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and N :: 'a and P' :: "('a, 'b, 'c) psi" and p :: "name prm" and yvec :: "name list" and zvec :: "name list" assumes "\ \ P \M\\*xvec\\N\ \ P'" and S: "set p \ set yvec \ set zvec" and "yvec \* P" and "zvec \* P" shows "(p \ \) \ P \(p \ M)\\*xvec\\N\ \ P'" proof - { fix xvec N P' Xs YS assume "\ \ P \M\\*xvec\\N\ \ P'" and "xvec \* M" and "xvec \* yvec" and "xvec \* zvec" have "(p \ \) \ P \(p \ M)\\*xvec\\N\ \ P'" using S proof(induct p) case Nil from \\ \ P \M\\*xvec\\N\ \ P'\ show ?case by simp next case(Cons a p) from \set(a#p) \ set yvec \ set zvec\ have "set p \ set yvec \ set zvec" by auto then have Trans: "(p \ \) \ P \(p \ M)\\*xvec\\N\ \ P'" by(rule Cons) from \set(a#p) \ set yvec \ set zvec\ show ?case proof(cases a, clarsimp) fix x y note Trans moreover from \xvec \* yvec\ \xvec \* zvec\ \set p \ set yvec \ set zvec\ \xvec \* M\ have "xvec \* (p \ M)" by(simp add: freshChainSimps) moreover assume "x \ set yvec" and "y \ set zvec" with \yvec \* P\ \zvec \* P\ have "x \ P" and "y \ P" by(auto simp add: fresh_star_def) ultimately show "([(x, y)] \ p \ \) \ P \([(x, y)] \ p \ M)\\*xvec\\N\ \ P'" by(rule outputSwapFrameSubject) qed qed } note Goal = this obtain q::"name prm" where "(q \ xvec) \* yvec" and "(q \ xvec) \* zvec" and "(q \ xvec) \* xvec" and "(q \ xvec) \* N" and "(q \ xvec) \* P'" and "(q \ xvec) \* M" and Sq: "(set q) \ (set xvec) \ (set(q \ xvec))" by(rule_tac xvec=xvec and c="(P, xvec, yvec, zvec, N, M, P')" in name_list_avoiding) auto with \\ \ P \M\\*xvec\\N\ \ P'\ have "\ \ P \M\\*(q \ xvec)\\(q \ N)\ \ (q \ P')" by(simp add: boundOutputChainAlpha'' residualInject) hence "(p \ \) \ P \(p \ M)\\*(q \ xvec)\\(q \ N)\ \ (q \ P')" using \(q \ xvec) \* M\ \(q \ xvec) \* yvec\ \(q \ xvec) \* zvec\ by(rule Goal) with \(q \ xvec) \* N\ \(q \ xvec) \* P'\ Sq show ?thesis by(simp add: boundOutputChainAlpha'' residualInject) qed lemma outputSwapSubject: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and B :: "('a, 'b, 'c) boundOutput" and x :: name and y :: name assumes "\ \ P \M\\*xvec\\N\ \ P'" and "xvec \* M" and "x \ P" and "y \ P" and "x \ \" and "y \ \" shows "\ \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" proof - from \\ \ P \M\\*xvec\\N\ \ P'\ \xvec \* M\ \x \ P\ \y \ P\ have "([(x, y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" by(rule outputSwapFrameSubject) with \x \ \\ \y \ \\ show ?thesis by simp qed lemma outputPermSubject: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and B :: "('a, 'b, 'c) boundOutput" and p :: "name prm" and yvec :: "name list" and zvec :: "name list" assumes "\ \ P \M\\*xvec\\N\ \ P'" and S: "set p \ set yvec \ set zvec" and "yvec \* P" and "zvec \* P" and "yvec \* \" and "zvec \* \" shows "\ \ P \(p \ M)\\*xvec\\N\ \ P'" proof - from assms have "(p \ \) \ P \(p \ M)\\*xvec\\N\ \ P'" by(rule_tac outputPermFrameSubject) with S \yvec \* \\ \zvec \* \\ show ?thesis by simp qed lemma outputSwapFrame: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and B :: "('a, 'b, 'c) boundOutput" and x :: name and y :: name assumes "\ \ P \M\\*xvec\\N\ \ P'" and "xvec \* M" and "x \ P" and "y \ P" and "x \ M" and "y \ M" shows "([(x, y)] \ \) \ P \M\\*xvec\\N\ \ P'" proof - from \\ \ P \M\\*xvec\\N\ \ P'\ \xvec \* M\ \x \ P\ \y \ P\ have "([(x, y)] \ \) \ P \([(x, y)] \ M)\\*xvec\\N\ \ P'" by(rule outputSwapFrameSubject) with \x \ M\ \y \ M\ show ?thesis by simp qed lemma outputPermFrame: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and B :: "('a, 'b, 'c) boundOutput" and p :: "name prm" and yvec :: "name list" and zvec :: "name list" assumes "\ \ P \M\\*xvec\\N\ \ P'" and S: "set p \ set yvec \ set zvec" and "yvec \* P" and "zvec \* P" and "yvec \* M" and "zvec \* M" shows "(p \ \) \ P \M\\*xvec\\N\ \ P'" proof - from assms have "(p \ \) \ P \(p \ M)\\*xvec\\N\ \ P'" by(rule_tac outputPermFrameSubject) with S \yvec \* M\ \zvec \* M\ show ?thesis by simp qed lemma Comm1: fixes \ :: 'b and \\<^sub>Q :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and Q :: "('a, 'b, 'c) psi" and K :: 'a and xvec :: "name list" and Q' :: "('a, 'b, 'c) psi" and A\<^sub>Q :: "name list" assumes "\ \ \\<^sub>Q \ P \M\N\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'" and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" and "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* M" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* K" and "xvec \* P" shows "\ \ P \ Q \\ \ \\*xvec\(P' \ Q')" proof - { fix \ :: 'b and \\<^sub>Q :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and Q :: "('a, 'b, 'c) psi" and K :: 'a and xvec :: "name list" and Q' :: "('a, 'b, 'c) psi" and A\<^sub>Q :: "name list" assume "\ \ \\<^sub>Q \ P \M\N\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'" and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>Q" and "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" and "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* M" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* K" and "xvec \* P" have "\ \ P \ Q \\ \ \\*xvec\(P' \ Q')" proof - obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" and "(r \ xvec) \* K" and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" by(rule_tac xvec=xvec and c="(\, P, Q, M, K, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q')" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* K" and "(q \ A\<^sub>Q) \* (r \ N)" and "(q \ A\<^sub>Q) \* (r \ xvec)" and "(q \ A\<^sub>Q) \* (r \ Q')" and "(q \ A\<^sub>Q) \* (r \ P')" and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" and Sq: "set q \ set A\<^sub>Q \ set(q \ A\<^sub>Q)" by(rule_tac xvec=A\<^sub>Q and c="(\, P, Q, K, r \ N, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, r \ Q', r \ P')" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" and "(p \ A\<^sub>P) \* (r \ N)" and "(p \ A\<^sub>P) \* (r \ xvec)" and "(p \ A\<^sub>P) \* (r \ Q')" and "(p \ A\<^sub>P) \* (r \ P')" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* A\<^sub>Q" and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" by(rule_tac xvec=A\<^sub>P and c="(\, P, Q, M, r \ N, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, r \ Q', r \ P')" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" by(drule_tac extractFrameFreshChain) auto from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" by(drule_tac extractFrameFreshChain) auto from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* A\<^sub>P\ Sp have "(r \ xvec) \* (p \ A\<^sub>P)" by(simp add: freshChainSimps) from \\ \ \\<^sub>Q \ P \M\N\ \ P'\ Sr \distinctPerm r\ \xvec \* P\ \(r \ xvec) \* P\ have "\ \ \\<^sub>Q \ P \M\(r \ N)\ \ (r \ P')" by(rule inputAlpha) hence "(q \ (\ \ \\<^sub>Q)) \ P \(q \ M)\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ by(rule_tac inputPermFrameSubject) (assumption | simp)+ hence PTrans: "\ \ (q \ \\<^sub>Q) \ P \(q \ M)\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ by(simp add: eqvts) moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" by(simp add: frameChainAlpha) moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp moreover from \\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* Q'\ have "\ \ \\<^sub>P \ Q \K\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" by(simp add: boundOutputChainAlpha'' create_residual.simps) hence "(p \ (\ \ \\<^sub>P)) \ Q \(p \ K)\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ \(r \ xvec) \* K\ \(r \ xvec) \* A\<^sub>P\ \(r \ xvec) \* (p \ A\<^sub>P)\ by(rule_tac outputPermFrameSubject) (assumption | auto) hence QTrans: "\ \ (p \ \\<^sub>P) \ Q \(p \ K)\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ by(simp add: eqvts) moreover hence "distinct(r \ xvec)" by(force dest: boundOutputDistinct) moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" by(simp add: frameChainAlpha) moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp moreover from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(p \ q \ (\ \ \\<^sub>P \ \\<^sub>Q)) \ (p \ q \ M) \ (p \ q \ K)" by(rule_tac chanEqClosed)+ with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ \A\<^sub>Q \* \\<^sub>P\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \A\<^sub>P \* \\<^sub>Q\ \(p \ A\<^sub>P) \* \\<^sub>Q\ \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \A\<^sub>Q \* K\ \(q \ A\<^sub>Q) \* K\ \A\<^sub>P \* A\<^sub>Q\ \(p \ A\<^sub>P) \* A\<^sub>Q\ Sp Sq have "\ \ (p \ \\<^sub>P) \ (q \ \\<^sub>Q) \ (q \ M) \ (p \ K)" by(simp add: eqvts freshChainSimps) moreover note \(p \ A\<^sub>P) \* \\ moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* P\ moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* M\ Sq have "(p \ A\<^sub>P) \* (q \ M)" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* (r \ N)\ \(p \ A\<^sub>P) \* (r \ P')\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* (r \ Q')\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: freshChainSimps) moreover note \(q \ A\<^sub>Q) \* P\ \(q \ A\<^sub>Q) \* (r \ N)\\(q \ A\<^sub>Q) \* (r \ P')\ \(q \ A\<^sub>Q) \* Q\ moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* K\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ K)" by(simp add: freshChainSimps) moreover note \(q \ A\<^sub>Q) \* (r \ Q')\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" by(simp add: freshChainSimps) moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" by(simp add: freshChainSimps) moreover note \(r \ xvec) \* P\ moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* M\ Sq have "(r \ xvec) \* (q \ M)" by(simp add: freshChainSimps) moreover note \(r \ xvec) \* Q\ moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* K\ Sp have "(r \ xvec) \* (p \ K)" by(simp add: freshChainSimps) ultimately have "\ \ P \ Q \\ \ \\*(r \ xvec)\((r \ P') \ (r \ Q'))" by(rule_tac cComm1) with \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ Sr show ?thesis by(subst resChainAlpha) auto qed } note Goal = this note \\ \ \\<^sub>Q \ P \M\N\ \ P'\ \\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ obtain A\<^sub>P' where "extractFrame P = \A\<^sub>P', \\<^sub>P\" and "distinct A\<^sub>P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* P" and "A\<^sub>P' \* Q" and "A\<^sub>P' \* M" and "A\<^sub>P' \* A\<^sub>Q" by(rule_tac C="(\, P, Q, M, A\<^sub>Q)" in distinctFrame) auto moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>P' \* A\<^sub>Q\ obtain A\<^sub>Q' where "extractFrame Q = \A\<^sub>Q', \\<^sub>Q\" and "distinct A\<^sub>Q'" and "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q' \* K" and "A\<^sub>P' \* A\<^sub>Q'" apply(rule_tac C="(\, P, Q, K, A\<^sub>P')" in distinctFrame) by auto ultimately show ?thesis using \xvec \* P\ by(rule_tac Goal) qed lemma Comm2: fixes \ :: 'b and \\<^sub>Q :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and N :: 'a and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and Q :: "('a, 'b, 'c) psi" and K :: 'a and Q' :: "('a, 'b, 'c) psi" and A\<^sub>Q :: "name list" assumes "\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "\ \ \\<^sub>P \ Q \K\N\ \ Q'" and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" and "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* M" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* K" and "xvec \* Q" shows "\ \ P \ Q \\ \ \\*xvec\(P' \ Q')" proof - { fix \ :: 'b and \\<^sub>Q :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and N :: 'a and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and Q :: "('a, 'b, 'c) psi" and K :: 'a and Q' :: "('a, 'b, 'c) psi" and A\<^sub>Q :: "name list" assume "\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "\ \ \\<^sub>P \ Q \K\N\ \ Q'" and "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>Q" and "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" and "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* M" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* K" and "xvec \* Q" have "\ \ P \ Q \\ \ \\*xvec\(P' \ Q')" proof - obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" and "(r \ xvec) \* K" and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" by(rule_tac xvec=xvec and c="(\, P, Q, M, K, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q')" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* K" and "(q \ A\<^sub>Q) \* (r \ N)" and "(q \ A\<^sub>Q) \* (r \ xvec)" and "(q \ A\<^sub>Q) \* (r \ Q')" and "(q \ A\<^sub>Q) \* (r \ P')" and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" and Sq: "set q \ set A\<^sub>Q \ set(q \ A\<^sub>Q)" by(rule_tac xvec=A\<^sub>Q and c="(\, P, Q, K, r \ N, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, r \ Q', r \ P')" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" and "(p \ A\<^sub>P) \* (r \ N)" and "(p \ A\<^sub>P) \* (r \ xvec)" and "(p \ A\<^sub>P) \* (r \ Q')" and "(p \ A\<^sub>P) \* (r \ P')" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* A\<^sub>Q" and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" by(rule_tac xvec=A\<^sub>P and c="(\, P, Q, M, r \ N, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, r \ Q', r \ P')" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" by(drule_tac extractFrameFreshChain) auto from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" by(drule_tac extractFrameFreshChain) auto from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* A\<^sub>Q\ Sq have "(r \ xvec) \* (q \ A\<^sub>Q)" by(simp add: freshChainSimps) from \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* P'\ have "\ \ \\<^sub>Q \ P \M\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" by(simp add: boundOutputChainAlpha'' create_residual.simps) hence "(q \ (\ \ \\<^sub>Q)) \ P \(q \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ \(r \ xvec) \* M\ \(r \ xvec) \* A\<^sub>Q\ \(r \ xvec) \* (q \ A\<^sub>Q)\ by(rule_tac outputPermFrameSubject) (assumption | auto) hence PTrans: "\ \ (q \ \\<^sub>Q) \ P \(q \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ by(simp add: eqvts) moreover hence "distinct(r \ xvec)" by(force dest: boundOutputDistinct) moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" by(simp add: frameChainAlpha) moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp moreover from \\ \ \\<^sub>P \ Q \K\N\ \ Q'\ Sr \distinctPerm r\ \xvec \* Q\ \(r \ xvec) \* Q\ have "\ \ \\<^sub>P \ Q \K\(r \ N)\ \ (r \ Q')" by(rule inputAlpha) hence "(p \ (\ \ \\<^sub>P)) \ Q \(p \ K)\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ by(rule_tac inputPermFrameSubject) (assumption | simp)+ hence QTrans: "\ \ (p \ \\<^sub>P) \ Q \(p \ K)\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ by(simp add: eqvts) moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" by(simp add: frameChainAlpha) moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp moreover from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(p \ q \ (\ \ \\<^sub>P \ \\<^sub>Q)) \ (p \ q \ M) \ (p \ q \ K)" by(rule_tac chanEqClosed)+ with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ \A\<^sub>Q \* \\<^sub>P\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \A\<^sub>P \* \\<^sub>Q\ \(p \ A\<^sub>P) \* \\<^sub>Q\ \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \A\<^sub>Q \* K\ \(q \ A\<^sub>Q) \* K\ \A\<^sub>P \* A\<^sub>Q\ \(p \ A\<^sub>P) \* A\<^sub>Q\ Sp Sq have "\ \ (p \ \\<^sub>P) \ (q \ \\<^sub>Q) \ (q \ M) \ (p \ K)" by(simp add: eqvts freshChainSimps) moreover note \(p \ A\<^sub>P) \* \\ moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* P\ moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* M\ Sq have "(p \ A\<^sub>P) \* (q \ M)" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* (r \ N)\ \(p \ A\<^sub>P) \* (r \ P')\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* (r \ Q')\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: freshChainSimps) moreover note \(q \ A\<^sub>Q) \* P\ \(q \ A\<^sub>Q) \* (r \ N)\\(q \ A\<^sub>Q) \* (r \ P')\ \(q \ A\<^sub>Q) \* Q\ moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* K\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ K)" by(simp add: freshChainSimps) moreover note \(q \ A\<^sub>Q) \* (r \ Q')\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" by(simp add: freshChainSimps) moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" by(simp add: freshChainSimps) moreover note \(r \ xvec) \* P\ moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* M\ Sq have "(r \ xvec) \* (q \ M)" by(simp add: freshChainSimps) moreover note \(r \ xvec) \* Q\ moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* K\ Sp have "(r \ xvec) \* (p \ K)" by(simp add: freshChainSimps) ultimately have "\ \ P \ Q \\ \ \\*(r \ xvec)\((r \ P') \ (r \ Q'))" by(rule_tac cComm2) with \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ Sr show ?thesis by(subst resChainAlpha) auto qed } note Goal = this note \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ \\ \ \\<^sub>P \ Q \K\N\ \ Q'\ \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* M\ \A\<^sub>P \* A\<^sub>Q\ obtain A\<^sub>P' where "extractFrame P = \A\<^sub>P', \\<^sub>P\" and "distinct A\<^sub>P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* P" and "A\<^sub>P' \* Q" and "A\<^sub>P' \* M" and "A\<^sub>P' \* A\<^sub>Q" by(rule_tac C="(\, P, Q, M, A\<^sub>Q)" in distinctFrame) auto moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>P' \* A\<^sub>Q\ obtain A\<^sub>Q' where "extractFrame Q = \A\<^sub>Q', \\<^sub>Q\" and "distinct A\<^sub>Q'" and "A\<^sub>Q' \* \" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Q" and "A\<^sub>Q' \* K" and "A\<^sub>P' \* A\<^sub>Q'" by(rule_tac C="(\, P, Q, K, A\<^sub>P')" in distinctFrame) auto ultimately show ?thesis using \xvec \* Q\ by(rule_tac Goal) qed lemma semanticsCasesAux[consumes 1, case_names cInput cOutput cCase cPar1 cPar2 cComm1 cComm2 cOpen cScope cBang]: fixes \ :: 'b and cP :: "('a, 'b, 'c) psi" and cRs :: "('a, 'b, 'c) residual" and C :: "'d::fs_name" and x :: name assumes "\ \ cP \ cRs" and rInput: "\M K xvec N Tvec P. \cP = M\\*xvec N\.P; cRs = K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]; \ \ M \ K; distinct xvec; set xvec \ supp N; length xvec=length Tvec; xvec \* Tvec; xvec \* \; xvec \* M; xvec \* K; xvec \* C\ \ Prop" and rOutput: "\M K N P. \cP = M\N\.P; cRs = K\N\ \ P; \ \ M \ K\ \ Prop" and rCase: "\Cs P \. \cP = Cases Cs; \ \ P \ cRs; (\, P) mem Cs; \ \ \; guarded P\ \ Prop" and rPar1: "\\\<^sub>Q P \ P' Q A\<^sub>Q. \cP = P \ Q; cRs = \ \ (P' \ Q); (\ \ \\<^sub>Q) \ P \ (\ \ P'); extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* \; A\<^sub>Q \* C; A\<^sub>Q \* P'; bn \ \* \; bn \ \* \\<^sub>Q; bn \ \* Q; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ Prop" and rPar2: "\\\<^sub>P Q \ Q' P A\<^sub>P. \cP = P \ Q; cRs = \ \ (P \ Q'); (\ \ \\<^sub>P) \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* \; A\<^sub>P \* C; A\<^sub>P \* Q'; bn \ \* \; bn \ \* \\<^sub>P; bn \ \* P; bn \ \* Q; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ Prop" and rComm1: "\\\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q. \cP = P \ Q; cRs = \ \ \\*xvec\P' \ Q'; \ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \ K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; distinct xvec\ \ Prop" and rComm2: "\\\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q. \cP = P \ Q; cRs = \ \ \\*xvec\P' \ Q'; \ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \ K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; distinct xvec\ \ Prop" and rOpen: "\P M xvec yvec N P' x. \cP = \\x\P; cRs = M\\*(xvec@x#yvec)\\N\ \ P'; \ \ P \ M\\*(xvec@yvec)\\N\ \ P'; x \ supp N; x \ xvec; x \ yvec; x \ M; x \ \; distinct xvec; distinct yvec; xvec \* \; xvec \* P; xvec \* M; xvec \* yvec; yvec \* \; yvec \* P; yvec \* M; xvec \* C; x \ C; yvec \* C\ \ Prop" and rScope: "\P \ P' x. \cP = \\x\P; cRs = \ \ \\x\P'; \ \ P \\ \ P'; x \ \; x \ \; x \ C; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ Prop" and rBang: "\P. \cP = !P; \ \ P \ !P \ cRs; guarded P\ \ Prop" shows Prop using \\ \ cP \ cRs\ proof(cases rule: semantics.cases) case(cInput M K xvec N Tvec P) obtain p::"name prm" where "(p \ xvec) \* \" and "(p \ xvec) \* M" and "(p \ xvec) \* N" and "(p \ xvec) \* K" and "(p \ xvec) \* Tvec" and "(p \ xvec) \* P" and "(p \ xvec) \* C" and S: "(set p) \ (set xvec) \ (set(p \ xvec))" and "distinctPerm p" by(rule_tac xvec=xvec and c="(\, M, K, N, P, C, Tvec)" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) from \cP = M\\*xvec N\.P\ \(p \ xvec) \* N\ \(p \ xvec) \* P\ S have "cP = M\\*(p \ xvec) (p \ N)\.(p \ P)" by(simp add: inputChainAlpha') moreover from \cRs = K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]\ \(p \ xvec) \* N\ \(p \ xvec) \* P\ S \length xvec = length Tvec\ \distinctPerm p\ have "cRs = K\((p \ N)[(p \ xvec)::=Tvec])\ \ (p \ P)[(p \ xvec)::=Tvec]" by(auto simp add: substTerm.renaming renaming residualInject) moreover note \\ \ M \ K\ moreover from \distinct xvec\ have "distinct(p \ xvec)" by simp moreover from \(set xvec) \ (supp N)\ have "(p \ (set xvec)) \ (p \ (supp N))" by simp hence "set(p \ xvec) \ supp(p \ N)" by(simp add: eqvts) moreover from \length xvec = length Tvec\ have "length(p \ xvec) = length Tvec" by simp ultimately show ?thesis using \(p \ xvec) \* Tvec\ \(p \ xvec) \* \\ \(p \ xvec) \* M\ \(p \ xvec) \* K\ \(p \ xvec) \* C\ by(rule rInput) next case(Output M K N P) thus ?thesis by(rule rOutput) next case(Case P \ Cs) thus ?thesis by(rule rCase) next case(cPar1 \\<^sub>Q P \ P' Q A\<^sub>Q) obtain q::"name prm" where "(bn(q \ \)) \* \" and "(bn(q \ \)) \* P" and "(bn(q \ \)) \* Q" and "(bn(q \ \)) \* \" and "(bn(q \ \)) \* A\<^sub>Q" and "(bn(q \ \)) \* P'" and "(bn(q \ \)) \* \\<^sub>Q" and "distinctPerm q" and "(bn(q \ \)) \* C" and Sq: "(set q) \ set(bn \) \ (set(bn(q \ \)))" by(rule_tac xvec="bn \" and c="(\, P, Q, \, A\<^sub>Q, \\<^sub>Q, P', C)" in name_list_avoiding) (auto simp add: eqvts) obtain p::"name prm" where "(p \ A\<^sub>Q) \* \" and "(p \ A\<^sub>Q) \* P" and "(p \ A\<^sub>Q) \* Q" and "(p \ A\<^sub>Q) \* \" and "(p \ A\<^sub>Q) \* (q \ \)" and "(p \ A\<^sub>Q) \* P'" and "(p \ A\<^sub>Q) \* (q \ P')" and "(p \ A\<^sub>Q) \* \\<^sub>Q" and "(p \ A\<^sub>Q) \* C" and Sp: "(set p) \ (set A\<^sub>Q) \ (set(p \ A\<^sub>Q))" and "distinctPerm p" by(rule_tac xvec=A\<^sub>Q and c="(\, P, Q, \, q \ \, P', (q \ P'), \\<^sub>Q, C)" in name_list_avoiding) auto from \A\<^sub>Q \* \\ \bn(q \ \) \* A\<^sub>Q\ Sq \distinctPerm q\ have "A\<^sub>Q \* (q \ \)" by(subst fresh_star_bij[symmetric, of _ _ q]) (simp add: eqvts) from \bn \ \* subject \\ \distinctPerm q\ have "bn(q \ \) \* subject(q \ \)" by(subst fresh_star_bij[symmetric, of _ _ q]) (simp add: eqvts) from \distinct(bn \)\ \distinctPerm q\ have "distinct(bn(q \ \))" by(subst distinctClosed[symmetric, of _ q]) (simp add: eqvts) note \cP = P \ Q\ moreover from \cRs = \ \ (P' \ Q)\ \bn \ \* subject \\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* P'\ \(bn(q \ \)) \* Q\ \bn \ \* Q\ Sq have "cRs = (q \ \) \ (q \ P') \ Q" by(force simp add: residualAlpha) moreover from \\ \ \\<^sub>Q \ P \\ \ P'\ \bn \ \* subject \\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* P'\ Sq have Trans: "\ \ \\<^sub>Q \ P \(q \ \) \ (q \ P')" by(force simp add: residualAlpha) hence "A\<^sub>Q \* (q \ P')" using \bn(q \ \) \* subject(q \ \)\ \distinct(bn(q \ \))\ \A\<^sub>Q \* P\ \A\<^sub>Q \* (q \ \)\ by(drule_tac freeFreshChainDerivative) auto from Trans have "(p \ (\ \ \\<^sub>Q)) \ (p \ P) \p \ ((q \ \) \ (q \ P'))" by(rule semantics.eqvt) with \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* (q \ \)\ \A\<^sub>Q \* (q \ P')\ \(bn(q \ \)) \* A\<^sub>Q\ \(p \ A\<^sub>Q) \* (q \ \)\ \(p \ A\<^sub>Q) \* \\ \(p \ A\<^sub>Q) \* P\ \(p \ A\<^sub>Q) \* (q \ P')\ Sp have "\ \ (p \ \\<^sub>Q) \ P \(q \ \) \ (q \ P')" by(simp add: eqvts) moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \(p \ A\<^sub>Q) \* \\<^sub>Q\ Sp have "extractFrame Q = \(p \ A\<^sub>Q), (p \ \\<^sub>Q)\" by(simp add: frameChainAlpha' eqvts) moreover from \(bn(q \ \)) \* \\<^sub>Q\ \(bn(q \ \)) \* A\<^sub>Q\ \(p \ A\<^sub>Q) \* (q \ \)\ Sp have "(bn(q \ \)) \* (p \ \\<^sub>Q)" by(simp add: freshAlphaPerm) moreover from \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>Q)" by simp ultimately show ?thesis using \(p \ A\<^sub>Q) \* P\ \(p \ A\<^sub>Q) \* Q\ \(p \ A\<^sub>Q) \* \\ \(p \ A\<^sub>Q) \* (q \ \)\ \(p \ A\<^sub>Q) \* (q \ P')\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q\ \(bn(q \ \)) \* P\ \(bn(q \ \)) \* C\ \(p \ A\<^sub>Q) \* C\ \bn (q \ \) \* subject (q \ \)\ \distinct(bn(q \ \))\ by(rule_tac rPar1) next case(cPar2 \\<^sub>P Q \ Q' P A\<^sub>P) obtain q::"name prm" where "(bn(q \ \)) \* \" and "(bn(q \ \)) \* P" and "(bn(q \ \)) \* Q" and "(bn(q \ \)) \* \" and "(bn(q \ \)) \* A\<^sub>P" and "(bn(q \ \)) \* Q'" and "(bn(q \ \)) \* \\<^sub>P" and "distinctPerm q" and "(bn(q \ \)) \* C" and Sq: "(set q) \ set(bn \) \ (set(bn(q \ \)))" by(rule_tac xvec="bn \" and c="(\, P, Q, \, A\<^sub>P, \\<^sub>P, Q', C)" in name_list_avoiding) (auto simp add: eqvts) obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* (q \ \)" and "(p \ A\<^sub>P) \* Q'" and "(p \ A\<^sub>P) \* (q \ Q')" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* C" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" and "distinctPerm p" by(rule_tac xvec=A\<^sub>P and c="(\, P, Q, \, q \ \, Q', (q \ Q'), \\<^sub>P, C)" in name_list_avoiding) auto from \A\<^sub>P \* \\ \bn(q \ \) \* A\<^sub>P\ Sq \distinctPerm q\ have "A\<^sub>P \* (q \ \)" by(subst fresh_star_bij[symmetric, of _ _ q]) (simp add: eqvts) from \bn \ \* subject \\ \distinctPerm q\ have "bn(q \ \) \* subject(q \ \)" by(subst fresh_star_bij[symmetric, of _ _ q]) (simp add: eqvts) from \distinct(bn \)\ \distinctPerm q\ have "distinct(bn(q \ \))" by(subst distinctClosed[symmetric, of _ q]) (simp add: eqvts) note \cP = P \ Q\ moreover from \cRs = \ \ (P \ Q')\ \bn \ \* subject \\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q'\ \(bn(q \ \)) \* P\ \bn \ \* P\ Sq have "cRs = (q \ \) \ P \ (q \ Q')" by(force simp add: residualAlpha) moreover from \\ \ \\<^sub>P \ Q \\ \ Q'\ \bn \ \* subject \\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q'\ Sq have Trans: "\ \ \\<^sub>P \ Q \(q \ \) \ (q \ Q')" by(force simp add: residualAlpha) hence "A\<^sub>P \* (q \ Q')" using \bn(q \ \) \* subject(q \ \)\ \distinct(bn(q \ \))\ \A\<^sub>P \* Q\ \A\<^sub>P \* (q \ \)\ by(drule_tac freeFreshChainDerivative) auto from Trans have "(p \ (\ \ \\<^sub>P)) \ (p \ Q) \p \ ((q \ \) \ (q \ Q'))" by(rule semantics.eqvt) with \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>P \* (q \ \)\ \A\<^sub>P \* (q \ Q')\ \(bn(q \ \)) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ \)\ \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* (q \ Q')\ Sp have "\ \ (p \ \\<^sub>P) \ Q \(q \ \) \ (q \ Q')" by(simp add: eqvts) moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \(p \ A\<^sub>P) \* \\<^sub>P\ Sp have "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" by(simp add: frameChainAlpha' eqvts) moreover from \(bn(q \ \)) \* \\<^sub>P\ \(bn(q \ \)) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ \)\ Sp have "(bn(q \ \)) \* (p \ \\<^sub>P)" by(simp add: freshAlphaPerm) moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp ultimately show ?thesis using \(p \ A\<^sub>P) \* P\ \(p \ A\<^sub>P) \* Q\ \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* (q \ \)\ \(p \ A\<^sub>P) \* (q \ Q')\ \(bn(q \ \)) \* \\ \(bn(q \ \)) \* Q\ \(bn(q \ \)) \* P\ \(bn(q \ \)) \* C\ \(p \ A\<^sub>P) \* C\ \bn (q \ \) \* subject (q \ \)\ \distinct(bn(q \ \))\ by(rule_tac rPar2) next case(cComm1 \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q) obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" and "(r \ xvec) \* K" and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" and "(r \ xvec) \* C" and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" by(rule_tac xvec=xvec and c="(\, P, Q, M, K, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q', C)" in name_list_avoiding) (auto simp add: eqvts) obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* K" and "(q \ A\<^sub>Q) \* N" and "(q \ A\<^sub>Q) \* xvec" and "(q \ A\<^sub>Q) \* Q'" and "(q \ A\<^sub>Q) \* P'" and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" and "(q \ A\<^sub>Q) \* (r \ xvec)" and "(q \ A\<^sub>Q) \* C" and Sq: "(set q) \ (set A\<^sub>Q) \ (set(q \ A\<^sub>Q))" by(rule_tac xvec=A\<^sub>Q and c="(\, P, Q, K, N, xvec, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, Q', P', C)" in name_list_avoiding) clarsimp obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" and "(p \ A\<^sub>P) \* N" and "(p \ A\<^sub>P) \* xvec" and "(p \ A\<^sub>P) \* Q'" and "(p \ A\<^sub>P) \* A\<^sub>Q" and "(p \ A\<^sub>P) \* P'" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" and "(p \ A\<^sub>P) \* C" and "(p \ A\<^sub>P) \* (r \ xvec)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" by(rule_tac xvec=A\<^sub>P and c="(\, P, Q, M, N, xvec, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, Q', P', C)" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" by(drule_tac extractFrameFreshChain) auto from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" by(drule_tac extractFrameFreshChain) auto note \cP = P \ Q\ moreover from \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ have "(r \ xvec) \* (P' \ Q')" by simp with \cRs = \ \ \\*xvec\(P' \ Q')\ \(r \ xvec) \* N\ Sr have "cRs = \ \ \\*(r \ xvec)\(r \ (P' \ Q'))" by(simp add: resChainAlpha residualInject) hence "cRs = \ \ \\*(r \ xvec)\((r \ P') \ (r \ Q'))" by simp moreover from \\ \ \\<^sub>Q \ P \M\N\ \ P'\ Sr \distinctPerm r\ \xvec \* P\ \(r \ xvec) \* P\ have "\ \ \\<^sub>Q \ P \M\(r \ N)\ \ (r \ P')" by(rule inputAlpha) hence "(q \ (\ \ \\<^sub>Q)) \ P \(q \ M)\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ by(rule_tac inputPermFrameSubject) (assumption | simp)+ hence PTrans: "\ \ (q \ \\<^sub>Q) \ P \(q \ M)\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ by(simp add: eqvts) moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" by(simp add: frameChainAlpha) moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp moreover from \\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* Q'\ have "\ \ \\<^sub>P \ Q \K\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" by(simp add: boundOutputChainAlpha'' residualInject) hence "(p \ (\ \ \\<^sub>P)) \ Q \(p \ K)\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ \(r \ xvec) \* K\\(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* A\<^sub>P\ by(rule_tac outputPermFrameSubject) (assumption | auto) hence QTrans: "\ \ (p \ \\<^sub>P) \ Q \(p \ K)\\*(r \ xvec)\\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ by(simp add: eqvts) moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" by(simp add: frameChainAlpha) moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp moreover from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(p \ q \ (\ \ \\<^sub>P \ \\<^sub>Q)) \ (p \ q \ M) \ (p \ q \ K)" by(rule_tac chanEqClosed)+ with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ \A\<^sub>Q \* \\<^sub>P\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \A\<^sub>P \* \\<^sub>Q\ \(p \ A\<^sub>P) \* \\<^sub>Q\ \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \A\<^sub>Q \* K\ \(q \ A\<^sub>Q) \* K\ \A\<^sub>P \* A\<^sub>Q\ \(p \ A\<^sub>P) \* A\<^sub>Q\ Sp Sq have "\ \ (p \ \\<^sub>P) \ (q \ \\<^sub>Q) \ (q \ M) \ (p \ K)" by(simp add: eqvts freshChainSimps) moreover note \(p \ A\<^sub>P) \* \\ moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* P\ moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* M\ Sq have "(p \ A\<^sub>P) \* (q \ M)" by(simp add: freshChainSimps) moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* N\ Sr have "(p \ A\<^sub>P) \* (r \ N)" by(simp add: freshChainSimps) moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* P'\ Sr have "(p \ A\<^sub>P) \* (r \ P')" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* Q\ moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* Q'\ Sr have "(p \ A\<^sub>P) \* (r \ Q')" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: freshChainSimps) moreover note \(q \ A\<^sub>Q) \* P\ moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* K\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ K)" by(simp add: freshChainSimps) moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* N\ Sr have "(q \ A\<^sub>Q) \* (r \ N)" by(simp add: freshChainSimps) moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* P'\ Sr have "(q \ A\<^sub>Q) \* (r \ P')" by(simp add: freshChainSimps) moreover note \(q \ A\<^sub>Q) \* Q\ moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* Q'\ Sr have "(q \ A\<^sub>Q) \* (r \ Q')" by(simp add: freshChainSimps) moreover note \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" by(simp add: freshChainSimps) moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" by(simp add: freshChainSimps) moreover note \(r \ xvec) \* P\ moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* M\ Sq have "(r \ xvec) \* (q \ M)" by(simp add: freshChainSimps) moreover note \(r \ xvec) \* Q\ moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* K\ Sp have "(r \ xvec) \* (p \ K)" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* C\ \(q \ A\<^sub>Q) \* C\ \(r \ xvec) \* C\ moreover from \distinct xvec\ have "distinct(r \ xvec)" by simp ultimately show ?thesis by(rule rComm1) next case(cComm2 \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q) obtain r::"name prm" where "(r \ xvec) \* \" and "(r \ xvec) \* P" and "(r \ xvec) \* Q" and "(r \ xvec) \* M" and "(r \ xvec) \* K" and "(r \ xvec) \* N" and "(r \ xvec) \* A\<^sub>P" and "(r \ xvec) \* A\<^sub>Q" and "(r \ xvec) \* P'" and "(r \ xvec) \* Q'" and "(r \ xvec) \* \\<^sub>P" and "(r \ xvec) \* \\<^sub>Q" and "(r \ xvec) \* C" and Sr: "(set r) \ (set xvec) \ (set(r \ xvec))" and "distinctPerm r" by(rule_tac xvec=xvec and c="(\, P, Q, M, K, N, A\<^sub>P, A\<^sub>Q, \\<^sub>P, \\<^sub>Q, P', Q', C)" in name_list_avoiding) (auto simp add: eqvts) obtain q::"name prm" where "(q \ A\<^sub>Q) \* \" and "(q \ A\<^sub>Q) \* P" and "(q \ A\<^sub>Q) \* Q" and "(q \ A\<^sub>Q) \* K" and "(q \ A\<^sub>Q) \* N" and "(q \ A\<^sub>Q) \* xvec" and "(q \ A\<^sub>Q) \* Q'" and "(q \ A\<^sub>Q) \* P'" and "(q \ A\<^sub>Q) \* \\<^sub>P" and "(q \ A\<^sub>Q) \* A\<^sub>P" and "(q \ A\<^sub>Q) \* \\<^sub>Q" and "(q \ A\<^sub>Q) \* (r \ xvec)" and "(q \ A\<^sub>Q) \* C" and Sq: "(set q) \ (set A\<^sub>Q) \ (set(q \ A\<^sub>Q))" by(rule_tac xvec=A\<^sub>Q and c="(\, P, Q, K, N, xvec, r \ xvec, \\<^sub>Q, A\<^sub>P, \\<^sub>P, Q', P', C)" in name_list_avoiding) clarsimp obtain p::"name prm" where "(p \ A\<^sub>P) \* \" and "(p \ A\<^sub>P) \* P" and "(p \ A\<^sub>P) \* Q" and "(p \ A\<^sub>P) \* M" and "(p \ A\<^sub>P) \* N" and "(p \ A\<^sub>P) \* xvec" and "(p \ A\<^sub>P) \* Q'" and "(p \ A\<^sub>P) \* A\<^sub>Q" and "(p \ A\<^sub>P) \* P'" and "(p \ A\<^sub>P) \* \\<^sub>P" and "(p \ A\<^sub>P) \* \\<^sub>Q" and "(p \ A\<^sub>P) \* (q \ A\<^sub>Q)" and "(p \ A\<^sub>P) \* C" and "(p \ A\<^sub>P) \* (r \ xvec)" and Sp: "(set p) \ (set A\<^sub>P) \ (set(p \ A\<^sub>P))" by(rule_tac xvec=A\<^sub>P and c="(\, P, Q, M, N, xvec, r \ xvec, A\<^sub>Q, q \ A\<^sub>Q, \\<^sub>Q, \\<^sub>P, Q', P', C)" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact from \A\<^sub>P \* Q\ FrQ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* \\<^sub>Q" by(drule_tac extractFrameFreshChain) auto from \A\<^sub>Q \* P\ FrP \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* \\<^sub>P" by(drule_tac extractFrameFreshChain) auto note \cP = P \ Q\ moreover from \(r \ xvec) \* P'\ \(r \ xvec) \* Q'\ have "(r \ xvec) \* (P' \ Q')" by simp with \cRs = \ \ \\*xvec\(P' \ Q')\ \(r \ xvec) \* N\ Sr have "cRs = \ \ \\*(r \ xvec)\(r \ (P' \ Q'))" by(simp add: resChainAlpha residualInject) hence "cRs = \ \ \\*(r \ xvec)\((r \ P') \ (r \ Q'))" by simp moreover from \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ Sr \(r \ xvec) \* N\ \(r \ xvec) \* P'\ have "\ \ \\<^sub>Q \ P \M\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" by(simp add: boundOutputChainAlpha'' residualInject) hence "(q \ (\ \ \\<^sub>Q)) \ P \(q \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* P\ \(q \ A\<^sub>Q) \* P\ \(r \ xvec) \* M\ \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ by(rule_tac outputPermFrameSubject) (assumption | auto) hence PTrans: "\ \ (q \ \\<^sub>Q) \ P \(q \ M)\\*(r \ xvec)\\(r \ N)\ \ (r \ P')" using Sq \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ by(simp add: eqvts) moreover from \extractFrame P = \A\<^sub>P, \\<^sub>P\\ Sp \(p \ A\<^sub>P) \* \\<^sub>P\ have FrP: "extractFrame P = \(p \ A\<^sub>P), (p \ \\<^sub>P)\" by(simp add: frameChainAlpha) moreover from \distinct A\<^sub>P\ have "distinct(p \ A\<^sub>P)" by simp moreover from \\ \ \\<^sub>P \ Q \K\N\ \ Q'\ Sr \distinctPerm r\ \xvec \* Q\ \(r \ xvec) \* Q\ have "\ \ \\<^sub>P \ Q \K\(r \ N)\ \ (r \ Q')" by(rule inputAlpha) hence "(p \ (\ \ \\<^sub>P)) \ Q \(p \ K)\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* Q\ \(p \ A\<^sub>P) \* Q\ by(rule_tac inputPermFrameSubject) (assumption | simp)+ hence QTrans: "\ \ (p \ \\<^sub>P) \ Q \(p \ K)\(r \ N)\ \ (r \ Q')" using Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ by(simp add: eqvts) moreover from \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ Sq \(q \ A\<^sub>Q) \* \\<^sub>Q\ have FrQ: "extractFrame Q = \(q \ A\<^sub>Q), (q \ \\<^sub>Q)\" by(simp add: frameChainAlpha) moreover from \distinct A\<^sub>Q\ have "distinct(q \ A\<^sub>Q)" by simp moreover from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(p \ q \ (\ \ \\<^sub>P \ \\<^sub>Q)) \ (p \ q \ M) \ (p \ q \ K)" by(rule_tac chanEqClosed)+ with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>Q \* \\ \(q \ A\<^sub>Q) \* \\ \A\<^sub>Q \* \\<^sub>P\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \A\<^sub>P \* \\<^sub>Q\ \(p \ A\<^sub>P) \* \\<^sub>Q\ \A\<^sub>P \* M\ \(p \ A\<^sub>P) \* M\ \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \A\<^sub>Q \* K\ \(q \ A\<^sub>Q) \* K\ \A\<^sub>P \* A\<^sub>Q\ \(p \ A\<^sub>P) \* A\<^sub>Q\ Sp Sq have "\ \ (p \ \\<^sub>P) \ (q \ \\<^sub>Q) \ (q \ M) \ (p \ K)" by(simp add: eqvts freshChainSimps) moreover note \(p \ A\<^sub>P) \* \\ moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* \\<^sub>Q\ Sq have "(p \ A\<^sub>P) \* (q \ \\<^sub>Q)" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* P\ moreover from \(p \ A\<^sub>P) \* A\<^sub>Q\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* M\ Sq have "(p \ A\<^sub>P) \* (q \ M)" by(simp add: freshChainSimps) moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* N\ Sr have "(p \ A\<^sub>P) \* (r \ N)" by(simp add: freshChainSimps) moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* P'\ Sr have "(p \ A\<^sub>P) \* (r \ P')" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* Q\ moreover from \(p \ A\<^sub>P) \* xvec\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(p \ A\<^sub>P) \* Q'\ Sr have "(p \ A\<^sub>P) \* (r \ Q')" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* \\ moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* \\<^sub>P\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: freshChainSimps) moreover note \(q \ A\<^sub>Q) \* P\ moreover from \(q \ A\<^sub>Q) \* A\<^sub>P\ \(p \ A\<^sub>P) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* K\ \(p \ A\<^sub>P) \* (q \ A\<^sub>Q)\ Sp Sq have "(q \ A\<^sub>Q) \* (p \ K)" by(simp add: freshChainSimps) moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* N\ Sr have "(q \ A\<^sub>Q) \* (r \ N)" by(simp add: freshChainSimps) moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* P'\ Sr have "(q \ A\<^sub>Q) \* (r \ P')" by(simp add: freshChainSimps) moreover note \(q \ A\<^sub>Q) \* Q\ moreover from \(q \ A\<^sub>Q) \* xvec\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(q \ A\<^sub>Q) \* Q'\ Sr have "(q \ A\<^sub>Q) \* (r \ Q')" by(simp add: freshChainSimps) moreover note \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\ moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>P\ Sp have "(r \ xvec) \* (p \ \\<^sub>P)" by(simp add: freshChainSimps) moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* \\<^sub>Q\ Sq have "(r \ xvec) \* (q \ \\<^sub>Q)" by(simp add: freshChainSimps) moreover note \(r \ xvec) \* P\ moreover from \(r \ xvec) \* A\<^sub>Q\ \(q \ A\<^sub>Q) \* (r \ xvec)\ \(r \ xvec) \* M\ Sq have "(r \ xvec) \* (q \ M)" by(simp add: freshChainSimps) moreover note \(r \ xvec) \* Q\ moreover from \(r \ xvec) \* A\<^sub>P\ \(p \ A\<^sub>P) \* (r \ xvec)\ \(r \ xvec) \* K\ Sp have "(r \ xvec) \* (p \ K)" by(simp add: freshChainSimps) moreover note \(p \ A\<^sub>P) \* C\ \(q \ A\<^sub>Q) \* C\ \(r \ xvec) \* C\ moreover from \distinct xvec\ have "distinct(r \ xvec)" by simp ultimately show ?thesis by(rule rComm2) next case(cOpen P M xvec yvec N P' x) from \\ \ P \ M\\*(xvec@yvec)\\N\ \ P'\ have "distinct(xvec@yvec)" by(force dest: boundOutputDistinct) hence "xvec \* yvec" by(induct xvec) auto obtain p where "(p \ yvec) \* \" and "(p \ yvec) \* P" and "(p \ yvec) \* M" and "(p \ yvec) \* yvec" and "(p \ yvec) \* N" and "(p \ yvec) \* P'" and "x \ (p \ yvec)" and "(p \ yvec) \* xvec" and "(p \ yvec) \* C" and Sp: "(set p) \ (set yvec) \ (set(p \ yvec))" by(rule_tac xvec=yvec and c="(\, P, M, xvec, yvec, N, P', x, C)" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) obtain q where "(q \ xvec) \* \" and "(q \ xvec) \* P" and "(q \ xvec) \* M" and "(q \ xvec) \* xvec" and "(q \ xvec) \* N" and "(q \ xvec) \* P'" and "x \ (q \ xvec)" and "(q \ xvec) \* yvec" and "(q \ xvec) \* p" and "(q \ xvec) \* (p \ yvec)" and "(q \ xvec) \* C" and Sq: "(set q) \ (set xvec) \ (set(q \ xvec))" by(rule_tac xvec=xvec and c="(\, P, M, xvec, yvec, p \ yvec, N, P', x, p, C)" in name_list_avoiding) (auto simp add: eqvts fresh_star_prod) obtain y::name where "y \ P" and "y \ C" and "y \ xvec" and "y \ yvec" and "y \ x" and "y \ N" and "y \ (q \ xvec)" and "y \ (p \ yvec)" and "y \ M" and "y \ \" and "y \ P'" by(generate_fresh "name") (auto simp add: freshChainSimps) from \cP = \\x\P\ \y \ P\ have "cP = \\y\([(x, y)] \ P)" by(simp add: alphaRes) moreover have "cRs = M\\*((q \ xvec)@y#(p \ yvec))\\((q@(x, y)#p) \ N)\ \ ((q@(x, y)#p) \ P')" proof - note \cRs = M\\*(xvec@x#yvec)\\N\ \ P'\ moreover have "\\*(xvec@x#yvec)\N \' P' = \\*xvec\(\\x\(\\*yvec\N \' P'))" by(simp add: boundOutputApp) moreover from \(p \ yvec) \* N\ \(p \ yvec) \* P'\ Sp have "\ = \\*xvec\(\\x\(\\*(p \ yvec)\(p \ N) \' (p \ P')))" by(simp add: boundOutputChainAlpha'') moreover with \y \ N\ \y \ P'\ \y \ (p \ yvec)\ \y \ yvec\ \x \ yvec\ \x \ (p \ yvec)\ Sp have "\ = \\*xvec\(\\y\(\\*(p \ yvec)\(([(x, y)] \ p \ N) \' ([(x, y)] \ p \ P'))))" by(subst alphaBoundOutput[where y=y]) (simp add: freshChainSimps eqvts)+ moreover hence "\ = \\*xvec\(\\y\(\\*(p \ yvec)\((((x, y)#p) \ N) \' (((x, y)#p) \ P'))))" by simp moreover from \(q \ xvec) \* N\ \(q \ xvec) \* P'\ \xvec \* yvec\ \(p \ yvec) \* xvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \y \ xvec\ \y \ (q \ xvec)\ \x \ xvec\ \x \ (q \ xvec)\ Sp Sq have "\ = \\*(q \ xvec)\(\\y\(\\*(p \ yvec)\((q \ ((x, y)#p) \ N) \' (q \ ((x, y)#p) \ P'))))" apply(subst boundOutputChainAlpha[where p=q and xvec=xvec and yvec="xvec"]) defer apply assumption apply simp apply(simp add: eqvts) apply(simp add: eqvts) apply(simp add: boundOutputFreshSet(4)) apply(rule conjI) apply(simp add: freshChainSimps) apply(simp add: freshChainSimps) done moreover hence "\ = \\*(q \ xvec@y#(p \ yvec))\((q@(x, y)#p) \ N) \' ((q@(x, y)#p) \ P')" by(simp only: pt2[OF pt_name_inst] boundOutputApp BOresChain.simps) ultimately show ?thesis by(simp only: residualInject) qed moreover have "\ \ ([(x, y)] \ P) \M\\*((q \ xvec)@(p \ yvec))\\((q@(x, y)#p) \ N)\ \ ((q@(x, y)#p) \ P')" proof - note\\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ moreover from \(p \ yvec) \* N\ \(q \ xvec) \* N\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq have "((q@p) \ (xvec @ yvec)) \* N" apply(simp only: eqvts) apply(simp only: pt2[OF pt_name_inst]) by simp moreover from \(p \ yvec) \* P'\ \(q \ xvec) \* P'\ \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ Sp Sq have "((q@p) \ (xvec @ yvec)) \* P'" by(simp del: freshAlphaPerm add: eqvts pt2[OF pt_name_inst]) moreover from Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ have Spq: "set(q@p) \ set(xvec@yvec) \ set((q@p) \ (xvec@yvec))" by(simp add: pt2[OF pt_name_inst] eqvts) blast ultimately have "\ \ P \M\\*((q@p) \ (xvec@yvec))\\((q@p) \ N)\ \ ((q@p) \ P')" apply(simp only: residualInject) by(erule_tac rev_mp) (subst boundOutputChainAlpha, auto) with Sp Sq \xvec \* yvec\ \(q \ xvec) \* yvec\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* xvec\ have "\ \ P \M\\*((q \ xvec)@(p \ yvec))\\((q@p) \ N)\ \ ((q@p) \ P')" by(simp add: eqvts pt2[OF pt_name_inst] del: freshAlphaPerm) hence "([(x, y)] \ \) \ ([(x, y)] \ P) \ [(x, y)] \ (M\\*((q \ xvec)@(p \ yvec))\\((q@p) \ N)\ \ ((q@p) \ P'))" by(rule semantics.eqvt) with \x \ \\ \y \ \\ \x \ M\ \y \ M\ \x \ xvec\ \y \ xvec\ \x \ (q \ xvec)\ \y \ (q \ xvec) \\x \ yvec\ \y \ yvec\ \x \ (p \ yvec)\ \y \ (p \ yvec)\ Sp Sq show ?thesis apply(simp add: eqvts pt2[OF pt_name_inst]) by(subst perm_compose[of q], simp)+ qed moreover from \x \ supp N\ have "((q@(x, y)#p) \ x) \ ((q@(x, y)#p) \ (supp N))" by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) with \x \ xvec\ \x \ yvec\ \x \ (q \ xvec)\ \x \ (p \ yvec)\ \y \ xvec\ \y \ (q \ xvec)\ Sp Sq have "y \ supp((q@(x, y)#p)\ N)" by(simp add: pt2[OF pt_name_inst] calc_atm eqvts) moreover from \distinct xvec\ have "distinct(q \ xvec)" by simp moreover from \distinct yvec\ have "distinct(p \ yvec)" by simp moreover note \x \ (q \ xvec)\ \x \ (p \ yvec)\ \x \ M\ \x \ \\ \(q \ xvec) \* \\ \(q \ xvec) \* P\ \(q \ xvec) \* M\ \(q \ xvec) \* (p \ yvec)\ \(p \ yvec) \* \\ \(p \ yvec) \* P\ \(p \ yvec) \* M\ \y \ (q \ xvec)\ \y \ (p \ yvec)\ \y \ M\ \y \ C\ \y \ \\ \(p \ yvec) \* C\ \(q \ xvec) \* C\ ultimately show Prop by(rule_tac rOpen) (assumption | simp)+ next case(cScope P \ P' x) obtain p::"name prm" where "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" and "x \ bn(p \ \)" and "distinctPerm p" and "(bn(p \ \)) \* C" and Sp: "(set p) \ set(bn \) \ (set(bn(p \ \)))" by(rule_tac xvec="bn \" and c="(\, P, \, x, P', C)" in name_list_avoiding) (auto simp add: eqvts) obtain y::name where "y \ \" and "y \ P" and "y \ (p \ P')" and "y \ (p \ \)" and "y \ C" by(generate_fresh "name") (auto simp add: freshChainSimps simp del: actionFresh) from \bn \ \* subject \\ \distinctPerm p\ have "bn(p \ \) \* subject(p \ \)" by(subst fresh_star_bij[symmetric, of _ _ p]) (simp add: eqvts) from \distinct(bn \)\ \distinctPerm p\ have "distinct(bn(p \ \))" by(subst distinctClosed[symmetric, of _ p]) (simp add: eqvts) from \x \ \\ \x \ (bn(p \ \))\ \distinctPerm p\ Sp have "x \ (p \ \)" by(subst fresh_bij[symmetric, of _ _ p]) (simp add: eqvts freshChainSimps) moreover from \cP = \\x\P\ \y \ P\ have "cP = \\y\([(x, y)] \ P)" by(simp add: alphaRes) moreover from \cRs = \ \ \\x\P'\ \bn \ \* subject \\ \(bn(p \ \)) \* \\ \x \ bn(p \ \)\ \(bn(p \ \)) \* P'\ \x \ \\ Sp have "cRs = (p \ \) \ \\x\(p \ P')" by(force simp add: residualAlpha) with \y \ (p \ P')\ have "cRs = (p \ \) \ \\y\([(x, y)] \ p \ P')" by(simp add: alphaRes) moreover from \\ \ P \\ \ P'\ \bn \ \* subject \\ \(bn(p \ \)) \* \\ \(bn(p \ \)) \* P'\ Sp have "\ \ P \(p \ \) \ (p \ P')" by(force simp add: residualAlpha) hence"([(x, y)] \ \) \ ([(x, y)] \ P) \[(x, y)] \ ((p \ \) \ (p \ P'))" by(rule eqvts) with \x \ \\ \y \ \\ \y \ (p \ \)\ \x \ (p \ \)\ Sp \distinctPerm p\ have "\ \ ([(x, y)] \ P) \(p \ \) \ ([(x, y)] \ p \ P')" by(simp add: eqvts) moreover from \bn(p \ \) \* P\ \y \ (p \ \)\ \y \ P\ have "bn(p \ \) \* ([(x, y)] \ P)" by(auto simp add: fresh_star_def fresh_left calc_atm) (simp add: fresh_def name_list_supp) moreover from \distinct(bn \)\ have "distinct(p \ bn \)" by simp hence "distinct(bn(p \ \))" by(simp add: eqvts) ultimately show ?thesis using \y \ \\ \y \ (p \ \)\ \y \ C\ \bn(p \ \) \* \\ \bn(p \ \) \* subject(p \ \)\ \bn(p \ \) \* C\ by(rule_tac rScope) next case(Bang P) thus ?thesis by(rule_tac rBang) qed nominal_primrec inputLength :: "('a::fs_name, 'b::fs_name, 'c::fs_name) psi \ nat" and inputLength' :: "('a::fs_name, 'b::fs_name, 'c::fs_name) input \ nat" and inputLength'' :: "('a::fs_name, 'b::fs_name, 'c::fs_name) psiCase \ nat" where "inputLength (\) = 0" | "inputLength (M\N\.P) = 0" | "inputLength (M\I) = inputLength' I" | "inputLength (Case C) = 0" | "inputLength (P \ Q) = 0" | "inputLength (\\x\P) = 0" | "inputLength (\\\) = 0" | "inputLength (!P) = 0" | "inputLength' (Trm M P) = 0" | "inputLength' (\ y I) = 1 + (inputLength' I)" | "inputLength'' (\\<^sub>c) = 0" | "inputLength'' (\\ \ P C) = 0" apply(finite_guess)+ apply(rule TrueI)+ by(fresh_guess add: fresh_nat)+ nominal_primrec boundOutputLength :: "('a, 'b, 'c) boundOutput \ nat" where "boundOutputLength (BOut M P) = 0" | "boundOutputLength (BStep x B) = (boundOutputLength B) + 1" apply(finite_guess)+ apply(rule TrueI)+ by(fresh_guess add: fresh_nat)+ nominal_primrec residualLength :: "('a, 'b, 'c) residual \ nat" where "residualLength (RIn M N P) = 0" | "residualLength (ROut M B) = boundOutputLength B" | "residualLength (RTau P) = 0" by(rule TrueI)+ lemma inputLengthProc[simp]: shows "inputLength(M\\*xvec N\.P) = length xvec" by(induct xvec) auto lemma boundOutputLengthSimp[simp]: shows "residualLength(M\\*xvec\\N\ \ P) = length xvec" by(induct xvec) (auto simp add: residualInject) lemma boundOuputLengthSimp2[simp]: shows "residualLength(\ \ P) = length(bn \)" by(nominal_induct \ rule: action.strong_induct, auto) (auto simp add: residualInject) lemmas [simp del] = inputLength_inputLength'_inputLength''.simps residualLength.simps boundOutputLength.simps lemma constructPerm: fixes xvec :: "name list" and yvec :: "name list" assumes "length xvec = length yvec" and "xvec \* yvec" and "distinct xvec" and "distinct yvec" obtains p where "set p \ set xvec \ set(p \ xvec)" and "distinctPerm p" and "yvec = p \ xvec" proof - assume "\p. \set p \ set xvec \ set (p \ xvec); distinctPerm p; yvec = p \ xvec\ \ thesis" moreover obtain n where "n = length xvec" by auto with assms have "\p. (set p) \ (set xvec) \ set (yvec) \ distinctPerm p \ yvec = p \ xvec" proof(induct n arbitrary: xvec yvec) case(0 xvec yvec) thus ?case by simp next case(Suc n xvec yvec) from \Suc n = length xvec\ obtain x xvec' where "xvec = x#xvec'" and "length xvec' = n" by(case_tac xvec) auto from \length xvec = length yvec\ \xvec = x # xvec'\ obtain y yvec' where "length xvec' = length yvec'" and "yvec = y#yvec'" by(case_tac yvec) auto from \xvec = x#xvec'\ \yvec=y#yvec'\ \xvec \* yvec\ have "x \ y" and "xvec' \* yvec'" and "x \ yvec'" and "y \ xvec'" by(auto simp add: fresh_list_cons) from \distinct xvec\ \distinct yvec\ \xvec=x#xvec'\ \yvec=y#yvec'\ have "x \ xvec'" and "y \ yvec'" and "distinct xvec'" and "distinct yvec'" by simp+ from \Suc n = length xvec\ \xvec=x#xvec'\ have "n = length xvec'" by simp with \length xvec' = length yvec'\ \xvec' \* yvec'\ \distinct xvec'\ \distinct yvec'\ obtain p where S: "set p \ set xvec' \ set yvec'" and "distinctPerm p" and "yvec' = p \ xvec'" by(drule_tac Suc) auto from S have "set((x, y)#p) \ set(x#xvec') \ set(y#yvec')" by auto moreover from \x \ xvec'\ \x \ yvec'\ \y \ xvec'\ \y \ yvec'\ S have "x \ p" and "y \ p" apply(induct p) by(auto simp add: fresh_list_nil fresh_list_cons fresh_prod name_list_supp) (auto simp add: fresh_def) with S \distinctPerm p\ \x \ y\ have "distinctPerm((x, y)#p)" by auto moreover from \yvec' = p \ xvec'\ \x \ p\ \y \ p\ \x \ xvec'\ \y \ xvec'\ have "(y#yvec') = ((x, y)#p) \ (x#xvec')" by(simp add: calc_atm freshChainSimps) ultimately show ?case using \xvec=x#xvec'\ \yvec=y#yvec'\ by blast qed ultimately show ?thesis by blast qed lemma distinctApend[simp]: fixes xvec :: "name list" and yvec :: "name list" shows "(set xvec \ set yvec = {}) = xvec \* yvec" by(auto simp add: fresh_star_def name_list_supp fresh_def) lemma lengthAux: fixes xvec :: "name list" and y :: name and yvec :: "name list" assumes "length xvec = length(y#yvec)" obtains z zvec where "xvec = z#zvec" and "length zvec = length yvec" using assms by(induct xvec arbitrary: yvec y) auto lemma lengthAux2: fixes xvec :: "name list" and yvec :: "name list" and zvec :: "name list" assumes "length xvec = length(yvec@y#zvec)" obtains xvec1 x xvec2 where "xvec=xvec1@x#xvec2" and "length xvec1 = length yvec" and "length xvec2 = length zvec" proof - assume "\xvec1 x xvec2. \xvec = xvec1 @ x # xvec2; length xvec1 = length yvec; length xvec2 = length zvec\ \ thesis" moreover from assms have "\xvec1 x xvec2. xvec=xvec1@x#xvec2 \ length xvec1 = length yvec \ length xvec2 = length zvec" apply(rule_tac x="take (length yvec) xvec" in exI) apply(rule_tac x="hd(drop (length yvec) xvec)" in exI) apply(rule_tac x="tl(drop (length yvec) xvec)" in exI) by auto ultimately show ?thesis by blast qed lemma semanticsCases[consumes 11, case_names cInput cCase cPar1 cPar2 cComm1 cComm2 cScope cBang]: fixes \ :: 'b and cP :: "('a, 'b, 'c) psi" and cRs :: "('a, 'b, 'c) residual" and C :: "'d::fs_name" and x1 :: name and x2 :: name and xvec1 :: "name list" and xvec2 :: "name list" and xvec3 :: "name list" and xvec4 :: "name list" and xvec5 :: "name list" assumes "\ \ cP \cRs" and "length xvec1 = inputLength cP" and "distinct xvec1" and "length xvec2 = residualLength cRs" and "distinct xvec2" and "length xvec3 = residualLength cRs" and "distinct xvec3" and "length xvec4 = residualLength cRs" and "distinct xvec4" and "length xvec5 = residualLength cRs" and "distinct xvec5" and rInput: "\M K N Tvec P. (\xvec1 \* \; xvec1 \* cP; xvec1 \* cRs\ \ cP = M\\*xvec1 N\.P \ cRs = K\(N[xvec1::=Tvec])\ \ P[xvec1::=Tvec] \ \ \ M \ K \ distinct xvec1 \ set xvec1 \ supp N \ length xvec1=length Tvec \ xvec1 \* Tvec \ xvec1 \* \ \ xvec1 \* M \ xvec1 \* K) \ Prop" and rOutput: "\M K N P. \cP = M\N\.P; cRs = K\N\ \ P; \ \ M \ K\ \ Prop" and rCase: "\Cs P \. \cP = Cases Cs; \ \ P \ cRs; (\, P) mem Cs; \ \ \; guarded P\ \ Prop" and rPar1: "\\\<^sub>Q P \ P' Q A\<^sub>Q. (\xvec2 \* \; xvec2 \* cP; xvec2 \* cRs\ \ cP = P \ Q \ cRs = \ \ (P' \ Q) \ xvec2 = bn \ \ \ \ \\<^sub>Q \ P \\ \ P' \ extractFrame Q = \A\<^sub>Q, \\<^sub>Q\ \ distinct A\<^sub>Q \ A\<^sub>Q \* P \ A\<^sub>Q \* Q \ A\<^sub>Q \* \ \ A\<^sub>Q \* \ \ A\<^sub>Q \* P' \ A\<^sub>Q \* C) \ Prop" and rPar2: "\\\<^sub>P Q \ Q' P A\<^sub>P. (\xvec3 \* \; xvec3 \* cP; xvec3 \* cRs\ \ cP = P \ Q \ cRs = \ \ (P \ Q') \ xvec3 = bn \ \ \ \ \\<^sub>P \ Q \\ \ Q' \ extractFrame P = \A\<^sub>P, \\<^sub>P\ \ distinct A\<^sub>P \ A\<^sub>P \* P \ A\<^sub>P \* Q \ A\<^sub>P \* \ \ A\<^sub>P \* \ \ A\<^sub>P \* Q' \ A\<^sub>P \* C) \ Prop" and rComm1: "\\\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q. \cP = P \ Q; cRs = \ \ \\*xvec\P' \ Q'; \ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \ K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; distinct xvec\ \ Prop" and rComm2: "\\\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q. \cP = P \ Q; cRs = \ \ \\*xvec\P' \ Q'; \ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \ K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* xvec; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* Q'; A\<^sub>Q \* xvec; xvec \* \; xvec \* \\<^sub>P; xvec \* \\<^sub>Q; xvec \* P; xvec \* M; xvec \* Q; xvec \* K; A\<^sub>P \* C; A\<^sub>Q \* C; xvec \* C; distinct xvec\ \ Prop" and rOpen: "\P M xvec y yvec N P'. (\xvec4 \* \; xvec4 \* cP; xvec4 \* cRs; x1 \ \; x1 \ cP; x1 \ cRs; x1 \ xvec4\ \ cP = \\x1\P \ cRs = M\\*(xvec@x1#yvec)\\N\ \ P' \ xvec4=xvec@y#yvec \ \ \ P \M\\*(xvec@yvec)\\N\ \ P' \ x1 \ supp N \ x1 \ xvec \ x1 \ yvec \ distinct xvec \ distinct yvec \ xvec \* \ \ xvec \* P \ xvec \* M \ xvec \* yvec \ yvec \* \ \ yvec \* P \ yvec \* M) \ Prop" and rScope: "\P \ P'. (\xvec5 \* \; xvec5 \* cP; xvec5 \* cRs; x2 \ \; x2 \ cP; x2 \ cRs; x2 \ xvec5\ \ cP = \\x2\P \ cRs = \ \ \\x2\P' \ xvec5 = bn \ \ \ \ P \\ \ P' \ x2 \ \ \ x2 \ \ \ bn \ \* subject \ \ distinct(bn \)) \ Prop" and rBang: "\P. \cP = !P; \ \ P \ !P \ cRs; guarded P\ \ Prop" shows Prop using \\ \ cP \cRs\ proof(cases rule: semanticsCasesAux[where C="(xvec1, xvec2, xvec3, xvec4, xvec5, x1, x2, cP, cRs, C)"]) case(cInput M K xvec N Tvec P) have B: "cP = M\\*xvec N\.P" and C: "cRs = K\(N[xvec::=Tvec])\ \ (P[xvec::=Tvec])" by fact+ from \xvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, x1, x2, cP, cRs, C)\ have "xvec \* xvec1" by simp from \length xvec1 = inputLength cP\ B have "length xvec1 = length xvec" by simp then obtain p where S: "set p \ set xvec \ set(p \ xvec)" and "distinctPerm p" and "xvec1 = p \ xvec" using \xvec \* xvec1\ \distinct xvec\ \distinct xvec1\ by(rule_tac constructPerm[where xvec=xvec and yvec=xvec1]) auto show ?thesis proof(rule rInput[where M=M and K=K and N = "p \ N" and Tvec=Tvec and P="p \ P"], goal_cases) case 1 from B \xvec \* xvec1\ \xvec1 \* cP\ have "xvec1 \* N" and "xvec1 \* P" by(auto simp add: fresh_star_def inputChainFresh name_list_supp) (auto simp add: fresh_def) from \cP = M\\*xvec N\.P\ S \xvec1 \* N\ \xvec1 \* P\ \xvec1 = p \ xvec\ have "cP = M\\*xvec1 (p \ N)\.(p \ P)" apply simp by(subst inputChainAlpha) auto moreover from \cRs = K\(N[xvec::=Tvec])\ \ P[xvec::=Tvec]\ S \xvec1 \* N\ \xvec1 \* P\ \xvec1 = p \ xvec\ \length xvec = length Tvec\ \distinctPerm p\ have "cRs = K\((p \ N)[xvec1::=Tvec])\ \ (p \ P)[xvec1::=Tvec]" by(simp add: renaming substTerm.renaming) moreover note \\ \ M \ K\ moreover from \distinct xvec\ \xvec1 = p \ xvec\ have "distinct xvec1" by simp moreover from \set xvec \ supp N\ have "(p \ set xvec) \ (p \ (supp N))" by(simp add: eqvts) with \xvec1 = p \ xvec\ have "set xvec1 \ supp(p \ N)" by(simp add: eqvts) moreover from \length xvec = length Tvec\ \xvec1 = p \ xvec\ have "length xvec1 = length Tvec" by simp moreover from \xvec1 \* cRs\ C \length xvec = length Tvec\ \distinct xvec\ \set xvec \ supp N\ have "(set xvec1) \* Tvec" by(rule_tac substTerm.subst3Chain[where T=N]) auto hence "xvec1 \* Tvec" by simp moreover from \xvec \* Tvec\ have "(p \ xvec) \* (p \ Tvec)" by(simp add: fresh_star_bij) with S \xvec \* Tvec\ \xvec1 \* Tvec\ \xvec1 = p \ xvec\ have "xvec1 \* Tvec" by simp moreover from \xvec \* M\ have "(p \ xvec) \* (p \ M)" by(simp add: fresh_star_bij) with S \xvec \* M\ \xvec1 \* cP\ B \xvec1 = p \ xvec\ have "xvec1 \* M" by simp moreover from \xvec \* K\ have "(p \ xvec) \* (p \ K)" by(simp add: fresh_star_bij) with S \xvec \* K\ \xvec1 \* cRs\ C \xvec1 = p \ xvec\ have "xvec1 \* K" by simp ultimately show ?case using \xvec1 \* \\ by blast qed next case(cOutput M K N P) thus ?thesis by(rule rOutput) next case(cCase Cs P \) thus ?thesis by(rule rCase) next case(cPar1 \\<^sub>Q P \ P' Q A\<^sub>Q) have B: "cP = P \ Q" and C: "cRs = \ \ P' \ Q" by fact+ from \bn \ \* (xvec1, xvec2, xvec3, xvec4, xvec5, x1, x2, cP, cRs, C)\ have "bn \ \* xvec2" by simp from \A\<^sub>Q \* (xvec1, xvec2, xvec3, xvec4, xvec5, x1, x2, cP, cRs, C)\ have "A\<^sub>Q \* xvec2" and "A\<^sub>Q \* C" by simp+ from \length xvec2 = residualLength cRs\ C have "length xvec2 = length(bn \)" by simp then obtain p where S: "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "xvec2= bn(p \ \)" using \bn \ \* xvec2\ \distinct(bn \)\ \distinct xvec2\ by(rule_tac constructPerm[where xvec="bn \" and yvec=xvec2]) (auto simp add: eqvts) show ?thesis proof(rule rPar1[where P=P and Q=Q and \="p \ \" and P'="p \ P'" and A\<^sub>Q=A\<^sub>Q and \\<^sub>Q=\\<^sub>Q], goal_cases) case 1 note \cP = P \ Q\ moreover from B C S \bn \ \* xvec2\ \xvec2 \* cRs\ \xvec2 = bn(p \ \)\ \bn \ \* subject \\ \xvec2 \* cP\ \bn \ \* Q\ have "cRs = (p \ \) \ (p \ P') \ Q" apply auto by(subst residualAlpha[where p=p]) auto moreover note \xvec2 = bn(p \ \)\ moreover from \\ \ \\<^sub>Q \ P \\ \ P'\ S B C S \bn \ \* xvec2\ \xvec2 \* cRs\ \xvec2 = bn(p \ \)\ \bn \ \* subject \\ \xvec2 \* cP\ have "\ \ \\<^sub>Q \ P \(p \ \) \ (p \ P')" by(subst residualAlpha[symmetric]) auto moreover note \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\ moreover from \A\<^sub>Q \* \\ \A\<^sub>Q \* xvec2\ S \xvec2 = bn(p \ \)\ \distinctPerm p\ have "A\<^sub>Q \* (p \ \)" by(subst fresh_star_bij[symmetric, where pi=p]) simp moreover from \A\<^sub>Q \* P'\ \A\<^sub>Q \* \\ \A\<^sub>Q \* xvec2\ S \xvec2 = bn(p \ \)\ \distinctPerm p\ have "A\<^sub>Q \* (p \ P')" by(subst fresh_star_bij[symmetric, where pi=p]) simp moreover note \A\<^sub>Q \* C\ ultimately show ?case by blast qed next case(cPar2 \\<^sub>P Q \ Q' P A\<^sub>P) have B: "cP = P \ Q" and C: "cRs = \ \ P \ Q'" by fact+ from \bn \ \* (xvec1, xvec2, xvec3, xvec4, xvec5, x1, x2, cP, cRs, C)\ have "bn \ \* xvec3" by simp from \A\<^sub>P \* (xvec1, xvec2, xvec3, xvec4, xvec5, x1, x2, cP, cRs, C)\ have "A\<^sub>P \* xvec3" and "A\<^sub>P \* C" by simp+ from \length xvec3 = residualLength cRs\ C have "length xvec3 = length(bn \)" by simp then obtain p where S: "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "xvec3 = bn(p \ \)" using \bn \ \* xvec3\ \distinct(bn \)\ \distinct xvec3\ by(rule_tac constructPerm[where xvec="bn \" and yvec=xvec3]) (auto simp add: eqvts) show ?thesis proof(rule rPar2[where P=P and Q=Q and \="p \ \" and Q'="p \ Q'" and A\<^sub>P=A\<^sub>P and \\<^sub>P=\\<^sub>P], goal_cases) case 1 note \cP = P \ Q\ moreover from B C S \bn \ \* xvec3\ \xvec3 \* cRs\ \xvec3 = bn(p \ \)\ \bn \ \* subject \\ \xvec3 \* cP\ \bn \ \* P\ have "cRs = (p \ \) \ P \ (p \ Q')" apply auto by(subst residualAlpha[where p=p]) auto moreover note \xvec3 = bn(p \ \)\ moreover from \\ \ \\<^sub>P \ Q \\ \ Q'\ S B C S \bn \ \* xvec3\ \xvec3 \* cRs\ \xvec3 = bn(p \ \)\ \bn \ \* subject \\ \xvec3 \* cP\ have "\ \ \\<^sub>P \ Q \(p \ \) \ (p \ Q')" by(subst residualAlpha[symmetric]) auto moreover note \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ \A\<^sub>P \* \\ moreover from \A\<^sub>P \* \\ \A\<^sub>P \* xvec3\ S \xvec3 = bn(p \ \)\ \distinctPerm p\ have "A\<^sub>P \* (p \ \)" by(subst fresh_star_bij[symmetric, where pi=p]) simp moreover from \A\<^sub>P \* Q'\ \A\<^sub>P \* \\ \A\<^sub>P \* xvec3\ S \xvec3 = bn(p \ \)\ \distinctPerm p\ have "A\<^sub>P \* (p \ Q')" by(subst fresh_star_bij[symmetric, where pi=p]) simp moreover note \A\<^sub>P \* C\ ultimately show ?case by blast qed next case(cComm1 \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q) thus ?thesis by(rule_tac rComm1[where P=P and Q=Q]) (assumption | simp)+ next case(cComm2 \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q) thus ?thesis by(rule_tac rComm2[where P=P and Q=Q]) (assumption | simp)+ next case(cOpen P M xvec yvec N P' x) have B: "cP = \\x\P" and C: "cRs = M\\*(xvec@x#yvec)\\N\ \ P'" by fact+ from \xvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, x1, x2, cP, cRs, C)\ have "xvec \* xvec4" and "xvec \* cP" and "xvec \* cRs" and "x1 \ xvec" by simp+ from \x \ (xvec1, xvec2, xvec3, xvec4, xvec5, x1, x2, cP, cRs, C)\ have "x \ xvec4" and "x \ cP" and "x \ cRs" and "x \ x1" by simp+ from \yvec \* (xvec1, xvec2, xvec3, xvec4, xvec5, x1, x2, cP, cRs, C)\ have "yvec \* xvec4" and "yvec \* cP" and "yvec \* cRs" and "x1 \ yvec" by simp+ from \xvec \* cRs\ \x \ cRs\ \yvec \* cRs\ C have "(xvec@x#yvec) \* M" by simp from \xvec \* \\ \x \ \\ \yvec \* \\ have "(xvec@x#yvec) \* \" by simp from \length xvec4 = residualLength cRs\ C obtain xvec' y yvec' where D: "xvec4 = xvec'@y#yvec'" and "length xvec' = length xvec" and "length yvec' = length yvec" by(rule_tac lengthAux2) auto with \distinct xvec\ \distinct yvec\ \x \ xvec\ \x \ yvec\ \xvec \* yvec\ \xvec \* xvec4\ \yvec \* xvec4\ \x \ xvec4\ \distinct xvec4\ have "distinct xvec'" and "distinct yvec'" and "xvec' \* yvec'" and "x \ y" and "y \ xvec'" and "y \ yvec'" and "x \ xvec'" and "x \ yvec'" and "y \ xvec" and "y \ yvec" and "xvec \* xvec'" and "yvec \* yvec'" by auto from \length xvec' = length xvec\ \xvec \* xvec'\ \distinct xvec\ \distinct xvec'\ obtain p where Sp: "set p \ set xvec \ set(p \ xvec)" and "distinctPerm p" and E: "xvec' = p \ xvec" by(metis constructPerm) from \length yvec' = length yvec\ \yvec \* yvec'\ \distinct yvec\ \distinct yvec'\ obtain q where Sq: "set q \ set yvec \ set(q \ yvec)" and "distinctPerm q" and F: "yvec' = q \ yvec" by(metis constructPerm) show ?thesis proof(rule rOpen[where P="([(x, x1)] \ P)" and xvec="p \ xvec" and y="y" and yvec="q \ yvec" and N="(p@(x1, x)#q) \ N" and P'="(p@(x1, x)#q) \ P'" and M=M], goal_cases) case 1 from \xvec \* xvec4\ \x \ xvec4\ \x1 \ xvec4\ \yvec \* xvec4\ D E F have "x \ y" and "x1 \ y" and "x1 \ p \ xvec" and "x1 \ q \ yvec" by simp+ from \xvec4 \* cRs\ \x1 \ cRs\ C have "xvec4 \* M" and "x1 \ M" by simp+ from \cP = \\x\P\ \x \ cP\ \x \ x1\ have "([(x, x1)] \ cP) = [(x, x1)] \ \\x\P" by simp with \x \ cP\ \x1 \ cP\ have "cP = \\x1\([(x, x1)] \ P)" by(simp add: eqvts calc_atm) moreover from C have "((p@(x1, x)#q) \ cRs) = (p@(x1, x)#q) \ (M\\*(xvec@x#yvec)\\N\ \ P')" by(simp add: fresh_star_bij) with Sp Sq \xvec4 \* cRs\ D E F \xvec \* cRs\ \x \ cRs\ \yvec \* cRs\ \xvec4 \* M\ \(xvec@x#yvec) \* M\ \xvec \* xvec4\ \x \ xvec4\ \yvec \* xvec4\ \xvec \* yvec\ \x \ xvec\ \x \ yvec\ \y \ xvec'\ \y \ yvec'\ \xvec' \* yvec'\ \x1 \ xvec\ \x1 \ yvec\ \x1 \ y\ \x1 \ xvec4\ \x1 \ cRs\ \x1 \ cRs\ \x \ x1\ \x1 \ M\ have "cRs = M\\*((p \ xvec)@x1#(q \ yvec))\\((p@(x1, x)#q) \ N)\ \ ((p@(x1, x)#q) \ P')" by(simp add: eqvts pt2[OF pt_name_inst] calc_atm) moreover from D E F have "xvec4 = (p \ xvec)@y#(q \ yvec)" by simp moreover from \\ \ P \M\\*(xvec@yvec)\\N\ \ P'\ have "((p@(x1, x)#q) \ \) \ ((p@(x1, x)#q) \ P) \((p@(x1, x)#q) \ (M\\*(xvec@yvec)\\N\ \ P'))" by(intro eqvts) with Sp Sq B C D E F \xvec4 \* \\ \(xvec@x#yvec) \* \\ \xvec4 \* cRs\ \x \ xvec4\ C D \x \ cRs\ \yvec \* cRs\ \xvec4 \* M\ \(xvec@x#yvec) \* M\ \x \ M\ \x1 \ cRs\ \x \ x1\ \x1 \ xvec\ \x1 \ yvec\ \xvec \* xvec4\ \yvec \* xvec4\ \x1 \ xvec4\ \x \ xvec\ \x \ yvec\ \x1 \ \\ \xvec4 \* cP\ \xvec \* P\ \yvec \* P\ \xvec' \* yvec'\ \x1 \ xvec4\ \xvec4 \* cP\ \yvec \* xvec4\ \xvec \* xvec4\ \x \ x1\ \xvec \* yvec\ have "\ \ ([(x, x1)] \ P) \M\\*((p \ xvec)@(q \ yvec))\\((p@(x1, x)#q) \ N)\ \ ((p@(x1, x)#q) \ P')" by(simp add: eqvts pt_fresh_bij[OF pt_name_inst, OF at_name_inst] pt2[OF pt_name_inst] name_swap) moreover from \x \ supp N\ have "((p@(x1, x)#q) \ x) \ ((p@(x1, x)#q) \ supp N)" by(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) hence "x1 \ supp((p@(x1, x)#q) \ N)" using \x \ xvec\ \x \ yvec\ \x1 \ xvec\ \x1 \ yvec\ \x \ xvec4\ \x1 \ xvec4\ \xvec \* xvec4\ \yvec \* xvec4\ \xvec' \* yvec'\ D E F Sp Sq \x \ x1\ by(simp add: eqvts pt2[OF pt_name_inst] calc_atm) moreover from \x1 \ xvec4\ D E F have "x1 \ (p \ xvec)" and "x1 \ (q \ yvec)" by simp+ moreover from \distinct xvec'\ \distinct yvec'\ E F have "distinct(p \ xvec)" and "distinct(q \ yvec)" by simp+ moreover from \xvec' \* yvec'\ E F have "(p \ xvec) \* (q \ yvec)" by auto moreover from \xvec \* \\ have "(p \ xvec) \* (p \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with Sp D E \xvec4 \* \\ \xvec \* \\ have "(p \ xvec) \* \" by(simp add: eqvts) moreover from \yvec \* \\ have "(p \ yvec) \* (p \ \)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with Sq D F \xvec4 \* \\ \yvec \* \\ have "(q \ yvec) \* \" by(simp add: eqvts) moreover from \xvec4 \* cP\ \x \ xvec4\ \x1 \ xvec4\ B D E F have "(p \ xvec) \* ([(x, x1)] \ P)" and "(q \ yvec) \* ([(x, x1)] \ P)" by simp+ moreover from \xvec4 \* M\ C D E F have "(p \ xvec) \* M" and "(q \ yvec) \* M" by simp+ ultimately show ?case by blast qed next case(cScope P \ P' x) have B: "cP = \\x\P" and C: "cRs = \ \ \\x\P'" by fact+ from \bn \ \* (xvec1, xvec2, xvec3, xvec4, xvec5, x1, x2, cP, cRs, C)\ have "bn \ \* xvec5" and "x2 \ bn \" by simp+ from \x \ (xvec1, xvec2, xvec3, xvec4, xvec5, x1, x2, cP, cRs, C)\ have "x \ xvec5" and "x \ x2" and "x \ cRs" by simp+ from \length xvec5 = residualLength cRs\ C have "length xvec5 = length(bn \)" by simp then obtain p where S: "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "xvec5= bn(p \ \)" using \bn \ \* xvec5\ \distinct(bn \)\ \distinct xvec5\ by(rule_tac constructPerm[where xvec="bn \" and yvec=xvec5]) (auto simp add: eqvts) show ?thesis proof(rule rScope[where P="[(x, x2)] \ P" and \="[(x, x2)] \ p \ \" and P'="[(x, x2)] \ p \ P'"], goal_cases) case 1 from \x2 \ cRs\ C \x2 \ bn \\ \x \ x2\ have "x2 \ \" and "x2 \ P'" by(auto simp add: abs_fresh) from \cP = \\x\P\ \x2 \ cP\ \x \ x2\ have "cP = \\x2\([(x, x2)] \ P)" by(simp add: alphaRes abs_fresh) moreover from B C S \bn \ \* xvec5\ \xvec5 \* cRs\ \xvec5 = bn(p \ \)\ \bn \ \* subject \\ \xvec5 \* cP\ \x \ \\ \x \ xvec5\ have "cRs = (p \ \) \ \\x\(p \ P')" apply auto by(subst residualAlpha[where p=p] alphaRes) (auto simp del: actionFresh) hence "([(x, x2)] \ cRs) = [(x, x2)] \ ((p \ \) \ \\x\(p \ P'))" by simp with \x2 \ cRs\ \x \ cRs\ have "cRs = ([(x, x2)] \ p \ \) \ \\x2\([(x, x2)] \ p \ P')" by(simp add: eqvts calc_atm) moreover from \xvec5= bn(p \ \)\ have "([(x, x2)] \ xvec5) = ([(x, x2)] \ bn(p \ \))" by simp with \x \ xvec5\ \x2 \ xvec5\ have "xvec5 = bn([(x, x2)] \ p \ \)" by(simp add: eqvts) moreover from \\ \ P \\ \ P'\ S B C S \bn \ \* xvec5\ \xvec5 \* cRs\ \xvec5 = bn(p \ \)\ \bn \ \* subject \\ \xvec5 \* cP\ \x \ xvec5\ have "\ \ P \(p \ \) \ (p \ P')" by(subst residualAlpha[symmetric]) auto hence "([(x, x2)] \ \) \ ([(x, x2)] \ P) \([(x, x2)] \ ((p \ \) \ (p \ P')))" by(rule eqvt) with \x \ \\ \x2 \ \\ have "\ \ ([(x, x2)] \ P) \([(x, x2)] \ p \ \) \ ([(x, x2)] \ p \ P')" by(simp add: eqvts) moreover note \x2 \ \\ moreover from \x \ \\ \x2 \ \\ \x \ xvec5\ \x2 \ xvec5\ S \x \ x2\ \xvec5 = bn(p \ \)\ have "x2 \ [(x, x2)] \ p \ \" apply(subgoal_tac "x \ p \ x2 \ p") apply(simp add: perm_compose freshChainSimps del: actionFresh) by(auto dest: freshAlphaSwap) moreover from \bn \ \* subject \\ have "([(x, x2)] \ p \ (bn \)) \* ([(x, x2)] \ p \ (subject \))" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) hence "bn([(x, x2)] \ p \ \) \* subject([(x, x2)] \ p \ \)" by(simp add: eqvts) moreover from \distinct(bn \)\ have "distinct([(x, x2)] \ p \ (bn \))" by simp hence "distinct(bn([(x, x2)] \ p \ \))" by(simp add: eqvts) ultimately show ?case by blast qed next case(cBang P) thus ?thesis by(rule_tac rBang) auto qed lemma parCases[consumes 5, case_names cPar1 cPar2 cComm1 cComm2]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" and \ :: "'a action" and T :: "('a, 'b, 'c) psi" and C :: "'d::fs_name" assumes Trans: "\ \ P \ Q \\ \ T" and "bn \ \* \" and "bn \ \* P" and "bn \ \* Q" and "bn \ \* subject \" and rPar1: "\P' A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* \; A\<^sub>Q \* P'; A\<^sub>Q \* C\ \ Prop \ (P' \ Q)" and rPar2: "\Q' A\<^sub>P \\<^sub>P. \\ \ \\<^sub>P \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* Q'; A\<^sub>P \* C\ \ Prop \ (P \ Q')" and rComm1: "\\\<^sub>Q M N P' A\<^sub>P \\<^sub>P K xvec Q' A\<^sub>Q. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* M; xvec \* K; xvec \* Q; xvec \* \\<^sub>Q; xvec \* C\ \ Prop (\) (\\*xvec\(P' \ Q'))" and rComm2: "\\\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P K Q' A\<^sub>Q. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; distinct xvec; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* M; xvec \* K; xvec \* Q; xvec \* \\<^sub>Q; xvec \* C\ \ Prop (\) (\\*xvec\(P' \ Q'))" shows "Prop \ T" proof - from Trans have "distinct(bn \)" by(auto dest: boundOutputDistinct) have "length(bn \) = residualLength(\ \ T)" by simp note Trans moreover have "length [] = inputLength(P \ Q)" and "distinct []" by(auto simp add: inputLength_inputLength'_inputLength''.simps) moreover note \length(bn \) = residualLength(\ \ T)\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ T)\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ T)\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ T)\ \distinct(bn \)\ moreover obtain x::name where "x \ \" and "x \ P" and "x \ Q" and "x \ \" and "x \ T" by(generate_fresh "name") auto ultimately show ?thesis using \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ apply(cases rule: semanticsCases[of _ _ _ _ _ _ _ _ _ C x x]) apply(auto simp add: psi.inject) apply(force simp add: residualInject residualInject' intro: rPar1) apply(force simp add: residualInject residualInject' intro: rPar2) apply(fastforce simp add: residualInject residualInject' intro: rComm1) by(fastforce simp add: residualInject residualInject' intro: rComm2) qed lemma parInputCases[consumes 1, case_names cPar1 cPar2]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and R :: "('a, 'b, 'c) psi" and C :: "'d::fs_name" assumes Trans: "\ \ P \ Q \M\N\ \ R" and rPar1: "\P' A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>Q \* N; A\<^sub>Q \* C\ \ Prop (P' \ Q)" and rPar2: "\Q' A\<^sub>P \\<^sub>P. \\ \ \\<^sub>P \ Q \M\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* C\ \ Prop (P \ Q')" shows "Prop R" proof - from Trans obtain \ where "\ \ P \ Q \\ \ R" and "bn \ \* \" and "bn \ \* P" and "bn \ \* Q" and "bn \ \* subject \" and "\ = M\N\" by auto thus ?thesis using rPar1 rPar2 by(induct rule: parCases) (auto simp add: residualInject) qed lemma parOutputCases[consumes 5, case_names cPar1 cPar2]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and N :: 'a and R :: "('a, 'b, 'c) psi" and C :: "'d::fs_name" assumes Trans: "\ \ P \ Q \M\\*xvec\\N\ \ R" and "xvec \* \" and "xvec \* P" and "xvec \* Q" and "xvec \* M" and rPar1: "\P' A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>Q \* xvec; A\<^sub>Q \* N; A\<^sub>Q \* C; A\<^sub>Q \* xvec; distinct xvec\ \ Prop (P' \ Q)" and rPar2: "\Q' A\<^sub>P \\<^sub>P. \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>P \* xvec; A\<^sub>P \* N; A\<^sub>P \* C; A\<^sub>P \* xvec; distinct xvec\ \ Prop (P \ Q')" shows "Prop R" proof - from Trans have "distinct xvec" by(auto dest: boundOutputDistinct) obtain \ where "\=M\\*xvec\\N\" by simp with Trans \xvec \* \\ \xvec \* P\ \xvec \* Q\ \xvec \* M\ have "\ \ P \ Q \\ \ R" and "bn \ \* \" and "bn \ \* P" and "bn \ \* Q" "bn \ \* subject \" by simp+ thus ?thesis using \\=M\\*xvec\\N\\ rPar1 rPar2 \distinct xvec\ by(induct rule: parCases[where C="(xvec, C)"]) (auto simp add: residualInject) qed lemma theEqvt[eqvt_force]: fixes p :: "name prm" and \ :: "'a action" assumes "\ \ \" shows "(p \ the(subject \)) = the(p \ (subject \))" using assms by(induct rule: actionCases[where \=\]) auto lemma theSubjectFresh[simp]: fixes \ :: "'a action" and x :: name assumes "\ \ \" shows "x \ the(subject \) = x \ subject \" using assms by(cases rule: actionCases) auto lemma theSubjectFreshChain[simp]: fixes \ :: "'a action" and xvec :: "name list" assumes "\ \ \" shows "xvec \* the(subject \) = xvec \* subject \" using assms by(cases rule: actionCases) auto lemma obtainPrefix: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and B :: "name list" assumes "\ \ P \\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "bn \ \* subject \" and "distinct(bn \)" and "\ \ \" and "B \* P" and "A\<^sub>P \* \" and "A\<^sub>P \* B" and "A\<^sub>P \* P" and "A\<^sub>P \* subject \" obtains M where "\ \ \\<^sub>P \ the(subject \) \ M" and "B \* M" using assms proof(nominal_induct avoiding: B arbitrary: thesis rule: semanticsFrameInduct') case(cAlpha \ P \ P' p A\<^sub>P \\<^sub>P B) then obtain M where subjEq: "\ \ \\<^sub>P \ the(subject \) \ M" and "B \* M" by(rule_tac cAlpha) auto from \set p \ set(bn \) \ set(bn(p \ \))\ \bn \ \* subject \\ \bn(p \ \) \* \\ subjEq have "\ \ \\<^sub>P \ the(subject(p \ \)) \ M" by(simp add: subjectEqvt[symmetric]) thus ?case using cAlpha \B \* M\ by auto next case(cFrameAlpha \ P A\<^sub>P \\<^sub>P p \ P' B) then obtain M where subjEq: "\ \ \\<^sub>P \ the(subject \) \ M" and "B \* M" by(rule_tac cFrameAlpha) auto have S: "set p \ set A\<^sub>P \ set (p \ A\<^sub>P)" by fact from subjEq have "(p \ (\ \ \\<^sub>P)) \ (p \ the(subject \)) \ (p \ M)" by(rule chanEqClosed) with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* subject \\ S \\ \ \\ \A\<^sub>P \* \\ have "\ \ (p \ \\<^sub>P) \ the(subject \) \ (p \ M)" by(simp add: eqvts del: subjectEqvt) moreover from \B \* M\ have "(p \ B) \* (p \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* B\ \(p \ A\<^sub>P) \* B\ S have "B \* (p \ M)" by(simp add: eqvts) ultimately show ?case by(rule cFrameAlpha) next case(cInput \ M K xvec N Tvec P B) from \\ \ M \ K\ have "\ \ \ \ M \ K" by(blast intro: statEqEnt AssertionStatEqSym[OF Identity]) hence "\ \ \ \ K \ M" by(rule chanEqSym) moreover from \B \* (M\\*xvec N\.P)\ have "B \* M" by simp ultimately show ?case by(rule_tac cInput) auto next case(cOutput \ M K N P B) from \\ \ M \ K\ have "\ \ \ \ M \ K" by(blast intro: statEqEnt AssertionStatEqSym[OF Identity]) hence "\ \ \ \ K \ M" by(rule chanEqSym) moreover from \B \* (M\N\.P)\ have "B \* M" by simp ultimately show ?case by(rule_tac cOutput) auto next case(cCase \ P \ P' \ Cs A\<^sub>P \\<^sub>P B) then obtain M where "\ \ \\<^sub>P \ the(subject \) \ M" and "B \* M" by(rule_tac cCase) (auto dest: memFreshChain) with \\\<^sub>P \ \\ show ?case by(blast intro: cCase statEqEnt compositionSym Identity) next case(cPar1 \ \\<^sub>Q P \ P' A\<^sub>Q Q A\<^sub>P \\<^sub>P B) then obtain M where "(\ \ \\<^sub>Q) \ \\<^sub>P \ the(subject \) \ M" and "B \* M" apply(rule_tac cPar1) by assumption auto thus ?case by(metis cPar1 statEqEnt Associativity Commutativity AssertionStatEqTrans Composition) next case(cPar2 \ \\<^sub>P Q \ Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q B) then obtain M where "(\ \ \\<^sub>P) \ \\<^sub>Q \ the(subject \) \ M" and "B \* M" by(rule_tac cPar2) auto thus ?case by(metis cPar2 statEqEnt Associativity) next case cComm1 thus ?case by simp next case cComm2 thus ?case by simp next case(cOpen \ P M xvec yvec N P' x A\<^sub>P \\<^sub>P B) then obtain K where "\ \ \\<^sub>P \ M \ K" and "B \* K" apply(rule_tac cOpen) by force auto thus ?case by(fastforce intro: cOpen) next case(cScope \ P \ P' x A\<^sub>P \\<^sub>P B) then obtain M where "\ \ \\<^sub>P \ the(subject \) \ M" and "B \* M" by(rule_tac cScope) auto thus ?case by(fastforce intro: cScope) next case(cBang \ P \ P' A\<^sub>P \\<^sub>P B) then obtain K where "\ \ \\<^sub>P \ \ \ the(subject \) \ K" and "B \* K" by(rule_tac cBang) auto with \\\<^sub>P \ \\ show ?case by(metis cBang statEqEnt compositionSym Identity) qed lemma inputRenameSubject: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b assumes "\ \ P \M\N\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "\ \ \\<^sub>P \ M \ K" and "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* M" and "A\<^sub>P \* K" shows "\ \ P \K\N\ \ P'" using assms proof(nominal_induct avoiding: K rule: inputFrameInduct) case(cAlpha \ P M N P' A\<^sub>P \\<^sub>P p K) have S: "set p \ set A\<^sub>P \ set (p \ A\<^sub>P)" by fact from \\ \ (p \ \\<^sub>P) \ M \ K\ have "(p \ (\ \ (p \ \\<^sub>P))) \ (p \ M) \ (p \ K)" by(rule chanEqClosed) with S \distinctPerm p\ \A\<^sub>P \* \\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* M\ \(p \ A\<^sub>P) \* K\ have "\ \ \\<^sub>P \ M \ K" by(simp add: eqvts) with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ \\\ \ \\<^sub>P \ M \ K; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ P \K\N\ \ P'\ show ?case by blast next case(cInput \ M K xvec N Tvec P K') from \\ \ \ \ K \ K'\ have "\ \ K \ K'" by(blast intro: statEqEnt Identity) with \\ \ M \ K\ have "\ \ M \ K'" by(rule chanEqTrans) thus ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ by(rule Input) next case(cCase \ P M N P' \ Cs A\<^sub>P \\<^sub>P K) from \\ \ \ \ M \ K\ \\\<^sub>P \ \\ have "\ \ \\<^sub>P \ M \ K" by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym) with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ \\K. \\ \ \\<^sub>P \ M \ K; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ P \K\N\ \ P'\ have "\ \ P \K\N\ \ P'" by force thus ?case using \(\, P) mem Cs\ \\ \ \\ \guarded P\ by(rule Case) next case(cPar1 \ \\<^sub>Q P M N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P K) from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K" by(metis statEqEnt Associativity Composition AssertionStatEqTrans Commutativity) with \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ \\K. \(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K; A\<^sub>P \* (\ \ \\<^sub>Q); A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ \\<^sub>Q \ P \K\N\ \ P'\ have "\ \ \\<^sub>Q \ P \K\N\ \ P'" by force thus ?case using \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* K\ \A\<^sub>Q \* N\ by(rule_tac Par1) auto next case(cPar2 \ \\<^sub>P Q M N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q K) from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(\ \ \\<^sub>P) \ \\<^sub>Q \ M \ K" by(rule statEqEnt[OF AssertionStatEqSym[OF Associativity]]) with \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \A\<^sub>Q \* K\ \\K. \(\ \ \\<^sub>P) \ \\<^sub>Q \ M \ K; A\<^sub>Q \* (\ \ \\<^sub>P); A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>Q \* K\ \ \ \ \\<^sub>P \ Q \K\N\ \ Q'\ have "\ \ \\<^sub>P \ Q \K\N\ \ Q'" by force thus ?case using \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>P \* K\ \A\<^sub>P \* N\ by(rule_tac Par2) auto next case(cScope \ P M N P' x A\<^sub>P \\<^sub>P) hence "\ \ P \K\N\ \ P'" by force with \x \ \\ \x \ K\ \x \ N\ show ?case by(rule_tac Scope) auto next case(cBang \ P M N P' A\<^sub>P \\<^sub>P K) from \\ \ \ \ M \ K\ \\\<^sub>P \ \\ have "\ \ \\<^sub>P \ \ \ M \ K" by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym) with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ \\K. \\ \ \\<^sub>P \ \ \ M \ K; A\<^sub>P \* \; A\<^sub>P \* (P \ !P); A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ P \ !P \K\N\ \ P'\ have "\ \ P \ !P \K\N\ \ P'" by force thus ?case using \guarded P\ by(rule Bang) qed lemma outputRenameSubject: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and N :: 'a and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b assumes "\ \ P \M\\*xvec\\N\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "\ \ \\<^sub>P \ M \ K" and "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* M" and "A\<^sub>P \* K" shows "\ \ P \K\\*xvec\\N\ \ P'" using assms apply(simp add: residualInject) proof(nominal_induct avoiding: K rule: outputFrameInduct) case(cAlpha \ P M A\<^sub>P \\<^sub>P p B K) have S: "set p \ set A\<^sub>P \ set(p \ A\<^sub>P)" by fact from \\ \ (p \ \\<^sub>P) \ M \ K\ have "(p \ (\ \ (p \ \\<^sub>P))) \ (p \ M) \ (p \ K)" by(rule chanEqClosed) with S \distinctPerm p\ \A\<^sub>P \* \\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ \(p \ A\<^sub>P) \* \\ \(p \ A\<^sub>P) \* M\ \(p \ A\<^sub>P) \* K\ have "\ \ \\<^sub>P \ M \ K" by(simp add: eqvts) with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ show ?case by(blast intro: cAlpha) next case(cOutput \ M K N P K') from \\ \ \ \ K \ K'\ have "\ \ K \ K'" by(blast intro: statEqEnt Identity) with \\ \ M \ K\ have "\ \ M \ K'" by(rule chanEqTrans) thus ?case using Output by(force simp add: residualInject) next case(cCase \ P M B \ Cs A\<^sub>P \\<^sub>P K) from \\ \ \ \ M \ K\ \\\<^sub>P \ \\ have "\ \ \\<^sub>P \ M \ K" by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym) with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ \\K. \\ \ \\<^sub>P \ M \ K; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ P \(ROut K B)\ have "\ \ P \ROut K B" by force thus ?case using \(\, P) mem Cs\ \\ \ \\ \guarded P\ by(rule Case) next case(cPar1 \ \\<^sub>Q P M xvec N P' A\<^sub>Q Q A\<^sub>P \\<^sub>P K) from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K" by(metis statEqEnt Associativity Composition AssertionStatEqTrans Commutativity) with \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ \\K. \(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K; A\<^sub>P \* (\ \ \\<^sub>Q); A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ \\<^sub>Q \ P \(ROut K (\\*xvec\N \' P'))\ have "\ \ \\<^sub>Q \ P \K\\*xvec\\N\ \ P'" by(force simp add: residualInject) thus ?case using \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \xvec \* Q\ \A\<^sub>Q \* \\ \A\<^sub>Q \* P\ \A\<^sub>Q \* K\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* N\ Par1[where \="K\\*xvec\\N\"] by(auto simp add: residualInject) next case(cPar2 \ \\<^sub>P Q M xvec N Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q K) from \\ \ \\<^sub>P \ \\<^sub>Q \ M \ K\ have "(\ \ \\<^sub>P) \ \\<^sub>Q \ M \ K" by(rule statEqEnt[OF AssertionStatEqSym[OF Associativity]]) with \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M\ \A\<^sub>Q \* K\ \\K. \(\ \ \\<^sub>P) \ \\<^sub>Q \ M \ K; A\<^sub>Q \* (\ \ \\<^sub>P); A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>Q \* K\ \ \ \ \\<^sub>P \ Q \ROut K (\\*xvec\N \' Q')\ have "\ \ \\<^sub>P \ Q \ROut K (\\*xvec\N \' Q')" by force thus ?case using \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \xvec \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* Q\ \A\<^sub>P \* K\ \A\<^sub>P \* xvec\ \A\<^sub>P \* N\ Par2[where \="K\\*xvec\\N\"] by(auto simp add: residualInject) next case(cOpen \ P M xvec yvec N P' x A\<^sub>P \\<^sub>P) hence "\ \ P \K\\*(xvec@yvec)\\N\ \ P'" by(force simp add: residualInject) with \x \ supp N\ \x \ \\ \x \ K\ \x \ xvec\ \x \ yvec\ Open show ?case by(auto simp add: residualInject) next case(cScope \ P M xvec N P' x A\<^sub>P \\<^sub>P) hence "\ \ P \K\\*xvec\\N\ \ P'" by(force simp add: residualInject) with \x \ \\ \x \ K\ \x \ xvec\ \x \ N\ Scope[where \="K\\*xvec\\N\"] show ?case by(auto simp add: residualInject) next case(cBang \ P M B A\<^sub>P \\<^sub>P K) from \\ \ \ \ M \ K\ \\\<^sub>P \ \\ have "\ \ \\<^sub>P \ \ \ M \ K" by(blast intro: statEqEnt Identity compositionSym AssertionStatEqSym) with \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K\ \\K. \\ \ \\<^sub>P \ \ \ M \ K; A\<^sub>P \* \; A\<^sub>P \* (P \ !P); A\<^sub>P \* M; A\<^sub>P \* K\ \ \ \ P \ !P \ROut K B\ have "\ \ P \ !P \ROut K B" by force thus ?case using \guarded P\ by(rule Bang) qed lemma parCasesSubject[consumes 7, case_names cPar1 cPar2 cComm1 cComm2]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" and \ :: "'a action" and R :: "('a, 'b, 'c) psi" and C :: "'d::fs_name" and yvec :: "name list" assumes Trans: "\ \ P \ Q \\ \ R" and "bn \ \* \" and "bn \ \* P" and "bn \ \* Q" and "bn \ \* subject \" and "yvec \* P" and "yvec \* Q" and rPar1: "\P' A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \\ \ P'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* \; A\<^sub>Q \* C\ \ Prop \ (P' \ Q)" and rPar2: "\Q' A\<^sub>P \\<^sub>P. \\ \ \\<^sub>P \ Q \\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; A\<^sub>P \* \; A\<^sub>P \* Q; A\<^sub>P \* \; A\<^sub>P \* C\ \ Prop \ (P \ Q')" and rComm1: "\\\<^sub>Q M N P' A\<^sub>P \\<^sub>P K xvec Q' A\<^sub>Q. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; yvec \* M; yvec \* K; distinct xvec; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* M; xvec \* K; xvec \* Q; xvec \* \\<^sub>Q; xvec \* C\ \ Prop (\) (\\*xvec\(P' \ Q'))" and rComm2: "\\\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P K Q' A\<^sub>Q. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; distinct A\<^sub>P; \ \ \\<^sub>P \ Q \K\N\ \ Q'; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>Q; \ \ \\<^sub>P \ \\<^sub>Q \ M \ K; yvec \* M; yvec \* K; distinct xvec; A\<^sub>P \* \; A\<^sub>P \* \\<^sub>Q; A\<^sub>P \* P; A\<^sub>P \* M; A\<^sub>P \* N; A\<^sub>P \* P'; A\<^sub>P \* Q; A\<^sub>P \* xvec; A\<^sub>P \* Q'; A\<^sub>P \* A\<^sub>Q; A\<^sub>P \* C; A\<^sub>Q \* \; A\<^sub>Q \* \\<^sub>P; A\<^sub>Q \* P; A\<^sub>Q \* K; A\<^sub>Q \* N; A\<^sub>Q \* P'; A\<^sub>Q \* Q; A\<^sub>Q \* xvec; A\<^sub>Q \* Q'; A\<^sub>Q \* C; xvec \* \; xvec \* \\<^sub>P; xvec \* P; xvec \* M; xvec \* K; xvec \* Q; xvec \* \\<^sub>Q; xvec \* C\ \ Prop (\) (\\*xvec\(P' \ Q'))" shows "Prop \ R" using Trans \bn \ \* \\ \bn \ \* P\ \bn \ \* Q\ \bn \ \* subject \\ proof(induct rule: parCases[where C="(C, yvec)"]) case(cPar1 P' A\<^sub>Q \\<^sub>Q) thus ?case by(rule_tac rPar1) auto next case(cPar2 Q' A\<^sub>P \\<^sub>P) thus ?case by(rule_tac rPar2) auto next case(cComm1 \\<^sub>Q M N P' A\<^sub>P \\<^sub>P K xvec Q' A\<^sub>Q) from \A\<^sub>P \* (C, yvec)\ \A\<^sub>Q \* (C, yvec)\ \xvec \* (C, yvec)\ have "A\<^sub>P \* C" and "A\<^sub>Q \* C" and "xvec \* C" and "A\<^sub>P \* yvec" and "A\<^sub>Q \* yvec" and "xvec \* yvec" by simp+ have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and MeqK: "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" by fact+ from \\ \ \\<^sub>Q \ P \M\N\ \ P'\ FrP \distinct A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \yvec \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* yvec\ \A\<^sub>P \* xvec\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \xvec \* P\ \A\<^sub>P \* \\<^sub>Q\ obtain M' where MeqM': "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "xvec \* M'" and "yvec \* M'" and "A\<^sub>Q \* M'" by(rule_tac B="xvec@yvec@A\<^sub>Q" in obtainPrefix) (assumption | force)+ from \\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \yvec \* Q\ \A\<^sub>Q \* \\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* yvec\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \xvec \* Q\ \A\<^sub>Q \* \\<^sub>P\ \xvec \* K\ \distinct xvec\ obtain K' where KeqK': "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ K'" and "xvec \* K'" and "yvec \* K'" and "A\<^sub>P \* K'" by(rule_tac B="xvec@yvec@A\<^sub>P" in obtainPrefix) (assumption | force | metis freshChainSym)+ from MeqK KeqK' have "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K'" by(metis statEqEnt Associativity Commutativity Composition chanEqTrans) with \\ \ \\<^sub>Q \ P \M\N\ \ P'\ FrP \distinct A\<^sub>P\ have "\ \ \\<^sub>Q \ P \K'\N\ \ P'" using \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K'\ by(rule_tac inputRenameSubject) (assumption | force)+ moreover note FrP \distinct A\<^sub>P\ moreover from MeqK MeqM' have "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym) with \\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ have "\ \ \\<^sub>P \ Q \M'\\*xvec\\N\ \ Q'" using \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>Q \* M'\ by(rule_tac outputRenameSubject) (assumption | force)+ moreover note FrQ \distinct A\<^sub>Q\ moreover from MeqM' KeqK' MeqK have "\ \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym) moreover note \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* K'\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* Q\ \A\<^sub>P \* xvec\ \A\<^sub>P \* Q'\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* C\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M'\ \A\<^sub>Q \* N\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* P\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* C\ \xvec \* \\ \xvec \* \\<^sub>P\ \xvec \* P\ \xvec \* M'\ \xvec \* K'\ \xvec \* Q\ \xvec \* \\<^sub>Q\ \xvec \* C\ \yvec \* M'\ \yvec \* K'\ \distinct xvec\ ultimately show ?case by(rule_tac rComm1) next case(cComm2 \\<^sub>Q M xvec N P' A\<^sub>P \\<^sub>P K Q' A\<^sub>Q) from \A\<^sub>P \* (C, yvec)\ \A\<^sub>Q \* (C, yvec)\ \xvec \* (C, yvec)\ have "A\<^sub>P \* C" and "A\<^sub>Q \* C" and "xvec \* C" and "A\<^sub>P \* yvec" and "A\<^sub>Q \* yvec" and "xvec \* yvec" by simp+ have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and MeqK: "\ \ \\<^sub>P \ \\<^sub>Q \ M \ K" by fact+ from \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ FrP \distinct A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \yvec \* P\ \A\<^sub>P \* \\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* yvec\ \A\<^sub>P \* xvec\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \xvec \* P\ \A\<^sub>P \* \\<^sub>Q\ \xvec \* M\ \distinct xvec\ obtain M' where MeqM': "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "xvec \* M'" and "yvec \* M'" and "A\<^sub>Q \* M'" by(rule_tac B="xvec@yvec@A\<^sub>Q" in obtainPrefix) (assumption | force)+ from \\ \ \\<^sub>P \ Q \K\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \yvec \* Q\ \A\<^sub>Q \* \\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* yvec\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \xvec \* Q\ \A\<^sub>Q \* \\<^sub>P\ obtain K' where KeqK': "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ K'" and "xvec \* K'" and "yvec \* K'" and "A\<^sub>P \* K'" by(rule_tac B="xvec@yvec@A\<^sub>P" in obtainPrefix) (assumption | force | metis freshChainSym)+ from MeqK KeqK' have "(\ \ \\<^sub>Q) \ \\<^sub>P \ M \ K'" by(metis statEqEnt Associativity Commutativity Composition chanEqTrans) with \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ FrP \distinct A\<^sub>P\ have "\ \ \\<^sub>Q \ P \K'\\*xvec\\N\ \ P'" using \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K'\ by(rule_tac outputRenameSubject) (assumption | force)+ moreover note FrP \distinct A\<^sub>P\ moreover from MeqK MeqM' have "(\ \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym) with \\ \ \\<^sub>P \ Q \K\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ have "\ \ \\<^sub>P \ Q \M'\N\ \ Q'" using \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>Q \* M'\ by(rule_tac inputRenameSubject) (assumption | force)+ moreover note FrQ \distinct A\<^sub>Q\ moreover from MeqM' KeqK' MeqK have "\ \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" by(metis statEqEnt Associativity Commutativity Composition chanEqTrans chanEqSym) moreover note \A\<^sub>P \* \\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* K'\ \A\<^sub>P \* N\ \A\<^sub>P \* P'\ \A\<^sub>P \* Q\ \A\<^sub>P \* xvec\ \A\<^sub>P \* Q'\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* C\ \A\<^sub>Q \* \\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M'\ \A\<^sub>Q \* N\ \A\<^sub>Q \* Q'\ \A\<^sub>Q \* P\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* P'\ \A\<^sub>Q \* C\ \xvec \* \\ \xvec \* \\<^sub>P\ \xvec \* P\ \xvec \* M'\ \xvec \* K'\ \xvec \* Q\ \xvec \* \\<^sub>Q\ \xvec \* C\ \yvec \* M'\ \yvec \* K'\ \distinct xvec\ ultimately show ?case by(rule_tac rComm2) qed lemma inputCases[consumes 1, case_names cInput]: fixes \ :: 'b and M :: 'a and xvec :: "name list" and N :: 'a and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" assumes Trans: "\ \ M\\*xvec N\.P \\ \ P'" and rInput: "\K Tvec. \\ \ M \ K; set xvec \ supp N; length xvec = length Tvec; distinct xvec\ \ Prop (K\N[xvec::=Tvec]\) (P[xvec::=Tvec])" shows "Prop \ P'" proof - { fix xvec N P assume Trans: "\ \ M\\*xvec N\.P \\ \ P'" and "xvec \* \" and "xvec \* M" and "xvec \* \" and "xvec \* P'" and "distinct xvec" and rInput: "\K Tvec. \\ \ M \ K; set xvec \ supp N; length xvec = length Tvec; distinct xvec\ \ Prop (K\N[xvec::=Tvec]\) (P[xvec::=Tvec])" from Trans have "bn \ = []" apply - by(ind_cases "\ \ M\\*xvec N\.P \\ \ P'") (auto simp add: residualInject) from Trans have "distinct(bn \)" by(auto dest: boundOutputDistinct) have "length(bn \) = residualLength(\ \ P')" by simp note Trans moreover have "length xvec = inputLength(M\\*xvec N\.P)" by auto moreover note \distinct xvec\ moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ moreover obtain x::name where "x \ \" and "x \ P" and "x \ M" and "x \ xvec" and "x \ \" and "x \ P'" and "x \ N" by(generate_fresh "name") auto ultimately have "Prop \ P'" using \bn \ = []\ \xvec \* \\\xvec \* M\ \xvec \* \\ \xvec \* P'\ apply(cases rule: semanticsCases[of _ _ _ _ _ _ _ _ _ C x]) apply(force simp add: residualInject psi.inject rInput) by(fastforce simp add: residualInject psi.inject inputChainFresh)+ } note Goal = this moreover obtain p :: "name prm" where "(p \ xvec) \* \" and "(p \ xvec) \* M" and "(p \ xvec) \* N" and "(p \ xvec) \* P" and "(p \ xvec) \* \" and "(p \ xvec) \* P'" and S: "set p \ set xvec \ set(p \ xvec)" and "distinctPerm p" by(rule_tac xvec=xvec and c="(\, M, N, P, \, P')" in name_list_avoiding) auto from Trans \(p \ xvec) \* N\ \(p \ xvec) \* P\ S have "\ \ M\\*(p \ xvec) (p \ N)\.(p \ P) \\ \ P'" by(simp add: inputChainAlpha') moreover { fix K Tvec assume "\ \ M \ K" moreover assume "set(p \ xvec) \ supp(p \ N)" hence "(p \ set(p \ xvec)) \ (p \ supp(p \ N))" by simp with \distinctPerm p\ have "set xvec \ supp N" by(simp add: eqvts) moreover assume "length(p \ xvec) = length(Tvec::'a list)" hence "length xvec = length Tvec" by simp moreover assume "distinct xvec" ultimately have "Prop (K\N[xvec::=Tvec]\) (P[xvec::=Tvec])" by(rule rInput) with \length xvec = length Tvec\ S \distinctPerm p\ \(p \ xvec) \* N\ \(p \ xvec) \* P\ have "Prop (K\(p \ N)[(p \ xvec)::=Tvec]\) ((p \ P)[(p \ xvec)::=Tvec])" by(simp add: renaming substTerm.renaming) } moreover from Trans have "distinct xvec" by(rule inputDistinct) hence "distinct(p \ xvec)" by simp ultimately show ?thesis using \(p \ xvec) \* \\ \(p \ xvec) \* M\ \(p \ xvec) \* \\ \(p \ xvec) \* P'\ \distinct xvec\ by(rule_tac Goal) assumption+ qed lemma outputCases[consumes 1, case_names cOutput]: fixes \ :: 'b and M :: 'a and N :: 'a and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" assumes "\ \ M\N\.P \\ \ P'" and "\K. \ \ M \ K \ Prop (K\N\) P" shows "Prop \ P'" using assms by(cases rule: semantics.cases) (auto simp add: residualInject psi.inject) lemma caseCases[consumes 1, case_names cCase]: fixes \ :: 'b and Cs :: "('c \ ('a, 'b, 'c) psi) list" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" assumes Trans: "\ \ (Cases Cs) \ Rs" and rCase: "\\ P. \\ \ P \ Rs; (\, P) mem Cs; \ \ \; guarded P\ \ Prop" shows "Prop" using assms by(cases rule: semantics.cases) (auto simp add: residualInject psi.inject) lemma resCases[consumes 7, case_names cOpen cRes]: fixes \ :: 'b and x :: name and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and C :: "'d::fs_name" assumes Trans: "\ \ \\x\P \\ \ P'" and "x \ \" and "x \ \" and "x \ P'" and "bn \ \* \" and "bn \ \* P" and "bn \ \* subject \" and rOpen: "\M xvec yvec y N P'. \\ \ P \M\\*(xvec@yvec)\\([(x, y)] \ N)\ \ ([(x, y)] \ P'); y \ supp N; x \ N; x \ P'; x \ y; y \ xvec; y \ yvec; y \ M; distinct xvec; distinct yvec; xvec \* \; y \ \; yvec \* \; xvec \* P; y \ P; yvec \* P; xvec \* M; y \ M; yvec \* M; xvec \* yvec\ \ Prop (M\\*(xvec@y#yvec)\\N\) P'" and rScope: "\P'. \\ \ P \\ \ P'\ \ Prop \ (\\x\P')" shows "Prop \ P'" proof - from Trans have "distinct(bn \)" by(auto dest: boundOutputDistinct) have "length(bn \) = residualLength(\ \ P')" by simp note Trans moreover have "length [] = inputLength(\\x\P)" and "distinct []" by(auto simp add: inputLength_inputLength'_inputLength''.simps) moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ ultimately show ?thesis using \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \x \ \\ \x \ \\ \x \ P'\ \distinct(bn \)\ apply(cases rule: semanticsCases[of _ _ _ _ _ _ _ _ _ _ x x]) apply(auto simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts) apply(subgoal_tac "y \ supp Na") apply(rule_tac rOpen) apply(auto simp add: residualInject boundOutputApp) apply(drule_tac pi="[(x, y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) apply(simp add: calc_atm eqvts) by(rule rScope) qed lemma resCases'[consumes 7, case_names cOpen cRes]: fixes \ :: 'b and x :: name and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and C :: "'d::fs_name" assumes Trans: "\ \ \\x\P \\ \ P'" and "x \ \" and "x \ \" and "x \ P'" and "bn \ \* \" and "bn \ \* P" and "bn \ \* subject \" and rOpen: "\M xvec yvec y N P'. \\ \ ([(x, y)] \ P) \M\\*(xvec@yvec)\\N\ \ P'; y \ supp N; x \ N; x \ P'; x \ y; y \ xvec; y \ yvec; y \ M; distinct xvec; distinct yvec; xvec \* \; y \ \; yvec \* \; xvec \* P; y \ P; yvec \* P; xvec \* M; y \ M; yvec \* M; xvec \* yvec\ \ Prop (M\\*(xvec@y#yvec)\\N\) P'" and rScope: "\P'. \\ \ P \\ \ P'\ \ Prop \ (\\x\P')" shows "Prop \ P'" proof - from Trans have "distinct(bn \)" by(auto dest: boundOutputDistinct) have "length(bn \) = residualLength(\ \ P')" by simp note Trans moreover have "length [] = inputLength(\\x\P)" and "distinct []" by(auto simp add: inputLength_inputLength'_inputLength''.simps) moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ moreover note \length(bn \) = residualLength(\ \ P')\ \distinct(bn \)\ ultimately show ?thesis using \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \x \ \\ \x \ \\ \x \ P'\ \distinct(bn \)\ apply(cases rule: semanticsCases[of _ _ _ _ _ _ _ _ _ _ x x]) apply(auto simp add: psi.inject alpha abs_fresh residualInject boundOutputApp boundOutput.inject eqvts) apply(subgoal_tac "y \ supp Na") apply(rule_tac rOpen) apply(auto simp add: residualInject boundOutputApp) apply(drule_tac pi="[(x, y)]" in semantics.eqvt) apply(simp add: calc_atm eqvts) apply(drule_tac pi="[(x, y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) apply(simp add: calc_atm eqvts) by(rule rScope) qed abbreviation statImpJudge ("_ \ _" [80, 80] 80) where "\ \ \' \ AssertionStatImp \ \'" lemma statEqTransition: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Rs :: "('a, 'b, 'c) residual" and \' :: 'b assumes "\ \ P \ Rs" and "\ \ \'" shows "\' \ P \ Rs" using assms proof(nominal_induct avoiding: \' rule: semantics.strong_induct) case(cInput \ M K xvec N Tvec P \') from \\ \ \'\ \\ \ M \ K\ have "\' \ M \ K" by(simp add: AssertionStatImp_def AssertionStatEq_def) thus ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ by(rule Input) next case(Output \ M K N P \') from \\ \ \'\ \\ \ M \ K\ have "\' \ M \ K" by(simp add: AssertionStatImp_def AssertionStatEq_def) thus ?case by(rule semantics.Output) next case(Case \ P Rs \ Cs \') then have "\' \ P \ Rs" by(rule_tac Case) moreover note \(\, P) mem Cs\ moreover from \\ \ \'\ \\ \ \\ have "\' \ \" by(simp add: AssertionStatImp_def AssertionStatEq_def) ultimately show ?case using \guarded P\ by(rule semantics.Case) next case(cPar1 \ \Q P \ P' xvec Q \') thus ?case by(rule_tac Par1) (auto intro: Composition) next case(cPar2 \ \P Q \ Q' xvec P \') thus ?case by(rule_tac Par2) (auto intro: Composition) next case(cComm1 \ \Q P M N P' xvec \P Q K zvec Q' yvec \') thus ?case by(clarsimp, rule_tac Comm1) (blast intro: Composition statEqEnt)+ next case(cComm2 \ \Q P M zvec N P' xvec \P Q K Q' yvec \') thus ?case by(clarsimp, rule_tac Comm2) (blast intro: Composition statEqEnt)+ next case(cOpen \ P M xvec N P' x yvec \') thus ?case by(force intro: Open) next case(cScope \ P \ P' x \') thus ?case by(force intro: Scope) next case(Bang \ P Rs \') thus ?case by(force intro: semantics.Bang) qed lemma actionPar1Dest': fixes \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and \ :: "('a::fs_name) action" and Q :: "('a, 'b, 'c) psi" and R :: "('a, 'b, 'c) psi" assumes "\ \ P = \ \ (Q \ R)" and "bn \ \* R" and "bn \ \* R" obtains T where "P = T \ R" and "\ \ T = \ \ Q" using assms apply(cases rule: actionCases[where \=\]) apply(auto simp add: residualInject) by(drule_tac boundOutputPar1Dest) auto lemma actionPar2Dest': fixes \ :: "('a::fs_name) action" and P :: "('a, 'b::fs_name, 'c::fs_name) psi" and \ :: "('a::fs_name) action" and Q :: "('a, 'b, 'c) psi" and R :: "('a, 'b, 'c) psi" assumes "\ \ P = \ \ (Q \ R)" and "bn \ \* Q" and "bn \ \* Q" obtains T where "P = Q \ T" and "\ \ T = \ \ R" using assms apply(cases rule: actionCases[where \=\]) apply(auto simp add: residualInject) by(drule_tac boundOutputPar2Dest) auto lemma expandNonTauFrame: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and C :: "'d::fs_name" and C' :: "'e::fs_name" assumes Trans: "\ \ P \\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "bn \ \* subject \" and "A\<^sub>P \* P" and "A\<^sub>P \* \" and "A\<^sub>P \* C" and "A\<^sub>P \* C'" and "bn \ \* P" and "bn \ \* C'" and "\ \ \" obtains p \' A\<^sub>P' \\<^sub>P' where "set p \ set(bn \) \ set(bn(p \ \))" and "(p \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinctPerm p" and "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (p \ \)" and "A\<^sub>P' \* C" and "(bn(p \ \)) \* C'" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" and "distinct A\<^sub>P'" proof - assume A: "\p \' \\<^sub>P' A\<^sub>P'. \set p \ set(bn \) \ set(bn(p \ \)); (p \ \\<^sub>P) \ \' \ \\<^sub>P'; distinctPerm p; extractFrame P' = \A\<^sub>P', \\<^sub>P'\; A\<^sub>P' \* P'; A\<^sub>P' \* \; A\<^sub>P' \* (p \ \); A\<^sub>P' \* C; (bn(p \ \)) \* C'; (bn(p \ \)) \* \; (bn(p \ \)) \* P'; distinct A\<^sub>P'\ \ thesis" from Trans have "distinct(bn \)" by(auto dest: boundOutputDistinct) with Trans \bn \ \* subject \\ \A\<^sub>P \* P\ \A\<^sub>P \* \\ have "A\<^sub>P \* P'" by(drule_tac freeFreshChainDerivative) auto { fix X :: "name list" and Y :: "'b list" and Z :: "('a, 'b, 'c) psi list" assume "bn \ \* X" and "bn \ \* Y" and "bn \ \* Z" and "A\<^sub>P \* X" and "A\<^sub>P \* Y" and "A\<^sub>P \* Z" with assms obtain p \' A\<^sub>P' \\<^sub>P' where "set p \ set(bn \) \ set(bn(p \ \))" and "(p \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinctPerm p" and "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (p \ \)" and "A\<^sub>P' \* C" and "(bn(p \ \)) \* C'" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "distinct A\<^sub>P'" and "(bn(p \ \)) \* X" and "(bn(p \ \)) \* Y" and "(bn(p \ \)) \* Z" using \A\<^sub>P \* P'\ \distinct(bn \)\ proof(nominal_induct \ P Rs=="\ \ P'" A\<^sub>P \\<^sub>P avoiding: C C' \ P' X Y Z arbitrary: thesis rule: semanticsFrameInduct) case(cAlpha \ P A\<^sub>P \\<^sub>P p C C' \ P' X Y Z) then obtain q \' A\<^sub>P' \\<^sub>P' where Sq: "set q \ set(bn \) \ set(bn(q \ \))" and PeqP': "(q \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinctPerm q" and FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (q \ \)" and "A\<^sub>P' \* C" and "(bn(q \ \)) \* C'" and "(bn(q \ \)) \* \" and "(bn(q \ \)) \* P'" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "distinct A\<^sub>P'" and "(bn(q \ \)) \* X" and "(bn(q \ \)) \* Y" and "(bn(q \ \)) \* Z" by metis have Sp: "set p \ set A\<^sub>P \ set (p \ A\<^sub>P)" by fact from Sq have "(p \ set q) \ p \ (set(bn \) \ set(bn(q \ \)))" by(simp add: subsetClosed) hence "set(p \ q) \ set(bn(p \ \)) \ set(p \ bn(q \ \))" by(simp add: eqvts) with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ Sp have "set(p \ q) \ set(bn \) \ set(bn((p \ q) \ \))" by(simp add: perm_compose bnEqvt[symmetric]) moreover from PeqP' have "(p \ (q \ \\<^sub>P) \ \') \ (p \ \\<^sub>P')" by(simp add: AssertionStatEqClosed) hence "((p \ q) \ p \ \\<^sub>P) \ (p \ \') \ (p \ \\<^sub>P')" apply(subst perm_compose[symmetric]) by(simp add: eqvts) moreover from \distinctPerm q\ have "distinctPerm (p \ q)" by simp moreover from \(bn(q \ \)) \* C'\ have "(p \ bn(q \ \)) \* (p \ C')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* C'\ \(p \ A\<^sub>P) \* C'\ Sp have "bn((p \ q) \ \) \* C'" by(simp add: perm_compose bnEqvt[symmetric]) moreover from FrP' have "(p \ extractFrame P') = p \ \A\<^sub>P', \\<^sub>P'\" by simp with \A\<^sub>P \* P'\ \(p \ A\<^sub>P) \* P'\ Sp have "extractFrame P' = \p \ A\<^sub>P', p \ \\<^sub>P'\" by(simp add: eqvts) moreover from \A\<^sub>P' \* P'\ have "(p \ A\<^sub>P') \* (p \ P')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* P'\ \(p \ A\<^sub>P) \* P'\ Sp have "(p \ A\<^sub>P') \* P'" by simp moreover from \A\<^sub>P' \* \\ have "(p \ A\<^sub>P') \* (p \ \)" by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ Sp have "(p \ A\<^sub>P') \* \" by simp moreover from \A\<^sub>P' \* C\ have "(p \ A\<^sub>P') \* (p \ C)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* C\ \(p \ A\<^sub>P) \* C\ Sp have "(p \ A\<^sub>P') \* C" by simp moreover from \(bn(q \ \)) \* \\ have "(p \ bn(q \ \)) \* (p \ \)" by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ Sp have "bn((p \ q) \ \) \* \" by(simp add: perm_compose eqvts) moreover from \(bn(q \ \)) \* P'\ have "(p \ bn(q \ \)) \* (p \ P')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* P'\ \(p \ A\<^sub>P) \* P'\ Sp have "bn((p \ q) \ \) \* P'" by(simp add: perm_compose eqvts) moreover from \A\<^sub>P' \* (q \ \)\ have "(p \ A\<^sub>P') \* (p \ q \ \)" by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with Sp \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ have "(p \ A\<^sub>P') \* ((p \ q) \ \)" by(simp add: perm_compose) moreover from \A\<^sub>P' \* X\ have "(p \ A\<^sub>P') \* (p \ X)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* X\ \(p \ A\<^sub>P) \* X\ Sp have "(p \ A\<^sub>P') \* X" by simp moreover from \A\<^sub>P' \* Y\ have "(p \ A\<^sub>P') \* (p \ Y)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* Y\ \(p \ A\<^sub>P) \* Y\ Sp have "(p \ A\<^sub>P') \* Y" by simp moreover from \A\<^sub>P' \* Z\ have "(p \ A\<^sub>P') \* (p \ Z)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* Z\ \(p \ A\<^sub>P) \* Z\ Sp have "(p \ A\<^sub>P') \* Z" by simp moreover from \(bn(q \ \)) \* X\ have "(p \ bn(q \ \)) \* (p \ X)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* X\ \(p \ A\<^sub>P) \* X\ Sp have "bn((p \ q) \ \) \* X" by(simp add: perm_compose eqvts) moreover from \(bn(q \ \)) \* Y\ have "(p \ bn(q \ \)) \* (p \ Y)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* Y\ \(p \ A\<^sub>P) \* Y\ Sp have "bn((p \ q) \ \) \* Y" by(simp add: perm_compose eqvts) moreover from \(bn(q \ \)) \* Z\ have "(p \ bn(q \ \)) \* (p \ Z)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* \\ \(p \ A\<^sub>P) \* \\ \A\<^sub>P \* Z\ \(p \ A\<^sub>P) \* Z\ Sp have "bn((p \ q) \ \) \* Z" by(simp add: perm_compose eqvts) moreover from \distinct A\<^sub>P'\ have "distinct(p \ A\<^sub>P')" by simp ultimately show ?case by(rule_tac cAlpha) next case(cInput \ M K xvec N Tvec P C C' \ P' X Y Z) moreover obtain A\<^sub>P \\<^sub>P where "extractFrame(P[xvec::=Tvec]) = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "A\<^sub>P \* (C, P[xvec::=Tvec], \, P', X, Y, Z, N)" by(rule freshFrame) moreover have "\ \ \\<^sub>P \ \\<^sub>P" by(blast intro: Identity Commutativity AssertionStatEqTrans) ultimately show ?case by(rule_tac cInput) (assumption | simp add: residualInject)+ next case(cOutput \ M K N P C C' \ P' X Y Z) moreover obtain A\<^sub>P \\<^sub>P where "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "A\<^sub>P \* (C, C', P, \, N, P', X, Y, Z)" by(rule freshFrame) moreover have "\ \ \\<^sub>P \ \\<^sub>P" by(blast intro: Identity Commutativity AssertionStatEqTrans) ultimately show ?case by(simp add: residualInject) next case(cCase \ P \ Cs A\<^sub>P \\<^sub>P C C' \ P' X Y Z) moreover from \bn \ \* (Cases Cs)\ \(\, P) mem Cs\ have "bn \ \* P" by(auto dest: memFreshChain) ultimately obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set(bn \) \ set(bn(p \ \))" and FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and PeqP': "(p \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinct A\<^sub>P'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (p \ \)" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "distinctPerm p" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" and "(bn(p \ \)) \* C'" and "(bn(p \ \)) \* X" and "(bn(p \ \)) \* Y" and "(bn(p \ \)) \* Z" by(rule_tac cCase) (assumption | simp (no_asm_use))+ moreover from \\\<^sub>P \ \\ have "(p \ \\<^sub>P) \ (p \ \)" by(simp add: AssertionStatEqClosed) hence "(p \ \\<^sub>P) \ \" by(simp add: permBottom) with PeqP' have "(\ \ \') \ \\<^sub>P'" by(metis Identity AssertionStatEqTrans composition' Commutativity Associativity AssertionStatEqSym) ultimately show ?case using cCase \bn \ \* P\ by(rule_tac cCase(20)) (assumption | simp)+ next case(cPar1 \ \\<^sub>Q P \ P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C C' \' PQ' X Y Z) have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact+ note \bn \' \* subject \'\ moreover from \bn \' \* (P \ Q)\ have "bn \' \* P" and "bn \' \* Q" by simp+ moreover with FrP FrQ \A\<^sub>P \* \'\ \A\<^sub>Q \* \'\ have "bn \' \* \\<^sub>P" and "bn \' \* \\<^sub>Q" by(force dest: extractFrameFreshChain)+ moreover from \bn \' \* X\ \A\<^sub>Q \* \'\ have "bn \' \* (X@A\<^sub>Q)" by simp moreover from \bn \' \* Y\ \bn \' \* \\<^sub>Q\ have "bn \' \* (\\<^sub>Q#Y)" by simp moreover from \bn \' \* Z\ \bn \' \* Q\ have "bn \' \* (Q#Z)" by simp moreover from \A\<^sub>P \* X\ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>P \* (X@A\<^sub>Q)" by simp moreover from \A\<^sub>P \* Y\ \A\<^sub>P \* \\<^sub>Q\ have "A\<^sub>P \* (\\<^sub>Q#Y)" by force moreover from \A\<^sub>P \* Z\ \A\<^sub>P \* Q\ have "A\<^sub>P \* (Q#Z)" by simp moreover from \\ \ (P' \ Q) = \' \ PQ'\ \bn \ \* Q\ \bn \' \* Q\ \bn \ \* \'\ obtain P'' where A: "\ \ P' = \' \ P''" and "PQ' = P'' \ Q" by(metis actionPar1Dest') moreover from \A\<^sub>P \* PQ'\ \PQ' = P'' \ Q\ have "A\<^sub>P \* P''" by simp ultimately obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set(bn \') \ set (bn(p \ \'))" and PeqP': "((p \ \\<^sub>P) \ \') \ \\<^sub>P'" and "distinctPerm p" and "(bn(p \ \')) \* C'" and FrP': "extractFrame P'' = \A\<^sub>P', \\<^sub>P'\" and "A\<^sub>P' \* P''" and "A\<^sub>P' \* \'" "A\<^sub>P' \* (p \ \')" and "A\<^sub>P' \* C" and "(bn(p \ \')) \* \'" and "(bn(p \ \')) \* P''" and "distinct A\<^sub>P'" and "A\<^sub>P' \* (X @ A\<^sub>Q)" and "A\<^sub>P' \* (\\<^sub>Q#Y)" and "A\<^sub>P' \* (Q#Z)" and "(bn(p \ \')) \* (X @ A\<^sub>Q)" and "(bn(p \ \')) \* (\\<^sub>Q#Y)" and "(bn(p \ \')) \* (Q#Z)" using cPar1 by(rule_tac cPar1) then have "A\<^sub>P' \* Q" and "A\<^sub>P' \* Z" and "A\<^sub>P' \* A\<^sub>Q" and "A\<^sub>P' \* X" and "A\<^sub>P' \* \\<^sub>Q" and "A\<^sub>P' \* Y" and "(bn(p \ \')) \* A\<^sub>Q" and "(bn(p \ \')) \* X" and "(bn(p \ \')) \* Y" and "(bn(p \ \')) \* Z" and "(bn(p \ \')) \* \\<^sub>Q" and "(bn(p \ \')) \* Q" by(simp del: freshChainSimps)+ from \A\<^sub>Q \* PQ'\ \PQ' = P'' \ Q\ \A\<^sub>P' \* A\<^sub>Q\ FrP' have "A\<^sub>Q \* \\<^sub>P'" by(force dest: extractFrameFreshChain) note S moreover from PeqP' have "((p \ (\\<^sub>P \ \\<^sub>Q)) \ \') \ \\<^sub>P' \ (p \ \\<^sub>Q)" by(simp add: eqvts) (metis Composition Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity) with \(bn(p \ \')) \* \\<^sub>Q\ \bn \' \* \\<^sub>Q\ S have "((p \ (\\<^sub>P \ \\<^sub>Q)) \ \') \ \\<^sub>P' \ \\<^sub>Q" by simp moreover from \PQ' = P'' \ Q\ \A\<^sub>P' \* A\<^sub>Q\ \A\<^sub>P' \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P'\ \A\<^sub>Q \* PQ'\ FrP' FrQ have "extractFrame PQ' = \(A\<^sub>P'@A\<^sub>Q), \\<^sub>P' \ \\<^sub>Q\" by simp moreover note \distinctPerm p\ \(bn(p \ \')) \* C'\ moreover from \A\<^sub>P' \* P''\ \A\<^sub>P' \* Q\ \PQ' = P'' \ Q\ have "A\<^sub>P' \* PQ'" by simp moreover note \A\<^sub>Q \* PQ'\ \A\<^sub>P' \* \'\ \A\<^sub>Q \* \'\ \A\<^sub>P' \* C\ \A\<^sub>Q \* C\ \(bn(p \ \')) \* \'\ moreover from \bn \' \* Q\ have "(bn(p \ \')) \* (p \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] bnEqvt[symmetric]) with \bn \ \* Q\ \(bn(p \ \')) \* Q\ S have "(bn(p \ \')) \* Q" by simp with \(bn(p \ \')) \* P''\ \PQ' = P'' \ Q\ have "(bn(p \ \')) \* PQ'" by simp moreover from \A\<^sub>P' \* \'\ \A\<^sub>Q \* \'\ have "(A\<^sub>P'@A\<^sub>Q) \* \'" by simp moreover from \A\<^sub>P' \* C\ \A\<^sub>Q \* C\ have "(A\<^sub>P'@A\<^sub>Q) \* C" by simp moreover from \A\<^sub>P' \* X\ \A\<^sub>Q \* X\ have "(A\<^sub>P'@A\<^sub>Q) \* X" by simp moreover from \A\<^sub>P' \* Y\ \A\<^sub>Q \* Y\ have "(A\<^sub>P'@A\<^sub>Q) \* Y" by simp moreover from \A\<^sub>P' \* Z\ \A\<^sub>Q \* Z\ have "(A\<^sub>P'@A\<^sub>Q) \* Z" by simp moreover from \A\<^sub>P' \* PQ'\ \A\<^sub>Q \* PQ'\ have "(A\<^sub>P'@A\<^sub>Q) \* PQ'" by simp moreover from \A\<^sub>P' \* \'\ \A\<^sub>Q \* \'\ have "(A\<^sub>P'@A\<^sub>Q) \* \'" by simp moreover from \A\<^sub>Q \* \'\ have "(p \ A\<^sub>Q) \* (p \ \')" by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] bnEqvt[symmetric]) with S \A\<^sub>Q \* \'\ \bn(p \ \') \* A\<^sub>Q\ have "A\<^sub>Q \* (p \ \')" by simp with \A\<^sub>P' \* (p \ \')\ \A\<^sub>Q \* \'\ \bn(p \ \') \* A\<^sub>Q\ S have "(A\<^sub>P'@A\<^sub>Q) \* (p \ \')" by simp moreover from \A\<^sub>P' \* A\<^sub>Q\ \distinct A\<^sub>P'\ \distinct A\<^sub>Q\ have "distinct(A\<^sub>P'@A\<^sub>Q)" by auto moreover note \(bn(p \ \')) \* X\ \(bn(p \ \')) \* Y\ \(bn(p \ \')) \* Z\ ultimately show ?case using cPar1 by metis next case(cPar2 \ \\<^sub>P Q \ Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C C' \' PQ' X Y Z) have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact+ note \bn \' \* subject \'\ moreover from \bn \' \* (P \ Q)\ have "bn \' \* Q" and "bn \' \* P" by simp+ moreover with FrP FrQ \A\<^sub>P \* \'\ \A\<^sub>Q \* \'\ have "bn \' \* \\<^sub>P" and "bn \' \* \\<^sub>Q" by(force dest: extractFrameFreshChain)+ moreover from \bn \' \* X\ \A\<^sub>P \* \'\ have "bn \' \* (X@A\<^sub>P)" by simp moreover from \bn \' \* Y\ \bn \' \* \\<^sub>P\ have "bn \' \* (\\<^sub>P#Y)" by simp moreover from \bn \' \* Z\ \bn \' \* P\ have "bn \' \* (P#Z)" by simp moreover from \A\<^sub>Q \* X\ \A\<^sub>P \* A\<^sub>Q\ have "A\<^sub>Q \* (X@A\<^sub>P)" by simp moreover from \A\<^sub>Q \* Y\ \A\<^sub>Q \* \\<^sub>P\ have "A\<^sub>Q \* (\\<^sub>P#Y)" by force moreover from \A\<^sub>Q \* Z\ \A\<^sub>Q \* P\ have "A\<^sub>Q \* (P#Z)" by simp moreover from \\ \ (P \ Q') = \' \ PQ'\ \bn \ \* P\ \bn \' \* P\ \bn \ \* \'\ obtain Q'' where A: "\ \ Q' = \' \ Q''" and "PQ' = P \ Q''" by(metis actionPar2Dest') moreover from \A\<^sub>Q \* PQ'\ \PQ' = P \ Q''\ have "A\<^sub>Q \* Q''" by simp ultimately obtain p \' A\<^sub>Q' \\<^sub>Q' where S: "set p \ set(bn \') \ set (bn(p \ \'))" and QeqQ': "((p \ \\<^sub>Q) \ \') \ \\<^sub>Q'" and "distinctPerm p" and "(bn(p \ \')) \* C'" and FrQ': "extractFrame Q'' = \A\<^sub>Q', \\<^sub>Q'\" and "A\<^sub>Q' \* Q''" and "A\<^sub>Q' \* \'" "A\<^sub>Q' \* (p \ \')" and "A\<^sub>Q' \* C" and "(bn(p \ \')) \* \'" and "(bn(p \ \')) \* Q''" and "distinct A\<^sub>Q'" and "A\<^sub>Q' \* (X @ A\<^sub>P)" and "A\<^sub>Q' \* (\\<^sub>P#Y)" and "A\<^sub>Q' \* (P#Z)" and "(bn(p \ \')) \* (X @ A\<^sub>P)" and "(bn(p \ \')) \* (\\<^sub>P#Y)" and "(bn(p \ \')) \* (P#Z)" using cPar2 by(rule_tac cPar2) then have "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Z" and "A\<^sub>Q' \* A\<^sub>P" and "A\<^sub>Q' \* X" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* Y" and "(bn(p \ \')) \* A\<^sub>P" and "(bn(p \ \')) \* X" and "(bn(p \ \')) \* Y" and "(bn(p \ \')) \* Z" and "(bn(p \ \')) \* \\<^sub>P" and "(bn(p \ \')) \* P" by(simp del: freshChainSimps)+ from \A\<^sub>P \* PQ'\ \PQ' = P \ Q''\ \A\<^sub>Q' \* A\<^sub>P\ FrQ' have "A\<^sub>P \* \\<^sub>Q'" by(force dest: extractFrameFreshChain) note S moreover from QeqQ' have "((p \ (\\<^sub>P \ \\<^sub>Q)) \ \') \ (p \ \\<^sub>P) \ \\<^sub>Q'" by(simp add: eqvts) (metis Composition Associativity AssertionStatEqTrans AssertionStatEqSym Commutativity) with \(bn(p \ \')) \* \\<^sub>P\ \bn \' \* \\<^sub>P\ S have "((p \ (\\<^sub>P \ \\<^sub>Q)) \ \') \ \\<^sub>P \ \\<^sub>Q'" by simp moreover from \PQ' = P \ Q''\ \A\<^sub>Q' \* A\<^sub>P\ \A\<^sub>Q' \* \\<^sub>P\ \A\<^sub>P \* \\<^sub>Q'\ \A\<^sub>P \* PQ'\ FrQ' FrP have "extractFrame PQ' = \(A\<^sub>P@A\<^sub>Q'), \\<^sub>P \ \\<^sub>Q'\" by simp moreover note \distinctPerm p\ \(bn(p \ \')) \* C'\ moreover from \A\<^sub>Q' \* Q''\ \A\<^sub>Q' \* P\ \PQ' = P \ Q''\ have "A\<^sub>Q' \* PQ'" by simp moreover note \A\<^sub>Q \* PQ'\ \A\<^sub>Q' \* \'\ \A\<^sub>Q \* \'\ \A\<^sub>Q' \* C\ \A\<^sub>Q \* C\ \(bn(p \ \')) \* \'\ moreover from \bn \' \* Q\ have "(bn(p \ \')) \* (p \ Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] bnEqvt[symmetric]) with \bn \ \* P\ \(bn(p \ \')) \* P\ S have "(bn(p \ \')) \* P" by simp with \(bn(p \ \')) \* Q''\ \PQ' = P \ Q''\ have "(bn(p \ \')) \* PQ'" by simp moreover from \A\<^sub>Q' \* \'\ \A\<^sub>P \* \'\ have "(A\<^sub>P@A\<^sub>Q') \* \'" by simp moreover from \A\<^sub>Q' \* C\ \A\<^sub>P \* C\ have "(A\<^sub>P@A\<^sub>Q') \* C" by simp moreover from \A\<^sub>Q' \* X\ \A\<^sub>P \* X\ have "(A\<^sub>P@A\<^sub>Q') \* X" by simp moreover from \A\<^sub>Q' \* Y\ \A\<^sub>P \* Y\ have "(A\<^sub>P@A\<^sub>Q') \* Y" by simp moreover from \A\<^sub>Q' \* Z\ \A\<^sub>P \* Z\ have "(A\<^sub>P@A\<^sub>Q') \* Z" by simp moreover from \A\<^sub>Q' \* PQ'\ \A\<^sub>P \* PQ'\ have "(A\<^sub>P@A\<^sub>Q') \* PQ'" by simp moreover from \A\<^sub>Q' \* \'\ \A\<^sub>P \* \'\ have "(A\<^sub>P@A\<^sub>Q') \* \'" by simp moreover from \A\<^sub>P \* \'\ have "(p \ A\<^sub>P) \* (p \ \')" by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst] bnEqvt[symmetric]) with S \A\<^sub>P \* \'\ \bn(p \ \') \* A\<^sub>P\ have "A\<^sub>P \* (p \ \')" by simp with \A\<^sub>Q' \* (p \ \')\ \A\<^sub>P \* \'\ \bn(p \ \') \* A\<^sub>P\ S have "(A\<^sub>P@A\<^sub>Q') \* (p \ \')" by simp moreover note \(bn(p \ \')) \* X\ \(bn(p \ \')) \* Y\ \(bn(p \ \')) \* Z\ moreover from \A\<^sub>Q' \* A\<^sub>P\ \distinct A\<^sub>P\ \distinct A\<^sub>Q'\ have "distinct(A\<^sub>P@A\<^sub>Q')" by auto ultimately show ?case using cPar2 by metis next case cComm1 thus ?case by(simp add: residualInject) next case cComm2 thus ?case by(simp add: residualInject) next case(cOpen \ P M xvec1 xvec2 N P' x A\<^sub>P \\<^sub>P C C' \ P'' X Y Z) from \M\\*(xvec1@x#xvec2)\\N\ \ P' = \ \ P''\ \x \ xvec1\ \x \ xvec2\ \x \ \\ \x \ P''\ \distinct(bn \)\ \A\<^sub>P \* \\ \x \ \\ obtain yvec1 y yvec2 N' where yvecEq: "bn \ = yvec1@y#yvec2" and P'eqP'': "\\*(xvec1@xvec2)\N \' P' = \\*(yvec1@yvec2)\([(x, y)] \ N') \' ([(x, y)] \ P'')" and "A\<^sub>P \* N'" and Subj: "subject \ = Some M" and "x \ N'" and \eq: "\ = M\\*(yvec1@y#yvec2)\\N'\" apply(cases rule: actionCases[where \=\]) apply(auto simp add: residualInject) apply(rule boundOutputOpenDest) by assumption auto note \A\<^sub>P \* P\ \A\<^sub>P \* M\ moreover from Subj yvecEq \bn \ \* subject \\ have "yvec1 \* M" "yvec2 \* M" by simp+ moreover from yvecEq \A\<^sub>P \* \\ have "A\<^sub>P \* (yvec1@yvec2)" by simp moreover note \A\<^sub>P \* C\ moreover from yvecEq \bn \ \* \\x\P\ \x \ \\ have "(yvec1@yvec2) \* P" by simp moreover from yvecEq \bn \ \* C'\ \bn \ \* X\ \bn \ \* Y\ \bn \ \* Z\ \distinct(bn \)\ \x \ \\ have "(yvec1@yvec2) \* C'" and "(yvec1@yvec2) \* (x#y#X)" and "(yvec1@yvec2) \* Y" and "(yvec1@yvec2) \* Z" by simp+ moreover from \A\<^sub>P \* X\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ yvecEq have "A\<^sub>P \* (x#y#X)" by simp moreover note \A\<^sub>P \* Y\ \A\<^sub>P \* Z\ moreover from \A\<^sub>P \* N'\ \A\<^sub>P \* P''\ \x \ A\<^sub>P\ \A\<^sub>P \* \\ yvecEq have "A\<^sub>P \* ([(x, y)] \ N')" and "A\<^sub>P \* ([(x, y)] \ P'')" by simp+ moreover from yvecEq \distinct(bn \)\ have "distinct(yvec1@yvec2)" by simp moreover from P'eqP'' have "M\\*(xvec1@xvec2)\\N\ \ P' = M\\*(yvec1@yvec2)\\([(x, y)] \ N')\ \ ([(x, y)] \ P'')" by(simp add: residualInject) ultimately obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set (yvec1@yvec2) \ set (p \ (yvec1@yvec2))" and PeqP': "((p \ \\<^sub>P) \ \') \ \\<^sub>P'" and "distinctPerm p" and "(p \ (yvec1@yvec2)) \* C'" and FrP': "extractFrame([(x, y)] \ P'') = \A\<^sub>P', \\<^sub>P'\" and "A\<^sub>P' \* ([(x, y)] \ P'')" and "A\<^sub>P' \* ([(x, y)] \ N')" and "A\<^sub>P' \* C" and "(p \ (yvec1@yvec2)) \* ([(x, y)] \ N')" and "A\<^sub>P' \* M" and "(p \ (yvec1@yvec2)) \* (yvec1@yvec2)" and "(p \ (yvec1@yvec2)) \* M" and "distinct A\<^sub>P'" and "(p \ (yvec1@yvec2)) \* ([(x, y)] \ P'')" and "(yvec1@yvec2) \* A\<^sub>P'" and "(p \ (yvec1@yvec2)) \* A\<^sub>P'" and "A\<^sub>P' \* (x#y#X)" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "(p \ (yvec1@yvec2)) \* (x#y#X)" and "(p \ (yvec1@yvec2)) \* Y" and "(p \ (yvec1@yvec2)) \* Z" using \A\<^sub>P \* C'\ by(rule_tac cOpen(4)[where bd="x#y#X"]) (assumption | simp_all)+ from \A\<^sub>P' \* (x#y#X)\ have "x \ A\<^sub>P'" and "y \ A\<^sub>P'" and "A\<^sub>P' \* X" by simp+ from \(p \ (yvec1@yvec2)) \* (x#y#X)\ have "x \ (p \ (yvec1@yvec2))" and "y \ (p \ (yvec1@yvec2))" and "(p \ (yvec1@yvec2)) \* X" by simp+ from \x \ \\ yvecEq have "x \ yvec1" and "x \ y" and "x \ yvec2" by simp+ from \distinct(bn \)\ yvecEq have "yvec1 \* yvec2" and "y \ yvec1" and "y \ yvec2" by simp+ from \bn \ \* C'\ yvecEq have "yvec1 \* C'" and "y \ C'" and "yvec2 \* C'" by simp+ from S \x \ \\ \x \ p \ (yvec1@yvec2)\ yvecEq have "x \ p" by(rule_tac freshAlphaSwap) (assumption | simp)+ from S \distinct(bn \)\ \y \ p \ (yvec1@yvec2)\ yvecEq have "y \ p" by(rule_tac freshAlphaSwap) (assumption | simp)+ from yvecEq S \x \ yvec1\ \x \ yvec2\ \y \ yvec1\ \y \ yvec2\ \x \ p \ (yvec1@yvec2)\ \y \ p \ (yvec1@yvec2)\ have "set ((y, x)#p) \ set(bn \) \ set(bn(((y, x)#p) \ \))" apply(simp add: bnEqvt[symmetric]) by(auto simp add: eqvts calc_atm) moreover from PeqP' have "([(y, x)] \ ((p \ \\<^sub>P) \ \')) \ [(y, x)] \ \\<^sub>P'" by(simp add: AssertionStatEqClosed) hence "(((y, x)#p) \ \\<^sub>P) \ ([(y, x)] \ \') \ ([(y, x)] \ \\<^sub>P')" by(simp add: eqvts) moreover from \distinctPerm p\ S \x \ y\ \x \ p\ \y \ p\ have "distinctPerm((y, x)#p)" by simp moreover from FrP' have "([(x, y)] \ (extractFrame([(x, y)] \ P''))) = ([(x, y)] \ \A\<^sub>P', \\<^sub>P'\)" by simp with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ have "extractFrame P'' = \A\<^sub>P', ([(y, x)] \ \\<^sub>P')\" by(simp add: eqvts name_swap) moreover from \x \ p\ \y \ p\ \x \ N'\ \(p \ (yvec1@yvec2)) \* ([(x, y)] \ N')\ have "([(y, x)] \ p \ (yvec1@yvec2)) \* ([(y, x)] \ [(x, y)] \ N')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) hence "(((y, x)#p) \ (yvec1@yvec2)) \* N'" by(simp add: name_swap) with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ C\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \x \ N'\ yvecEq have "bn(((y, x)#p) \ \) \* N'" by(simp add: bnEqvt[symmetric]) (simp add: eqvts perm_compose calc_atm freshChainSimps) moreover from \x \ p\ \y \ p\ \x \ N'\ \(p \ (yvec1@yvec2)) \* ([(x, y)] \ P'')\ have "([(y, x)] \ p \ (yvec1@yvec2)) \* ([(y, x)] \ [(x, y)] \ P'')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) hence "(((y, x)#p) \ (yvec1@yvec2)) \* P''" by(simp add: name_swap) with \x \ yvec1\ \x \ yvec2\ \x \ y\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \x \ P''\ yvecEq have "bn(((y, x)#p) \ \) \* P''" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) moreover from \x \ p\ \y \ p\ \x \ A\<^sub>P'\ \(p \ (yvec1@yvec2)) \* A\<^sub>P'\ have "(p \ (yvec1@x#yvec2)) \* A\<^sub>P'" by(simp add: eqvts freshChainSimps) hence "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ A\<^sub>P')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ A\<^sub>P'\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \y \ A\<^sub>P'\ yvecEq have "bn(((y, x)#p) \ \) \* A\<^sub>P'" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) moreover from \x \ p\ \y \ p\ \x \ C'\ \(p \ (yvec1@yvec2)) \* C'\ have "(p \ (yvec1@x#yvec2)) \* C'" by(simp add: eqvts freshChainSimps) hence "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ C')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ C'\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \y \ C'\ yvecEq have "bn(((y, x)#p) \ \) \* C'" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) moreover from \x \ p\ \y \ p\ \x \ X\ \(p \ (yvec1@yvec2)) \* X\ have "(p \ (yvec1@x#yvec2)) \* X" by(simp add: eqvts freshChainSimps) hence "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ X)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ X\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* X\ yvecEq have "bn(((y, x)#p) \ \) \* X" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) moreover from \x \ p\ \y \ p\ \x \ Y\ \(p \ (yvec1@yvec2)) \* Y\ have "(p \ (yvec1@x#yvec2)) \* Y" by(simp add: eqvts freshChainSimps) hence "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ Y)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ Y\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* Y\ yvecEq have "bn(((y, x)#p) \ \) \* Y" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) moreover from \x \ p\ \y \ p\ \x \ Z\ \(p \ (yvec1@yvec2)) \* Z\ have "(p \ (yvec1@x#yvec2)) \* Z" by(simp add: eqvts freshChainSimps) hence "([(y, x)] \ p \ (yvec1@x#yvec2)) \* ([(y, x)] \ Z)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \x \ yvec1\ \x \ yvec2\ \x \ y\ \x \ Z\ \y \ yvec1\ \y \ yvec2\ \x \ p\ \y \ p\ \bn \ \* Z\ yvecEq have "bn(((y, x)#p) \ \) \* Z" by(simp add: bnEqvt[symmetric]) (simp add: perm_compose calc_atm eqvts freshChainSimps) moreover from \(yvec1@yvec2) \* A\<^sub>P'\ \y \ A\<^sub>P'\ yvecEq have "bn \ \* A\<^sub>P'" by simp moreover from \A\<^sub>P' \* ([(x, y)] \ N')\ have "([(x, y)] \ A\<^sub>P') \* ([(x, y)] \ [(x, y)] \ N')" by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ have "A\<^sub>P' \* N'" by simp with \A\<^sub>P' \* M\ \(yvec1@yvec2) \* A\<^sub>P'\ \y \ A\<^sub>P'\ \eq have "A\<^sub>P' \* \" by simp moreover hence "(((y, x)#p) \ A\<^sub>P') \* (((y, x)#p) \ \)" by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ S \(yvec1@yvec2) \* A\<^sub>P'\ \(p \ (yvec1@yvec2)) \* A\<^sub>P'\ have "A\<^sub>P' \* (((y, x)#p) \ \)" by(simp add: eqvts) moreover from \A\<^sub>P' \* ([(x, y)] \ P'')\ have "([(x, y)] \ A\<^sub>P') \* ([(x, y)] \ [(x, y)] \ P'')" by(simp only: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \x \ A\<^sub>P'\ \y \ A\<^sub>P'\ have "A\<^sub>P' \* P''" by simp moreover from yvecEq \eq \(p \ (yvec1@yvec2)) \* (yvec1@yvec2)\ \y \ p\ \x \ \\ S \(p \ (yvec1@yvec2)) \* M\\(p \ (yvec1@yvec2)) \* ([(x, y)] \ N')\ \y \ yvec1\\y \ yvec2\ \x \ p\ have "bn(((y, x)#p) \ \) \* \" apply(simp add: eqvts del: set_append) apply(intro conjI) apply(simp add: perm_compose eqvts del: set_append) apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) apply(simp add: perm_compose eqvts del: set_append) apply(simp add: perm_compose eqvts swapStarFresh del: set_append) apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) apply(simp add: perm_compose freshChainSimps(6) swapStarFresh calc_atm eqvts del: set_append) apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) apply(subst pt_fresh_star_bij[symmetric, OF pt_name_inst, OF at_name_inst, where pi="[(x, y)]"]) apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) apply(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) apply(subst pt_fresh_star_bij[symmetric, OF pt_name_inst, OF at_name_inst, where pi="[(x, y)]"]) by(simp add: perm_compose freshChainSimps(6) calc_atm eqvts del: set_append) moreover note \A\<^sub>P' \* C\ \A\<^sub>P' \* X\ \A\<^sub>P' \* Y\ \A\<^sub>P' \* Z\ \distinct A\<^sub>P'\ ultimately show ?case by(rule_tac cOpen) next case(cScope \ P \ P' x A\<^sub>P \\<^sub>P C C' \' P'' X Y Z) from \\ \ \\x\P' = \' \ P''\ \x \ \\ \x \ \'\ obtain P''' where "\ \ P' = \' \ P'''" and "P'' = \\x\P'''" apply(cases rule: actionCases[where \=\]) apply(auto simp add: residualInject) by(metis boundOutputScopeDest) then obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set(bn \') \ set(bn(p \ \'))" and PeqP': "((p \ \\<^sub>P) \ \') \ \\<^sub>P'" and "distinctPerm p" and "bn(p \ \') \* C'" and FrP': "extractFrame P''' = \A\<^sub>P', \\<^sub>P'\" and "A\<^sub>P' \* P'''" and "A\<^sub>P' \* \'" and "A\<^sub>P' \* (p \ \')" and "A\<^sub>P' \* C" and "distinct A\<^sub>P'" and "bn(p \ \') \* P'''" and "A\<^sub>P' \* (x#X)" and "A\<^sub>P' \* Y" and "bn(p \ \') \* \'" and "A\<^sub>P' \* Z" and "bn(p \ \') \* (x#X)" and "bn(p \ \') \* Y" and "bn(p \ \') \* Z" using cScope by(rule_tac cScope) (assumption | simp)+ from \A\<^sub>P' \* (x#X)\ have "x \ A\<^sub>P'" and "A\<^sub>P' \* X" by simp+ from \bn(p \ \') \* (x#X)\ have "x \ bn(p \ \')" and "bn(p \ \') \* X" by simp+ note S PeqP' \distinctPerm p\ \bn(p \ \') \* C'\ moreover from FrP' \P'' = \\x\P'''\ have "extractFrame P'' = \(x#A\<^sub>P'), \\<^sub>P'\" by simp moreover from \A\<^sub>P' \* P'''\ \P'' = \\x\P'''\ \x \ A\<^sub>P'\ have "(x#A\<^sub>P') \* P''" by(simp add: abs_fresh) moreover from \A\<^sub>P' \* \'\ \A\<^sub>P' \* C\ \x \ \'\ \x \ C\ have "(x#A\<^sub>P') \* \'" and "(x#A\<^sub>P') \* C" by simp+ moreover note \bn(p \ \') \* \'\ moreover from \bn(p \ \') \* P'''\ \P'' = \\x\P'''\ \x \ bn(p \ \')\ have "bn(p \ \') \* P''" by simp moreover from \A\<^sub>P' \* \'\ \x \ \'\ have "(x#A\<^sub>P') \* \'" by simp moreover from \A\<^sub>P' \* (p \ \')\ \x \ \'\ S \x \ bn(p \ \')\ have "(x#A\<^sub>P') \* (p \ \')" by(simp add: subjectEqvt[symmetric] bnEqvt[symmetric] okjectEqvt[symmetric] freshChainSimps) moreover from \A\<^sub>P' \* X\ \x \ X\ have "(x#A\<^sub>P') \* X" by simp+ moreover from \A\<^sub>P' \* Y\ \x \ Y\ have "(x#A\<^sub>P') \* Y" by simp+ moreover from \A\<^sub>P' \* Z\ \x \ Z\ have "(x#A\<^sub>P') \* Z" by simp+ moreover note \bn(p \ \') \* X\ \bn(p \ \') \* Y\ \bn(p \ \') \* Z\ moreover from \distinct A\<^sub>P'\ \x \ A\<^sub>P'\ have "distinct(x#A\<^sub>P')" by simp ultimately show ?case by(rule_tac cScope) next case(cBang \ P A\<^sub>P \\<^sub>P C C' \ P' X Y Z) then obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set(bn \) \ set(bn(p \ \))" and FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and PeqP': "(p \ (\\<^sub>P \ \)) \ \' \ \\<^sub>P'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (p \ \)" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "distinct A\<^sub>P'" and "distinctPerm p" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" and "(bn(p \ \)) \* C'" and "(bn(p \ \)) \* X" and "(bn(p \ \)) \* Y" and "(bn(p \ \)) \* Z" by(rule_tac cBang) (assumption | simp (no_asm_use))+ moreover from \\\<^sub>P \ \\ have "(p \ \\<^sub>P) \ (p \ \)" by(simp add: AssertionStatEqClosed) hence "(p \ \\<^sub>P) \ \" by(simp add: permBottom) with PeqP' have "(\ \ \') \ \\<^sub>P'" by(simp add: eqvts permBottom) (metis Identity AssertionStatEqTrans composition' Commutativity Associativity AssertionStatEqSym) ultimately show ?case using cBang apply(rule_tac cBang(18)) by(assumption | simp)+ qed with A have ?thesis by blast } moreover have "bn \ \* ([]::name list)" and "bn \ \* ([]::'b list)" and "bn \ \* ([]::('a, 'b, 'c) psi list)" and "A\<^sub>P \* ([]::name list)" and "A\<^sub>P \* ([]::'b list)" and "A\<^sub>P \* ([]::('a, 'b, 'c) psi list)" by simp+ ultimately show ?thesis by blast qed lemma expandTauFrame: fixes \ :: 'b and P :: "('a, 'b,'c) psi" and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and C :: "'d::fs_name" assumes "\ \ P \\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "A\<^sub>P \* P" and "A\<^sub>P \* C" obtains \' A\<^sub>P' \\<^sub>P' where "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "\\<^sub>P \ \' \ \\<^sub>P'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "distinct A\<^sub>P'" proof - assume A: "\A\<^sub>P' \\<^sub>P' \'. \extractFrame P' = \A\<^sub>P', \\<^sub>P'\; \\<^sub>P \ \' \ \\<^sub>P'; A\<^sub>P' \* C; A\<^sub>P' \* P'; distinct A\<^sub>P'\ \ thesis" from \\ \ P \\ \P'\ \A\<^sub>P \* P\ have "A\<^sub>P \* P'" by(rule tauFreshChainDerivative) { fix X :: "name list" and Y :: "'b list" and Z :: "('a, 'b, 'c) psi list" assume "A\<^sub>P \* X" and "A\<^sub>P \* Y" and "A\<^sub>P \* Z" with assms \A\<^sub>P \* P'\ obtain \' A\<^sub>P' \\<^sub>P' where "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "\\<^sub>P \ \' \ \\<^sub>P'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" and "distinct A\<^sub>P'" proof(nominal_induct avoiding: C X Y Z arbitrary: thesis rule: tauFrameInduct) case(cAlpha \ P P' A\<^sub>P \\<^sub>P p C X Y Z) then obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "\\<^sub>P \ \' \ \\<^sub>P'" and "distinct A\<^sub>P'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" by metis have S: "set p \ set A\<^sub>P \ set(p \ A\<^sub>P)" by fact from FrP' have "(p \ extractFrame P') = p \ \A\<^sub>P', \\<^sub>P'\" by simp with \A\<^sub>P \* P'\ \(p \ A\<^sub>P) \* P'\ S have "extractFrame P' = \(p \ A\<^sub>P'), (p \ \\<^sub>P')\" by(simp add: eqvts) moreover from \\\<^sub>P \ \' \ \\<^sub>P'\ have "(p \ (\\<^sub>P \ \')) \ (p \ \\<^sub>P')" by(rule AssertionStatEqClosed) hence "(p \ \\<^sub>P) \ (p \ \') \ (p \ \\<^sub>P')" by(simp add: eqvts) moreover from \A\<^sub>P' \* C\ have "(p \ A\<^sub>P') \* (p \ C)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* C\ \(p \ A\<^sub>P) \* C\ S have "(p \ A\<^sub>P') \* C" by simp moreover from \A\<^sub>P' \* P'\ have "(p \ A\<^sub>P') \* (p \ P')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* P'\ \(p \ A\<^sub>P) \* P'\ S have "(p \ A\<^sub>P') \* P'" by simp moreover from \A\<^sub>P' \* X\ have "(p \ A\<^sub>P') \* (p \ X)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* X\ \(p \ A\<^sub>P) \* X\ S have "(p \ A\<^sub>P') \* X" by simp moreover from \A\<^sub>P' \* Y\ have "(p \ A\<^sub>P') \* (p \ Y)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* Y\ \(p \ A\<^sub>P) \* Y\ S have "(p \ A\<^sub>P') \* Y" by simp moreover from \A\<^sub>P' \* Z\ have "(p \ A\<^sub>P') \* (p \ Z)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>P \* Z\ \(p \ A\<^sub>P) \* Z\ S have "(p \ A\<^sub>P') \* Z" by simp moreover from \distinct A\<^sub>P'\ have "distinct(p \ A\<^sub>P')" by simp ultimately show ?case by(rule cAlpha) next case(cCase \ P P' \ Cs A\<^sub>P \\<^sub>P C B Y Z thesis) then obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "\\<^sub>P \ \' \ \\<^sub>P'" and "distinct A\<^sub>P'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* B" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" by(rule_tac cCase) (assumption | simp (no_asm_use))+ with \\\<^sub>P \ \\ \\\<^sub>P \ \' \ \\<^sub>P'\ have "\ \ \' \ \\<^sub>P'" by(metis Identity AssertionStatEqTrans composition' Commutativity Associativity AssertionStatEqSym) thus ?case using FrP' \A\<^sub>P' \* P'\ \A\<^sub>P' \* C\ \A\<^sub>P' \* B\ \A\<^sub>P' \* Y\ \A\<^sub>P' \* Z\ \distinct A\<^sub>P'\ using cCase by force next case(cPar1 \ \\<^sub>Q P P' A\<^sub>Q Q A\<^sub>P \\<^sub>P C X Y Z) moreover from \A\<^sub>P \* X\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* Y\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* Q\ \A\<^sub>P \* Z\ have "A\<^sub>P \* (X@A\<^sub>Q)" and "A\<^sub>P \* (\\<^sub>Q#Y)" and "A\<^sub>P \* (Q#Z)" by simp+ ultimately obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "distinct A\<^sub>P'" and "\\<^sub>P \ \' \ \\<^sub>P'" and "A\<^sub>P \* P" and "A\<^sub>P \* C" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* (X@A\<^sub>Q)" and "A\<^sub>P' \* (\\<^sub>Q#Y)" and "A\<^sub>P' \* (Q#Z)" by metis hence "A\<^sub>P' \* X" and "A\<^sub>P' \* A\<^sub>Q" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* \\<^sub>Q" and "A\<^sub>P' \* Q" and "A\<^sub>P' \* Z" by simp+ from \A\<^sub>P' \* A\<^sub>Q\ \A\<^sub>Q \* P'\ FrP' have "A\<^sub>Q \* \\<^sub>P'" by(force dest: extractFrameFreshChain) with \A\<^sub>P' \* \\<^sub>Q\ \A\<^sub>P' \* A\<^sub>Q\ \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ FrP' have "extractFrame(P' \ Q) = \(A\<^sub>P'@A\<^sub>Q), \\<^sub>P' \ \\<^sub>Q\" by simp moreover from \\\<^sub>P \ \' \ \\<^sub>P'\have "(\\<^sub>P \ \\<^sub>Q) \ \' \ \\<^sub>P' \ \\<^sub>Q" by(metis Associativity Commutativity Composition AssertionStatEqTrans AssertionStatEqSym) moreover from \A\<^sub>P' \* C\ \A\<^sub>Q \* C\ have "(A\<^sub>P'@A\<^sub>Q) \* C" by simp moreover from \A\<^sub>P' \* P'\ \A\<^sub>Q \* P'\ \A\<^sub>P' \* Q\ \A\<^sub>Q \* Q\ have "(A\<^sub>P'@A\<^sub>Q) \* (P' \ Q)" by simp moreover from \A\<^sub>P' \* X\ \A\<^sub>Q \* X\ have "(A\<^sub>P'@A\<^sub>Q) \* X" by simp moreover from \A\<^sub>P' \* Y\ \A\<^sub>Q \* Y\ have "(A\<^sub>P'@A\<^sub>Q) \* Y" by simp moreover from \A\<^sub>P' \* Z\ \A\<^sub>Q \* Z\ have "(A\<^sub>P'@A\<^sub>Q) \* Z" by simp moreover from \A\<^sub>P' \* A\<^sub>Q\ \distinct A\<^sub>P'\ \distinct A\<^sub>Q\ have "distinct(A\<^sub>P'@A\<^sub>Q)" by simp ultimately show ?case by(rule cPar1) next case(cPar2 \ \\<^sub>P Q Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q C X Y Z) moreover from \A\<^sub>Q \* X\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* Y\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Z\ have "A\<^sub>Q \* (X@A\<^sub>P)" and "A\<^sub>Q \* (\\<^sub>P#Y)" and "A\<^sub>Q \* (P#Z)" by(simp add: freshChainSimps)+ ultimately obtain \' A\<^sub>Q' \\<^sub>Q' where FrQ': "extractFrame Q' = \A\<^sub>Q', \\<^sub>Q'\" and "distinct A\<^sub>Q'" and "\\<^sub>Q \ \' \ \\<^sub>Q'"and "A\<^sub>Q' \* C" and "A\<^sub>Q' \* Q'" and "A\<^sub>Q' \* (X@A\<^sub>P)" and "A\<^sub>Q' \* (\\<^sub>P#Y)" and "A\<^sub>Q' \* (P#Z)" by metis hence "A\<^sub>Q' \* X" and "A\<^sub>Q' \* A\<^sub>P" and "A\<^sub>Q' \* Y" and "A\<^sub>Q' \* \\<^sub>P" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* Z" by simp+ from \A\<^sub>Q' \* A\<^sub>P\ \A\<^sub>P \* Q'\ FrQ' have "A\<^sub>P \* \\<^sub>Q'" by(force dest: extractFrameFreshChain) with \A\<^sub>Q' \* \\<^sub>P\ \A\<^sub>Q' \* A\<^sub>P\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ FrQ' have "extractFrame(P \ Q') = \(A\<^sub>P@A\<^sub>Q'), \\<^sub>P \ \\<^sub>Q'\" by simp moreover from \\\<^sub>Q \ \' \ \\<^sub>Q'\have "(\\<^sub>P \ \\<^sub>Q) \ \' \ \\<^sub>P \ \\<^sub>Q'" by(metis Associativity Commutativity Composition AssertionStatEqTrans AssertionStatEqSym) moreover from \A\<^sub>P \* C\ \A\<^sub>Q' \* C\ have "(A\<^sub>P@A\<^sub>Q') \* C" by simp moreover from \A\<^sub>P \* P\ \A\<^sub>Q' \* P\ \A\<^sub>P \* Q'\ \A\<^sub>Q' \* Q'\ have "(A\<^sub>P@A\<^sub>Q') \* (P \ Q')" by simp moreover from \A\<^sub>P \* X\ \A\<^sub>Q' \* X\ have "(A\<^sub>P@A\<^sub>Q') \* X" by simp moreover from \A\<^sub>P \* Y\ \A\<^sub>Q' \* Y\ have "(A\<^sub>P@A\<^sub>Q') \* Y" by simp moreover from \A\<^sub>P \* Z\ \A\<^sub>Q' \* Z\ have "(A\<^sub>P@A\<^sub>Q') \* Z" by simp moreover from \A\<^sub>Q' \* A\<^sub>P\ \distinct A\<^sub>P\ \distinct A\<^sub>Q'\ have "distinct(A\<^sub>P@A\<^sub>Q')" by simp ultimately show ?case by(rule cPar2) next case(cComm1 \ \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q C X Y Z) have PTrans: "\ \ \\<^sub>Q \ P \M\N\ \ P'" and QTrans: "\ \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'" by fact+ from PTrans \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>P \* N\ \A\<^sub>P \* M\ \A\<^sub>P \* Q'\ \A\<^sub>P \* C\ \A\<^sub>P \* X\ \A\<^sub>P \* Y\ \A\<^sub>P \* Z\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* xvec\ obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and PeqP': "\\<^sub>P \ \' \ \\<^sub>P'" and "A\<^sub>P' \* Q'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "distinct A\<^sub>P'" and "A\<^sub>P' \* Z" and "A\<^sub>P' \* A\<^sub>Q" and "A\<^sub>P' \* xvec" and "A\<^sub>P' \* P'" by(rule_tac C="(Q', C, X, Y, Z, A\<^sub>Q, xvec)" and C'="(Q', C, X, Y, Z, A\<^sub>Q, xvec)" in expandNonTauFrame) auto moreover from QTrans have "distinct xvec" by(auto dest: boundOutputDistinct) from QTrans \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* K\ \xvec \* K\ \distinct xvec\ \A\<^sub>P' \* A\<^sub>Q\ \A\<^sub>P' \* xvec\ \A\<^sub>Q \* Q\ \xvec \* Q\ \A\<^sub>Q \* \\<^sub>P\ \xvec \* \\<^sub>P\ \A\<^sub>Q \* N\ \A\<^sub>Q \* C\ \A\<^sub>Q \* X\ \A\<^sub>Q \* Y\ \A\<^sub>Q \* Z\ \xvec \* P\ \xvec \* C\ \xvec \* X\ \xvec \* Y\ \xvec \* Z\ obtain p \'' A\<^sub>Q' \\<^sub>Q' where S: "set p \ set xvec \ set(p \ xvec)" and QeqQ': "(p \ \\<^sub>Q) \ \'' \ \\<^sub>Q'" and FrQ': "extractFrame Q' = \A\<^sub>Q', \\<^sub>Q'\" and "A\<^sub>Q' \* A\<^sub>P'" and "A\<^sub>Q' \* C" and "A\<^sub>Q' \* X" and "A\<^sub>Q' \* Y" and "A\<^sub>Q' \* Z" and "A\<^sub>Q' \* P" and "A\<^sub>Q' \* N" and "distinct A\<^sub>Q'" and "(p \ xvec) \* A\<^sub>P'" and "(p \ xvec) \* C" and "(p \ xvec) \* X" and "(p \ xvec) \* Y" and "(p \ xvec) \* P" and "(p \ xvec) \* Z" and "(p \ xvec) \* N" and "(p \ xvec) \* \\<^sub>P" and "(p \ xvec) \* A\<^sub>Q'"and "(p \ xvec) \* Q'" and "distinctPerm p" and "A\<^sub>Q' \* xvec" and "A\<^sub>Q' \* Q'" by(rule_tac C="(P, C, X, Y, Z, A\<^sub>P', \\<^sub>P)" and C'="(P, C, X, Y, Z, A\<^sub>P', \\<^sub>P)" in expandNonTauFrame) (assumption | simp)+ from PTrans \A\<^sub>Q' \* P\ \A\<^sub>Q' \* N\ \(p \ xvec) \* P\ \(p \ xvec) \* N\ have "A\<^sub>Q' \* P'" and "(p \ xvec) \* P'" by(force dest: inputFreshChainDerivative)+ with FrP' \A\<^sub>Q' \* A\<^sub>P'\ \(p \ xvec) \* A\<^sub>P'\ have "A\<^sub>Q' \* \\<^sub>P'" and "(p \ xvec) \* \\<^sub>P'" by(force dest: extractFrameFreshChain)+ from FrQ' \A\<^sub>Q' \* A\<^sub>P'\ \A\<^sub>P' \* Q'\ \(p \ xvec) \* A\<^sub>Q'\ \(p \ xvec) \* Q'\ have "A\<^sub>P' \* \\<^sub>Q'" and "(p \ xvec) \* \\<^sub>Q'" by(force dest: extractFrameFreshChain)+ have "extractFrame(\\*xvec\(P' \ Q')) = \((p \ xvec)@A\<^sub>P'@A\<^sub>Q'), (p \ \\<^sub>P') \ (p \ \\<^sub>Q')\" proof - from FrP' FrQ' \A\<^sub>P' \* \\<^sub>Q'\ \A\<^sub>Q' \* A\<^sub>P'\ \A\<^sub>Q' \* \\<^sub>P'\ have "extractFrame(P' \ Q') = \(A\<^sub>P'@A\<^sub>Q'), \\<^sub>P' \ \\<^sub>Q'\" by simp hence "extractFrame(\\*xvec\(P' \ Q')) = \(xvec@A\<^sub>P'@A\<^sub>Q'), \\<^sub>P' \ \\<^sub>Q'\" by(induct xvec) auto moreover from \(p \ xvec) \* \\<^sub>P'\ \(p \ xvec) \* \\<^sub>Q'\ S have "\\*xvec\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert (\\<^sub>P' \ \\<^sub>Q'))) = \\*(p \ xvec)\(p \ \\*(A\<^sub>P'@A\<^sub>Q')\(FAssert(\\<^sub>P' \ \\<^sub>Q')))" by(rule_tac frameChainAlpha) (auto simp add: fresh_star_def frameResChainFresh) hence "\\*xvec\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert (\\<^sub>P' \ \\<^sub>Q'))) = \\*(p \ xvec)\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert((p \ \\<^sub>P') \ (p \ \\<^sub>Q'))))" using \A\<^sub>P' \* xvec\ \(p \ xvec) \* A\<^sub>P'\ \A\<^sub>Q' \* xvec\ \(p \ xvec) \* A\<^sub>Q'\ S by(fastforce simp add: eqvts) ultimately show ?thesis by(simp add: frameChainAppend) qed moreover have "(\\<^sub>P \ \\<^sub>Q) \ ((p \ \') \ (p \ \'')) \ (p \ \\<^sub>P') \ (p \ \\<^sub>Q')" proof - have "(\\<^sub>P \ (p \ \\<^sub>Q)) \ (\' \ \'') \ (\\<^sub>P \ \') \ ((p \ \\<^sub>Q) \ \'')" by(metis Associativity Commutativity Composition AssertionStatEqTrans) moreover from PeqP' QeqQ' have "(\\<^sub>P \ \') \ ((p \ \\<^sub>Q) \ \'') \ \\<^sub>P' \ \\<^sub>Q'" by(metis Associativity Commutativity Composition AssertionStatEqTrans) ultimately have "(\\<^sub>P \ (p \ \\<^sub>Q)) \ (\' \ \'') \ \\<^sub>P' \ \\<^sub>Q'" by(metis AssertionStatEqTrans) hence "(p \ ((\\<^sub>P \ (p \ \\<^sub>Q)) \ (\' \ \''))) \ (p \ (\\<^sub>P' \ \\<^sub>Q'))" by(rule AssertionStatEqClosed) with \xvec \* \\<^sub>P\ \(p \ xvec) \* \\<^sub>P\ S \distinctPerm p\ show ?thesis by(simp add: eqvts) qed moreover from \(p \ xvec) \* C\ \A\<^sub>P' \* C\ \A\<^sub>Q' \* C\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* C" by simp moreover from \(p \ xvec) \* X\ \A\<^sub>P' \* X\ \A\<^sub>Q' \* X\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* X" by simp moreover from \(p \ xvec) \* Y\ \A\<^sub>P' \* Y\ \A\<^sub>Q' \* Y\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* Y" by simp moreover from \(p \ xvec) \* Z\ \A\<^sub>P' \* Z\ \A\<^sub>Q' \* Z\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* Z" by simp moreover from \(p \ xvec) \* P'\ \(p \ xvec) \* Q'\ \A\<^sub>P' \* P'\ \A\<^sub>P' \* Q' \\A\<^sub>Q' \* P'\ \A\<^sub>Q' \* Q'\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* (\\*xvec\(P' \ Q'))" by(auto simp add: resChainFresh fresh_star_def) moreover from \(p \ xvec) \* A\<^sub>P'\ \(p \ xvec) \* A\<^sub>Q'\ \A\<^sub>Q' \* A\<^sub>P'\ \distinct xvec\ \distinct A\<^sub>P'\ \distinct A\<^sub>Q'\ have "distinct((p \ xvec)@A\<^sub>P'@A\<^sub>Q')" by auto (simp add: name_list_supp fresh_star_def fresh_def)+ ultimately show ?case using cComm1 by metis next case(cComm2 \ \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q C X Y Z) have PTrans: "\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'" and QTrans: "\ \ \\<^sub>P \ Q \K\N\ \ Q'" by fact+ from PTrans have "distinct xvec" by(auto dest: boundOutputDistinct) from PTrans \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \xvec \* Q\ \distinct xvec\ \xvec \* M\ \A\<^sub>P \* C\ \A\<^sub>P \* X\ \A\<^sub>P \* Y\ \A\<^sub>P \* Z\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* xvec\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* M\ \A\<^sub>P \* N\ \xvec \* P\ \xvec \* C\ \xvec \* X\ \xvec \* Y\ \xvec \* Z\ \A\<^sub>Q \* xvec\ \xvec \* \\<^sub>Q\ obtain p \' A\<^sub>P' \\<^sub>P' where S: "set p \ set xvec \ set(p \ xvec)" and FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and PeqP': "(p \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinct A\<^sub>P'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* X" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* N" and "A\<^sub>P' \* Q" and "(p \ xvec) \* Q" and "A\<^sub>P' \* Z" and "A\<^sub>P' \* A\<^sub>Q" and "A\<^sub>P' \* xvec" and "A\<^sub>P' \* P'" and "(p \ xvec) \* N" and "(p \ xvec) \* \\<^sub>Q" and "(p \ xvec) \* A\<^sub>P'" and "(p \ xvec) \* C" and "(p \ xvec) \* X" and "(p \ xvec) \* A\<^sub>Q" and "(p \ xvec) \* Y" and "(p \ xvec) \* Z" and "(p \ xvec) \* P'" and "distinctPerm p" by(rule_tac C="(C, X, Y, Z, A\<^sub>Q, Q, \\<^sub>Q)" and C'="(C, X, Y, Z, A\<^sub>Q, Q, \\<^sub>Q)" in expandNonTauFrame) (assumption | simp)+ from QTrans \extractFrame Q = \A\<^sub>Q, \\<^sub>Q\\ \distinct A\<^sub>Q\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* xvec\ \A\<^sub>Q \* P'\ \(p \ xvec) \* A\<^sub>Q\ \A\<^sub>P' \* A\<^sub>Q\ \A\<^sub>Q \* P\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* C\ \A\<^sub>Q \* X\ \A\<^sub>Q \* Y\ \A\<^sub>Q \* Z\ \A\<^sub>Q \* N\ \A\<^sub>Q \* K\ obtain \'' A\<^sub>Q' \\<^sub>Q' where QeqQ': "\\<^sub>Q \ \'' \ \\<^sub>Q'" and FrQ': "extractFrame Q' = \A\<^sub>Q', \\<^sub>Q'\" and "distinct A\<^sub>Q'" and "A\<^sub>Q' \* xvec" and "A\<^sub>Q' \* Q'" and "A\<^sub>Q' \* xvec" and "A\<^sub>Q' \* P'" and "A\<^sub>Q' \* (p \ xvec)" and "A\<^sub>Q' \* A\<^sub>P'" and "A\<^sub>Q' \* C" and "A\<^sub>Q' \* X" and "A\<^sub>Q' \* Y" and "A\<^sub>Q' \* Z" and "A\<^sub>Q' \* P" by(rule_tac C="(P, C, P', X, Y, Z, A\<^sub>P', xvec, (p \ xvec), \\<^sub>P)" and C'="(P, C, P', X, Y, Z, A\<^sub>P', xvec, (p \ xvec), \\<^sub>P)" in expandNonTauFrame) (assumption | simp)+ from QTrans \A\<^sub>P' \* Q\ \A\<^sub>P' \* N\ \(p \ xvec) \* Q\ \(p \ xvec) \* N\ have "A\<^sub>P' \* Q'" and "(p \ xvec) \* Q'" by(force dest: inputFreshChainDerivative)+ with FrQ' \A\<^sub>Q' \* A\<^sub>P'\ \A\<^sub>Q' \* (p \ xvec)\ have "A\<^sub>P' \* \\<^sub>Q'" and "(p \ xvec) \* \\<^sub>Q'" by(force dest: extractFrameFreshChain)+ from FrP' \A\<^sub>Q' \* A\<^sub>P'\ \A\<^sub>Q' \* P'\ \(p \ xvec) \* A\<^sub>P'\ \(p \ xvec) \* P'\ have "A\<^sub>Q' \* \\<^sub>P'" and "(p \ xvec) \* \\<^sub>P'" by(force dest: extractFrameFreshChain)+ have "extractFrame(\\*xvec\(P' \ Q')) = \((p \ xvec)@A\<^sub>P'@A\<^sub>Q'), (p \ \\<^sub>P') \ (p \ \\<^sub>Q')\" proof - from FrP' FrQ' \A\<^sub>P' \* \\<^sub>Q'\ \A\<^sub>Q' \* A\<^sub>P'\ \A\<^sub>Q' \* \\<^sub>P'\ have "extractFrame(P' \ Q') = \(A\<^sub>P'@A\<^sub>Q'), \\<^sub>P' \ \\<^sub>Q'\" by simp hence "extractFrame(\\*xvec\(P' \ Q')) = \(xvec@A\<^sub>P'@A\<^sub>Q'), \\<^sub>P' \ \\<^sub>Q'\" by(induct xvec) auto moreover from \(p \ xvec) \* \\<^sub>P'\ \(p \ xvec) \* \\<^sub>Q'\ S have "\\*xvec\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert (\\<^sub>P' \ \\<^sub>Q'))) = \\*(p \ xvec)\(p \ \\*(A\<^sub>P'@A\<^sub>Q')\(FAssert(\\<^sub>P' \ \\<^sub>Q')))" by(rule_tac frameChainAlpha) (auto simp add: fresh_star_def frameResChainFresh) hence "\\*xvec\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert (\\<^sub>P' \ \\<^sub>Q'))) = \\*(p \ xvec)\(\\*(A\<^sub>P'@A\<^sub>Q')\(FAssert((p \ \\<^sub>P') \ (p \ \\<^sub>Q'))))" using \A\<^sub>P' \* xvec\ \(p \ xvec) \* A\<^sub>P'\ \A\<^sub>Q' \* xvec\ \A\<^sub>Q' \* (p \ xvec)\ S by(fastforce simp add: eqvts) ultimately show ?thesis by(simp add: frameChainAppend) qed moreover have "(\\<^sub>P \ \\<^sub>Q) \ ((p \ \') \ (p \ \'')) \ (p \ \\<^sub>P') \ (p \ \\<^sub>Q')" proof - have "((p \ \\<^sub>P) \ \\<^sub>Q) \ (\' \ \'') \ ((p \ \\<^sub>P) \ \') \ (\\<^sub>Q \ \'')" by(metis Associativity Commutativity Composition AssertionStatEqTrans) moreover from PeqP' QeqQ' have "((p \ \\<^sub>P) \ \') \ (\\<^sub>Q \ \'') \ \\<^sub>P' \ \\<^sub>Q'" by(metis Associativity Commutativity Composition AssertionStatEqTrans) ultimately have "((p \ \\<^sub>P) \ \\<^sub>Q) \ (\' \ \'') \ \\<^sub>P' \ \\<^sub>Q'" by(metis AssertionStatEqTrans) hence "(p \ ((p \ \\<^sub>P) \ \\<^sub>Q) \ (\' \ \'')) \ (p \ (\\<^sub>P' \ \\<^sub>Q'))" by(rule AssertionStatEqClosed) with \xvec \* \\<^sub>Q\ \(p \ xvec) \* \\<^sub>Q\ S \distinctPerm p\ show ?thesis by(simp add: eqvts) qed moreover from \(p \ xvec) \* C\ \A\<^sub>P' \* C\ \A\<^sub>Q' \* C\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* C" by simp moreover from \(p \ xvec) \* X\ \A\<^sub>P' \* X\ \A\<^sub>Q' \* X\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* X" by simp moreover from \(p \ xvec) \* Y\ \A\<^sub>P' \* Y\ \A\<^sub>Q' \* Y\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* Y" by simp moreover from \(p \ xvec) \* Z\ \A\<^sub>P' \* Z\ \A\<^sub>Q' \* Z\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* Z" by simp moreover from \(p \ xvec) \* P'\ \(p \ xvec) \* Q'\ \A\<^sub>P' \* P'\ \A\<^sub>P' \* Q' \\A\<^sub>Q' \* P'\ \A\<^sub>Q' \* Q'\ have "((p \ xvec)@A\<^sub>P'@A\<^sub>Q') \* (\\*xvec\(P' \ Q'))" by(auto simp add: resChainFresh fresh_star_def) moreover from \(p \ xvec) \* A\<^sub>P'\ \ A\<^sub>Q' \* (p \ xvec)\ \A\<^sub>Q' \* A\<^sub>P'\ \distinct xvec\ \distinct A\<^sub>P'\ \distinct A\<^sub>Q'\ have "distinct((p \ xvec)@A\<^sub>P'@A\<^sub>Q')" by auto (simp add: name_list_supp fresh_star_def fresh_def)+ ultimately show ?case using cComm2 by metis next case(cScope \ P P' x A\<^sub>P \\<^sub>P C X Y Z) then obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "distinct A\<^sub>P'" and "\\<^sub>P \ \' \ \\<^sub>P'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* (x#X)" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" by(rule_tac cScope(4)[where ba="x#X"]) auto from \A\<^sub>P' \* (x#X)\ have "x \ A\<^sub>P'" and "A\<^sub>P' \* X" by simp+ moreover from FrP' have "extractFrame(\\x\P') = \(x#A\<^sub>P'), \\<^sub>P'\" by simp moreover note \\\<^sub>P \ \' \ \\<^sub>P'\ moreover from \x \ C\ \A\<^sub>P' \* C\ have "(x#A\<^sub>P') \* C" by simp moreover from \A\<^sub>P' \* P'\ have "(x#A\<^sub>P') \* (\\x\P')" by(simp add: abs_fresh fresh_star_def) moreover from \x \ X\ \A\<^sub>P' \* X\ have "(x#A\<^sub>P') \* X" by simp moreover from \x \ Y\ \A\<^sub>P' \* Y\ have "(x#A\<^sub>P') \* Y" by simp moreover from \x \ Z\ \A\<^sub>P' \* Z\ have "(x#A\<^sub>P') \* Z" by simp moreover from \x \ A\<^sub>P'\ \distinct A\<^sub>P'\ have "distinct(x#A\<^sub>P')" by simp ultimately show ?case by(rule_tac cScope) next case(cBang \ P P' A\<^sub>P \\<^sub>P C B Y Z) then obtain \' A\<^sub>P' \\<^sub>P' where FrP': "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "(\\<^sub>P \ \) \ \' \ \\<^sub>P'" and "A\<^sub>P' \* C" and "A\<^sub>P' \* P'" and "distinct A\<^sub>P'" and "A\<^sub>P' \* B" and "A\<^sub>P' \* Y" and "A\<^sub>P' \* Z" by(rule_tac cBang) (assumption | simp)+ with \\\<^sub>P \ \\ \(\\<^sub>P \ \) \ \' \ \\<^sub>P'\ have "\ \ \' \ \\<^sub>P'" by(metis Identity AssertionStatEqTrans composition' Commutativity Associativity AssertionStatEqSym) thus ?case using FrP' \A\<^sub>P' \* P'\ \A\<^sub>P' \* C\ \A\<^sub>P' \* B\ \A\<^sub>P' \* Y\ \A\<^sub>P' \* Z\ \distinct A\<^sub>P'\ by(rule_tac cBang) qed with A have ?thesis by simp } moreover have "A\<^sub>P \* ([]::name list)" and "A\<^sub>P \* ([]::'b list)" and "A\<^sub>P \* ([]::('a, 'b, 'c) psi list)" by simp+ ultimately show ?thesis by blast qed lemma expandFrame: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and C :: "'d::fs_name" and C' :: "'e::fs_name" assumes Trans: "\ \ P \\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "bn \ \* subject \" and "distinct(bn \)" and "A\<^sub>P \* \" and "A\<^sub>P \* P" and "A\<^sub>P \* C" and "A\<^sub>P \* C'" and "bn \ \* P" and "bn \ \* C'" obtains p \' A\<^sub>P' \\<^sub>P' where "set p \ set(bn \) \ set(bn(p \ \))" and "(p \ \\<^sub>P) \ \' \ \\<^sub>P'" and "distinctPerm p" and "extractFrame P' = \A\<^sub>P', \\<^sub>P'\" and "A\<^sub>P' \* P'" and "A\<^sub>P' \* \" and "A\<^sub>P' \* (p \ \)" and "A\<^sub>P' \* C" and "(bn(p \ \)) \* C'" and "(bn(p \ \)) \* \" and "(bn(p \ \)) \* P'" and "distinct A\<^sub>P'" using assms apply(cases "\=\") by(auto intro: expandTauFrame[where C=C] expandNonTauFrame[where C=C and C'=C']) abbreviation frameImpJudge ("_ \\<^sub>F _" [80, 80] 80) where "F \\<^sub>F G \ FrameStatImp F G" lemma FrameStatEqImpCompose: fixes F :: "'b frame" and G :: "'b frame" and H :: "'b frame" and I :: "'b frame" assumes "F \\<^sub>F G" and "G \\<^sub>F H" and "H \\<^sub>F I" shows "F \\<^sub>F I" using assms by(auto simp add: FrameStatEq_def) (blast intro: FrameStatImpTrans) lemma transferNonTauFrame: fixes \\<^sub>F :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and A\<^sub>F :: "name list" and A\<^sub>G :: "name list" and \\<^sub>G :: 'b assumes "\\<^sub>F \ P \\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "distinct(bn \)" and "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" and "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* subject \" and "A\<^sub>G \* subject \" and "A\<^sub>P \* A\<^sub>F" and "A\<^sub>P \* A\<^sub>G" and "A\<^sub>P \* \\<^sub>F" and "A\<^sub>P \* \\<^sub>G" and "\ \ \" shows "\\<^sub>G \ P \\ \ P'" using assms proof(nominal_induct \\<^sub>F P Rs=="\ \ P'" A\<^sub>P \\<^sub>P avoiding: \ P' \\<^sub>G A\<^sub>F A\<^sub>G rule: semanticsFrameInduct) case(cAlpha \\<^sub>F P A\<^sub>P \\<^sub>P p \ P' \\<^sub>G A\<^sub>F A\<^sub>G) from \\A\<^sub>F, \\<^sub>F \ (p \ \\<^sub>P)\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ (p \ \\<^sub>P)\\ have "(p \ (\A\<^sub>F, \\<^sub>F \ (p \ \\<^sub>P)\)) \\<^sub>F (p \ (\A\<^sub>G, \\<^sub>G \ (p \ \\<^sub>P)\))" by(rule FrameStatImpClosed) with \A\<^sub>P \* A\<^sub>F\ \(p \ A\<^sub>P) \* A\<^sub>F\ \A\<^sub>P \* \\<^sub>F\ \(p \ A\<^sub>P) \* \\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \(p \ A\<^sub>P) \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>G\ \(p \ A\<^sub>P) \* \\<^sub>G\ \distinctPerm p\ \set p \ set A\<^sub>P \ set (p \ A\<^sub>P)\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" by(simp add: eqvts) with cAlpha show ?case by force next case(cInput \\<^sub>F M K xvec N Tvec P \ P' \\<^sub>G A\<^sub>F A\<^sub>G) from cInput have "A\<^sub>F \* K" and "A\<^sub>G \* K" by(auto simp add: residualInject) from \A\<^sub>F \* (M\\*xvec N\.P)\ \A\<^sub>G \* (M\\*xvec N\.P)\ have "A\<^sub>F \* M" and "A\<^sub>G \* M" by simp+ from \\\<^sub>F \ M \ K\ have "\\<^sub>F \ \ \ M \ K" by(blast intro: statEqEnt Identity AssertionStatEqSym) with \A\<^sub>F \* M\ \A\<^sub>F \* K\ have "(\A\<^sub>F, \\<^sub>F \ \\) \\<^sub>F M \ K" by(force intro: frameImpI) with \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ have "(\A\<^sub>G, \\<^sub>G \ \\) \\<^sub>F M \ K" by(simp add: FrameStatEq_def FrameStatImp_def) with \A\<^sub>G \* M\ \A\<^sub>G \* K\ have "\\<^sub>G \ \ \ M \ K" by(force dest: frameImpE) hence "\\<^sub>G \ M \ K" by(blast intro: statEqEnt Identity) thus ?case using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ using cInput Input by(force simp add: residualInject) next case(cOutput \\<^sub>F M K N P \ P' \\<^sub>G A\<^sub>F A\<^sub>G) from cOutput have "A\<^sub>F \* K" and "A\<^sub>G \* K" by(auto simp add: residualInject) from \A\<^sub>F \* (M\N\.P)\ \A\<^sub>G \* (M\N\.P)\ have "A\<^sub>F \* M" and "A\<^sub>G \* M" by simp+ from \\\<^sub>F \ M \ K\ have "\\<^sub>F \ \ \ M \ K" by(blast intro: statEqEnt Identity AssertionStatEqSym) with \A\<^sub>F \* M\ \A\<^sub>F \* K\ have "(\A\<^sub>F, \\<^sub>F \ \\) \\<^sub>F M \ K" by(force intro: frameImpI) with \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ have "(\A\<^sub>G, \\<^sub>G \ \\) \\<^sub>F M \ K" by(simp add: FrameStatImp_def) with \A\<^sub>G \* M\ \A\<^sub>G \* K\ have "\\<^sub>G \ \ \ M \ K" by(force dest: frameImpE) hence "\\<^sub>G \ M \ K" by(blast intro: statEqEnt Identity) thus ?case using cOutput Output by(force simp add: residualInject) next case(cCase \\<^sub>F P \ Cs A\<^sub>P \\<^sub>P \ P' \\<^sub>G A\<^sub>F A\<^sub>G) from \\\<^sub>P \ \\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\" by(metis frameIntCompositionSym Identity AssertionStatEqTrans) moreover note \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ moreover from \\\<^sub>P \ \\ have "\A\<^sub>G, \\<^sub>G \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym) ultimately have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" by(rule FrameStatEqImpCompose) with cCase have "\\<^sub>G \ P \ \ \ P'" by(fastforce dest: memFreshChain) moreover note \(\, P) mem Cs\ moreover from \A\<^sub>F \* (Cases Cs)\\A\<^sub>G \* (Cases Cs)\ \(\, P) mem Cs\ have "A\<^sub>F \* \" and "A\<^sub>G \* \" by(auto dest: memFreshChain) from \\\<^sub>F \ \\ have "\\<^sub>F \ \ \ \" by(blast intro: statEqEnt Identity AssertionStatEqSym) with \A\<^sub>F \* \\ have "(\A\<^sub>F, \\<^sub>F \ \\) \\<^sub>F \" by(force intro: frameImpI) with \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ have "(\A\<^sub>G, \\<^sub>G \ \\) \\<^sub>F \" by(simp add: FrameStatImp_def) with \A\<^sub>G \* \\ have "\\<^sub>G \ \ \ \" by(force dest: frameImpE) hence "\\<^sub>G \ \" by(blast intro: statEqEnt Identity) ultimately show ?case using \guarded P\ cCase Case by(force intro: residualInject) next case(cPar1 \\<^sub>F \\<^sub>Q P \ P' A\<^sub>Q Q A\<^sub>P \\<^sub>P \' PQ' \\<^sub>G A\<^sub>F A\<^sub>G) from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" by simp+ have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) moreover note \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\ " by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) ultimately have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" by(rule FrameStatEqImpCompose) moreover from \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ FrQ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>Q" and "A\<^sub>G \* \\<^sub>Q" by(force dest: extractFrameFreshChain)+ moreover note \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>F \* subject \'\ \A\<^sub>G \* subject \'\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>G\ moreover from \\ \ P' \ Q = \' \ PQ'\ \bn \ \* \'\ obtain p P'' where "\ \ P' = \' \ P''" and "set p \ set(bn \') \ set(bn \)" and "PQ' = P'' \ (p \ Q)" apply(drule_tac sym) by(rule_tac actionPar1Dest) (assumption | simp | blast dest: sym)+ ultimately have "\\<^sub>G \ \\<^sub>Q \ P \\ \ P'" using \\' \ \\ by(force intro: cPar1) thus ?case using FrQ \(bn \) \* Q\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* P\ \A\<^sub>Q \* \\ using cPar1 by(metis Par1) next case(cPar2 \\<^sub>F \\<^sub>P Q \ Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q \' PQ' \\<^sub>G A\<^sub>F A\<^sub>G) from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" by simp+ have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" by(metis Associativity frameResChainPres frameNilStatEq) moreover note \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\ " by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq) ultimately have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" by(rule FrameStatEqImpCompose) moreover from \A\<^sub>F \* P\ \A\<^sub>G \* P\ FrP \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>P" and "A\<^sub>G \* \\<^sub>P" by(force dest: extractFrameFreshChain)+ moreover note \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>F \* subject \'\ \A\<^sub>G \* subject \'\ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>G\ moreover from \\ \ P \ Q' = \' \ PQ'\ \bn \ \* \'\ obtain p Q'' where "\ \ Q' = \' \ Q''" and "set p \ set(bn \') \ set(bn \)" and "PQ' = (p \ P) \ Q''" apply(drule_tac sym) by(rule_tac actionPar2Dest) (assumption | simp | blast dest: sym)+ ultimately have "\\<^sub>G \ \\<^sub>P \ Q \\ \ Q'" using \\' \ \\ by(force intro: cPar2) thus ?case using FrP \(bn \) \* P\ \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* Q\ \A\<^sub>P \* \\ using cPar2 by(metis Par2) next case cComm1 thus ?case by(simp add: residualInject) next case cComm2 thus ?case by(simp add: residualInject) next case(cOpen \\<^sub>F P M xvec yvec N P' x A\<^sub>P \\<^sub>P \ P'' \\<^sub>G A\<^sub>F A\<^sub>G) from \M\\*(xvec @ x # yvec)\\N\ \ P' = \ \ P''\ \x \ xvec\ \x \ yvec\ \x \ \\ \x \ P''\ \distinct(bn \)\ obtain xvec' x' yvec' N' where "M\\*(xvec@yvec)\\N\ \ P' = M\\*(xvec'@yvec')\\([(x, x')] \ N')\ \ ([(x, x')] \ P'')" and "\ = M\\*(xvec'@x'#yvec')\\N'\" apply(cases rule: actionCases[where \=\]) apply(auto simp add: residualInject) apply(rule boundOutputOpenDest) by assumption auto then have "\\<^sub>G \ P \M\\*(xvec@yvec)\\N\ \ P'" using cOpen by(rule_tac cOpen) (assumption | simp)+ with \x \ supp N\ \x \ \\<^sub>G\ \x \ M\ \x \ xvec\ \x \ yvec\ have "\\<^sub>G \ \\x\P \M\\*(xvec@x#yvec)\\N\ \ P'" by(rule_tac Open) thus ?case using \\ = M\\*(xvec'@x'#yvec')\\N'\\ \M\\*(xvec @ x # yvec)\\N\ \ P' = \ \ P''\ by simp next case(cScope \\<^sub>F P \ P' x A\<^sub>P \\<^sub>P \' xP \\<^sub>G A\<^sub>F A\<^sub>G) from \\ \ \\x\P' = \' \ xP\ \x \ \\ \x \ \'\ obtain P'' where "xP = \\x\P''" and "\ \ P' = \' \ P''" by(drule_tac sym) (force intro: actionScopeDest) then have "\\<^sub>G \ P \\ \ P'" using cScope by auto with \x \ \\<^sub>G\ \x \ \'\ \\ \ P' = \' \ P''\ \xP = \\x\P''\ show ?case by(metis Scope) next case(cBang \\<^sub>F P A\<^sub>P \\<^sub>P \ P' \\<^sub>G A\<^sub>F A\<^sub>G) from \\\<^sub>P \ \\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\" by(metis frameIntCompositionSym Identity AssertionStatEqTrans) moreover note \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ moreover from \\\<^sub>P \ \\ have "\A\<^sub>G, \\<^sub>G \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\" by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym) ultimately have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\" by(rule FrameStatEqImpCompose) with cBang have "\\<^sub>G \ P \ !P \ \ \ P'" by force thus ?case using \guarded P\ using cBang by(metis Bang) qed lemma transferTauFrame: fixes \\<^sub>F :: 'b and P :: "('a, 'b, 'c) psi" and P' :: "('a, 'b, 'c) psi" and A\<^sub>F :: "name list" and A\<^sub>G :: "name list" and \\<^sub>G :: 'b assumes "\\<^sub>F \ P \\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" and "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>P \* A\<^sub>F" and "A\<^sub>P \* A\<^sub>G" and "A\<^sub>P \* \\<^sub>F" and "A\<^sub>P \* \\<^sub>G" shows "\\<^sub>G \ P \\ \ P'" using assms proof(nominal_induct avoiding: \\<^sub>G A\<^sub>F A\<^sub>G rule: tauFrameInduct) case(cAlpha \\<^sub>F P P' A\<^sub>P \\<^sub>P p \\<^sub>G A\<^sub>F A\<^sub>G) from \\A\<^sub>F, \\<^sub>F \ (p \ \\<^sub>P)\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ (p \ \\<^sub>P)\\ have "(p \ (\A\<^sub>F, \\<^sub>F \ (p \ \\<^sub>P)\)) \\<^sub>F (p \ (\A\<^sub>G, \\<^sub>G \ (p \ \\<^sub>P)\))" by(rule FrameStatImpClosed) with \A\<^sub>P \* A\<^sub>F\ \(p \ A\<^sub>P) \* A\<^sub>F\ \A\<^sub>P \* \\<^sub>F\ \(p \ A\<^sub>P) \* \\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \(p \ A\<^sub>P) \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>G\ \(p \ A\<^sub>P) \* \\<^sub>G\ \distinctPerm p\ \set p \ set A\<^sub>P \ set (p \ A\<^sub>P)\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" by(simp add: eqvts) with cAlpha show ?case by blast next case(cCase \\<^sub>F P P' \ Cs A\<^sub>P \\<^sub>P \\<^sub>G A\<^sub>F A\<^sub>G) from \\\<^sub>P \ \\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\" by(metis frameIntCompositionSym Identity AssertionStatEqTrans) moreover note \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ moreover from \\\<^sub>P \ \\ have "\A\<^sub>G, \\<^sub>G \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym) ultimately have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" by(rule FrameStatEqImpCompose) with cCase have "\\<^sub>G \ P \\ \ P'" by(fastforce dest: memFreshChain) moreover note \(\, P) mem Cs\ moreover from \A\<^sub>F \* (Cases Cs)\\A\<^sub>G \* (Cases Cs)\ \(\, P) mem Cs\ have "A\<^sub>F \* \" and "A\<^sub>G \* \" by(auto dest: memFreshChain) from \\\<^sub>F \ \\ have "\\<^sub>F \ \ \ \" by(blast intro: statEqEnt Identity AssertionStatEqSym) with \A\<^sub>F \* \\ have "(\A\<^sub>F, \\<^sub>F \ \\) \\<^sub>F \" by(force intro: frameImpI) with \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ have "(\A\<^sub>G, \\<^sub>G \ \\) \\<^sub>F \" by(simp add: FrameStatImp_def) with \A\<^sub>G \* \\ have "\\<^sub>G \ \ \ \" by(force dest: frameImpE) hence "\\<^sub>G \ \" by(blast intro: statEqEnt Identity) ultimately show ?case using \guarded P\ by(rule Case) next case(cPar1 \\<^sub>F \\<^sub>Q P P' A\<^sub>Q Q A\<^sub>P \\<^sub>P \\<^sub>G A\<^sub>F A\<^sub>G) from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" by simp+ have IH: "\\ A\<^sub>F A\<^sub>G. \\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \ \ \\<^sub>P\; A\<^sub>F \* P; A\<^sub>G \* P; A\<^sub>P \* A\<^sub>F; A\<^sub>P \* A\<^sub>G; A\<^sub>P \* (\\<^sub>F \ \\<^sub>Q); A\<^sub>P \* \\ \ \ \ P \\ \ P'" by fact have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) moreover note \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\ " by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) ultimately have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" by(rule FrameStatEqImpCompose) moreover from \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ FrQ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>Q" and "A\<^sub>G \* \\<^sub>Q" by(force dest: extractFrameFreshChain)+ moreover note \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>G\ ultimately have "\\<^sub>G \ \\<^sub>Q \ P \\ \ P'" by(rule_tac IH) (assumption | auto)+ thus ?case using FrQ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* P\ by(rule_tac Par1) auto next case(cPar2 \\<^sub>F \\<^sub>P Q Q' A\<^sub>P P A\<^sub>Q \\<^sub>Q \\<^sub>G A\<^sub>F A\<^sub>G) from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" by simp+ have IH: "\\ A\<^sub>F A\<^sub>G. \\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \ \ \\<^sub>Q\; A\<^sub>F \* Q; A\<^sub>G \* Q; A\<^sub>Q \* A\<^sub>F; A\<^sub>Q \* A\<^sub>G; A\<^sub>Q \* (\\<^sub>F \ \\<^sub>P); A\<^sub>Q \* \\ \ \ \ Q \\ \ Q'" by fact have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" by(metis Associativity frameResChainPres frameNilStatEq) moreover note \\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\\ moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\ " by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq) ultimately have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" by(rule FrameStatEqImpCompose) moreover from \A\<^sub>F \* P\ \A\<^sub>G \* P\ FrP \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>P" and "A\<^sub>G \* \\<^sub>P" by(force dest: extractFrameFreshChain)+ moreover note \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>G\ ultimately have "\\<^sub>G \ \\<^sub>P \ Q \\ \ Q'" by(rule_tac IH) (assumption | auto)+ thus ?case using FrP \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* Q\ by(rule_tac Par2) auto next case(cComm1 \\<^sub>F \\<^sub>Q P M N P' A\<^sub>P \\<^sub>P Q K xvec Q' A\<^sub>Q \\<^sub>G A\<^sub>F A\<^sub>G) have FimpG: "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\" by fact from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" by simp+ have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact from \A\<^sub>F \* P\ \A\<^sub>G \* P\ FrP \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>P" and "A\<^sub>G \* \\<^sub>P" by(force dest: extractFrameFreshChain)+ from \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ FrQ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>Q" and "A\<^sub>G \* \\<^sub>Q" by(force dest: extractFrameFreshChain)+ from \\\<^sub>F \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ obtain K' where KeqK': "(\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q \ K \ K'" and "A\<^sub>P \* K'" and "A\<^sub>F \* K'" and "A\<^sub>G \* K'" using \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \xvec \* K\ \distinct xvec\ by(rule_tac B="A\<^sub>P@A\<^sub>F@A\<^sub>G" in obtainPrefix) force+ have "\\<^sub>G \ \\<^sub>Q \ P \K'\N\ \ P'" proof - from KeqK' have "\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ K \ K'" by(rule statEqEnt[OF Associativity]) with \\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K\ have "\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K'" by(rule chanEqTrans) hence "(\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P \ M \ K'" by(metis statEqEnt AssertionStatEqSym Associativity AssertionStatEqTrans compositionSym Commutativity) with \\\<^sub>F \ \\<^sub>Q \ P \M\N\ \ P'\ FrP \distinct A\<^sub>P\ have "\\<^sub>F \ \\<^sub>Q \ P \K'\N\ \ P'" using \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K'\ by(force intro: inputRenameSubject) moreover have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" proof - have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\ " by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) ultimately show ?thesis using FimpG by(rule_tac FrameStatEqImpCompose) qed ultimately show ?thesis using \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>F \* K'\ \A\<^sub>G \* K'\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* \\<^sub>Q\ FrP \distinct A\<^sub>P\ by(auto intro: transferNonTauFrame) qed moreover from FrP \\\<^sub>F \ \\<^sub>Q \ P \M\N\ \ P'\ \distinct A\<^sub>P\ obtain M' where MeqM': "(\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "A\<^sub>Q \* M'" and "A\<^sub>F \* M'" and "A\<^sub>G \* M'" using \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ by(rule_tac B="A\<^sub>Q@A\<^sub>F@A\<^sub>G" in obtainPrefix) force+ have "\\<^sub>G \ \\<^sub>P \ Q \M'\\*xvec\\N\ \ Q'" proof - from MeqM' have "\\<^sub>F \ (\\<^sub>Q \ \\<^sub>P) \ M \ M'" by(rule statEqEnt[OF Associativity]) with \\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K\ have "\\<^sub>F \ (\\<^sub>Q \ \\<^sub>P) \ K \ M'" by(blast intro: chanEqTrans chanEqSym compositionSym Commutativity statEqEnt) hence "(\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" by(blast intro: statEqEnt AssertionStatEqSym Associativity AssertionStatEqTrans compositionSym Commutativity) with \\\<^sub>F \ \\<^sub>P \ Q \K\\*xvec\\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ have "\\<^sub>F \ \\<^sub>P \ Q \M'\\*xvec\\N\ \ Q'" using \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>Q \* M'\ by(force intro: outputRenameSubject) moreover have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" proof - have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" by(metis Associativity frameResChainPres frameNilStatEq) moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\ " by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq) ultimately show ?thesis using FimpG by(rule_tac FrameStatEqImpCompose) qed ultimately show ?thesis using \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>F \* M'\ \A\<^sub>G \* M'\ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* \\<^sub>P\ FrQ \distinct A\<^sub>Q\ \distinct xvec\ by(auto intro: transferNonTauFrame) qed moreover have "\\<^sub>G \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" proof - from MeqM' have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ M' \ M" by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym) moreover from KeqK' have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ K \ K'" by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym) ultimately have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" using \\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ M \ K\ by(blast intro: chanEqSym chanEqTrans) thus ?thesis using \A\<^sub>F \* M'\ \A\<^sub>F \* K'\ \A\<^sub>G \* M'\ \A\<^sub>G \* K'\ FimpG apply(auto simp add: FrameStatImp_def) apply(erule_tac x="SChanEq' K' M'" in allE) by(force intro: frameImpI dest: frameImpE) qed ultimately show ?case using \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* K'\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M'\ \xvec \* P\ FrP FrQ by(rule_tac Comm1) (assumption | simp)+ next case(cComm2 \\<^sub>F \\<^sub>Q P M xvec N P' A\<^sub>P \\<^sub>P Q K Q' A\<^sub>Q \\<^sub>G A\<^sub>F A\<^sub>G) have FimpG: "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\" by fact from \A\<^sub>F \* (P \ Q)\ \A\<^sub>G \* (P \ Q)\ have "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* Q" and "A\<^sub>G \* Q" by simp+ have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact from \A\<^sub>F \* P\ \A\<^sub>G \* P\ FrP \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>P" and "A\<^sub>G \* \\<^sub>P" by(force dest: extractFrameFreshChain)+ from \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ FrQ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ have "A\<^sub>F \* \\<^sub>Q" and "A\<^sub>G \* \\<^sub>Q" by(force dest: extractFrameFreshChain)+ from \\\<^sub>F \ \\<^sub>P \ Q \K\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ obtain K' where KeqK': "(\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q \ K \ K'" and "A\<^sub>P \* K'" and "A\<^sub>F \* K'" and "A\<^sub>G \* K'" using \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ by(rule_tac B="A\<^sub>P@A\<^sub>F@A\<^sub>G" in obtainPrefix) force+ have "\\<^sub>G \ \\<^sub>Q \ P \K'\\*xvec\\N\ \ P'" proof - from KeqK' have "\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ K \ K'" by(rule statEqEnt[OF Associativity]) with \\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K\ have "\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K'" by(rule chanEqTrans) hence "(\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P \ M \ K'" by(metis statEqEnt AssertionStatEqSym Associativity AssertionStatEqTrans compositionSym Commutativity) with \\\<^sub>F \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ FrP \distinct A\<^sub>P\ have "\\<^sub>F \ \\<^sub>Q \ P \K'\\*xvec\\N\ \ P'" using \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* K'\ by(force intro: outputRenameSubject) moreover have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\" proof - have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>Q) \ \\<^sub>P\ " by(metis Associativity Composition AssertionStatEqSym AssertionStatEqTrans Commutativity frameResChainPres frameNilStatEq) ultimately show ?thesis using FimpG by(rule_tac FrameStatEqImpCompose) qed ultimately show ?thesis using \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>F \* K'\ \A\<^sub>G \* K'\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* \\<^sub>Q\ FrP \distinct A\<^sub>P\ \distinct xvec\ by(auto intro: transferNonTauFrame) qed moreover from FrP \\\<^sub>F \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ \distinct A\<^sub>P\ obtain M' where MeqM': "(\\<^sub>F \ \\<^sub>Q) \ \\<^sub>P \ M \ M'" and "A\<^sub>Q \* M'" and "A\<^sub>F \* M'" and "A\<^sub>G \* M'" using \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>F\ \A\<^sub>P \* A\<^sub>G\ \A\<^sub>F \* P\ \A\<^sub>G \* P\ \A\<^sub>P \* \\<^sub>F\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \xvec \* M\ \distinct xvec\ by(rule_tac B="A\<^sub>Q@A\<^sub>F@A\<^sub>G" in obtainPrefix) force+ have "\\<^sub>G \ \\<^sub>P \ Q \M'\N\ \ Q'" proof - from MeqM' have "\\<^sub>F \ (\\<^sub>Q \ \\<^sub>P) \ M \ M'" by(rule statEqEnt[OF Associativity]) with \\\<^sub>F \ (\\<^sub>P \ \\<^sub>Q) \ M \ K\ have "\\<^sub>F \ (\\<^sub>Q \ \\<^sub>P) \ K \ M'" by(blast intro: chanEqTrans chanEqSym compositionSym Commutativity statEqEnt) hence "(\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q \ K \ M'" by(blast intro: statEqEnt AssertionStatEqSym Associativity AssertionStatEqTrans compositionSym Commutativity) with \\\<^sub>F \ \\<^sub>P \ Q \K\N\ \ Q'\ FrQ \distinct A\<^sub>Q\ have "\\<^sub>F \ \\<^sub>P \ Q \M'\N\ \ Q'" using \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* K\ \A\<^sub>Q \* M'\ by(force intro: inputRenameSubject) moreover have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\" proof - have "\A\<^sub>F, (\\<^sub>F \ \\<^sub>P) \ \\<^sub>Q\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\<^sub>Q\" by(metis Associativity frameResChainPres frameNilStatEq) moreover have "\A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\<^sub>Q\ \\<^sub>F \A\<^sub>G, (\\<^sub>G \ \\<^sub>P) \ \\<^sub>Q\ " by(metis Associativity AssertionStatEqSym frameResChainPres frameNilStatEq) ultimately show ?thesis using FimpG by(rule_tac FrameStatEqImpCompose) qed ultimately show ?thesis using \A\<^sub>F \* Q\ \A\<^sub>G \* Q\ \A\<^sub>F \* M'\ \A\<^sub>G \* M'\ \A\<^sub>Q \* A\<^sub>F\ \A\<^sub>Q \* A\<^sub>G\ \A\<^sub>Q \* \\<^sub>F\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* \\<^sub>P\ FrQ \distinct A\<^sub>Q\ \distinct xvec\ by(auto intro: transferNonTauFrame) qed moreover have "\\<^sub>G \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" proof - from MeqM' have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ M' \ M" by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym) moreover from KeqK' have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ K \ K'" by(blast intro: chanEqSym Associativity statEqEnt Commutativity compositionSym) ultimately have "\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ K' \ M'" using \\\<^sub>F \ \\<^sub>P \ \\<^sub>Q \ M \ K\ by(blast intro: chanEqSym chanEqTrans) thus ?thesis using \A\<^sub>F \* M'\ \A\<^sub>F \* K'\ \A\<^sub>G \* M'\ \A\<^sub>G \* K'\ FimpG apply(auto simp add: FrameStatImp_def) apply(erule_tac x="SChanEq' K' M'" in allE) by(force intro: frameImpI dest: frameImpE) qed ultimately show ?case using \A\<^sub>P \* \\<^sub>G\ \A\<^sub>P \* P\ \A\<^sub>P \* Q\ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* K'\ \A\<^sub>Q \* \\<^sub>G\ \A\<^sub>Q \* P\ \A\<^sub>Q \* Q\ \A\<^sub>Q \* M'\ \xvec \* Q\ FrP FrQ by(rule_tac Comm2) (assumption | simp)+ next case(cScope \\<^sub>F P P' x A\<^sub>P \\<^sub>P \\<^sub>G A\<^sub>F A\<^sub>G) then have "\\<^sub>G \ P \\ \ P'" by auto with \x \ \\<^sub>G\ show ?case by(rule_tac Scope) auto next case(cBang \\<^sub>F P P' A\<^sub>P \\<^sub>P \\<^sub>G A\<^sub>F A\<^sub>G) from \\\<^sub>P \ \\ have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\ \\<^sub>F \A\<^sub>F, \\<^sub>F \ \\" by(metis frameIntCompositionSym Identity AssertionStatEqTrans) moreover note \\A\<^sub>F, \\<^sub>F \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\\ moreover from \\\<^sub>P \ \\ have "\A\<^sub>G, \\<^sub>G \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\" by(metis frameIntCompositionSym Identity AssertionStatEqTrans AssertionStatEqSym) ultimately have "\A\<^sub>F, \\<^sub>F \ \\<^sub>P \ \\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P \ \\" by(rule FrameStatEqImpCompose) with cBang have "\\<^sub>G \ P \ !P \\ \ P'" by force thus ?case using \guarded P\ by(rule Bang) qed lemma transferFrame: fixes \\<^sub>F :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and A\<^sub>F :: "name list" and A\<^sub>G :: "name list" and \\<^sub>G :: 'b assumes "\\<^sub>F \ P \\ \ P'" and "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" and "\A\<^sub>F, \\<^sub>F \ \\<^sub>P\ \\<^sub>F \A\<^sub>G, \\<^sub>G \ \\<^sub>P\" and "A\<^sub>F \* P" and "A\<^sub>G \* P" and "A\<^sub>F \* subject \" and "A\<^sub>G \* subject \" and "A\<^sub>P \* A\<^sub>F" and "A\<^sub>P \* A\<^sub>G" and "A\<^sub>P \* \\<^sub>F" and "A\<^sub>P \* \\<^sub>G" shows "\\<^sub>G \ P \\ \ P'" using assms proof - from \\\<^sub>F \ P \\ \ P'\ have "distinct(bn \)" by(auto dest: boundOutputDistinct) thus ?thesis using assms by(cases "\ = \") (auto intro: transferNonTauFrame transferTauFrame) qed lemma parCasesInputFrame[consumes 7, case_names cPar1 cPar2]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and N :: 'a and T :: "('a, 'b, 'c) psi" and C :: "'d::fs_name" assumes Trans: "\ \ P \ Q \M\N\ \ T" and "extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" and "distinct A\<^sub>P\<^sub>Q" and "A\<^sub>P\<^sub>Q \* \" and "A\<^sub>P\<^sub>Q \* P" and "A\<^sub>P\<^sub>Q \* Q" and "A\<^sub>P\<^sub>Q \* M" and rPar1: "\P' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \M\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P' \ Q)" and rPar2: "\Q' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>P \ Q \M\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P \ Q')" shows "Prop T" using Trans proof(induct rule: parInputCases[of _ _ _ _ _ _ "(A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)"]) case(cPar1 P' A\<^sub>Q \\<^sub>Q) from \A\<^sub>Q \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P\<^sub>Q" by simp+ obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" "A\<^sub>P \* (P, Q, \, M, A\<^sub>Q, A\<^sub>P\<^sub>Q, \\<^sub>Q)" by(rule freshFrame) hence "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* \" and "A\<^sub>P \* M" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" by simp+ have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact from \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>Q\ FrP have "A\<^sub>Q \* \\<^sub>P" by(force dest: extractFrameFreshChain) from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" by(auto simp add: fresh_star_def fresh_def name_list_supp) ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ by(rule_tac frameChainEq') (assumption | simp add: eqvts)+ from \\ \ \\<^sub>Q \ P \M\N\ \ P'\ S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq have "(p \ (\ \ \\<^sub>Q)) \ P \M\N\ \ P'" by(rule_tac inputPermFrame) auto with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>Q) \ P \M\N\ \ P'" by(simp add: eqvts) moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" by(simp add: eqvts) moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" by(simp add: eqvts) moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" by simp+ moreover from \A\<^sub>P \* A\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq by(rule_tac rPar1) (assumption | simp)+ next case(cPar2 Q' A\<^sub>P \\<^sub>P) from \A\<^sub>P \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>P\<^sub>Q" by simp+ obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>Q" "A\<^sub>Q \* (P, Q, \, M, A\<^sub>P, A\<^sub>P\<^sub>Q, \\<^sub>P)" by(rule freshFrame) hence "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* M" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P" by simp+ have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact from \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ FrQ have "A\<^sub>P \* \\<^sub>Q" by(force dest: extractFrameFreshChain) from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\ have "distinct(A\<^sub>P@A\<^sub>Q)" by(auto simp add: fresh_star_def fresh_def name_list_supp) ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ by(rule_tac frameChainEq') (assumption | simp add: eqvts)+ from \\ \ \\<^sub>P \ Q \M\N\ \ Q'\ S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq have "(p \ (\ \ \\<^sub>P)) \ Q \M\N\ \ Q'" by(rule_tac inputPermFrame) auto with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>P) \ Q \M\N\ \ Q'" by(simp add: eqvts) moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" by(simp add: eqvts) moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" by(simp add: eqvts) moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" by simp+ moreover from \A\<^sub>Q \* A\<^sub>P\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq by(rule_tac rPar2) (assumption | simp)+ qed lemma parCasesOutputFrame[consumes 11, case_names cPar1 cPar2]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Q :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and N :: 'a and T :: "('a, 'b, 'c) psi" and C :: "'d::fs_name" assumes Trans: "\ \ P \ Q \M\\*xvec\\N\ \ T" and "xvec \* \" and "xvec \* P" and "xvec \* Q" and "xvec \* M" and "extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" and "distinct A\<^sub>P\<^sub>Q" and "A\<^sub>P\<^sub>Q \* \" and "A\<^sub>P\<^sub>Q \* P" and "A\<^sub>P\<^sub>Q \* Q" and "A\<^sub>P\<^sub>Q \* M" and rPar1: "\P' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P' \ Q)" and rPar2: "\Q' A\<^sub>P \\<^sub>P A\<^sub>Q \\<^sub>Q. \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'; extractFrame P = \A\<^sub>P, \\<^sub>P\; extractFrame Q = \A\<^sub>Q, \\<^sub>Q\; distinct A\<^sub>P; distinct A\<^sub>Q; A\<^sub>P \* \; A\<^sub>P \* P; A\<^sub>P \* Q; A\<^sub>P \* M; A\<^sub>Q \* \; A\<^sub>Q \* P; A\<^sub>Q \* Q; A\<^sub>Q \* M; A\<^sub>P \* \\<^sub>Q; A\<^sub>Q \* \\<^sub>P; A\<^sub>P \* A\<^sub>Q; A\<^sub>P\<^sub>Q = A\<^sub>P@A\<^sub>Q; \\<^sub>P\<^sub>Q = \\<^sub>P \ \\<^sub>Q\ \ Prop (P \ Q')" shows "Prop T" using Trans \xvec \* \\ \xvec \* P\ \xvec \* Q\ \xvec \* M\ proof(induct rule: parOutputCases[of _ _ _ _ _ _ _ "(A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)"]) case(cPar1 P' A\<^sub>Q \\<^sub>Q) from \A\<^sub>Q \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P\<^sub>Q" by simp+ obtain A\<^sub>P \\<^sub>P where FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and "distinct A\<^sub>P" "A\<^sub>P \* (P, Q, \, M, A\<^sub>Q, A\<^sub>P\<^sub>Q, \\<^sub>Q)" by(rule freshFrame) hence "A\<^sub>P \* P" and "A\<^sub>P \* Q" and "A\<^sub>P \* \" and "A\<^sub>P \* M" and "A\<^sub>P \* A\<^sub>Q" and "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>Q" by simp+ have FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" by fact from \A\<^sub>Q \* P\ \A\<^sub>P \* A\<^sub>Q\ FrP have "A\<^sub>Q \* \\<^sub>P" by(force dest: extractFrameFreshChain) from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>P \* A\<^sub>Q\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>P \* A\<^sub>Q\ have "distinct(A\<^sub>P@A\<^sub>Q)" by(auto simp add: fresh_star_def fresh_def name_list_supp) ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ by(rule_tac frameChainEq') (assumption | simp add: eqvts)+ from \\ \ \\<^sub>Q \ P \M\\*xvec\\N\ \ P'\ S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq have "(p \ (\ \ \\<^sub>Q)) \ P \M\\*xvec\\N\ \ P'" by(rule_tac outputPermFrame) (assumption | simp)+ with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>Q) \ P \M\\*xvec\\N\ \ P'" by(simp add: eqvts) moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" by(simp add: eqvts) moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" by(simp add: eqvts) moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" by simp+ moreover from \A\<^sub>P \* A\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq by(rule_tac rPar1) (assumption | simp)+ next case(cPar2 Q' A\<^sub>P \\<^sub>P) from \A\<^sub>P \* (A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q)\ have "A\<^sub>P \* A\<^sub>P\<^sub>Q" and "A\<^sub>P \* \\<^sub>P\<^sub>Q" by simp+ obtain A\<^sub>Q \\<^sub>Q where FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>Q" "A\<^sub>Q \* (P, Q, \, M, A\<^sub>P, A\<^sub>P\<^sub>Q, \\<^sub>P)" by(rule freshFrame) hence "A\<^sub>Q \* P" and "A\<^sub>Q \* Q" and "A\<^sub>Q \* \" and "A\<^sub>Q \* M" and "A\<^sub>Q \* A\<^sub>P" and "A\<^sub>Q \* A\<^sub>P\<^sub>Q" and "A\<^sub>Q \* \\<^sub>P" by simp+ have FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" by fact from \A\<^sub>P \* Q\ \A\<^sub>Q \* A\<^sub>P\ FrQ have "A\<^sub>P \* \\<^sub>Q" by(force dest: extractFrameFreshChain) from \extractFrame(P \ Q) = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\\ FrP FrQ \A\<^sub>Q \* A\<^sub>P\ \A\<^sub>P \* \\<^sub>Q\ \A\<^sub>Q \* \\<^sub>P\ have "\(A\<^sub>P@A\<^sub>Q), \\<^sub>P \ \\<^sub>Q\ = \A\<^sub>P\<^sub>Q, \\<^sub>P\<^sub>Q\" by simp moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\ have "distinct(A\<^sub>P@A\<^sub>Q)" by(auto simp add: fresh_star_def fresh_def name_list_supp) ultimately obtain p where S: "set p \ set(A\<^sub>P@A\<^sub>Q) \ set((p \ A\<^sub>P)@(p \ A\<^sub>Q))" and "distinctPerm p" and \eq: "\\<^sub>P\<^sub>Q = (p \ \\<^sub>P) \ (p \ \\<^sub>Q)" and Aeq: "A\<^sub>P\<^sub>Q = (p \ A\<^sub>P)@(p \ A\<^sub>Q)" using \A\<^sub>P \* A\<^sub>P\<^sub>Q\ \A\<^sub>Q \* A\<^sub>P\<^sub>Q\ \distinct A\<^sub>P\<^sub>Q\ by(rule_tac frameChainEq') (assumption | simp add: eqvts)+ from \\ \ \\<^sub>P \ Q \M\\*xvec\\N\ \ Q'\ S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ Aeq have "(p \ (\ \ \\<^sub>P)) \ Q \M\\*xvec\\N\ \ Q'" by(rule_tac outputPermFrame) (assumption | simp)+ with S \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P \* \\ \A\<^sub>Q \* \\ Aeq have "\ \ (p \ \\<^sub>P) \ Q \M\\*xvec\\N\ \ Q'" by(simp add: eqvts) moreover from FrP have "(p \ extractFrame P) = p \ \A\<^sub>P, \\<^sub>P\" by simp with S \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P \* P\ \A\<^sub>Q \* P\ Aeq have "extractFrame P = \(p \ A\<^sub>P), p \ \\<^sub>P\" by(simp add: eqvts) moreover from FrQ have "(p \ extractFrame Q) = p \ \A\<^sub>Q, \\<^sub>Q\" by simp with S \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P \* Q\ \A\<^sub>Q \* Q\ Aeq have "extractFrame Q = \(p \ A\<^sub>Q), p \ \\<^sub>Q\" by(simp add: eqvts) moreover from \distinct A\<^sub>P\ \distinct A\<^sub>Q\ have "distinct(p \ A\<^sub>P)" and "distinct(p \ A\<^sub>Q)" by simp+ moreover from \A\<^sub>Q \* A\<^sub>P\ have "(p \ A\<^sub>P) \* (p \ A\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) moreover from \A\<^sub>P \* \\<^sub>Q\ have "(p \ A\<^sub>P) \* (p \ \\<^sub>Q)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) moreover from \A\<^sub>Q \* \\<^sub>P\ have "(p \ A\<^sub>Q) \* (p \ \\<^sub>P)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) ultimately show ?case using \A\<^sub>P\<^sub>Q \* \\ \A\<^sub>P\<^sub>Q \* P\ \A\<^sub>P\<^sub>Q \* Q\ \A\<^sub>P\<^sub>Q \* M\ Aeq \eq by(rule_tac rPar2) (assumption | simp)+ qed inductive bangPred :: "('a, 'b, 'c) psi \ ('a, 'b, 'c) psi \ bool" where aux1: "bangPred P (!P)" | aux2: "bangPred P (P \ !P)" lemma bangInduct[consumes 1, case_names cPar1 cPar2 cComm1 cComm2 cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and Rs :: "('a, 'b, 'c) residual" and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) residual \ bool" and C :: 'd assumes "\ \ !P \ Rs" and rPar1: "\\ P' C. \\ \ P \\ \ P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ Prop C \ (P \ !P) (\ \ (P' \ !P))" and rPar2: "\\ P' C. \\ \ !P \\ \ P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \); \C. Prop C \ (!P) (\ \ P')\ \ Prop C \ (P \ !P) (\ \ (P \ P'))" and rComm1: "\M N P' K xvec P'' C. \\ \ P \M\N\ \ P'; \ \ !P \K\\*xvec\\N\ \ P''; \C. Prop C \ (!P) (K\\*xvec\\N\ \ P''); \ \ M \ K; xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" and rComm2: "\M xvec N P' K P'' C. \\ \ P \M\\*xvec\\N\ \ P'; \ \ !P \K\N\ \ P''; \C. Prop C \ (!P) (K\N\ \ P''); \ \ M \ K; xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" and rBang: "\Rs C. \\ \ P \ !P \ Rs; \C. Prop C \ (P \ !P) Rs; guarded P\ \ Prop C \ (!P) Rs" shows "Prop C \ (!P) Rs" proof - from \\ \ !P \ Rs\ have "guarded P" by(nominal_induct \ P=="!P" Rs rule: semantics.strong_induct) (auto simp add: psi.inject) { fix Q :: "('a, 'b, 'c) psi" and \' :: 'b assume "\' \ Q \ Rs" and "guarded Q" and "bangPred P Q" and "\ \ \'" hence "Prop C \ Q Rs" using rPar1 rPar1 rPar2 rPar2 rComm1 rComm2 rBang proof(nominal_induct avoiding: \ C rule: semantics.strong_induct) case(cInput \' M K xvec N Tvec Q \ C) thus ?case by - (ind_cases "bangPred P (M\\*xvec N\.Q)") next case(Output \' M K N Q \ C) thus ?case by - (ind_cases "bangPred P (M\N\.Q)") next case(Case \' Q Rs \ Cs \ C) thus ?case by - (ind_cases "bangPred P (Cases Cs)") next case(cPar1 \' \\<^sub>R Q \ P' R A\<^sub>R \ C) have rPar1: "\\ P' C. \\ \ P \\ \ P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ Prop C \ (P \ !P) (\ \ (P' \ !P))" by fact from \bangPred P (Q \ R)\ have "Q = P" and "R = !P" by - (ind_cases "bangPred P (Q \ R)", auto simp add: psi.inject)+ from \R = !P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>R = []" and "\\<^sub>R = \" by auto from \\' \ \\<^sub>R \ Q \\ \ P'\ \Q = P\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ P \\ \ P'" by(metis statEqTransition Identity AssertionStatEqSym) hence "Prop C \ (P \ !P) (\ \ (P' \ !P))" using \bn \ \* \\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \Q = P\ \distinct(bn \)\ by(rule_tac rPar1) auto with \R = !P\ \Q = P\ show ?case by simp next case(cPar2 \' \\<^sub>P R \ P' Q A\<^sub>P \ C) have rPar2: "\\ P' C. \\ \ !P \\ \ P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \); \C. Prop C \ (!P) (\ \ P')\ \ Prop C \ (P \ !P) (\ \ (P \ P'))" by fact from \bangPred P (Q \ R)\ have "Q = P" and "R = !P" by - (ind_cases "bangPred P (Q \ R)", auto simp add: psi.inject)+ from \Q = P\ \extractFrame Q = \A\<^sub>P, \\<^sub>P\\ \guarded P\ have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" by(blast dest: guardedStatEq)+ from \\' \ \\<^sub>P \ R \\ \ P'\ \R = !P\ \\ \ \'\ \\\<^sub>P \ \\ have "\ \ !P \\ \ P'" by(metis statEqTransition Identity Composition Commutativity AssertionStatEqSym) moreover { fix C have "bangPred P (!P)" by(rule aux1) moreover from \\ \ \'\ \\\<^sub>P \ \\ have "\ \ \' \ \\<^sub>P" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans) ultimately have "\C. Prop C \ (!P) (\ \ P')" using cPar2 \R = !P\ \guarded P\ by simp } ultimately have "Prop C \ (P \ !P) (\ \ (P \ P'))" using \bn \ \* \\ \bn \ \* Q\ \bn \ \* subject \\ \bn \ \* C\ \Q = P\ \distinct(bn \)\ by(rule_tac rPar2) auto with \R = !P\ \Q = P\ show ?case by simp next case(cComm1 \' \\<^sub>R Q M N P' A\<^sub>P \\<^sub>P R K xvec P'' A\<^sub>R \ C) have rComm1: "\M N P' K xvec P'' C. \\ \ P \M\N\ \ P'; \ \ !P \K\\*xvec\\N\ \ P''; \C. Prop C \ (!P) (K\\*xvec\\N\ \ P''); \ \ M \ K; xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" by fact from \bangPred P (Q \ R)\ have "Q = P" and "R = !P" by - (ind_cases "bangPred P (Q \ R)", auto simp add: psi.inject)+ from \R = !P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>R = []" and "\\<^sub>R = \" by auto from \\' \ \\<^sub>R \ Q \M\N\ \ P'\ \Q = P\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ P \M\N\ \ P'" by(metis statEqTransition Identity AssertionStatEqSym) moreover from \Q = P\ \extractFrame Q = \A\<^sub>P, \\<^sub>P\\ \guarded P\ have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" by(blast dest: guardedStatEq)+ moreover from \\' \ \\<^sub>P \ R \K\\*xvec\\N\ \ P''\ \R = !P\ \\\<^sub>P \ \\ \\ \ \'\ have "\ \ !P \K\\*xvec\\N\ \ P''" by(metis statEqTransition Identity Composition Commutativity AssertionStatEqSym) moreover { fix C have "bangPred P (!P)" by(rule aux1) moreover from \\ \ \'\ \\\<^sub>P \ \\ have "\ \ \' \ \\<^sub>P" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans) ultimately have "\C. Prop C \ (!P) (K\\*xvec\\N\ \ P'')" using cComm1 \R = !P\ \guarded P\ by simp } moreover from \\' \ \\<^sub>P \ \\<^sub>R \ M \ K\ \\\<^sub>P \ \\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ M \ K" by(metis statEqEnt Identity Composition Commutativity AssertionStatEqSym) ultimately have "Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" using \xvec \* \\ \xvec \* Q\ \xvec \* M\ \xvec \* K\ \xvec \* C\ \Q = P\ \distinct xvec\ by(rule_tac rComm1) (assumption | auto)+ with \R = !P\ \Q = P\ show ?case by simp next case(cComm2 \' \\<^sub>R Q M xvec N P' A\<^sub>P \\<^sub>P R K P'' A\<^sub>R \ C) have rComm2: "\M xvec N P' K P'' C. \\ \ P \M\\*xvec\\N\ \ P'; \ \ !P \K\N\ \ P''; \C. Prop C \ (!P) (K\N\ \ P''); \ \ M \ K; xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" by fact from \bangPred P (Q \ R)\ have "Q = P" and "R = !P" by - (ind_cases "bangPred P (Q \ R)", auto simp add: psi.inject)+ from \R = !P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>R = []" and "\\<^sub>R = \" by auto from \\' \ \\<^sub>R \ Q \M\\*xvec\\N\ \ P'\ \Q = P\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ P \M\\*xvec\\N\ \ P'" by(metis statEqTransition Identity AssertionStatEqSym) moreover from \Q = P\ \extractFrame Q = \A\<^sub>P, \\<^sub>P\\ \guarded P\ have "\\<^sub>P \ \" and "supp \\<^sub>P = ({}::name set)" by(blast dest: guardedStatEq)+ moreover from \\' \ \\<^sub>P \ R \K\N\ \ P''\ \R = !P\ \\\<^sub>P \ \\ \\ \ \'\ have "\ \ !P \K\N\ \ P''" by(metis statEqTransition Identity Composition Commutativity AssertionStatEqSym) moreover { fix C have "bangPred P (!P)" by(rule aux1) moreover from \\ \ \'\ \\\<^sub>P \ \\ have "\ \ \' \ \\<^sub>P" by(metis Composition Identity Commutativity AssertionStatEqSym AssertionStatEqTrans) ultimately have "\C. Prop C \ (!P) (K\N\ \ P'')" using cComm2 \R = !P\ \guarded P\ by simp } moreover from \\' \ \\<^sub>P \ \\<^sub>R \ M \ K\ \\\<^sub>P \ \\ \\ \ \'\ \\\<^sub>R = \\ have "\ \ M \ K" by(metis statEqEnt Identity Composition Commutativity AssertionStatEqSym) ultimately have "Prop C \ (P \ !P) (\ \ \\*xvec\(P' \ P''))" using \xvec \* \\ \xvec \* Q\ \xvec \* M\ \xvec \* K\ \xvec \* C\ \Q = P\ \distinct xvec\ by(rule_tac rComm2) (assumption | auto)+ with \R = !P\ \Q = P\ show ?case by simp next case(cOpen \ Q M xvec yvec N P' x C) thus ?case by - (ind_cases "bangPred P (\\x\Q)") next case(cScope \ Q \ P' x C) thus ?case by - (ind_cases "bangPred P (\\x\Q)") next case(Bang \' Q Rs \ C) have rBang: "\Rs C. \\ \ P \ !P \ Rs; \C. Prop C \ (P \ !P) Rs; guarded P\ \ Prop C \ (!P) Rs" by fact from \bangPred P (!Q)\ have "P = Q" by - (ind_cases "bangPred P (!Q)", auto simp add: psi.inject) with \\' \ Q \ !Q \ Rs\ \\ \ \'\ have "\ \ P \ !P \ Rs" by(metis statEqTransition AssertionStatEqSym) moreover { fix C have "bangPred P (P \ !P)" by(rule aux2) with Bang \P = Q\ have "\C. Prop C \ (P \ !P) Rs" by simp } moreover from \guarded Q\ \P = Q\ have "guarded P" by simp ultimately have "Prop C \ (!P) Rs" by(rule rBang) with \P = Q\ show ?case by simp qed } with \guarded P\ \\ \ !P \ Rs\ show ?thesis by(force intro: aux1) qed lemma bangInputInduct[consumes 1, case_names cPar1 cPar2 cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and N :: 'a and P' :: "('a, 'b, 'c) psi" and Prop :: "'b \ ('a, 'b, 'c) psi \ 'a \ 'a \ ('a, 'b, 'c) psi \ bool" assumes "\ \ !P \M\N\ \ P'" and rPar1: "\P'. \ \ P \M\N\ \ P' \ Prop \ (P \ !P) M N (P' \ !P)" and rPar2: "\P'. \\ \ !P \M\N\ \ P'; Prop \ (!P) M N P'\ \ Prop \ (P \ !P) M N (P \ P')" and rBang: "\P'. \\ \ P \ !P \M\N\ \ P'; Prop \ (P \ !P) M N P'; guarded P\ \ Prop \ (!P) M N P'" shows "Prop \ (!P) M N P'" using \\ \ !P \M\N\ \ P'\ by(nominal_induct \ P Rs=="M\N\ \ P'" arbitrary: P' rule: bangInduct) (auto simp add: residualInject intro: rPar1 rPar2 rBang) lemma bangOutputInduct[consumes 1, case_names cPar1 cPar2 cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and N :: 'a and P' :: "('a, 'b, 'c) psi" and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a \ ('a, 'b, 'c) boundOutput \ bool" and C :: 'd assumes "\ \ !P \M\\*xvec\\N\ \ P'" and rPar1: "\xvec N P' C. \\ \ P \M\\*xvec\\N\ \ P'; xvec \* \; xvec \* P; xvec \* M; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) M (\\*xvec\N \' (P' \ !P))" and rPar2: "\xvec N P' C. \\ \ !P \M\\*xvec\\N\ \ P'; \C. Prop C \ (!P) M (\\*xvec\N \' P'); xvec \* \; xvec \* P; xvec \* M; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) M (\\*xvec\N \' (P \ P'))" and rBang: "\B C. \\ \ P \ !P \(ROut M B); \C. Prop C \ (P \ !P) M B; guarded P\ \ Prop C \ (!P) M B" shows "Prop C \ (!P) M (\\*xvec\N \' P')" using \\ \ !P \M\\*xvec\\N\ \ P'\ apply(auto simp add: residualInject) by(nominal_induct \ P Rs=="ROut M (\\*xvec\N \' P')" avoiding: C arbitrary: xvec N P' rule: bangInduct) (force simp add: residualInject intro: rPar1 rPar2 rBang)+ lemma bangTauInduct[consumes 1, case_names cPar1 cPar2 cComm1 cComm2 cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and P' :: "('a, 'b, 'c) psi" and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ ('a, 'b, 'c) psi \ bool" and C :: 'd assumes "\ \ !P \\ \ P'" and rPar1: "\P' C. \ \ P \\ \ P' \ Prop C \ (P \ !P) (P' \ !P)" and rPar2: "\P' C. \\ \ !P \\ \ P'; \C. Prop C \ (!P) P'\ \ Prop C \ (P \ !P) (P \ P')" and rComm1: "\M N P' K xvec P'' C. \\ \ P \M\N\ \ P'; \ \ !P \K\\*xvec\\N\ \ P''; \ \ M \ K; xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C\ \ Prop C \ (P \ !P) (\\*xvec\(P' \ P''))" and rComm2: "\M N P' K xvec P'' C. \\ \ P \M\\*xvec\\N\ \ P'; \ \ !P \K\N\ \ P''; \ \ M \ K; xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C\ \ Prop C \ (P \ !P) (\\*xvec\(P' \ P''))" and rBang: "\P' C. \\ \ P \ !P \\ \ P'; \C. Prop C \ (P \ !P) P'; guarded P\ \ Prop C \ (!P) P'" shows "Prop C \ (!P) P'" using \\ \ !P \\ \ P'\ by(nominal_induct \ P Rs=="\ \ P'" avoiding: C arbitrary: P' rule: bangInduct) (auto simp add: residualInject intro: rPar1 rPar2 rComm1 rComm2 rBang) lemma bangInduct'[consumes 2, case_names cAlpha cPar1 cPar2 cComm1 cComm2 cBang]: fixes \ :: 'b and P :: "('a, 'b, 'c) psi" and \ :: "'a action" and P' :: "('a, 'b, 'c) psi" and Prop :: "'d::fs_name \ 'b \ ('a, 'b, 'c) psi \ 'a action \ ('a, 'b, 'c) psi \ bool" and C :: "'d::fs_name" assumes "\ \ !P \\ \ P'" and "bn \ \* subject \" and rAlpha: "\\ P' p C. \bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; set p \ set(bn \) \ set(bn(p \ \)); distinctPerm p; bn(p \ \) \* \; bn(p \ \) \* P'; Prop C \ (P \ !P) \ P'\ \ Prop C \ (P \ !P) (p \ \) (p \ P')" and rPar1: "\\ P' C. \\ \ P \\ \ P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ Prop C \ (P \ !P) \ (P' \ !P)" and rPar2: "\\ P' C. \\ \ !P \\ \ P'; \C. Prop C \ (!P) \ P'; bn \ \* \; bn \ \* P; bn \ \* subject \; bn \ \* C; distinct(bn \)\ \ Prop C \ (P \ !P) \ (P \ P')" and rComm1: "\M N P' K xvec P'' C. \\ \ P \M\N\ \ P'; \ \ !P \K\\*xvec\\N\ \ P''; \C. Prop C \ (!P) (K\\*xvec\\N\) P''; \ \ M \ K; xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\) (\\*xvec\(P' \ P''))" and rComm2: "\M xvec N P' K P'' C. \\ \ P \M\\*xvec\\N\ \ P'; \ \ !P \K\N\ \ P''; \C. Prop C \ (!P) (K\N\) P''; \ \ M \ K; xvec \* \; xvec \* P; xvec \* M; xvec \* K; xvec \* C; distinct xvec\ \ Prop C \ (P \ !P) (\) (\\*xvec\(P' \ P''))" and rBang: "\\ P' C. \\ \ P \ !P \\ \ P'; guarded P; \C. Prop C \ (P \ !P) \ P'; guarded P; distinct(bn \)\ \ Prop C \ (!P) \ P'" shows "Prop C \ (!P) \ P'" proof - from \\ \ !P \\ \ P'\ have "distinct(bn \)" by(rule boundOutputDistinct) with \\ \ !P \\ \ P'\ \bn \ \* subject \\ show ?thesis proof(nominal_induct \ P Rs=="\ \ P'" avoiding: C \ P' rule: bangInduct) case(cPar1 \ P' C \' P'') note \\ \ (P' \ !P) = \' \ P''\ moreover from \bn \ \* \'\ have "bn \ \* bn \'" by simp moreover note \distinct(bn \)\ \distinct(bn \')\ moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P' \ !P)" by simp moreover from \bn \' \* subject \'\ have "bn \' \* (\' \ P'')" by simp ultimately obtain p where S: "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "\' = p \ \" and P'eq: "P'' = p \ (P' \ !P)" and "bn(p \ \) \* \" and "bn(p \ \) \* (P' \ !P)" by(rule_tac residualEq) from \\ \ P \\ \ P'\ \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ \distinct(bn \)\ have "Prop C \ (P \ !P) \ (P' \ !P)" by(rule rPar1) with \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ S \distinctPerm p\ \bn \ \* \'\ \bn \ \* P''\ \\' = (p \ \)\ P'eq \bn(p \ \) \* \\ \bn(p \ \) \* (P' \ !P)\ have "Prop C \ (P \ !P) (p \ \) (p \ (P' \ !P))" by(rule_tac rAlpha) with P'eq \\' = p \ \\ \distinctPerm p\ show ?case by simp next case(cPar2 \ P' C \' P'') note \\ \ (P \ P') = \' \ P''\ moreover from \bn \ \* \'\ have "bn \ \* bn \'" by simp moreover note \distinct(bn \)\ \distinct(bn \')\ moreover from \bn \ \* subject \\ have "bn \ \* (\ \ P \ P')" by simp moreover from \bn \' \* subject \'\ have "bn \' \* (\' \ P'')" by simp ultimately obtain p where S: "set p \ set(bn \) \ set(bn(p \ \))" and "distinctPerm p" and "\' = p \ \" and P'eq: "P'' = p \ (P \ P')" and "bn(p \ \) \* \" and "bn(p \ \) \* (P \ P')" by(rule_tac residualEq) note \\ \ !P \\ \ P'\ moreover from \bn \ \* subject \\ \distinct(bn \)\ have "\C. Prop C \ (!P) \ P'" by(rule_tac cPar2) auto moreover note \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ \distinct(bn \)\ ultimately have "Prop C \ (P \ !P) \ (P \ P')" by(rule rPar2) with \bn \ \* \\ \bn \ \* P\ \bn \ \* subject \\ \bn \ \* C\ S \distinctPerm p\ \bn \ \* \'\ \bn \ \* P''\ \\' = (p \ \)\ P'eq \bn(p \ \) \* \\ \bn(p \ \) \* (P \ P')\ have "Prop C \ (P \ !P) (p \ \) (p \ (P \ P'))" by(rule_tac rAlpha) with P'eq \\' = p \ \\ show ?case by simp next case(cComm1 M N P' K xvec P'' C \ P''') hence "Prop C \ (P \ !P) (\) (\\*xvec\(P' \ P''))" by(rule_tac rComm1) (assumption | simp)+ thus ?case using \\ \ \\*xvec\(P' \ P'') = \ \ P'''\ by(simp add: residualInject) next case(cComm2 M xvec N P' K P'' C \ P''') hence "Prop C \ (P \ !P) (\) (\\*xvec\(P' \ P''))" by(rule_tac rComm2) (assumption | simp)+ thus ?case using \\ \ \\*xvec\(P' \ P'') = \ \ P'''\ by(simp add: residualInject) next case(cBang C \ P') thus ?case by(auto intro: rBang) qed qed lemma comm1Aux: fixes \ :: 'b and \\<^sub>Q :: 'b and R :: "('a, 'b, 'c) psi" and K :: 'a and xvec :: "name list" and N :: 'a and R' :: "('a, 'b, 'c) psi" and A\<^sub>R :: "name list" and \\<^sub>R :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and L :: 'a and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and A\<^sub>Q :: "name list" assumes RTrans: "\ \ \\<^sub>Q \ R \K\\*xvec\\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" and PTrans: "\ \ \\<^sub>R \ P \M\L\ \ P'" and MeqK: "\ \ \\<^sub>Q \ \\<^sub>R \ M \ K" and PeqQ: "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>P" and "distinct A\<^sub>R" and "A\<^sub>R \* A\<^sub>P" and "A\<^sub>R \* A\<^sub>Q" and "A\<^sub>R \* \" and "A\<^sub>R \* P" and "A\<^sub>R \* Q" and "A\<^sub>R \* R" and "A\<^sub>R \* K" and "A\<^sub>P \* \" and "A\<^sub>P \* R" and "A\<^sub>P \* P" and "A\<^sub>P \* M" and "A\<^sub>Q \* R" and "A\<^sub>Q \* M" obtains K' where "\ \ \\<^sub>P \ R \K'\\*xvec\\N\ \ R'" and "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "A\<^sub>R \* K'" proof - from \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ FrP FrQ have "A\<^sub>R \* \\<^sub>P" and "A\<^sub>R \* \\<^sub>Q" by(force dest: extractFrameFreshChain)+ assume Assumptions: "\K'. \\ \ \\<^sub>P \ R \K'\\*xvec\\N\ \ R'; \ \ \\<^sub>P \ \\<^sub>R \ M \ K'; A\<^sub>R \* K'\ \ thesis" { fix \' ::'b and zvec ::"name list" assume A: "\ \ \\<^sub>Q \ \'" assume "\' \ R \K\\*xvec\\N\ \ R'" hence "\' \ R \ROut K (\\*xvec\N \' R')" by(simp add: residualInject) moreover note FrR \distinct A\<^sub>R\ PTrans moreover from \\' \ R \K\\*xvec\\N\ \ R'\ have "distinct xvec" by(auto dest: boundOutputDistinct) moreover assume "\' \ \\<^sub>R \ M \ K" and "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" and "A\<^sub>R \* zvec" and "A\<^sub>P \* zvec" and "zvec \* R" and "zvec \* P" and "A\<^sub>R \* \'" ultimately have "\K'. \ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R') \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using FrP \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \distinct A\<^sub>P\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* R\ \A\<^sub>P \* M\ \A\<^sub>Q \* M\ \A\<^sub>R \* K\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ \A\<^sub>R \* \\<^sub>P\ proof(nominal_induct \' R K B=="\\*xvec\N \' R'" A\<^sub>R \\<^sub>R avoiding: \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec N R' arbitrary: M rule: outputFrameInduct) case(cAlpha \' R K A\<^sub>R \\<^sub>R p \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec N R' M) have S: "set p \ set A\<^sub>R \ set (p \ A\<^sub>R)" by fact from \\' \ (p \ \\<^sub>R) \ M \ K\ have "(p \ (\' \ (p \ \\<^sub>R))) \ (p \ M) \ (p \ K)" by(rule chanEqClosed) with \A\<^sub>R \* \'\ \(p \ A\<^sub>R) \* \'\ \A\<^sub>R \* K\ \(p \ A\<^sub>R) \* K\ S \distinctPerm p\ have "\' \ \\<^sub>R \ (p \ M) \ K" by(simp add: eqvts) moreover from \\ \ (p \ \\<^sub>R) \ P \M\L\ \ P'\ S \A\<^sub>R \* P\ \(p \ A\<^sub>R) \* P\ have "(p \ (\ \ (p \ \\<^sub>R))) \ P \(p \ M)\L\ \ P'" by(rule_tac inputPermFrameSubject) auto with \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ S \distinctPerm p\ have "\ \ \\<^sub>R \ P \(p \ M)\L\ \ P'" by(simp add: eqvts) moreover from \A\<^sub>P \* M\ have "(p \ A\<^sub>P) \* (p \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>R \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>P\ S have "A\<^sub>P \* (p \ M)" by simp moreover from \A\<^sub>Q \* M\ have "(p \ A\<^sub>Q) \* (p \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>R \* A\<^sub>Q\ \(p \ A\<^sub>R) \* A\<^sub>Q\ S have "A\<^sub>Q \* (p \ M)" by simp moreover from \\A\<^sub>Q, \' \ (p \ \\<^sub>R)\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\ \ have "(p \ \A\<^sub>Q, \' \ (p \ \\<^sub>R)\) \\<^sub>F (p \ \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\)" by(rule FrameStatImpClosed) with \A\<^sub>R \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>P\ \A\<^sub>R \* \'\ \(p \ A\<^sub>R) \* \'\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ \(p \ A\<^sub>R) \* A\<^sub>Q\ \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ S \distinctPerm p\ have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" by(simp add: eqvts) ultimately obtain K' where "\ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R')" and "\ \ \\<^sub>P \ \\<^sub>R \ (p \ M) \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" using cAlpha by metis from \\ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R')\ S \A\<^sub>R \* R\ \(p \ A\<^sub>R) \* R\ have "(p \ (\ \ \\<^sub>P)) \ R \(ROut (p \ K') (\\*xvec\N \' R'))" using outputPermFrameSubject by(auto simp add: residualInject) with S \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ have "\ \ \\<^sub>P \ R \(ROut (p \ K') (\\*xvec\N \' R'))" by(simp add: eqvts) moreover from \\ \ \\<^sub>P \ \\<^sub>R \ (p \ M) \ K'\ have "(p \ (\ \ \\<^sub>P \ \\<^sub>R))\ (p \ p \ M) \ (p \ K')" by(rule chanEqClosed) with S \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ \distinctPerm p\ have "\ \ \\<^sub>P \ (p \ \\<^sub>R) \ M \ (p \ K')" by(simp add: eqvts) moreover from \zvec \* K'\ have "(p \ zvec) \* (p \ K')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>R \* zvec\ \(p \ A\<^sub>R) \* zvec\ S have "zvec \* (p \ K')" by simp moreover from \A\<^sub>R \* K'\ have "(p \ A\<^sub>R) \* (p \ K')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) ultimately show ?case by blast next case(cOutput \' M' K N R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec N' R' M) from \A\<^sub>P \* (M'\N\.R)\ \A\<^sub>Q \* (M'\N\.R)\ \zvec \* (M'\N\.R)\ have "A\<^sub>P \* M'" and "A\<^sub>Q \* M'" and "zvec \* M'" by simp+ from \\' \ M' \ K\ have "\' \ \ \ M' \ K" by(blast intro: statEqEnt Identity AssertionStatEqSym) hence "\' \ \ \ M' \ M'" by(blast intro: chanEqSym chanEqTrans) with \A\<^sub>Q \* M'\ have "(\A\<^sub>Q, \' \ \\) \\<^sub>F M' \ M'" by(force intro: frameImpI) with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F M' \ M'" by(simp add: FrameStatImp_def) with \A\<^sub>P \* M'\ have "(\ \ \\<^sub>P) \ \ \ M' \ M'" by(force dest: frameImpE) hence "\ \ \\<^sub>P \ M' \ M'" by(blast intro: statEqEnt Identity) hence "\ \ \\<^sub>P \ M'\N\.R \M'\N\ \ R" by(rule Output) moreover from \\' \ \ \ M \ K\ \\' \ \ \ M' \ K\ have "\' \ \ \ M \ M'" by(metis chanEqSym chanEqTrans) with \A\<^sub>Q \* M\ \A\<^sub>Q \* M'\ have "(\A\<^sub>Q, \' \ \\) \\<^sub>F M \ M'" by(force intro: frameImpI) with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F M \ M'" by(simp add: FrameStatImp_def) with \A\<^sub>P \* M\ \A\<^sub>P \* M'\ have "(\ \ \\<^sub>P) \ \ \ M \ M'" by(force dest: frameImpE) hence "\ \ \\<^sub>P \ \ \ M \ M'" by(metis statEqEnt Associativity) ultimately show ?case using cOutput by(auto simp add: residualInject) next case(cCase \' R M' \ Cs A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec N R' M) from \guarded R\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "\\<^sub>R \ \" by(metis guardedStatEq) with \\' \ \ \ M \ M'\ have "\' \ \\<^sub>R \ M \ M'" by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition) moreover have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" proof - from \\\<^sub>R \ \\ have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, \' \ \\" by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) moreover note \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ moreover from \\\<^sub>R \ \\ have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) ultimately show ?thesis by(rule FrameStatEqImpCompose) qed moreover from \\ \ \ \ P \M\L\ \ P'\ \\\<^sub>R \ \\ have "\ \ \\<^sub>R \ P \M\L\ \ P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition) moreover from \zvec \* (Cases Cs)\ \A\<^sub>P \* (Cases Cs)\ \A\<^sub>Q \* (Cases Cs)\ \(\, R) mem Cs\ have "A\<^sub>P \* R" and "A\<^sub>Q \* R" and "zvec \* R" and "A\<^sub>P \* \" and "A\<^sub>Q \* \" by(auto dest: memFreshChain) ultimately have "\K'. \ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R') \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using cCase by(rule_tac cCase) (assumption | simp)+ then obtain K' where RTrans: "\ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R')" and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" by metis note RTrans \(\, R) mem Cs\ moreover from \\' \ \\ have "\' \ \ \ \" by(blast intro: statEqEnt Identity AssertionStatEqSym) with \A\<^sub>Q \* \\ have "(\A\<^sub>Q, \' \ \\) \\<^sub>F \" by(force intro: frameImpI) with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F \" by(simp add: FrameStatImp_def) with \A\<^sub>P \* \\ have "(\ \ \\<^sub>P) \ \ \ \" by(force dest: frameImpE) hence "\ \ \\<^sub>P \ \" by(blast intro: statEqEnt Identity) ultimately have "\ \ \\<^sub>P \ Cases Cs \ROut K' (\\*xvec\N \' R')" using \guarded R\ by(rule Case) moreover from MeqK' \\\<^sub>R \ \\ have "\ \ \\<^sub>P \ \ \ M \ K'" by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans) ultimately show ?case using \zvec \* K'\ by fastforce next case(cPar1 \' \\<^sub>R\<^sub>2 R\<^sub>1 M' xvec N' R\<^sub>1' A\<^sub>R\<^sub>2 R\<^sub>2 A\<^sub>R\<^sub>1 \\<^sub>R\<^sub>1 \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec yvec N R' M) have FrR2: "extractFrame R\<^sub>2 = \A\<^sub>R\<^sub>2, \\<^sub>R\<^sub>2\" by fact from \\' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ M'\ have "(\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1 \ M \ M'" by(metis statEqEnt Associativity Composition Commutativity) moreover have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P) \ \\<^sub>R\<^sub>1\" proof - have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1\ \\<^sub>F \A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\" by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) moreover note \\A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\\ moreover have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P) \ \\<^sub>R\<^sub>1\" by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) ultimately show ?thesis by(rule FrameStatEqImpCompose) qed moreover from \\ \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ P \M\L\ \ P'\ have "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1 \ P \M\L\ \ P'" by(metis statEqTransition Associativity Composition Commutativity) moreover from \A\<^sub>R\<^sub>1 \* A\<^sub>P\ \A\<^sub>R\<^sub>2 \* A\<^sub>P\ \A\<^sub>P \* (R\<^sub>1 \ R\<^sub>2)\ \extractFrame R\<^sub>1 = \A\<^sub>R\<^sub>1, \\<^sub>R\<^sub>1\\ FrR2 have "A\<^sub>P \* \\<^sub>R\<^sub>1" and "A\<^sub>P \* \\<^sub>R\<^sub>2" by(force dest: extractFrameFreshChain)+ moreover from \\\*xvec\N' \' (R\<^sub>1' \ R\<^sub>2) = \\*yvec\N \' R'\ \xvec \* yvec\ obtain p T where "\\*xvec\N' \' R\<^sub>1' = \\*yvec\N \' T" and "R' = T \ (p \ R\<^sub>2)" and "set p \ set yvec \ set xvec" apply(drule_tac sym) by(rule_tac boundOutputPar1Dest') (assumption | simp | blast dest: sym)+ ultimately have "\K'. (\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ R\<^sub>1 \ROut K' (\\*yvec\N \' T) \ (\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ M \ K' \ (A\<^sub>R\<^sub>2@zvec) \* K' \ A\<^sub>R\<^sub>1 \* K'" using cPar1 apply(rule_tac cPar1(6)) by(assumption | simp | fastforce)+ then obtain K' where RTrans: "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ R\<^sub>1 \K'\\*xvec\\N'\ \ R\<^sub>1'" and MeqK': "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ M \ K'" and "A\<^sub>R\<^sub>2 \* K'" and "A\<^sub>R\<^sub>1 \* K'" and "zvec \* K'" using \\\*xvec\N' \' R\<^sub>1' = \\*yvec\N \' T\ by(auto simp add: residualInject) from RTrans have "(\ \ \\<^sub>P) \ \\<^sub>R\<^sub>2 \ R\<^sub>1 \K'\\*xvec\\N'\ \ R\<^sub>1'" by(metis statEqTransition Associativity Composition Commutativity) hence "\ \ \\<^sub>P \ (R\<^sub>1 \ R\<^sub>2) \K'\\*xvec\\N'\ \ (R\<^sub>1' \ R\<^sub>2)" using FrR2 \xvec \* R\<^sub>2\ \A\<^sub>R\<^sub>2 \* \\ \A\<^sub>R\<^sub>2 \* \\<^sub>P\ \A\<^sub>R\<^sub>2 \* K'\ \A\<^sub>R\<^sub>2 \* R\<^sub>1\ \A\<^sub>R\<^sub>2 \* xvec\ \A\<^sub>R\<^sub>2 \* N'\ by(force intro: Par1) moreover from MeqK' have "\ \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ K'" by(metis statEqEnt Associativity Composition Commutativity) ultimately show ?case using \zvec \* K'\ \A\<^sub>R\<^sub>1 \* K'\ \A\<^sub>R\<^sub>2 \* K'\ \\\*xvec\N' \' (R\<^sub>1' \ R\<^sub>2) = \\*yvec\N \' R'\ by(auto simp add: residualInject) next case(cPar2 \' \\<^sub>R\<^sub>1 R\<^sub>2 M' xvec N' R\<^sub>2' A\<^sub>R\<^sub>1 R\<^sub>1 A\<^sub>R\<^sub>2 \\<^sub>R\<^sub>2 \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec yvec N R' M) have FrR1: "extractFrame R\<^sub>1 = \A\<^sub>R\<^sub>1, \\<^sub>R\<^sub>1\" by fact from \\' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ M'\ have "(\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2 \ M \ M'" by(metis statEqEnt Associativity Composition Commutativity) moreover have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P) \ \\<^sub>R\<^sub>2\" proof - have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\" by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) moreover note \\A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\\ moreover have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P) \ \\<^sub>R\<^sub>2\" by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) ultimately show ?thesis by(rule FrameStatEqImpCompose) qed moreover from \\ \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ P \M\L\ \ P'\ have "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2 \ P \M\L\ \ P'" by(metis statEqTransition Associativity Composition Commutativity) moreover from \A\<^sub>R\<^sub>1 \* A\<^sub>P\ \A\<^sub>R\<^sub>2 \* A\<^sub>P\ \A\<^sub>P \* (R\<^sub>1 \ R\<^sub>2)\ FrR1 \extractFrame R\<^sub>2 = \A\<^sub>R\<^sub>2, \\<^sub>R\<^sub>2\\ have "A\<^sub>P \* \\<^sub>R\<^sub>1" and "A\<^sub>P \* \\<^sub>R\<^sub>2" by(force dest: extractFrameFreshChain)+ moreover from \\\*xvec\N' \' (R\<^sub>1 \ R\<^sub>2') = \\*yvec\N \' R'\ \xvec \* yvec\ obtain p T where "\\*xvec\N' \' R\<^sub>2' = \\*yvec\N \' T" and "R' = (p \ R\<^sub>1) \ T" and "set p \ set yvec \ set xvec" apply(drule_tac sym) by(rule_tac boundOutputPar2Dest') (assumption | simp | blast dest: sym)+ ultimately have "\K'. (\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ R\<^sub>2 \ROut K' (\\*yvec\N \' T) \ (\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ \\<^sub>R\<^sub>2 \ M \ K' \ (A\<^sub>R\<^sub>1@zvec) \* K' \ A\<^sub>R\<^sub>2 \* K'" using cPar2 by(rule_tac cPar2(6)) (assumption | simp | fastforce)+ then obtain K' where RTrans: "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ R\<^sub>2 \K'\\*xvec\\N'\ \ R\<^sub>2'" and MeqK': "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ \\<^sub>R\<^sub>2 \ M \ K'" and "A\<^sub>R\<^sub>1 \* K'" and "zvec \* K'" and "A\<^sub>R\<^sub>2 \* K'" using \\\*xvec\N' \' R\<^sub>2' = \\*yvec\N \' T\ by(auto simp add: residualInject) from RTrans have "(\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ R\<^sub>2 \K'\\*xvec\\N'\ \ R\<^sub>2'" by(metis statEqTransition Associativity Composition Commutativity) hence "\ \ \\<^sub>P \ (R\<^sub>1 \ R\<^sub>2) \K'\\*xvec\\N'\ \ (R\<^sub>1 \ R\<^sub>2')" using FrR1 \xvec \* R\<^sub>1\ \A\<^sub>R\<^sub>1 \* \\ \A\<^sub>R\<^sub>1 \* \\<^sub>P\ \A\<^sub>R\<^sub>1 \* K'\\A\<^sub>R\<^sub>1 \* xvec\ \A\<^sub>R\<^sub>1 \* N'\ \A\<^sub>R\<^sub>1 \* R\<^sub>2\ by(force intro: Par2) moreover from MeqK' have "\ \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ K'" by(metis statEqEnt Associativity Composition Commutativity) ultimately show ?case using \zvec \* K'\ \A\<^sub>R\<^sub>1 \* K'\ \A\<^sub>R\<^sub>2 \* K'\ \\\*xvec\N' \' (R\<^sub>1 \ R\<^sub>2') = \\*yvec\N \' R'\ by(auto simp add: residualInject) next case(cOpen \' R M' xvec yvec N' R' x A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec zvec2 N R'' M) from \\\*(xvec @ x # yvec)\N' \' R' = \\*zvec2\N \' R''\ \x \ xvec\ \x \ yvec\ \x \ zvec2\ \x \ R''\ \x \ N\ \distinct zvec2\ obtain xvec' x' yvec' where A: "\\*(xvec@yvec)\N' \' R' = \\*(xvec'@yvec')\([(x, x')] \ N) \' ([(x, x')] \ R'')" and B: "zvec2 = (xvec'@x'#yvec')" by(rule_tac boundOutputOpenDest) auto then have "\K'. \ \ \\<^sub>P \ R \ROut K' (\\*(xvec'@yvec')\([(x, x')] \ N) \' ([(x, x')] \ R'')) \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ (zvec) \* K' \ A\<^sub>R \* K'" using cOpen by(rule_tac cOpen(4)) (assumption | simp)+ then obtain K' where RTrans: "\ \ \\<^sub>P \ R \K'\\*(xvec@yvec)\\N'\ \ R'" and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" using A by(auto simp add: residualInject) from \A\<^sub>R \* A\<^sub>P\ \A\<^sub>P \* (\\x\R)\ \x \ A\<^sub>P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>P \* \\<^sub>R" by(force dest: extractFrameFreshChain)+ from \\ \ \\<^sub>R \ P \M\L\ \ P'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \zvec \* P\ \A\<^sub>P \* \\<^sub>R\ \x \ A\<^sub>P\ \A\<^sub>P \* M\ \A\<^sub>P \* P\ \A\<^sub>P \* zvec\ \A\<^sub>P \* \\ \A\<^sub>P \* zvec\ \A\<^sub>R \* P\ \A\<^sub>R \* A\<^sub>P\ \x \ A\<^sub>P\ \x \ P\ obtain K'' where MeqK'': "(\ \ \\<^sub>R) \ \\<^sub>P \ M \ K''" and "A\<^sub>R \* K''" and "zvec \* K''" and "x \ K''" by(rule_tac B="(x#A\<^sub>R@zvec)" in obtainPrefix) (assumption | simp | force)+ from MeqK'' MeqK' have KeqK'': "(\ \ \\<^sub>P) \ \\<^sub>R \ K' \ K''" by(metis statEqEnt Associativity Composition Commutativity chanEqSym chanEqTrans) with RTrans \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \distinct A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* K'\ \A\<^sub>R \* K''\ \A\<^sub>R \* R\ have "\ \ \\<^sub>P \ R \K''\\*(xvec@yvec)\\N'\ \ R'" by(rule_tac outputRenameSubject) (assumption | force)+ hence "\ \ \\<^sub>P \ \\x\R \K''\\*(xvec@x#yvec)\\N'\ \ R'" using \x \ \\ \x \ \\<^sub>P\ \x \ K''\ \x \ xvec\ \ x \ yvec\ \x \ supp N'\ \xvec \* \\ \xvec \* \\<^sub>P\ \xvec \* R\ \x \ K''\ by(rule_tac Open) (assumption | force)+ moreover from MeqK'' have "\ \ \\<^sub>P \ \\<^sub>R \ M \ K''" by(metis statEqEnt Associativity Composition Commutativity) ultimately show ?case using \zvec \* K''\ \x \ K''\ \A\<^sub>R \* K''\ B \\\*(xvec @ x # yvec)\N' \' R' = \\*zvec2\N \' R''\ by(auto simp add: residualInject) next case(cScope \' R M' xvec N' R' x A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec yvec N R'' M) from \\\*xvec\N' \' \\x\R' = \\*yvec\N \' R''\ \x \ xvec\ \x \ yvec\ obtain R''' where "R'' = \\x\R'''" and "\\*xvec\N' \' R' = \\*yvec\N \' R'''" apply(drule_tac sym) by(rule boundOutputScopeDest) (assumption | auto)+ then have "\K'. \ \ \\<^sub>P \ R \ROut K' (\\*yvec\N \' R''') \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using cScope by(rule_tac cScope(4)) (assumption | simp)+ then obtain K' where RTrans: "\ \ \\<^sub>P \ R \K'\\*xvec\\N'\ \ R'" and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" using \\\*xvec\N' \' R' = \\*yvec\N \' R'''\ by(auto simp add: residualInject) from \A\<^sub>R \* A\<^sub>P\ \A\<^sub>P \* (\\x\R)\ \x \ A\<^sub>P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>P \* \\<^sub>R" by(force dest: extractFrameFreshChain)+ from \\ \ \\<^sub>R \ P \M\L\ \ P'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \x \ P\ \zvec \* P\ \A\<^sub>P \* \\<^sub>R\ \x \ A\<^sub>P\ \A\<^sub>P \* M\ \A\<^sub>P \* P\ \A\<^sub>P \* zvec\ \A\<^sub>P \* \\ \A\<^sub>P \* zvec\ \A\<^sub>R \* P\ \A\<^sub>R \* A\<^sub>P\ obtain K'' where MeqK'': "(\ \ \\<^sub>R) \ \\<^sub>P \ M \ K''" and "x \ K''" and "A\<^sub>R \* K''" and "zvec \* K''" by(rule_tac B="(x#A\<^sub>R@zvec)" in obtainPrefix) (assumption | force)+ from MeqK'' MeqK' have KeqK'': "(\ \ \\<^sub>P) \ \\<^sub>R \ K' \ K''" by(metis statEqEnt Associativity Composition Commutativity chanEqSym chanEqTrans) with RTrans \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \distinct A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* K'\ \A\<^sub>R \* K''\ \A\<^sub>R \* R\ have "\ \ \\<^sub>P \ R \K''\\*xvec\\N'\ \ R'" by(rule_tac outputRenameSubject) (assumption | force)+ hence "\ \ \\<^sub>P \ \\x\R \K''\\*xvec\\N'\ \ \\x\R'" using \x \ \\ \x \ \\<^sub>P\ \x \ K''\ \x \ xvec\ \x \ N'\ \xvec \* \\ \xvec \* \\<^sub>P\ \xvec \* R\ \x \ K''\ by(rule_tac Scope) (assumption | force)+ moreover from MeqK'' have "\ \ \\<^sub>P \ \\<^sub>R \ M \ K''" by(metis statEqEnt Associativity Composition Commutativity) ultimately show ?case using \zvec \* K''\ \x \ K''\ \A\<^sub>R \* K''\ \\\*xvec\N' \' \\x\R' = \\*yvec\N \' R''\ by(auto simp add: residualInject) next case(cBang \' R M' A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec N R' M) from \guarded R\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "\\<^sub>R \ \" by(metis guardedStatEq) with \\' \ \ \ M \ M'\ have "\' \ \\<^sub>R \ \ \ M \ M'" by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition) moreover have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" proof - from \\\<^sub>R \ \\ have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>Q, \' \ \\" by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) moreover note \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ moreover from \\\<^sub>R \ \\ have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) ultimately show ?thesis by(rule FrameStatEqImpCompose) qed moreover from \\ \ \ \ P \M\L\ \ P'\ \\\<^sub>R \ \\ have "\ \ \\<^sub>R \ \ \ P \M\L\ \ P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition) ultimately have "\K'. \ \ \\<^sub>P \ R \ !R \ROut K' (\\*xvec\N \' R') \ \ \ \\<^sub>P \ \\<^sub>R \ \ \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using cBang by(rule_tac cBang(5)) (assumption |simp)+ then obtain K' where RTrans: "\ \ \\<^sub>P \ R \ !R \ROut K' (\\*xvec\N \' R')" and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ \ \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" by metis from RTrans \guarded R\ have "\ \ \\<^sub>P \ !R \ROut K' (\\*xvec\N \' R')" by(rule Bang) moreover from MeqK' \\\<^sub>R \ \\ have "\ \ \\<^sub>P \ \ \ M \ K'" by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans) ultimately show ?case using \zvec \* K'\ by force qed } note Goal = this have "\ \ \\<^sub>Q \ \ \ \\<^sub>Q" by simp moreover note RTrans moreover from MeqK have "(\ \ \\<^sub>Q) \ \\<^sub>R \ M \ K" by(metis statEqEnt Associativity Commutativity) moreover note PeqQ \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>Q\ ultimately have "\K'. \ \ \\<^sub>P \ R \ROut K' (\\*xvec\N \' R') \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ ([]::name list) \* K' \ A\<^sub>R \* K'" by(rule_tac Goal) (assumption | force simp add: residualInject)+ with Assumptions show ?thesis by(force simp add: residualInject) qed lemma comm2Aux: fixes \ :: 'b and \\<^sub>Q :: 'b and R :: "('a, 'b, 'c) psi" and K :: 'a and N :: 'a and R' :: "('a, 'b, 'c) psi" and A\<^sub>R :: "name list" and \\<^sub>R :: 'b and P :: "('a, 'b, 'c) psi" and M :: 'a and xvec :: "name list" and P' :: "('a, 'b, 'c) psi" and A\<^sub>P :: "name list" and \\<^sub>P :: 'b and A\<^sub>Q :: "name list" assumes RTrans: "\ \ \\<^sub>Q \ R \K\N\ \ R'" and FrR: "extractFrame R = \A\<^sub>R, \\<^sub>R\" and PTrans: "\ \ \\<^sub>R \ P \M\\*xvec\\N\ \ P'" and MeqK: "\ \ \\<^sub>Q \ \\<^sub>R \ M \ K" and QimpP: "\A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" and FrP: "extractFrame P = \A\<^sub>P, \\<^sub>P\" and FrQ: "extractFrame Q = \A\<^sub>Q, \\<^sub>Q\" and "distinct A\<^sub>P" and "distinct A\<^sub>R" and "A\<^sub>R \* A\<^sub>P" and "A\<^sub>R \* A\<^sub>Q" and "A\<^sub>R \* \" and "A\<^sub>R \* P" and "A\<^sub>R \* Q" and "A\<^sub>R \* R" and "A\<^sub>R \* K" and "A\<^sub>P \* \" and "A\<^sub>P \* R" and "A\<^sub>P \* P" and "A\<^sub>P \* M" and "A\<^sub>Q \* R" and "A\<^sub>Q \* M" and "A\<^sub>R \* xvec" and "xvec \* M" obtains K' where "\ \ \\<^sub>P \ R \K'\N\ \ R'" and "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "A\<^sub>R \* K'" proof - from \A\<^sub>R \* P\ \A\<^sub>R \* Q\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ FrP FrQ have "A\<^sub>R \* \\<^sub>P" and "A\<^sub>R \* \\<^sub>Q" by(force dest: extractFrameFreshChain)+ assume Assumptions: "\K'. \\ \ \\<^sub>P \ R \K'\N\ \ R'; \ \ \\<^sub>P \ \\<^sub>R \ M \ K'; A\<^sub>R \* K'\ \ thesis" { fix \'::'b fix zvec::"name list" assume "A\<^sub>R \* \'" assume "A\<^sub>R \* zvec" assume "A\<^sub>P \* zvec" assume "zvec \* R" assume "zvec \* P" assume A: "\ \ \\<^sub>Q \ \'" with RTrans have "\' \ R \K\N\ \ R'" by(rule statEqTransition) moreover note FrR \distinct A\<^sub>R\ moreover from \\ \ \\<^sub>Q \ \\<^sub>R \ M \ K\ have "(\ \ \\<^sub>Q) \ \\<^sub>R \ M \ K" by(blast intro: statEqEnt Associativity AssertionStatEqSym) with A have "\' \ \\<^sub>R \ M \ K" by(rule statEqEnt[OF Composition]) moreover have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, (\ \ \\<^sub>Q) \ \\<^sub>R\" using A by(blast dest: frameIntComposition FrameStatEqTrans FrameStatEqSym) with QimpP have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" by(force intro: FrameStatEqImpCompose) moreover from PTrans have "distinct xvec" by(auto dest: boundOutputDistinct) ultimately have "\K'. \ \ \\<^sub>P \ R \K'\N\ \ R' \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using PTrans FrP \A\<^sub>R \* K\ \A\<^sub>R \* \'\ \A\<^sub>R \* \\ \A\<^sub>R \* P\ \A\<^sub>R \* R\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>P \* R\ \A\<^sub>Q \* R\ \A\<^sub>P \* \\ \A\<^sub>P \* P\ \A\<^sub>P \* M\ \A\<^sub>P \* zvec\ \A\<^sub>Q \* M\ \A\<^sub>R \* zvec\ \zvec \* R\ \zvec \* P\ \distinct A\<^sub>P\ \A\<^sub>R \* A\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ \A\<^sub>R \* xvec\ \xvec \* M\ proof(nominal_induct avoiding: \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec arbitrary: M rule: inputFrameInduct) case(cAlpha \' R K N R' A\<^sub>R \\<^sub>R p \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) have S: "set p \ set A\<^sub>R \ set (p \ A\<^sub>R)" by fact from \\' \ (p \ \\<^sub>R) \ M \ K\ have "(p \ (\' \ (p \ \\<^sub>R))) \ (p \ M) \ (p \ K)" by(rule chanEqClosed) with \A\<^sub>R \* \'\ \(p \ A\<^sub>R) \* \'\ \A\<^sub>R \* K\ \(p \ A\<^sub>R) \* K\ S \distinctPerm p\ have "\' \ \\<^sub>R \ (p \ M) \ K" by(simp add: eqvts) moreover from \\ \ (p \ \\<^sub>R) \ P \M\\*xvec\\N\ \ P'\ S \A\<^sub>R \* P\ \(p \ A\<^sub>R) \* P\ \A\<^sub>R \* xvec\ \(p \ A\<^sub>R) \* xvec\ \xvec \* M\ have "(p \ (\ \ (p \ \\<^sub>R))) \ P \(p \ M)\\*xvec\\N\ \ P'" using outputPermFrameSubject by(auto simp add: residualInject) with \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ S \distinctPerm p\ have "\ \ \\<^sub>R \ P \(p \ M)\\*xvec\\N\ \ P'" by(simp add: eqvts) moreover from \A\<^sub>P \* M\ have "(p \ A\<^sub>P) \* (p \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>R \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>P\ S have "A\<^sub>P \* (p \ M)" by simp moreover from \A\<^sub>Q \* M\ have "(p \ A\<^sub>Q) \* (p \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>R \* A\<^sub>Q\ \(p \ A\<^sub>R) \* A\<^sub>Q\ S have "A\<^sub>Q \* (p \ M)" by simp moreover from \\A\<^sub>Q, \' \ (p \ \\<^sub>R)\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\\ have "(p \ \A\<^sub>Q, \' \ (p \ \\<^sub>R)\) \\<^sub>F (p \ \A\<^sub>P, (\ \ \\<^sub>P) \ (p \ \\<^sub>R)\)" by(rule FrameStatImpClosed) with \A\<^sub>R \* A\<^sub>P\ \(p \ A\<^sub>R) \* A\<^sub>P\ \A\<^sub>R \* \'\ \(p \ A\<^sub>R) \* \'\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ \A\<^sub>R \* A\<^sub>Q\ \(p \ A\<^sub>R) \* A\<^sub>Q\ \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ S \distinctPerm p\ have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" by(simp add: eqvts) moreover from \xvec \* M\ have "(p \ xvec) \* (p \ M)" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with S \A\<^sub>R \* xvec\ \(p \ A\<^sub>R) \* xvec\ have "xvec \* (p \ M)" by simp ultimately obtain K' where "\ \ \\<^sub>P \ R \K'\N\ \ R'" and "\ \ \\<^sub>P \ \\<^sub>R \ (p \ M) \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" using cAlpha by(fastforce simp del: freshChainSimps) from \\ \ \\<^sub>P \ R \K'\N\ \ R'\ S \A\<^sub>R \* R\ \(p \ A\<^sub>R) \* R\ have "(p \ (\ \ \\<^sub>P)) \ R \(p \ K')\N\ \ R'" by(rule_tac inputPermFrameSubject) auto with S \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ have "\ \ \\<^sub>P \ R \(p \ K')\N\ \ R'" by(simp add: eqvts) moreover from \\ \ \\<^sub>P \ \\<^sub>R \ (p \ M) \ K'\ have "(p \ (\ \ \\<^sub>P \ \\<^sub>R))\ (p \ p \ M) \ (p \ K')" by(rule chanEqClosed) with S \A\<^sub>R \* \\ \(p \ A\<^sub>R) \* \\ \A\<^sub>R \* \\<^sub>P\ \(p \ A\<^sub>R) \* \\<^sub>P\ \distinctPerm p\ have "\ \ \\<^sub>P \ (p \ \\<^sub>R) \ M \ (p \ K')" by(simp add: eqvts) moreover from \zvec \* K'\ have "(p \ zvec) \* (p \ K')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) with \A\<^sub>R \* zvec\ \(p \ A\<^sub>R) \* zvec\ S have "zvec \* (p \ K')" by simp moreover from \A\<^sub>R \* K'\ have "(p \ A\<^sub>R) \* (p \ K')" by(simp add: pt_fresh_star_bij[OF pt_name_inst, OF at_name_inst]) ultimately show ?case by blast next case(cInput \' M' K xvec N Tvec R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec yvec M) from \A\<^sub>P \* (M'\\*xvec N\.R)\ \A\<^sub>Q \* (M'\\*xvec N\.R)\ \zvec \* (M'\\*xvec N\.R)\ have "A\<^sub>P \* M'" and "A\<^sub>Q \* M'" and "zvec \* M'" by simp+ from \\' \ M' \ K\ have "\' \ \ \ M' \ K" by(blast intro: statEqEnt Identity AssertionStatEqSym) hence "\' \ \ \ M' \ M'" by(blast intro: chanEqSym chanEqTrans) with \A\<^sub>Q \* M'\ have "(\A\<^sub>Q, \' \ \\) \\<^sub>F M' \ M'" by(force intro: frameImpI) with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F M' \ M'" by(simp add: FrameStatImp_def) with \A\<^sub>P \* M'\ have "(\ \ \\<^sub>P) \ \ \ M' \ M'" by(force dest: frameImpE) hence "\ \ \\<^sub>P \ M' \ M'" by(blast intro: statEqEnt Identity) hence "\ \ \\<^sub>P \ M'\\*xvec N\.R \M'\(N[xvec::=Tvec])\ \ R[xvec::=Tvec]" using \distinct xvec\ \set xvec \ supp N\ \length xvec = length Tvec\ by(rule Input) moreover from \\' \ \ \ M \ K\ \\' \ \ \ M' \ K\ have "\' \ \ \ M \ M'" by(metis chanEqSym chanEqTrans) with \A\<^sub>Q \* M\ \A\<^sub>Q \* M'\ have "(\A\<^sub>Q, \' \ \\) \\<^sub>F M \ M'" by(force intro: frameImpI) with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\ \ have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F M \ M'" by(simp add: FrameStatImp_def) with \A\<^sub>P \* M\ \A\<^sub>P \* M'\ have "(\ \ \\<^sub>P) \ \ \ M \ M'" by(force dest: frameImpE) hence "\ \ \\<^sub>P \ \ \ M \ M'" by(metis statEqEnt Associativity) ultimately show ?case using \zvec \* M'\ by force next case(cCase \' R M' N R' \ Cs A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) from \guarded R\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "\\<^sub>R \ \" by(metis guardedStatEq) with \\' \ \ \ M \ M'\ have "\' \ \\<^sub>R \ M \ M'" by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition) moreover have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" proof - from \\\<^sub>R \ \\ have "\A\<^sub>Q, \' \ \\<^sub>R\ \\<^sub>F \A\<^sub>Q, \' \ \\" by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) moreover note \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ moreover from \\\<^sub>R \ \\ have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\" by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) ultimately show ?thesis by(rule FrameStatEqImpCompose) qed moreover from \\ \ \ \ P \M\\*xvec\\N\ \ P'\ \\\<^sub>R \ \\ have "\ \ \\<^sub>R \ P \M\\*xvec\\N\ \ P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition) moreover from \zvec \* (Cases Cs)\ \A\<^sub>P \* (Cases Cs)\ \A\<^sub>Q \* (Cases Cs)\ \(\, R) mem Cs\ have "A\<^sub>P \* R" and "A\<^sub>Q \* R" and "zvec \* R" and "A\<^sub>P \* \" and "A\<^sub>Q \* \" by(auto dest: memFreshChain) ultimately have "\K'. \ \ \\<^sub>P \ R \K'\N\ \ R'\ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using cCase by(rule_tac cCase) (assumption |simp)+ then obtain K' where RTrans: "\ \ \\<^sub>P \ R \K'\N\ \ R'" and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" by metis note RTrans \(\, R) mem Cs\ moreover from \\' \ \\ have "\' \ \ \ \" by(blast intro: statEqEnt Identity AssertionStatEqSym) with \A\<^sub>Q \* \\ have "(\A\<^sub>Q, \' \ \\) \\<^sub>F \" by(force intro: frameImpI) with \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ have "(\A\<^sub>P, (\ \ \\<^sub>P) \ \\) \\<^sub>F \" by(simp add: FrameStatImp_def) with \A\<^sub>P \* \\ have "(\ \ \\<^sub>P) \ \ \ \" by(force dest: frameImpE) hence "\ \ \\<^sub>P \ \" by(blast intro: statEqEnt Identity) ultimately have "\ \ \\<^sub>P \ Cases Cs \K'\N\ \ R'" using \guarded R\ by(rule Case) moreover from MeqK' \\\<^sub>R \ \\ have "\ \ \\<^sub>P \ \ \ M \ K'" by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans) ultimately show ?case using \zvec \* K'\ by force next case(cPar1 \' \\<^sub>R\<^sub>2 R\<^sub>1 M' N R\<^sub>1' A\<^sub>R\<^sub>2 R\<^sub>2 A\<^sub>R\<^sub>1 \\<^sub>R\<^sub>1 \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) have FrR2: "extractFrame R\<^sub>2 = \A\<^sub>R\<^sub>2, \\<^sub>R\<^sub>2\" by fact from \\' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ M'\ have "(\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1 \ M \ M'" by(metis statEqEnt Associativity Composition Commutativity) moreover have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P) \ \\<^sub>R\<^sub>1\" proof - have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1\ \\<^sub>F \A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\" by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) moreover note \\A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\\ moreover have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P) \ \\<^sub>R\<^sub>1\" by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) ultimately show ?thesis by(rule FrameStatEqImpCompose) qed moreover from \\ \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ P \M\\*xvec\\N\ \ P'\ have "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>R\<^sub>1 \ P \M\\*xvec\\N\ \ P'" by(metis statEqTransition Associativity Composition Commutativity) moreover from \A\<^sub>R\<^sub>1 \* A\<^sub>P\ \A\<^sub>R\<^sub>2 \* A\<^sub>P\ \A\<^sub>P \* (R\<^sub>1 \ R\<^sub>2)\ \extractFrame R\<^sub>1 = \A\<^sub>R\<^sub>1, \\<^sub>R\<^sub>1\\ FrR2 have "A\<^sub>P \* \\<^sub>R\<^sub>1" and "A\<^sub>P \* \\<^sub>R\<^sub>2" by(force dest: extractFrameFreshChain)+ moreover note \distinct xvec\ ultimately have "\K'. (\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ R\<^sub>1 \K'\N\ \ R\<^sub>1' \ (\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ M \ K' \ (A\<^sub>R\<^sub>2@zvec) \* K' \ A\<^sub>R\<^sub>1 \* K'" using cPar1 by(rule_tac cPar1(6)[where bf=xvec]) (assumption | simp | fastforce)+ then obtain K' where RTrans: "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ R\<^sub>1 \K'\N\ \ R\<^sub>1'" and MeqK': "(\ \ \\<^sub>R\<^sub>2) \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ M \ K'" and "A\<^sub>R\<^sub>2 \* K'" and "zvec \* K'" and "A\<^sub>R\<^sub>1 \* K'" by force from RTrans have "(\ \ \\<^sub>P) \ \\<^sub>R\<^sub>2 \ R\<^sub>1 \K'\N\ \ R\<^sub>1'" by(metis statEqTransition Associativity Composition Commutativity) hence "\ \ \\<^sub>P \ (R\<^sub>1 \ R\<^sub>2) \K'\N\ \ (R\<^sub>1' \ R\<^sub>2)" using FrR2 \A\<^sub>R\<^sub>2 \* \\ \A\<^sub>R\<^sub>2 \* \\<^sub>P\ \A\<^sub>R\<^sub>2 \* K'\ \A\<^sub>R\<^sub>2 \* R\<^sub>1\ \A\<^sub>R\<^sub>2 \* N\ by(force intro: Par1) moreover from MeqK' have "\ \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ K'" by(metis statEqEnt Associativity Composition Commutativity) ultimately show ?case using \zvec \* K'\ \A\<^sub>R\<^sub>1 \* K'\ \A\<^sub>R\<^sub>2 \* K'\ by force next case(cPar2 \' \\<^sub>R\<^sub>1 R\<^sub>2 M' N R\<^sub>2' A\<^sub>R\<^sub>1 R\<^sub>1 A\<^sub>R\<^sub>2 \\<^sub>R\<^sub>2 \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) have FrR1: "extractFrame R\<^sub>1 = \A\<^sub>R\<^sub>1, \\<^sub>R\<^sub>1\" by fact from \\' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ M'\ have "(\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2 \ M \ M'" by(metis statEqEnt Associativity Composition Commutativity) moreover have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P) \ \\<^sub>R\<^sub>2\" proof - have "\A\<^sub>Q, (\' \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\" by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) moreover note \\A\<^sub>Q, \' \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\\ moreover have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2\ \\<^sub>F \A\<^sub>P, ((\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P) \ \\<^sub>R\<^sub>2\" by(metis Associativity Composition Commutativity AssertionStatEqTrans AssertionStatEqSym frameNilStatEq frameResChainPres) ultimately show ?thesis by(rule FrameStatEqImpCompose) qed moreover from \\ \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ P \M\\*xvec\\N\ \ P'\ have "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>R\<^sub>2 \ P \M\\*xvec\\N\ \ P'" by(metis statEqTransition Associativity Composition Commutativity) moreover from \A\<^sub>R\<^sub>1 \* A\<^sub>P\ \A\<^sub>R\<^sub>2 \* A\<^sub>P\ \A\<^sub>P \* (R\<^sub>1 \ R\<^sub>2)\ FrR1 \extractFrame R\<^sub>2 = \A\<^sub>R\<^sub>2, \\<^sub>R\<^sub>2\\ have "A\<^sub>P \* \\<^sub>R\<^sub>1" and "A\<^sub>P \* \\<^sub>R\<^sub>2" by(force dest: extractFrameFreshChain)+ ultimately have "\K'. (\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ R\<^sub>2 \K'\N\ \ R\<^sub>2' \ (\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ \\<^sub>R\<^sub>2 \ M \ K' \ (A\<^sub>R\<^sub>1@zvec) \* K' \ A\<^sub>R\<^sub>2 \* K'" using \distinct xvec\ cPar2 by(rule_tac cPar2(6)) (assumption | simp | fastforce)+ then obtain K' where RTrans: "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ R\<^sub>2 \K'\N\ \ R\<^sub>2'" and MeqK': "(\ \ \\<^sub>R\<^sub>1) \ \\<^sub>P \ \\<^sub>R\<^sub>2 \ M \ K'" and "A\<^sub>R\<^sub>1 \* K'" and "zvec \* K'" and "A\<^sub>R\<^sub>2 \* K'" by force from RTrans have "(\ \ \\<^sub>P) \ \\<^sub>R\<^sub>1 \ R\<^sub>2 \K'\N\ \ R\<^sub>2'" by(metis statEqTransition Associativity Composition Commutativity) hence "\ \ \\<^sub>P \ (R\<^sub>1 \ R\<^sub>2) \K'\N\ \ (R\<^sub>1 \ R\<^sub>2')" using FrR1 \A\<^sub>R\<^sub>1 \* \\ \A\<^sub>R\<^sub>1 \* \\<^sub>P\ \A\<^sub>R\<^sub>1 \* K'\ \A\<^sub>R\<^sub>1 \* R\<^sub>2\ \A\<^sub>R\<^sub>1 \* N\ by(force intro: Par2) moreover from MeqK' have "\ \ \\<^sub>P \ \\<^sub>R\<^sub>1 \ \\<^sub>R\<^sub>2 \ M \ K'" by(metis statEqEnt Associativity Composition Commutativity) ultimately show ?case using \zvec \* K'\ \A\<^sub>R\<^sub>1 \* K'\ \A\<^sub>R\<^sub>2 \* K'\ by force next case(cScope \' R M' N R' x A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) then have "\K'. \ \ \\<^sub>P \ R \K'\N\ \ R' \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" by(rule_tac cScope(4)) (assumption | simp del: freshChainSimps)+ then obtain K' where RTrans: "\ \ \\<^sub>P \ R \K'\N\ \ R'" and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" by metis from \A\<^sub>R \* A\<^sub>P\ \A\<^sub>P \* (\\x\R)\ \x \ A\<^sub>P\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "A\<^sub>P \* \\<^sub>R" by(force dest: extractFrameFreshChain)+ from \\ \ \\<^sub>R \ P \M\\*xvec\\N\ \ P'\ \extractFrame P = \A\<^sub>P, \\<^sub>P\\ \distinct A\<^sub>P\ \x \ P\ \zvec \* P\ \A\<^sub>P \* \\<^sub>R\ \x \ A\<^sub>P\ \A\<^sub>P \* M\ \A\<^sub>P \* P\ \A\<^sub>P \* zvec\ \A\<^sub>P \* \\ \A\<^sub>P \* zvec\ \A\<^sub>R \* P\ \A\<^sub>R \* A\<^sub>P\ \xvec \* M\ \distinct xvec\ obtain K'' where MeqK'': "(\ \ \\<^sub>R) \ \\<^sub>P \ M \ K''" and "x \ K''" and "A\<^sub>R \* K''" and "zvec \* K''" by(rule_tac B="(x#A\<^sub>R@zvec)" in obtainPrefix) (assumption | simp | force | metis freshChainSym)+ from MeqK'' MeqK' have KeqK'': "(\ \ \\<^sub>P) \ \\<^sub>R \ K' \ K''" by(metis statEqEnt Associativity Composition Commutativity chanEqSym chanEqTrans) with RTrans \extractFrame R = \A\<^sub>R, \\<^sub>R\\ \distinct A\<^sub>R\ \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>P\ \A\<^sub>R \* K'\ \A\<^sub>R \* K''\ \A\<^sub>R \* R\ have "\ \ \\<^sub>P \ R \K''\N\ \ R'" by(rule_tac inputRenameSubject) (assumption | force)+ hence "\ \ \\<^sub>P \ \\x\R \K''\N\ \ \\x\R'" using \x \ \\ \x \ \\<^sub>P\ \x \ K''\ \x \ N\ by(rule_tac Scope) (assumption | force)+ moreover from MeqK'' have "\ \ \\<^sub>P \ \\<^sub>R \ M \ K''" by(metis statEqEnt Associativity Composition Commutativity) ultimately show ?case using \zvec \* K''\ \x \ K''\ \A\<^sub>R \* K''\ by force next case(cBang \' R M' N R' A\<^sub>R \\<^sub>R \ P A\<^sub>P \\<^sub>P A\<^sub>Q zvec xvec M) from \guarded R\ \extractFrame R = \A\<^sub>R, \\<^sub>R\\ have "\\<^sub>R \ \" by(metis guardedStatEq) with \\' \ \ \ M \ M'\ have "\' \ \\<^sub>R \ \ \ M \ M'" by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition) moreover have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" proof - from \\\<^sub>R \ \\ have "\A\<^sub>Q, \' \ \\<^sub>R \ \\ \\<^sub>F \A\<^sub>Q, \' \ \\" by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) moreover note \\A\<^sub>Q, \' \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\\ moreover from \\\<^sub>R \ \\ have "\A\<^sub>P, (\ \ \\<^sub>P) \ \\ \\<^sub>F \A\<^sub>P, (\ \ \\<^sub>P) \ \\<^sub>R \ \\" by(metis Identity Commutativity AssertionStatEqSym Composition frameResChainPres frameNilStatEq AssertionStatEqTrans) ultimately show ?thesis by(rule FrameStatEqImpCompose) qed moreover from \\ \ \ \ P \M\\*xvec\\N\ \ P'\ \\\<^sub>R \ \\ have "\ \ \\<^sub>R \ \ \ P \M\\*xvec\\N\ \ P'" by(metis statEqTransition Identity Commutativity AssertionStatEqSym Composition) ultimately have "\K'. \ \ \\<^sub>P \ R \ !R \K'\N\ \ R'\ \ \ \\<^sub>P \ \\<^sub>R \ \ \ M \ K' \ zvec \* K' \ A\<^sub>R \* K'" using cBang by(rule_tac cBang(5)) (assumption | simp del: freshChainSimps)+ then obtain K' where RTrans: "\ \ \\<^sub>P \ R \ !R \K'\N\ \ R'" and MeqK': "\ \ \\<^sub>P \ \\<^sub>R \ \ \ M \ K'" and "zvec \* K'" and "A\<^sub>R \* K'" by metis from RTrans \guarded R\ have "\ \ \\<^sub>P \ !R \K'\N\ \ R'" by(rule Bang) moreover from MeqK' \\\<^sub>R \ \\ have "\ \ \\<^sub>P \ \ \ M \ K'" by(metis Identity Commutativity statEqEnt AssertionStatEqSym Composition AssertionStatEqTrans) ultimately show ?case using \zvec \* K'\ by force qed } note Goal = this have "\ \ \\<^sub>Q \ \ \ \\<^sub>Q" by(simp add: AssertionStatEqRefl) moreover from \A\<^sub>R \* \\ \A\<^sub>R \* \\<^sub>Q\ have "A\<^sub>R \* (\ \ \\<^sub>Q)" by force ultimately have "\K'. \ \ \\<^sub>P \ R \K'\N\ \ R' \ \ \ \\<^sub>P \ \\<^sub>R \ M \ K' \ ([]::name list) \* K' \ A\<^sub>R \* K'" by(rule_tac Goal) (assumption | force)+ with Assumptions show ?thesis by blast qed end end diff --git a/thys/Relation_Algebra/Relation_Algebra.thy b/thys/Relation_Algebra/Relation_Algebra.thy --- a/thys/Relation_Algebra/Relation_Algebra.thy +++ b/thys/Relation_Algebra/Relation_Algebra.thy @@ -1,300 +1,300 @@ (* Title: Relation Algebra Author: Alasdair Armstrong, Simon Foster, Georg Struth, Tjark Weber Maintainer: Georg Struth Tjark Weber *) section \Relation Algebra\ theory Relation_Algebra imports More_Boolean_Algebra Kleene_Algebra.Kleene_Algebra begin text \We follow Tarski's original article and Maddux's book, in particular we use their notation. In contrast to Schmidt and Str\"ohlein we do not assume that the Boolean algebra is complete and we do not consider the Tarski rule in this development. A main reason for using complete Boolean algebras seems to be that the Knaster-Tarski fixpoint theorem becomes available for defining notions of iteration. In fact, several chapters of Schmidt and Str\"ohlein's book deal with iteration. We capture iteration in an alternative way by linking relation algebras with Kleene algebras (cf.~\emph{relation-algebra-rtc}).\ class relation_algebra = boolean_algebra + fixes composition :: "'a \ 'a \ 'a" (infixl ";" 75) and converse :: "'a \ 'a" ("(_\<^sup>\)" [1000] 999) - and unit :: "'a" ("1'") + and unit :: "'a" ("1''") assumes comp_assoc: "(x ; y) ; z = x ; (y ; z)" and comp_unitr [simp]: "x ; 1' = x" and comp_distr: "(x + y) ; z = x ; z + y ; z" and conv_invol [simp]: "(x\<^sup>\)\<^sup>\ = x" and conv_add [simp]: "(x + y)\<^sup>\ = x\<^sup>\ + y\<^sup>\" and conv_contrav [simp]: "(x ; y)\<^sup>\ = y\<^sup>\ ; x\<^sup>\" and comp_res: "x\<^sup>\ ; -(x ; y) \ -y" text \We first show that every relation algebra is a dioid. We do not yet treat the zero (the minimal element of the boolean reduct) since the proof of the annihilation laws is rather tricky to automate. Following Maddux we derive them from properties of Boolean algebras with operators.\ sublocale relation_algebra \ dioid_one "(+)" "(;)" "(\) " "(<)" "1'" proof fix x y z :: 'a show "x ; y ; z = x ; (y ; z)" by (fact comp_assoc) show "x + y + z = x + (y + z)" by (metis sup.commute sup.left_commute) show "x + y = y + x" by (fact sup.commute) show "(x + y) ; z = x ; z + y ; z" by (fact comp_distr) show "x + x = x" by (fact sup.idem) show "x ; (y + z) = x ; y + x ; z" by (metis conv_invol conv_add conv_contrav comp_distr) show "x ; 1' = x" by (fact comp_unitr) show "1' ; x = x" by (metis comp_unitr conv_contrav conv_invol) show "x \ y \ x + y = y" by (metis sup.commute sup_absorb1 sup_ge1) show "x < y \ x \ y \ x \ y" by (fact less_le) qed context relation_algebra begin text \First we prove some basic facts about joins and meets.\ lemma meet_interchange: "(w \ x) ; (y \ z) \ w ; y \ x ; z" by (metis inf_le1 inf_le2 le_infI mult_isol_var) lemma join_interchange: "w ; x + y ; z \ (w + y) ; (x + z)" using local.mult_isol_var local.sup.bounded_iff local.sup.cobounded2 local.sup_ge1 by presburger text \We now prove some simple facts about conversion.\ lemma conv_iso: "x \ y \ x\<^sup>\ \ y\<^sup>\" by (metis conv_add conv_invol le_iff_sup) lemma conv_zero [simp]: "0\<^sup>\ = 0" by (metis conv_add conv_invol sup_bot_right sup_eq_bot_iff) lemma conv_one [simp]: "1\<^sup>\ = 1" by (metis conv_add conv_invol sup_top_left sup_top_right) lemma conv_compl_aux: "(-x)\<^sup>\ = (-x)\<^sup>\ + (-x\<^sup>\)" by (metis aux9 conv_add conv_one double_compl galois_aux4 inf.commute less_eq_def sup.commute sup_top_left) lemma conv_compl: "(-x)\<^sup>\ = -(x\<^sup>\)" by (metis add_commute conv_add conv_compl_aux conv_invol) lemma comp_res_aux [simp]: "x\<^sup>\ ; -(x ; y) \ y = 0" by (metis comp_res galois_aux) lemma conv_e [simp]: "1'\<^sup>\ = 1'" by (metis comp_unitr conv_contrav conv_invol) lemma conv_times [simp]: "(x \ y)\<^sup>\ = x\<^sup>\ \ y\<^sup>\" by (metis compl_inf double_compl conv_add conv_compl) text \The next lemmas show that conversion is self-conjugate in the sense of Boolean algebra with operators.\ lemma conv_self_conjugate: "x\<^sup>\ \ y = 0 \ x \ y\<^sup>\ = 0" by (metis conv_invol conv_times conv_zero) lemma conv_self_conjugate_var: "is_conjugation converse converse" by (metis conv_self_conjugate is_conjugation_def) text \The following lemmas link the relative product and meet.\ lemma one_idem_mult [simp]: "1 ; 1 = 1" by (metis compl_eq_compl_iff galois_aux2 inf.commute inf_top_right mult_1_left mult_isor top_greatest) lemma mult_subdistl: "x ; (y \ z) \ x ; y" by (metis inf_le1 mult_isol) lemma mult_subdistr: "(x \ y) ; z \ x ; z" by (metis inf_le1 mult_isor) lemma mult_subdistr_var: "(x \ y) ; z \ x ; z \ y ; z" by (metis inf.commute le_inf_iff mult_subdistr) text \The following lemmas deal with variants of the Peirce law, the Schr\"oder laws and the Dedekind law. Some of them are obtained from Boolean algebras with operators by instantiation, using conjugation properties. However, Isabelle does not always pick up this relationship.\ lemma peirce_1: "x ; y \ z\<^sup>\ = 0 \ y ; z \ x\<^sup>\ = 0" by (metis compl_le_swap1 conv_contrav conv_self_conjugate galois_aux comp_res conv_invol galois_aux mult_isol order_trans) lemma peirce: "x ; y \ z\<^sup>\ = 0 \ y ; z \ x\<^sup>\ = 0" by (metis peirce_1) lemma schroeder_1: "x ; y \ z = 0 \ y \ x\<^sup>\ ; z = 0" by (metis conv_invol peirce conv_contrav conv_invol conv_self_conjugate inf.commute) lemma schroeder_2: "y ; x \ z = 0 \ y \ z ; x\<^sup>\ = 0" by (metis conv_invol peirce schroeder_1) text \The following two conjugation properties between multiplication with elements and their converses are used for deriving modular laws of relation algebra from those of Boolean algebras with operators.\ lemma schroeder_1_var: "is_conjugation (composition x) (composition (x\<^sup>\))" by (metis schroeder_1 is_conjugation_def) lemma schroeder_2_var: "is_conjugation (\x. x ; y) (\x. x ; y\<^sup>\)" by (unfold is_conjugation_def, metis schroeder_2) text \The following Galois connections define residuals. They link relation algebras with action algebras. This could be further explored and formalised. \ lemma conv_galois_1: "x ; y \ z \ y \ -(x\<^sup>\ ; -z)" by (metis galois_aux galois_aux2 schroeder_1) lemma conv_galois_2: "y ; x \ z \ y \ -(-z ; x\<^sup>\)" by (metis galois_aux galois_aux2 schroeder_2) text \Variants of the modular law for relation algebras can now be instantiated from Boolean algebras with operators.\ lemma modular_1_aux': "x ; (y \ -(x\<^sup>\ ; z)) \ z = 0" by (metis schroeder_1_var modular_1_aux) lemma modular_2_aux': "(y \ -(z ; x\<^sup>\)) ; x \ z = 0" by (metis modular_1_aux schroeder_2_var) lemma modular_1': "x ; y \ z = x ; (y \ x\<^sup>\ ; z) \ z" by (metis schroeder_1_var modular_1) lemma modular_2': "y ; x \ z = (y \ z ; x\<^sup>\) ; x \ z" proof - have "y ; x \ z = (y \ z ; x\<^sup>\) ; x \ z + (y \ -(z ; x\<^sup>\)) ; x \ z" by (metis aux4 distrib_right inf.commute inf_sup_distrib1) thus ?thesis by (metis sup_bot_right modular_2_aux') qed lemma modular_1_var: "x ; y \ z \ x ; (y \ x\<^sup>\ ; z)" by (metis inf.commute inf_le2 modular_1') lemma modular_2_var: "x ; y \ z \ (x \ z ; y\<^sup>\) ; y" by (metis inf.commute inf_le2 modular_2') lemma modular_var_2: "x ; y \ x ; (y \ x\<^sup>\ ; 1)" by (metis inf_top_right modular_1_var) lemma modular_var_3: "x ; y \ (x \ 1 ; y\<^sup>\) ; y" by (metis inf_top_right modular_2_var) text \The modular laws are used to prove the Dedekind rule.\ lemma dedekind: "x ; y \ z \ (x \ z ; y\<^sup>\) ; (y \ x\<^sup>\ ; z)" proof - have "x ; y \ z \ (x \ z ; y\<^sup>\) ; (y \ ((x \ z ; y\<^sup>\)\<^sup>\ ; z))" by (metis modular_2' modular_1_var) also have "\ \ (x \ z ; y\<^sup>\) ; (y \ x\<^sup>\ ; z)" by (metis conv_iso inf_le1 inf_mono mult_isol mult_isor order_refl) thus ?thesis by (metis calculation order_trans) qed lemma dedekind_var_1: "x ; y \ (x \ 1 ; y\<^sup>\) ; (y \ x\<^sup>\ ; 1)" by (metis dedekind inf.commute inf_top_left) end (* relation_algebra *) text \The Schr\"oder laws allow us, finally, to prove the annihilation laws for zero. We formalise this by proving that relation algebras form dioids with zero.\ sublocale relation_algebra < dioid_one_zero "(+)" "(;)" "1'" 0 "(\)" "(<)" proof fix x :: 'a show "0 + x = x" by (fact sup_bot_left) show "0 ; x = 0" by (metis f_strict schroeder_2_var) show "x ; 0 = 0" by (metis f_strict schroeder_1_var) qed context relation_algebra begin text \Next we prove miscellaneous properties which we found in the books of Maddux and Schmidt and Str\"ohlein. Most of them do not carry any meaningful names.\ lemma ra_1: "(x \ y ; 1) ; z = x ; z \ y ; 1" proof (rule antisym) show "x ; z \ y ; 1 \ (x \ y ; 1) ; z" by (metis modular_2_var comp_assoc order_trans eq_refl inf_mono inf_top_left mult_isor mult_subdistl) show "(x \ y ; 1) ; z \ x ; z \ y ; 1" by (metis inf.commute inf_greatest inf_top_left mult.assoc mult_subdistl mult_subdistr order_trans) qed lemma ra_2: "x ; (z \ y ; 1) = (x \ (y ; 1)\<^sup>\) ; z" proof (rule antisym) have "(x \ 1 ; (z \ y ; 1)\<^sup>\) ; z \ (x \ (y ; 1)\<^sup>\) ; z" by (metis conv_contrav conv_invol conv_one conv_times inf.idem inf.left_commute le_iff_inf meet_assoc mult_isor ra_1) thus "x ; (z \ y ; 1) \ (x \ (y ; 1)\<^sup>\) ; z" by (metis modular_var_3 mult_subdistl order_trans) next have "x ; (z \ ((x \ (y ; 1)\<^sup>\)\<^sup>\ ; 1)) \ x ; (z \ y ; 1)" by (metis conv_invol conv_times eq_refl inf_le2 inf_mono mult.assoc mult_isol mult_isor one_idem_mult) thus "x ; (z \ y ; 1) \ (x \ (y ; 1)\<^sup>\) ; z" by (metis modular_var_2 mult_subdistr order_trans) qed lemma one_conv: "1' \ x ; 1 = 1' \ x ; x\<^sup>\" by (metis inf.commute inf_top_left modular_1' mult.right_neutral) lemma maddux_12: "-(y ; x) ; x\<^sup>\ \ -y" by (metis galois_aux inf.commute inf_compl_bot schroeder_2) lemma maddux_141: "x ; y \ z = 0 \ x\<^sup>\ ; z \ y = 0" by (metis inf.commute schroeder_1) lemma maddux_142: "x\<^sup>\ ; z \ y = 0 \ z ; y\<^sup>\ \ x = 0" by (metis inf.commute schroeder_1 schroeder_2) lemmas maddux_16 = modular_1_var lemmas maddux_17 = modular_2_var lemma maddux_20: "x \ x ; 1" by (metis inf_top_left mult.right_neutral mult_subdistl) lemma maddux_21: "x \ 1 ; x" by (metis mult_isor mult_onel top_greatest) lemma maddux_23: "x ; y \ -(x ; z) = x ; (y \ -z) \ -(x ; z)" apply (rule antisym) apply (metis local.aux6 local.aux6_var local.aux9 local.compl_inf_bot local.compl_sup_top local.compl_unique local.distrib_left local.galois_2 local.sup_ge2) using local.meet_iso local.mult_subdistl by blast lemma maddux_24: "-(x ; y) + x ; z = -(x ; (y \ -z)) + x ; z" by (metis de_morgan_3 double_compl maddux_23) lemma one_compl: "-(x ; 1) ; 1 = -(x ; 1)" by (metis antisym conv_one maddux_12 mult.assoc one_idem_mult maddux_20) lemma ss_p18: "x ; 1 = 0 \ x = 0" by (metis annil le_bot maddux_20) end (* relation_algebra *) text \This finishes our development of the basic laws of relation algebras. The next sections are devoted to special elements such as vectors, test or subidentities, and, in particular, functions.\ end diff --git a/thys/Residuated_Lattices/Involutive_Residuated.thy b/thys/Residuated_Lattices/Involutive_Residuated.thy --- a/thys/Residuated_Lattices/Involutive_Residuated.thy +++ b/thys/Residuated_Lattices/Involutive_Residuated.thy @@ -1,138 +1,138 @@ (* Title: Involutive Residuated Structures Author: Victor Gomes Maintainer: Victor Gomes *) section \Involutive Residuated Structures\ theory Involutive_Residuated imports Residuated_Lattices begin class uminus' = - fixes uminus' :: "'a \ 'a" ("-' _" [81] 80) + fixes uminus' :: "'a \ 'a" ("-'' _" [81] 80) text \ Involutive posets is a structure where the double negation property holds for the negation operations, and a Galois connection for negations exists. \ class involutive_order = order + uminus + uminus' + assumes gn: "x \ -'y \ y \ -x" and dn1[simp]: "-'(-x) = x" and dn2[simp]: "-(-'x) = x" (* The involutive pair (-', -) is compatible with multiplication *) class involutive_pogroupoid = order + times + involutive_order + assumes ipg1: "x\y \ z \ (-z)\x \ -y" and ipg2: "x\y \ z \ y\(-'z) \ -'x" begin lemma neg_antitone: "x \ y \ -y \ -x" by (metis local.dn1 local.gn) lemma neg'_antitone: "x \ y \ -'y \ -'x" by (metis local.dn2 local.gn) subclass pogroupoid proof fix x y z assume assm: "x \ y" show "x \ z \ y \ z" by (metis assm local.ipg2 local.order_refl local.order_trans neg'_antitone) show "z \ x \ z \ y" by (metis assm local.dual_order.trans local.ipg1 local.order_refl neg_antitone) qed abbreviation inv_resl :: "'a \ 'a \ 'a" where "inv_resl y x \ -(x\(-'y))" abbreviation inv_resr :: "'a \ 'a \ 'a" where "inv_resr x y \ -'((-y)\x)" sublocale residuated_pogroupoid _ _ _ inv_resl inv_resr proof fix x y z show "(x \ - (y \ -' z)) = (x \ y \ z)" by (metis local.gn local.ipg2) show "(x \ y \ z) = (y \ -' (- z \ x))" by (metis local.gn local.ipg1) qed end class division_order = order + residual_l_op + residual_r_op + assumes div_galois: "x \ z \ y \ y \ x \ z" class involutive_division_order = division_order + involutive_order + assumes contraposition: "y \ -x = -'y \ x" context involutive_pogroupoid begin sublocale involutive_division_order _ _ inv_resl inv_resr proof fix x y z show "(x \ - (y \ -' z)) = (y \ -' (- z \ x))" by (metis local.gn local.ipg1 local.ipg2) show "-' (- (- x) \ y) = - (x \ -' (-' y))" by (metis local.dn1 local.dn2 local.eq_iff local.gn local.jipsen1l local.jipsen1r local.resl_galois local.resr_galois) qed lemma inv_resr_neg [simp]: "inv_resr (-x) (-y) = inv_resl x y" by (metis local.contraposition local.dn1) lemma inv_resl_neg' [simp]: "inv_resl (-'x) (-'y) = inv_resr x y" by (metis local.contraposition local.dn2) lemma neg'_mult_resl: "-'((-y)\(-x)) = inv_resl x (-'y)" by (metis inv_resr_neg local.dn2) lemma neg_mult_resr: "-((-'y)\(-'x)) = inv_resr (-x) y" by (metis neg'_mult_resl) lemma resr_de_morgan1: "-'(inv_resr (-y) (-x)) = -'(inv_resl y x)" by (metis local.dn1 neg_mult_resr) lemma resr_de_morgan2: "-(inv_resl (-'x) (-'y)) = -(inv_resr x y)" by (metis inv_resl_neg') end text \ We prove that an involutive division poset is equivalent to an involutive po-groupoid by a lemma to avoid cyclic definitions \ lemma (in involutive_division_order) inv_pogroupoid: "class.involutive_pogroupoid (\x y. -(y \ -'x)) uminus uminus' (\) (<)" proof fix x y z have "(- (y \ -' x) \ z) = (-z \ -y \ x)" by (metis local.contraposition local.dn1 local.dn2 local.gn local.div_galois) thus "(- (y \ -' x) \ z) = (- (x \ -' (- z)) \ - y)" by (metis local.contraposition local.div_galois local.dn1 local.dn2 local.gn) moreover have "(- (x \ -' (- z)) \ - y) = (- (-' z \ -' y) \ -' x)" apply (auto, metis local.contraposition local.div_galois local.dn1 local.dn2 local.gn) by (metis local.contraposition local.div_galois local.dn1 local.dn2 local.gn) ultimately show "(- (y \ -' x) \ z) = (- (-' z \ -' y) \ -' x)" by metis qed context involutive_pogroupoid begin definition negation_constant :: "'a \ bool" where "negation_constant a \ \x. -'x = inv_resr x a \ -x = inv_resl a x" definition division_unit :: "'a \ bool" where "division_unit a \ \x. x = inv_resr a x \ x = inv_resl x a" lemma neg_iff_div_unit: "(\a. negation_constant a) \ (\b. division_unit b)" unfolding negation_constant_def division_unit_def apply safe apply (rule_tac x="-a" in exI, auto) apply (metis local.dn1 local.dn2) apply (metis local.dn2) apply (rule_tac x="-b" in exI, auto) apply (metis local.contraposition) apply (metis local.dn2) done end end diff --git a/thys/Stone_Relation_Algebras/Relation_Algebras.thy b/thys/Stone_Relation_Algebras/Relation_Algebras.thy --- a/thys/Stone_Relation_Algebras/Relation_Algebras.thy +++ b/thys/Stone_Relation_Algebras/Relation_Algebras.thy @@ -1,1817 +1,1817 @@ (* Title: Relation Algebras Author: Walter Guttmann Maintainer: Walter Guttmann *) section \Relation Algebras\ text \ The main structures introduced in this theory are Stone relation algebras. They generalise Tarski's relation algebras \cite{Tarski1941} by weakening the Boolean algebra lattice structure to a Stone algebra. Our motivation is to generalise relation-algebraic methods from unweighted graphs to weighted graphs. Unlike unweighted graphs, weighted graphs do not form a Boolean algebra because there is no complement operation on the edge weights. However, edge weights form a Stone algebra, and matrices over edge weights (that is, weighted graphs) form a Stone relation algebra. The development in this theory is described in our papers \cite{Guttmann2016c,Guttmann2017b}. Our main application there is the verification of Prim's minimum spanning tree algorithm. Related work about fuzzy relations \cite{Goguen1967,Winter2001b}, Dedekind categories \cite{KawaharaFurusawa2001} and rough relations \cite{Comer1993,Pawlak1996} is also discussed in these papers. In particular, Stone relation algebras do not assume that the underlying lattice is complete or a Heyting algebra, and they do not assume that composition has residuals. We proceed in two steps. First, we study the positive fragment in the form of single-object bounded distributive allegories \cite{FreydScedrov1990}. Second, we extend these structures by a pseudocomplement operation with additional axioms to obtain Stone relation algebras. Tarski's relation algebras are then obtained by a simple extension that imposes a Boolean algebra. See, for example, \cite{BirdMoor1997,HirschHodkinson2002,Maddux1996,Maddux2006,Schmidt2011,SchmidtStroehlein1993} for further details about relations and relation algebras, and \cite{AndrekaMikulas2011,BredihinSchein1978} for algebras of relations with a smaller signature. \ theory Relation_Algebras imports Stone_Algebras.P_Algebras Semirings begin subsection \Single-Object Bounded Distributive Allegories\ text \ We start with developing bounded distributive allegories. The following definitions concern properties of relations that require converse in addition to lattice and semiring operations. \ class conv = fixes conv :: "'a \ 'a" ("_\<^sup>T" [100] 100) class bounded_distrib_allegory_signature = inf + sup + times + conv + bot + top + one + ord begin subclass times_one_ord . subclass times_top . abbreviation total_var :: "'a \ bool" where "total_var x \ 1 \ x * x\<^sup>T" abbreviation surjective_var :: "'a \ bool" where "surjective_var x \ 1 \ x\<^sup>T * x" abbreviation univalent :: "'a \ bool" where "univalent x \ x\<^sup>T * x \ 1" abbreviation injective :: "'a \ bool" where "injective x \ x * x\<^sup>T \ 1" abbreviation mapping :: "'a \ bool" where "mapping x \ univalent x \ total x" abbreviation bijective :: "'a \ bool" where "bijective x \ injective x \ surjective x" abbreviation point :: "'a \ bool" where "point x \ vector x \ bijective x" abbreviation arc :: "'a \ bool" where "arc x \ bijective (x * top) \ bijective (x\<^sup>T * top)" (* earlier name: atom *) abbreviation symmetric :: "'a \ bool" where "symmetric x \ x\<^sup>T = x" abbreviation antisymmetric :: "'a \ bool" where "antisymmetric x \ x \ x\<^sup>T \ 1" abbreviation asymmetric :: "'a \ bool" where "asymmetric x \ x \ x\<^sup>T = bot" abbreviation linear :: "'a \ bool" where "linear x \ x \ x\<^sup>T = top" abbreviation equivalence :: "'a \ bool" where "equivalence x \ preorder x \ symmetric x" abbreviation order :: "'a \ bool" where "order x \ preorder x \ antisymmetric x" abbreviation linear_order :: "'a \ bool" where "linear_order x \ order x \ linear x" end text \ We reuse the relation algebra axioms given in \cite{Maddux1996} except for one -- see lemma \conv_complement_sub\ below -- which we replace with the Dedekind rule (or modular law) \dedekind_1\. The Dedekind rule or variants of it are known from \cite{BirdMoor1997,FreydScedrov1990,KawaharaFurusawaMori1999,SchmidtStroehlein1993}. We add \comp_left_zero\, which follows in relation algebras but not in the present setting. The main change is that only a bounded distributive lattice is required, not a Boolean algebra. \ class bounded_distrib_allegory = bounded_distrib_lattice + times + one + conv + assumes comp_associative : "(x * y) * z = x * (y * z)" assumes comp_right_dist_sup : "(x \ y) * z = (x * z) \ (y * z)" assumes comp_left_zero [simp]: "bot * x = bot" assumes comp_left_one [simp]: "1 * x = x" assumes conv_involutive [simp]: "x\<^sup>T\<^sup>T = x" assumes conv_dist_sup : "(x \ y)\<^sup>T = x\<^sup>T \ y\<^sup>T" assumes conv_dist_comp : "(x * y)\<^sup>T = y\<^sup>T * x\<^sup>T" assumes dedekind_1 : "x * y \ z \ x * (y \ (x\<^sup>T * z))" begin subclass bounded_distrib_allegory_signature . text \ Many properties of relation algebras already follow in bounded distributive allegories. \ lemma conv_isotone: "x \ y \ x\<^sup>T \ y\<^sup>T" by (metis conv_dist_sup le_iff_sup) lemma conv_order: "x \ y \ x\<^sup>T \ y\<^sup>T" using conv_isotone by fastforce lemma conv_bot [simp]: "bot\<^sup>T = bot" using conv_order bot_unique by force lemma conv_top [simp]: "top\<^sup>T = top" by (metis conv_involutive conv_order eq_iff top_greatest) lemma conv_dist_inf: "(x \ y)\<^sup>T = x\<^sup>T \ y\<^sup>T" apply (rule antisym) using conv_order apply simp by (metis conv_order conv_involutive inf.boundedI inf.cobounded1 inf.cobounded2) lemma conv_inf_bot_iff: "bot = x\<^sup>T \ y \ bot = x \ y\<^sup>T" using conv_dist_inf conv_bot by fastforce lemma conv_one [simp]: "1\<^sup>T = 1" by (metis comp_left_one conv_dist_comp conv_involutive) lemma comp_left_dist_sup: "(x * y) \ (x * z) = x * (y \ z)" by (metis comp_right_dist_sup conv_involutive conv_dist_sup conv_dist_comp) lemma comp_right_isotone: "x \ y \ z * x \ z * y" by (simp add: comp_left_dist_sup sup.absorb_iff1) lemma comp_left_isotone: "x \ y \ x * z \ y * z" by (metis comp_right_dist_sup le_iff_sup) lemma comp_isotone: "x \ y \ w \ z \ x * w \ y * z" using comp_left_isotone comp_right_isotone order.trans by blast lemma comp_left_subdist_inf: "(x \ y) * z \ x * z \ y * z" by (simp add: comp_left_isotone) lemma comp_left_increasing_sup: "x * y \ (x \ z) * y" by (simp add: comp_left_isotone) lemma comp_right_subdist_inf: "x * (y \ z) \ x * y \ x * z" by (simp add: comp_right_isotone) lemma comp_right_increasing_sup: "x * y \ x * (y \ z)" by (simp add: comp_right_isotone) lemma comp_right_zero [simp]: "x * bot = bot" by (metis comp_left_zero conv_dist_comp conv_involutive) lemma comp_right_one [simp]: "x * 1 = x" by (metis comp_left_one conv_dist_comp conv_involutive) lemma comp_left_conjugate: "conjugate (\y . x * y) (\y . x\<^sup>T * y)" apply (unfold conjugate_def, intro allI) by (metis comp_right_zero bot.extremum_unique conv_involutive dedekind_1 inf.commute) lemma comp_right_conjugate: "conjugate (\y . y * x) (\y . y * x\<^sup>T)" apply (unfold conjugate_def, intro allI) by (metis comp_left_conjugate[unfolded conjugate_def] conv_inf_bot_iff conv_dist_comp conv_involutive) text \ We still obtain a semiring structure. \ subclass bounded_idempotent_semiring by (unfold_locales) (auto simp: comp_right_isotone comp_right_dist_sup comp_associative comp_left_dist_sup) sublocale inf: semiring_0 sup bot inf by (unfold_locales, auto simp: inf_sup_distrib2 inf_sup_distrib1 inf_assoc) lemma schroeder_1: "x * y \ z = bot \ x\<^sup>T * z \ y = bot" using abel_semigroup.commute comp_left_conjugate conjugate_def inf.abel_semigroup_axioms by fastforce lemma schroeder_2: "x * y \ z = bot \ z * y\<^sup>T \ x = bot" by (metis comp_right_conjugate conjugate_def inf_commute) lemma comp_additive: "additive (\y . x * y) \ additive (\y . x\<^sup>T * y) \ additive (\y . y * x) \ additive (\y . y * x\<^sup>T)" by (simp add: comp_left_dist_sup additive_def comp_right_dist_sup) lemma dedekind_2: "y * x \ z \ (y \ (z * x\<^sup>T)) * x" by (metis conv_dist_inf conv_order conv_dist_comp dedekind_1) text \ The intersection with a vector can still be exported from the first argument of a composition, and many other properties of vectors and covectors continue to hold. \ lemma vector_inf_comp: "vector x \ (x \ y) * z = x \ (y * z)" apply (rule antisym) apply (metis comp_left_subdist_inf comp_right_isotone inf.sup_left_isotone order_lesseq_imp top_greatest) by (metis comp_left_isotone comp_right_isotone dedekind_2 inf_commute inf_mono order_refl order_trans top_greatest) lemma vector_inf_closed: "vector x \ vector y \ vector (x \ y)" by (simp add: vector_inf_comp) lemma vector_inf_one_comp: "vector x \ (x \ 1) * y = x \ y" by (simp add: vector_inf_comp) lemma covector_inf_comp_1: assumes "vector x" shows "(y \ x\<^sup>T) * z = (y \ x\<^sup>T) * (x \ z)" proof - have "(y \ x\<^sup>T) * z \ (y \ x\<^sup>T) * (z \ ((y\<^sup>T \ x) * top))" by (metis inf_top_right dedekind_1 conv_dist_inf conv_involutive) also have "... \ (y \ x\<^sup>T) * (x \ z)" by (metis assms comp_left_isotone comp_right_isotone inf_le2 inf_mono order_refl inf_commute) finally show ?thesis by (simp add: comp_right_isotone antisym) qed lemma covector_inf_comp_2: assumes "vector x" shows "y * (x \ z) = (y \ x\<^sup>T) * (x \ z)" proof - have "y * (x \ z) \ (y \ (top * (x \ z)\<^sup>T)) * (x \ z)" by (metis dedekind_2 inf_top_right) also have "... \ (y \ x\<^sup>T) * (x \ z)" by (metis assms comp_left_isotone conv_dist_comp conv_order conv_top eq_refl inf_le1 inf_mono) finally show ?thesis using comp_left_subdist_inf antisym by auto qed lemma covector_inf_comp_3: "vector x \ (y \ x\<^sup>T) * z = y * (x \ z)" by (metis covector_inf_comp_1 covector_inf_comp_2) lemma covector_inf_closed: "covector x \ covector y \ covector (x \ y)" by (metis comp_right_subdist_inf inf.antisym top_left_mult_increasing) lemma vector_conv_covector: "vector v \ covector (v\<^sup>T)" by (metis conv_dist_comp conv_involutive conv_top) lemma covector_conv_vector: "covector v \ vector (v\<^sup>T)" by (simp add: vector_conv_covector) lemma covector_comp_inf: "covector z \ x * (y \ z) = x * y \ z" apply (rule antisym) apply (metis comp_isotone comp_right_subdist_inf inf.boundedE inf.boundedI inf.cobounded2 top.extremum) by (metis comp_left_isotone comp_right_isotone dedekind_1 inf_commute inf_mono order_refl order_trans top_greatest) lemma vector_restrict_comp_conv: "vector x \ x \ y \ x\<^sup>T * y" by (metis covector_inf_comp_3 eq_refl inf.sup_monoid.add_commute inf_top_right le_supE sup.orderE top_left_mult_increasing) lemma covector_restrict_comp_conv: "covector x \ y \ x \ y * x\<^sup>T" by (metis conv_dist_comp conv_dist_inf conv_order conv_top inf.sup_monoid.add_commute vector_restrict_comp_conv) lemma covector_comp_inf_1: "covector x \ (y \ x) * z = y * (x\<^sup>T \ z)" using covector_conv_vector covector_inf_comp_3 by fastforce text \ We still have two ways to represent surjectivity and totality. \ lemma surjective_var: "surjective x \ surjective_var x" proof assume "surjective x" thus "surjective_var x" by (metis dedekind_2 comp_left_one inf_absorb2 top_greatest) next assume "surjective_var x" hence "x\<^sup>T * (x * top) = top" by (metis comp_left_isotone comp_associative comp_left_one top_le) thus "surjective x" by (metis comp_right_isotone conv_top conv_dist_comp conv_involutive top_greatest top_le) qed lemma total_var: "total x \ total_var x" by (metis conv_top conv_dist_comp conv_involutive surjective_var) lemma surjective_conv_total: "surjective x \ total (x\<^sup>T)" by (metis conv_top conv_dist_comp conv_involutive) lemma total_conv_surjective: "total x \ surjective (x\<^sup>T)" by (simp add: surjective_conv_total) lemma injective_conv_univalent: "injective x \ univalent (x\<^sup>T)" by simp lemma univalent_conv_injective: "univalent x \ injective (x\<^sup>T)" by simp text \ We continue with studying further closure properties. \ lemma univalent_bot_closed: "univalent bot" by simp lemma univalent_one_closed: "univalent 1" by simp lemma univalent_inf_closed: "univalent x \ univalent (x \ y)" by (metis comp_left_subdist_inf comp_right_subdist_inf conv_dist_inf inf.cobounded1 order_lesseq_imp) lemma univalent_mult_closed: assumes "univalent x" and "univalent y" shows "univalent (x * y)" proof - have "(x * y)\<^sup>T * x \ y\<^sup>T" by (metis assms(1) comp_left_isotone comp_right_one conv_one conv_order comp_associative conv_dist_comp conv_involutive) thus ?thesis by (metis assms(2) comp_left_isotone comp_associative dual_order.trans) qed lemma injective_bot_closed: "injective bot" by simp lemma injective_one_closed: "injective 1" by simp lemma injective_inf_closed: "injective x \ injective (x \ y)" by (metis conv_dist_inf injective_conv_univalent univalent_inf_closed) lemma injective_mult_closed: "injective x \ injective y \ injective (x * y)" by (metis injective_conv_univalent conv_dist_comp univalent_mult_closed) lemma mapping_one_closed: "mapping 1" by simp lemma mapping_mult_closed: "mapping x \ mapping y \ mapping (x * y)" by (simp add: comp_associative univalent_mult_closed) lemma bijective_one_closed: "bijective 1" by simp lemma bijective_mult_closed: "bijective x \ bijective y \ bijective (x * y)" by (metis injective_mult_closed comp_associative) lemma bijective_conv_mapping: "bijective x \ mapping (x\<^sup>T)" by (simp add: surjective_conv_total) lemma mapping_conv_bijective: "mapping x \ bijective (x\<^sup>T)" by (simp add: total_conv_surjective) lemma reflexive_inf_closed: "reflexive x \ reflexive y \ reflexive (x \ y)" by simp lemma reflexive_conv_closed: "reflexive x \ reflexive (x\<^sup>T)" using conv_isotone by force lemma coreflexive_inf_closed: "coreflexive x \ coreflexive (x \ y)" by (simp add: le_infI1) lemma coreflexive_conv_closed: "coreflexive x \ coreflexive (x\<^sup>T)" using conv_order by force lemma coreflexive_symmetric: "coreflexive x \ symmetric x" by (metis comp_right_one comp_right_subdist_inf conv_dist_inf conv_dist_comp conv_involutive dedekind_1 inf.absorb1 inf_absorb2) lemma transitive_inf_closed: "transitive x \ transitive y \ transitive (x \ y)" by (meson comp_left_subdist_inf inf.cobounded1 inf.sup_mono inf_le2 mult_right_isotone order.trans) lemma transitive_conv_closed: "transitive x \ transitive (x\<^sup>T)" using conv_order conv_dist_comp by fastforce lemma dense_conv_closed: "dense_rel x \ dense_rel (x\<^sup>T)" using conv_order conv_dist_comp by fastforce lemma idempotent_conv_closed: "idempotent x \ idempotent (x\<^sup>T)" by (metis conv_dist_comp) lemma preorder_inf_closed: "preorder x \ preorder y \ preorder (x \ y)" using transitive_inf_closed by auto lemma preorder_conv_closed: "preorder x \ preorder (x\<^sup>T)" by (simp add: reflexive_conv_closed transitive_conv_closed) lemma symmetric_bot_closed: "symmetric bot" by simp lemma symmetric_one_closed: "symmetric 1" by simp lemma symmetric_top_closed: "symmetric top" by simp lemma symmetric_inf_closed: "symmetric x \ symmetric y \ symmetric (x \ y)" by (simp add: conv_dist_inf) lemma symmetric_sup_closed: "symmetric x \ symmetric y \ symmetric (x \ y)" by (simp add: conv_dist_sup) lemma symmetric_conv_closed: "symmetric x \ symmetric (x\<^sup>T)" by simp lemma one_inf_conv: "1 \ x = 1 \ x\<^sup>T" by (metis conv_dist_inf coreflexive_symmetric inf.cobounded1 symmetric_one_closed) lemma antisymmetric_bot_closed: "antisymmetric bot" by simp lemma antisymmetric_one_closed: "antisymmetric 1" by simp lemma antisymmetric_inf_closed: "antisymmetric x \ antisymmetric (x \ y)" by (rule order_trans[where y="x \ x\<^sup>T"]) (simp_all add: conv_isotone inf.coboundedI2 inf.sup_assoc) lemma antisymmetric_conv_closed: "antisymmetric x \ antisymmetric (x\<^sup>T)" by (simp add: inf_commute) lemma asymmetric_bot_closed: "asymmetric bot" by simp lemma asymmetric_inf_closed: "asymmetric x \ asymmetric (x \ y)" by (metis conv_dist_inf inf.mult_zero_left inf.left_commute inf_assoc) lemma asymmetric_conv_closed: "asymmetric x \ asymmetric (x\<^sup>T)" by (simp add: inf_commute) lemma linear_top_closed: "linear top" by simp lemma linear_sup_closed: "linear x \ linear (x \ y)" by (metis conv_dist_sup sup_assoc sup_commute sup_top_right) lemma linear_reflexive: "linear x \ reflexive x" by (metis one_inf_conv inf.distrib_left inf.cobounded2 inf.orderE reflexive_top_closed sup.idem) lemma linear_conv_closed: "linear x \ linear (x\<^sup>T)" by (simp add: sup_commute) lemma linear_comp_closed: assumes "linear x" and "linear y" shows "linear (x * y)" proof - have "reflexive y" by (simp add: assms(2) linear_reflexive) hence "x \ x\<^sup>T \ x * y \ y\<^sup>T * x\<^sup>T" by (metis case_split_left case_split_right le_supI sup.cobounded1 sup.cobounded2 sup.idem reflexive_conv_closed) thus ?thesis by (simp add: assms(1) conv_dist_comp top_le) qed lemma equivalence_one_closed: "equivalence 1" by simp lemma equivalence_top_closed: "equivalence top" by simp lemma equivalence_inf_closed: "equivalence x \ equivalence y \ equivalence (x \ y)" using conv_dist_inf preorder_inf_closed by auto lemma equivalence_conv_closed: "equivalence x \ equivalence (x\<^sup>T)" by simp lemma order_one_closed: "order 1" by simp lemma order_inf_closed: "order x \ order y \ order (x \ y)" using antisymmetric_inf_closed transitive_inf_closed by auto lemma order_conv_closed: "order x \ order (x\<^sup>T)" by (simp add: inf_commute reflexive_conv_closed transitive_conv_closed) lemma linear_order_conv_closed: "linear_order x \ linear_order (x\<^sup>T)" using equivalence_top_closed conv_dist_sup inf_commute reflexive_conv_closed transitive_conv_closed by force text \ We show a fact about equivalences. \ lemma equivalence_comp_dist_inf: "equivalence x \ x * y \ x * z = x * (y \ x * z)" by (metis antisym comp_right_subdist_inf dedekind_1 eq_iff inf.absorb1 inf.absorb2 mult_1_right mult_assoc) text \ The following result generalises the fact that composition with a test amounts to intersection with the corresponding vector. Both tests and vectors can be used to represent sets as relations. \ lemma coreflexive_comp_top_inf: "coreflexive x \ x * top \ y = x * y" apply (rule antisym) apply (metis comp_left_isotone comp_left_one coreflexive_symmetric dedekind_1 inf_top_left order_trans) using comp_left_isotone comp_right_isotone by fastforce lemma coreflexive_comp_top_inf_one: "coreflexive x \ x * top \ 1 = x" by (simp add: coreflexive_comp_top_inf) lemma coreflexive_comp_inf: "coreflexive x \ coreflexive y \ x * y = x \ y" by (metis (full_types) coreflexive_comp_top_inf coreflexive_comp_top_inf_one inf.mult_assoc inf.absorb2) lemma coreflexive_comp_inf_comp: assumes "coreflexive x" and "coreflexive y" shows "(x * z) \ (y * z) = (x \ y) * z" proof - have "(x * z) \ (y * z) = x * top \ z \ y * top \ z" using assms coreflexive_comp_top_inf inf_assoc by auto also have "... = x * top \ y * top \ z" by (simp add: inf.commute inf.left_commute) also have "... = (x \ y) * top \ z" by (metis assms coreflexive_comp_inf coreflexive_comp_top_inf mult_assoc) also have "... = (x \ y) * z" by (simp add: assms(1) coreflexive_comp_top_inf coreflexive_inf_closed) finally show ?thesis . qed lemma coreflexive_idempotent: "coreflexive x \ idempotent x" by (simp add: coreflexive_comp_inf) lemma coreflexive_univalent: "coreflexive x \ univalent x" by (simp add: coreflexive_idempotent coreflexive_symmetric) lemma coreflexive_injective: "coreflexive x \ injective x" by (simp add: coreflexive_idempotent coreflexive_symmetric) lemma coreflexive_commutative: "coreflexive x \ coreflexive y \ x * y = y * x" by (simp add: coreflexive_comp_inf inf.commute) lemma coreflexive_dedekind: "coreflexive x \ coreflexive y \ coreflexive z \ x * y \ z \ x * (y \ x * z)" by (simp add: coreflexive_comp_inf inf.coboundedI1 inf.left_commute) text \ Also the equational version of the Dedekind rule continues to hold. \ lemma dedekind_eq: "x * y \ z = (x \ (z * y\<^sup>T)) * (y \ (x\<^sup>T * z)) \ z" proof (rule antisym) have "x * y \ z \ x * (y \ (x\<^sup>T * z)) \ z" by (simp add: dedekind_1) also have "... \ (x \ (z * (y \ (x\<^sup>T * z))\<^sup>T)) * (y \ (x\<^sup>T * z)) \ z" by (simp add: dedekind_2) also have "... \ (x \ (z * y\<^sup>T)) * (y \ (x\<^sup>T * z)) \ z" by (metis comp_left_isotone comp_right_isotone inf_mono conv_order inf.cobounded1 order_refl) finally show "x * y \ z \ (x \ (z * y\<^sup>T)) * (y \ (x\<^sup>T * z)) \ z" . next show "(x \ (z * y\<^sup>T)) * (y \ (x\<^sup>T * z)) \ z \ x * y \ z" using comp_isotone inf.sup_left_isotone by auto qed lemma dedekind: "x * y \ z \ (x \ (z * y\<^sup>T)) * (y \ (x\<^sup>T * z))" by (metis dedekind_eq inf.cobounded1) lemma vector_export_comp: "(x * top \ y) * z = x * top \ y * z" proof - have "vector (x * top)" by (simp add: comp_associative) thus ?thesis by (simp add: vector_inf_comp) qed lemma vector_export_comp_unit: "(x * top \ 1) * y = x * top \ y" by (simp add: vector_export_comp) text \ We solve a few exercises from \cite{SchmidtStroehlein1993}. \ lemma ex231a [simp]: "(1 \ x * x\<^sup>T) * x = x" by (metis inf.cobounded1 inf.idem inf_right_idem comp_left_one conv_one coreflexive_comp_top_inf dedekind_eq) lemma ex231b [simp]: "x * (1 \ x\<^sup>T * x) = x" by (metis conv_dist_comp conv_dist_inf conv_involutive conv_one ex231a) lemma ex231c: "x \ x * x\<^sup>T * x" by (metis comp_left_isotone ex231a inf_le2) lemma ex231d: "x \ x * top * x" by (metis comp_left_isotone comp_right_isotone top_greatest order_trans ex231c) lemma ex231e [simp]: "x * top * x * top = x * top" by (metis ex231d antisym comp_associative mult_right_isotone top.extremum) lemma arc_injective: "arc x \ injective x" by (metis conv_dist_inf conv_involutive inf.absorb2 top_right_mult_increasing univalent_inf_closed) lemma arc_conv_closed: "arc x \ arc (x\<^sup>T)" by simp lemma arc_univalent: "arc x \ univalent x" using arc_conv_closed arc_injective univalent_conv_injective by blast text \ The following result generalises \cite[Exercise 2]{Oliveira2009}. It is used to show that the while-loop preserves injectivity of the constructed tree. \ lemma injective_sup: assumes "injective t" and "e * t\<^sup>T \ 1" and "injective e" shows "injective (t \ e)" proof - have "(t \ e) * (t \ e)\<^sup>T = t * t\<^sup>T \ t * e\<^sup>T \ e * t\<^sup>T \ e * e\<^sup>T" by (simp add: comp_left_dist_sup conv_dist_sup semiring.distrib_right sup.assoc) thus ?thesis using assms coreflexive_symmetric conv_dist_comp by fastforce qed lemma injective_inv: "injective t \ e * t\<^sup>T = bot \ arc e \ injective (t \ e)" using arc_injective injective_sup bot_least by blast lemma univalent_sup: "univalent t \ e\<^sup>T * t \ 1 \ univalent e \ univalent (t \ e)" by (metis injective_sup conv_dist_sup conv_involutive) lemma point_injective: "arc x \ x\<^sup>T * top * x \ 1" by (metis conv_top comp_associative conv_dist_comp conv_involutive vector_top_closed) lemma vv_transitive: "vector v \ (v * v\<^sup>T) * (v * v\<^sup>T) \ v * v\<^sup>T" by (metis comp_associative comp_left_isotone comp_right_isotone top_greatest) lemma epm_3: assumes "e \ w" and "injective w" shows "e = w \ top * e" proof - have "w \ top * e \ w * e\<^sup>T * e" by (metis (no_types, lifting) inf.absorb2 top.extremum dedekind_2 inf.commute) also have "... \ w * w\<^sup>T * e" by (simp add: assms(1) conv_isotone mult_left_isotone mult_right_isotone) also have "... \ e" using assms(2) coreflexive_comp_top_inf inf.sup_right_divisibility by blast finally show ?thesis by (simp add: assms(1) top_left_mult_increasing antisym) qed lemma comp_inf_vector: "x * (y \ z * top) = (x \ top * z\<^sup>T) * y" by (metis conv_top covector_inf_comp_3 comp_associative conv_dist_comp inf.commute vector_top_closed) lemma inf_vector_comp: "(x \ y * top) * z = y * top \ x * z" using inf.commute vector_export_comp by auto lemma comp_inf_covector: "x * (y \ top * z) = x * y \ top * z" by (simp add: covector_comp_inf covector_mult_closed) text \ Well-known distributivity properties of univalent and injective relations over meet continue to hold. \ lemma univalent_comp_left_dist_inf: assumes "univalent x" shows "x * (y \ z) = x * y \ x * z" proof (rule antisym) show "x * (y \ z) \ x * y \ x * z" by (simp add: comp_right_isotone) next have "x * y \ x * z \ (x \ x * z * y\<^sup>T) * (y \ x\<^sup>T * x * z)" by (metis comp_associative dedekind) also have "... \ x * (y \ x\<^sup>T * x * z)" by (simp add: comp_left_isotone) also have "... \ x * (y \ 1 * z)" using assms comp_left_isotone comp_right_isotone inf.sup_right_isotone by blast finally show "x * y \ x * z \ x * (y \ z)" by simp qed lemma injective_comp_right_dist_inf: "injective z \ (x \ y) * z = x * z \ y * z" by (metis univalent_comp_left_dist_inf conv_dist_comp conv_involutive conv_dist_inf) lemma vector_covector: "vector v \ vector w \ v \ w\<^sup>T = v * w\<^sup>T" by (metis covector_comp_inf inf_top_left vector_conv_covector) lemma comp_inf_vector_1: "(x \ top * y) * z = x * (z \ (top * y)\<^sup>T)" by (simp add: comp_inf_vector conv_dist_comp) text \ The shunting properties for bijective relations and mappings continue to hold. \ lemma shunt_bijective: assumes "bijective z" shows "x \ y * z \ x * z\<^sup>T \ y" proof assume "x \ y * z" hence "x * z\<^sup>T \ y * z * z\<^sup>T" by (simp add: mult_left_isotone) also have "... \ y" using assms comp_associative mult_right_isotone by fastforce finally show "x * z\<^sup>T \ y" . next assume 1: "x * z\<^sup>T \ y" have "x = x \ top * z" by (simp add: assms) also have "... \ x * z\<^sup>T * z" by (metis dedekind_2 inf_commute inf_top.right_neutral) also have "... \ y * z" using 1 by (simp add: mult_left_isotone) finally show "x \ y * z" . qed lemma shunt_mapping: "mapping z \ x \ z * y \ z\<^sup>T * x \ y" by (metis shunt_bijective mapping_conv_bijective conv_order conv_dist_comp conv_involutive) lemma bijective_reverse: assumes "bijective p" and "bijective q" shows "p \ r * q \ q \ r\<^sup>T * p" proof - have "p \ r * q \ p * q\<^sup>T \ r" by (simp add: assms(2) shunt_bijective) also have "... \ q\<^sup>T \ p\<^sup>T * r" by (metis assms(1) conv_dist_comp conv_involutive conv_order shunt_bijective) also have "... \ q \ r\<^sup>T * p" using conv_dist_comp conv_isotone by fastforce finally show ?thesis by simp qed lemma arc_expanded: "arc x \ x * top * x\<^sup>T \ 1 \ x\<^sup>T * top * x \ 1 \ top * x * top = top" by (metis conv_top comp_associative conv_dist_comp conv_involutive vector_top_closed) lemma arc_top_arc: assumes "arc x" shows "x * top * x = x" by (metis assms epm_3 top_right_mult_increasing vector_inf_comp vector_mult_closed vector_top_closed) lemma arc_top_edge: assumes "arc x" shows "x\<^sup>T * top * x = x\<^sup>T * x" proof - have "x\<^sup>T = x\<^sup>T * top \ top * x\<^sup>T" using assms epm_3 top_right_mult_increasing by simp thus ?thesis by (metis comp_inf_vector_1 conv_dist_comp conv_involutive conv_top inf.absorb1 top_right_mult_increasing) qed end subsection \Single-Object Pseudocomplemented Distributive Allegories\ text \ We extend single-object bounded distributive allegories by a pseudocomplement operation. The following definitions concern properties of relations that require a pseudocomplement. \ class relation_algebra_signature = bounded_distrib_allegory_signature + uminus begin abbreviation irreflexive :: "'a \ bool" where "irreflexive x \ x \ -1" abbreviation strict_linear :: "'a \ bool" where "strict_linear x \ x \ x\<^sup>T = -1" abbreviation strict_order :: "'a \ bool" where "strict_order x \ irreflexive x \ transitive x" abbreviation linear_strict_order :: "'a \ bool" where "linear_strict_order x \ strict_order x \ strict_linear x" text \ The following variants are useful for the graph model. \ abbreviation pp_mapping :: "'a \ bool" where "pp_mapping x \ univalent x \ total (--x)" abbreviation pp_bijective :: "'a \ bool" where "pp_bijective x \ injective x \ surjective (--x)" abbreviation pp_point :: "'a \ bool" where "pp_point x \ vector x \ pp_bijective x" abbreviation pp_arc :: "'a \ bool" where "pp_arc x \ pp_bijective (x * top) \ pp_bijective (x\<^sup>T * top)" end class pd_allegory = bounded_distrib_allegory + p_algebra begin subclass relation_algebra_signature . subclass pd_algebra .. lemma conv_complement_1: "-(x\<^sup>T) \ (-x)\<^sup>T = (-x)\<^sup>T" by (metis conv_dist_inf conv_order bot_least conv_involutive pseudo_complement sup.absorb2 sup.cobounded2) lemma conv_complement: "(-x)\<^sup>T = -(x\<^sup>T)" by (metis conv_complement_1 conv_dist_sup conv_involutive sup_commute) lemma conv_complement_sub_inf [simp]: "x\<^sup>T * -(x * y) \ y = bot" by (metis comp_left_zero conv_dist_comp conv_involutive dedekind_1 inf_import_p inf_p inf_right_idem ppp pseudo_complement regular_closed_bot) lemma conv_complement_sub_leq: "x\<^sup>T * -(x * y) \ -y" using pseudo_complement conv_complement_sub_inf by blast lemma conv_complement_sub [simp]: "x\<^sup>T * -(x * y) \ -y = -y" by (simp add: conv_complement_sub_leq sup.absorb2) lemma complement_conv_sub: "-(y * x) * x\<^sup>T \ -y" by (metis conv_complement conv_complement_sub_leq conv_order conv_dist_comp) text \ The following so-called Schr\"oder equivalences, or De Morgan's Theorem K, hold only with a pseudocomplemented element on both right-hand sides. \ lemma schroeder_3_p: "x * y \ -z \ x\<^sup>T * z \ -y" using pseudo_complement schroeder_1 by auto lemma schroeder_4_p: "x * y \ -z \ z * y\<^sup>T \ -x" using pseudo_complement schroeder_2 by auto lemma comp_pp_semi_commute: "x * --y \ --(x * y)" using conv_complement_sub_leq schroeder_3_p by fastforce text \ The following result looks similar to a property of (anti)domain. \ lemma p_comp_pp [simp]: "-(x * --y) = -(x * y)" using comp_pp_semi_commute comp_right_isotone inf.eq_iff p_antitone pp_increasing by fastforce lemma pp_comp_semi_commute: "--x * y \ --(x * y)" using complement_conv_sub schroeder_4_p by fastforce lemma p_pp_comp [simp]: "-(--x * y) = -(x * y)" using pp_comp_semi_commute comp_left_isotone inf.eq_iff p_antitone pp_increasing by fastforce lemma pp_comp_subdist: "--x * --y \ --(x * y)" by (simp add: p_antitone_iff) lemma theorem24xxiii: "x * y \ -(x * z) = x * (y \ -z) \ -(x * z)" proof - have "x * y \ -(x * z) \ x * (y \ (x\<^sup>T * -(x * z)))" by (simp add: dedekind_1) also have "... \ x * (y \ -z)" using comp_right_isotone conv_complement_sub_leq inf.sup_right_isotone by auto finally show ?thesis using comp_right_subdist_inf antisym inf.coboundedI2 inf.commute by auto qed text \ Even in Stone relation algebras, we do not obtain the backward implication in the following result. \ lemma vector_complement_closed: "vector x \ vector (-x)" by (metis complement_conv_sub conv_top eq_iff top_right_mult_increasing) lemma covector_complement_closed: "covector x \ covector (-x)" by (metis conv_complement_sub_leq conv_top eq_iff top_left_mult_increasing) lemma covector_vector_comp: "vector v \ -v\<^sup>T * v = bot" by (metis conv_bot conv_complement conv_complement_sub_inf conv_dist_comp conv_involutive inf_top.right_neutral) lemma irreflexive_bot_closed: "irreflexive bot" by simp lemma irreflexive_inf_closed: "irreflexive x \ irreflexive (x \ y)" by (simp add: le_infI1) lemma irreflexive_sup_closed: "irreflexive x \ irreflexive y \ irreflexive (x \ y)" by simp lemma irreflexive_conv_closed: "irreflexive x \ irreflexive (x\<^sup>T)" using conv_complement conv_isotone by fastforce lemma reflexive_complement_irreflexive: "reflexive x \ irreflexive (-x)" by (simp add: p_antitone) lemma irreflexive_complement_reflexive: "irreflexive x \ reflexive (-x)" by (simp add: p_antitone_iff) lemma symmetric_complement_closed: "symmetric x \ symmetric (-x)" by (simp add: conv_complement) lemma asymmetric_irreflexive: "asymmetric x \ irreflexive x" by (metis inf.mult_not_zero inf.left_commute inf.right_idem inf.sup_monoid.add_commute pseudo_complement one_inf_conv) lemma linear_asymmetric: "linear x \ asymmetric (-x)" using conv_complement p_top by force lemma strict_linear_sup_closed: "strict_linear x \ strict_linear y \ strict_linear (x \ y)" by (metis (mono_tags, hide_lams) conv_dist_sup sup.right_idem sup_assoc sup_commute) lemma strict_linear_irreflexive: "strict_linear x \ irreflexive x" using sup_left_divisibility by blast lemma strict_linear_conv_closed: "strict_linear x \ strict_linear (x\<^sup>T)" by (simp add: sup_commute) lemma strict_order_var: "strict_order x \ asymmetric x \ transitive x" by (metis asymmetric_irreflexive comp_right_one irreflexive_conv_closed conv_dist_comp dual_order.trans pseudo_complement schroeder_3_p) lemma strict_order_bot_closed: "strict_order bot" by simp lemma strict_order_inf_closed: "strict_order x \ strict_order y \ strict_order (x \ y)" using inf.coboundedI1 transitive_inf_closed by auto lemma strict_order_conv_closed: "strict_order x \ strict_order (x\<^sup>T)" using irreflexive_conv_closed transitive_conv_closed by blast lemma order_strict_order: assumes "order x" shows "strict_order (x \ -1)" proof (rule conjI) show 1: "irreflexive (x \ -1)" by simp have "antisymmetric (x \ -1)" using antisymmetric_inf_closed assms by blast hence "(x \ -1) * (x \ -1) \ 1 \ (x \ -1 \ (x \ -1)\<^sup>T) * (x \ -1 \ (x \ -1)\<^sup>T)" using 1 by (metis (no_types) coreflexive_symmetric irreflexive_inf_closed coreflexive_transitive dedekind_1 inf_idem mult_1_right semiring.mult_not_zero strict_order_var) also have "... = (x \ x\<^sup>T \ -1) * (x \ x\<^sup>T \ -1)" by (simp add: conv_complement conv_dist_inf inf.absorb2 inf.sup_monoid.add_assoc) also have "... = bot" using assms inf.antisym reflexive_conv_closed by fastforce finally have "(x \ -1) * (x \ -1) \ -1" using le_bot pseudo_complement by blast thus "transitive (x \ -1)" by (meson assms comp_isotone inf.boundedI inf.cobounded1 inf.order_lesseq_imp) qed lemma strict_order_order: "strict_order x \ order (x \ 1)" apply (unfold strict_order_var, intro conjI) apply simp apply (simp add: mult_left_dist_sup mult_right_dist_sup sup.absorb2) using conv_dist_sup coreflexive_bot_closed sup.absorb2 sup_inf_distrib2 by fastforce lemma linear_strict_order_conv_closed: "linear_strict_order x \ linear_strict_order (x\<^sup>T)" by (simp add: irreflexive_conv_closed sup_monoid.add_commute transitive_conv_closed) lemma linear_order_strict_order: "linear_order x \ linear_strict_order (x \ -1)" apply (rule conjI) using order_strict_order apply simp by (metis conv_complement conv_dist_inf coreflexive_symmetric eq_iff inf.absorb2 inf.distrib_left inf.sup_monoid.add_commute top.extremum) lemma regular_conv_closed: "regular x \ regular (x\<^sup>T)" by (metis conv_complement) text \ We show a number of facts about equivalences. \ lemma equivalence_comp_left_complement: "equivalence x \ x * -x = -x" apply (rule antisym) apply (metis conv_complement_sub_leq preorder_idempotent) using mult_left_isotone by fastforce lemma equivalence_comp_right_complement: "equivalence x \ -x * x = -x" by (metis equivalence_comp_left_complement conv_complement conv_dist_comp) text \ The pseudocomplement of tests is given by the following operation. \ -abbreviation coreflexive_complement :: "'a \ 'a" ("_ '" [80] 80) +abbreviation coreflexive_complement :: "'a \ 'a" ("_ ''" [80] 80) where "x ' \ -x \ 1" lemma coreflexive_comp_top_coreflexive_complement: "coreflexive x \ (x * top)' = x '" by (metis coreflexive_comp_top_inf_one inf.commute inf_import_p) lemma coreflexive_comp_inf_complement: "coreflexive x \ (x * y) \ -z = (x * y) \ -(x * z)" by (metis coreflexive_comp_top_inf inf.sup_relative_same_increasing inf_import_p inf_le1) lemma double_coreflexive_complement: "x '' = (-x)'" using inf.sup_monoid.add_commute inf_import_p by auto lemma coreflexive_pp_dist_comp: assumes "coreflexive x" and "coreflexive y" shows "(x * y)'' = x '' * y ''" proof - have "(x * y)'' = --(x * y) \ 1" by (simp add: double_coreflexive_complement) also have "... = --x \ --y \ 1" by (simp add: assms coreflexive_comp_inf) also have "... = (--x \ 1) * (--y \ 1)" by (simp add: coreflexive_comp_inf inf.left_commute inf.sup_monoid.add_assoc) also have "... = x '' * y ''" by (simp add: double_coreflexive_complement) finally show ?thesis . qed lemma coreflexive_pseudo_complement: "coreflexive x \ x \ y = bot \ x \ y '" by (simp add: pseudo_complement) lemma pp_bijective_conv_mapping: "pp_bijective x \ pp_mapping (x\<^sup>T)" by (simp add: conv_complement surjective_conv_total) lemma pp_arc_expanded: "pp_arc x \ x * top * x\<^sup>T \ 1 \ x\<^sup>T * top * x \ 1 \ top * --x * top = top" proof assume 1: "pp_arc x" have 2: "x * top * x\<^sup>T \ 1" using 1 by (metis comp_associative conv_dist_comp equivalence_top_closed vector_top_closed) have 3: "x\<^sup>T * top * x \ 1" using 1 by (metis conv_dist_comp conv_involutive equivalence_top_closed vector_top_closed mult_assoc) have 4: "x\<^sup>T \ x\<^sup>T * x * x\<^sup>T" by (metis conv_involutive ex231c) have "top = --(top * x) * top" using 1 by (metis conv_complement conv_dist_comp conv_involutive equivalence_top_closed) also have "... \ --(top * x\<^sup>T * top * x) * top" using 1 by (metis eq_refl mult_assoc p_comp_pp p_pp_comp) also have "... = (top * --(x * top) \ --(top * x\<^sup>T * top * x)) * top" using 1 by simp also have "... = top * (--(x * top) \ --(top * x\<^sup>T * top * x)) * top" by (simp add: covector_complement_closed covector_comp_inf covector_mult_closed) also have "... = top * --(x * top \ top * x\<^sup>T * top * x) * top" by simp also have "... = top * --(x * top * x\<^sup>T * top * x) * top" by (metis comp_associative comp_inf_covector inf_top.left_neutral) also have "... \ top * --(x * top * x\<^sup>T * x * x\<^sup>T * top * x) * top" using 4 by (metis comp_associative comp_left_isotone comp_right_isotone pp_isotone) also have "... \ top * --(x * x\<^sup>T * top * x) * top" using 2 by (metis comp_associative comp_left_isotone comp_right_isotone pp_isotone comp_left_one) also have "... \ top * --x * top" using 3 by (metis comp_associative comp_left_isotone comp_right_isotone pp_isotone comp_right_one) finally show "x * top * x\<^sup>T \ 1 \ x\<^sup>T * top * x \ 1 \ top * --x * top = top" using 2 3 top_le by blast next assume "x * top * x\<^sup>T \ 1 \ x\<^sup>T * top * x \ 1 \ top * --x * top = top" thus "pp_arc x" apply (intro conjI) apply (metis comp_associative conv_dist_comp equivalence_top_closed vector_top_closed) apply (metis comp_associative mult_right_isotone top_le pp_comp_semi_commute) apply (metis conv_dist_comp coreflexive_symmetric vector_conv_covector vector_top_closed mult_assoc) by (metis conv_complement conv_dist_comp equivalence_top_closed inf.orderE inf_top.left_neutral mult_right_isotone pp_comp_semi_commute) qed text \ The following operation represents states with infinite executions of non-strict computations. \ abbreviation N :: "'a \ 'a" where "N x \ -(-x * top) \ 1" lemma N_comp: "N x * y = -(-x * top) \ y" by (simp add: vector_mult_closed vector_complement_closed vector_inf_one_comp) lemma N_comp_top [simp]: "N x * top = -(-x * top)" by (simp add: N_comp) lemma vector_N_pp: "vector x \ N x = --x \ 1" by (simp add: vector_complement_closed) lemma N_vector_pp [simp]: "N (x * top) = --(x * top) \ 1" by (simp add: comp_associative vector_complement_closed) lemma N_vector_top_pp [simp]: "N (x * top) * top = --(x * top)" by (metis N_comp_top comp_associative vector_top_closed vector_complement_closed) lemma N_below_inf_one_pp: "N x \ --x \ 1" using inf.sup_left_isotone p_antitone top_right_mult_increasing by auto lemma N_below_pp: "N x \ --x" using N_below_inf_one_pp by auto lemma N_comp_N: "N x * N y = -(-x * top) \ -(-y * top) \ 1" by (simp add: N_comp inf.mult_assoc) lemma N_bot [simp]: "N bot = bot" by simp lemma N_top [simp]: "N top = 1" by simp lemma n_split_omega_mult_pp: "xs * --xo = xo \ vector xo \ N top * xo = xs * N xo * top" by (metis N_top N_vector_top_pp comp_associative comp_left_one) text \ Many of the following results have been derived for verifying Prim's minimum spanning tree algorithm. \ lemma ee: assumes "vector v" and "e \ v * -v\<^sup>T" shows "e * e = bot" proof - have "e * v \ bot" by (metis assms covector_vector_comp comp_associative mult_left_isotone mult_right_zero) thus ?thesis by (metis assms(2) bot_unique comp_associative mult_right_isotone semiring.mult_not_zero) qed lemma et: assumes "vector v" and "e \ v * -v\<^sup>T" and "t \ v * v\<^sup>T" shows "e * t = bot" and "e * t\<^sup>T = bot" proof - have "e * t \ v * -v\<^sup>T * v * v\<^sup>T" using assms(2-3) comp_isotone mult_assoc by fastforce thus "e * t = bot" by (simp add: assms(1) covector_vector_comp le_bot mult_assoc) next have "t\<^sup>T \ v * v\<^sup>T" using assms(3) conv_order conv_dist_comp by fastforce hence "e * t\<^sup>T \ v * -v\<^sup>T * v * v\<^sup>T" by (metis assms(2) comp_associative comp_isotone) thus "e * t\<^sup>T = bot" by (simp add: assms(1) covector_vector_comp le_bot mult_assoc) qed lemma ve_dist: assumes "e \ v * -v\<^sup>T" and "vector v" and "arc e" shows "(v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T = v * v\<^sup>T \ v * v\<^sup>T * e \ e\<^sup>T * v * v\<^sup>T \ e\<^sup>T * e" proof - have "e \ v * top" using assms(1) comp_right_isotone dual_order.trans top_greatest by blast hence "v * top * e = v * top * (v * top \ e)" by (simp add: inf.absorb2) also have "... = (v * top \ top * v\<^sup>T) * e" using assms(2) covector_inf_comp_3 vector_conv_covector by force also have "... = v * top * v\<^sup>T * e" by (metis assms(2) inf_top_right vector_inf_comp) also have "... = v * v\<^sup>T * e" by (simp add: assms(2)) finally have 1: "v * top * e = v * v\<^sup>T * e" . have "e\<^sup>T * top * e \ e\<^sup>T * top * e * e\<^sup>T * e" using ex231c comp_associative mult_right_isotone by auto also have "... \ e\<^sup>T * e" by (metis assms(3) coreflexive_comp_top_inf le_infE mult_semi_associative point_injective) finally have 2: "e\<^sup>T * top * e = e\<^sup>T * e" by (simp add: inf.antisym mult_left_isotone top_right_mult_increasing) have "(v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T = (v \ e\<^sup>T * top) * (v\<^sup>T \ top * e)" by (simp add: conv_dist_comp conv_dist_sup) also have "... = v * v\<^sup>T \ v * top * e \ e\<^sup>T * top * v\<^sup>T \ e\<^sup>T * top * top * e" by (metis semiring.distrib_left semiring.distrib_right sup_assoc mult_assoc) also have "... = v * v\<^sup>T \ v * top * e \ (v * top * e)\<^sup>T \ e\<^sup>T * top * e" by (simp add: comp_associative conv_dist_comp) also have "... = v * v\<^sup>T \ v * v\<^sup>T * e \ (v * v\<^sup>T * e)\<^sup>T \ e\<^sup>T * e" using 1 2 by simp finally show ?thesis by (simp add: comp_associative conv_dist_comp) qed lemma ev: "vector v \ e \ v * -v\<^sup>T \ e * v = bot" by (metis covector_vector_comp antisym bot_least comp_associative mult_left_isotone mult_right_zero) lemma vTeT: "vector v \ e \ v * -v\<^sup>T \ v\<^sup>T * e\<^sup>T = bot" using conv_bot ev conv_dist_comp by fastforce text \ The following result is used to show that the while-loop of Prim's algorithm preserves that the constructed tree is a subgraph of g. \ lemma prim_subgraph_inv: assumes "e \ v * -v\<^sup>T \ g" and "t \ v * v\<^sup>T \ g" shows "t \ e \ ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T) \ g" proof (rule sup_least) have "t \ ((v \ e\<^sup>T * top) * v\<^sup>T) \ g" using assms(2) le_supI1 mult_right_dist_sup by auto also have "... \ ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T) \ g" using comp_right_isotone conv_dist_sup inf.sup_left_isotone by auto finally show "t \ ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T) \ g" . next have "e \ v * top" by (meson assms(1) inf.boundedE mult_right_isotone order.trans top.extremum) hence "e \ v * top \ top * e" by (simp add: top_left_mult_increasing) also have "... = v * top * e" by (metis inf_top_right vector_export_comp) finally have "e \ v * top * e \ g" using assms(1) by auto also have "... = v * (e\<^sup>T * top)\<^sup>T \ g" by (simp add: comp_associative conv_dist_comp) also have "... \ v * (v \ e\<^sup>T * top)\<^sup>T \ g" by (simp add: conv_dist_sup mult_left_dist_sup sup.assoc sup.orderI) also have "... \ (v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T \ g" using inf.sup_left_isotone mult_right_sub_dist_sup_left by auto finally show "e \ ((v \ e\<^sup>T * top) * (v \ e\<^sup>T * top)\<^sup>T) \ g" . qed text \ The following result shows how to apply the Schr\"oder equivalence to the middle factor in a composition of three relations. Again the elements on the right-hand side need to be pseudocomplemented. \ lemma triple_schroeder_p: "x * y * z \ -w \ x\<^sup>T * w * z\<^sup>T \ -y" using mult_assoc p_antitone_iff schroeder_3_p schroeder_4_p by auto text \ The rotation versions of the Schr\"oder equivalences continue to hold, again with pseudocomplemented elements on the right-hand side. \ lemma schroeder_5_p: "x * y \ -z \ y * z\<^sup>T \ -x\<^sup>T" using schroeder_3_p schroeder_4_p by auto lemma schroeder_6_p: "x * y \ -z \ z\<^sup>T * x \ -y\<^sup>T" using schroeder_3_p schroeder_4_p by auto lemma vector_conv_compl: "vector v \ top * -v\<^sup>T = -v\<^sup>T" by (simp add: covector_complement_closed vector_conv_covector) text \ Composition commutes, relative to the diversity relation. \ lemma comp_commute_below_diversity: "x * y \ -1 \ y * x \ -1" by (metis comp_right_one conv_dist_comp conv_one schroeder_3_p schroeder_4_p) lemma comp_injective_below_complement: "injective y \ -x * y \ -(x * y)" by (metis p_antitone_iff comp_associative comp_right_isotone comp_right_one schroeder_4_p) lemma comp_univalent_below_complement: "univalent x \ x * -y \ -(x * y)" by (metis p_inf pseudo_complement semiring.mult_zero_right univalent_comp_left_dist_inf) text \ Bijective relations and mappings can be exported from a pseudocomplement. \ lemma comp_bijective_complement: "bijective y \ -x * y = -(x * y)" using comp_injective_below_complement complement_conv_sub antisym shunt_bijective by blast lemma comp_mapping_complement: "mapping x \ x * -y = -(x * y)" by (metis (full_types) comp_bijective_complement conv_complement conv_dist_comp conv_involutive total_conv_surjective) text \ The following facts are used in the correctness proof of Kruskal's minimum spanning tree algorithm. \ lemma kruskal_injective_inv: assumes "injective f" and "covector q" and "q * f\<^sup>T \ q" and "e \ q" and "q * f\<^sup>T \ -e" and "injective e" and "q\<^sup>T * q \ f\<^sup>T * f \ 1" shows "injective ((f \ -q) \ (f \ q)\<^sup>T \ e)" proof - have 1: "(f \ -q) * (f \ -q)\<^sup>T \ 1" by (simp add: assms(1) injective_inf_closed) have 2: "(f \ -q) * (f \ q) \ 1" proof - have 21: "bot = q * f\<^sup>T \ - q" by (metis assms(3) inf.sup_monoid.add_assoc inf.sup_right_divisibility inf_import_p inf_p) have "(f \ -q) * (f \ q) \ -q * f \ q" by (metis assms(2) comp_inf_covector comp_isotone inf.cobounded2 inf.left_idem) also have "... = bot" using 21 schroeder_2 by auto finally show ?thesis by (simp add: bot_unique) qed have 3: "(f \ -q) * e\<^sup>T \ 1" proof - have "(f \ -q) * e\<^sup>T \ -q * e\<^sup>T" by (simp add: mult_left_isotone) also have "... = bot" by (metis assms(2,4) bot_unique conv_bot conv_complement covector_complement_closed p_antitone p_bot regular_closed_bot schroeder_5_p) finally show ?thesis by (simp add: bot_unique) qed have 4: "(f \ q)\<^sup>T * (f \ -q)\<^sup>T \ 1" using 2 conv_dist_comp conv_isotone by force have 5: "(f \ q)\<^sup>T * (f \ q) \ 1" proof - have "(f \ q)\<^sup>T * (f \ q) \ q\<^sup>T * q \ f\<^sup>T * f" by (simp add: conv_isotone mult_isotone) also have "... \ 1" by (simp add: assms(7)) finally show ?thesis by simp qed have 6: "(f \ q)\<^sup>T * e\<^sup>T \ 1" proof - have "f\<^sup>T * e\<^sup>T \ -q\<^sup>T" using assms(5) schroeder_5_p by simp hence "(f \ q)\<^sup>T * e\<^sup>T = bot" by (metis assms(2,5) conv_bot conv_dist_comp covector_comp_inf inf.absorb1 inf.cobounded2 inf.sup_monoid.add_commute inf_left_commute inf_p schroeder_4_p) thus ?thesis by (simp add: bot_unique) qed have 7: "e * (f \ -q)\<^sup>T \ 1" using 3 conv_dist_comp coreflexive_symmetric by fastforce have 8: "e * (f \ q) \ 1" using 6 conv_dist_comp coreflexive_symmetric by fastforce have 9: "e * e\<^sup>T \ 1" by (simp add: assms(6)) have "((f \ -q) \ (f \ q)\<^sup>T \ e) * ((f \ -q) \ (f \ q)\<^sup>T \ e)\<^sup>T = (f \ -q) * (f \ -q)\<^sup>T \ (f \ -q) * (f \ q) \ (f \ -q) * e\<^sup>T \ (f \ q)\<^sup>T * (f \ -q)\<^sup>T \ (f \ q)\<^sup>T * (f \ q) \ (f \ q)\<^sup>T * e\<^sup>T \ e * (f \ -q)\<^sup>T \ e * (f \ q) \ e * e\<^sup>T" using comp_left_dist_sup comp_right_dist_sup conv_dist_sup sup.assoc by simp also have "... \ 1" using 1 2 3 4 5 6 7 8 9 by simp finally show ?thesis by simp qed lemma kruskal_exchange_injective_inv_1: assumes "injective f" and "covector q" and "q * f\<^sup>T \ q" and "q\<^sup>T * q \ f\<^sup>T * f \ 1" shows "injective ((f \ -q) \ (f \ q)\<^sup>T)" using kruskal_injective_inv[where e=bot] by (simp add: assms) lemma kruskal_exchange_acyclic_inv_3: assumes "injective w" and "d \ w" shows "(w \ -d) * d\<^sup>T * top = bot" proof - have "(w \ -d) * d\<^sup>T * top = (w \ -d \ (d\<^sup>T * top)\<^sup>T) * top" by (simp add: comp_associative comp_inf_vector_1 conv_dist_comp) also have "... = (w \ top * d \ -d) * top" by (simp add: conv_dist_comp inf_commute inf_left_commute) finally show ?thesis using assms epm_3 by simp qed lemma kruskal_subgraph_inv: assumes "f \ --(-h \ g)" and "e \ --g" and "symmetric h" and "symmetric g" shows "(f \ -q) \ (f \ q)\<^sup>T \ e \ --(-(h \ -e \ -e\<^sup>T) \ g)" proof - let ?f = "(f \ -q) \ (f \ q)\<^sup>T \ e" let ?h = "h \ -e \ -e\<^sup>T" have 1: "f \ -q \ -h \ --g" using assms(1) inf.coboundedI1 by simp have "(f \ q)\<^sup>T \ (-h \ --g)\<^sup>T" using assms(1) inf.coboundedI1 conv_isotone by simp also have "... = -h \ --g" using assms(3,4) conv_complement conv_dist_inf by simp finally have "?f \ (-h \ --g) \ (e \ --g)" using 1 assms(2) inf.absorb1 semiring.add_right_mono by simp also have "... \ (-h \ --e) \ --g" by (simp add: inf.coboundedI1 le_supI2 pp_increasing) also have "... \ -?h \ --g" using inf.sup_left_isotone order_trans p_antitone_inf p_supdist_inf by blast finally show "?f \ --(-?h \ g)" using inf_pp_semi_commute order_lesseq_imp by blast qed end subsection \Stone Relation Algebras\ text \ We add \pp_dist_comp\ and \pp_one\, which follow in relation algebras but not in the present setting. The main change is that only a Stone algebra is required, not a Boolean algebra. \ class stone_relation_algebra = pd_allegory + stone_algebra + assumes pp_dist_comp : "--(x * y) = --x * --y" assumes pp_one [simp]: "--1 = 1" begin text \ The following property is a simple consequence of the Stone axiom. We cannot hope to remove the double complement in it. \ lemma conv_complement_0_p [simp]: "(-x)\<^sup>T \ (--x)\<^sup>T = top" by (metis conv_top conv_dist_sup stone) lemma theorem24xxiv_pp: "-(x * y) \ --(x * z) = -(x * (y \ -z)) \ --(x * z)" by (metis p_dist_inf theorem24xxiii) lemma asymmetric_linear: "asymmetric x \ linear (-x)" by (metis conv_complement inf.distrib_left inf_p maddux_3_11_pp p_bot p_dist_inf) lemma strict_linear_asymmetric: "strict_linear x \ antisymmetric (-x)" by (metis conv_complement eq_refl p_dist_sup pp_one) lemma regular_complement_top: "regular x \ x \ -x = top" by (metis stone) lemma regular_mult_closed: "regular x \ regular y \ regular (x * y)" by (simp add: pp_dist_comp) lemma regular_one_closed: "regular 1" by simp text \ The following variants of total and surjective are useful for graphs. \ lemma pp_total: "total (--x) \ -(x*top) = bot" by (simp add: dense_pp pp_dist_comp) lemma pp_surjective: "surjective (--x) \ -(top*x) = bot" by (metis p_bot p_comp_pp p_top pp_dist_comp) text \ Bijective elements and mappings are necessarily regular, that is, invariant under double-complement. This implies that points are regular. Moreover, also arcs are regular. \ lemma bijective_regular: "bijective x \ regular x" by (metis comp_bijective_complement mult_left_one regular_one_closed) lemma mapping_regular: "mapping x \ regular x" by (metis bijective_regular conv_complement conv_involutive total_conv_surjective) lemma arc_regular: assumes "arc x" shows "regular x" proof - have "--x \ --(x * top \ top * x)" by (simp add: pp_isotone top_left_mult_increasing top_right_mult_increasing) also have "... = --(x * top) \ --(top * x)" by simp also have "... = x * top \ top * x" by (metis assms bijective_regular conv_top conv_dist_comp conv_involutive mapping_regular) also have "... \ x * x\<^sup>T * top * x" by (metis comp_associative dedekind_1 inf.commute inf_top.right_neutral) also have "... \ x" by (metis assms comp_right_one conv_top comp_associative conv_dist_comp conv_involutive mult_right_isotone vector_top_closed) finally show ?thesis by (simp add: antisym pp_increasing) qed (* lemma conv_complement_0 [simp]: "x\<^sup>T \ (-x)\<^sup>T = top" nitpick [expect=genuine] oops lemma schroeder_3: "x * y \ z \ x\<^sup>T * -z \ -y" nitpick [expect=genuine] oops lemma schroeder_4: "x * y \ z \ -z * y\<^sup>T \ -x" nitpick [expect=genuine] oops lemma theorem24xxiv: "-(x * y) \ (x * z) = -(x * (y \ -z)) \ (x * z)" nitpick [expect=genuine] oops lemma vector_N: "x = x * top \ N(x) = x \ 1" nitpick [expect=genuine] oops lemma N_vector [simp]: "N(x * top) = x * top \ 1" nitpick [expect=genuine] oops lemma N_vector_top [simp]: "N(x * top) * top = x * top" nitpick [expect=genuine] oops lemma N_below_inf_one: "N(x) \ x \ 1" nitpick [expect=genuine] oops lemma N_below: "N(x) \ x" nitpick [expect=genuine] oops lemma n_split_omega_mult: "xs * xo = xo \ xo * top = xo \ N(top) * xo = xs * N(xo) * top" nitpick [expect=genuine] oops lemma complement_vector: "vector v \ vector (-v)" nitpick [expect=genuine] oops lemma complement_covector: "covector v \ covector (-v)" nitpick [expect=genuine] oops lemma triple_schroeder: "x * y * z \ w \ x\<^sup>T * -w * z\<^sup>T \ -y" nitpick [expect=genuine] oops lemma schroeder_5: "x * y \ z \ y * -z\<^sup>T \ -x\<^sup>T" nitpick [expect=genuine] oops lemma schroeder_6: "x * y \ z \ -z\<^sup>T * x \ -y\<^sup>T" nitpick [expect=genuine] oops *) end text \ Every Stone algebra can be expanded to a Stone relation algebra by identifying the semiring and lattice structures and taking identity as converse. \ sublocale stone_algebra < comp_inf: stone_relation_algebra where one = top and times = inf and conv = id proof (unfold_locales, goal_cases) case 7 show ?case by (simp add: inf_commute) qed (auto simp: inf.assoc inf_sup_distrib2 inf_left_commute) text \ Every bounded linear order can be expanded to a Stone algebra, which can be expanded to a Stone relation algebra by reusing some of the operations. In particular, composition is meet, its identity is \top\ and converse is the identity function. \ class linorder_stone_relation_algebra_expansion = linorder_stone_algebra_expansion + times + conv + one + assumes times_def [simp]: "x * y = min x y" assumes conv_def [simp]: "x\<^sup>T = x" assumes one_def [simp]: "1 = top" begin lemma times_inf [simp]: "x * y = x \ y" by simp subclass stone_relation_algebra apply unfold_locales using comp_inf.mult_right_dist_sup inf_commute inf_assoc inf_left_commute pp_dist_inf min_def by simp_all lemma times_dense: "x \ bot \ y \ bot \ x * y \ bot" using inf_dense min_inf times_def by presburger end subsection \Relation Algebras\ text \ For a relation algebra, we only require that the underlying lattice is a Boolean algebra. In fact, the only missing axiom is that double-complement is the identity. \ class relation_algebra = boolean_algebra + stone_relation_algebra begin lemma conv_complement_0 [simp]: "x\<^sup>T \ (-x)\<^sup>T = top" by (simp add: conv_complement) text \ We now obtain the original formulations of the Schr\"oder equivalences. \ lemma schroeder_3: "x * y \ z \ x\<^sup>T * -z \ -y" by (simp add: schroeder_3_p) lemma schroeder_4: "x * y \ z \ -z * y\<^sup>T \ -x" by (simp add: schroeder_4_p) lemma theorem24xxiv: "-(x * y) \ (x * z) = -(x * (y \ -z)) \ (x * z)" using theorem24xxiv_pp by auto lemma vector_N: "vector x \ N(x) = x \ 1" by (simp add: vector_N_pp) lemma N_vector [simp]: "N(x * top) = x * top \ 1" by simp lemma N_vector_top [simp]: "N(x * top) * top = x * top" using N_vector_top_pp by simp lemma N_below_inf_one: "N(x) \ x \ 1" using N_below_inf_one_pp by simp lemma N_below: "N(x) \ x" using N_below_pp by simp lemma n_split_omega_mult: "xs * xo = xo \ xo * top = xo \ N(top) * xo = xs * N(xo) * top" using n_split_omega_mult_pp by simp lemma complement_vector: "vector v \ vector (-v)" using vector_complement_closed by fastforce lemma complement_covector: "covector v \ covector (-v)" using covector_complement_closed by force lemma triple_schroeder: "x * y * z \ w \ x\<^sup>T * -w * z\<^sup>T \ -y" by (simp add: triple_schroeder_p) lemma schroeder_5: "x * y \ z \ y * -z\<^sup>T \ -x\<^sup>T" by (simp add: conv_complement schroeder_5_p) lemma schroeder_6: "x * y \ z \ -z\<^sup>T * x \ -y\<^sup>T" by (simp add: conv_complement schroeder_5_p) end text \ We briefly look at the so-called Tarski rule. In some models of Stone relation algebras it only holds for regular elements, so we add this as an assumption. \ class stone_relation_algebra_tarski = stone_relation_algebra + assumes tarski: "regular x \ x \ bot \ top * x * top = top" begin text \ We can then show, for example, that every arc is contained in a pseudocomplemented relation or its pseudocomplement. \ lemma arc_in_partition: assumes "arc x" shows "x \ -y \ x \ --y" proof - have 1: "x * top * x\<^sup>T \ 1 \ x\<^sup>T * top * x \ 1" using assms arc_expanded by auto have "\ x \ --y \ x \ -y" proof assume "\ x \ --y" hence "x \ -y \ bot" using pseudo_complement by simp hence "top * (x \ -y) * top = top" using assms arc_regular tarski by auto hence "x = x \ top * (x \ -y) * top" by simp also have "... \ x \ x * ((x \ -y) * top)\<^sup>T * (x \ -y) * top" by (metis dedekind_2 inf.cobounded1 inf.boundedI inf_commute mult_assoc inf.absorb2 top.extremum) also have "... = x \ x * top * (x\<^sup>T \ -y\<^sup>T) * (x \ -y) * top" by (simp add: comp_associative conv_complement conv_dist_comp conv_dist_inf) also have "... \ x \ x * top * x\<^sup>T * (x \ -y) * top" using inf.sup_right_isotone mult_left_isotone mult_right_isotone by auto also have "... \ x \ 1 * (x \ -y) * top" using 1 by (metis comp_associative comp_isotone inf.sup_right_isotone mult_1_left mult_semi_associative) also have "... = x \ (x \ -y) * top" by simp also have "... \ (x \ -y) * ((x \ -y)\<^sup>T * x)" by (metis dedekind_1 inf_commute inf_top_right) also have "... \ (x \ -y) * (x\<^sup>T * x)" by (simp add: conv_dist_inf mult_left_isotone mult_right_isotone) also have "... \ (x \ -y) * (x\<^sup>T * top * x)" by (simp add: mult_assoc mult_right_isotone top_left_mult_increasing) also have "... \ x \ -y" using 1 by (metis mult_right_isotone mult_1_right) finally show "x \ -y" by simp qed thus ?thesis by auto qed lemma non_bot_arc_in_partition_xor: assumes "arc x" and "x \ bot" shows "(x \ -y \ \ x \ --y) \ (\ x \ -y \ x \ --y)" proof - have "x \ -y \ x \ --y \ False" by (simp add: assms(2) inf_absorb1 shunting_1_pp) thus ?thesis using assms(1) arc_in_partition by auto qed end class relation_algebra_tarski = relation_algebra + stone_relation_algebra_tarski text \ Finally, the above axioms of relation algebras do not imply that they contain at least two elements. This is necessary, for example, to show that arcs are not empty. \ class stone_relation_algebra_consistent = stone_relation_algebra + assumes consistent: "bot \ top" begin lemma arc_not_bot: "arc x \ x \ bot" using consistent mult_right_zero by auto end class relation_algebra_consistent = relation_algebra + stone_relation_algebra_consistent class stone_relation_algebra_tarski_consistent = stone_relation_algebra_tarski + stone_relation_algebra_consistent begin lemma arc_in_partition_xor: "arc x \ (x \ -y \ \ x \ --y) \ (\ x \ -y \ x \ --y)" by (simp add: non_bot_arc_in_partition_xor arc_not_bot) end end