(* Title: strand ID: $Id: strand.thy,v 1.00 2008/07/02 13:41:34 Yongjian Li ver1.0 $ Author: Yongjian Li, Institute of software, Chinese academy of sciences Copyright 2008 Institute of software, Chinese academy of sciences Datatypes of strands, bundles, components, authentication tests; the main results of strand space theories *) header{*Theory of strand spaces for Security Protocols*} theory strand05=Event+Public: section{*Main defintions*} datatype Sign=positive ("+" 100) |negative ("-" 100) typedecl sigma types signed_msg="Sign× msg" types node=" sigma × nat" types strand_space="sigma => signed_msg list " consts Sigma_set::"sigma set" ("Σ") SP::"strand_space" constdefs Domain::"node set" "Domain == {(n1,i). n1 ∈ Σ ∧ i < length (SP n1)}" constdefs strand::"node => sigma" "strand n==fst n" constdefs index ::"node => nat" "index n == snd n" constdefs node_sign ::"node => Sign" "node_sign n ==fst (nth (SP(fst(n))) (snd(n)) )" constdefs node_term ::"node=>msg" "node_term n== snd (nth (SP(fst(n))) (snd(n)) )" constdefs casual1:: "( node × node ) set " "casual1 == { (n1,n2) . n1 ∈ Domain ∧ n2 ∈ Domain ∧ node_sign n1= + ∧ node_sign n2= - ∧ node_term n1= node_term n2 ∧ strand n1 ≠ strand n2 } " syntax "_casual1"::" node =>node=>bool" (infix "->" 100) translations "n1->n2 "=="(n1 ,n2) ∈ casual1" constdefs casual2::"(node × node) set" "casual2 == { (n1,n2) . n1 ∈ Domain ∧ n2 ∈ Domain ∧ (strand n1)= (strand n2 ) ∧ Suc (index n1)=index n2} " syntax "_casual2"::" node =>node=>bool" (infix "=>" 50) translations "n1=>n2 "=="(n1 ,n2):casual2" syntax "_casual2Trans"::" node =>node=>bool" (infix "=>+" 50) translations "n1=>+n2 "=="(n1 ,n2):casual2^+" syntax "_casual2TransReflex"::" node =>node=>bool" (infix "=>*" 50) translations "n1=>*n2 "=="(n1 ,n2):casual2^*" constdefs casual3::"(node × node) set" "casual3 == { (n1,n2) . n1->n2 ∨(n1, n2):(casual2^+) }" syntax "_casual3"::" node =>node=>bool" (infix "\<mapsto> " 50) translations "n1\<mapsto> n2 "=="(n1 ,n2):casual3" types edge="node × node" graph="node set × edge set" consts attr::"sigma=> agent" (* "attr s== (fst (SP s))"*) constdefs Atoms::"msg set" "Atoms=={a. (∃ ag. a=Agent ag) ∨ (∃ n. a=Number n) ∨(∃ n. a=Nonce n) | (∃ k. a=Key k)}" syntax Is_atom::"msg=> bool" translations "Is_atom m"=="m ∈ Atoms" consts KP::"key set" defs KP_def:"KP=={k. ∃ A. (k=pubK A)|(A ∈ bad ∧(k=priK A |k=shrK A))}" consts T:: "msg set" defs T_is_only_text: "T=={t. t: Atoms∧ (∀ k. t ≠ Key k)}" constdefs Is_K_strand::"sigma => bool " "Is_K_strand KS == (∃ k. k∈ KP ∧ (SP KS)=[(+, Key k)])" constdefs Is_T_strand::"sigma => bool" "Is_T_strand KS== (∃ t. t ∈ T ∧ (SP KS)=[(+, t)]) " constdefs Is_E_strand ::"sigma => bool " "Is_E_strand KS == (∃ k. ∃ h. (SP KS)=[(-, (Key k)),(-,h),(+, (Crypt k h))])" constdefs Is_D_strand::"sigma =>bool" "Is_D_strand KS== (∃ k. ∃ k'. ∃ h. k'=invKey k∧ (SP KS)=[(-, (Key k')),(-, (Crypt k h)),(+,h)])" constdefs Is_Cat_strand::"sigma =>bool" "Is_Cat_strand KS== (∃ g. ∃ h. (SP KS)=[(-, g),(-, h),(+, MPair g h)])" constdefs Is_Sep_strand::"sigma =>bool" "Is_Sep_strand KS== (∃ g. ∃ h. (SP KS)=[(-, MPair g h),(+, g),(+, h)])" constdefs Is_Flush_strand::"sigma =>bool" "Is_Flush_strand KS== (∃ g. (SP KS)=[(-, g )])" constdefs Is_Tee_strand::"sigma =>bool" "Is_Tee_strand KS== (∃ g. (SP KS)=[(-, g),(+, g),(+, g )])" consts Is_penetrator_strand:: "sigma => bool" defs penetrator_trace: "Is_penetrator_strand s== ( Is_Tee_strand s | Is_Flush_strand s | Is_Cat_strand s | Is_Sep_strand s | Is_E_strand s | Is_D_strand s | Is_T_strand s | Is_K_strand s)" constdefs Is_regular_strand :: "sigma =>bool" "Is_regular_strand s==¬ ( Is_penetrator_strand s)" constdefs nodes::"graph => node set" "nodes b== fst b" constdefs edges::"graph => edge set" "edges b==snd b" consts bundles ::" graph set" inductive "bundles" intros Nil: "( {} , {}) : bundles" Add_positive1: "[| b ∈ bundles; (node_sign n2) = +; n2 ∈ Domain; n2 ∉ (nodes b); 0 < index n2 ; n1 ∈ nodes b; n1=> n2 |]==> ({n2} Un (nodes b), {(n1, n2)} Un (edges b)) ∈ bundles" Add_positive2: "[| b ∈ bundles; node_sign n2=+; n2 ∉ (nodes b); n2 ∈ Domain; index n2=0 |]==> ({n2} Un nodes b, edges b) ∈ bundles" Add_negtive1: "[| b ∈ bundles; node_sign n2=-; n2 ∉ nodes b; ((strand n1 ≠ strand n2)∧ (n1 -> n2) ∧ (n1 ∈ nodes b) ∧ (∀ n3. ( (n3 ∈ nodes b)--> (n1,n3) ∉ edges b)) ); 0 < index n2 ; n1' ∈ nodes b; n1'=>n2 |]==> ({n2} Un nodes b, {(n1, n2), (n1' , n2)} Un edges b) ∈ bundles" Add_negtive2: "[| b ∈ bundles; node_sign n2=-; n2 ∉ nodes b; ((strand n1 ≠ strand n2)∧ (n1 -> n2) ∧ (n1 ∈ nodes b) ∧ (∀ n3. ( (n3 ∈ nodes b)--> (n1,n3) ∉ edges b)) ); index n2=0 |]==> ({n2} Un nodes b, {(n1, n2)} Un edges b) ∈ bundles" (*syntax "_edgebTrans"::" node=>node=>graph=>bool" ( "("_ \<prec> _ _ " ) " translations "n1\<prec>b n2 "=="(n1 ,n2):(edge b)^+" syntax "_edgebTransReflex"::" node=>node=>graph=>bool" (infix "\<preceq> b" 50) translations "n1\<preceq> b n2 "=="(n1 ,n2):(edge b)^*"*) constdefs parts_relation:: "msg=>msg=>bool"(infix "\<sqsubset> " 55) "parts_relation t1 t2==(t1 ∈ (Message.parts {t2}))" constdefs proper_subterm::"msg=>msg=>bool" "proper_subterm t1 t2 == t1 \<sqsubset> t2 ∧ t1≠t2" constdefs originate:: "msg=>node=>bool" "originate t n==(node_sign n)=+ ∧ t \<sqsubset> node_term n ∧ (∀ i. i < index n --> ¬ t \<sqsubset> node_term (strand n,i))" constdefs unique_originate:: "msg=>node=>bool" "unique_originate t n== originate t n ∧ (∀ n'. originate t n'--> n=n')" constdefs non_originate::" msg=>sigma=>bool" "non_originate t s==∀ n. (strand n=s--> (¬t \<sqsubset> node_term n| node_sign n=- | (∃ i. i< index n ∧ t \<sqsubset> node_term (strand n,i) )))" constdefs regularK::"nat => graph=>bool" "regularK k b== ∀ n∈ (nodes b).(node_term n=Key k )-->Is_regular_strand (strand n)" constdefs component::"msg => graph =>bool" "component t b== (∀ g h. t≠(MPair g h)) ∧ (∀ k h . (t=Crypt k h) --> (regularK (invKey k) b))" constdefs suite::"msg set =>msg=>graph =>bool" "suite M a b==(∀ t∈ M. (a \<sqsubset> t --> component t b)) ∧ (∀ t. (¬a\<sqsubset> t --> t∈ M))" constdefs path:: "msg=>node list=>graph=>bool" "path a p b== a \<sqsubset> node_term ( p ! (length p - 1)) ∧ ( ∀ i. (i <length p - 1 --> ( a \<sqsubset> node_term ( p ! i) ∧ ((p!i, p!(i+1)) ∈ (edges b Int casual1) ∨ ( (p!i, p!(i+1)) ∈ (edges b Int casual2)^+ )))) ) " constdefs orig_path:: "node list =>graph =>bool" "orig_path p b == p=[] | (∃ n. p=[n]∧ n ∈ (nodes b)) | (∀ i. (i <length p - 1 --> ( (p!i, p!(i+1)) ∈ (edges b Int casual1) ∨ (p!i, p!(i+1)) ∈ (edges b Int casual2)^+ ) ))" constdefs first_node_in_nonorigi_strand::"msg=>node=>node=>bool" "first_node_in_nonorigi_strand a n m== node_sign m=- ∧ a \<sqsubset> node_term m ∧ strand m=strand n ∧ (∀ i0< index m. ¬ a \<sqsubset> node_term (strand m , i0))" consts slice_arr_cons::"sigma=> nat =>nat => node list" primrec "slice_arr_cons s i (Suc len) = (s,i )# (slice_arr_cons s (Suc i) len)" "slice_arr_cons s j 0=[(s,j)]" consts complete_Path ::"msg => graph=>((node list) set)" inductive "complete_Path a b" intros (*uniqueOriginate: "unique_originate a m ==> [m] ∈ (complete_Path a b)"*) addPos1:"[| non_originate a (strand m2); node_sign m2=+; a\<sqsubset>node_term m2; first_node_in_nonorigi_strand a m2 m; p ∈ complete_Path a b ; last p=m |] ==>p@(slice_arr_cons (strand m) (index m) (index m2 - index m)) ∈ complete_Path a b" addPos2:"[| unique_originate a m; strand m2=strand m;a\<sqsubset>node_term m2; node_sign m2=+ |] ==> (slice_arr_cons (strand m) (index m) (index m2 - index m)) ∈ complete_Path a b" addnege: "[|m1->m2; m1 ∈ nodes b; (m1,m2) ∈ edges b; p ∈ complete_Path a b; a\<sqsubset>node_term m2; last p=m1|] ==>(p@[m2]) ∈ complete_Path a b" syntax "_casualBundleTrancl"::" node =>graph =>node=> bool" ("(_ \<prec>_ _)" [70,65,65] 61) translations "n1\<prec>b n2 "=="(n1 ,n2):(edges b)^+" syntax "_casualBundleTranclRefl"::" node =>graph =>node=> bool" ("(_ \<preceq>_ _)" [70,65,65] 61) translations "n1\<preceq>b n2 "=="(n1 ,n2):(edges b)^*" section{*basic lemmas*} subsection{*basic lemmas on causal relations*} lemma casual2_rel_more:" [|n1=>n2|]==>n2=(strand n1, (index n1) + 1)" apply (unfold strand_def index_def casual2_def) apply simp done lemma casual2_aux1 :" strand n=a ∧ index n=b ==>n=(a,b) " apply (unfold strand_def index_def,auto) done lemma casual2_rel_less:" [|n1=>n2|]==>n1=(strand n2, (snd n2) - 1)" apply (unfold strand_def index_def casual2_def) apply (simp split:nat_diff_split) apply auto apply (rule casual2_aux1) apply (unfold strand_def index_def,simp) done lemma casual2_trancl_less: "m=>+m'==> index m < index m' ∧strand m=strand m'" apply( erule trancl_induct) apply(unfold casual2_def,unfold index_def, auto) done lemma casual2_rtrancl1: "[|n2 ∈ Domain; i<=index n2|] ==> (strand n2,index n2 - i) ∈ Domain ∧ ((strand n2,index n2 - i)=>* n2)" (is " [|?A;?B i|]==>?P i") proof(unfold strand_def index_def , induct i) case 0 show ?case by auto next fix n assume IH:"[|n2 ∈ Domain; n ≤ snd n2|] ==> (fst n2, snd n2 - n) ∈ Domain ∧ ((fst n2, snd n2 - n)=>* n2) " and a1: " n2 ∈ Domain" and a2:"Suc n ≤ snd n2" show "(fst n2, snd n2 - Suc n) ∈ Domain ∧ ((fst n2, snd n2 - Suc n)=>* n2) " (is "?P' n") proof - from a2 have B4:" n ≤ snd n2 " by auto from this and IH and a1 have B5: "?P n" by(unfold strand_def index_def,auto) from B5 and a1 and a2 have B7:"(fst n2, snd n2 - Suc n) ∈ Domain∧((fst n2, snd n2 - Suc n), (fst n2, snd n2 -n)) ∈ casual2" apply- apply(unfold casual2_def, unfold Domain_def strand_def index_def) apply auto apply arith apply arith done from this and B5 show "?P' ( n)" by (unfold strand_def index_def, auto intro: converse_rtrancl_into_rtrancl) qed qed lemma casual2_rtrancl2:"[|n2 ∈ Domain; i<=index n2|] ==> (strand n2, i) ∈ Domain ∧ ((strand n2, i)=>* n2) " apply(unfold strand_def index_def) apply(subgoal_tac "snd n2 - i<=snd n2") prefer 2 apply simp apply(subgoal_tac "snd n2 - (snd n2 - i) =i") prefer 2 apply simp apply(fold strand_def index_def) apply(drule_tac i="snd n2 - i" in casual2_rtrancl1) apply(unfold strand_def index_def) apply auto done lemma casual2_trancl:"[|n2 ∈ Domain; i<snd n2|] ==> (fst n2, i) ∈ Domain ∧ ((fst n2, i) =>+n2)" thm casual2_rtrancl2 apply(drule_tac i="i" and ?n2.0="n2" in casual2_rtrancl2) apply(unfold strand_def index_def, simp) apply(erule conjE, drule rtranclD ) apply(erule disjE) prefer 2 apply blast apply(cut_tac surjective_pairing[where p="n2"]) apply(drule_tac s="n2" and t="(fst n2, snd n2)" and P="λ s.(fst n2, i) = s" in subst) apply simp apply( subgoal_tac "i=snd n2") apply simp apply(thin_tac "(fst n2, i) = n2") apply (simp only:Pair_eq) done lemma casual1_casual2_disjoint : "[|n1->n2 ∧ n3=>n4|] ==> n1≠n3 ∨ n2≠n4" apply(unfold casual1_def, unfold casual2_def) apply(auto ) done lemma send_node_is_plus: "[|(n1->n2)|]==> node_sign n1=+ ∧ n1 ∈ Domain" apply( unfold node_sign_def , unfold casual1_def, unfold Domain_def) apply(unfold node_sign_def node_term_def strand_def,auto ) done lemma recv_node_is_minus: "[|(n2->n1)|]==> node_sign n1=- ∧ n1 ∈ Domain" apply(unfold casual1_def, unfold node_sign_def, unfold Domain_def) apply( auto) done lemma casual2_rel_n1_n2_less:" [|n1=>n2|]==> (strand n1= strand n2) ∧ index n1= (index n2) - 1" apply (unfold casual2_def index_def strand_def) apply (simp split:nat_diff_split) done lemma casual_diff:" [|n1-> n2 | n1=>n2|] ==> n1≠n2" apply(unfold casual1_def, unfold casual2_def) apply( auto dest:recv_node_is_minus casual2_rel_less) done lemma is_casual2 : "[| m:Domain; n ∈ Domain; strand n= strand m ; index n= Suc (index m)|]==>m=>n" apply(unfold casual2_def ,unfold Domain_def) apply auto done lemma casual1_judge_sign : "[|n1 -> n2|] ==>strand n1 ≠strand n2" apply (unfold casual1_def, auto) done lemma casual2_judge_sign : "[|n1 => n2|] ==>strand n1 = strand n2" apply (unfold casual2_def, auto) done lemma casul2_trancl_sign: "[|(n1 , n2): casual2^+|] ==>strand n2 = strand n1 ∧ (index n1 < index n2)" apply(erule trancl_trans_induct) apply(unfold casual2_def) apply auto done lemma node_equal : " (n=n')= (strand n=strand n' ∧ index n=index n')" apply(unfold strand_def index_def) apply(simp add:Pair_fst_snd_eq) done lemma node_sign_is_Plus_or_minus:" ((n ::Sign)= + | n=-)" (*apply(rule allI)*) apply(case_tac n) apply blast+ done lemma rtrancl_mono1:" [|x:r^*; r⊆ s|]==> x ∈ s^*" apply(blast dest:rtrancl_mono) done lemma part_MPair:" [|(a≠{|g,h|}); a \<sqsubset> {|g,h|}|]==> a \<sqsubset> g | a \<sqsubset> h" apply(unfold parts_relation_def, simp) apply(simp only:parts_insert2) apply blast done subsection{*basic lemmas on subterm relation*} lemma subterm_atom:"[|a \<sqsubset> t; t ∈ Atoms |]==> a=t" apply(unfold Atoms_def) apply(simp, erule disjE) apply(unfold parts_relation_def) apply(erule exE,simp add:parts_insert_Agent) apply(erule disjE,erule exE,simp add:parts_insert_Number) apply(erule disjE,erule exE,simp add:parts_insert_Nonce) apply(erule exE,simp add:parts_insert_Key) done lemma subterm_trans:" [|g \<sqsubset> h; h \<sqsubset> l|] ==> g \<sqsubset> l" apply(unfold parts_relation_def) thm parts_mono apply( blast dest:parts_trans) done lemma subterm_itself:" h \<sqsubset> h" apply(unfold parts_relation_def) apply(rule_tac H="{h}" and X="h" in parts.Inj) apply simp done lemma body_subterm:"t'' \<sqsubset> Crypt k h ==> t'' = Crypt k h | t'' \<sqsubset> h" apply(unfold parts_relation_def) apply(erule parts.induct) apply blast apply(fold parts_relation_def) apply(subgoal_tac "X \<sqsubset> {|X,Y|}") prefer 2 apply(unfold parts_relation_def,blast intro:parts.Fst parts.Snd parts.Body,fold parts_relation_def) apply(erule disjE,simp, blast dest:subterm_trans) apply(subgoal_tac "Y \<sqsubset> {|X,Y|}") prefer 2 apply(unfold parts_relation_def,blast intro:parts.Fst parts.Snd parts.Body,fold parts_relation_def) apply(erule disjE,simp, blast dest:subterm_trans) apply(subgoal_tac "X \<sqsubset> Crypt K X") prefer 2 apply(unfold parts_relation_def,blast intro:parts.Fst parts.Snd parts.Body,fold parts_relation_def) apply(erule disjE,simp) prefer 2 apply( blast dest:subterm_trans) apply(blast intro:subterm_itself) done section{*the soundness of the inductive bundle model*} lemma bundle_node_in_Domain: assumes a1:"b ∈ bundles" shows c:"n ∈ nodes b --> n ∈ Domain" using a1 proof( induct) case Nil show ?case by (unfold nodes_def,auto) next fix b n1 n2 assume a1: " b ∈ bundles" and a2:" n ∈ nodes b --> n ∈ Domain" and a3:" node_sign n2 = +" and a4:" n2 ∈ Domain" and a5:"n2 ∉ nodes b" and a6:"0 < index n2" and a7:"n1 ∈ nodes b" and a8:" n1 => n2" show " n ∈ nodes ({n2} ∪ nodes b, {(n1, n2)} ∪ edges b) --> n ∈ Domain" proof - from a1 and a2 and a3 and a4 and a5 and a6 and a7 and a8 show ?thesis apply - apply(unfold nodes_def edges_def,rule impI) apply( simp only:fst_conv) apply simp apply blast done qed next fix b n1 n2 assume a1: " b ∈ bundles" and a2:" n ∈ nodes b --> n ∈ Domain" and a3:" node_sign n2 = +" and a4:" n2 ∈ Domain" and a5:"n2 ∉ nodes b" and a6:" index n2=0" show " n ∈ nodes ({n2} ∪ nodes b, edges b) --> n ∈ Domain" proof - from a1 and a2 and a3 and a4 and a5 and a6 show ?thesis apply - apply(unfold nodes_def edges_def,rule impI) apply( simp only:fst_conv) apply simp apply blast done qed next fix b n1 n1' n2 assume a1: " b ∈ bundles" and a2:" n ∈ nodes b --> n ∈ Domain" and a3:" node_sign n2 = -" and a5:"n2 ∉ nodes b" and a6:" 0<index n2" and a7:"strand n1 ≠ strand n2 ∧ n1 -> n2 ∧ n1 ∈ nodes b ∧ (∀n3. n3 ∈ nodes b --> (n1, n3) ∉ edges b)" and a8:"n1' ∈ nodes b" and a9:" n1' => n2" show " n ∈ nodes ({n2} ∪ nodes b, {(n1, n2), (n1', n2)} Un edges b) --> n ∈ Domain" proof - from a1 and a2 and a3 and a5 and a6 and a7 and a8 and a9 show ?thesis apply - apply(unfold nodes_def edges_def strand_def,rule impI) apply( simp only:fst_conv) apply((erule conjE)+, unfold casual1_def, simp) apply blast done qed next fix b n1 n2 assume a1: " b ∈ bundles" and a2:" n ∈ nodes b --> n ∈ Domain" and a3:" node_sign n2 = -" and a5:"n2 ∉ nodes b" and a6:" index n2=0" and a7:"strand n1 ≠ strand n2 ∧ n1 -> n2 ∧ n1 ∈ nodes b ∧ (∀n3. n3 ∈ nodes b --> (n1, n3) ∉ edges b)" show " n ∈ nodes ({n2} ∪ nodes b, {(n1, n2)} Un edges b) --> n ∈ Domain" proof - from a1 and a2 and a3 and a5 and a6 and a7 show ?thesis apply - apply(unfold nodes_def edges_def strand_def,rule impI) apply( simp only:fst_conv) apply((erule conjE)+, unfold casual1_def, simp) apply blast done qed qed lemma bundle_nodes_strand_in_space: assumes a1:"b ∈ bundles" and a2:"m ∈ nodes b" shows "strand m ∈ Σ" apply - apply(cut_tac a1 a2,drule_tac n="m" in bundle_node_in_Domain) apply(simp) apply(unfold Domain_def strand_def,auto) done inductive_cases pair_is_a_bundle [elim!]: "(a,b) ∈ bundles" lemma bundle_edge_is_exist: "[|b ∈ bundles |]==> ((node_sign n2=-) ∧ n2 ∈ (nodes b))-->(∃ n1. (n1->n2) ∧ n1 ∈ (nodes b) ∧(n1, n2) ∈ (edges b))" thm bundles.induct apply(erule bundles.induct, unfold nodes_def edges_def strand_def index_def) apply simp (*goal 1*) apply (rule impI) apply ( simp only:fst_conv snd_conv) apply(simp) apply(erule conjE) apply(erule disjE) apply simp apply simp apply(erule exE)+ thm exI apply (rule_tac ?x="a" in exI) apply (rule_tac ?x="ba" in exI) apply simp (*goal 2*) apply (rule impI) apply ( simp only:fst_conv snd_conv) apply(simp) apply(erule conjE) apply(erule disjE) apply simp apply simp apply(erule exE)+ thm exI apply (rule_tac ?x="a" in exI) apply (rule_tac ?x="ba" in exI) apply simp (*goal 3*) apply (rule impI) apply ( simp only:fst_conv snd_conv) apply(erule conjE)+ apply simp apply(erule disjE) apply (rule_tac ?x="fst n1" in exI) apply (rule_tac ?x="snd n1" in exI) apply simp apply simp thm exE apply(erule exE)+ thm exI apply (rule_tac ?x="a" in exI) apply (rule_tac ?x="ba" in exI) apply simp (*goal 4*) apply (rule impI) apply ( simp only:fst_conv snd_conv) apply(erule conjE)+ apply simp apply(erule disjE) apply (rule_tac ?x="fst n1" in exI) apply (rule_tac ?x="snd n1" in exI) apply simp apply simp thm exE apply(erule exE)+ thm exI apply (rule_tac ?x="a" in exI) apply (rule_tac ?x="ba" in exI) apply simp done lemma bundle_is_cl_under_edge: "[|b:bundles|]==>(n1,n2) ∈ edges b -->n1 ∈ nodes b ∧ n2 ∈ nodes b" apply(erule bundles.induct,unfold nodes_def edges_def) apply auto done lemma bundle_is_cl_under_edge1: "[|b ∈ bundles;(n1,n2) ∈ edges b|]==>n1 ∈ nodes b ∧ n2 ∈ nodes b" apply(auto dest!: bundle_is_cl_under_edge) done lemma bundle_is_not_close_edge: "[|b ∈ bundles|]==>n1:nodes b ∧ n2∉nodes b -->(n1, n2)∉edges b" apply(rule impI) apply(rule notI) apply(auto dest: bundle_is_cl_under_edge) done lemma bundle_is_not_close_edge1: "[|b ∈ bundles;n1 ∈ nodes b ; n2∉nodes b|]==>(n1, n2)∉edges b" apply(auto dest: bundle_is_not_close_edge) done lemma bundle_is_closed:"[|b ∈ bundles|]==> ((y, z) ∈ edges b)-->y ∈ nodes b ∧ z ∈ nodes b" apply(erule bundles.induct, unfold nodes_def edges_def) apply (auto dest:send_node_is_plus recv_node_is_minus ) done lemma bundle_is_closed1:"[|b ∈ bundles|]==> ((y, z) ∈ edges b)-->y ≠ z" apply(erule bundles.induct,unfold edges_def) apply(auto) done lemma bundle_casual1_prec_is_unique: assumes A:"b ∈ bundles " shows "(node_sign n2=-) ∧ n1->n2 ∧ n1' -> n2∧ (n1, n2) ∈ (edges b)∧(n1', n2) ∈ (edges b)-->n1=n1'" proof - have Goal0:"node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ edges b ∧ (n1', n2) ∈ edges b ∧ n2 ∈ nodes b --> n1 = n1'" using A proof( induct ) case Nil show ?case by(unfold nodes_def, auto) next fix "b" "n1a" "n2a" assume "b ∈ bundles" and b:" node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ edges b ∧ (n1', n2) ∈ edges b ∧ n2 ∈ nodes b --> n1 = n1'" and a:" node_sign n2a = +" and " n2a ∈ Domain" and " n2a ∉nodes b" and " 0 < index n2a" and " n1a ∈nodes b" and " n1a => n2a" show "node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ edges ({n2a} ∪nodes b, {(n1a, n2a)} ∪ edges b) ∧ (n1', n2) ∈ edges ({n2a} ∪nodes b, {(n1a, n2a)} ∪ edges b) ∧ n2 ∈ nodes ({n2a} ∪nodes b, {(n1a, n2a)} ∪ edges b) --> n1 = n1'" proof (unfold edges_def nodes_def,rule impI) assume c0:"node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ snd ({n2a} ∪ fst b, {(n1a, n2a)} ∪ snd b) ∧ (n1', n2) ∈ snd ({n2a} ∪ fst b, {(n1a, n2a)} ∪ snd b) ∧ n2 ∈ fst ({n2a} ∪ fst b, {(n1a, n2a)} ∪ snd b)" show "n1=n1'" proof (cases "n2=n2a") case True assume c:"n2=n2a" show ?thesis proof (rule FalseE) from a and c and c0 show False by auto (*from a and c have d:"node_sign n2=+ " by simp from c0 have e: "node_sign n2= - " by simp from e and d show False by simp*) qed next case False assume c:"n2 ≠n2a" show "n1=n1'" proof - from c0 have "n2 ∈ fst ({n2a} ∪ fst b, {(n1a, n2a)} ∪ snd b)" by simp then have f:"n2 : {n2a} Un fst b" by simp from c and f have g0':" n2 ∈ fst b" by simp from c0 have g:"(n1, n2) ∈ snd ({n2a} ∪ fst b, {(n1a, n2a)} ∪ snd b)" by simp from g and c have g0:"(n1, n2) ∈ snd b" by simp from c0 and c have g':"(n1', n2) ∈ snd b" by simp from c0 and g0 and g' and g0' and b show ?thesis by (unfold edges_def nodes_def,simp) qed qed qed next fix "b" "n1a" "n2a" assume "b ∈ bundles" and b:" node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ edges b ∧ (n1', n2) ∈ edges b ∧ n2 ∈ nodes b --> n1 = n1'" and a:" node_sign n2a = +" and " n2a ∈ Domain" and " n2a ∉ nodes b" and " index n2a =0" show "node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ edges ({n2a} ∪ nodes b, edges b) ∧ (n1', n2) ∈ edges ({n2a} ∪ nodes b, edges b) ∧ n2 ∈ nodes ({n2a} ∪ nodes b, edges b) --> n1 = n1'" proof(unfold edges_def nodes_def,rule impI) assume c0:"node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ snd ({n2a} ∪ fst b,snd b) ∧ (n1', n2) ∈ snd ({n2a} ∪ fst b,snd b) ∧ n2 ∈ fst ({n2a} ∪ fst b,snd b)" show "n1=n1'" proof (cases "n2=n2a") case True assume c:"n2=n2a" show ?thesis proof (rule FalseE) from a and c and c0 show False by auto (*from a and c have d:"node_sign n2=+ " by simp from c0 have e: "node_sign n2= - " by simp from e and d show False by simp*) qed next case False assume c:"n2 ≠n2a" show "n1=n1'" proof - from c0 have "n2 ∈ fst ({n2a} ∪ fst b, snd b)" by simp then have f:"n2 : {n2a} Un fst b" by simp from c and f have g0':" n2:fst b" by simp from c0 have g:"(n1, n2) ∈ snd b" by simp from g and c have g0:"(n1, n2) ∈snd b" by simp from c0 and c have g':"(n1', n2) ∈ snd b" by simp from c0 and g and g0' and b show ?thesis by(unfold nodes_def edges_def, simp) qed qed qed next fix "b" "n1a" "n1'a" "n2a" assume bb:"b ∈ bundles" and IH:" node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ edges b ∧ (n1', n2) ∈ edges b ∧ n2 ∈ nodes b --> n1 = n1'" and " node_sign n2a = -" and as0:"n2a ∉ nodes b" and n1ab: "strand n1a ≠ strand n2a ∧ n1a -> n2a ∧ n1a ∈ nodes b ∧ (∀n3. n3 ∈ nodes b --> (n1a, n3) ∉ edges b)" and " 0 < index n2a" and " n1'a ∈ nodes b" and n1'ab:"n1'a => n2a" show "(node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ edges ({n2a} ∪ nodes b, {(n1a, n2a), (n1'a, n2a)} ∪ edges b) ∧ (n1', n2) ∈ edges ({n2a} ∪ nodes b, {(n1a, n2a), (n1'a, n2a)} ∪ edges b) ∧ n2 ∈ nodes ({n2a} ∪ nodes b, {(n1a, n2a), (n1'a, n2a)} ∪ edges b) --> n1 = n1')" proof(rule impI, unfold nodes_def edges_def index_def strand_def) from n1ab have ass0: "n1a->n2a" by simp assume as1: " node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ snd ({n2a} ∪ fst b, {(n1a, n2a), (n1'a, n2a)} ∪ snd b) ∧ (n1', n2) ∈ snd ({n2a} ∪ fst b, {(n1a, n2a), (n1'a, n2a)} ∪ snd b) ∧ n2 ∈ fst ({n2a} ∪ fst b, {(n1a, n2a), (n1'a, n2a)} ∪ snd b)" show "n1=n1'" proof (cases "n2=n2a") case True from this have ascase: "n2=n2a" . from as1 have as4:"(n1, n2) ∈ {(n1a, n2a), (n1'a, n2a)} ∪ snd b" by simp from as1 have as5:"(n1', n2) ∈ {(n1a, n2a), (n1'a, n2a)} ∪ snd b" by simp from bb and as0 and ascase and as4 have as6:"(n1, n2) ∈ {(n1a, n2a), (n1'a, n2a)}" apply simp apply(erule disjE,simp)+ thm bundle_is_closed apply(fold nodes_def edges_def, drule_tac y="n1 " and z="n2a" in bundle_is_closed) apply simp done from as1 have as7: "n1 -> n2" by simp from as6 and as0 and n1'ab and as7 have as8:"n1=n1a" apply simp apply(erule disjE) apply simp apply(blast dest:casual1_judge_sign casual2_judge_sign) done from bb and as0 and ascase and as5 have as6:"(n1', n2) ∈ {(n1a, n2a), (n1'a, n2a)}" apply simp apply(erule disjE,simp)+ apply(fold nodes_def edges_def,drule_tac y="n1' " and z="n2a" in bundle_is_closed) apply simp done from as1 have as9: "n1'-> n2" by simp from as6 and as0 and n1'ab and as9 have as10: "n1'=n1a" apply simp apply(erule disjE) apply simp apply(blast dest:casual1_judge_sign casual2_judge_sign) done from as8 and as10 show ?thesis by simp next case False assume c:"n2 ≠n2a" show "n1=n1'" proof - from as1 have "n2 ∈ fst ({n2a} ∪ fst b, snd b)" by simp then have f:"n2 ∈ {n2a} Un fst b" by simp from c and f have g0':" n2 ∈ fst b" by simp from as1 and c have g:"(n1, n2) ∈ snd b" by simp from as1 and c have g':"(n1', n2) ∈ snd b" by simp from as1 and g and g' and g0' and IH show ?thesis by (unfold nodes_def edges_def,simp) qed qed qed next fix "b" "n1a" "n2a" assume bb:"b ∈ bundles" and IH:" node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ edges b ∧ (n1', n2) ∈ edges b ∧ n2 ∈ nodes b --> n1 = n1'" and " node_sign n2a = -" and as0:"n2a ∉ nodes b" and n1ab: "strand n1a ≠ strand n2a ∧ n1a -> n2a ∧ n1a ∈ nodes b ∧ (∀n3. n3 ∈ nodes b --> (n1a, n3) ∉ edges b)" and "index n2a=0" show " node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ edges ({n2a} ∪ nodes b, {(n1a, n2a)} ∪ edges b) ∧ (n1', n2) ∈edges ({n2a} ∪ nodes b, {(n1a, n2a)} ∪ edges b) ∧ n2 ∈ nodes ({n2a} ∪ nodes b, {(n1a, n2a)} ∪ edges b) --> n1 = n1'" proof(rule impI, unfold nodes_def edges_def index_def strand_def) from n1ab have ass0: "n1a->n2a" by simp assume as1: "node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ snd ({n2a} ∪ fst b, {(n1a, n2a)} ∪ snd b) ∧ (n1', n2) ∈ snd ({n2a} ∪ fst b, {(n1a, n2a)} ∪ snd b) ∧ n2 ∈ fst ({n2a} ∪ fst b, {(n1a, n2a)} ∪ snd b)" show "n1=n1'" proof (cases "n2=n2a") case True from this have ascase: "n2=n2a" . from as1 have as4:"(n1, n2) ∈ {(n1a, n2a)} ∪ snd b" by simp from as1 have as5:"(n1', n2) ∈ {(n1a, n2a)} ∪ snd b" by simp from bb and as0 and ascase and as4 have as6:"(n1, n2) ∈ {(n1a, n2a)}" apply simp apply(erule disjE,simp)+ apply(fold nodes_def edges_def) apply(drule_tac y="n1 " and z="n2a" in bundle_is_closed) apply simp done from as1 have as7: "n1 -> n2" by simp from as6 and as0 and as7 have as8:"n1=n1a" apply simp done from bb and as0 and ascase and as5 have as6:"(n1', n2) ∈ {(n1a, n2a)}" apply simp apply(erule disjE,simp)+ apply(fold nodes_def edges_def) apply(drule_tac y="n1' " and z="n2a" in bundle_is_closed) apply simp done from as1 have as9: "n1'-> n2" by simp from as6 and as0 and as9 have as10: "n1'=n1a" apply simp done from as8 and as10 show ?thesis by simp next case False assume c:"n2 ≠n2a" show "n1=n1'" proof - from as1 have "n2 ∈ fst ({n2a} ∪ fst b, snd b)" by simp then have f:"n2 : {n2a} Un fst b" by simp from c and f have g0':" n2 ∈ fst b" by simp from as1 and c have g:"(n1, n2) ∈ snd b" by simp from as1 and c have g':"(n1', n2) ∈ snd b" by simp from as1 and g and g' and g0' and IH show ?thesis by (unfold nodes_def edges_def,simp) qed qed qed qed show " node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ edges b ∧ (n1', n2) ∈ edges b --> n1 = n1'" proof(rule impI,unfold edges_def nodes_def) assume b:"node_sign n2 = - ∧ n1 -> n2 ∧ n1' -> n2 ∧ (n1, n2) ∈ snd b ∧ (n1', n2) ∈ snd b" show "n1=n1'" proof - from b have "(n1', n2) ∈ snd b" by simp from A and this have "n2 ∈ fst b" thm bundle_is_cl_under_edge1 apply(fold edges_def nodes_def) apply(drule_tac b="b" and ?n1.0="n1'" and ?n2.0="n2" in bundle_is_cl_under_edge1) apply auto done from this and Goal0 and b show ?thesis by (unfold nodes_def edges_def,auto) qed qed qed lemma bundle_casual2_prec_is_unique: "[|b: bundles |]==> ( n1=>n2 ∧ n2 ∈ nodes b )-->((n1,n2) ∈ (edges b) ∧ n1 ∈ nodes b)" apply(erule bundles.induct,unfold nodes_def edges_def ) apply simp apply(rule impI, simp only:fst_conv snd_conv) apply(erule conjE)+ apply(subgoal_tac "n2 ∈ {n2a} | n2 ∈ fst b") prefer 2 apply blast apply(erule disjE) apply( simp) apply(subgoal_tac "n1=n1a") apply simp apply(unfold casual2_def strand_def index_def,simp add:node_equal) apply(unfold strand_def index_def,simp) apply blast apply(rule impI, simp only:fst_conv snd_conv ) apply(subgoal_tac "n2a≠n2") apply blast apply(rule ccontr) apply simp apply(rule impI, simp only:fst_conv snd_conv) apply(erule conjE)+ apply(subgoal_tac "n2 ∈ {n2a} | n2 ∈ fst b") prefer 2 apply blast apply(erule disjE) apply( simp) apply(subgoal_tac "n1=n1'") apply simp apply(simp add:node_equal) apply(unfold strand_def index_def,simp) apply blast apply(rule impI, simp only:fst_conv snd_conv) apply(subgoal_tac "n2a≠n2") apply blast apply(rule ccontr) apply simp done lemma bundle_node_is_in_Domains:"[|b ∈ bundles|]==>n ∈ nodes b-->n ∈ Domain" apply(erule bundles.induct) apply(unfold nodes_def casual1_def) apply auto done lemma bundle_casual2_prec_is_unique_intro1: "[|b ∈ bundles ; n1=>n2 ; n2 ∈ nodes b|]==> n1 ∈ (nodes b) ∧ n1 ∈ Domain " apply(subgoal_tac " n1 ∈ nodes b") apply simp apply(blast dest:bundle_node_is_in_Domains) apply(blast dest:bundle_casual2_prec_is_unique ) done lemma bundle_casual2_closed_aux1: "[|0< index n2 - n; n2 ∈ Domain|] ==> (strand n2, (index n2) - Suc n) =>(strand n2, (index n2) - n)" apply(unfold index_def strand_def casual2_def Domain_def) apply (auto ) apply arith+ done lemma bundle_casual2_closed: "[|b ∈ bundles;n2 ∈ nodes b; i≤ index n2|]==> (strand n2, index n2 - i) ∈ nodes b ∧ (strand n2, index n2 - i) ∈ Domain " apply(induct i) thm bundle_node_is_in_Domains apply(unfold index_def strand_def) apply(force dest: bundle_node_is_in_Domains) apply(fold index_def strand_def) thm bundle_casual2_prec_is_unique_intro1 apply(rule_tac ?b="b" and ?n1.0="(strand n2, index n2 - Suc i)" and ?n2.0="(strand n2, index n2 - i)" in bundle_casual2_prec_is_unique_intro1) apply simp apply(rule bundle_casual2_closed_aux1) apply arith apply(force dest: bundle_node_is_in_Domains) apply simp done lemma bundle_casual2_closed2: "[|b ∈ bundles; n2 ∈ nodes b; i≤index n2|]==> ( (strand n2, i) ∈ nodes b)∧ (strand n2, i) ∈ Domain" apply(frule_tac i="index n2 - i" in bundle_casual2_closed) apply simp+ done lemma bundle_casual2_closed222: shows "[|b ∈ bundles; n2 ∈ nodes b; i≤index n2|] ==>((strand n2,index n2 - i), n2) ∈ (edges b ∩ casual2)*" (is "[|?A1; ?A2;?A3|] ==> ?P i") proof(induct i ) show "?P 0" by(unfold nodes_def strand_def edges_def index_def,simp) next fix n assume IH:"[|b ∈ bundles; n2 ∈ nodes b; n ≤ index n2|] ==>?P n" and B1:"b ∈ bundles" and B2:" n2 ∈ nodes b" and B3:" Suc n ≤ index n2" show "?P(Suc n)" proof - from B3 have B4:" n ≤ index n2 " by auto from this and IH and B1 and B2 have B5: "?P n" by auto from B2 and this and B1 have B6: "(strand n2, index n2 - n) ∈ nodes b ∧ (strand n2, index n2 - n) ∈ Domain " apply - apply(drule_tac i="index n2 - n" and ?n2.0="n2" and b="b" in bundle_casual2_closed2) apply auto done from B3 and B4 and B6 have B7:"((strand n2, index n2 - Suc n), (strand n2, index n2 -n)) ∈ casual2" apply- apply(unfold casual2_def, unfold Domain_def) apply auto apply arith apply(unfold strand_def,simp) apply(unfold index_def) apply simp apply arith done from this and B1 and B6 have "((strand n2, index n2 - Suc n), (strand n2, index n2 -n)) ∈ edges b" by (auto dest:bundle_casual2_prec_is_unique) from this and B7 have B8:"((strand n2, index n2 - Suc n), (strand n2, index n2 -n)) ∈ (edges b Int casual2)" by auto from B5 and this show "?P (Suc n)" by ( auto intro: converse_rtrancl_into_rtrancl) qed qed lemma bundle_casual2_closed3: assumes A1 : "b ∈ bundles" and A2:" n2 ∈ nodes b" and A3:" i≤index n2" (is "?A i") shows "(strand n2, i) ∈ nodes b ∧ ( (strand n2, i) ,n2): (edges b Int casual2)^*" proof(rule conjI) from A1 and A2 and A3 show " (strand n2, i) ∈ nodes b" apply - apply(blast dest:bundle_casual2_closed2) done next from A3 have C2:"index n2 - i <= index n2" by(unfold index_def, arith) from A1 and A2 and A3 and C2 show "((strand n2, i), n2) ∈ (edges b ∩ casual2)*" apply - thm bundle_casual2_closed222 apply(drule_tac b="b" and ?n2.0="n2" and i="index n2 - i" in bundle_casual2_closed222) apply simp apply simp apply auto done qed lemma bundle_casual2_trancl: " [|b ∈ bundles; n' ∈ (nodes b); (m=>+ n')|]==>m ∈ (nodes b) ∧(m\<prec>b n')" apply - apply(subgoal_tac "index m < index n' ∧strand m=strand n'") prefer 2 apply(blast dest:casual2_trancl_less) apply(subgoal_tac "index m <= index n' ") prefer 2 apply simp apply(subgoal_tac "(strand n', index m) ∈ nodes b ∧ ( (strand n', index m) ,n') ∈ (edges b Int casual2)^*") prefer 2 apply(blast dest:bundle_casual2_closed3) apply(subgoal_tac "(strand n', index m)=m") apply simp apply(subgoal_tac "(m, n') ∈ (edges b)^*") apply(drule rtranclD) apply(erule disjE) apply simp apply simp apply(subgoal_tac "(edges b ∩ casual2)* ⊆ (edges b)*") apply blast apply(blast intro!:rtrancl_mono) apply(simp add:node_equal) apply(unfold strand_def index_def,auto) done lemma bundle_edge_is_unique: "[|b ∈ bundles |]==> ((node_sign n2=-) ∧ n2 ∈ (nodes b))--> (∃! n1. (n1->n2) ∧ n1 ∈ (nodes b) ∧(n1, n2) ∈ (edges b))" thm bundles.induct (*apply(erule bundles.induct)*) apply(blast dest: bundle_edge_is_exist bundle_casual1_prec_is_unique ) done lemma bundle_is_trans_closed2: "[|b ∈ bundles |]==>((y, z) ∈ (edges b)^+)-->y ∈ nodes b ∧ z ∈ nodes b" apply(rule impI) (*apply(erule rtrancl_induct)*) apply( erule trancl_trans_induct) apply (auto dest:bundle_is_closed bundle_is_closed1 ) done lemma bundle_is_trans_closed22: "[|b ∈ bundles; y \<prec>b z |]==>y ∈ nodes b ∧ z ∈ nodes b" apply (auto dest:bundle_is_trans_closed2) done lemma bundle_is_rtrans_closed23: "[|b ∈ bundles; z ∈ nodes b; (y \<preceq>b z)|]==>y ∈ nodes b ∧ z ∈ nodes b" apply(drule rtranclD) apply(auto dest: bundle_is_trans_closed22) done lemma bundle_is_rtrans_closed24: "[|b ∈ bundles; z∉ nodes b;(y \<preceq>b z) |]==>y = z" apply(drule rtranclD) apply(auto dest: bundle_is_trans_closed22) done thm trancl_insert lemma bundle_edge_is_anti2: assumes A:"b ∈ bundles" shows "y \<prec>b z--> y≠z" using A proof induct case Nil show ?case by(unfold edges_def, auto) next fix b n1 n2 assume a1:"b ∈ bundles" and IH:" (y, z) ∈ (edges b)+ --> y ≠ z" and a2:" node_sign n2 = +" and a3:" n2 ∈ Domain" and a4:" n2 ∉ nodes b" and a5:" 0 < index n2" and a6:" n1 ∈ nodes b" and a7:" n1 => n2" show "(y, z) ∈ (edges ({n2} ∪ nodes b, {(n1, n2)} ∪ edges b))+ --> y≠z" proof assume a8:"(y, z) ∈ (edges ({n2} ∪ nodes b, {(n1, n2)} ∪ edges b))+" show "y≠z" proof(cases "z=n2") case True from this have casehyp:"z=n2" . from a8 have a9:"(y,z):({(n1, n2)} ∪ edges b)+" by(unfold edges_def, simp) from a9 have a10:"(y,z) ∈ ((edges b)^+ Un {(a,c). (a,n1) ∈ (edges b)^* ∧ ((n2,c) ∈ (edges b)^*)})" by (simp add:trancl_insert) from a10 have a12:"(y,z) ∈ (edges b)^+ Un {(a,n2). (a,n1) ∈ (edges b)^* }" proof - show "(y, z) ∈ (edges b)+ ∪ {(a, c). (a, n1) ∈ (edges b)* ∧ (n2, c) ∈ (edges b)*} ==> (y, z) ∈ (edges b)+ ∪ {(a, n2). (a, n1) ∈ (edges b)*}" apply simp apply(erule disjE) apply simp apply simp done qed from a12 and a6 and a4 and IH and casehyp and A show ?thesis proof - show " [|(y, z) ∈ (edges b)+ ∪ {(a, n2). (a, n1) ∈ (edges b)*}; n1 ∈ nodes b; n2 ∉ nodes b; (y, z) ∈ (edges b)+ --> y ≠ z; z = n2; b ∈ bundles|] ==> y ≠ z" apply simp apply(erule disjE) apply simp apply(frule_tac b="b" and y="y" and z="n1" in bundle_is_rtrans_closed23) apply auto done qed next case False from this have casehyp:"z≠n2" . from a8 have a9:"(y,z) ∈ ({(n1, n2)} ∪ edges b)+" by (unfold edges_def,simp) from A and a8 and casehyp and a4 have a10: "(y,z) ∈ (edges b)+" proof - show "[|b ∈ bundles; z ≠ n2; (y, z) ∈ (edges ({n2} ∪ nodes b, {(n1, n2)} ∪ edges b))+; n2 ∉ nodes b|] ==> (y, z) ∈ (edges b)+ " apply (simp add:edges_def trancl_insert) apply(erule disjE) apply simp apply(erule conjE) apply(drule_tac a="n2" and b="z" in rtranclD) apply(fold edges_def) apply(auto dest:bundle_is_trans_closed22) done qed from a10 and IH show ?thesis by simp qed qed next fix b n2 assume a1: "b ∈ bundles" and a2: "(y, z) ∈ (edges b)+ --> y ≠ z" and a3:" node_sign n2 = +" and a4:" n2 ∉ nodes b" and a5:" index n2 = 0" show "(y, z) ∈ (edges ({n2} ∪ nodes b, edges b))+ --> y ≠ z" proof assume a6:"(y, z) ∈ (edges ({n2} ∪ nodes b, edges b))+ " show "y≠z" proof - from a6 have a7:"(y, z) ∈ (edges b)^+" by(unfold edges_def, auto) from a2 and a6 show ?thesis by(unfold edges_def, simp) qed qed next fix b n1 n1' n2 assume a1:"b ∈ bundles" and IH:" (y, z) ∈ (edges b)+ --> y ≠ z" and a2:" node_sign n2 = -" and a3:" strand n1 ≠ strand n2 ∧ n1 -> n2 ∧ n1 ∈ nodes b ∧ (∀n3. n3 ∈ nodes b --> (n1, n3) ∉ edges b)" and a4:" n2 ∉ nodes b" and a5:" 0 < index n2" and a6:" n1' ∈ nodes b" and a7:" n1' => n2" show "(y, z) ∈ (edges ({n2} ∪ nodes b, {(n1, n2), (n1', n2)} ∪ edges b))+ --> y ≠ z" proof assume a8:"(y, z) ∈ (edges ({n2} ∪ nodes b, {(n1, n2),(n1',n2)} ∪ edges b))+" show "y≠z" proof(cases "z=n2") case True from this have casehyp:"z=n2" . from a8 have a9:"(y,z) ∈ ({(n1, n2),(n1',n2)} ∪ edges b)+" by(unfold edges_def,simp) from a9 have a10:"(y,z) ∈ (({(n1',n2)} Un edges b)^+ Un {(a,c). (a,n1) ∈ ({(n1',n2)} Un edges b)^* ∧ ((n2,c) ∈ ({(n1',n2)} Un edges b)^*)})" by (simp add:edges_def trancl_insert) from a10 have a12:"(y,z) ∈ (edges b)^+ Un {(a,n2). (a,n1) ∈ (edges b)^* | (a,n1') ∈ (edges b)^* }" proof - show "(y, z) ∈ ({(n1', n2)} ∪ edges b)+ ∪ {(a, c). (a, n1) ∈ ({(n1', n2)} ∪ edges b)* ∧ (n2, c) ∈ ({(n1', n2)} ∪ edges b)*} ==> (y, z) ∈ (edges b)+ ∪ {(a, n2). (a, n1) ∈ (edges b)* ∨ (a, n1') ∈ (edges b)*}" apply simp apply(erule disjE) apply(simp add:trancl_insert) apply blast apply(erule conjE) apply(drule rtranclD) apply(drule rtranclD) apply(simp add:trancl_insert) apply auto done qed from a12 and a6 and a3 and a4 and IH and casehyp and A show ?thesis proof - show " [|(y, z) ∈ (edges b)+ ∪ {(a, n2). (a, n1) ∈ (edges b)* ∨ (a, n1') ∈ (edges b)*}; n1' ∈ nodes b; strand n1 ≠ strand n2 ∧ n1 -> n2 ∧ n1 ∈ nodes b ∧ (∀n3. n3 ∈ nodes b --> (n1, n3) ∉ edges b); n2 ∉ nodes b; (y, z) ∈ (edges b)+ --> y ≠ z; z = n2; b ∈ bundles|] ==> y ≠ z" apply simp apply(erule conjE)+ apply(erule disjE) apply simp apply(auto dest:bundle_is_rtrans_closed23) done qed next case False from this have casehyp:"z≠n2" . from a8 have a9:"(y,z) ∈ ({(n1, n2),(n1',n2)} ∪ edges b)+" by(unfold edges_def, simp) from A and a8 and casehyp and a4 have a10: "(y,z) ∈ (edges b)+" proof - show "[|b ∈ bundles; (y, z) ∈ (edges ({n2} ∪ nodes b, {(n1, n2), (n1', n2)} ∪ edges b))+; z ≠ n2; n2 ∉ nodes b|] ==> (y, z) ∈ (edges b)^+" apply - apply (simp add:edges_def trancl_insert) apply(erule disjE) apply simp apply(erule disjE, erule conjE) apply(drule rtranclD)+ apply(fold edges_def) apply(auto dest:bundle_is_trans_closed22) apply(drule rtranclD)+ apply(simp add:trancl_insert)+ apply(auto dest:rtranclD bundle_is_trans_closed22) done qed from a10 and IH show ?thesis by simp qed qed next fix b n1 n2 assume a1:"b ∈ bundles" and IH:"(y, z) ∈ (edges b)+ --> y ≠ z" and a2:" node_sign n2 = -" and a3:" strand n1 ≠ strand n2 ∧ n1 -> n2 ∧ n1 ∈ nodes b ∧ (∀n3. n3 ∈ nodes b --> (n1, n3) ∉ edges b)" and a4:" n2 ∉ nodes b" and a5:"index n2=0" show "(y, z) ∈ (edges ({n2} ∪ nodes b, {(n1, n2)} ∪ edges b))+ --> y ≠ z" proof assume a8:"(y, z) ∈ (edges ({n2} ∪ nodes b, {(n1, n2)} ∪ edges b))+" show "y≠z" proof(cases "z=n2") case True from this have casehyp:"z=n2" . from a8 have a9:"(y,z):({(n1, n2)} ∪ edges b)+" by(unfold edges_def, simp) from a9 have a12:"(y,z): (edges b)^+ Un {(a,n2). (a,n1) ∈ (edges b)^* }" proof - show "(y, z) ∈ ({(n1, n2)} ∪ edges b)+ ==> (y, z) ∈ (edges b)+ ∪ {(a, n2). (a, n1) ∈ (edges b)*}" apply(simp add:edges_def trancl_insert) apply blast done qed from a12 and a3 and a4 and IH and casehyp and A show ?thesis proof - show " [|(y, z) ∈ (edges b)+ ∪ {(a, n2). (a, n1) ∈ (edges b)*}; strand n1 ≠ strand n2 ∧ n1 -> n2 ∧ n1 ∈ nodes b ∧ (∀n3. n3 ∈ nodes b --> (n1, n3) ∉ edges b); n2 ∉ nodes b; (y, z) ∈ (edges b)+ --> y ≠ z; z = n2; b ∈ bundles|] ==> y ≠ z" apply simp apply(erule conjE)+ apply(erule disjE) apply simp apply(auto dest:bundle_is_rtrans_closed23) done qed next case False from this have casehyp:"z≠n2" . from a8 have a9:"(y,z) ∈ ({(n1, n2)} ∪ edges b)+" by(unfold edges_def, simp) from A and a8 and casehyp and a4 have a10: "(y,z): (edges b)+" proof - show "[|b ∈ bundles; (y, z) ∈ (edges ({n2} ∪ nodes b, {(n1, n2)} ∪ edges b))+; z ≠ n2; n2 ∉ nodes b|] ==> (y, z) ∈ (edges b)^+" apply (simp add:edges_def trancl_insert) apply(erule disjE) apply simp apply(fold edges_def) apply(auto dest:rtranclD bundle_is_trans_closed22) done qed from a10 and IH show ?thesis by simp qed qed qed lemma bundle_is_wf:"[|b ∈ bundles |]==> wf (edges b)" apply(rule finite_acyclic_wf) apply(erule bundles.induct) apply(unfold edges_def, simp+) apply( unfold acyclic_def,fold edges_def,auto dest:bundle_edge_is_anti2) done lemma bundle_trancl_wf: "[|b ∈ bundles |]==> wf ((edges b)^+)" apply(blast dest: bundle_is_wf wf_trancl) done lemma bundle_minimal: "[|b ∈ bundles; x ∈ Q|] ==>∃z∈Q. (∀y. (y\<prec>b z) --> y ∉ Q)" apply(frule bundle_trancl_wf) apply( simp only:wf_eq_minimal) done section{*lemmas on origination assumption*} lemma non_originate_node:"[|unique_originate a n; n≠n'|]==> strand n= strand n' ∧ (a \<sqsubset> node_term n' --> index n < index n') | strand n≠ strand n' ∧ non_originate a (strand n')" apply(unfold unique_originate_def, unfold originate_def) apply(case_tac "strand n=strand n'") apply(erule conjE)+ apply (case_tac "(index n') < (index n)") apply(drule_tac ?x= " n'" in spec) apply(drule_tac ?x="index n'" in spec)apply(unfold strand_def index_def, simp) apply(case_tac "index n= index n'") apply (unfold index_def,simp ) apply(subgoal_tac "n=n'") apply blast apply(simp add:node_equal strand_def index_def) apply simp apply(unfold non_originate_def strand_def index_def) apply (rule disjI2) apply(rule conjI) apply simp apply(rule allI) apply(rule impI) apply(rule ccontr) apply(erule conjE)+ apply(drule_tac x="na" in spec) apply( auto ) apply(cut_tac n="node_sign (fst n', b)" in node_sign_is_Plus_or_minus) apply simp done lemma arith_aux_non_originate1:"[| (∃ (i::nat). P i)|]==> ∃ i. (P i ∧ (∀ j<i. ¬ P j))" apply(erule_tac exE) thm exI apply(rule_tac ?x=" Least P " in exI) apply(rule conjI) apply(rule LeastI) apply simp apply (rule allI) apply (auto dest: not_less_Least) done lemma non_originate_strand:" non_originate a s ==> (∃ i. node_sign (s,i)=- ∧ a \<sqsubset> node_term (s,i) ∧ (∀ i0< i. ¬ a \<sqsubset> node_term (s , i0))) | (∀ i. ¬ a \<sqsubset> node_term (s,i))" apply(unfold non_originate_def) apply( case_tac "∃i. a \<sqsubset> node_term (s,i) ") prefer 2 apply simp apply(drule arith_aux_non_originate1) apply(erule exE) apply(drule_tac ?x="(s,i)" in spec) apply(unfold strand_def index_def, simp add:fst_conv) apply(erule disjE) apply blast apply auto done lemma non_originate_node1: "[|unique_originate a n; strand n≠strand n';a\<sqsubset> node_term n'|]==> ∃ m. first_node_in_nonorigi_strand a n' m " apply - apply(subgoal_tac "n≠n'") prefer 2 apply(rule ccontr) apply(simp add:node_equal) apply(unfold first_node_in_nonorigi_strand_def) apply(drule_tac n'="n'" in non_originate_node,blast) apply(erule disjE) apply simp apply(erule conjE)+ apply(drule non_originate_strand) apply(erule disjE) apply(erule exE) thm exI apply(rule_tac x="(strand n',i)" in exI) apply(unfold strand_def index_def,simp) apply(drule_tac x="index n'" in spec) apply(fold strand_def) apply(subgoal_tac "n'=(strand n', index n')") apply auto apply(simp add:node_equal) apply(unfold strand_def index_def) apply auto done lemma nonOrigNode: assumes a1:"node_sign m2=+" and a2:" non_originate a (strand m2)" and a4:" first_node_in_nonorigi_strand a m2 m" and a8:"a \<sqsubset> node_term m2" shows "index m < index m2" proof - from a1 a4 have b1:"index m ≠index m2" apply - apply (rule ccontr) apply(unfold first_node_in_nonorigi_strand_def) apply(subgoal_tac "m=m2") prefer 2 apply(simp add:node_equal) apply auto done from a2 a4 a8 have "index m <= index m2" apply - apply(unfold non_originate_def) apply(unfold first_node_in_nonorigi_strand_def) apply(rule ccontr) apply(drule_tac x="m" in spec) apply(erule conjE)+ apply(drule_tac x="index m2" in spec) apply simp apply(subgoal_tac "m2=(strand m2, index m2)") apply(unfold strand_def index_def) apply auto done with b1 show ?thesis by auto qed lemma nonOrigNode1: assumes a2:" strand m1=strand m2" and a4:" first_node_in_nonorigi_strand a m2 m" and a8:"a \<sqsubset> node_term m1" shows "index m <= index m1" proof - from a2 a4 a8 show ?thesis apply - apply(unfold first_node_in_nonorigi_strand_def) apply(rule ccontr) apply(erule conjE)+ apply(drule_tac x="index m1" in spec) apply simp apply(subgoal_tac "m1=(strand m2, index m1)") apply simp apply(simp add:node_equal) apply(unfold strand_def index_def) apply auto done qed section{*lemmas on the penetrator's knowledge closure property*} lemma synth_sep_suite: "[|suite M a b; t ∈ synth M; t=MPair g h|]==> g ∈ synth M ∧ h ∈ synth M" apply(case_tac "a\<sqsubset> t") apply(unfold suite_def component_def) apply(subgoal_tac "MPair g h ∉ M") apply blast apply(rule ccontr) apply blast apply(subgoal_tac "¬ a \<sqsubset> g ∧ ¬ a \<sqsubset> h") apply(subgoal_tac "g ∈ M ∧ h ∈ M") apply blast apply blast apply(rule ccontr, unfold parts_relation_def) apply( simp) apply(subgoal_tac "a ∈ parts {g,h}") apply blast apply(erule disjE) apply(subgoal_tac "parts {g} ⊆ parts {g,h}",blast) thm parts_mono apply( rule parts_mono, blast) apply(subgoal_tac "parts {h} ⊆ parts {g,h}",blast) thm parts_mono apply( rule parts_mono, blast) done lemma penetrator_strand: assumes a1:" Is_penetrator_strand (strand m')" and a2: "node_sign m'=+" and a3: "suite M a b" and a5:"∃ m. (m=>+m')" and a6:"b ∈ bundles" and a7:" ∀ m.(m=>+m') ∧node_sign m=--->(node_term m ∈ (synth M))" and a8:" m' ∈ nodes b" shows c:" node_term m' ∈ (synth M) " proof - from a8 and a6 have a8:"m' ∈ Domain" thm bundle_node_in_Domain by(blast dest: bundle_node_in_Domain) let ?s="strand m'" from a1 have "( Is_Tee_strand ?s | Is_Flush_strand ?s | Is_Cat_strand ?s |Is_Sep_strand ?s| Is_E_strand ?s | Is_D_strand ?s |Is_T_strand ?s | Is_K_strand ?s)" apply - apply ( simp add:penetrator_trace) done moreover { assume a12:"Is_E_strand (strand m')" from a2 and a8 and a12 have M0:"snd m'=2" apply - apply( unfold Is_E_strand_def, unfold Domain_def) apply(unfold node_sign_def strand_def,auto) apply(case_tac "y=0") apply simp apply (case_tac "y=1" ) apply simp apply auto done from a12 and this have M1:"∃ k h . node_term (fst m',0)=Key k ∧ node_term (fst m', 1)=h ∧ node_term m'=Crypt k h" (is "∃ k h. ?P k h") by (unfold node_term_def, unfold Is_E_strand_def strand_def, auto) then obtain k and h where M01:" ?P k h" by auto from a12 have M00:"node_sign (fst m', 0)=-∧ node_sign (fst m', 1)=-" apply - apply(unfold Is_E_strand_def,unfold node_sign_def strand_def, auto) done from M0 and a8 have M000:"(fst m',1) ∈ Domain∧((fst m', 1) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from M0 and a8 have M001:"(fst m',0) ∈ Domain∧((fst m', 0) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from M000 and a7 and M00 and M01 have M3:" node_term (fst m',1) ∈ (synth ( M))" apply- apply((erule conjE)+,drule_tac x="(fst m',1)" in spec, drule mp) apply auto done from this and M01 have M4:"h ∈ synth( ( M))" by simp from M0 and a8 have M001:"(fst m',0) ∈ Domain∧((fst m', 0) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from M001 and a7 and M00 and M01 have M3:" node_term (fst m',0) ∈ (synth ( M))" apply- apply((erule conjE)+,drule_tac x="(fst m',0)" in spec, drule mp) apply auto done with M01 have M6:"Key k ∈ M" by simp from this and M4 have M7:"(Crypt k h) ∈ synth ( synth ( M))" by blast from this and M01 have ?thesis apply - apply((erule conjE)+ ,simp) apply( blast dest:synth_synthD) done } moreover { assume a12: "Is_D_strand (fst m')" from a2 and a8 and a12 have M0:"snd m'=2" apply - apply( unfold Is_D_strand_def, unfold Domain_def) apply(unfold node_sign_def strand_def,auto) apply(case_tac "y=0") apply simp apply (case_tac "y=1" ) apply simp apply auto done from a12 and this have M1:"∃ k. ∃ k'. ∃ h. k'=invKey k∧ node_term (fst m',0)= (Key k') ∧node_term (fst m',1)= (Crypt k h)∧ node_term m'=h" (is "∃ k k' h. ?P k k' h") by (unfold node_term_def, unfold Is_D_strand_def, auto) then obtain k and k' and h where M01:" ?P k k' h" by auto from a12 have M00:"node_sign (fst m', 0)=-∧ node_sign (fst m', 1)=-" apply - apply(unfold Is_D_strand_def,unfold node_sign_def, auto) done from M0 and a8 have M000:"(fst m',1)∈Domain∧((fst m', 1) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from M0 and a8 have M001:"(fst m',0)∈Domain∧((fst m', 0) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done with a7 M00 M01 have M002:"Key (invKey k) ∈ synth M" apply - apply(drule_tac x="(fst m',0)" in spec) apply auto done then have M002:"Key (invKey k) ∈ M" apply - apply blast done from M000 M00 a7 M01 have M003:"Crypt k h ∈ synth M" apply - apply(drule_tac x="(fst m',1)" in spec) apply auto done have "a \<sqsubset> Crypt k h | ¬ a \<sqsubset> Crypt k h" by blast moreover {assume N0:"a \<sqsubset> Crypt k h" from M000 and a6 and a8 and M0 have M51: "(strand m',0)∈ nodes b " apply - apply( subgoal_tac "((fst m',1), m'):casual2^+") apply( drule_tac i="0" in bundle_casual2_closed3 ) apply assumption+ apply simp apply blast apply(blast intro:trancl_into_rtrancl) done from a3 M002 N0 M51 a1 M01 have M6: "¬ Crypt k h∈M" apply - apply(unfold suite_def) apply (rule ccontr) apply(erule conjE) apply(drule_tac x="Crypt k h" in bspec) apply simp apply(unfold component_def regularK_def) apply simp apply(drule_tac x="(strand m',0)" in bspec) apply simp apply(unfold Is_regular_strand_def strand_def) apply auto done from this and M003 and M01 have ?thesis by blast } moreover {assume N0:"¬ a \<sqsubset> Crypt k h" from N0 have "¬ a\<sqsubset> h" apply - apply(rule ccontr) apply (unfold parts_relation_def) apply auto done with a3 have "h∈M" apply - apply(unfold suite_def) apply blast done then have "h ∈synth M" by blast with M01 have ?thesis by blast } ultimately have ?thesis by blast } moreover {assume a12:"Is_Cat_strand (fst m')" from a2 and a8 and a12 have M0:"snd m'=2" apply - apply( unfold Is_Cat_strand_def, unfold Domain_def, unfold node_sign_def,auto) apply(case_tac "y=0") apply simp apply (case_tac "y=1" ) apply simp apply auto done from a12 and this have M1:"∃ g h . node_term (fst m',0)=g ∧ node_term (fst m', 1)=h ∧ node_term m'={|g,h|}" (is "∃ g h. ?P g h") by (unfold node_term_def, unfold Is_Cat_strand_def, auto) then obtain g and h where M01:" ?P g h" by auto from a12 have M00:"node_sign (fst m', 0)=-∧ node_sign (fst m', 1)=-" apply - apply(unfold Is_Cat_strand_def,unfold node_sign_def, auto) done from M0 and a8 have M000:"(fst m',1)∈Domain∧((fst m', 1) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from M000 and a7 and M00 and M01 and a3 have M3:" h ∈(synth ( M))" apply- apply(case_tac "a\<sqsubset> node_term (fst m', 1)") apply((erule conjE)+,drule_tac x="(fst m',1)" in spec, drule mp) apply auto done from M0 and a8 have M001:"(fst m',0)∈Domain∧((fst m', 0) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from M001 and a7 and M00 and M01 and a3 have M5:" g∈ (synth (M))" apply- apply(case_tac "a\<sqsubset> node_term (fst m', 0)") apply((erule conjE)+,drule_tac x="(fst m',0)" in spec, drule mp) apply auto done from M3 and M5 and M01 have ?thesis by (simp, blast dest:synth_synthD) } moreover {assume a12:"Is_Sep_strand (fst m')" from a2 and a8 and a12 have M0:"snd m'=2|snd m'=1" apply - apply( unfold Is_Sep_strand_def) apply(unfold Domain_def, unfold node_sign_def,auto) apply(case_tac "y=0") apply simp apply(case_tac "y=1" ) apply simp apply auto done from a12 have M00:"node_sign (fst m', 0)=-∧ node_sign (fst m', 1)=+∧ node_sign (fst m', 2)=+" apply - apply(unfold Is_Sep_strand_def) apply(unfold node_sign_def, auto) done from a12 and this have M1:"∃ g h . node_term (fst m',0)={|g,h|} ∧ node_term (fst m', 1)=g ∧ node_term (fst m',2)=h" (is "∃ g h. ?P g h") by (unfold node_term_def, unfold Is_Sep_strand_def, auto) then obtain g and h where M01:" ?P g h" by auto from M0 have M0:"snd m'=2|snd m'=1" by simp moreover {assume a20:"snd m'=2" from this and a8 have M000:"(fst m',0)∈Domain∧((fst m', 0) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from a20 and M01 have M02: "node_term m'= h" apply - apply ((erule conjE)+,simp) apply(drule sym) apply simp done from M000 and a7 and M00 and M01 and a3 have M3:" MPair g h ∈(synth ( M))" apply- apply((erule conjE)+,drule_tac x="(fst m',0)" in spec, drule mp) apply auto done with a3 have "h∈synth M" by (blast dest:synth_sep_suite) from this and M02 have ?thesis apply - apply blast done } moreover {assume a20:"snd m'=1" from this and a8 have M000:"(fst m',0)∈Domain∧((fst m', 0) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from a20 and M01 have M02: "node_term m'= g" apply - apply ((erule conjE)+,simp) apply(drule sym) apply simp done from M000 and a7 and M00 and M01 and a3 have M3:" MPair g h ∈(synth ( M))" apply- apply((erule conjE)+,drule_tac x="(fst m',0)" in spec, drule mp) apply auto done with a3 have "g∈synth M" by (blast dest:synth_sep_suite) from this and M02 have ?thesis by blast } ultimately have ?thesis by blast } moreover {assume a9:"Is_Tee_strand ?s" from a8 and a2 and a9 have M0: "snd m' =2 | snd m'=1" apply - apply( unfold Is_Tee_strand_def) apply(unfold Domain_def strand_def, auto) apply(case_tac "y=0") apply(unfold node_sign_def, simp) apply auto done from a9 and this have "∃ g . node_term (fst m',0)=g ∧ node_term (fst m', 1)=g ∧ node_term m'=g" (is "∃ g. ?P g") by (unfold node_term_def Is_Tee_strand_def strand_def, auto) then obtain g where M01:" ?P g" .. from a9 have M00:"node_sign (fst m', 0)=-" apply - apply(unfold Is_Tee_strand_def node_sign_def) apply(unfold strand_def, auto) done have ?thesis proof - from M0 have M0: "snd m' =2 | snd m'=1" by simp moreover {assume M1:" snd m'=2" from this and a8 have M000:"(fst m',0)∈Domain∧((fst m', 0) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from M1 and M01 have M02: "node_term m'= g" apply - apply ((erule conjE)+,simp) done from M000 and a7 and M00 and M01 and a3 have M3:" g ∈(synth ( M))" apply- apply((erule conjE)+,drule_tac x="(fst m',0)" in spec, drule mp) apply auto done from this and M02 have ?thesis by blast } moreover {assume M2:" snd m'=1" from this and a8 have M000:"(fst m',0)∈Domain∧((fst m', 0) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from M2 and M01 have M02: "node_term m'= g" apply - apply ((erule conjE)+,simp) done from M000 and a7 and M00 and M01 and a3 have M3:" g ∈(synth ( M))" apply- apply((erule conjE)+,drule_tac x="(fst m',0)" in spec, drule mp) apply auto done from this and M02 have ?thesis by blast } ultimately show ?thesis by blast qed } moreover {assume a9:"Is_T_strand ?s" from a2 and a9 and a8 have "snd m'=0" apply - apply (unfold Is_T_strand_def, unfold node_sign_def) apply(unfold Domain_def strand_def) apply auto done from this and a5 have ?thesis apply - apply (erule exE ) apply(drule casual2_trancl_less) apply(unfold index_def,simp) done } moreover {assume a10 :"Is_K_strand (fst m')" from a2 and a10 and a8 have "snd m'=0" apply - apply (unfold Is_K_strand_def, unfold node_sign_def) apply(unfold Domain_def ) apply auto done from this and a5 have ?thesis apply - apply (erule exE ) apply(drule casual2_trancl_less) apply(unfold index_def, arith) done } moreover {assume a11:"Is_Flush_strand (fst m')" from a2 and a11 and a8 have ?thesis apply- apply(unfold Is_Flush_strand_def,unfold Domain_def, unfold node_sign_def, auto) done } ultimately show ?thesis by(unfold strand_def, blast) qed lemma another_penetrator_strand: assumes a1:" node_term m'∉ (synth M)" and a2: "node_sign m'=+" and a3: "suite M a b" and a5:"∃ m. (m=>+m')" and a6:"b:bundles" and a7:" ∀ m. (m=>+m')∧ node_sign m=--->(node_term m: (synth M))" and a8:" m'∈ nodes b" shows c:" Is_regular_strand (strand m') " proof - from a1 a2 a3 a5 a6 a7 a8 show " Is_regular_strand (strand m') " apply - apply(rule ccontr) apply(unfold Is_regular_strand_def) apply (blast dest: penetrator_strand) done qed section{*lemmas on path existence *} lemma slice_arr_not_nil:" ∃ n ns. slice_arr_cons s i len= n# ns ∧ n=(s,i)" apply(induct_tac len) apply auto done lemma slic_arr_ith: "∀ start i. i<=len -->(slice_arr_cons s start len)!i=(s,start+i)" apply(induct_tac len) apply simp apply(rule allI)+ apply(rule impI) apply auto apply(case_tac "i=0") apply simp apply(drule_tac x=" Suc start" in spec) apply(drule_tac x="i- 1 " in spec) apply simp apply(subgoal_tac " ((s, start) # slice_arr_cons s (Suc start) n)= ([(s, start)]@ slice_arr_cons s (Suc start) n)") apply(simp only:nth_append) apply(subgoal_tac "0<i") apply(drule mp) apply arith apply (simp ) apply simp apply simp done lemma slice_arr_ith:"i<=len ==>(slice_arr_cons s start len)!i=(s,start+i)" apply - apply(cut_tac slic_arr_ith, blast) done lemma slice_arr_length: "∀ start. length (slice_arr_cons s start len) =len +1" apply(induct_tac len) apply auto done lemma slice_arr_length1[simp]: " length (slice_arr_cons s start len) =len +1" apply(cut_tac slice_arr_length, auto) done lemma slice_arr_last[simp]: " last (slice_arr_cons s start len)= (s,(start + len))" apply(subgoal_tac "(slice_arr_cons s start len) ≠[]") prefer 2 apply(cut_tac slice_arr_not_nil[where s="s" and i="start" and len="len"]) apply(erule exE)+ apply simp apply(cut_tac slice_arr_length[where s="s" and len="len"]) apply (drule_tac x="start" in spec) apply(cut_tac last_conv_nth[where xs="(slice_arr_cons s start len)"]) apply auto apply(cut_tac slic_arr_ith[where s="s" and len="len"]) apply(drule_tac x="start" in spec) apply(drule_tac x="len" in spec) apply blast done lemma slice_arr_last2[simp]: "[|index m <= index m2;strand m=strand m2|]==>last ( slice_arr_cons (strand m) (index m) (index m2 - index m))=m2" apply(simp) apply(unfold strand_def index_def, simp add:node_equal) done lemma slice_arr_no_empty[simp]:"slice_arr_cons s i len ≠[]" apply(cut_tac slice_arr_not_nil[where s="s" and i="i" and len="len"]) apply(erule exE)+ apply simp done lemma slice_index: assumes a: "m∈set (slice_arr_cons (strand m) start i) " shows "index m <= start +i " proof - let ?xs="(slice_arr_cons (strand m) start i)" from a have b1:"∃ i < length ?xs. ?xs!i = m" apply - apply(cut_tac a,simp only:in_set_conv_nth) done then obtain ia where b1:"ia< length ?xs ∧ ?xs!ia = m" by blast from b1 have "slice_arr_cons (strand m) start i ! ia =(strand m,start+ia)" apply - apply(rule slice_arr_ith) by simp with b1 have "m=(strand m,start+ia)" by simp with b1 show ?thesis apply - apply(simp add:node_equal) apply(unfold index_def) apply auto done qed lemma slice_arr_close: assumes a1:" (m', ma) ∈ casual2+" and a2:"ma ∈ set (slice_arr_cons (strand m) (index m) (index m2 - index m))" and a3:"index m < index m2" and a4:"a \<sqsubset> node_term m'" and a5:"first_node_in_nonorigi_strand a m2 m" shows "m'∈set (slice_arr_cons (strand m) (index m) (index m2 - index m))" proof - let ?xs="(slice_arr_cons (strand m) (index m) (index m2 - index m))" from a2 have "(∃ i < length ?xs. ?xs!i = ma)" apply - apply( simp add:set_conv_nth,blast) done then obtain i where c1:"i < length ?xs ∧ ?xs!i = ma" by blast from a3 have c2:"length ?xs = (index m2 - index m) +1" apply - apply(cut_tac ?s="strand m" and len="index m2 -index m" in slice_arr_length) apply(drule_tac x="index m" in spec) by simp with c1 have c3:"(strand m, (index m + i))=ma" apply - apply simp apply(cut_tac s="(strand m)" and len="(index m2 - index m) " in slic_arr_ith) apply(drule_tac x="index m" in spec) apply(drule_tac x="i" in spec) apply auto done thm casual2_trancl_less from a1 have c30:"index m' < index ma ∧ strand m' = strand ma" apply - apply(blast dest:casual2_trancl_less) done with a1 c3 have c4:"strand m'=strand m" apply - apply (simp add:node_equal) apply(unfold strand_def,auto) done from a5 this a4 have c5:"index m <=index m'" apply - apply(rule nonOrigNode1) apply(unfold first_node_in_nonorigi_strand_def,auto) done from c4 c30 have c6:"strand ma=strand m" by auto from a2 c6 have c7:"index ma<=index m + (index m2 - index m)" apply - apply(rule_tac start="index m" and i="(index m2 - index m)" in slice_index) by simp from a3 c7 have c8:"index ma<=index m2" by simp with c5 c8 have "(∃ i < length ?xs. ?xs!i = m')" apply - apply(rule_tac x="index m' - index m" in exI) apply(rule conjI) apply(cut_tac ?s="strand m" and len="index m2 -index m" in slice_arr_length) apply(drule_tac x="index m" in spec) apply simp apply(cut_tac c30) apply arith apply(subgoal_tac "(slice_arr_cons (strand m) (index m) (index m2 - index m))! (index m' - index m)=(strand m, index m+ (index m' - index m))") prefer 2 apply(rule slice_arr_ith) apply(cut_tac c8 c30) apply arith apply simp apply(cut_tac c4, simp add:node_equal) apply(unfold strand_def index_def, simp) done from this show ?thesis by(simp add:in_set_conv_nth) qed lemma slice_arr_close1: assumes a1:" (m', ma) ∈ casual2+" and a2:"ma ∈ set (slice_arr_cons (strand m) (index m) (index m2 - index m))" and a3:"index m <= index m2" and a4:"a \<sqsubset> node_term m'" and a5:"unique_originate a m" shows "m'∈set (slice_arr_cons (strand m) (index m) (index m2 - index m))" proof - let ?xs="(slice_arr_cons (strand m) (index m) (index m2 - index m))" from a2 have "(∃ i < length ?xs. ?xs!i = ma)" apply - apply( simp add:set_conv_nth,blast) done then obtain i where c1:"i < length ?xs ∧ ?xs!i = ma" by blast from a3 have c2:"length ?xs = (index m2 - index m) +1" apply - apply(cut_tac ?s="strand m" and len="index m2 -index m" in slice_arr_length) apply(drule_tac x="index m" in spec) by simp with c1 have c3:"(strand m, (index m + i))=ma" apply - apply simp apply(cut_tac s="(strand m)" and len="(index m2 - index m) " in slic_arr_ith) apply(drule_tac x="index m" in spec) apply(drule_tac x="i" in spec) apply auto done thm casual2_trancl_less from a1 have c30:"index m' < index ma ∧ strand m' = strand ma" apply - apply(blast dest:casual2_trancl_less) done with a1 c3 have c4:"strand m'=strand m" apply - apply (simp add:node_equal) apply(unfold strand_def,auto) done from a5 this a4 have c5:"index m <=index m'" apply - apply(unfold unique_originate_def originate_def ) apply(rule ccontr) apply(erule conjE)+ apply(drule_tac x="index m'" in spec) apply simp apply(subgoal_tac "m'=(strand m, index m')") apply simp apply(simp add:node_equal) apply(unfold strand_def index_def, simp) done from c4 c30 have c6:"strand ma=strand m" by auto from a2 c6 have c7:"index ma<=index m + (index m2 - index m)" apply - apply(rule_tac start="index m" and i="(index m2 - index m)" in slice_index) by simp from a3 c7 have c8:"index ma<=index m2" by simp with c5 c8 have "(∃ i < length ?xs. ?xs!i = m')" apply - apply(rule_tac x="index m' - index m" in exI) apply(rule conjI) apply(cut_tac ?s="strand m" and len="index m2 -index m" in slice_arr_length) apply(drule_tac x="index m" in spec) apply simp apply(cut_tac c30) apply arith apply(subgoal_tac "(slice_arr_cons (strand m) (index m) (index m2 - index m))! (index m' - index m)=(strand m, index m+ (index m' - index m))") prefer 2 apply(rule slice_arr_ith) apply(cut_tac c8 c30) apply arith apply simp apply(cut_tac c4, simp add:node_equal) apply(unfold strand_def index_def, simp) done from this show ?thesis by(simp add:in_set_conv_nth) qed lemma Complete_path_mono: "[|p∈complete_Path a g1 |]==> nodes g1 ⊆ nodes g2--> edges g1 ⊆ edges g2--> p∈complete_Path a g2" apply(erule complete_Path.induct) apply (blast intro: addPos1 addPos2 addnege )+ done lemma completePathProp: assumes a1:"p∈complete_Path a b" shows "unique_originate a n -->(b∈bundles) -->(∀ n'.((n'∈ nodes b) ∧last p=n' -->(n'=n | (n,n'):(edges b)^+)))" (is "?cons1 p ") and "p ≠[]" (is "?cons2 p") and "∀m. m∈(set p)∧node_sign m=+ ∧ a\<sqsubset> node_term m --> (∀m'.(m'=>+ m)∧a\<sqsubset>node_term m'--> m'∈(set p))" (is "?cons3 p ") using a1 proof induct fix m m2 p let ?p="p @ slice_arr_cons (strand m) (index m) (index m2 - index m)" assume a2:" non_originate a (strand m2)" and a3:"node_sign m2 = +" and a4:" first_node_in_nonorigi_strand a m2 m" and a5:"p ∈ complete_Path a b" and a6:"?cons1 p ∧ ?cons2 p ∧?cons3 p " and a7:"last p = m" and a8:"a \<sqsubset> node_term m2" show " ?cons1 ?p ∧ ?cons2 ?p ∧?cons3 ?p " proof(rule conjI,(rule impI)+,rule allI) fix n' assume b1:"unique_originate a n" and b2:"b ∈ bundles" show "n' ∈ nodes b ∧ last (p @ slice_arr_cons (strand m) (index m) (index m2 -index m)) = n' --> n' = n ∨ (n, n') ∈ (edges b)+" proof(rule impI, (erule conjE)+) assume b3:" n' ∈ nodes b" and b4:"last (p @ slice_arr_cons (strand m) (index m) (index m2 - index m)) = n'" from a2 a3 a4 a8 have c1:"index m <= index m2" by(blast intro:nonOrigNode1) show "n' = n ∨ (n, n') ∈ (edges b)+" proof - from a4 have c2:"strand m2=strand m" by(unfold first_node_in_nonorigi_strand_def,simp) from c1 c2 have "last (p @ slice_arr_cons (strand m) (index m) (index m2 - index m))=m2" apply - apply(subgoal_tac "slice_arr_cons (strand m) (index m) (index m2 - index m)≠[]") prefer 2 apply simp apply(simp add:last_appendR) apply(simp add:node_equal) apply(unfold index_def strand_def, auto) done with b4 have c4:"n'=m2" by simp thm bundle_casual2_closed3 with b2 c1 c2 b3 have c3:"m:nodes b∧(m,m2):(edges b Int casual2)^*" apply - apply(subgoal_tac "(strand m2, index m)∈nodes b ∧ ((strand m2, index m),m2)∈(edges b Int casual2)^*") prefer 2 apply (blast dest: bundle_casual2_closed3) apply(simp add:node_equal) apply(unfold strand_def index_def,auto) done with b1 b2 a6 a7 have "m=n | (n,m): (edges b)^+" apply - apply blast done with c3 have "m2=n | (n,m2): (edges b)^+" apply - apply(subgoal_tac "(m, m2) ∈ (edges b)^*") apply(drule rtranclD) apply(erule disjE) apply(erule disjE) apply simp apply simp apply(erule disjE) apply simp apply(blast dest:transitive_closure_trans) apply(subgoal_tac "(edges b ∩ casual2)* ⊆ (edges b)*") apply(blast ) thm rtrancl_mono apply(blast intro!: rtrancl_mono) done with c4 show ?thesis by simp qed qed next show "?cons2 ?p ∧?cons3 ?p " proof show "?cons2 ?p" by simp next show "?cons3 ?p" proof((rule allI,rule impI)+,(erule conjE)+) fix ma m' assume c1:"(m', ma) ∈ casual2+" and c2:"a \<sqsubset> node_term m'" and c3:"ma∈set ?p" and c4:" node_sign ma = +" and c5:" a \<sqsubset> node_term ma" show "m' ∈ set ?p" proof - from c3 have "ma:set p | ma∈set (slice_arr_cons (strand m) (index m) (index m2 - index m))" by simp moreover { assume d1:"ma:set p " from a6 d1 c4 c5 c1 c2 have "m':set p" apply - apply(erule conjE)+ apply(drule_tac x="ma" in spec) apply blast done then have ?thesis by simp } moreover { assume d1:"ma:set (slice_arr_cons (strand m) (index m) (index m2 - index m))" from a2 a3 a4 a8 have d2:"index m < index m2" by(blast intro:nonOrigNode) from c1 d1 c2 a4 d2 have "m':set (slice_arr_cons (strand m) (index m) (index m2 - index m))" apply - thm slice_arr_close apply(rule_tac m'="m'" and ma="ma" and ?m2.0="m2" and m="m" in slice_arr_close) apply auto done then have ?thesis by simp } ultimately show ?thesis by blast qed qed qed qed next fix m m2 let ?p="slice_arr_cons (strand m) (index m) (index m2 - index m)" assume a2:"unique_originate a m" and a3:"node_sign m2 = +" and a4:"strand m2 = strand m" and a8:"a \<sqsubset> node_term m2" show " ?cons1 ?p ∧ ?cons2 ?p ∧?cons3 ?p " proof(rule conjI,(rule impI)+,rule allI) fix n' assume b1:"unique_originate a n" and b2:"b ∈ bundles" show "n' ∈ nodes b ∧ last ( slice_arr_cons (strand m) (index m) (index m2 -index m)) = n' --> n' = n ∨ (n, n') ∈ (edges b)+" proof(rule impI, (erule conjE)+) assume b3:" n' ∈ nodes b" and b4:"last ( slice_arr_cons (strand m) (index m) (index m2 - index m)) = n'" show "n' = n ∨ (n, n') ∈ (edges b)+" proof - from a3 a2 a4 a8 have c1:"index m <= index m2" apply - apply(unfold unique_originate_def originate_def) apply(rule ccontr) apply(erule conjE)+ apply(drule_tac x="index m2" in spec) apply simp apply(subgoal_tac "strand m=strand m2") apply(thin_tac "strand m2 = strand m") apply simp apply(subgoal_tac "m2=(strand m2,index m2)") apply (simp add:node_equal) prefer 2 apply simp apply(unfold strand_def index_def,simp) done from c1 a4 have "last ( slice_arr_cons (strand m) (index m) (index m2 - index m))=m2" apply - apply (simp only:node_equal) apply(unfold strand_def index_def,auto) done with b4 have c4:"n'=m2" by simp with b2 c1 a4 b3 have c3:"m∈nodes b∧(m,m2):(edges b Int casual2)^*" apply - apply(subgoal_tac "(strand m2, index m):nodes b ∧ ((strand m2, index m),m2):(edges b Int casual2)^*") prefer 2 apply (blast dest: bundle_casual2_closed3) apply(rule conjI) apply(subgoal_tac "m=(strand m,index m)") apply(simp add:node_equal) apply(unfold strand_def index_def,simp) apply(simp add:node_equal) done from b1 a2 have "m=n" apply - apply(unfold unique_originate_def,blast ) done with c3 c4 have "m=n | (n,m): (edges b)^+" apply - apply blast done with c3 have "m2=n | (n,m2): (edges b)^+" apply - apply(subgoal_tac "(m, m2) ∈ (edges b)^*") apply(drule rtranclD) apply(erule disjE) apply(erule disjE) apply simp apply simp apply(erule disjE) apply simp apply(blast dest:transitive_closure_trans) apply(subgoal_tac "(edges b ∩ casual2)* ⊆ (edges b)*") apply(blast ) thm rtrancl_mono apply(blast intro!: rtrancl_mono) done with c4 show ?thesis by simp qed qed next show "?cons2 ?p ∧?cons3 ?p " proof show "?cons2 ?p" by simp next show "?cons3 ?p" proof((rule allI,rule impI)+,(erule conjE)+) fix ma m' assume c1:"(m', ma) ∈ casual2+" and c2:"a \<sqsubset> node_term m'" and c3:"ma:set ?p" and c4:" node_sign ma = +" and c5:" a \<sqsubset> node_term ma" show "m' ∈ set ?p" proof - from a2 a4 a8 have d2:"index m <= index m2" apply - apply(unfold unique_originate_def originate_def ) apply(rule ccontr) apply(erule conjE)+ apply(drule_tac x="index m2" in spec) apply simp apply(subgoal_tac "m2=(strand m, index m2)") apply simp apply(simp add:node_equal) apply(unfold strand_def index_def, auto) done (*from a4 have "index m≠index m2" apply - apply(rule ccontr) apply(simp add:node_equal) done with d2 have "index m < index m2" by simp*) from c1 c2 c3 a2 a4 this thm slice_arr_close1 show ?thesis apply - thm slice_arr_close1 apply(rule_tac m'="m'" and ma="ma" and ?m2.0="m2" and m="m" in slice_arr_close1) apply auto done qed qed qed qed next fix m1 m2 p assume a2:"m1-> m2" and a3:"m1:nodes b" and a4:"(m1,m2)∈edges b" and a5:"p ∈ complete_Path a b" and a6:"?cons1 p ∧ ?cons2 p ∧ ?cons3 p" and a7:"last p = m1" and a8:"a \<sqsubset> node_term m2" let ?p = "p @ [m2]" show " ?cons1 ?p ∧ ?cons2 ?p ∧?cons3 ?p " proof(rule conjI,(rule impI)+,rule allI) fix n' assume b1:"unique_originate a n" and b2:"b ∈ bundles" show "n' ∈ nodes b ∧ last (p @ [m2]) = n' --> n' = n ∨ (n, n') ∈ (edges b)+" proof(rule impI, (erule conjE)+) assume b3:" n' ∈ nodes b" and b4:"last (p @ [m2] ) = n'" show "n' = n ∨ (n, n') ∈ (edges b)+" proof - from a6 have "last (p @ [m2] ) =m2" by simp with b4 have c0:"m2=n'" by simp from b3 b4 have c1: "m2∈nodes b" by simp from a3 a6 a7 b1 b2 have "m1 = n ∨ (n, m1) ∈ (edges b)+" by blast thm bundle_edge_is_unique with a4 have "(n,m2):(edges b)^+" by (blast dest:transitive_closure_trans) with c0 show ?thesis by simp qed qed next show "?cons2 ?p ∧?cons3 ?p " proof show "?cons2 ?p" by simp next show "?cons3 ?p" proof((rule allI,rule impI)+,(erule conjE)+) fix ma m' assume c1:"(m', ma) ∈ casual2+" and c2:"a \<sqsubset> node_term m'" and c3:"ma∈set ?p" and c4:" node_sign ma = +" and c5:" a \<sqsubset> node_term ma" show "m' ∈ set ?p" proof - from c3 have "ma:set p | ma=m2" by auto moreover { assume d1:"ma∈set p " from a6 d1 c4 c5 c1 c2 have "m':set p" apply - apply(erule conjE)+ apply(drule_tac x="ma" in spec) apply blast done then have ?thesis by simp } moreover { assume d1:"ma=m2" from a2 have "node_sign m2=-" by (blast dest:recv_node_is_minus) from this d1 c4 have ?thesis by simp } ultimately show ?thesis by blast qed qed qed qed qed lemma path_exist: assumes A:"b∈bundles" shows C: "unique_originate a n--> (∀ n'.( n'∈ nodes b∧ a \<sqsubset> node_term n'--> (∃ p. p∈ (complete_Path a b) ∧ (last p)=n'))) " using A proof induct case Nil show ?case by(unfold nodes_def, auto) next fix b n1 n2 assume a1: "b ∈ bundles" and a2: " unique_originate a n --> (∀n'. n' ∈ nodes b ∧ a \<sqsubset> node_term n' --> (∃p. p∈ (complete_Path a b) ∧ (last p)=n') ) " and a3: "node_sign n2 = +" and a4:" n2 ∈ Domain" and a5:" n2 ∉ nodes b" and a6:" 0 <index n2" and a7:" n1 ∈ nodes b" and a8:" n1 => n2" let ?b'="({n2} ∪ nodes b, {(n1, n2)} ∪ edges b)" show "unique_originate a n --> (∀n'. n' ∈ nodes ?b' ∧ a \<sqsubset> node_term n' --> (∃p. p ∈ complete_Path a ?b' ∧ last p = n'))" proof(rule impI,rule allI,rule impI,erule conjE) fix n' assume b1:" unique_originate a n" and b2:"n' ∈ nodes ?b'" and b3:"a \<sqsubset> node_term n'" show "∃p. p ∈ complete_Path a ?b' ∧ last p = n'" proof - from b2 have "n'∈nodes b |n'=n2" by(unfold nodes_def,auto) moreover { assume c1:"n'∈nodes b " from c1 b1 b2 b3 a2 have "(∃p. p∈ (complete_Path a b) ∧ (last p)=n') " (is "∃ p. ?P p b") by blast then obtain p where c2:"?P p b" by blast then have c3:"?P p ?b'" thm Complete_path_mono apply - apply(subgoal_tac "(nodes b ) ⊆ nodes ?b' ") apply(subgoal_tac "(edges b ) ⊆ edges ?b'") apply(erule conjE) apply (blast dest:Complete_path_mono) apply(unfold nodes_def edges_def,auto) done then have ?thesis by blast } moreover { assume c1:"n'=n2" have "strand n2=strand n | strand n2≠strand n" by simp moreover { assume d1:"strand n2=strand n" have "n2=n|n2≠n" by simp moreover { assume e1:"n2=n" have e2: "(slice_arr_cons (strand n) (index n) (index n - index n))∈complete_Path a ?b'" apply - apply(rule addPos2) apply(assumption)+ prefer 2 apply(cut_tac b1,unfold unique_originate_def originate_def,simp) prefer 2 apply(cut_tac b1,unfold unique_originate_def originate_def,simp) apply(cut_tac c1 b3,simp) done have "(slice_arr_cons (strand n) (index n) (index n - index n))=[n]" apply (simp add:node_equal) apply(unfold strand_def index_def,auto) done with c1 e2 e1 have ?thesis apply - apply(rule_tac x="[n]" in exI) by simp } moreover { assume e2:"n2≠n" from e2 b1 c1 b3 d1 have e0:"index n<=index n2 " apply - thm non_originate_node apply(drule_tac ?n'="n2" in non_originate_node) apply blast apply(unfold non_originate_def) apply auto done have "(slice_arr_cons (strand n) (index n) (index n2 - index n))∈complete_Path a ?b'" apply - apply(rule addPos2) apply(assumption)+ prefer 2 apply assumption apply(cut_tac c1 b3,simp) done with e0 c1 d1 have ?thesis apply - apply(rule_tac x="slice_arr_cons (strand n) (index n) (index n2 - index n)" in exI) apply simp apply(simp add:node_equal) apply(unfold strand_def index_def, auto) done } ultimately have ?thesis by blast } moreover { assume d2:" strand n2≠strand n" from c1 b3 have d3:"a \<sqsubset> node_term n2" by simp from d2 b1 this have d4:"∃ m. first_node_in_nonorigi_strand a n2 m" thm non_originate_node1 apply - apply(rule non_originate_node1) by auto then obtain m where d4:"first_node_in_nonorigi_strand a n2 m" by blast from d2 have "n2≠n" apply - apply(rule ccontr) apply(simp add:node_equal) done from this d2 b1 d3 have d5:" non_originate a (strand n2)" apply - apply(drule non_originate_node) apply auto done from d5 d3 have "index m < index n2" apply - apply(rule_tac a="a" in nonOrigNode) apply assumption+ done with a3 d4 have d01:"index m < index n2" apply - apply(subgoal_tac "index m ≠index n2") apply arith apply(rule ccontr) apply(unfold first_node_in_nonorigi_strand_def) apply(subgoal_tac "m=n2") apply(simp add:node_equal)+ done from this and a8 have "index m <=index n1" apply - apply(unfold casual2_def,simp) done with a7 a1 a8 have "m∈nodes b ∧ a\<sqsubset> node_term m" apply - apply(subgoal_tac "(strand n1, index m)∈nodes b") apply(cut_tac d4) apply(unfold first_node_in_nonorigi_strand_def) apply(unfold casual2_def) apply(subgoal_tac "(strand n1, index m)=m") apply simp apply(simp add:node_equal) apply(unfold strand_def index_def, simp) apply(fold strand_def index_def) apply(blast dest:bundle_casual2_closed3) done from a2 this b1 have " (∃p. p: (complete_Path a b) ∧ (last p)=m) " (is "∃ p. ?P p b m") by blast then obtain p where d50 :"?P p b m" by blast then have d6:"?P p ?b' m" thm Complete_path_mono apply - apply(subgoal_tac "(nodes b ) ⊆ nodes ?b' ") apply(subgoal_tac "(edges b ) ⊆ edges ?b'") apply(erule conjE) apply (blast dest:Complete_path_mono) apply(unfold nodes_def edges_def,auto) done let ?p'="p@(slice_arr_cons (strand m) (index m) (index n2 - index m))" from d5 a3 d50 d6 b3 c1 d4 have "?P ?p' ?b' n2" apply - apply(rule conjI) apply(rule addPos1) prefer 7 apply(subgoal_tac "slice_arr_cons (strand m) (index m) (index n2 - index m)≠[]") apply( simp add:node_equal) apply(cut_tac d01 d4 ) apply(unfold strand_def index_def first_node_in_nonorigi_strand_def) apply auto done with c1 have ?thesis by blast } ultimately have ?thesis by blast } ultimately show ?thesis by blast qed qed next fix b n2 assume a1: "b ∈ bundles" and a2: " unique_originate a n --> (∀n'. n' ∈ nodes b ∧ a \<sqsubset> node_term n' --> (∃p. p∈ (complete_Path a b) ∧ (last p)=n') ) " and a3: "node_sign n2 = +" and a5:" n2 ∉ nodes b" and a7:" n2 ∈Domain" and a8:"index n2 = 0" show " unique_originate a n --> (∀n'. n' ∈ nodes ({n2} ∪ nodes b, edges b) ∧ a \<sqsubset> node_term n' --> (∃p. p ∈ complete_Path a ({n2} ∪ nodes b, edges b) ∧ last p = n'))" proof(rule impI,rule allI,rule impI,erule conjE) fix n' assume b1:" unique_originate a n" and b2:"n' ∈ nodes ({n2} ∪ nodes b, edges b)" and b3:"a \<sqsubset> node_term n'" show "∃p. p ∈ complete_Path a ({n2} ∪ nodes b, edges b) ∧ last p = n'" proof - let ?b'="({n2} ∪ nodes b, edges b)" from b2 have "n':nodes b |n'=n2" by(unfold nodes_def,auto) moreover { assume c1:"n'∈nodes b " from c1 b1 b2 b3 a2 have "(∃p. p∈ (complete_Path a b) ∧ (last p)=n') " (is "∃ p. ?P p b") by blast then obtain p where c2:"?P p b" by blast then have c3:"?P p ?b'" thm Complete_path_mono apply - apply(subgoal_tac "(nodes b ) ⊆ nodes ?b' ") apply(subgoal_tac "(edges b ) ⊆ edges ?b'") apply(erule conjE) apply (blast dest:Complete_path_mono) apply(unfold nodes_def edges_def,auto) done then have ?thesis by blast } moreover { assume c1:"n'=n2" from a3 b1 c1 b3 a8 have c2:"n=n2" apply - apply(subgoal_tac "originate a n2") apply(unfold unique_originate_def,blast) apply(unfold originate_def) apply auto done have e2: "(slice_arr_cons (strand n) (index n) (index n - index n))∈complete_Path a ?b'" apply - apply(rule addPos2) apply(assumption)+ prefer 2 apply(cut_tac b1,unfold unique_originate_def originate_def,simp) prefer 2 apply(cut_tac b1,unfold unique_originate_def originate_def,simp) apply(cut_tac c1 b3,simp) done have "(slice_arr_cons (strand n) (index n) (index n - index n))=[n]" apply (simp add:node_equal) apply(unfold strand_def index_def,auto) done with c1 c2 e2 have ?thesis apply - apply(rule_tac x="[n]" in exI) by simp } ultimately show ?thesis by blast qed qed next fix b n1 n1' n2 assume a1: "b ∈ bundles" and a2: " unique_originate a n --> (∀n'. n' ∈ nodes b ∧ a \<sqsubset> node_term n' --> (∃p. p∈ (complete_Path a b) ∧ (last p)=n') ) " and a3: "node_sign n2 = - " and a4:"strand n1 ≠ strand n2 ∧ n1 -> n2 ∧ n1 ∈ nodes b ∧ (∀n3. n3 ∈ nodes b --> (n1, n3) ∉ edges b)" and a5:" n2 ∉ nodes b" and a6:" 0 <index n2" and a7:" n1' ∈ nodes b" and a8:" n1' => n2" show "unique_originate a n --> (∀n'. n' ∈ nodes ({n2} ∪ nodes b, {(n1, n2), (n1', n2)} ∪ edges b) ∧ a \<sqsubset> node_term n' --> (∃p. p ∈ complete_Path a ({n2} ∪ nodes b, {(n1, n2), (n1', n2)} ∪ edges b) ∧ last p = n'))" proof(rule impI,rule allI,rule impI,erule conjE) fix n' assume b1:" unique_originate a n" and b2:"n' ∈ nodes ({n2} ∪ nodes b, {(n1, n2), (n1', n2)} ∪ edges b)" and b3:"a \<sqsubset> node_term n'" show "∃p. p ∈ complete_Path a ({n2} ∪ nodes b, {(n1, n2), (n1', n2)} ∪ edges b) ∧ last p = n'" proof - let ?b'="({n2} ∪ nodes b, {(n1, n2), (n1', n2)} ∪ edges b)" from b2 have "n':nodes b |n'=n2" by(unfold nodes_def,auto) moreover { assume c1:"n'∈nodes b " from c1 b1 b2 b3 a2 have "(∃p. p∈ (complete_Path a b) ∧ (last p)=n') " (is "∃ p. ?P p b") by blast then obtain p where c2:"?P p b" by blast then have c3:"?P p ?b'" thm Complete_path_mono apply - apply(subgoal_tac "(nodes b ) ⊆ nodes ?b' ") apply(subgoal_tac "(edges b ) ⊆ edges ?b'") apply(erule conjE) apply (blast dest:Complete_path_mono) apply(unfold nodes_def edges_def,auto) done then have ?thesis by blast } moreover { assume c1:"n'=n2" from b3 a4 c1 have c2:"a\<sqsubset> node_term n1" apply - apply( unfold casual1_def,auto) done with b1 a3 b3 a4 a2 have "(∃p. p: (complete_Path a b) ∧ (last p)=n1) " (is "∃ p. ?P p b n1 ") by blast then obtain p where c3:"?P p b n1" by blast then have c4:"?P p ?b' n1" thm Complete_path_mono apply - apply(subgoal_tac "(nodes b ) ⊆ nodes ?b' ") apply(subgoal_tac "(edges b ) ⊆ edges ?b'") apply(erule conjE) apply (blast dest:Complete_path_mono) apply(unfold nodes_def edges_def,auto) done have "?P (p@[n2]) ?b' n2" apply - apply(rule conjI) apply(rule_tac ?m1.0="n1" in addnege) apply(cut_tac a4) apply simp apply(cut_tac a4) apply(unfold nodes_def edges_def) apply simp apply simp apply(cut_tac c4) apply(unfold nodes_def edges_def,simp) apply(cut_tac c1 b3,simp) apply(cut_tac c4,simp) apply simp done with c1 have ?thesis by blast } ultimately show ?thesis by blast qed qed next fix b n1 n2 assume a1: "b ∈ bundles" and a2: " unique_originate a n --> (∀n'. n' ∈ nodes b ∧ a \<sqsubset> node_term n' --> (∃p. p∈ (complete_Path a b) ∧ (last p)=n') ) " and a3: "node_sign n2 = - " and a4:"strand n1 ≠ strand n2 ∧ n1 -> n2 ∧ n1 ∈ nodes b ∧ (∀n3. n3 ∈ nodes b --> (n1, n3) ∉ edges b)" and a5:" n2 ∉ nodes b" and a6:" index n2=0" show "unique_originate a n --> (∀n'. n' ∈ nodes ({n2} ∪ nodes b, {(n1, n2)} ∪ edges b) ∧ a \<sqsubset> node_term n' --> (∃p. p ∈ complete_Path a ({n2} ∪ nodes b, {(n1, n2)} ∪ edges b) ∧ last p = n'))" proof(rule impI,rule allI,rule impI,erule conjE) fix n' assume b1:" unique_originate a n" and b2:"n' ∈ nodes ({n2} ∪ nodes b, {(n1, n2)} ∪ edges b)" and b3:"a \<sqsubset> node_term n'" show "∃p. p ∈ complete_Path a ({n2} ∪ nodes b, {(n1, n2)} ∪ edges b) ∧ last p = n'" proof - let ?b'="({n2} ∪ nodes b, {(n1, n2)} ∪ edges b)" from b2 have "n'∈nodes b |n'=n2" by(unfold nodes_def,auto) moreover { assume c1:"n'∈nodes b " from c1 b1 b2 b3 a2 have "(∃p. p∈ (complete_Path a b) ∧ (last p)=n') " (is "∃ p. ?P p b") by blast then obtain p where c2:"?P p b" by blast then have c3:"?P p ?b'" thm Complete_path_mono apply - apply(subgoal_tac "(nodes b ) ⊆ nodes ?b' ") apply(subgoal_tac "(edges b ) ⊆ edges ?b'") apply(erule conjE) apply (blast dest:Complete_path_mono) apply(unfold nodes_def edges_def,auto) done then have ?thesis by blast } moreover { assume c1:"n'=n2" from b3 a4 c1 have c2:"a\<sqsubset> node_term n1" apply - apply( unfold casual1_def,auto) done with b1 a3 b3 a4 a2 have "(∃p. p∈ (complete_Path a b) ∧ (last p)=n1) " (is "∃ p. ?P p b n1 ") by blast then obtain p where c3:"?P p b n1" by blast then have c4:"?P p ?b' n1" thm Complete_path_mono apply - apply(subgoal_tac "(nodes b ) ⊆ nodes ?b' ") apply(subgoal_tac "(edges b ) ⊆ edges ?b'") apply(erule conjE) apply (blast dest:Complete_path_mono) apply(unfold nodes_def edges_def,auto) done have "?P (p@[n2]) ?b' n2" apply - apply(rule conjI) apply(rule_tac ?m1.0="n1" in addnege) apply(cut_tac a4) apply simp apply(cut_tac a4) apply(unfold nodes_def edges_def) apply simp apply simp apply(cut_tac c4) apply(unfold nodes_def edges_def,simp) apply(cut_tac c1 b3,simp) apply(cut_tac c4,simp) apply simp done with c1 have ?thesis by blast } ultimately show ?thesis by blast qed qed qed section{*main lemmas on gerneral authentication tests*} lemma general_auth_test: assumes a1:"b∈bundles" and a2:" unique_originate a n" and a3:"suite M a b" and a4:"n'∈nodes b" and a5:"a\<sqsubset>node_term n'" and a6:" node_term n'∉ (synth M)" and a8:"node_term n :(synth M)" shows "∃ m' m.( n \<preceq>b m ∧ m' \<preceq>b n' ∧ m =>+ m' ∧ node_sign m'=+ ∧( strand m'≠strand n--> node_sign m=- ) ∧ a \<sqsubset> node_term m ∧ a \<sqsubset> node_term m' ∧ node_term m'∉ (synth M) ∧ node_term m∈ (synth M) ∧Is_regular_strand (strand m') ∧(∀ y.(y \<prec>b m') --> ¬ a\<sqsubset> node_term y|(node_term y)∈(synth M))) " (*¬ a\<sqsubset> node_term y|*) proof - let ?P="{ x. (x \<preceq>b n') ∧ a \<sqsubset> node_term x∧ node_term x∉ (synth M) }" from a4 a5 a6 have M0:"n'∈ ?P" by blast from a1 and this have M1:"∃ z. z ∈ ?P ∧ (∀ y.(y \<prec>b z) --> y∉?P)" (is "∃z. z ∈ ?P ∧ ?Q z") apply - apply (drule_tac x="n'" and Q="?P" in bundle_minimal) apply simp apply blast done from this obtain m' where M2: "m'∈?P ∧?Q m'" by auto from a1 a4 M2 have M2a:"m'∈nodes b" apply - apply(simp,(erule conjE)+) apply(blast dest:bundle_is_rtrans_closed23) done from this have M3:" node_sign m'= +" proof(rule_tac P="node_sign m' = +" in ccontr) assume a00:"node_sign m' ≠ +" from a00 have M4:"node_sign m'=-" apply - apply (cut_tac node_sign_is_Plus_or_minus[where n="node_sign m'"], blast) done from a1 and M2 and M4 and M2a have M5:"∃ m''. m''∈nodes b ∧ m'' -> m'∧(m'', m') ∈ edges b" apply - apply ((erule conjE)+,simp only:mem_Collect_eq) apply(erule conjE) thm bundle_edge_is_exist apply( drule_tac b="b" and ?n2.0="m'" in bundle_edge_is_exist) apply blast done from this obtain m'' where M6:"m''∈nodes b ∧ m'' -> m'∧(m'', m') ∈ edges b" by auto from this have M7:"node_term m'=node_term m''" by (unfold casual1_def,simp) from M6 M2 have "(m'',n'):(edges b)^*" apply - apply(blast intro:transitive_closure_trans) done from M7 and M2 and M6 and this have " m''∈?P" apply - apply(simp ) done from this and M2 and M6 show "False" apply - thm r_into_trancl apply(erule conjE)+ apply(drule_tac x="m''" in spec) apply (blast dest:r_into_trancl) done qed from M2 have M4:"a\<sqsubset> node_term m'" by simp have M8a:"∀ m. ((m,m'):casual2^+∧node_sign m=- --> node_term m∈synth M)"(is "∀ m. ?P m") proof(rule allI,rule impI) fix ma assume N1:"(ma, m') ∈ casual2+∧node_sign ma=- " show "node_term ma∈synth M" proof - from N1 a1 M2a have N2:" ma∈ nodes b ∧(ma,m'):(edges b)^+" thm bundle_casual2_trancl apply - apply(rule bundle_casual2_trancl) apply auto done with M2 have N2a:"(ma,n'): (edges b)^*" apply - apply simp apply(rule transitive_closure_trans) apply auto done have "a\<sqsubset> node_term ma | ¬ a \<sqsubset>node_term ma " by blast moreover { assume N3:"a\<sqsubset> node_term ma" from N2 M2 N2a N3 have ?thesis apply - apply(erule conjE)+ apply(drule_tac x="ma" in spec) apply auto done } moreover { assume N3:"¬ a\<sqsubset> node_term ma" from N3 a3 have "node_term ma ∈ M" apply - apply(unfold suite_def,auto) done then have ?thesis by blast } ultimately show ?thesis by blast qed qed from M2 a8 have "m'≠n" apply - apply(rule ccontr,auto) done from this and a2 have "strand n= strand m' ∧ (a \<sqsubset> node_term m' --> index n < index m') | strand n≠ strand m' ∧ non_originate a (strand m')" apply - apply(rule non_originate_node) by auto moreover {assume M9:" strand n≠ strand m' ∧ non_originate a (strand m')" thm non_originate_node1 have "∃m. first_node_in_nonorigi_strand a m' m" apply - apply(cut_tac a2 M9 M4) apply(blast intro!:non_originate_node1) done then obtain m where M10:"first_node_in_nonorigi_strand a m' m" by blast then have M11:"a \<sqsubset> node_term m" by(unfold first_node_in_nonorigi_strand_def,auto) from M10 M4 have M12:"index m <=index m'" apply - apply(rule nonOrigNode1) apply auto done thm bundle_casual2_closed3 have "(strand m', index m) ∈ nodes b ∧ ((strand m', index m), m') ∈ (edges b ∩ casual2)*" apply - apply(cut_tac a1 M12 M2a) apply(rule bundle_casual2_closed3) by auto with M10 have M13:"m∈nodes b∧(m,m'):(edges b ∩ casual2)*" apply - apply(subgoal_tac "(strand m', index m)=m",simp) apply(unfold first_node_in_nonorigi_strand_def) apply(simp add:node_equal) apply(unfold strand_def index_def,auto) done then have M14:"(m,m'):( casual2)*" apply - thm rtrancl_mono apply(subgoal_tac "(edges b ∩ casual2)^*⊆ casual2^*") apply (blast intro!: rtrancl_mono)+done have M15: "m≠m'" apply - apply(rule ccontr) apply(cut_tac M3 M10) apply(unfold first_node_in_nonorigi_strand_def, auto) done with M14 have M16:"(m,m'):( casual2)+" apply - apply(blast dest:rtranclD ) done from M13 have M17:"(m,m'):(edges b)*" apply - thm rtrancl_mono apply(subgoal_tac "(edges b ∩ casual2)^*⊆ (edges b)^*") apply (blast intro!: rtrancl_mono)+done with M15 have M18:"(m,m'):( edges b)+" apply - apply(blast dest:rtranclD ) done from M17 M2 M14 have M19:"(m,n'):(edges b)^*" by(blast dest:transitive_closure_trans) from M19 M18 M11 M2 have M20:"node_term m∈synth M" apply - apply(erule conjE) apply(drule_tac x="m" in spec) by blast from M16 have M21:"∃ m.(m,m'):casual2^+" by blast have M22:"Is_regular_strand (strand m')" apply - thm another_penetrator_strand apply(rule another_penetrator_strand) apply(cut_tac M2) apply blast apply(cut_tac M3) apply simp apply assumption+ apply(cut_tac M21, assumption+) apply(cut_tac M8a, assumption+) apply(cut_tac M2a, auto) done from M11 and a1 and a2 and M13 have M23: "(∃ p. p∈ (complete_Path a b) ∧ (last p)=m)" (is "∃ p. ?R p") thm path_exist by(blast dest:path_exist) from this obtain p where "?R p" by auto from a1 this a2 M13 have M24: "n=m|(n,m):(edges b)^+" thm completePathProp apply - apply((erule conjE), drule completePathProp) apply blast done from this have M24:"(n,m):(edges b)^*" by (blast dest:trancl_into_rtrancl) have ?thesis apply - apply(rule_tac x="m'" in exI) apply(rule_tac x="m" in exI) apply(cut_tac M22 M20 M19 M16 M2 M3 M10 M24 M10) apply(unfold first_node_in_nonorigi_strand_def) apply(rule conjI, blast)+ apply(rule allI,rule impI) apply(subgoal_tac "(y,n'):(edges b)^*") apply blast apply(subgoal_tac " (y, n') ∈ (edges b)+") apply simp thm transitive_closure_trans apply simp apply(rule_tac a="y" and b="m'" and c="n'" in trancl_rtrancl_trancl) apply auto done } moreover { assume M9:"strand n= strand m' ∧ (a \<sqsubset> node_term m' --> index n < index m') " from M4 M9 have M10:"index n < index m'" by simp from this have M12:"index n <=index m'" by simp thm bundle_casual2_closed3 have "(strand m', index n) ∈ nodes b ∧ ((strand m', index n), m') ∈ (edges b ∩ casual2)*" apply - apply(cut_tac a1 M12 M2a) apply(rule bundle_casual2_closed3) by auto with M9 have M13:"n:nodes b∧(n,m'):(edges b ∩ casual2)*" apply - apply(subgoal_tac "(strand m', index n)=n",simp) apply(simp add:node_equal) apply(unfold strand_def index_def,auto) done then have M14:"(n,m'):( casual2)*" apply - thm rtrancl_mono apply(subgoal_tac "(edges b ∩ casual2)^*⊆ casual2^*") apply (blast intro!: rtrancl_mono)+done have M15: "n≠m'" apply - apply(rule ccontr) apply(cut_tac M3 M10) apply( auto) done with M14 have M16:"(n,m'):( casual2)+" apply - apply(blast dest:rtranclD ) done then have M17:"∃ m.(m,m'):( casual2)+" by blast have M22:"Is_regular_strand (strand m')" apply - thm another_penetrator_strand apply(rule another_penetrator_strand) apply(cut_tac M2) apply blast apply(cut_tac M3) apply simp apply assumption+ apply(cut_tac M17, assumption+) apply(cut_tac M8a, assumption+) apply(cut_tac M2a, auto) done have ?thesis apply - apply(rule_tac x="m'" in exI) apply(rule_tac x="n" in exI) apply(cut_tac M22 M16 M2 M3 M9 M10 a2 a8) apply(unfold unique_originate_def originate_def) apply(rule conjI,simp)+ apply(rule allI,rule impI) apply(subgoal_tac "(y,n'):(edges b)^*") apply blast apply(subgoal_tac " (y, n') ∈ (edges b)+") apply simp thm transitive_closure_trans apply simp apply(rule_tac a="y" and b="m'" and c="n'" in trancl_rtrancl_trancl) apply auto done } ultimately show ?thesis by blast qed lemma general_auth_test1: assumes a1:"b∈ bundles" and a2:" unique_originate a n" and a3:"suite M a b" and a4:"n'∈ nodes b" and a5:"a\<sqsubset>node_term n'" and a6:" node_term n'∉ (synth M)" and a8:"node_term n ∈(synth M)" and a9:"∀ n0. strand n0=strand n∧ node_sign n0=+∧ a \<sqsubset> node_term n0 ∧ n0∈ nodes b--> (node_term n0∈ (synth M))" shows "∃ m' m.((n,m):(edges b)^* ∧ (m',n'):(edges b)^*∧(m, m'):casual2^+ ∧ node_sign m'=+ ∧( strand m'≠strand n ∧ node_sign m=- ) ∧ a \<sqsubset> node_term m ∧ a \<sqsubset> node_term m' ∧ node_term m'∉ (synth M) ∧ node_term m∈ (synth M) ∧Is_regular_strand (strand m') ∧(∀ y.(y,m'):(edges b)^+ -->¬ a\<sqsubset> node_term y| (node_term y)∈ (synth M)))" apply - apply(cut_tac a1 a2 a3 a4 a5 a6 a8) apply(drule general_auth_test) apply assumption+ apply(erule exE)+ apply(rule_tac x="m'" in exI) apply(rule_tac x="m" in exI) apply(cut_tac a9) apply(case_tac "strand m'=strand n") apply(erule conjE)+ apply(drule_tac x="m'" in spec) apply(subgoal_tac "m'∈ nodes b") apply blast apply(cut_tac a1 a4) apply(blast dest:bundle_is_rtrans_closed23) by (blast ) lemma general_auth_test2: assumes a1:"b∈bundles" and a2:" unique_originate a n" and a3:"suite M a b" and a4:"n'∈nodes b" and a5:"a\<sqsubset>node_term n'" and a6:" node_term n'∉ (synth M)" and a8:"node_term n :(synth M)" shows "∃ m' m.( n \<preceq>b m ∧ m' \<preceq>b n' ∧ m =>+ m' ∧ node_sign m'=+ ∧( strand m'≠strand n--> node_sign m=- ) ∧ a \<sqsubset> node_term m ∧ a \<sqsubset> node_term m' ∧ node_term m'∉ (synth M) ∧ node_term m∈ (synth M) ∧Is_regular_strand (strand m') ∧(∀ y.(y \<prec>b m') --> (node_term y)∈(synth M))) " proof - have "∃ m' m.( n \<preceq>b m ∧ m' \<preceq>b n' ∧ m =>+ m' ∧ node_sign m'=+ ∧( strand m'≠strand n--> node_sign m=- ) ∧ a \<sqsubset> node_term m ∧ a \<sqsubset> node_term m' ∧ node_term m'∉ (synth M) ∧ node_term m∈ (synth M) ∧Is_regular_strand (strand m') ∧(∀ y.(y \<prec>b m') --> ¬ a\<sqsubset> node_term y|(node_term y)∈(synth M))) " (is "∃ m' m. ?P m' m") proof(rule general_auth_test)qed then obtain m' and m where b1:"?P m' m" by blast from this show ?thesis proof(rule_tac x="m'" in exI,rule_tac x="m" in exI,(rule_tac conjI,simp)+) show "∀y. y \<prec>b m' --> node_term y ∈ synth M" proof(rule allI,rule impI) fix y assume c1:" y \<prec>b m'" from b1 c1 have "¬ a\<sqsubset> node_term y|(node_term y)∈(synth M)" by blast then show "node_term y ∈ synth M" apply - apply(erule disjE) apply(subgoal_tac "node_term y ∈ M") apply blast apply(cut_tac a3, unfold suite_def,auto) done qed qed qed lemma synthM: assumes a8:"m ∈(synth M)" and a9:"a=Nonce N ∨ a=Key k'" and a11:"M= {Crypt k h} ∪ {m. ¬ a \<sqsubset> m} " shows " a \<sqsubset>m --> Crypt k h \<sqsubset>m" (is "?P m ") using a8 proof induct fix X assume b1:"X ∈ M " show "?P X" apply(cut_tac b1 a11,unfold parts_relation_def,auto) done next fix agt show "?P (Agent agt)" apply(cut_tac a9 a11,unfold parts_relation_def,auto) done next fix n show "?P (Number n)" apply(cut_tac a9 a11,unfold parts_relation_def,auto) done next fix X show "?P (Hash X)" apply(cut_tac a9 a11,unfold parts_relation_def,auto) done next fix X Y assume c1:"a \<sqsubset> X --> Crypt k h \<sqsubset> X" and c2:" a \<sqsubset> Y --> Crypt k h \<sqsubset> Y" show "?P \<lbrace>X, Y\<rbrace>" proof assume d1:" a \<sqsubset> \<lbrace>X, Y\<rbrace>" have "a \<sqsubset> X | a \<sqsubset> Y " apply(cut_tac a9 d1 ,unfold parts_relation_def,simp) apply(erule disjE) apply simp apply(subgoal_tac "parts {X, Y}=parts{X} ∪ parts{Y}") apply simp apply(force simp add:parts_insert2) apply(subgoal_tac "parts {X, Y}=parts{X} ∪ parts{Y}") apply simp apply(force simp add:parts_insert2) done with c1 c2 show " Crypt k h \<sqsubset> \<lbrace>X, Y\<rbrace>" apply - apply(erule disjE, unfold parts_relation_def) apply(subgoal_tac "parts {X} ⊆ parts {X,Y}",force,simp add:parts_mono ) apply(subgoal_tac "parts {Y} ⊆ parts {X,Y}",force,simp add:parts_mono ) done qed next fix K X assume c1:"a \<sqsubset> X --> Crypt k h \<sqsubset> X" show "?P (Crypt K X)" proof assume d1:" a \<sqsubset> Crypt K X" have "a \<sqsubset> X " apply(cut_tac a9 d1 ,unfold parts_relation_def,auto) done with c1 show " Crypt k h \<sqsubset> (Crypt K X)" by(unfold parts_relation_def,auto) qed qed lemma synthM1: assumes a8:"m ∈(synth M)" and a9:"a=Nonce N ∨ a=Key k'" and a11:"M= {Crypt k h} ∪ {m. ¬ a \<sqsubset> m} " and a12:" a \<sqsubset>m " shows " Crypt k h \<sqsubset>m" (is "?P m ") apply - apply(cut_tac a8 a9 a11 a12,blast dest: synthM) done lemma out_auth_tast: assumes a1:"b∈ bundles" and a2:" unique_originate a n" and a3:"regularK (invKey k) b" and a4:"n'∈ nodes b" and a5:"a\<sqsubset>node_term n'" and a6:"¬ Crypt k h \<sqsubset> node_term n'" and a7:"node_sign n'= -" and a8:"node_term n ∈(synth M)" and a9:"a=Nonce N ∨ a=Key k'" and a10:"M= {Crypt k h} ∪ {m. ¬ a \<sqsubset> m} " and a11:"a\<sqsubset>h" and a12:"n=>n'" and a14:"node_sign n'=-" shows "∃ m' m.( n \<preceq>b m ∧ m' \<preceq>b n' ∧ m =>+ m' ∧ node_sign m'=+ ∧( node_sign m=- ) ∧ a \<sqsubset> node_term m ∧ a \<sqsubset> node_term m' ∧ node_term m'∉ (synth M) ∧ node_term m∈ (synth M) ∧Is_regular_strand (strand m') ∧(∀ y.(y \<prec>b m') -->(node_term y)∈(synth M))) " proof - have b1:"suite M a b" apply(cut_tac a10 a3,unfold suite_def component_def) by auto have b2:" node_term n' ∉ synth M" proof(rule ccontr) assume c1:"¬ node_term n' ∉ synth M " from c1 have c2:"node_term n'∈ synth M " by simp have "Crypt k h \<sqsubset> node_term n'" proof(rule synthM1,cut_tac c2)qed with a6 show False by simp qed have "∃ m' m.( n \<preceq>b m ∧ m' \<preceq>b n' ∧ m =>+ m' ∧ node_sign m'=+ ∧( strand m'≠strand n--> node_sign m=- ) ∧ a \<sqsubset> node_term m ∧ a \<sqsubset> node_term m' ∧ node_term m'∉ (synth M) ∧ node_term m∈ (synth M) ∧Is_regular_strand (strand m') ∧(∀ y.(y \<prec>b m') --> (node_term y)∈(synth M))) " (is "∃ m' m. ?P m' m") proof(cut_tac b1 b2, rule general_auth_test2 )qed then obtain m' and m where b3:"?P m' m" by blast have "strand m'≠ strand n" proof(rule ccontr) assume c1:"¬ strand m'≠ strand n" from c1 have c2:"strand m'=strand n" by simp have c2a:"m' ≠ n' " apply(rule ccontr) apply(cut_tac b3 a14,simp) done have c2b:"m' \<preceq>b n' " apply(cut_tac b3 ,simp) done with c2a have c3:"m' \<prec>b n' " apply - thm rtrancl_eq_or_trancl apply(simp add: rtrancl_eq_or_trancl) done have c4:"index m' < index n'" proof(rule ccontr) assume d1:" ¬ index m' < index n'" with a1 a4 c2b have d2:"m' ∈ nodes b" by (blast dest:bundle_is_rtrans_closed23) from d1 have d3:"index n' <=index m'" by simp thm bundle_casual2_closed3 with d2 a1 have "((strand m', index n'),m') ∈ (edges b ∩ casual2)*" thm bundle_casual2_closed3 proof(drule_tac bundle_casual2_closed3,auto)qed then have e1:"((strand m', index n'),m') ∈ (edges b)*" apply - apply(subgoal_tac " (edges b ∩ casual2)* ⊆ (edges b)*") apply blast apply(rule rtrancl_mono) by blast have e2:"(strand m', index n')=n'" apply(simp add:node_equal) apply(cut_tac c2 a12, unfold casual2_def strand_def index_def,auto) done with e1 have "n' \<preceq>b m'" by simp with c3 have "n' \<prec>b n'" by (blast intro:transitive_closure_trans) with a1 show False by (blast dest:bundle_edge_is_anti2) qed have c5:"index n'=Suc (index n)" by(cut_tac a12,unfold casual2_def, simp) with c4 have c6:"index m' <= index n" by arith have c7:"m'=(strand n, index m')" apply(cut_tac c1,simp add:node_equal) apply(unfold strand_def index_def, auto) done have c8:"index n <=index m'" apply(rule ccontr) apply(cut_tac b3 a2 c1 c2, unfold unique_originate_def originate_def) apply(subgoal_tac "¬ a \<sqsubset>node_term (strand n, index m')") prefer 2 apply simp apply(cut_tac c7) by (simp ) with c6 have "index n =index m'" by simp with c2 have "n=m'" by (simp add:node_equal) with a8 b3 show False by simp qed with b3 show ?thesis apply - apply(rule_tac x="m'" in exI) apply(rule_tac x="m" in exI) by blast qed lemma unsolicited_test: assumes a1:"b∈ bundles" and a2:"n∈ nodes b" and a3:"Crypt k h \<sqsubset> node_term n" and a4:" regularK k b" shows c:"∃ n'. (n',n):(edges b)^* ∧n'∈ (nodes b) ∧ Crypt k h \<sqsubset> node_term n'∧ Is_regular_strand (strand n') ∧ node_sign n'=+ ∧ (∀ y.(y,n'):(edges b)^+ -->¬ Crypt k h \<sqsubset> node_term y)" proof - let ?t="Crypt k h" let ?P="{ x. (x,n):(edges b)^* ∧x∈ nodes b ∧ ?t \<sqsubset> node_term x }" from a1 and a2 and a3 have M1:"∃ z. z ∈ ?P ∧ (∀ y.(y,z)∈ (edges b)^+ -->y∉?P)" (is "∃ z. z ∈ ?P ∧ ?Q z") apply - apply (drule_tac x="n" and Q="?P" in bundle_minimal) apply simp apply blast done from this obtain m' where M2: "m'∈ ?P ∧?Q m'" by auto from this have M3:" node_sign m'= +" proof(rule_tac P="node_sign m' = +" in ccontr) assume a00:"node_sign m' ≠ +" from a00 have M4:"node_sign m'=-" apply - apply (cut_tac node_sign_is_Plus_or_minus[where n="node_sign m'"], blast) done from a1 and M2 and M4 have M5:"∃ m''. m''∈nodes b ∧ m'' -> m'∧(m'', m') ∈ edges b" apply - apply ((erule conjE)+,simp only:mem_Collect_eq) apply(erule conjE) thm bundle_edge_is_exist apply( drule_tac b="b" and ?n2.0="m'" in bundle_edge_is_exist) apply(subgoal_tac "m'∈ nodes b") apply blast apply(cut_tac a1 a2) thm bundle_is_rtrans_closed23 apply(blast dest: bundle_is_rtrans_closed23) done from this obtain m'' where M6: "m''∈nodes b ∧ m'' -> m'∧(m'', m') ∈ edges b" by auto from this have M7:"node_term m'=node_term m''" by (unfold casual1_def,auto) from this and M2 and M6 have " m''∈?P" apply - apply(simp only:mem_Collect_eq) apply(blast dest:transitive_closure_trans ) done from this and M2 and M6 show "False" apply - thm r_into_trancl apply(erule conjE)+ apply(drule_tac x="m''" in spec) apply (blast dest:r_into_trancl)done qed from M2 and a1 thm penetrator_strand have a8:"m'∈Domain" thm bundle_node_in_Domain by(blast dest: bundle_node_in_Domain) from this and a4 have M03:"Is_regular_strand (strand m')" thm ccontr proof(rule_tac ?P="Is_regular_strand (strand m')" in ccontr) assume a6:"¬ Is_regular_strand (strand m') " show False proof - let ?s="strand m'" from a6 have ab1:"Is_penetrator_strand ?s" by(unfold Is_regular_strand_def, simp) from ab1 have "( Is_Tee_strand ?s | Is_Flush_strand ?s | Is_Cat_strand ?s |Is_Sep_strand ?s| Is_E_strand ?s | Is_D_strand ?s |Is_T_strand ?s | Is_K_strand ?s)" by (unfold penetrator_trace,simp) moreover { assume a12:"Is_E_strand (strand m')" from M3 and a8 and a12 have M0:"snd m'=2" apply - apply( unfold Is_E_strand_def, unfold Domain_def) apply( unfold node_sign_def strand_def index_def,auto) apply(case_tac "y=0") apply simp apply (case_tac "y=1" ) apply simp apply auto done from a12 and this have N1:"∃ k h . node_term (fst m',0)=Key k ∧ node_term (fst m', 1)=h ∧ node_term m'=Crypt k h" (is "∃ k h. ?P k h") by (unfold node_term_def, unfold Is_E_strand_def strand_def index_def, auto) then obtain k0 and h0 where N01:" ?P k0 h0" by auto from a12 have N00:"node_sign (fst m', 0)=-∧ node_sign (fst m', 1)=-" apply - apply(unfold Is_E_strand_def, unfold node_sign_def strand_def index_def, auto) done from M0 and a8 have N000:"(fst m',1)∈Domain∧((fst m', 1) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from a1 and M0 and M2 have N2:"(fst m',0)∈ fst b∧ ((fst m',0),m'): (edges b)^+" apply - apply (simp only:mem_Collect_eq,(erule conjE)+, fold strand_def index_def nodes_def, drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply( subgoal_tac "((strand m', 0), m') ∈ (edges b )^*" ) prefer 2 apply(blast intro:rtrancl_mono1) apply(thin_tac "(m', n) ∈ (edges b)*") apply(drule_tac R="edges b" in rtranclD) apply (erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done with M2 have N2a:"((fst m',0),n): (edges b)^*" apply - apply(subgoal_tac "((fst m',0),n): (edges b)^+") prefer 2 apply(blast intro:transitive_closure_trans ) apply(blast dest:trancl_into_rtrancl) done from a1 and M0 and M2 have N3:"(fst m',1)∈ fst b ∧ ((fst m',1),m'): (edges b)^+" apply - apply (simp only:mem_Collect_eq,(erule conjE)+, fold strand_def index_def nodes_def, drule_tac i="1" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply( subgoal_tac "((strand m', 1), m') ∈ (edges b )^*" ) prefer 2 apply(blast intro:rtrancl_mono1) apply(thin_tac "(m', n) ∈ (edges b)*") apply(drule_tac R="edges b" in rtranclD) apply (erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done with M2 have N3a:"((fst m',1),n): (edges b)^*" apply - apply(subgoal_tac "((fst m',1),n): (edges b)^+") prefer 2 apply(blast intro:transitive_closure_trans ) apply(blast dest:trancl_into_rtrancl) done from N01 and M2 have "?t\<sqsubset> h0 | ?t=Crypt k0 h0" apply- apply simp apply(erule conjE)+ apply(drule body_subterm) apply( auto) done moreover {assume a13:" ?t \<sqsubset> h0" from N01 and a13 have N4:" ?t \<sqsubset> node_term (fst m', 1)" by simp from this and M2 and N3 N3a have False by (unfold nodes_def,blast ) from this have ?thesis by blast } moreover {assume a13:" ?t=Crypt k0 h0" from N01 and a13 have N4:"node_term ( fst m', 0)=Key k" by simp from this and a4 and ab1 and N2 N2a have False apply - apply(unfold regularK_def Is_regular_strand_def nodes_def strand_def) apply (drule_tac x="(fst m',0)" in bspec,blast ) apply auto done from this have ?thesis by blast } ultimately have ?thesis .. } moreover { assume a12:"Is_D_strand (strand m')" from M3 and a8 and a12 have M0:"snd m'=2" apply - apply( unfold Is_D_strand_def, unfold Domain_def, unfold node_sign_def strand_def index_def,auto) apply(case_tac "y=0") apply simp apply (case_tac "y=1" ) apply simp apply auto done from a12 and this have M1:"∃ k. ∃ k'. ∃ h. k'=invKey k∧ node_term (fst m',0)= (Key k')∧node_term (fst m',1)= (Crypt k h)∧ node_term m'=h" (is "∃ k k' h. ?P k k' h") by (unfold node_term_def, unfold Is_D_strand_def strand_def index_def, auto) then obtain k0 and k0' and h0 where M01:" ?P k0 k0' h0" by auto from a12 have M00:"node_sign (fst m', 0)=-∧ node_sign (fst m', 1)=-" apply - apply(unfold Is_D_strand_def,unfold node_sign_def strand_def index_def, auto) done from M0 and a8 have M000:"(fst m',1)∈Domain∧((fst m', 1) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from a1 and M0 and M2 have N3:"(fst m',1)∈ fst b ∧ ((fst m',1),m'): (edges b)^+" apply - apply (simp only:mem_Collect_eq,(erule conjE)+, fold strand_def index_def nodes_def, drule_tac i="1" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply( subgoal_tac "((strand m', 1), m') ∈ (edges b )^*" ) prefer 2 apply(blast intro:rtrancl_mono1) apply(thin_tac "(m', n) ∈ (edges b)*") apply(drule_tac R="edges b" in rtranclD) apply (erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done with M2 have N3a:"((fst m',1),n): (edges b)^*" apply - apply(subgoal_tac "((fst m',1),n): (edges b)^+") prefer 2 apply(blast intro:transitive_closure_trans ) apply(blast dest:trancl_into_rtrancl) done from M01 and M2 have "?t \<sqsubset> h0" by simp from this and M01 have "?t \<sqsubset> node_term (fst m',1)" by (unfold parts_relation_def,simp) from this and M2 and N3 N3a have False apply - apply(unfold regularK_def Is_regular_strand_def nodes_def strand_def) apply(erule conjE)+ apply(drule_tac x="(fst m',1)" in spec ) apply auto done from this have ?thesis by blast } moreover {assume a12:"Is_Cat_strand (strand m')" from M3 and a8 and a12 have M0:"snd m'=2" apply - apply( unfold Is_Cat_strand_def, unfold Domain_def, unfold node_sign_def strand_def index_def,auto) apply(case_tac "y=0") apply simp apply (case_tac "y=1" ) apply simp apply auto done from a12 and this have M1:"∃ g h . node_term (fst m',0)=g ∧ node_term (fst m', 1)=h ∧ node_term m'={|g,h|}" (is "∃ g h. ?P g h") by (unfold node_term_def, unfold Is_Cat_strand_def strand_def index_def, auto) then obtain g0 and h0 where M01:" ?P g0 h0" by auto from M0 and a8 have N000:"(fst m',1)∈Domain∧((fst m', 1) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from a1 and M0 and M2 have N2:"(fst m',0)∈ fst b∧ ((fst m',0),m'): (edges b)^+" apply - apply (simp only:mem_Collect_eq,(erule conjE)+, fold strand_def index_def nodes_def, drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply( subgoal_tac "((strand m', 0), m') ∈ (edges b )^*" ) prefer 2 apply(blast intro:rtrancl_mono1) apply(thin_tac "(m', n) ∈ (edges b)*") apply(drule_tac R="edges b" in rtranclD) apply (erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done with M2 have N2a:"((fst m',0),n): (edges b)^*" apply - apply(subgoal_tac "((fst m',0),n): (edges b)^+") prefer 2 apply(blast intro:transitive_closure_trans ) apply(blast dest:trancl_into_rtrancl) done from a1 and M0 and M2 have N3:"(fst m',1)∈ fst b ∧ ((fst m',1),m'): (edges b)^+" apply - apply (simp only:mem_Collect_eq,(erule conjE)+, fold strand_def index_def nodes_def, drule_tac i="1" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply( subgoal_tac "((strand m', 1), m') ∈ (edges b )^*" ) prefer 2 apply(blast intro:rtrancl_mono1) apply(thin_tac "(m', n) ∈ (edges b)*") apply(drule_tac R="edges b" in rtranclD) apply (erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done with M2 have N3a:"((fst m',1),n): (edges b)^*" apply - apply(subgoal_tac "((fst m',1),n): (edges b)^+") prefer 2 apply(blast intro:transitive_closure_trans ) apply(blast dest:trancl_into_rtrancl) done have "?t ≠ {|g0,h0|}" by simp from this and M01 and M2 have "?t\<sqsubset> g0 | ?t \<sqsubset> h0" apply- apply (simp only:mem_Collect_eq) apply(erule conjE)+ apply(drule part_MPair) apply( auto) done moreover {assume a14:"?t \<sqsubset> g0" from M01 and a14 have N4:" ?t \<sqsubset> node_term (fst m', 0)" by simp from this and M2 and N2 N2a have False apply - apply(unfold regularK_def Is_regular_strand_def nodes_def strand_def) apply(erule conjE)+ apply(drule_tac x="(fst m',0)" in spec ) apply auto done from this have ?thesis by blast } moreover {assume a13:" ?t \<sqsubset> h0" from M01 and a13 have N4:"?t \<sqsubset> node_term ( fst m', 1)" by simp from this and M2 and N3 N3a have False apply - apply(unfold regularK_def Is_regular_strand_def nodes_def strand_def) apply(erule conjE)+ apply(drule_tac x="(fst m',1)" in spec ) apply auto done from this have ?thesis by blast } ultimately have ?thesis by blast } moreover {assume a12:"Is_Sep_strand (strand m')" from M3 and a8 and a12 have M0:"snd m'=2|snd m'=1" apply - apply( unfold Is_Sep_strand_def, unfold Domain_def, unfold node_sign_def strand_def index_def,auto) apply(case_tac "y=0") apply simp apply (case_tac "y=1" ) apply simp apply auto done from a12 have M00:"node_sign (fst m', 0)=-∧ node_sign (fst m', 1)=+∧ node_sign (fst m', 2)=+" apply - apply(unfold Is_Sep_strand_def, unfold node_sign_def strand_def index_def, auto) done from a12 and this have M1:"∃ g h . node_term (fst m',0) ={|g,h|} ∧ node_term (fst m', 1)=g ∧ node_term (fst m',2)=h" (is "∃ g h. ?P g h") by (unfold node_term_def, unfold Is_Sep_strand_def strand_def index_def, auto) then obtain g0 and h0 where M01:" ?P g0 h0" by auto from M0 have M0:"snd m'=2|snd m'=1" by simp moreover {assume a20:"snd m'=2" from a20 and a8 have N000:"(fst m',1):Domain∧((fst m', 1) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from a1 and M0 and M2 a20 have N2:"(fst m',0): fst b∧ ((fst m',0),m'): (edges b)^+" apply - apply (simp only:mem_Collect_eq,(erule conjE)+, fold strand_def index_def nodes_def, drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply( subgoal_tac "((strand m', 0), m') ∈ (edges b )^*" ) prefer 2 apply(blast intro:rtrancl_mono1) apply(thin_tac "(m', n) ∈ (edges b)*") apply(drule_tac R="edges b" in rtranclD) apply (erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done with M2 have N2a:"((fst m',0),n): (edges b)^*" apply - apply(subgoal_tac "((fst m',0),n): (edges b)^+") prefer 2 apply(blast intro:transitive_closure_trans ) apply(blast dest:trancl_into_rtrancl) done from M2 and M01 and a20 have "?t \<sqsubset> h0" by (subgoal_tac "(fst m',2)=m'",simp, simp add:node_equal, unfold strand_def index_def,auto) from this and M2 and M01 and a20 have "?t \<sqsubset> node_term (fst m',0)" apply - apply(simp) apply( subgoal_tac "h0 \<sqsubset> node_term (fst m',0)") thm subterm_trans apply(drule_tac g="Crypt k h" and l=" node_term (fst m', 0)" in subterm_trans) apply(erule conjE)+ apply simp+ apply(unfold parts_relation_def, simp add:parts_insert_MPair parts_insert2) apply blast done from this and M2 and N2 N2a have False apply - apply(unfold regularK_def Is_regular_strand_def nodes_def strand_def) apply(erule conjE)+ apply(drule_tac x="(fst m',0)" in spec ) apply auto done from this have ?thesis by blast } moreover {assume a20:"snd m'=1" from a20 and a8 have N000:"(fst m',0):Domain∧((fst m', 0) ,m'):casual2^+" apply- apply( rule casual2_trancl) apply simp apply simp done from a1 and M0 and M2 a20 have N2:"(fst m',0): fst b∧ ((fst m',0),m'): (edges b)^+" apply - apply (simp only:mem_Collect_eq,(erule conjE)+, fold strand_def index_def nodes_def, drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply( subgoal_tac "((strand m', 0), m') ∈ (edges b )^*" ) prefer 2 apply(blast intro:rtrancl_mono1) apply(thin_tac "(m', n) ∈ (edges b)*") apply(drule_tac R="edges b" in rtranclD) apply simp apply (erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done with M2 have N2a:"((fst m',0),n): (edges b)^*" apply - apply(subgoal_tac "((fst m',0),n): (edges b)^+") prefer 2 apply(blast intro:transitive_closure_trans ) apply(blast dest:trancl_into_rtrancl) done from M2 and M01 and a20 have "?t \<sqsubset> g0" apply - apply(subgoal_tac "(fst m',1)=m'",simp, simp add:node_equal, unfold strand_def index_def,auto) done from this and M2 and M01 and a20 have "?t \<sqsubset> node_term (fst m',0)" apply - apply(simp) apply( subgoal_tac "g0 \<sqsubset> node_term (fst m',0)") thm subterm_trans apply(drule_tac g="Crypt k h" and l=" node_term (fst m', 0)" in subterm_trans) apply(erule conjE)+ apply simp+ apply(unfold parts_relation_def, simp add:parts_insert_MPair parts_insert2) apply blast done from this and M2 and N2 N2a have False apply - apply(unfold regularK_def Is_regular_strand_def nodes_def strand_def) apply(erule conjE)+ apply(drule_tac x="(fst m',0)" in spec ) apply auto done from this have ?thesis by blast } ultimately have ?thesis by blast } moreover {assume a9:"Is_Tee_strand ?s" have N0:"(fst m',0): fst b∧ ((fst m',0),m'): (edges b)^+" proof - from M3 and a8 and a9 have M0: "snd m' =2 | snd m'=1" apply - apply( unfold Is_Tee_strand_def, unfold Domain_def strand_def index_def, auto) apply(case_tac "y=0") apply(unfold node_sign_def, simp) apply auto done moreover {assume a10:"snd m'=2" from a1 and M0 and M2 and a10 have ?thesis apply - apply (simp only:mem_Collect_eq,(erule conjE)+, fold strand_def index_def nodes_def, drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply( subgoal_tac "((strand m', 0), m') ∈ (edges b )^*" ) prefer 2 apply(blast intro:rtrancl_mono1) apply(thin_tac "(m', n) ∈ (edges b)*") apply(drule_tac R="edges b" in rtranclD) apply simp apply (erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done } moreover {assume a10:"snd m'=1" from a1 and M0 and M2 a10 have ?thesis apply - apply (simp only:mem_Collect_eq,(erule conjE)+, fold strand_def index_def nodes_def, drule_tac i="0" and ?n2.0="m'" in bundle_casual2_closed3) apply( simp, arith) apply( subgoal_tac "((strand m', 0), m') ∈ (edges b )^*" ) prefer 2 apply(blast intro:rtrancl_mono1) apply(thin_tac "(m', n) ∈ (edges b)*") apply(drule_tac R="edges b" in rtranclD) apply simp apply (erule disjE) apply( simp add:node_equal) apply(unfold strand_def index_def) apply auto done } ultimately show ?thesis by blast qed with M2 have N2a:"((fst m',0),n): (edges b)^*" apply - apply(subgoal_tac "((fst m',0),n): (edges b)^+") prefer 2 apply(blast intro:transitive_closure_trans ) apply(blast dest:trancl_into_rtrancl) done from M3 and a8 and a9 have M0: "snd m' =2 | snd m'=1" apply - apply( unfold Is_Tee_strand_def, unfold Domain_def strand_def index_def , auto) apply(case_tac "y=0") apply(unfold node_sign_def, simp) apply auto done from a9 and this have "∃ g . node_term (fst m',0)=g ∧ node_term (fst m', 1)=g ∧ node_term m'=g" (is "∃ g. ?P g") apply - apply (unfold node_term_def, unfold Is_Tee_strand_def strand_def index_def , auto) done then obtain g where M01:" ?P g" .. from M2 and this have "?t \<sqsubset> node_term (fst m', 0)" by simp from this and N0 and M2 N2a have False apply - apply(unfold regularK_def Is_regular_strand_def nodes_def strand_def) apply(erule conjE)+ apply(drule_tac x="(fst m',0)" in spec ) apply auto done then have ?thesis by blast } moreover {assume a9:"Is_T_strand ?s" from M3 and a9 and a8 have N0:"snd m'=0" apply - apply (unfold Is_T_strand_def, unfold node_sign_def strand_def index_def, unfold Domain_def ) apply auto done from a9 and this have "∃ g. g ∈ T ∧ node_term (fst m',0)=g " (is "∃ g. g ∈ T ∧ ?P g") apply - apply (unfold node_term_def, unfold Is_T_strand_def strand_def index_def, auto) done then obtain g where M01:" g:T∧?P g" by blast from N0 and M2 and this have "?t \<sqsubset> g" apply - apply (subgoal_tac "(fst m',0)=m'", simp, simp add:node_equal, unfold strand_def index_def,auto) done from this and M01 have ?thesis apply - apply (simp only:T_is_only_text) apply(unfold Atoms_def, unfold parts_relation_def,auto) done } moreover {assume a10 :"Is_K_strand (strand m')" from M3 and a10 and a8 have N0:"snd m'=0" apply - apply (unfold Is_K_strand_def, unfold node_sign_def strand_def index_def, unfold Domain_def strand_def index_def) apply auto done from a10 and this have "∃ k. k∈ KP ∧ node_term (fst m',0)=Key k " (is "∃ k. k∈ KP ∧ ?P k") apply - apply (unfold node_term_def, unfold Is_K_strand_def index_def strand_def , auto) done then obtain k0 where M01:" k0:KP∧?P k0" by blast from N0 and M2 and this have "?t \<sqsubset> Key k0" apply - apply (subgoal_tac "(fst m',0)=m'", simp, simp add:node_equal, unfold strand_def index_def, auto) done from this and M01 have ?thesis apply - apply(unfold parts_relation_def,auto) done } moreover { assume a11:"Is_Flush_strand (strand m')" from M3 and a11 and a8 have ?thesis apply - apply(unfold Is_Flush_strand_def,unfold Domain_def, unfold node_sign_def, unfold Domain_def strand_def index_def,auto) done } ultimately show ?thesis by blast qed qed from a1 and M2 and M3 and this show ?thesis apply - apply(rule_tac x="m'" in exI ) apply(rule conjI,simp)+ apply(rule allI)+ apply(erule conjE)+ apply(drule_tac x="y" in spec) apply(rule impI) thm bundle_is_trans_closed22 apply(subgoal_tac "(y,n):(edges b)^*") apply(drule bundle_is_trans_closed22 ) apply simp apply simp thm transitive_closure_trans apply(subgoal_tac "(y,m'):(edges b)^*") apply(blast intro:transitive_closure_trans ) apply(blast dest:trancl_into_rtrancl) done qed lemma in_auth_test: assumes a1:"b:bundles" and a2:" unique_originate a n" and a3:"suite M a b" and a4:"n':nodes b" and a5:"a\<sqsubset>h" and a6:"Crypt k h \<sqsubset> node_term n'" and a8:"¬Crypt k h \<sqsubset> node_term n" and a9:"regularK k b" shows "∃ m' m.((n,m):(edges b)^* ∧ (m',n'):(edges b)^*∧ m=>+m' ∧ node_sign m'=+ ∧( strand m'≠strand n--> node_sign m=- ) ∧ ¬ Crypt k h \<sqsubset> node_term m ∧ Crypt k h \<sqsubset> node_term m' ∧ ¬ Crypt k h \<sqsubset> node_term m ∧ Crypt k h \<sqsubset> node_term m' ∧Is_regular_strand (strand m')) " proof - have b1:"∃ m'. (m',n'):(edges b)^* ∧m': (nodes b) ∧ Crypt k h \<sqsubset> node_term m'∧ Is_regular_strand (strand m') ∧ node_sign m'=+ ∧ (∀ y.(y,m'):(edges b)^+ -->¬ Crypt k h \<sqsubset> node_term y)" (is "∃ m'. ?P m'") by(rule unsolicited_test) then obtain m' where b1:"?P m'" by blast from b1 a5 have M4:"a \<sqsubset> node_term m'" thm subterm_trans apply(rule_tac h="h" in subterm_trans) apply auto apply(rule_tac h="Crypt k h" in subterm_trans) apply auto apply(unfold parts_relation_def,auto) done from b1 a8 have b2:"m'≠n" apply - apply(rule ccontr,auto) done from this and a2 have "strand n= strand m' ∧ (a \<sqsubset> node_term m' --> index n < index m') | strand n≠ strand m' ∧ non_originate a (strand m')" apply - apply(rule non_originate_node) by auto moreover {assume M9:" strand n≠ strand m' ∧ non_originate a (strand m')" thm non_originate_node1 have "∃m. first_node_in_nonorigi_strand a m' m" apply - apply(cut_tac a2 M9 M4) apply(blast intro!:non_originate_node1) done then obtain m where M10:"first_node_in_nonorigi_strand a m' m" by blast then have M11:"a \<sqsubset> node_term m" by(unfold first_node_in_nonorigi_strand_def,auto) from M10 M4 have M12:"index m <=index m'" apply - apply(rule nonOrigNode1) apply auto done thm bundle_casual2_closed3 have "(strand m', index m) ∈ nodes b ∧ ((strand m', index m), m') ∈ (edges b ∩ casual2)*" apply - apply(cut_tac a1 M12 b1) apply(rule bundle_casual2_closed3) by auto with M10 have M13:"m:nodes b∧(m,m'):(edges b ∩ casual2)*" apply - apply(subgoal_tac "(strand m', index m)=m",simp) apply(unfold first_node_in_nonorigi_strand_def) apply(simp add:node_equal) apply(unfold strand_def index_def,auto) done then have M14:"(m,m'):( casual2)*" apply - thm rtrancl_mono apply(subgoal_tac "(edges b ∩ casual2)^*⊆ casual2^*") apply (blast intro!: rtrancl_mono)+done have M15: "m≠m'" apply - apply(rule ccontr) apply(cut_tac b1 M10) apply(unfold first_node_in_nonorigi_strand_def, auto) done with M14 have M16:"(m,m'):( casual2)+" apply - apply(blast dest:rtranclD ) done from M13 have M17:"(m,m'):(edges b)*" apply - thm rtrancl_mono apply(subgoal_tac "(edges b ∩ casual2)^*⊆ (edges b)^*") apply (blast intro!: rtrancl_mono)+done with M15 have M18:"(m,m'):( edges b)+" apply - apply(blast dest:rtranclD ) done from M17 b1 M14 have M19:"(m,n'):(edges b)^*" by(blast dest:transitive_closure_trans) from M19 M18 M11 b1 have M20:"¬ Crypt k h \<sqsubset> node_term m" apply - apply(erule conjE)+ apply(drule_tac x="m" in spec) by blast from M16 have M21:"∃ m.(m,m'):casual2^+" by blast have M22:"Is_regular_strand (strand m')" apply(cut_tac b1) apply blast done from M11 and a1 and a2 and M13 have M23: "(∃ p. p: (complete_Path a b) ∧ (last p)=m)" (is "∃ p. ?R p") thm path_exist by(blast dest:path_exist) from this obtain p where "?R p" by auto from a1 this a2 M13 have M24: "n=m|(n,m):(edges b)^+" thm completePathProp apply - apply((erule conjE), drule completePathProp) apply blast done from this have M24:"(n,m):(edges b)^*" by (blast dest:trancl_into_rtrancl) have ?thesis apply - apply(rule_tac x="m'" in exI) apply(rule_tac x="m" in exI) apply(cut_tac M22 M20 M19 M16 b1 M10 M24 M10) apply(unfold first_node_in_nonorigi_strand_def) apply(rule conjI, blast)+ apply auto done } moreover { assume M9:"strand n= strand m' ∧ (a \<sqsubset> node_term m' --> index n < index m') " from M4 M9 have M10:"index n < index m'" by simp from this have M12:"index n <=index m'" by simp thm bundle_casual2_closed3 have "(strand m', index n) ∈ nodes b ∧ ((strand m', index n), m') ∈ (edges b ∩ casual2)*" apply - apply(cut_tac a1 M12 b1) apply(rule bundle_casual2_closed3) by auto with M9 have M13:"n∈ nodes b∧(n,m'):(edges b ∩ casual2)*" apply - apply(subgoal_tac "(strand m', index n)=n",simp) apply(simp add:node_equal) apply(unfold strand_def index_def,auto) done then have M14:"(n,m'):( casual2)*" apply - thm rtrancl_mono apply(subgoal_tac "(edges b ∩ casual2)^*⊆ casual2^*") apply (blast intro!: rtrancl_mono)+done have M15: "n≠m'" apply - apply(rule ccontr) apply(cut_tac b1 M10) apply( auto) done with M14 have M16:"(n,m'):( casual2)+" apply - apply(blast dest:rtranclD ) done then have M17:"∃ m.(m,m'):( casual2)+" by blast have M22:"Is_regular_strand (strand m')" apply(cut_tac b1) apply blast done have ?thesis apply - apply(rule_tac x="m'" in exI) apply(rule_tac x="n" in exI) apply(cut_tac M22 M16 b1 M9 M10 a2 a8) apply(unfold unique_originate_def originate_def) apply(rule conjI,simp)+ apply auto done } ultimately show ?thesis by blast qed end
lemma casual2_rel_more:
n1.0 => n2.0 ==> n2.0 = (strand n1.0, index n1.0 + 1)
lemma casual2_aux1:
strand n = a ∧ index n = b ==> n = (a, b)
lemma casual2_rel_less:
n1.0 => n2.0 ==> n1.0 = (strand n2.0, snd n2.0 - 1)
lemma casual2_trancl_less:
m =>+ m' ==> index m < index m' ∧ strand m = strand m'
lemma casual2_rtrancl1:
[|n2.0 ∈ strand05.Domain; i ≤ index n2.0|] ==> (strand n2.0, index n2.0 - i) ∈ strand05.Domain ∧ (strand n2.0, index n2.0 - i) =>* n2.0
lemma casual2_rtrancl2:
[|n2.0 ∈ strand05.Domain; i ≤ index n2.0|] ==> (strand n2.0, i) ∈ strand05.Domain ∧ (strand n2.0, i) =>* n2.0
lemma casual2_trancl:
[|n2.0 ∈ strand05.Domain; i < snd n2.0|] ==> (fst n2.0, i) ∈ strand05.Domain ∧ (fst n2.0, i) =>+ n2.0
lemma casual1_casual2_disjoint:
n1.0 -> n2.0 ∧ n3.0 => n4.0 ==> n1.0 ≠ n3.0 ∨ n2.0 ≠ n4.0
lemma send_node_is_plus:
n1.0 -> n2.0 ==> node_sign n1.0 = + ∧ n1.0 ∈ strand05.Domain
lemma recv_node_is_minus:
n2.0 -> n1.0 ==> node_sign n1.0 = - ∧ n1.0 ∈ strand05.Domain
lemma casual2_rel_n1_n2_less:
n1.0 => n2.0 ==> strand n1.0 = strand n2.0 ∧ index n1.0 = index n2.0 - 1
lemma casual_diff:
n1.0 -> n2.0 ∨ n1.0 => n2.0 ==> n1.0 ≠ n2.0
lemma is_casual2:
[|m ∈ strand05.Domain; n ∈ strand05.Domain; strand n = strand m; index n = Suc (index m)|] ==> m => n
lemma casual1_judge_sign:
n1.0 -> n2.0 ==> strand n1.0 ≠ strand n2.0
lemma casual2_judge_sign:
n1.0 => n2.0 ==> strand n1.0 = strand n2.0
lemma casul2_trancl_sign:
n1.0 =>+ n2.0 ==> strand n2.0 = strand n1.0 ∧ index n1.0 < index n2.0
lemma node_equal:
(n = n') = (strand n = strand n' ∧ index n = index n')
lemma node_sign_is_Plus_or_minus:
n = + ∨ n = -
lemma rtrancl_mono1:
[|x ∈ r*; r ⊆ s|] ==> x ∈ s*
lemma part_MPair:
[|a ≠ \<lbrace>g, h\<rbrace>; a \<sqsubset> \<lbrace>g, h\<rbrace>|] ==> a \<sqsubset> g ∨ a \<sqsubset> h
lemma subterm_atom:
[|a \<sqsubset> t; Is_atom t|] ==> a = t
lemma subterm_trans:
[|g \<sqsubset> h; h \<sqsubset> l|] ==> g \<sqsubset> l
lemma subterm_itself:
h \<sqsubset> h
lemma body_subterm:
t'' \<sqsubset> Crypt k h ==> t'' = Crypt k h ∨ t'' \<sqsubset> h
lemma c:
b ∈ bundles ==> n ∈ nodes b --> n ∈ strand05.Domain
lemma bundle_nodes_strand_in_space:
[|b ∈ bundles; m ∈ nodes b|] ==> strand m ∈ Σ
lemmas pair_is_a_bundle:
[|(a, b) ∈ bundles; [|a = {}; b = {}|] ==> P; !!b n1 n2. [|b ∈ bundles; node_sign n2 = +; n2 ∈ strand05.Domain; n2 ∉ nodes b; 0 < index n2; n1 ∈ nodes b; n1 => n2; a = insert n2 (nodes b); b = insert (n1, n2) (edges b)|] ==> P; !!b n2. [|b ∈ bundles; node_sign n2 = +; n2 ∉ nodes b; n2 ∈ strand05.Domain; index n2 = 0; a = insert n2 (nodes b); b = edges b|] ==> P; !!b n1 n1' n2. [|b ∈ bundles; node_sign n2 = -; n2 ∉ nodes b; 0 < index n2; n1' ∈ nodes b; n1' => n2; strand n1 ≠ strand n2; n1 -> n2; n1 ∈ nodes b; ∀a ba. (a, ba) ∈ nodes b --> (n1, a, ba) ∉ edges b; a = insert n2 (nodes b); b = insert (n1, n2) (insert (n1', n2) (edges b))|] ==> P; !!b n1 n2. [|b ∈ bundles; node_sign n2 = -; n2 ∉ nodes b; index n2 = 0; strand n1 ≠ strand n2; n1 -> n2; n1 ∈ nodes b; ∀a ba. (a, ba) ∈ nodes b --> (n1, a, ba) ∉ edges b; a = insert n2 (nodes b); b = insert (n1, n2) (edges b)|] ==> P|] ==> P
lemma bundle_edge_is_exist:
b ∈ bundles ==> node_sign n2.0 = - ∧ n2.0 ∈ nodes b --> (∃n1. n1 -> n2.0 ∧ n1 ∈ nodes b ∧ (n1, n2.0) ∈ edges b)
lemma bundle_is_cl_under_edge:
b ∈ bundles ==> (n1.0, n2.0) ∈ edges b --> n1.0 ∈ nodes b ∧ n2.0 ∈ nodes b
lemma bundle_is_cl_under_edge1:
[|b ∈ bundles; (n1.0, n2.0) ∈ edges b|] ==> n1.0 ∈ nodes b ∧ n2.0 ∈ nodes b
lemma bundle_is_not_close_edge:
b ∈ bundles ==> n1.0 ∈ nodes b ∧ n2.0 ∉ nodes b --> (n1.0, n2.0) ∉ edges b
lemma bundle_is_not_close_edge1:
[|b ∈ bundles; n1.0 ∈ nodes b; n2.0 ∉ nodes b|] ==> (n1.0, n2.0) ∉ edges b
lemma bundle_is_closed:
b ∈ bundles ==> (y, z) ∈ edges b --> y ∈ nodes b ∧ z ∈ nodes b
lemma bundle_is_closed1:
b ∈ bundles ==> (y, z) ∈ edges b --> y ≠ z
lemma bundle_casual1_prec_is_unique:
b ∈ bundles ==> node_sign n2.0 = - ∧ n1.0 -> n2.0 ∧ n1' -> n2.0 ∧ (n1.0, n2.0) ∈ edges b ∧ (n1', n2.0) ∈ edges b --> n1.0 = n1'
lemma bundle_casual2_prec_is_unique:
b ∈ bundles ==> n1.0 => n2.0 ∧ n2.0 ∈ nodes b --> (n1.0, n2.0) ∈ edges b ∧ n1.0 ∈ nodes b
lemma bundle_node_is_in_Domains:
b ∈ bundles ==> n ∈ nodes b --> n ∈ strand05.Domain
lemma bundle_casual2_prec_is_unique_intro1:
[|b ∈ bundles; n1.0 => n2.0; n2.0 ∈ nodes b|] ==> n1.0 ∈ nodes b ∧ n1.0 ∈ strand05.Domain
lemma bundle_casual2_closed_aux1:
[|0 < index n2.0 - n; n2.0 ∈ strand05.Domain|] ==> ((strand n2.0, index n2.0 - Suc n), strand n2.0, index n2.0 - n) ∈ casual2
lemma bundle_casual2_closed:
[|b ∈ bundles; n2.0 ∈ nodes b; i ≤ index n2.0|] ==> (strand n2.0, index n2.0 - i) ∈ nodes b ∧ (strand n2.0, index n2.0 - i) ∈ strand05.Domain
lemma bundle_casual2_closed2:
[|b ∈ bundles; n2.0 ∈ nodes b; i ≤ index n2.0|] ==> (strand n2.0, i) ∈ nodes b ∧ (strand n2.0, i) ∈ strand05.Domain
lemma bundle_casual2_closed222:
[|b ∈ bundles; n2.0 ∈ nodes b; i ≤ index n2.0|] ==> ((strand n2.0, index n2.0 - i), n2.0) ∈ (edges b ∩ casual2)*
lemma bundle_casual2_closed3:
[|b ∈ bundles; n2.0 ∈ nodes b; i ≤ index n2.0|] ==> (strand n2.0, i) ∈ nodes b ∧ ((strand n2.0, i), n2.0) ∈ (edges b ∩ casual2)*
lemma bundle_casual2_trancl:
[|b ∈ bundles; n' ∈ nodes b; m =>+ n'|] ==> m ∈ nodes b ∧ m \<prec>\<^sub>b n'
lemma bundle_edge_is_unique:
b ∈ bundles ==> node_sign n2.0 = - ∧ n2.0 ∈ nodes b --> (∃!n1. n1 -> n2.0 ∧ n1 ∈ nodes b ∧ (n1, n2.0) ∈ edges b)
lemma bundle_is_trans_closed2:
b ∈ bundles ==> y \<prec>\<^sub>b z --> y ∈ nodes b ∧ z ∈ nodes b
lemma bundle_is_trans_closed22:
[|b ∈ bundles; y \<prec>\<^sub>b z|] ==> y ∈ nodes b ∧ z ∈ nodes b
lemma bundle_is_rtrans_closed23:
[|b ∈ bundles; z ∈ nodes b; y \<preceq>\<^sub>b z|] ==> y ∈ nodes b ∧ z ∈ nodes b
lemma bundle_is_rtrans_closed24:
[|b ∈ bundles; z ∉ nodes b; y \<preceq>\<^sub>b z|] ==> y = z
lemma bundle_edge_is_anti2:
b ∈ bundles ==> y \<prec>\<^sub>b z --> y ≠ z
lemma bundle_is_wf:
b ∈ bundles ==> wf (edges b)
lemma bundle_trancl_wf:
b ∈ bundles ==> wf ((edges b)+)
lemma bundle_minimal:
[|b ∈ bundles; x ∈ Q|] ==> ∃z∈Q. ∀y. y \<prec>\<^sub>b z --> y ∉ Q
lemma non_originate_node:
[|unique_originate a n; n ≠ n'|] ==> strand n = strand n' ∧ (a \<sqsubset> node_term n' --> index n < index n') ∨ strand n ≠ strand n' ∧ non_originate a (strand n')
lemma arith_aux_non_originate1:
∃i. P i ==> ∃i. P i ∧ (∀j<i. ¬ P j)
lemma non_originate_strand:
non_originate a s ==> (∃i. node_sign (s, i) = - ∧ a \<sqsubset> node_term (s, i) ∧ (∀i0<i. ¬ a \<sqsubset> node_term (s, i0))) ∨ (∀i. ¬ a \<sqsubset> node_term (s, i))
lemma non_originate_node1:
[|unique_originate a n; strand n ≠ strand n'; a \<sqsubset> node_term n'|] ==> ∃m. first_node_in_nonorigi_strand a n' m
lemma nonOrigNode:
[|node_sign m2.0 = +; non_originate a (strand m2.0); first_node_in_nonorigi_strand a m2.0 m; a \<sqsubset> node_term m2.0|] ==> index m < index m2.0
lemma nonOrigNode1:
[|strand m1.0 = strand m2.0; first_node_in_nonorigi_strand a m2.0 m; a \<sqsubset> node_term m1.0|] ==> index m ≤ index m1.0
lemma synth_sep_suite:
[|suite M a b; t ∈ synth M; t = \<lbrace>g, h\<rbrace>|] ==> g ∈ synth M ∧ h ∈ synth M
lemma c:
[|Is_penetrator_strand (strand m'); node_sign m' = +; suite M a b; ∃m. m =>+ m'; b ∈ bundles; ∀m. m =>+ m' ∧ node_sign m = - --> node_term m ∈ synth M; m' ∈ nodes b|] ==> node_term m' ∈ synth M
lemma c:
[|node_term m' ∉ synth M; node_sign m' = +; suite M a b; ∃m. m =>+ m'; b ∈ bundles; ∀m. m =>+ m' ∧ node_sign m = - --> node_term m ∈ synth M; m' ∈ nodes b|] ==> Is_regular_strand (strand m')
lemma slice_arr_not_nil:
∃n ns. slice_arr_cons s i len = n # ns ∧ n = (s, i)
lemma slic_arr_ith:
∀start i. i ≤ len --> slice_arr_cons s start len ! i = (s, start + i)
lemma slice_arr_ith:
i ≤ len ==> slice_arr_cons s start len ! i = (s, start + i)
lemma slice_arr_length:
∀start. length (slice_arr_cons s start len) = len + 1
lemma slice_arr_length1:
length (slice_arr_cons s start len) = len + 1
lemma slice_arr_last:
last (slice_arr_cons s start len) = (s, start + len)
lemma slice_arr_last2:
[|index m ≤ index m2.0; strand m = strand m2.0|] ==> last (slice_arr_cons (strand m) (index m) (index m2.0 - index m)) = m2.0
lemma slice_arr_no_empty:
slice_arr_cons s i len ≠ []
lemma slice_index:
m ∈ set (slice_arr_cons (strand m) start i) ==> index m ≤ start + i
lemma slice_arr_close:
[|m' =>+ ma; ma ∈ set (slice_arr_cons (strand m) (index m) (index m2.0 - index m)); index m < index m2.0; a \<sqsubset> node_term m'; first_node_in_nonorigi_strand a m2.0 m|] ==> m' ∈ set (slice_arr_cons (strand m) (index m) (index m2.0 - index m))
lemma slice_arr_close1:
[|m' =>+ ma; ma ∈ set (slice_arr_cons (strand m) (index m) (index m2.0 - index m)); index m ≤ index m2.0; a \<sqsubset> node_term m'; unique_originate a m|] ==> m' ∈ set (slice_arr_cons (strand m) (index m) (index m2.0 - index m))
lemma Complete_path_mono:
p ∈ complete_Path a g1.0 ==> nodes g1.0 ⊆ nodes g2.0 --> edges g1.0 ⊆ edges g2.0 --> p ∈ complete_Path a g2.0
lemma completePathProp(1):
p ∈ complete_Path a b ==> unique_originate a n --> b ∈ bundles --> (∀n'. n' ∈ nodes b ∧ last p = n' --> n' = n ∨ n \<prec>\<^sub>b n')
and completePathProp(2):
p ∈ complete_Path a b ==> p ≠ []
and completePathProp(3):
p ∈ complete_Path a b ==> ∀m. m ∈ set p ∧ node_sign m = + ∧ a \<sqsubset> node_term m --> (∀m'. m' =>+ m ∧ a \<sqsubset> node_term m' --> m' ∈ set p)
lemma C:
b ∈ bundles ==> unique_originate a n --> (∀n'. n' ∈ nodes b ∧ a \<sqsubset> node_term n' --> (∃p. p ∈ complete_Path a b ∧ last p = n'))
lemma general_auth_test:
[|b ∈ bundles; unique_originate a n; suite M a b; n' ∈ nodes b; a \<sqsubset> node_term n'; node_term n' ∉ synth M; node_term n ∈ synth M|] ==> ∃m' m. n \<preceq>\<^sub>b m ∧ m' \<preceq>\<^sub>b n' ∧ m =>+ m' ∧ node_sign m' = + ∧ (strand m' ≠ strand n --> node_sign m = -) ∧ a \<sqsubset> node_term m ∧ a \<sqsubset> node_term m' ∧ node_term m' ∉ synth M ∧ node_term m ∈ synth M ∧ Is_regular_strand (strand m') ∧ (∀y. y \<prec>\<^sub>b m' --> ¬ a \<sqsubset> node_term y ∨ node_term y ∈ synth M)
lemma general_auth_test1:
[|b ∈ bundles; unique_originate a n; suite M a b; n' ∈ nodes b; a \<sqsubset> node_term n'; node_term n' ∉ synth M; node_term n ∈ synth M; ∀n0. strand n0 = strand n ∧ node_sign n0 = + ∧ a \<sqsubset> node_term n0 ∧ n0 ∈ nodes b --> node_term n0 ∈ synth M|] ==> ∃m' m. n \<preceq>\<^sub>b m ∧ m' \<preceq>\<^sub>b n' ∧ m =>+ m' ∧ node_sign m' = + ∧ (strand m' ≠ strand n ∧ node_sign m = -) ∧ a \<sqsubset> node_term m ∧ a \<sqsubset> node_term m' ∧ node_term m' ∉ synth M ∧ node_term m ∈ synth M ∧ Is_regular_strand (strand m') ∧ (∀y. y \<prec>\<^sub>b m' --> ¬ a \<sqsubset> node_term y ∨ node_term y ∈ synth M)
lemma general_auth_test2:
[|b ∈ bundles; unique_originate a n; suite M a b; n' ∈ nodes b; a \<sqsubset> node_term n'; node_term n' ∉ synth M; node_term n ∈ synth M|] ==> ∃m' m. n \<preceq>\<^sub>b m ∧ m' \<preceq>\<^sub>b n' ∧ m =>+ m' ∧ node_sign m' = + ∧ (strand m' ≠ strand n --> node_sign m = -) ∧ a \<sqsubset> node_term m ∧ a \<sqsubset> node_term m' ∧ node_term m' ∉ synth M ∧ node_term m ∈ synth M ∧ Is_regular_strand (strand m') ∧ (∀y. y \<prec>\<^sub>b m' --> node_term y ∈ synth M)
lemma synthM:
[|m ∈ synth M; a = Nonce N ∨ a = Key k'; M = {Crypt k h} ∪ {m. ¬ a \<sqsubset> m}|] ==> a \<sqsubset> m --> Crypt k h \<sqsubset> m
lemma synthM1:
[|m ∈ synth M; a = Nonce N ∨ a = Key k'; M = {Crypt k h} ∪ {m. ¬ a \<sqsubset> m}; a \<sqsubset> m|] ==> Crypt k h \<sqsubset> m
lemma out_auth_tast:
[|b ∈ bundles; unique_originate a n; regularK (invKey k) b; n' ∈ nodes b; a \<sqsubset> node_term n'; ¬ Crypt k h \<sqsubset> node_term n'; node_sign n' = -; node_term n ∈ synth M; a = Nonce N ∨ a = Key k'; M = {Crypt k h} ∪ {m. ¬ a \<sqsubset> m}; a \<sqsubset> h; n => n'; node_sign n' = -|] ==> ∃m' m. n \<preceq>\<^sub>b m ∧ m' \<preceq>\<^sub>b n' ∧ m =>+ m' ∧ node_sign m' = + ∧ node_sign m = - ∧ a \<sqsubset> node_term m ∧ a \<sqsubset> node_term m' ∧ node_term m' ∉ synth M ∧ node_term m ∈ synth M ∧ Is_regular_strand (strand m') ∧ (∀y. y \<prec>\<^sub>b m' --> node_term y ∈ synth M)
lemma c:
[|b ∈ bundles; n ∈ nodes b; Crypt k h \<sqsubset> node_term n; regularK k b|] ==> ∃n'. n' \<preceq>\<^sub>b n ∧ n' ∈ nodes b ∧ Crypt k h \<sqsubset> node_term n' ∧ Is_regular_strand (strand n') ∧ node_sign n' = + ∧ (∀y. y \<prec>\<^sub>b n' --> ¬ Crypt k h \<sqsubset> node_term y)
lemma in_auth_test:
[|b ∈ bundles; unique_originate a n; suite M a b; n' ∈ nodes b; a \<sqsubset> h; Crypt k h \<sqsubset> node_term n'; ¬ Crypt k h \<sqsubset> node_term n; regularK k b|] ==> ∃m' m. n \<preceq>\<^sub>b m ∧ m' \<preceq>\<^sub>b n' ∧ m =>+ m' ∧ node_sign m' = + ∧ (strand m' ≠ strand n --> node_sign m = -) ∧ ¬ Crypt k h \<sqsubset> node_term m ∧ Crypt k h \<sqsubset> node_term m' ∧ ¬ Crypt k h \<sqsubset> node_term m ∧ Crypt k h \<sqsubset> node_term m' ∧ Is_regular_strand (strand m')